diff --git a/src/Doc/Isar_Ref/Generic.thy b/src/Doc/Isar_Ref/Generic.thy --- a/src/Doc/Isar_Ref/Generic.thy +++ b/src/Doc/Isar_Ref/Generic.thy @@ -1,1826 +1,1826 @@ (*:maxLineLen=78:*) theory Generic imports Main Base begin chapter \Generic tools and packages \label{ch:gen-tools}\ section \Configuration options \label{sec:config}\ text \ Isabelle/Pure maintains a record of named configuration options within the theory or proof context, with values of type \<^ML_type>\bool\, \<^ML_type>\int\, \<^ML_type>\real\, or \<^ML_type>\string\. Tools may declare options in ML, and then refer to these values (relative to the context). Thus global reference variables are easily avoided. The user may change the value of a configuration option by means of an associated attribute of the same name. This form of context declaration works particularly well with commands such as @{command "declare"} or @{command "using"} like this: \ (*<*)experiment begin(*>*) declare [[show_main_goal = false]] notepad begin note [[show_main_goal = true]] end (*<*)end(*>*) text \ \begin{matharray}{rcll} @{command_def "print_options"} & : & \context \\ \\ \end{matharray} \<^rail>\ @@{command print_options} ('!'?) ; @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? \ \<^descr> @{command "print_options"} prints the available configuration options, with names, types, and current values; the ``\!\'' option indicates extra verbosity. \<^descr> \name = value\ as an attribute expression modifies the named option, with the syntax of the value depending on the option's type. For \<^ML_type>\bool\ the default value is \true\. Any attempt to change a global option in a local context is ignored. \ section \Basic proof tools\ subsection \Miscellaneous methods and attributes \label{sec:misc-meth-att}\ text \ \begin{matharray}{rcl} @{method_def unfold} & : & \method\ \\ @{method_def fold} & : & \method\ \\ @{method_def insert} & : & \method\ \\[0.5ex] @{method_def erule}\\<^sup>*\ & : & \method\ \\ @{method_def drule}\\<^sup>*\ & : & \method\ \\ @{method_def frule}\\<^sup>*\ & : & \method\ \\ @{method_def intro} & : & \method\ \\ @{method_def elim} & : & \method\ \\ @{method_def fail} & : & \method\ \\ @{method_def succeed} & : & \method\ \\ @{method_def sleep} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thms} ; (@@{method erule} | @@{method drule} | @@{method frule}) ('(' @{syntax nat} ')')? @{syntax thms} ; (@@{method intro} | @@{method elim}) @{syntax thms}? ; @@{method sleep} @{syntax real} \ \<^descr> @{method unfold}~\a\<^sub>1 \ a\<^sub>n\ and @{method fold}~\a\<^sub>1 \ a\<^sub>n\ expand (or fold back) the given definitions throughout all goals; any chained facts provided are inserted into the goal and subject to rewriting as well. Unfolding works in two stages: first, the given equations are used directly for rewriting; second, the equations are passed through the attribute @{attribute_ref abs_def} before rewriting --- to ensure that definitions are fully expanded, regardless of the actual parameters that are provided. \<^descr> @{method insert}~\a\<^sub>1 \ a\<^sub>n\ inserts theorems as facts into all goals of the proof state. Note that current facts indicated for forward chaining are ignored. \<^descr> @{method erule}~\a\<^sub>1 \ a\<^sub>n\, @{method drule}~\a\<^sub>1 \ a\<^sub>n\, and @{method frule}~\a\<^sub>1 \ a\<^sub>n\ are similar to the basic @{method rule} method (see \secref{sec:pure-meth-att}), but apply rules by elim-resolution, destruct-resolution, and forward-resolution, respectively \<^cite>\"isabelle-implementation"\. The optional natural number argument (default 0) specifies additional assumption steps to be performed here. Note that these methods are improper ones, mainly serving for experimentation and tactic script emulation. Different modes of basic rule application are usually expressed in Isar at the proof language level, rather than via implicit proof state manipulations. For example, a proper single-step elimination would be done using the plain @{method rule} method, with forward chaining of current facts. \<^descr> @{method intro} and @{method elim} repeatedly refine some goal by intro- or elim-resolution, after having inserted any chained facts. Exactly the rules given as arguments are taken into account; this allows fine-tuned decomposition of a proof problem, in contrast to common automated tools. \<^descr> @{method fail} yields an empty result sequence; it is the identity of the ``\|\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method succeed} yields a single (unchanged) result; it is the identity of the ``\,\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method sleep}~\s\ succeeds after a real-time delay of \s\ seconds. This is occasionally useful for demonstration and testing purposes. \begin{matharray}{rcl} @{attribute_def tagged} & : & \attribute\ \\ @{attribute_def untagged} & : & \attribute\ \\[0.5ex] @{attribute_def THEN} & : & \attribute\ \\ @{attribute_def unfolded} & : & \attribute\ \\ @{attribute_def folded} & : & \attribute\ \\ @{attribute_def abs_def} & : & \attribute\ \\[0.5ex] @{attribute_def rotated} & : & \attribute\ \\ @{attribute_def (Pure) elim_format} & : & \attribute\ \\ @{attribute_def no_vars}\\<^sup>*\ & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute tagged} @{syntax name} @{syntax name} ; @@{attribute untagged} @{syntax name} ; @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thm} ; (@@{attribute unfolded} | @@{attribute folded}) @{syntax thms} ; @@{attribute rotated} @{syntax int}? \ \<^descr> @{attribute tagged}~\name value\ and @{attribute untagged}~\name\ add and remove \<^emph>\tags\ of some theorem. Tags may be any list of string pairs that serve as formal comment. The first string is considered the tag name, the second its value. Note that @{attribute untagged} removes any tags of the same name. \<^descr> @{attribute THEN}~\a\ composes rules by resolution; it resolves with the first premise of \a\ (an alternative position may be also specified). See also \<^ML_infix>\RS\ in \<^cite>\"isabelle-implementation"\. \<^descr> @{attribute unfolded}~\a\<^sub>1 \ a\<^sub>n\ and @{attribute folded}~\a\<^sub>1 \ a\<^sub>n\ expand and fold back again the given definitions throughout a rule. \<^descr> @{attribute abs_def} turns an equation of the form \<^prop>\f x y \ t\ into \<^prop>\f \ \x y. t\, which ensures that @{method simp} steps always expand it. This also works for object-logic equality. \<^descr> @{attribute rotated}~\n\ rotate the premises of a theorem by \n\ (default 1). \<^descr> @{attribute (Pure) elim_format} turns a destruction rule into elimination rule format, by resolving with the rule \<^prop>\PROP A \ (PROP A \ PROP B) \ PROP B\. Note that the Classical Reasoner (\secref{sec:classical}) provides its own version of this operation. \<^descr> @{attribute no_vars} replaces schematic variables by free ones; this is mainly for tuning output of pretty printed theorems. \ subsection \Low-level equational reasoning\ text \ \begin{matharray}{rcl} @{method_def subst} & : & \method\ \\ @{method_def hypsubst} & : & \method\ \\ @{method_def split} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method subst} ('(' 'asm' ')')? \ ('(' (@{syntax nat}+) ')')? @{syntax thm} ; @@{method split} @{syntax thms} \ These methods provide low-level facilities for equational reasoning that are intended for specialized applications only. Normally, single step calculations would be performed in a structured text (see also \secref{sec:calculation}), while the Simplifier methods provide the canonical way for automated normalization (see \secref{sec:simplifier}). \<^descr> @{method subst}~\eq\ performs a single substitution step using rule \eq\, which may be either a meta or object equality. \<^descr> @{method subst}~\(asm) eq\ substitutes in an assumption. \<^descr> @{method subst}~\(i \ j) eq\ performs several substitutions in the conclusion. The numbers \i\ to \j\ indicate the positions to substitute at. Positions are ordered from the top of the term tree moving down from left to right. For example, in \(a + b) + (c + d)\ there are three positions where commutativity of \+\ is applicable: 1 refers to \a + b\, 2 to the whole term, and 3 to \c + d\. If the positions in the list \(i \ j)\ are non-overlapping (e.g.\ \(2 3)\ in \(a + b) + (c + d)\) you may assume all substitutions are performed simultaneously. Otherwise the behaviour of \subst\ is not specified. \<^descr> @{method subst}~\(asm) (i \ j) eq\ performs the substitutions in the assumptions. The positions refer to the assumptions in order from left to right. For example, given in a goal of the form \P (a + b) \ P (c + d) \ \\, position 1 of commutativity of \+\ is the subterm \a + b\ and position 2 is the subterm \c + d\. \<^descr> @{method hypsubst} performs substitution using some assumption; this only works for equations of the form \x = t\ where \x\ is a free or bound variable. \<^descr> @{method split}~\a\<^sub>1 \ a\<^sub>n\ performs single-step case splitting using the given rules. Splitting is performed in the conclusion or some assumption of the subgoal, depending of the structure of the rule. Note that the @{method simp} method already involves repeated application of split rules as declared in the current context, using @{attribute split}, for example. \ section \The Simplifier \label{sec:simplifier}\ text \ The Simplifier performs conditional and unconditional rewriting and uses contextual information: rule declarations in the background theory or local proof context are taken into account, as well as chained facts and subgoal premises (``local assumptions''). There are several general hooks that allow to modify the simplification strategy, or incorporate other proof tools that solve sub-problems, produce rewrite rules on demand etc. The rewriting strategy is always strictly bottom up, except for congruence rules, which are applied while descending into a term. Conditions in conditional rewrite rules are solved recursively before the rewrite rule is applied. The default Simplifier setup of major object logics (HOL, HOLCF, FOL, ZF) makes the Simplifier ready for immediate use, without engaging into the internal structures. Thus it serves as general-purpose proof tool with the main focus on equational reasoning, and a bit more than that. \ subsection \Simplification methods \label{sec:simp-meth}\ text \ \begin{tabular}{rcll} @{method_def simp} & : & \method\ \\ @{method_def simp_all} & : & \method\ \\ \Pure.\@{method_def (Pure) simp} & : & \method\ \\ \Pure.\@{method_def (Pure) simp_all} & : & \method\ \\ @{attribute_def simp_depth_limit} & : & \attribute\ & default \100\ \\ \end{tabular} \<^medskip> \<^rail>\ (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' ; @{syntax_def simpmod}: ('add' | 'del' | 'flip' | 'only' | 'split' (() | '!' | 'del') | 'cong' (() | 'add' | 'del')) ':' @{syntax thms} \ \<^descr> @{method simp} invokes the Simplifier on the first subgoal, after inserting chained facts as additional goal premises; further rule declarations may be included via \(simp add: facts)\. The proof method fails if the subgoal remains unchanged after simplification. Note that the original goal premises and chained facts are subject to simplification themselves, while declarations via \add\/\del\ merely follow the policies of the object-logic to extract rewrite rules from theorems, without further simplification. This may lead to slightly different behavior in either case, which might be required precisely like that in some boundary situations to perform the intended simplification step! \<^medskip> Modifier \flip\ deletes the following theorems from the simpset and adds their symmetric version (i.e.\ lhs and rhs exchanged). No warning is shown if the original theorem was not present. \<^medskip> The \only\ modifier first removes all other rewrite rules, looper tactics (including split rules), congruence rules, and then behaves like \add\. Implicit solvers remain, which means that trivial rules like reflexivity or introduction of \True\ are available to solve the simplified subgoals, but also non-trivial tools like linear arithmetic in HOL. The latter may lead to some surprise of the meaning of ``only'' in Isabelle/HOL compared to English! \<^medskip> The \split\ modifiers add or delete rules for the Splitter (see also \secref{sec:simp-strategies} on the looper). This works only if the Simplifier method has been properly setup to include the Splitter (all major object logics such HOL, HOLCF, FOL, ZF do this already). The \!\ option causes the split rules to be used aggressively: after each application of a split rule in the conclusion, the \safe\ tactic of the classical reasoner (see \secref{sec:classical:partial}) is applied to the new goal. The net effect is that the goal is split into the different cases. This option can speed up simplification of goals with many nested conditional or case expressions significantly. There is also a separate @{method_ref split} method available for single-step case splitting. The effect of repeatedly applying \(split thms)\ can be imitated by ``\(simp only: split: thms)\''. \<^medskip> The \cong\ modifiers add or delete Simplifier congruence rules (see also \secref{sec:simp-rules}); the default is to add. \<^descr> @{method simp_all} is similar to @{method simp}, but acts on all goals, working backwards from the last to the first one as usual in Isabelle.\<^footnote>\The order is irrelevant for goals without schematic variables, so simplification might actually be performed in parallel here.\ Chained facts are inserted into all subgoals, before the simplification process starts. Further rule declarations are the same as for @{method simp}. The proof method fails if all subgoals remain unchanged after simplification. \<^descr> @{attribute simp_depth_limit} limits the number of recursive invocations of the Simplifier during conditional rewriting. By default the Simplifier methods above take local assumptions fully into account, using equational assumptions in the subsequent normalization process, or simplifying assumptions themselves. Further options allow to fine-tune the behavior of the Simplifier in this respect, corresponding to a variety of ML tactics as follows.\<^footnote>\Unlike the corresponding Isar proof methods, the ML tactics do not insist in changing the goal state.\ \begin{center} \small \begin{tabular}{|l|l|p{0.3\textwidth}|} \hline Isar method & ML tactic & behavior \\\hline \(simp (no_asm))\ & \<^ML>\simp_tac\ & assumptions are ignored completely \\\hline \(simp (no_asm_simp))\ & \<^ML>\asm_simp_tac\ & assumptions are used in the simplification of the conclusion but are not themselves simplified \\\hline \(simp (no_asm_use))\ & \<^ML>\full_simp_tac\ & assumptions are simplified but are not used in the simplification of each other or the conclusion \\\hline \(simp)\ & \<^ML>\asm_full_simp_tac\ & assumptions are used in the simplification of the conclusion and to simplify other assumptions \\\hline \(simp (asm_lr))\ & \<^ML>\asm_lr_simp_tac\ & compatibility mode: an assumption is only used for simplifying assumptions which are to the right of it \\\hline \end{tabular} \end{center} \<^medskip> In Isabelle/Pure, proof methods @{method (Pure) simp} and @{method (Pure) simp_all} only know about meta-equality \\\. Any new object-logic needs to re-define these methods via \<^ML>\Simplifier.method_setup\ in ML: Isabelle/FOL or Isabelle/HOL may serve as blue-prints. \ subsubsection \Examples\ text \ We consider basic algebraic simplifications in Isabelle/HOL. The rather trivial goal \<^prop>\0 + (x + 0) = x + 0 + 0\ looks like a good candidate to be solved by a single call of @{method simp}: \ lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops text \ The above attempt \<^emph>\fails\, because \<^term>\0\ and \<^term>\(+)\ in the HOL library are declared as generic type class operations, without stating any algebraic laws yet. More specific types are required to get access to certain standard simplifications of the theory context, e.g.\ like this:\ lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp text \ \<^medskip> In many cases, assumptions of a subgoal are also needed in the simplification process. For example: \ lemma fixes x :: nat shows "x = 0 \ x + x = 0" by simp lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp text \ As seen above, local assumptions that shall contribute to simplification need to be part of the subgoal already, or indicated explicitly for use by the subsequent method invocation. Both too little or too much information can make simplification fail, for different reasons. In the next example the malicious assumption \<^prop>\\x::nat. f x = g (f (g x))\ does not contribute to solve the problem, but makes the default @{method simp} method loop: the rewrite rule \f ?x \ g (f (g ?x))\ extracted from the assumption does not terminate. The Simplifier notices certain simple forms of nontermination, but not this one. The problem can be solved nonetheless, by ignoring assumptions via special options as explained before: \ lemma "(\x::nat. f x = g (f (g x))) \ f 0 = f 0 + 0" by (simp (no_asm)) text \ The latter form is typical for long unstructured proof scripts, where the control over the goal content is limited. In structured proofs it is usually better to avoid pushing too many facts into the goal state in the first place. Assumptions in the Isar proof context do not intrude the reasoning if not used explicitly. This is illustrated for a toplevel statement and a local proof body as follows: \ lemma assumes "\x::nat. f x = g (f (g x))" shows "f 0 = f 0 + 0" by simp notepad begin assume "\x::nat. f x = g (f (g x))" have "f 0 = f 0 + 0" by simp end text \ \<^medskip> Because assumptions may simplify each other, there can be very subtle cases of nontermination. For example, the regular @{method simp} method applied to \<^prop>\P (f x) \ y = x \ f x = f y \ Q\ gives rise to the infinite reduction sequence \[ \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \P (f y)\ \stackrel{\y \ x\}{\longmapsto} \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \cdots \] whereas applying the same to \<^prop>\y = x \ f x = f y \ P (f x) \ Q\ terminates (without solving the goal): \ lemma "y = x \ f x = f y \ P (f x) \ Q" apply simp oops text \ See also \secref{sec:simp-trace} for options to enable Simplifier trace mode, which often helps to diagnose problems with rewrite systems. \ subsection \Declaring rules \label{sec:simp-rules}\ text \ \begin{matharray}{rcl} @{attribute_def simp} & : & \attribute\ \\ @{attribute_def split} & : & \attribute\ \\ @{attribute_def cong} & : & \attribute\ \\ @{command_def "print_simpset"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} \<^rail>\ (@@{attribute simp} | @@{attribute cong}) (() | 'add' | 'del') | @@{attribute split} (() | '!' | 'del') ; @@{command print_simpset} ('!'?) \ \<^descr> @{attribute simp} declares rewrite rules, by adding or deleting them from the simpset within the theory or proof context. Rewrite rules are theorems expressing some form of equality, for example: \Suc ?m + ?n = ?m + Suc ?n\ \\ \?P \ ?P \ ?P\ \\ \?A \ ?B \ {x. x \ ?A \ x \ ?B}\ \<^medskip> Conditional rewrites such as \?m < ?n \ ?m div ?n = 0\ are also permitted; the conditions can be arbitrary formulas. \<^medskip> Internally, all rewrite rules are translated into Pure equalities, theorems with conclusion \lhs \ rhs\. The simpset contains a function for extracting equalities from arbitrary theorems, which is usually installed when the object-logic is configured initially. For example, \\ ?x \ {}\ could be turned into \?x \ {} \ False\. Theorems that are declared as @{attribute simp} and local assumptions within a goal are treated uniformly in this respect. The Simplifier accepts the following formats for the \lhs\ term: \<^enum> First-order patterns, considering the sublanguage of application of constant operators to variable operands, without \\\-abstractions or functional variables. For example: \(?x + ?y) + ?z \ ?x + (?y + ?z)\ \\ \f (f ?x ?y) ?z \ f ?x (f ?y ?z)\ \<^enum> Higher-order patterns in the sense of \<^cite>\"nipkow-patterns"\. These are terms in \\\-normal form (this will always be the case unless you have done something strange) where each occurrence of an unknown is of the form \?F x\<^sub>1 \ x\<^sub>n\, where the \x\<^sub>i\ are distinct bound variables. For example, \(\x. ?P x \ ?Q x) \ (\x. ?P x) \ (\x. ?Q x)\ or its symmetric form, since the \rhs\ is also a higher-order pattern. \<^enum> Physical first-order patterns over raw \\\-term structure without \\\\\-equality; abstractions and bound variables are treated like quasi-constant term material. For example, the rule \?f ?x \ range ?f = True\ rewrites the term \g a \ range g\ to \True\, but will fail to match \g (h b) \ range (\x. g (h x))\. However, offending subterms (in our case \?f ?x\, which is not a pattern) can be replaced by adding new variables and conditions like this: \?y = ?f ?x \ ?y \ range ?f = True\ is acceptable as a conditional rewrite rule of the second category since conditions can be arbitrary terms. \<^descr> @{attribute split} declares case split rules. \<^descr> @{attribute cong} declares congruence rules to the Simplifier context. Congruence rules are equalities of the form @{text [display] "\ \ f ?x\<^sub>1 \ ?x\<^sub>n = f ?y\<^sub>1 \ ?y\<^sub>n"} This controls the simplification of the arguments of \f\. For example, some arguments can be simplified under additional assumptions: @{text [display] "?P\<^sub>1 \ ?Q\<^sub>1 \ (?Q\<^sub>1 \ ?P\<^sub>2 \ ?Q\<^sub>2) \ (?P\<^sub>1 \ ?P\<^sub>2) \ (?Q\<^sub>1 \ ?Q\<^sub>2)"} Given this rule, the Simplifier assumes \?Q\<^sub>1\ and extracts rewrite rules from it when simplifying \?P\<^sub>2\. Such local assumptions are effective for rewriting formulae such as \x = 0 \ y + x = y\. %FIXME %The local assumptions are also provided as theorems to the solver; %see \secref{sec:simp-solver} below. \<^medskip> The following congruence rule for bounded quantifiers also supplies contextual information --- about the bound variable: @{text [display] "(?A = ?B) \ (\x. x \ ?B \ ?P x \ ?Q x) \ (\x \ ?A. ?P x) \ (\x \ ?B. ?Q x)"} \<^medskip> This congruence rule for conditional expressions can supply contextual information for simplifying the arms: @{text [display] "?p = ?q \ (?q \ ?a = ?c) \ (\ ?q \ ?b = ?d) \ (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} A congruence rule can also \<^emph>\prevent\ simplification of some arguments. Here is an alternative congruence rule for conditional expressions that conforms to non-strict functional evaluation: @{text [display] "?p = ?q \ (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} Only the first argument is simplified; the others remain unchanged. This can make simplification much faster, but may require an extra case split over the condition \?q\ to prove the goal. \<^descr> @{command "print_simpset"} prints the collection of rules declared to the Simplifier, which is also known as ``simpset'' internally; the ``\!\'' option indicates extra verbosity. The implicit simpset of the theory context is propagated monotonically through the theory hierarchy: forming a new theory, the union of the simpsets of its imports are taken as starting point. Also note that definitional packages like @{command "datatype"}, @{command "primrec"}, @{command "fun"} routinely declare Simplifier rules to the target context, while plain @{command "definition"} is an exception in \<^emph>\not\ declaring anything. \<^medskip> It is up the user to manipulate the current simpset further by explicitly adding or deleting theorems as simplification rules, or installing other tools via simplification procedures (\secref{sec:simproc}). Good simpsets are hard to design. Rules that obviously simplify, like \?n + 0 \ ?n\ are good candidates for the implicit simpset, unless a special non-normalizing behavior of certain operations is intended. More specific rules (such as distributive laws, which duplicate subterms) should be added only for specific proof steps. Conversely, sometimes a rule needs to be deleted just for some part of a proof. The need of frequent additions or deletions may indicate a poorly designed simpset. \begin{warn} The union of simpsets from theory imports (as described above) is not always a good starting point for the new theory. If some ancestors have deleted simplification rules because they are no longer wanted, while others have left those rules in, then the union will contain the unwanted rules, and thus have to be deleted again in the theory body. \end{warn} \ subsection \Ordered rewriting with permutative rules\ text \ A rewrite rule is \<^emph>\permutative\ if the left-hand side and right-hand side are the equal up to renaming of variables. The most common permutative rule is commutativity: \?x + ?y = ?y + ?x\. Other examples include \(?x - ?y) - ?z = (?x - ?z) - ?y\ in arithmetic and \insert ?x (insert ?y ?A) = insert ?y (insert ?x ?A)\ for sets. Such rules are common enough to merit special attention. Because ordinary rewriting loops given such rules, the Simplifier employs a special strategy, called \<^emph>\ordered rewriting\. Permutative rules are detected and only applied if the rewriting step decreases the redex wrt.\ a given term ordering. For example, commutativity rewrites \b + a\ to \a + b\, but then stops, because the redex cannot be decreased further in the sense of the term ordering. The default is lexicographic ordering of term structure, but this could be also changed locally for special applications via @{define_ML Simplifier.set_term_ord} in Isabelle/ML. \<^medskip> Permutative rewrite rules are declared to the Simplifier just like other rewrite rules. Their special status is recognized automatically, and their application is guarded by the term ordering accordingly. \ subsubsection \Rewriting with AC operators\ text \ Ordered rewriting is particularly effective in the case of associative-commutative operators. (Associativity by itself is not permutative.) When dealing with an AC-operator \f\, keep the following points in mind: \<^item> The associative law must always be oriented from left to right, namely \f (f x y) z = f x (f y z)\. The opposite orientation, if used with commutativity, leads to looping in conjunction with the standard term order. \<^item> To complete your set of rewrite rules, you must add not just associativity (A) and commutativity (C) but also a derived rule \<^emph>\left-commutativity\ (LC): \f x (f y z) = f y (f x z)\. Ordered rewriting with the combination of A, C, and LC sorts a term lexicographically --- the rewriting engine imitates bubble-sort. \ experiment fixes f :: "'a \ 'a \ 'a" (infix "\" 60) assumes assoc: "(x \ y) \ z = x \ (y \ z)" assumes commute: "x \ y = y \ x" begin lemma left_commute: "x \ (y \ z) = y \ (x \ z)" proof - have "(x \ y) \ z = (y \ x) \ z" by (simp only: commute) then show ?thesis by (simp only: assoc) qed lemmas AC_rules = assoc commute left_commute text \ Thus the Simplifier is able to establish equalities with arbitrary permutations of subterms, by normalizing to a common standard form. For example: \ lemma "(b \ c) \ a = xxx" apply (simp only: AC_rules) txt \\<^subgoals>\ oops lemma "(b \ c) \ a = a \ (b \ c)" by (simp only: AC_rules) lemma "(b \ c) \ a = c \ (b \ a)" by (simp only: AC_rules) lemma "(b \ c) \ a = (c \ b) \ a" by (simp only: AC_rules) end text \ Martin and Nipkow \<^cite>\"martin-nipkow"\ discuss the theory and give many examples; other algebraic structures are amenable to ordered rewriting, such as Boolean rings. The Boyer-Moore theorem prover \<^cite>\bm88book\ also employs ordered rewriting. \ subsubsection \Re-orienting equalities\ text \Another application of ordered rewriting uses the derived rule @{thm [source] eq_commute}: @{thm [source = false] eq_commute} to reverse equations. This is occasionally useful to re-orient local assumptions according to the term ordering, when other built-in mechanisms of reorientation and mutual simplification fail to apply.\ subsection \Simplifier tracing and debugging \label{sec:simp-trace}\ text \ \begin{tabular}{rcll} @{attribute_def simp_trace} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_depth_limit} & : & \attribute\ & default \1\ \\ @{attribute_def simp_debug} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_new} & : & \attribute\ \\ @{attribute_def simp_break} & : & \attribute\ \\ \end{tabular} \<^medskip> \<^rail>\ @@{attribute simp_trace_new} ('interactive')? \ ('mode' '=' ('full' | 'normal'))? \ ('depth' '=' @{syntax nat})? ; @@{attribute simp_break} (@{syntax term}*) \ These attributes and configurations options control various aspects of Simplifier tracing and debugging. \<^descr> @{attribute simp_trace} makes the Simplifier output internal operations. This includes rewrite steps, but also bookkeeping like modifications of the simpset. \<^descr> @{attribute simp_trace_depth_limit} limits the effect of @{attribute simp_trace} to the given depth of recursive Simplifier invocations (when solving conditions of rewrite rules). \<^descr> @{attribute simp_debug} makes the Simplifier output some extra information about internal operations. This includes any attempted invocation of simplification procedures. \<^descr> @{attribute simp_trace_new} controls Simplifier tracing within Isabelle/PIDE applications, notably Isabelle/jEdit \<^cite>\"isabelle-jedit"\. This provides a hierarchical representation of the rewriting steps performed by the Simplifier. Users can configure the behaviour by specifying breakpoints, verbosity and enabling or disabling the interactive mode. In normal verbosity (the default), only rule applications matching a breakpoint will be shown to the user. In full verbosity, all rule applications will be logged. Interactive mode interrupts the normal flow of the Simplifier and defers the decision how to continue to the user via some GUI dialog. \<^descr> @{attribute simp_break} declares term or theorem breakpoints for @{attribute simp_trace_new} as described above. Term breakpoints are patterns which are checked for matches on the redex of a rule application. Theorem breakpoints trigger when the corresponding theorem is applied in a rewrite step. For example: \ (*<*)experiment begin(*>*) declare conjI [simp_break] declare [[simp_break "?x \ ?y"]] (*<*)end(*>*) subsection \Simplification procedures \label{sec:simproc}\ text \ Simplification procedures are ML functions that produce proven rewrite rules on demand. They are associated with higher-order patterns that approximate the left-hand sides of equations. The Simplifier first matches the current redex against one of the LHS patterns; if this succeeds, the corresponding ML function is invoked, passing the Simplifier context and redex term. Thus rules may be specifically fashioned for particular situations, resulting in a more powerful mechanism than term rewriting by a fixed set of rules. Any successful result needs to be a (possibly conditional) rewrite rule \t \ u\ that is applicable to the current redex. The rule will be applied just as any ordinary rewrite rule. It is expected to be already in \<^emph>\internal form\, bypassing the automatic preprocessing of object-level equivalences. \begin{matharray}{rcl} @{command_def "simproc_setup"} & : & \local_theory \ local_theory\ \\ simproc & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' @{syntax text}; @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) \ \<^descr> @{command "simproc_setup"} defines a named simplification procedure that is invoked by the Simplifier whenever any of the given term patterns match the current redex. The implementation, which is provided as ML source text, needs to be of type \<^ML_type>\morphism -> Proof.context -> cterm -> thm option\, where the \<^ML_type>\cterm\ represents the current redex \r\ and the result is supposed to be some proven rewrite rule \r \ r'\ (or a generalized version), or \<^ML>\NONE\ to indicate failure. The \<^ML_type>\Proof.context\ argument holds the full context of the current Simplifier invocation. The \<^ML_type>\morphism\ informs about the difference of the original compilation context wrt.\ the one of the actual application later on. Morphisms are only relevant for simprocs that are defined within a local target context, e.g.\ in a locale. \<^descr> \simproc add: name\ and \simproc del: name\ add or delete named simprocs to the current Simplifier context. The default is to add a simproc. Note that @{command "simproc_setup"} already adds the new simproc to the subsequent context. \ subsubsection \Example\ text \ The following simplification procedure for @{thm [source = false, show_types] unit_eq} in HOL performs fine-grained control over rule application, beyond higher-order pattern matching. Declaring @{thm unit_eq} as @{attribute simp} directly would make the Simplifier loop! Note that a version of this simplification procedure is already active in Isabelle/HOL. \ (*<*)experiment begin(*>*) simproc_setup unit ("x::unit") = - \fn _ => fn _ => fn ct => + \K (K (fn ct => if HOLogic.is_unit (Thm.term_of ct) then NONE - else SOME (mk_meta_eq @{thm unit_eq})\ + else SOME (mk_meta_eq @{thm unit_eq})))\ (*<*)end(*>*) text \ Since the Simplifier applies simplification procedures frequently, it is important to make the failure check in ML reasonably fast.\ subsection \Configurable Simplifier strategies \label{sec:simp-strategies}\ text \ The core term-rewriting engine of the Simplifier is normally used in combination with some add-on components that modify the strategy and allow to integrate other non-Simplifier proof tools. These may be reconfigured in ML as explained below. Even if the default strategies of object-logics like Isabelle/HOL are used unchanged, it helps to understand how the standard Simplifier strategies work.\ subsubsection \The subgoaler\ text \ \begin{mldecls} @{define_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) -> Proof.context -> Proof.context"} \\ @{define_ML Simplifier.prems_of: "Proof.context -> thm list"} \\ \end{mldecls} The subgoaler is the tactic used to solve subgoals arising out of conditional rewrite rules or congruence rules. The default should be simplification itself. In rare situations, this strategy may need to be changed. For example, if the premise of a conditional rule is an instance of its conclusion, as in \Suc ?m < ?n \ ?m < ?n\, the default strategy could loop. % FIXME !?? \<^descr> \<^ML>\Simplifier.set_subgoaler\~\tac ctxt\ sets the subgoaler of the context to \tac\. The tactic will be applied to the context of the running Simplifier instance. \<^descr> \<^ML>\Simplifier.prems_of\~\ctxt\ retrieves the current set of premises from the context. This may be non-empty only if the Simplifier has been told to utilize local assumptions in the first place (cf.\ the options in \secref{sec:simp-meth}). As an example, consider the following alternative subgoaler: \ ML_val \ fun subgoaler_tac ctxt = assume_tac ctxt ORELSE' resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE' asm_simp_tac ctxt \ text \ This tactic first tries to solve the subgoal by assumption or by resolving with with one of the premises, calling simplification only if that fails.\ subsubsection \The solver\ text \ \begin{mldecls} @{define_ML_type solver} \\ @{define_ML Simplifier.mk_solver: "string -> (Proof.context -> int -> tactic) -> solver"} \\ @{define_ML_infix setSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix setSSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSSolver: "Proof.context * solver -> Proof.context"} \\ \end{mldecls} A solver is a tactic that attempts to solve a subgoal after simplification. Its core functionality is to prove trivial subgoals such as \<^prop>\True\ and \t = t\, but object-logics might be more ambitious. For example, Isabelle/HOL performs a restricted version of linear arithmetic here. Solvers are packaged up in abstract type \<^ML_type>\solver\, with \<^ML>\Simplifier.mk_solver\ as the only operation to create a solver. \<^medskip> Rewriting does not instantiate unknowns. For example, rewriting alone cannot prove \a \ ?A\ since this requires instantiating \?A\. The solver, however, is an arbitrary tactic and may instantiate unknowns as it pleases. This is the only way the Simplifier can handle a conditional rewrite rule whose condition contains extra variables. When a simplification tactic is to be combined with other provers, especially with the Classical Reasoner, it is important whether it can be considered safe or not. For this reason a simpset contains two solvers: safe and unsafe. The standard simplification strategy solely uses the unsafe solver, which is appropriate in most cases. For special applications where the simplification process is not allowed to instantiate unknowns within the goal, simplification starts with the safe solver, but may still apply the ordinary unsafe one in nested simplifications for conditional rules or congruences. Note that in this way the overall tactic is not totally safe: it may instantiate unknowns that appear also in other subgoals. \<^descr> \<^ML>\Simplifier.mk_solver\~\name tac\ turns \tac\ into a solver; the \name\ is only attached as a comment and has no further significance. \<^descr> \ctxt setSSolver solver\ installs \solver\ as the safe solver of \ctxt\. \<^descr> \ctxt addSSolver solver\ adds \solver\ as an additional safe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^descr> \ctxt setSolver solver\ installs \solver\ as the unsafe solver of \ctxt\. \<^descr> \ctxt addSolver solver\ adds \solver\ as an additional unsafe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^medskip> The solver tactic is invoked with the context of the running Simplifier. Further operations may be used to retrieve relevant information, such as the list of local Simplifier premises via \<^ML>\Simplifier.prems_of\ --- this list may be non-empty only if the Simplifier runs in a mode that utilizes local assumptions (see also \secref{sec:simp-meth}). The solver is also presented the full goal including its assumptions in any case. Thus it can use these (e.g.\ by calling \<^ML>\assume_tac\), even if the Simplifier proper happens to ignore local premises at the moment. \<^medskip> As explained before, the subgoaler is also used to solve the premises of congruence rules. These are usually of the form \s = ?x\, where \s\ needs to be simplified and \?x\ needs to be instantiated with the result. Typically, the subgoaler will invoke the Simplifier at some point, which will eventually call the solver. For this reason, solver tactics must be prepared to solve goals of the form \t = ?x\, usually by reflexivity. In particular, reflexivity should be tried before any of the fancy automated proof tools. It may even happen that due to simplification the subgoal is no longer an equality. For example, \False \ ?Q\ could be rewritten to \\ ?Q\. To cover this case, the solver could try resolving with the theorem \\ False\ of the object-logic. \<^medskip> \begin{warn} If a premise of a congruence rule cannot be proved, then the congruence is ignored. This should only happen if the rule is \<^emph>\conditional\ --- that is, contains premises not of the form \t = ?x\. Otherwise it indicates that some congruence rule, or possibly the subgoaler or solver, is faulty. \end{warn} \ subsubsection \The looper\ text \ \begin{mldecls} @{define_ML_infix setloop: "Proof.context * (Proof.context -> int -> tactic) -> Proof.context"} \\ @{define_ML_infix addloop: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delloop: "Proof.context * string -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split_bang: " thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\ \end{mldecls} The looper is a list of tactics that are applied after simplification, in case the solver failed to solve the simplified goal. If the looper succeeds, the simplification process is started all over again. Each of the subgoals generated by the looper is attacked in turn, in reverse order. A typical looper is \<^emph>\case splitting\: the expansion of a conditional. Another possibility is to apply an elimination rule on the assumptions. More adventurous loopers could start an induction. \<^descr> \ctxt setloop tac\ installs \tac\ as the only looper tactic of \ctxt\. \<^descr> \ctxt addloop (name, tac)\ adds \tac\ as an additional looper tactic with name \name\, which is significant for managing the collection of loopers. The tactic will be tried after the looper tactics that had already been present in \ctxt\. \<^descr> \ctxt delloop name\ deletes the looper tactic that was associated with \name\ from \ctxt\. \<^descr> \<^ML>\Splitter.add_split\~\thm ctxt\ adds split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.add_split_bang\~\thm ctxt\ adds aggressive (see \S\ref{sec:simp-meth}) split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.del_split\~\thm ctxt\ deletes the split tactic corresponding to \thm\ from the looper tactics of \ctxt\. The splitter replaces applications of a given function; the right-hand side of the replacement can be anything. For example, here is a splitting rule for conditional expressions: @{text [display] "?P (if ?Q ?x ?y) \ (?Q \ ?P ?x) \ (\ ?Q \ ?P ?y)"} Another example is the elimination operator for Cartesian products (which happens to be called \<^const>\case_prod\ in Isabelle/HOL: @{text [display] "?P (case_prod ?f ?p) \ (\a b. ?p = (a, b) \ ?P (f a b))"} For technical reasons, there is a distinction between case splitting in the conclusion and in the premises of a subgoal. The former is done by \<^ML>\Splitter.split_tac\ with rules like @{thm [source] if_split} or @{thm [source] option.split}, which do not split the subgoal, while the latter is done by \<^ML>\Splitter.split_asm_tac\ with rules like @{thm [source] if_split_asm} or @{thm [source] option.split_asm}, which split the subgoal. The function \<^ML>\Splitter.add_split\ automatically takes care of which tactic to call, analyzing the form of the rules given as argument; it is the same operation behind \split\ attribute or method modifier syntax in the Isar source language. Case splits should be allowed only when necessary; they are expensive and hard to control. Case-splitting on if-expressions in the conclusion is usually beneficial, so it is enabled by default in Isabelle/HOL and Isabelle/FOL/ZF. \begin{warn} With \<^ML>\Splitter.split_asm_tac\ as looper component, the Simplifier may split subgoals! This might cause unexpected problems in tactic expressions that silently assume 0 or 1 subgoals after simplification. \end{warn} \ subsection \Forward simplification \label{sec:simp-forward}\ text \ \begin{matharray}{rcl} @{attribute_def simplified} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute simplified} opt? @{syntax thms}? ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' \ \<^descr> @{attribute simplified}~\a\<^sub>1 \ a\<^sub>n\ causes a theorem to be simplified, either by exactly the specified rules \a\<^sub>1, \, a\<^sub>n\, or the implicit Simplifier context if no arguments are given. The result is fully simplified by default, including assumptions and conclusion; the options \no_asm\ etc.\ tune the Simplifier in the same way as the for the \simp\ method. Note that forward simplification restricts the Simplifier to its most basic operation of term rewriting; solver and looper tactics (\secref{sec:simp-strategies}) are \<^emph>\not\ involved here. The @{attribute simplified} attribute should be only rarely required under normal circumstances. \ section \The Classical Reasoner \label{sec:classical}\ subsection \Basic concepts\ text \Although Isabelle is generic, many users will be working in some extension of classical first-order logic. Isabelle/ZF is built upon theory FOL, while Isabelle/HOL conceptually contains first-order logic as a fragment. Theorem-proving in predicate logic is undecidable, but many automated strategies have been developed to assist in this task. Isabelle's classical reasoner is a generic package that accepts certain information about a logic and delivers a suite of automatic proof tools, based on rules that are classified and declared in the context. These proof procedures are slow and simplistic compared with high-end automated theorem provers, but they can save considerable time and effort in practice. They can prove theorems such as Pelletier's \<^cite>\pelletier86\ problems 40 and 41 in a few milliseconds (including full proof reconstruction):\ lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" by blast lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" by blast text \The proof tools are generic. They are not restricted to first-order logic, and have been heavily used in the development of the Isabelle/HOL library and applications. The tactics can be traced, and their components can be called directly; in this manner, any proof can be viewed interactively.\ subsubsection \The sequent calculus\ text \Isabelle supports natural deduction, which is easy to use for interactive proof. But natural deduction does not easily lend itself to automation, and has a bias towards intuitionism. For certain proofs in classical logic, it can not be called natural. The \<^emph>\sequent calculus\, a generalization of natural deduction, is easier to automate. A \<^bold>\sequent\ has the form \\ \ \\, where \\\ and \\\ are sets of formulae.\<^footnote>\For first-order logic, sequents can equivalently be made from lists or multisets of formulae.\ The sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ is \<^bold>\valid\ if \P\<^sub>1 \ \ \ P\<^sub>m\ implies \Q\<^sub>1 \ \ \ Q\<^sub>n\. Thus \P\<^sub>1, \, P\<^sub>m\ represent assumptions, each of which is true, while \Q\<^sub>1, \, Q\<^sub>n\ represent alternative goals. A sequent is \<^bold>\basic\ if its left and right sides have a common formula, as in \P, Q \ Q, R\; basic sequents are trivially valid. Sequent rules are classified as \<^bold>\right\ or \<^bold>\left\, indicating which side of the \\\ symbol they operate on. Rules that operate on the right side are analogous to natural deduction's introduction rules, and left rules are analogous to elimination rules. The sequent calculus analogue of \(\I)\ is the rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\P, \ \ \, Q\} \] Applying the rule backwards, this breaks down some implication on the right side of a sequent; \\\ and \\\ stand for the sets of formulae that are unaffected by the inference. The analogue of the pair \(\I1)\ and \(\I2)\ is the single rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\\ \ \, P, Q\} \] This breaks down some disjunction on the right side, replacing it by both disjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic. To illustrate the use of multiple formulae on the right, let us prove the classical theorem \(P \ Q) \ (Q \ P)\. Working backwards, we reduce this formula to a basic sequent: \[ \infer[\(\R)\]{\\ (P \ Q) \ (Q \ P)\} {\infer[\(\R)\]{\\ (P \ Q), (Q \ P)\} {\infer[\(\R)\]{\P \ Q, (Q \ P)\} {\P, Q \ Q, P\}}} \] This example is typical of the sequent calculus: start with the desired theorem and apply rules backwards in a fairly arbitrary manner. This yields a surprisingly effective proof procedure. Quantifiers add only few complications, since Isabelle handles parameters and schematic variables. See \<^cite>\\Chapter 10\ in "paulson-ml2"\ for further discussion.\ subsubsection \Simulating sequents by natural deduction\ text \Isabelle can represent sequents directly, as in the object-logic LK. But natural deduction is easier to work with, and most object-logics employ it. Fortunately, we can simulate the sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ by the Isabelle formula \P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1\ where the order of the assumptions and the choice of \Q\<^sub>1\ are arbitrary. Elim-resolution plays a key role in simulating sequent proofs. We can easily handle reasoning on the left. Elim-resolution with the rules \(\E)\, \(\E)\ and \(\E)\ achieves a similar effect as the corresponding sequent rules. For the other connectives, we use sequent-style elimination rules instead of destruction rules such as \(\E1, 2)\ and \(\E)\. But note that the rule \(\L)\ has no effect under our representation of sequents! \[ \infer[\(\L)\]{\\ P, \ \ \\}{\\ \ \, P\} \] What about reasoning on the right? Introduction rules can only affect the formula in the conclusion, namely \Q\<^sub>1\. The other right-side formulae are represented as negated assumptions, \\ Q\<^sub>2, \, \ Q\<^sub>n\. In order to operate on one of these, it must first be exchanged with \Q\<^sub>1\. Elim-resolution with the \swap\ rule has this effect: \\ P \ (\ R \ P) \ R\ To ensure that swaps occur only when necessary, each introduction rule is converted into a swapped form: it is resolved with the second premise of \(swap)\. The swapped form of \(\I)\, which might be called \(\\E)\, is @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} Similarly, the swapped form of \(\I)\ is @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} Swapped introduction rules are applied using elim-resolution, which deletes the negated formula. Our representation of sequents also requires the use of ordinary introduction rules. If we had no regard for readability of intermediate goal states, we could treat the right side more uniformly by representing sequents as @{text [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} \ subsubsection \Extra rules for the sequent calculus\ text \As mentioned, destruction rules such as \(\E1, 2)\ and \(\E)\ must be replaced by sequent-style elimination rules. In addition, we need rules to embody the classical equivalence between \P \ Q\ and \\ P \ Q\. The introduction rules \(\I1, 2)\ are replaced by a rule that simulates \(\R)\: @{text [display] "(\ Q \ P) \ P \ Q"} The destruction rule \(\E)\ is replaced by @{text [display] "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} Quantifier replication also requires special rules. In classical logic, \\x. P x\ is equivalent to \\ (\x. \ P x)\; the rules \(\R)\ and \(\L)\ are dual: \[ \infer[\(\R)\]{\\ \ \, \x. P x\}{\\ \ \, \x. P x, P t\} \qquad \infer[\(\L)\]{\\x. P x, \ \ \\}{\P t, \x. P x, \ \ \\} \] Thus both kinds of quantifier may be replicated. Theorems requiring multiple uses of a universal formula are easy to invent; consider @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any \n > 1\. Natural examples of the multiple use of an existential formula are rare; a standard one is \\x. \y. P x \ P y\. Forgoing quantifier replication loses completeness, but gains decidability, since the search space becomes finite. Many useful theorems can be proved without replication, and the search generally delivers its verdict in a reasonable time. To adopt this approach, represent the sequent rules \(\R)\, \(\L)\ and \(\R)\ by \(\I)\, \(\E)\ and \(\I)\, respectively, and put \(\E)\ into elimination form: @{text [display] "\x. P x \ (P t \ Q) \ Q"} Elim-resolution with this rule will delete the universal formula after a single use. To replicate universal quantifiers, replace the rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} To replicate existential quantifiers, replace \(\I)\ by @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} All introduction rules mentioned above are also useful in swapped form. Replication makes the search space infinite; we must apply the rules with care. The classical reasoner distinguishes between safe and unsafe rules, applying the latter only when there is no alternative. Depth-first search may well go down a blind alley; best-first search is better behaved in an infinite search space. However, quantifier replication is too expensive to prove any but the simplest theorems. \ subsection \Rule declarations\ text \The proof tools of the Classical Reasoner depend on collections of rules declared in the context, which are classified as introduction, elimination or destruction and as \<^emph>\safe\ or \<^emph>\unsafe\. In general, safe rules can be attempted blindly, while unsafe rules must be used with care. A safe rule must never reduce a provable goal to an unprovable set of subgoals. The rule \P \ P \ Q\ is unsafe because it reduces \P \ Q\ to \P\, which might turn out as premature choice of an unprovable subgoal. Any rule whose premises contain new unknowns is unsafe. The elimination rule \\x. P x \ (P t \ Q) \ Q\ is unsafe, since it is applied via elim-resolution, which discards the assumption \\x. P x\ and replaces it by the weaker assumption \P t\. The rule \P t \ \x. P x\ is unsafe for similar reasons. The quantifier duplication rule \\x. P x \ (P t \ \x. P x \ Q) \ Q\ is unsafe in a different sense: since it keeps the assumption \\x. P x\, it is prone to looping. In classical first-order logic, all rules are safe except those mentioned above. The safe~/ unsafe distinction is vague, and may be regarded merely as a way of giving some rules priority over others. One could argue that \(\E)\ is unsafe, because repeated application of it could generate exponentially many subgoals. Induction rules are unsafe because inductive proofs are difficult to set up automatically. Any inference that instantiates an unknown in the proof state is unsafe --- thus matching must be used, rather than unification. Even proof by assumption is unsafe if it instantiates unknowns shared with other subgoals. \begin{matharray}{rcl} @{command_def "print_claset"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def intro} & : & \attribute\ \\ @{attribute_def elim} & : & \attribute\ \\ @{attribute_def dest} & : & \attribute\ \\ @{attribute_def rule} & : & \attribute\ \\ @{attribute_def iff} & : & \attribute\ \\ @{attribute_def swapped} & : & \attribute\ \\ \end{matharray} \<^rail>\ (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? ; @@{attribute rule} 'del' ; @@{attribute iff} (((() | 'add') '?'?) | 'del') \ \<^descr> @{command "print_claset"} prints the collection of rules declared to the Classical Reasoner, i.e.\ the \<^ML_type>\claset\ within the context. \<^descr> @{attribute intro}, @{attribute elim}, and @{attribute dest} declare introduction, elimination, and destruction rules, respectively. By default, rules are considered as \<^emph>\unsafe\ (i.e.\ not applied blindly without backtracking), while ``\!\'' classifies as \<^emph>\safe\. Rule declarations marked by ``\?\'' coincide with those of Isabelle/Pure, cf.\ \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps of the @{method rule} method). The optional natural number specifies an explicit weight argument, which is ignored by the automated reasoning tools, but determines the search order of single rule steps. Introduction rules are those that can be applied using ordinary resolution. Their swapped forms are generated internally, which will be applied using elim-resolution. Elimination rules are applied using elim-resolution. Rules are sorted by the number of new subgoals they will yield; rules that generate the fewest subgoals will be tried first. Otherwise, later declarations take precedence over earlier ones. Rules already present in the context with the same classification are ignored. A warning is printed if the rule has already been added with some other classification, but the rule is added anyway as requested. \<^descr> @{attribute rule}~\del\ deletes all occurrences of a rule from the classical context, regardless of its classification as introduction~/ elimination~/ destruction and safe~/ unsafe. \<^descr> @{attribute iff} declares logical equivalences to the Simplifier and the Classical reasoner at the same time. Non-conditional rules result in a safe introduction and elimination pair; conditional ones are considered unsafe. Rules with negative conclusion are automatically inverted (using \\\-elimination internally). The ``\?\'' version of @{attribute iff} declares rules to the Isabelle/Pure context only, and omits the Simplifier declaration. \<^descr> @{attribute swapped} turns an introduction rule into an elimination, by resolving with the classical swap principle \\ P \ (\ R \ P) \ R\ in the second position. This is mainly for illustrative purposes: the Classical Reasoner already swaps rules internally as explained above. \ subsection \Structured methods\ text \ \begin{matharray}{rcl} @{method_def rule} & : & \method\ \\ @{method_def contradiction} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method rule} @{syntax thms}? \ \<^descr> @{method rule} as offered by the Classical Reasoner is a refinement over the Pure one (see \secref{sec:pure-meth-att}). Both versions work the same, but the classical version observes the classical rule context in addition to that of Isabelle/Pure. Common object logics (HOL, ZF, etc.) declare a rich collection of classical rules (even if these would qualify as intuitionistic ones), but only few declarations to the rule context of Isabelle/Pure (\secref{sec:pure-meth-att}). \<^descr> @{method contradiction} solves some goal by contradiction, deriving any result from both \\ A\ and \A\. Chained facts, which are guaranteed to participate, may appear in either order. \ subsection \Fully automated methods\ text \ \begin{matharray}{rcl} @{method_def blast} & : & \method\ \\ @{method_def auto} & : & \method\ \\ @{method_def force} & : & \method\ \\ @{method_def fast} & : & \method\ \\ @{method_def slow} & : & \method\ \\ @{method_def best} & : & \method\ \\ @{method_def fastforce} & : & \method\ \\ @{method_def slowsimp} & : & \method\ \\ @{method_def bestsimp} & : & \method\ \\ @{method_def deepen} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method blast} @{syntax nat}? (@{syntax clamod} * ) ; @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) ; @@{method force} (@{syntax clasimpmod} * ) ; (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) ; (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) (@{syntax clasimpmod} * ) ; @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) ; @{syntax_def clamod}: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thms} ; @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | 'cong' (() | 'add' | 'del') | 'split' (() | '!' | 'del') | 'iff' (((() | 'add') '?'?) | 'del') | (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thms} \ \<^descr> @{method blast} is a separate classical tableau prover that uses the same classical rule declarations as explained before. Proof search is coded directly in ML using special data structures. A successful proof is then reconstructed using regular Isabelle inferences. It is faster and more powerful than the other classical reasoning tools, but has major limitations too. \<^item> It does not use the classical wrapper tacticals, such as the integration with the Simplifier of @{method fastforce}. \<^item> It does not perform higher-order unification, as needed by the rule @{thm [source=false] rangeI} in HOL. There are often alternatives to such rules, for example @{thm [source=false] range_eqI}. \<^item> Function variables may only be applied to parameters of the subgoal. (This restriction arises because the prover does not use higher-order unification.) If other function variables are present then the prover will fail with the message @{verbatim [display] \Function unknown's argument not a bound variable\} \<^item> Its proof strategy is more general than @{method fast} but can be slower. If @{method blast} fails or seems to be running forever, try @{method fast} and the other proof tools described below. The optional integer argument specifies a bound for the number of unsafe steps used in a proof. By default, @{method blast} starts with a bound of 0 and increases it successively to 20. In contrast, \(blast lim)\ tries to prove the goal using a search bound of \lim\. Sometimes a slow proof using @{method blast} can be made much faster by supplying the successful search bound to this proof method instead. \<^descr> @{method auto} combines classical reasoning with simplification. It is intended for situations where there are a lot of mostly trivial subgoals; it proves all the easy ones, leaving the ones it cannot prove. Occasionally, attempting to prove the hard ones may take a long time. The optional depth arguments in \(auto m n)\ refer to its builtin classical reasoning procedures: \m\ (default 4) is for @{method blast}, which is tried first, and \n\ (default 2) is for a slower but more general alternative that also takes wrappers into account. \<^descr> @{method force} is intended to prove the first subgoal completely, using many fancy proof tools and performing a rather exhaustive search. As a result, proof attempts may take rather long or diverge easily. \<^descr> @{method fast}, @{method best}, @{method slow} attempt to prove the first subgoal using sequent-style reasoning as explained before. Unlike @{method blast}, they construct proofs directly in Isabelle. There is a difference in search strategy and back-tracking: @{method fast} uses depth-first search and @{method best} uses best-first search (guided by a heuristic function: normally the total size of the proof state). Method @{method slow} is like @{method fast}, but conducts a broader search: it may, when backtracking from a failed proof attempt, undo even the step of proving a subgoal by assumption. \<^descr> @{method fastforce}, @{method slowsimp}, @{method bestsimp} are like @{method fast}, @{method slow}, @{method best}, respectively, but use the Simplifier as additional wrapper. The name @{method fastforce}, reflects the behaviour of this popular method better without requiring an understanding of its implementation. \<^descr> @{method deepen} works by exhaustive search up to a certain depth. The start depth is 4 (unless specified explicitly), and the depth is increased iteratively up to 10. Unsafe rules are modified to preserve the formula they act on, so that it be used repeatedly. This method can prove more goals than @{method fast}, but is much slower, for example if the assumptions have many universal quantifiers. Any of the above methods support additional modifiers of the context of classical (and simplifier) rules, but the ones related to the Simplifier are explicitly prefixed by \simp\ here. The semantics of these ad-hoc rule declarations is analogous to the attributes given before. Facts provided by forward chaining are inserted into the goal before commencing proof search. \ subsection \Partially automated methods\label{sec:classical:partial}\ text \These proof methods may help in situations when the fully-automated tools fail. The result is a simpler subgoal that can be tackled by other means, such as by manual instantiation of quantifiers. \begin{matharray}{rcl} @{method_def safe} & : & \method\ \\ @{method_def clarify} & : & \method\ \\ @{method_def clarsimp} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) ; @@{method clarsimp} (@{syntax clasimpmod} * ) \ \<^descr> @{method safe} repeatedly performs safe steps on all subgoals. It is deterministic, with at most one outcome. \<^descr> @{method clarify} performs a series of safe steps without splitting subgoals; see also @{method clarify_step}. \<^descr> @{method clarsimp} acts like @{method clarify}, but also does simplification. Note that if the Simplifier context includes a splitter for the premises, the subgoal may still be split. \ subsection \Single-step tactics\ text \ \begin{matharray}{rcl} @{method_def safe_step} & : & \method\ \\ @{method_def inst_step} & : & \method\ \\ @{method_def step} & : & \method\ \\ @{method_def slow_step} & : & \method\ \\ @{method_def clarify_step} & : & \method\ \\ \end{matharray} These are the primitive tactics behind the automated proof methods of the Classical Reasoner. By calling them yourself, you can execute these procedures one step at a time. \<^descr> @{method safe_step} performs a safe step on the first subgoal. The safe wrapper tacticals are applied to a tactic that may include proof by assumption or Modus Ponens (taking care not to instantiate unknowns), or substitution. \<^descr> @{method inst_step} is like @{method safe_step}, but allows unknowns to be instantiated. \<^descr> @{method step} is the basic step of the proof procedure, it operates on the first subgoal. The unsafe wrapper tacticals are applied to a tactic that tries @{method safe}, @{method inst_step}, or applies an unsafe rule from the context. \<^descr> @{method slow_step} resembles @{method step}, but allows backtracking between using safe rules with instantiation (@{method inst_step}) and using unsafe rules. The resulting search space is larger. \<^descr> @{method clarify_step} performs a safe step on the first subgoal; no splitting step is applied. For example, the subgoal \A \ B\ is left as a conjunction. Proof by assumption, Modus Ponens, etc., may be performed provided they do not instantiate unknowns. Assumptions of the form \x = t\ may be eliminated. The safe wrapper tactical is applied. \ subsection \Modifying the search step\ text \ \begin{mldecls} @{define_ML_type wrapper = "(int -> tactic) -> (int -> tactic)"} \\[0.5ex] @{define_ML_infix addSWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addSbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addSafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML_infix addWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML addSss: "Proof.context -> Proof.context"} \\ @{define_ML addss: "Proof.context -> Proof.context"} \\ \end{mldecls} The proof strategy of the Classical Reasoner is simple. Perform as many safe inferences as possible; or else, apply certain safe rules, allowing instantiation of unknowns; or else, apply an unsafe rule. The tactics also eliminate assumptions of the form \x = t\ by substitution if they have been set up to do so. They may perform a form of Modus Ponens: if there are assumptions \P \ Q\ and \P\, then replace \P \ Q\ by \Q\. The classical reasoning tools --- except @{method blast} --- allow to modify this basic proof strategy by applying two lists of arbitrary \<^emph>\wrapper tacticals\ to it. The first wrapper list, which is considered to contain safe wrappers only, affects @{method safe_step} and all the tactics that call it. The second one, which may contain unsafe wrappers, affects the unsafe parts of @{method step}, @{method slow_step}, and the tactics that call them. A wrapper transforms each step of the search, for example by attempting other tactics before or after the original step tactic. All members of a wrapper list are applied in turn to the respective step tactic. Initially the two wrapper lists are empty, which means no modification of the step tactics. Safe and unsafe wrappers are added to the context with the functions given below, supplying them with wrapper names. These names may be used to selectively delete wrappers. \<^descr> \ctxt addSWrapper (name, wrapper)\ adds a new wrapper, which should yield a safe tactic, to modify the existing safe step tactic. \<^descr> \ctxt addSbefore (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried \<^emph>\before\ each safe step of the search. \<^descr> \ctxt addSafter (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried when a safe step of the search would fail. \<^descr> \ctxt delSWrapper name\ deletes the safe wrapper with the given name. \<^descr> \ctxt addWrapper (name, wrapper)\ adds a new wrapper to modify the existing (unsafe) step tactic. \<^descr> \ctxt addbefore (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\before\ the result of each unsafe step. \<^descr> \ctxt addafter (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\after\ the result of each unsafe step. \<^descr> \ctxt delWrapper name\ deletes the unsafe wrapper with the given name. \<^descr> \addSss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, in a rather safe way, after each safe step of the search. \<^descr> \addss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, before the each unsafe step of the search. \ section \Object-logic setup \label{sec:object-logic}\ text \ \begin{matharray}{rcl} @{command_def "judgment"} & : & \theory \ theory\ \\ @{method_def atomize} & : & \method\ \\ @{attribute_def atomize} & : & \attribute\ \\ @{attribute_def rule_format} & : & \attribute\ \\ @{attribute_def rulify} & : & \attribute\ \\ \end{matharray} The very starting point for any Isabelle object-logic is a ``truth judgment'' that links object-level statements to the meta-logic (with its minimal language of \prop\ that covers universal quantification \\\ and implication \\\). Common object-logics are sufficiently expressive to internalize rule statements over \\\ and \\\ within their own language. This is useful in certain situations where a rule needs to be viewed as an atomic statement from the meta-level perspective, e.g.\ \\x. x \ A \ P x\ versus \\x \ A. P x\. From the following language elements, only the @{method atomize} method and @{attribute rule_format} attribute are occasionally required by end-users, the rest is for those who need to setup their own object-logic. In the latter case existing formulations of Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. Generic tools may refer to the information provided by object-logic declarations internally. \<^rail>\ @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? ; @@{attribute atomize} ('(' 'full' ')')? ; @@{attribute rule_format} ('(' 'noasm' ')')? \ \<^descr> @{command "judgment"}~\c :: \ (mx)\ declares constant \c\ as the truth judgment of the current object-logic. Its type \\\ should specify a coercion of the category of object-level propositions to \prop\ of the Pure meta-logic; the mixfix annotation \(mx)\ would typically just link the object language (internally of syntactic category \logic\) with that of \prop\. Only one @{command "judgment"} declaration may be given in any theory development. \<^descr> @{method atomize} (as a method) rewrites any non-atomic premises of a sub-goal, using the meta-level equations declared via @{attribute atomize} (as an attribute) beforehand. As a result, heavily nested goals become amenable to fundamental operations such as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``\(full)\'' option here means to turn the whole subgoal into an object-statement (if possible), including the outermost parameters and assumptions as well. A typical collection of @{attribute atomize} rules for a particular object-logic would provide an internalization for each of the connectives of \\\, \\\, and \\\. Meta-level conjunction should be covered as well (this is particularly important for locales, see \secref{sec:locale}). \<^descr> @{attribute rule_format} rewrites a theorem by the equalities declared as @{attribute rulify} rules in the current object-logic. By default, the result is fully normalized, including assumptions and conclusions at any depth. The \(no_asm)\ option restricts the transformation to the conclusion of a rule. In common object-logics (HOL, FOL, ZF), the effect of @{attribute rule_format} is to replace (bounded) universal quantification (\\\) and implication (\\\) by the corresponding rule statements over \\\ and \\\. \ section \Tracing higher-order unification\ text \ \begin{tabular}{rcll} @{attribute_def unify_trace_simp} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_types} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_bound} & : & \attribute\ & default \50\ \\ @{attribute_def unify_search_bound} & : & \attribute\ & default \60\ \\ \end{tabular} \<^medskip> Higher-order unification works well in most practical situations, but sometimes needs extra care to identify problems. These tracing options may help. \<^descr> @{attribute unify_trace_simp} controls tracing of the simplification phase of higher-order unification. \<^descr> @{attribute unify_trace_types} controls warnings of incompleteness, when unification is not considering all possible instantiations of schematic type variables. \<^descr> @{attribute unify_trace_bound} determines the depth where unification starts to print tracing information once it reaches depth; 0 for full tracing. At the default value, tracing information is almost never printed in practice. \<^descr> @{attribute unify_search_bound} prevents unification from searching past the given depth. Because of this bound, higher-order unification cannot return an infinite sequence, though it can return an exponentially long one. The search rarely approaches the default value in practice. If the search is cut off, unification prints a warning ``Unification bound exceeded''. \begin{warn} Options for unification cannot be modified in a local context. Only the global theory content is taken into account. \end{warn} \ end diff --git a/src/HOL/Bali/Eval.thy b/src/HOL/Bali/Eval.thy --- a/src/HOL/Bali/Eval.thy +++ b/src/HOL/Bali/Eval.thy @@ -1,1179 +1,1179 @@ (* Title: HOL/Bali/Eval.thy Author: David von Oheimb *) subsection \Operational evaluation (big-step) semantics of Java expressions and statements \ theory Eval imports State DeclConcepts begin text \ improvements over Java Specification 1.0: \begin{itemize} \item dynamic method lookup does not need to consider the return type (cf.15.11.4.4) \item throw raises a NullPointer exception if a null reference is given, and each throw of a standard exception yield a fresh exception object (was not specified) \item if there is not enough memory even to allocate an OutOfMemory exception, evaluation/execution fails, i.e. simply stops (was not specified) \item array assignment checks lhs (and may throw exceptions) before evaluating rhs \item fixed exact positions of class initializations (immediate at first active use) \end{itemize} design issues: \begin{itemize} \item evaluation vs. (single-step) transition semantics evaluation semantics chosen, because: \begin{itemize} \item[++] less verbose and therefore easier to read (and to handle in proofs) \item[+] more abstract \item[+] intermediate values (appearing in recursive rules) need not be stored explicitly, e.g. no call body construct or stack of invocation frames containing local variables and return addresses for method calls needed \item[+] convenient rule induction for subject reduction theorem \item[-] no interleaving (for parallelism) can be described \item[-] stating a property of infinite executions requires the meta-level argument that this property holds for any finite prefixes of it (e.g. stopped using a counter that is decremented to zero and then throwing an exception) \end{itemize} \item unified evaluation for variables, expressions, expression lists, statements \item the value entry in statement rules is redundant \item the value entry in rules is irrelevant in case of exceptions, but its full inclusion helps to make the rule structure independent of exception occurrence. \item as irrelevant value entries are ignored, it does not matter if they are unique. For simplicity, (fixed) arbitrary values are preferred over "free" values. \item the rule format is such that the start state may contain an exception. \begin{itemize} \item[++] faciliates exception handling \item[+] symmetry \end{itemize} \item the rules are defined carefully in order to be applicable even in not type-correct situations (yielding undefined values), e.g. \the_Addr (Val (Bool b)) = undefined\. \begin{itemize} \item[++] fewer rules \item[-] less readable because of auxiliary functions like \the_Addr\ \end{itemize} Alternative: "defensive" evaluation throwing some InternalError exception in case of (impossible, for correct programs) type mismatches \item there is exactly one rule per syntactic construct \begin{itemize} \item[+] no redundancy in case distinctions \end{itemize} \item halloc fails iff there is no free heap address. When there is only one free heap address left, it returns an OutOfMemory exception. In this way it is guaranteed that when an OutOfMemory exception is thrown for the first time, there is a free location on the heap to allocate it. \item the allocation of objects that represent standard exceptions is deferred until execution of any enclosing catch clause, which is transparent to the program. \begin{itemize} \item[-] requires an auxiliary execution relation \item[++] avoids copies of allocation code and awkward case distinctions (whether there is enough memory to allocate the exception) in evaluation rules \end{itemize} \item unfortunately \new_Addr\ is not directly executable because of Hilbert operator. \end{itemize} simplifications: \begin{itemize} \item local variables are initialized with default values (no definite assignment) \item garbage collection not considered, therefore also no finalizers \item stack overflow and memory overflow during class initialization not modelled \item exceptions in initializations not replaced by ExceptionInInitializerError \end{itemize} \ type_synonym vvar = "val \ (val \ state \ state)" type_synonym vals = "(val, vvar, val list) sum3" translations (type) "vvar" <= (type) "val \ (val \ state \ state)" (type) "vals" <= (type) "(val, vvar, val list) sum3" text \To avoid redundancy and to reduce the number of rules, there is only one evaluation rule for each syntactic term. This is also true for variables (e.g. see the rules below for \LVar\, \FVar\ and \AVar\). So evaluation of a variable must capture both possible further uses: read (rule \Acc\) or write (rule \Ass\) to the variable. Therefor a variable evaluates to a special value \<^term>\vvar\, which is a pair, consisting of the current value (for later read access) and an update function (for later write access). Because during assignment to an array variable an exception may occur if the types don't match, the update function is very generic: it transforms the full state. This generic update function causes some technical trouble during some proofs (e.g. type safety, correctness of definite assignment). There we need to prove some additional invariant on this update function to prove the assignment correct, since the update function could potentially alter the whole state in an arbitrary manner. This invariant must be carried around through the whole induction. So for future approaches it may be better not to take such a generic update function, but only to store the address and the kind of variable (array (+ element type), local variable or field) for later assignment. \ abbreviation dummy_res :: "vals" ("\") where "\ == In1 Unit" abbreviation (input) val_inj_vals ("\_\\<^sub>e" 1000) where "\e\\<^sub>e == In1 e" abbreviation (input) var_inj_vals ("\_\\<^sub>v" 1000) where "\v\\<^sub>v == In2 v" abbreviation (input) lst_inj_vals ("\_\\<^sub>l" 1000) where "\es\\<^sub>l == In3 es" definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \ vals" where "undefined3 = case_sum3 (In1 \ case_sum (\x. undefined) (\x. Unit)) (\x. In2 undefined) (\x. In3 undefined)" lemma [simp]: "undefined3 (In1l x) = In1 undefined" by (simp add: undefined3_def) lemma [simp]: "undefined3 (In1r x) = \" by (simp add: undefined3_def) lemma [simp]: "undefined3 (In2 x) = In2 undefined" by (simp add: undefined3_def) lemma [simp]: "undefined3 (In3 x) = In3 undefined" by (simp add: undefined3_def) subsubsection "exception throwing and catching" definition throw :: "val \ abopt \ abopt" where "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)" lemma throw_def2: "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)" apply (unfold throw_def) apply (simp (no_asm)) done definition fits :: "prog \ st \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) where "G,s\a' fits T = ((\rt. T=RefT rt) \ a'=Null \ G\obj_ty(lookup_obj s a')\T)" lemma fits_Null [simp]: "G,s\Null fits T" by (simp add: fits_def) lemma fits_Addr_RefT [simp]: "G,s\Addr a fits RefT t = G\obj_ty (the (heap s a))\RefT t" by (simp add: fits_def) lemma fitsD: "\X. G,s\a' fits T \ (\pt. T = PrimT pt) \ (\t. T = RefT t) \ a' = Null \ (\t. T = RefT t) \ a' \ Null \ G\obj_ty (lookup_obj s a')\T" apply (unfold fits_def) apply (case_tac "\pt. T = PrimT pt") apply simp_all apply (case_tac "T") defer apply (case_tac "a' = Null") apply simp_all done definition catch :: "prog \ state \ qtname \ bool" ("_,_\catch _"[61,61,61]60) where "G,s\catch C = (\xc. abrupt s=Some (Xcpt xc) \ G,store s\Addr (the_Loc xc) fits Class C)" lemma catch_Norm [simp]: "\G,Norm s\catch tn" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_XcptLoc [simp]: "G,(Some (Xcpt (Loc a)),s)\catch C = G,s\Addr a fits Class C" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_Jump [simp]: "\G,(Some (Jump j),s)\catch tn" apply (unfold catch_def) apply (simp (no_asm)) done lemma catch_Error [simp]: "\G,(Some (Error e),s)\catch tn" apply (unfold catch_def) apply (simp (no_asm)) done definition new_xcpt_var :: "vname \ state \ state" where "new_xcpt_var vn = (\(x,s). Norm (lupd(VName vn\Addr (the_Loc (the_Xcpt (the x)))) s))" lemma new_xcpt_var_def2 [simp]: "new_xcpt_var vn (x,s) = Norm (lupd(VName vn\Addr (the_Loc (the_Xcpt (the x)))) s)" apply (unfold new_xcpt_var_def) apply (simp (no_asm)) done subsubsection "misc" definition assign :: "('a \ state \ state) \ 'a \ state \ state" where "assign f v = (\(x,s). let (x',s') = (if x = None then f v else id) (x,s) in (x',if x' = None then s' else s))" (* lemma assign_Norm_Norm [simp]: "f v \abrupt=None,store=s\ = \abrupt=None,store=s'\ \ assign f v \abrupt=None,store=s\ = \abrupt=None,store=s'\" by (simp add: assign_def Let_def) *) lemma assign_Norm_Norm [simp]: "f v (Norm s) = Norm s' \ assign f v (Norm s) = Norm s'" by (simp add: assign_def Let_def) (* lemma assign_Norm_Some [simp]: "\abrupt (f v \abrupt=None,store=s\) = Some y\ \ assign f v \abrupt=None,store=s\ = \abrupt=Some y,store =s\" by (simp add: assign_def Let_def split_beta) *) lemma assign_Norm_Some [simp]: "\abrupt (f v (Norm s)) = Some y\ \ assign f v (Norm s) = (Some y,s)" by (simp add: assign_def Let_def split_beta) lemma assign_Some [simp]: "assign f v (Some x,s) = (Some x,s)" by (simp add: assign_def Let_def split_beta) lemma assign_Some1 [simp]: "\ normal s \ assign f v s = s" by (auto simp add: assign_def Let_def split_beta) lemma assign_supd [simp]: "assign (\v. supd (f v)) v (x,s) = (x, if x = None then f v s else s)" apply auto done lemma assign_raise_if [simp]: "assign (\v (x,s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) = (raise_if (b s v) xcpt x, if x=None \ \b s v then f v s else s)" apply (case_tac "x = None") apply auto done (* lemma assign_raise_if [simp]: "assign (\v s. \abrupt=(raise_if (b (store s) v) xcpt) (abrupt s), store = f v (store s)\) v s = \abrupt=raise_if (b (store s) v) xcpt (abrupt s), store= if (abrupt s)=None \ \b (store s) v then f v (store s) else (store s)\" apply (case_tac "abrupt s = None") apply auto done *) definition init_comp_ty :: "ty \ stmt" where "init_comp_ty T = (if (\C. T = Class C) then Init (the_Class T) else Skip)" lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip" apply (unfold init_comp_ty_def) apply (simp (no_asm)) done definition invocation_class :: "inv_mode \ st \ val \ ref_ty \ qtname" where "invocation_class m s a' statT = (case m of Static \ if (\ statC. statT = ClassT statC) then the_Class (RefT statT) else Object | SuperM \ the_Class (RefT statT) | IntVir \ obj_class (lookup_obj s a'))" definition invocation_declclass :: "prog \ inv_mode \ st \ val \ ref_ty \ sig \ qtname" where "invocation_declclass G m s a' statT sig = declclass (the (dynlookup G statT (invocation_class m s a' statT) sig))" lemma invocation_class_IntVir [simp]: "invocation_class IntVir s a' statT = obj_class (lookup_obj s a')" by (simp add: invocation_class_def) lemma dynclass_SuperM [simp]: "invocation_class SuperM s a' statT = the_Class (RefT statT)" by (simp add: invocation_class_def) lemma invocation_class_Static [simp]: "invocation_class Static s a' statT = (if (\ statC. statT = ClassT statC) then the_Class (RefT statT) else Object)" by (simp add: invocation_class_def) definition init_lvars :: "prog \ qtname \ sig \ inv_mode \ val \ val list \ state \ state" where "init_lvars G C sig mode a' pvs = (\(x,s). let m = mthd (the (methd G C sig)); l = \ k. (case k of EName e \ (case e of VNam v \ (Map.empty ((pars m)[\]pvs)) v | Res \ None) | This \ (if mode=Static then None else Some a')) in set_lvars l (if mode = Static then x else np a' x,s))" lemma init_lvars_def2: \ \better suited for simplification\ "init_lvars G C sig mode a' pvs (x,s) = set_lvars (\ k. (case k of EName e \ (case e of VNam v \ (Map.empty ((pars (mthd (the (methd G C sig))))[\]pvs)) v | Res \ None) | This \ (if mode=Static then None else Some a'))) (if mode = Static then x else np a' x,s)" apply (unfold init_lvars_def) apply (simp (no_asm) add: Let_def) done definition body :: "prog \ qtname \ sig \ expr" where "body G C sig = (let m = the (methd G C sig) in Body (declclass m) (stmt (mbody (mthd m))))" lemma body_def2: \ \better suited for simplification\ "body G C sig = Body (declclass (the (methd G C sig))) (stmt (mbody (mthd (the (methd G C sig)))))" apply (unfold body_def Let_def) apply auto done subsubsection "variables" definition lvar :: "lname \ st \ vvar" where "lvar vn s = (the (locals s vn), \v. supd (lupd(vn\v)))" definition fvar :: "qtname \ bool \ vname \ val \ state \ vvar \ state" where "fvar C stat fn a' s = (let (oref,xf) = if stat then (Stat C,id) else (Heap (the_Addr a'),np a'); n = Inl (fn,C); f = (\v. supd (upd_gobj oref n v)) in ((the (values (the (globs (store s) oref)) n),f),abupd xf s))" definition avar :: "prog \ val \ val \ state \ vvar \ state" where "avar G i' a' s = (let oref = Heap (the_Addr a'); i = the_Intg i'; n = Inr i; (T,k,cs) = the_Arr (globs (store s) oref); f = (\v (x,s). (raise_if (\G,s\v fits T) ArrStore x ,upd_gobj oref n v s)) in ((the (cs n),f),abupd (raise_if (\i in_bounds k) IndOutBound \ np a') s))" lemma fvar_def2: \ \better suited for simplification\ "fvar C stat fn a' s = ((the (values (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) (Inl (fn,C))) ,(\v. supd (upd_gobj (if stat then Stat C else Heap (the_Addr a')) (Inl (fn,C)) v))) ,abupd (if stat then id else np a') s) " apply (unfold fvar_def) apply (simp (no_asm) add: Let_def split_beta) done lemma avar_def2: \ \better suited for simplification\ "avar G i' a' s = ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) (Inr (the_Intg i'))) ,(\v (x,s'). (raise_if (\G,s'\v fits (fst(the_Arr (globs (store s) (Heap (the_Addr a')))))) ArrStore x ,upd_gobj (Heap (the_Addr a')) (Inr (the_Intg i')) v s'))) ,abupd (raise_if (\(the_Intg i') in_bounds (fst(snd(the_Arr (globs (store s) (Heap (the_Addr a'))))))) IndOutBound \ np a') s)" apply (unfold avar_def) apply (simp (no_asm) add: Let_def split_beta) done definition check_field_access :: "prog \ qtname \ qtname \ vname \ bool \ val \ state \ state" where "check_field_access G accC statDeclC fn stat a' s = (let oref = if stat then Stat statDeclC else Heap (the_Addr a'); dynC = case oref of Heap a \ obj_class (the (globs (store s) oref)) | Stat C \ C; f = (the (table_of (DeclConcepts.fields G dynC) (fn,statDeclC))) in abupd (error_if (\ G\Field fn (statDeclC,f) in dynC dyn_accessible_from accC) AccessViolation) s)" definition check_method_access :: "prog \ qtname \ ref_ty \ inv_mode \ sig \ val \ state \ state" where "check_method_access G accC statT mode sig a' s = (let invC = invocation_class mode (store s) a' statT; dynM = the (dynlookup G statT invC sig) in abupd (error_if (\ G\Methd sig dynM in invC dyn_accessible_from accC) AccessViolation) s)" subsubsection "evaluation judgments" inductive halloc :: "[prog,state,obj_tag,loc,state]\bool" ("_\_ \halloc _\_\ _"[61,61,61,61,61]60) for G::prog where \ \allocating objects on the heap, cf. 12.5\ Abrupt: "G\(Some x,s) \halloc oi\undefined\ (Some x,s)" | New: "\new_Addr (heap s) = Some a; (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi) else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\ \ G\Norm s \halloc oi\a\ (x,init_obj G oi' (Heap a) s)" inductive sxalloc :: "[prog,state,state]\bool" ("_\_ \sxalloc\ _"[61,61,61]60) for G::prog where \ \allocating exception objects for standard exceptions (other than OutOfMemory)\ Norm: "G\ Norm s \sxalloc\ Norm s" | Jmp: "G\(Some (Jump j), s) \sxalloc\ (Some (Jump j), s)" | Error: "G\(Some (Error e), s) \sxalloc\ (Some (Error e), s)" | XcptL: "G\(Some (Xcpt (Loc a) ),s) \sxalloc\ (Some (Xcpt (Loc a)),s)" | SXcpt: "\G\Norm s0 \halloc (CInst (SXcpt xn))\a\ (x,s1)\ \ G\(Some (Xcpt (Std xn)),s0) \sxalloc\ (Some (Xcpt (Loc a)),s1)" inductive eval :: "[prog,state,term,vals,state]\bool" ("_\_ \_\\ '(_, _')" [61,61,80,0,0]60) and exec ::"[prog,state,stmt ,state]\bool"("_\_ \_\ _" [61,61,65, 61]60) and evar ::"[prog,state,var ,vvar,state]\bool"("_\_ \_=\_\ _"[61,61,90,61,61]60) and eval'::"[prog,state,expr ,val ,state]\bool"("_\_ \_-\_\ _"[61,61,80,61,61]60) and evals::"[prog,state,expr list , val list ,state]\bool"("_\_ \_\\_\ _"[61,61,61,61,61]60) for G::prog where "G\s \c \ s' \ G\s \In1r c\\ (\, s')" | "G\s \e-\v \ s' \ G\s \In1l e\\ (In1 v, s')" | "G\s \e=\vf\ s' \ G\s \In2 e\\ (In2 vf, s')" | "G\s \e\\v \ s' \ G\s \In3 e\\ (In3 v, s')" \ \propagation of abrupt completion\ \ \cf. 14.1, 15.5\ | Abrupt: "G\(Some xc,s) \t\\ (undefined3 t, (Some xc, s))" \ \execution of statements\ \ \cf. 14.5\ | Skip: "G\Norm s \Skip\ Norm s" \ \cf. 14.7\ | Expr: "\G\Norm s0 \e-\v\ s1\ \ G\Norm s0 \Expr e\ s1" | Lab: "\G\Norm s0 \c \ s1\ \ G\Norm s0 \l\ c\ abupd (absorb l) s1" \ \cf. 14.2\ | Comp: "\G\Norm s0 \c1 \ s1; G\ s1 \c2 \ s2\ \ G\Norm s0 \c1;; c2\ s2" \ \cf. 14.8.2\ | If: "\G\Norm s0 \e-\b\ s1; G\ s1\(if the_Bool b then c1 else c2)\ s2\ \ G\Norm s0 \If(e) c1 Else c2 \ s2" \ \cf. 14.10, 14.10.1\ \ \A continue jump from the while body \<^term>\c\ is handled by this rule. If a continue jump with the proper label was invoked inside \<^term>\c\ this label (Cont l) is deleted out of the abrupt component of the state before the iterative evaluation of the while statement. A break jump is handled by the Lab Statement \Lab l (while\)\.\ | Loop: "\G\Norm s0 \e-\b\ s1; if the_Bool b then (G\s1 \c\ s2 \ G\(abupd (absorb (Cont l)) s2) \l\ While(e) c\ s3) else s3 = s1\ \ G\Norm s0 \l\ While(e) c\ s3" | Jmp: "G\Norm s \Jmp j\ (Some (Jump j), s)" \ \cf. 14.16\ | Throw: "\G\Norm s0 \e-\a'\ s1\ \ G\Norm s0 \Throw e\ abupd (throw a') s1" \ \cf. 14.18.1\ | Try: "\G\Norm s0 \c1\ s1; G\s1 \sxalloc\ s2; if G,s2\catch C then G\new_xcpt_var vn s2 \c2\ s3 else s3 = s2\ \ G\Norm s0 \Try c1 Catch(C vn) c2\ s3" \ \cf. 14.18.2\ | Fin: "\G\Norm s0 \c1\ (x1,s1); G\Norm s1 \c2\ s2; s3=(if (\ err. x1=Some (Error err)) then (x1,s1) else abupd (abrupt_if (x1\None) x1) s2) \ \ G\Norm s0 \c1 Finally c2\ s3" \ \cf. 12.4.2, 8.5\ | Init: "\the (class G C) = c; if inited C (globs s0) then s3 = Norm s0 else (G\Norm (init_class_obj G C s0) \(if C = Object then Skip else Init (super c))\ s1 \ G\set_lvars Map.empty s1 \init c\ s2 \ s3 = restore_lvars s1 s2)\ \ G\Norm s0 \Init C\ s3" \ \This class initialisation rule is a little bit inaccurate. Look at the exact sequence: (1) The current class object (the static fields) are initialised (\init_class_obj\), (2) the superclasses are initialised, (3) the static initialiser of the current class is invoked. More precisely we should expect another ordering, namely 2 1 3. But we can't just naively toggle 1 and 2. By calling \init_class_obj\ before initialising the superclasses, we also implicitly record that we have started to initialise the current class (by setting an value for the class object). This becomes crucial for the completeness proof of the axiomatic semantics \AxCompl.thy\. Static initialisation requires an induction on the number of classes not yet initialised (or to be more precise, classes were the initialisation has not yet begun). So we could first assign a dummy value to the class before superclass initialisation and afterwards set the correct values. But as long as we don't take memory overflow into account when allocating class objects, we can leave things as they are for convenience.\ \ \evaluation of expressions\ \ \cf. 15.8.1, 12.4.1\ | NewC: "\G\Norm s0 \Init C\ s1; G\ s1 \halloc (CInst C)\a\ s2\ \ G\Norm s0 \NewC C-\Addr a\ s2" \ \cf. 15.9.1, 12.4.1\ | NewA: "\G\Norm s0 \init_comp_ty T\ s1; G\s1 \e-\i'\ s2; G\abupd (check_neg i') s2 \halloc (Arr T (the_Intg i'))\a\ s3\ \ G\Norm s0 \New T[e]-\Addr a\ s3" \ \cf. 15.15\ | Cast: "\G\Norm s0 \e-\v\ s1; s2 = abupd (raise_if (\G,store s1\v fits T) ClassCast) s1\ \ G\Norm s0 \Cast T e-\v\ s2" \ \cf. 15.19.2\ | Inst: "\G\Norm s0 \e-\v\ s1; b = (v\Null \ G,store s1\v fits RefT T)\ \ G\Norm s0 \e InstOf T-\Bool b\ s1" \ \cf. 15.7.1\ | Lit: "G\Norm s \Lit v-\v\ Norm s" | UnOp: "\G\Norm s0 \e-\v\ s1\ \ G\Norm s0 \UnOp unop e-\(eval_unop unop v)\ s1" | BinOp: "\G\Norm s0 \e1-\v1\ s1; G\s1 \(if need_second_arg binop v1 then (In1l e2) else (In1r Skip)) \\ (In1 v2, s2) \ \ G\Norm s0 \BinOp binop e1 e2-\(eval_binop binop v1 v2)\ s2" \ \cf. 15.10.2\ | Super: "G\Norm s \Super-\val_this s\ Norm s" \ \cf. 15.2\ | Acc: "\G\Norm s0 \va=\(v,f)\ s1\ \ G\Norm s0 \Acc va-\v\ s1" \ \cf. 15.25.1\ | Ass: "\G\Norm s0 \va=\(w,f)\ s1; G\ s1 \e-\v \ s2\ \ G\Norm s0 \va:=e-\v\ assign f v s2" \ \cf. 15.24\ | Cond: "\G\Norm s0 \e0-\b\ s1; G\ s1 \(if the_Bool b then e1 else e2)-\v\ s2\ \ G\Norm s0 \e0 ? e1 : e2-\v\ s2" \ \The interplay of \<^term>\Call\, \<^term>\Methd\ and \<^term>\Body\: Method invocation is split up into these three rules: \begin{itemize} \item [\<^term>\Call\] Calculates the target address and evaluates the arguments of the method, and then performs dynamic or static lookup of the method, corresponding to the call mode. Then the \<^term>\Methd\ rule is evaluated on the calculated declaration class of the method invocation. \item [\<^term>\Methd\] A syntactic bridge for the folded method body. It is used by the axiomatic semantics to add the proper hypothesis for recursive calls of the method. \item [\<^term>\Body\] An extra syntactic entity for the unfolded method body was introduced to properly trigger class initialisation. Without class initialisation we could just evaluate the body statement. \end{itemize}\ \ \cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5\ | Call: "\G\Norm s0 \e-\a'\ s1; G\s1 \args\\vs\ s2; D = invocation_declclass G mode (store s2) a' statT \name=mn,parTs=pTs\; s3=init_lvars G D \name=mn,parTs=pTs\ mode a' vs s2; s3' = check_method_access G accC statT mode \name=mn,parTs=pTs\ a' s3; G\s3' \Methd D \name=mn,parTs=pTs\-\v\ s4\ \ G\Norm s0 \{accC,statT,mode}e\mn({pTs}args)-\v\ (restore_lvars s2 s4)" \ \The accessibility check is after \<^term>\init_lvars\, to keep it simple. \<^term>\init_lvars\ already tests for the absence of a null-pointer reference in case of an instance method invocation.\ | Methd: "\G\Norm s0 \body G D sig-\v\ s1\ \ G\Norm s0 \Methd D sig-\v\ s1" | Body: "\G\Norm s0 \Init D\ s1; G\s1 \c\ s2; s3 = (if (\ l. abrupt s2 = Some (Jump (Break l)) \ abrupt s2 = Some (Jump (Cont l))) then abupd (\ x. Some (Error CrossMethodJump)) s2 else s2)\ \ G\Norm s0 \Body D c-\the (locals (store s2) Result) \abupd (absorb Ret) s3" \ \cf. 14.15, 12.4.1\ \ \We filter out a break/continue in \<^term>\s2\, so that we can proof definite assignment correct, without the need of conformance of the state. By this the different parts of the typesafety proof can be disentangled a little.\ \ \evaluation of variables\ \ \cf. 15.13.1, 15.7.2\ | LVar: "G\Norm s \LVar vn=\lvar vn s\ Norm s" \ \cf. 15.10.1, 12.4.1\ | FVar: "\G\Norm s0 \Init statDeclC\ s1; G\s1 \e-\a\ s2; (v,s2') = fvar statDeclC stat fn a s2; s3 = check_field_access G accC statDeclC fn stat a s2' \ \ G\Norm s0 \{accC,statDeclC,stat}e..fn=\v\ s3" \ \The accessibility check is after \<^term>\fvar\, to keep it simple. \<^term>\fvar\ already tests for the absence of a null-pointer reference in case of an instance field\ \ \cf. 15.12.1, 15.25.1\ | AVar: "\G\ Norm s0 \e1-\a\ s1; G\s1 \e2-\i\ s2; (v,s2') = avar G i a s2\ \ G\Norm s0 \e1.[e2]=\v\ s2'" \ \evaluation of expression lists\ \ \cf. 15.11.4.2\ | Nil: "G\Norm s0 \[]\\[]\ Norm s0" \ \cf. 15.6.4\ | Cons: "\G\Norm s0 \e -\ v \ s1; G\ s1 \es\\vs\ s2\ \ G\Norm s0 \e#es\\v#vs\ s2" (* Rearrangement of premisses: [0,1(Abrupt),2(Skip),8(Jmp),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst), 17(Lit),18(UnOp),19(BinOp),20(Super),21(Acc),3(Expr),5(Comp),25(Methd),26(Body),23(Cond),6(If), 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),22(Ass),10(Try),28(FVar), 29(AVar),24(Call)] *) ML \ ML_Thms.bind_thm ("eval_induct", rearrange_prems [0,1,2,8,4,30,31,27,15,16, 17,18,19,20,21,3,5,25,26,23,6, 7,11,9,13,14,12,22,10,28, 29,24] @{thm eval.induct}) \ declare if_split [split del] if_split_asm [split del] option.split [split del] option.split_asm [split del] inductive_cases halloc_elim_cases: "G\(Some xc,s) \halloc oi\a\ s'" "G\(Norm s) \halloc oi\a\ s'" inductive_cases sxalloc_elim_cases: "G\ Norm s \sxalloc\ s'" "G\(Some (Jump j),s) \sxalloc\ s'" "G\(Some (Error e),s) \sxalloc\ s'" "G\(Some (Xcpt (Loc a )),s) \sxalloc\ s'" "G\(Some (Xcpt (Std xn)),s) \sxalloc\ s'" inductive_cases sxalloc_cases: "G\s \sxalloc\ s'" lemma sxalloc_elim_cases2: "\G\s \sxalloc\ s'; \s. \s' = Norm s\ \ P; \j s. \s' = (Some (Jump j),s)\ \ P; \e s. \s' = (Some (Error e),s)\ \ P; \a s. \s' = (Some (Xcpt (Loc a)),s)\ \ P \ \ P" apply cut_tac apply (erule sxalloc_cases) apply blast+ done declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *) declare split_paired_All [simp del] split_paired_Ex [simp del] setup \map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\ inductive_cases eval_cases: "G\s \t\\ (v, s')" inductive_cases eval_elim_cases [cases set]: "G\(Some xc,s) \t \\ (v, s')" "G\Norm s \In1r Skip \\ (x, s')" "G\Norm s \In1r (Jmp j) \\ (x, s')" "G\Norm s \In1r (l\ c) \\ (x, s')" "G\Norm s \In3 ([]) \\ (v, s')" "G\Norm s \In3 (e#es) \\ (v, s')" "G\Norm s \In1l (Lit w) \\ (v, s')" "G\Norm s \In1l (UnOp unop e) \\ (v, s')" "G\Norm s \In1l (BinOp binop e1 e2) \\ (v, s')" "G\Norm s \In2 (LVar vn) \\ (v, s')" "G\Norm s \In1l (Cast T e) \\ (v, s')" "G\Norm s \In1l (e InstOf T) \\ (v, s')" "G\Norm s \In1l (Super) \\ (v, s')" "G\Norm s \In1l (Acc va) \\ (v, s')" "G\Norm s \In1r (Expr e) \\ (x, s')" "G\Norm s \In1r (c1;; c2) \\ (x, s')" "G\Norm s \In1l (Methd C sig) \\ (x, s')" "G\Norm s \In1l (Body D c) \\ (x, s')" "G\Norm s \In1l (e0 ? e1 : e2) \\ (v, s')" "G\Norm s \In1r (If(e) c1 Else c2) \\ (x, s')" "G\Norm s \In1r (l\ While(e) c) \\ (x, s')" "G\Norm s \In1r (c1 Finally c2) \\ (x, s')" "G\Norm s \In1r (Throw e) \\ (x, s')" "G\Norm s \In1l (NewC C) \\ (v, s')" "G\Norm s \In1l (New T[e]) \\ (v, s')" "G\Norm s \In1l (Ass va e) \\ (v, s')" "G\Norm s \In1r (Try c1 Catch(tn vn) c2) \\ (x, s')" "G\Norm s \In2 ({accC,statDeclC,stat}e..fn) \\ (v, s')" "G\Norm s \In2 (e1.[e2]) \\ (v, s')" "G\Norm s \In1l ({accC,statT,mode}e\mn({pT}p)) \\ (v, s')" "G\Norm s \In1r (Init C) \\ (x, s')" declare not_None_eq [simp] (* IntDef.Zero_def [simp] *) declare split_paired_All [simp] split_paired_Ex [simp] declaration \K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac)))\ declare if_split [split] if_split_asm [split] option.split [split] option.split_asm [split] lemma eval_Inj_elim: "G\s \t\\ (w,s') \ case t of In1 ec \ (case ec of Inl e \ (\v. w = In1 v) | Inr c \ w = \) | In2 e \ (\v. w = In2 v) | In3 e \ (\v. w = In3 v)" apply (erule eval_cases) apply auto apply (induct_tac "t") apply (rename_tac a, induct_tac "a") apply auto done text \The following simplification procedures set up the proper injections of terms and their corresponding values in the evaluation relation: E.g. an expression (injection \<^term>\In1l\ into terms) always evaluates to ordinary values (injection \<^term>\In1\ into generalised values \<^term>\vals\). \ lemma eval_expr_eq: "G\s \In1l t\\ (w, s') = (\v. w=In1 v \ G\s \t-\v \ s')" by (auto, frule eval_Inj_elim, auto) lemma eval_var_eq: "G\s \In2 t\\ (w, s') = (\vf. w=In2 vf \ G\s \t=\vf\ s')" by (auto, frule eval_Inj_elim, auto) lemma eval_exprs_eq: "G\s \In3 t\\ (w, s') = (\vs. w=In3 vs \ G\s \t\\vs\ s')" by (auto, frule eval_Inj_elim, auto) lemma eval_stmt_eq: "G\s \In1r t\\ (w, s') = (w=\ \ G\s \t \ s')" by (auto, frule eval_Inj_elim, auto, frule eval_Inj_elim, auto) simproc_setup eval_expr ("G\s \In1l t\\ (w, s')") = \ - fn _ => fn _ => fn ct => - (case Thm.term_of ct of - (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm eval_expr_eq}))\ - -simproc_setup eval_var ("G\s \In2 t\\ (w, s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm eval_var_eq}))\ + | _ => SOME (mk_meta_eq @{thm eval_expr_eq}))))\ -simproc_setup eval_exprs ("G\s \In3 t\\ (w, s')") = \ - fn _ => fn _ => fn ct => +simproc_setup eval_var ("G\s \In2 t\\ (w, s')") = \ + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm eval_exprs_eq}))\ + | _ => SOME (mk_meta_eq @{thm eval_var_eq}))))\ + +simproc_setup eval_exprs ("G\s \In3 t\\ (w, s')") = \ + K (K (fn ct => + (case Thm.term_of ct of + (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE + | _ => SOME (mk_meta_eq @{thm eval_exprs_eq}))))\ simproc_setup eval_stmt ("G\s \In1r t\\ (w, s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm eval_stmt_eq}))\ + | _ => SOME (mk_meta_eq @{thm eval_stmt_eq}))))\ ML \ ML_Thms.bind_thms ("AbruptIs", sum3_instantiate \<^context> @{thm eval.Abrupt}) \ declare halloc.Abrupt [intro!] eval.Abrupt [intro!] AbruptIs [intro!] text\\Callee\,\InsInitE\, \InsInitV\, \FinA\ are only used in smallstep semantics, not in the bigstep semantics. So their is no valid evaluation of these terms \ lemma eval_Callee: "G\Norm s\Callee l e-\v\ s' = False" proof - { fix s t v s' assume eval: "G\s \t\\ (v,s')" and normal: "normal s" and callee: "t=In1l (Callee l e)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma eval_InsInitE: "G\Norm s\InsInitE c e-\v\ s' = False" proof - { fix s t v s' assume eval: "G\s \t\\ (v,s')" and normal: "normal s" and callee: "t=In1l (InsInitE c e)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma eval_InsInitV: "G\Norm s\InsInitV c w=\v\ s' = False" proof - { fix s t v s' assume eval: "G\s \t\\ (v,s')" and normal: "normal s" and callee: "t=In2 (InsInitV c w)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma eval_FinA: "G\Norm s\FinA a c\ s' = False" proof - { fix s t v s' assume eval: "G\s \t\\ (v,s')" and normal: "normal s" and callee: "t=In1r (FinA a c)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma eval_no_abrupt_lemma: "\s s'. G\s \t\\ (w,s') \ normal s' \ normal s" by (erule eval_cases, auto) lemma eval_no_abrupt: "G\(x,s) \t\\ (w,Norm s') = (x = None \ G\Norm s \t\\ (w,Norm s'))" apply auto apply (frule eval_no_abrupt_lemma, auto)+ done simproc_setup eval_no_abrupt ("G\(x,s) \e\\ (w,Norm s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ (Const (\<^const_name>\Pair\, _) $ (Const (\<^const_name>\None\, _)) $ _) $ _ $ _ $ _) => NONE - | _ => SOME (mk_meta_eq @{thm eval_no_abrupt})) + | _ => SOME (mk_meta_eq @{thm eval_no_abrupt})))) \ lemma eval_abrupt_lemma: "G\s \t\\ (v,s') \ abrupt s=Some xc \ s'= s \ v = undefined3 t" by (erule eval_cases, auto) lemma eval_abrupt: " G\(Some xc,s) \t\\ (w,s') = (s'=(Some xc,s) \ w=undefined3 t \ G\(Some xc,s) \t\\ (undefined3 t,(Some xc,s)))" apply auto apply (frule eval_abrupt_lemma, auto)+ done simproc_setup eval_abrupt ("G\(Some xc,s) \e\\ (w,s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ (Const (\<^const_name>\Pair\, _) $ (Const (\<^const_name>\Some\, _) $ _)$ _)) => NONE - | _ => SOME (mk_meta_eq @{thm eval_abrupt})) + | _ => SOME (mk_meta_eq @{thm eval_abrupt})))) \ lemma LitI: "G\s \Lit v-\(if normal s then v else undefined)\ s" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Lit) lemma SkipI [intro!]: "G\s \Skip\ s" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Skip) lemma ExprI: "G\s \e-\v\ s' \ G\s \Expr e\ s'" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Expr) lemma CompI: "\G\s \c1\ s1; G\s1 \c2\ s2\ \ G\s \c1;; c2\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Comp) lemma CondI: "\s1. \G\s \e-\b\ s1; G\s1 \(if the_Bool b then e1 else e2)-\v\ s2\ \ G\s \e ? e1 : e2-\(if normal s1 then v else undefined)\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Cond) lemma IfI: "\G\s \e-\v\ s1; G\s1 \(if the_Bool v then c1 else c2)\ s2\ \ G\s \If(e) c1 Else c2\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.If) lemma MethdI: "G\s \body G C sig-\v\ s' \ G\s \Methd C sig-\v\ s'" apply (case_tac "s", case_tac "a = None") by (auto intro!: eval.Methd) lemma eval_Call: "\G\Norm s0 \e-\a'\ s1; G\s1 \ps\\pvs\ s2; D = invocation_declclass G mode (store s2) a' statT \name=mn,parTs=pTs\; s3 = init_lvars G D \name=mn,parTs=pTs\ mode a' pvs s2; s3' = check_method_access G accC statT mode \name=mn,parTs=pTs\ a' s3; G\s3'\Methd D \name=mn,parTs=pTs\-\ v\ s4; s4' = restore_lvars s2 s4\ \ G\Norm s0 \{accC,statT,mode}e\mn({pTs}ps)-\v\ s4'" apply (drule eval.Call, assumption) apply (rule HOL.refl) apply simp+ done lemma eval_Init: "\if inited C (globs s0) then s3 = Norm s0 else G\Norm (init_class_obj G C s0) \(if C = Object then Skip else Init (super (the (class G C))))\ s1 \ G\set_lvars Map.empty s1 \(init (the (class G C)))\ s2 \ s3 = restore_lvars s1 s2\ \ G\Norm s0 \Init C\ s3" apply (rule eval.Init) apply auto done lemma init_done: "initd C s \ G\s \Init C\ s" apply (case_tac "s", simp) apply (case_tac "a") apply safe apply (rule eval_Init) apply auto done lemma eval_StatRef: "G\s \StatRef rt-\(if abrupt s=None then Null else undefined)\ s" apply (case_tac "s", simp) apply (case_tac "a = None") apply (auto del: eval.Abrupt intro!: eval.intros) done lemma SkipD [dest!]: "G\s \Skip\ s' \ s' = s" apply (erule eval_cases) by auto lemma Skip_eq [simp]: "G\s \Skip\ s' = (s = s')" by auto (*unused*) lemma init_retains_locals [rule_format (no_asm)]: "G\s \t\\ (w,s') \ (\C. t=In1r (Init C) \ locals (store s) = locals (store s'))" apply (erule eval.induct) apply (simp (no_asm_use) split del: if_split_asm option.split_asm)+ apply auto done lemma halloc_xcpt [dest!]: "\s'. G\(Some xc,s) \halloc oi\a\ s' \ s'=(Some xc,s)" apply (erule_tac halloc_elim_cases) by auto (* G\(x,(h,l)) \e\v\ (x',(h',l'))) \ l This = l' This" G\(x,(h,l)) \s \ (x',(h',l'))) \ l This = l' This" *) lemma eval_Methd: "G\s \In1l(body G C sig)\\ (w,s') \ G\s \In1l(Methd C sig)\\ (w,s')" apply (case_tac "s") apply (case_tac "a") apply clarsimp+ apply (erule eval.Methd) apply (drule eval_abrupt_lemma) apply force done lemma eval_Body: "\G\Norm s0 \Init D\ s1; G\s1 \c\ s2; res=the (locals (store s2) Result); s3 = (if (\ l. abrupt s2 = Some (Jump (Break l)) \ abrupt s2 = Some (Jump (Cont l))) then abupd (\ x. Some (Error CrossMethodJump)) s2 else s2); s4=abupd (absorb Ret) s3\ \ G\Norm s0 \Body D c-\res\s4" by (auto elim: eval.Body) lemma eval_binop_arg2_indep: "\ need_second_arg binop v1 \ eval_binop binop v1 x = eval_binop binop v1 y" by (cases binop) (simp_all add: need_second_arg_def) lemma eval_BinOp_arg2_indepI: assumes eval_e1: "G\Norm s0 \e1-\v1\ s1" and no_need: "\ need_second_arg binop v1" shows "G\Norm s0 \BinOp binop e1 e2-\(eval_binop binop v1 v2)\ s1" (is "?EvalBinOp v2") proof - from eval_e1 have "?EvalBinOp Unit" by (rule eval.BinOp) (simp add: no_need) moreover from no_need have "eval_binop binop v1 Unit = eval_binop binop v1 v2" by (simp add: eval_binop_arg2_indep) ultimately show ?thesis by simp qed subsubsection "single valued" lemma unique_halloc [rule_format (no_asm)]: "G\s \halloc oi\a \ s' \ G\s \halloc oi\a' \ s'' \ a' = a \ s'' = s'" apply (erule halloc.induct) apply (auto elim!: halloc_elim_cases split del: if_split if_split_asm) apply (drule trans [THEN sym], erule sym) defer apply (drule trans [THEN sym], erule sym) apply auto done lemma single_valued_halloc: "single_valued {((s,oi),(a,s')). G\s \halloc oi\a \ s'}" apply (unfold single_valued_def) by (clarsimp, drule (1) unique_halloc, auto) lemma unique_sxalloc [rule_format (no_asm)]: "G\s \sxalloc\ s' \ G\s \sxalloc\ s'' \ s'' = s'" apply (erule sxalloc.induct) apply (auto dest: unique_halloc elim!: sxalloc_elim_cases split del: if_split if_split_asm) done lemma single_valued_sxalloc: "single_valued {(s,s'). G\s \sxalloc\ s'}" apply (unfold single_valued_def) apply (blast dest: unique_sxalloc) done lemma split_pairD: "(x,y) = p \ x = fst p & y = snd p" by auto lemma unique_eval [rule_format (no_asm)]: "G\s \t\\ (w, s') \ (\w' s''. G\s \t\\ (w', s'') \ w' = w \ s'' = s')" apply (erule eval_induct) apply (tactic \ALLGOALS (EVERY' [strip_tac \<^context>, rotate_tac ~1, eresolve_tac \<^context> @{thms eval_elim_cases}])\) (* 31 subgoals *) prefer 28 (* Try *) apply (simp (no_asm_use) only: split: if_split_asm) (* 34 subgoals *) prefer 30 (* Init *) apply (case_tac "inited C (globs s0)", (simp only: if_True if_False simp_thms)+) prefer 26 (* While *) apply (simp (no_asm_use) only: split: if_split_asm, blast) (* 36 subgoals *) apply (blast dest: unique_sxalloc unique_halloc split_pairD)+ done (* unused *) lemma single_valued_eval: "single_valued {((s, t), (v, s')). G\s \t\\ (v, s')}" apply (unfold single_valued_def) by (clarify, drule (1) unique_eval, auto) end diff --git a/src/HOL/Bali/Evaln.thy b/src/HOL/Bali/Evaln.thy --- a/src/HOL/Bali/Evaln.thy +++ b/src/HOL/Bali/Evaln.thy @@ -1,832 +1,832 @@ (* Title: HOL/Bali/Evaln.thy Author: David von Oheimb and Norbert Schirmer *) subsection \Operational evaluation (big-step) semantics of Java expressions and statements \ theory Evaln imports TypeSafe begin text \ Variant of \<^term>\eval\ relation with counter for bounded recursive depth. In principal \<^term>\evaln\ could replace \<^term>\eval\. Validity of the axiomatic semantics builds on \<^term>\evaln\. For recursive method calls the axiomatic semantics rule assumes the method ok to derive a proof for the body. To prove the method rule sound we need to perform induction on the recursion depth. For the completeness proof of the axiomatic semantics the notion of the most general formula is used. The most general formula right now builds on the ordinary evaluation relation \<^term>\eval\. So sometimes we have to switch between \<^term>\evaln\ and \<^term>\eval\ and vice versa. To make this switch easy \<^term>\evaln\ also does all the technical accessibility tests \<^term>\check_field_access\ and \<^term>\check_method_access\ like \<^term>\eval\. If it would omit them \<^term>\evaln\ and \<^term>\eval\ would only be equivalent for welltyped, and definitely assigned terms. \ inductive evaln :: "[prog, state, term, nat, vals, state] \ bool" ("_\_ \_\\_\ '(_, _')" [61,61,80,61,0,0] 60) and evarn :: "[prog, state, var, vvar, nat, state] \ bool" ("_\_ \_=\_\_\ _" [61,61,90,61,61,61] 60) and eval_n:: "[prog, state, expr, val, nat, state] \ bool" ("_\_ \_-\_\_\ _" [61,61,80,61,61,61] 60) and evalsn :: "[prog, state, expr list, val list, nat, state] \ bool" ("_\_ \_\\_\_\ _" [61,61,61,61,61,61] 60) and execn :: "[prog, state, stmt, nat, state] \ bool" ("_\_ \_\_\ _" [61,61,65, 61,61] 60) for G :: prog where "G\s \c \n\ s' \ G\s \In1r c\\n\ (\ , s')" | "G\s \e-\v \n\ s' \ G\s \In1l e\\n\ (In1 v , s')" | "G\s \e=\vf \n\ s' \ G\s \In2 e\\n\ (In2 vf, s')" | "G\s \e\\v \n\ s' \ G\s \In3 e\\n\ (In3 v , s')" \ \propagation of abrupt completion\ | Abrupt: "G\(Some xc,s) \t\\n\ (undefined3 t,(Some xc,s))" \ \evaluation of variables\ | LVar: "G\Norm s \LVar vn=\lvar vn s\n\ Norm s" | FVar: "\G\Norm s0 \Init statDeclC\n\ s1; G\s1 \e-\a\n\ s2; (v,s2') = fvar statDeclC stat fn a s2; s3 = check_field_access G accC statDeclC fn stat a s2'\ \ G\Norm s0 \{accC,statDeclC,stat}e..fn=\v\n\ s3" | AVar: "\G\ Norm s0 \e1-\a\n\ s1 ; G\s1 \e2-\i\n\ s2; (v,s2') = avar G i a s2\ \ G\Norm s0 \e1.[e2]=\v\n\ s2'" \ \evaluation of expressions\ | NewC: "\G\Norm s0 \Init C\n\ s1; G\ s1 \halloc (CInst C)\a\ s2\ \ G\Norm s0 \NewC C-\Addr a\n\ s2" | NewA: "\G\Norm s0 \init_comp_ty T\n\ s1; G\s1 \e-\i'\n\ s2; G\abupd (check_neg i') s2 \halloc (Arr T (the_Intg i'))\a\ s3\ \ G\Norm s0 \New T[e]-\Addr a\n\ s3" | Cast: "\G\Norm s0 \e-\v\n\ s1; s2 = abupd (raise_if (\G,snd s1\v fits T) ClassCast) s1\ \ G\Norm s0 \Cast T e-\v\n\ s2" | Inst: "\G\Norm s0 \e-\v\n\ s1; b = (v\Null \ G,store s1\v fits RefT T)\ \ G\Norm s0 \e InstOf T-\Bool b\n\ s1" | Lit: "G\Norm s \Lit v-\v\n\ Norm s" | UnOp: "\G\Norm s0 \e-\v\n\ s1\ \ G\Norm s0 \UnOp unop e-\(eval_unop unop v)\n\ s1" | BinOp: "\G\Norm s0 \e1-\v1\n\ s1; G\s1 \(if need_second_arg binop v1 then (In1l e2) else (In1r Skip)) \\n\ (In1 v2,s2)\ \ G\Norm s0 \BinOp binop e1 e2-\(eval_binop binop v1 v2)\n\ s2" | Super: "G\Norm s \Super-\val_this s\n\ Norm s" | Acc: "\G\Norm s0 \va=\(v,f)\n\ s1\ \ G\Norm s0 \Acc va-\v\n\ s1" | Ass: "\G\Norm s0 \va=\(w,f)\n\ s1; G\ s1 \e-\v \n\ s2\ \ G\Norm s0 \va:=e-\v\n\ assign f v s2" | Cond: "\G\Norm s0 \e0-\b\n\ s1; G\ s1 \(if the_Bool b then e1 else e2)-\v\n\ s2\ \ G\Norm s0 \e0 ? e1 : e2-\v\n\ s2" | Call: "\G\Norm s0 \e-\a'\n\ s1; G\s1 \args\\vs\n\ s2; D = invocation_declclass G mode (store s2) a' statT \name=mn,parTs=pTs\; s3=init_lvars G D \name=mn,parTs=pTs\ mode a' vs s2; s3' = check_method_access G accC statT mode \name=mn,parTs=pTs\ a' s3; G\s3'\Methd D \name=mn,parTs=pTs\-\v\n\ s4 \ \ G\Norm s0 \{accC,statT,mode}e\mn({pTs}args)-\v\n\ (restore_lvars s2 s4)" | Methd:"\G\Norm s0 \body G D sig-\v\n\ s1\ \ G\Norm s0 \Methd D sig-\v\Suc n\ s1" | Body: "\G\Norm s0\Init D\n\ s1; G\s1 \c\n\ s2; s3 = (if (\ l. abrupt s2 = Some (Jump (Break l)) \ abrupt s2 = Some (Jump (Cont l))) then abupd (\ x. Some (Error CrossMethodJump)) s2 else s2)\\ G\Norm s0 \Body D c -\the (locals (store s2) Result)\n\abupd (absorb Ret) s3" \ \evaluation of expression lists\ | Nil: "G\Norm s0 \[]\\[]\n\ Norm s0" | Cons: "\G\Norm s0 \e -\ v \n\ s1; G\ s1 \es\\vs\n\ s2\ \ G\Norm s0 \e#es\\v#vs\n\ s2" \ \execution of statements\ | Skip: "G\Norm s \Skip\n\ Norm s" | Expr: "\G\Norm s0 \e-\v\n\ s1\ \ G\Norm s0 \Expr e\n\ s1" | Lab: "\G\Norm s0 \c \n\ s1\ \ G\Norm s0 \l\ c\n\ abupd (absorb l) s1" | Comp: "\G\Norm s0 \c1 \n\ s1; G\ s1 \c2 \n\ s2\ \ G\Norm s0 \c1;; c2\n\ s2" | If: "\G\Norm s0 \e-\b\n\ s1; G\ s1\(if the_Bool b then c1 else c2)\n\ s2\ \ G\Norm s0 \If(e) c1 Else c2 \n\ s2" | Loop: "\G\Norm s0 \e-\b\n\ s1; if the_Bool b then (G\s1 \c\n\ s2 \ G\(abupd (absorb (Cont l)) s2) \l\ While(e) c\n\ s3) else s3 = s1\ \ G\Norm s0 \l\ While(e) c\n\ s3" | Jmp: "G\Norm s \Jmp j\n\ (Some (Jump j), s)" | Throw:"\G\Norm s0 \e-\a'\n\ s1\ \ G\Norm s0 \Throw e\n\ abupd (throw a') s1" | Try: "\G\Norm s0 \c1\n\ s1; G\s1 \sxalloc\ s2; if G,s2\catch tn then G\new_xcpt_var vn s2 \c2\n\ s3 else s3 = s2\ \ G\Norm s0 \Try c1 Catch(tn vn) c2\n\ s3" | Fin: "\G\Norm s0 \c1\n\ (x1,s1); G\Norm s1 \c2\n\ s2; s3=(if (\ err. x1=Some (Error err)) then (x1,s1) else abupd (abrupt_if (x1\None) x1) s2)\ \ G\Norm s0 \c1 Finally c2\n\ s3" | Init: "\the (class G C) = c; if inited C (globs s0) then s3 = Norm s0 else (G\Norm (init_class_obj G C s0) \(if C = Object then Skip else Init (super c))\n\ s1 \ G\set_lvars Map.empty s1 \init c\n\ s2 \ s3 = restore_lvars s1 s2)\ \ G\Norm s0 \Init C\n\ s3" monos if_bool_eq_conj declare if_split [split del] if_split_asm [split del] option.split [split del] option.split_asm [split del] not_None_eq [simp del] split_paired_All [simp del] split_paired_Ex [simp del] setup \map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\ inductive_cases evaln_cases: "G\s \t\\n\ (v, s')" inductive_cases evaln_elim_cases: "G\(Some xc, s) \t \\n\ (v, s')" "G\Norm s \In1r Skip \\n\ (x, s')" "G\Norm s \In1r (Jmp j) \\n\ (x, s')" "G\Norm s \In1r (l\ c) \\n\ (x, s')" "G\Norm s \In3 ([]) \\n\ (v, s')" "G\Norm s \In3 (e#es) \\n\ (v, s')" "G\Norm s \In1l (Lit w) \\n\ (v, s')" "G\Norm s \In1l (UnOp unop e) \\n\ (v, s')" "G\Norm s \In1l (BinOp binop e1 e2) \\n\ (v, s')" "G\Norm s \In2 (LVar vn) \\n\ (v, s')" "G\Norm s \In1l (Cast T e) \\n\ (v, s')" "G\Norm s \In1l (e InstOf T) \\n\ (v, s')" "G\Norm s \In1l (Super) \\n\ (v, s')" "G\Norm s \In1l (Acc va) \\n\ (v, s')" "G\Norm s \In1r (Expr e) \\n\ (x, s')" "G\Norm s \In1r (c1;; c2) \\n\ (x, s')" "G\Norm s \In1l (Methd C sig) \\n\ (x, s')" "G\Norm s \In1l (Body D c) \\n\ (x, s')" "G\Norm s \In1l (e0 ? e1 : e2) \\n\ (v, s')" "G\Norm s \In1r (If(e) c1 Else c2) \\n\ (x, s')" "G\Norm s \In1r (l\ While(e) c) \\n\ (x, s')" "G\Norm s \In1r (c1 Finally c2) \\n\ (x, s')" "G\Norm s \In1r (Throw e) \\n\ (x, s')" "G\Norm s \In1l (NewC C) \\n\ (v, s')" "G\Norm s \In1l (New T[e]) \\n\ (v, s')" "G\Norm s \In1l (Ass va e) \\n\ (v, s')" "G\Norm s \In1r (Try c1 Catch(tn vn) c2) \\n\ (x, s')" "G\Norm s \In2 ({accC,statDeclC,stat}e..fn) \\n\ (v, s')" "G\Norm s \In2 (e1.[e2]) \\n\ (v, s')" "G\Norm s \In1l ({accC,statT,mode}e\mn({pT}p)) \\n\ (v, s')" "G\Norm s \In1r (Init C) \\n\ (x, s')" declare if_split [split] if_split_asm [split] option.split [split] option.split_asm [split] not_None_eq [simp] split_paired_All [simp] split_paired_Ex [simp] declaration \K (Simplifier.map_ss (fn ss => ss addloop ("split_all_tac", split_all_tac)))\ lemma evaln_Inj_elim: "G\s \t\\n\ (w,s') \ case t of In1 ec \ (case ec of Inl e \ (\v. w = In1 v) | Inr c \ w = \) | In2 e \ (\v. w = In2 v) | In3 e \ (\v. w = In3 v)" apply (erule evaln_cases , auto) apply (induct_tac "t") apply (rename_tac a, induct_tac "a") apply auto done text \The following simplification procedures set up the proper injections of terms and their corresponding values in the evaluation relation: E.g. an expression (injection \<^term>\In1l\ into terms) always evaluates to ordinary values (injection \<^term>\In1\ into generalised values \<^term>\vals\). \ lemma evaln_expr_eq: "G\s \In1l t\\n\ (w, s') = (\v. w=In1 v \ G\s \t-\v \n\ s')" by (auto, frule evaln_Inj_elim, auto) lemma evaln_var_eq: "G\s \In2 t\\n\ (w, s') = (\vf. w=In2 vf \ G\s \t=\vf\n\ s')" by (auto, frule evaln_Inj_elim, auto) lemma evaln_exprs_eq: "G\s \In3 t\\n\ (w, s') = (\vs. w=In3 vs \ G\s \t\\vs\n\ s')" by (auto, frule evaln_Inj_elim, auto) lemma evaln_stmt_eq: "G\s \In1r t\\n\ (w, s') = (w=\ \ G\s \t \n\ s')" by (auto, frule evaln_Inj_elim, auto, frule evaln_Inj_elim, auto) simproc_setup evaln_expr ("G\s \In1l t\\n\ (w, s')") = \ - fn _ => fn _ => fn ct => - (case Thm.term_of ct of - (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm evaln_expr_eq}))\ - -simproc_setup evaln_var ("G\s \In2 t\\n\ (w, s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm evaln_var_eq}))\ + | _ => SOME (mk_meta_eq @{thm evaln_expr_eq}))))\ -simproc_setup evaln_exprs ("G\s \In3 t\\n\ (w, s')") = \ - fn _ => fn _ => fn ct => +simproc_setup evaln_var ("G\s \In2 t\\n\ (w, s')") = \ + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq}))\ + | _ => SOME (mk_meta_eq @{thm evaln_var_eq}))))\ + +simproc_setup evaln_exprs ("G\s \In3 t\\n\ (w, s')") = \ + K (K (fn ct => + (case Thm.term_of ct of + (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE + | _ => SOME (mk_meta_eq @{thm evaln_exprs_eq}))))\ simproc_setup evaln_stmt ("G\s \In1r t\\n\ (w, s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ (Const _ $ _) $ _) => NONE - | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq}))\ + | _ => SOME (mk_meta_eq @{thm evaln_stmt_eq}))))\ ML \ML_Thms.bind_thms ("evaln_AbruptIs", sum3_instantiate \<^context> @{thm evaln.Abrupt})\ declare evaln_AbruptIs [intro!] lemma evaln_Callee: "G\Norm s\In1l (Callee l e)\\n\ (v,s') = False" proof - { fix s t v s' assume eval: "G\s \t\\n\ (v,s')" and normal: "normal s" and callee: "t=In1l (Callee l e)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma evaln_InsInitE: "G\Norm s\In1l (InsInitE c e)\\n\ (v,s') = False" proof - { fix s t v s' assume eval: "G\s \t\\n\ (v,s')" and normal: "normal s" and callee: "t=In1l (InsInitE c e)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma evaln_InsInitV: "G\Norm s\In2 (InsInitV c w)\\n\ (v,s') = False" proof - { fix s t v s' assume eval: "G\s \t\\n\ (v,s')" and normal: "normal s" and callee: "t=In2 (InsInitV c w)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma evaln_FinA: "G\Norm s\In1r (FinA a c)\\n\ (v,s') = False" proof - { fix s t v s' assume eval: "G\s \t\\n\ (v,s')" and normal: "normal s" and callee: "t=In1r (FinA a c)" then have "False" by induct auto } then show ?thesis by (cases s') fastforce qed lemma evaln_abrupt_lemma: "G\s \e\\n\ (v,s') \ fst s = Some xc \ s' = s \ v = undefined3 e" apply (erule evaln_cases , auto) done lemma evaln_abrupt: "\s'. G\(Some xc,s) \e\\n\ (w,s') = (s' = (Some xc,s) \ w=undefined3 e \ G\(Some xc,s) \e\\n\ (undefined3 e,(Some xc,s)))" apply auto apply (frule evaln_abrupt_lemma, auto)+ done simproc_setup evaln_abrupt ("G\(Some xc,s) \e\\n\ (w,s')") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ _ $ _ $ (Const (\<^const_name>\Pair\, _) $ (Const (\<^const_name>\Some\,_) $ _)$ _)) => NONE - | _ => SOME (mk_meta_eq @{thm evaln_abrupt})) + | _ => SOME (mk_meta_eq @{thm evaln_abrupt})))) \ lemma evaln_LitI: "G\s \Lit v-\(if normal s then v else undefined)\n\ s" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.Lit) lemma CondI: "\s1. \G\s \e-\b\n\ s1; G\s1 \(if the_Bool b then e1 else e2)-\v\n\ s2\ \ G\s \e ? e1 : e2-\(if normal s1 then v else undefined)\n\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.Cond) lemma evaln_SkipI [intro!]: "G\s \Skip\n\ s" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.Skip) lemma evaln_ExprI: "G\s \e-\v\n\ s' \ G\s \Expr e\n\ s'" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.Expr) lemma evaln_CompI: "\G\s \c1\n\ s1; G\s1 \c2\n\ s2\ \ G\s \c1;; c2\n\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.Comp) lemma evaln_IfI: "\G\s \e-\v\n\ s1; G\s1 \(if the_Bool v then c1 else c2)\n\ s2\ \ G\s \If(e) c1 Else c2\n\ s2" apply (case_tac "s", case_tac "a = None") by (auto intro!: evaln.If) lemma evaln_SkipD [dest!]: "G\s \Skip\n\ s' \ s' = s" by (erule evaln_cases, auto) lemma evaln_Skip_eq [simp]: "G\s \Skip\n\ s' = (s = s')" apply auto done subsubsection \evaln implies eval\ lemma evaln_eval: assumes evaln: "G\s0 \t\\n\ (v,s1)" shows "G\s0 \t\\ (v,s1)" using evaln proof (induct) case (Loop s0 e b n s1 c s2 l s3) note \G\Norm s0 \e-\b\ s1\ moreover have "if the_Bool b then (G\s1 \c\ s2) \ G\abupd (absorb (Cont l)) s2 \l\ While(e) c\ s3 else s3 = s1" using Loop.hyps by simp ultimately show ?case by (rule eval.Loop) next case (Try s0 c1 n s1 s2 C vn c2 s3) note \G\Norm s0 \c1\ s1\ moreover note \G\s1 \sxalloc\ s2\ moreover have "if G,s2\catch C then G\new_xcpt_var vn s2 \c2\ s3 else s3 = s2" using Try.hyps by simp ultimately show ?case by (rule eval.Try) next case (Init C c s0 s3 n s1 s2) note \the (class G C) = c\ moreover have "if inited C (globs s0) then s3 = Norm s0 else G\Norm ((init_class_obj G C) s0) \(if C = Object then Skip else Init (super c))\ s1 \ G\(set_lvars Map.empty) s1 \init c\ s2 \ s3 = (set_lvars (locals (store s1))) s2" using Init.hyps by simp ultimately show ?case by (rule eval.Init) qed (rule eval.intros,(assumption+ | assumption?))+ lemma Suc_le_D_lemma: "\Suc n <= m'; (\m. n <= m \ P (Suc m)) \ \ P m'" apply (frule Suc_le_D) apply fast done lemma evaln_nonstrict [rule_format (no_asm), elim]: "G\s \t\\n\ (w, s') \ \m. n\m \ G\s \t\\m\ (w, s')" apply (erule evaln.induct) apply (tactic \ALLGOALS (EVERY' [strip_tac \<^context>, TRY o eresolve_tac \<^context> @{thms Suc_le_D_lemma}, REPEAT o smp_tac \<^context> 1, resolve_tac \<^context> @{thms evaln.intros} THEN_ALL_NEW TRY o assume_tac \<^context>])\) (* 3 subgoals *) apply (auto split del: if_split) done lemmas evaln_nonstrict_Suc = evaln_nonstrict [OF _ le_refl [THEN le_SucI]] lemma evaln_max2: "\G\s1 \t1\\n1\ (w1, s1'); G\s2 \t2\\n2\ (w2, s2')\ \ G\s1 \t1\\max n1 n2\ (w1, s1') \ G\s2 \t2\\max n1 n2\ (w2, s2')" by (fast intro: max.cobounded1 max.cobounded2) corollary evaln_max2E [consumes 2]: "\G\s1 \t1\\n1\ (w1, s1'); G\s2 \t2\\n2\ (w2, s2'); \G\s1 \t1\\max n1 n2\ (w1, s1');G\s2 \t2\\max n1 n2\ (w2, s2') \ \ P \ \ P" by (drule (1) evaln_max2) simp lemma evaln_max3: "\G\s1 \t1\\n1\ (w1, s1'); G\s2 \t2\\n2\ (w2, s2'); G\s3 \t3\\n3\ (w3, s3')\ \ G\s1 \t1\\max (max n1 n2) n3\ (w1, s1') \ G\s2 \t2\\max (max n1 n2) n3\ (w2, s2') \ G\s3 \t3\\max (max n1 n2) n3\ (w3, s3')" apply (drule (1) evaln_max2, erule thin_rl) apply (fast intro!: max.cobounded1 max.cobounded2) done corollary evaln_max3E: "\G\s1 \t1\\n1\ (w1, s1'); G\s2 \t2\\n2\ (w2, s2'); G\s3 \t3\\n3\ (w3, s3'); \G\s1 \t1\\max (max n1 n2) n3\ (w1, s1'); G\s2 \t2\\max (max n1 n2) n3\ (w2, s2'); G\s3 \t3\\max (max n1 n2) n3\ (w3, s3') \ \ P \ \ P" by (drule (2) evaln_max3) simp lemma le_max3I1: "(n2::nat) \ max n1 (max n2 n3)" proof - have "n2 \ max n2 n3" by (rule max.cobounded1) also have "max n2 n3 \ max n1 (max n2 n3)" by (rule max.cobounded2) finally show ?thesis . qed lemma le_max3I2: "(n3::nat) \ max n1 (max n2 n3)" proof - have "n3 \ max n2 n3" by (rule max.cobounded2) also have "max n2 n3 \ max n1 (max n2 n3)" by (rule max.cobounded2) finally show ?thesis . qed declare [[simproc del: wt_expr wt_var wt_exprs wt_stmt]] subsubsection \eval implies evaln\ lemma eval_evaln: assumes eval: "G\s0 \t\\ (v,s1)" shows "\n. G\s0 \t\\n\ (v,s1)" using eval proof (induct) case (Abrupt xc s t) obtain n where "G\(Some xc, s) \t\\n\ (undefined3 t, (Some xc, s))" by (iprover intro: evaln.Abrupt) then show ?case .. next case Skip show ?case by (blast intro: evaln.Skip) next case (Expr s0 e v s1) then obtain n where "G\Norm s0 \e-\v\n\ s1" by (iprover) then have "G\Norm s0 \Expr e\n\ s1" by (rule evaln.Expr) then show ?case .. next case (Lab s0 c s1 l) then obtain n where "G\Norm s0 \c\n\ s1" by (iprover) then have "G\Norm s0 \l\ c\n\ abupd (absorb l) s1" by (rule evaln.Lab) then show ?case .. next case (Comp s0 c1 s1 c2 s2) then obtain n1 n2 where "G\Norm s0 \c1\n1\ s1" "G\s1 \c2\n2\ s2" by (iprover) then have "G\Norm s0 \c1;; c2\max n1 n2\ s2" by (blast intro: evaln.Comp dest: evaln_max2 ) then show ?case .. next case (If s0 e b s1 c1 c2 s2) then obtain n1 n2 where "G\Norm s0 \e-\b\n1\ s1" "G\s1 \(if the_Bool b then c1 else c2)\n2\ s2" by (iprover) then have "G\Norm s0 \If(e) c1 Else c2\max n1 n2\ s2" by (blast intro: evaln.If dest: evaln_max2) then show ?case .. next case (Loop s0 e b s1 c s2 l s3) from Loop.hyps obtain n1 where "G\Norm s0 \e-\b\n1\ s1" by (iprover) moreover from Loop.hyps obtain n2 where "if the_Bool b then (G\s1 \c\n2\ s2 \ G\(abupd (absorb (Cont l)) s2)\l\ While(e) c\n2\ s3) else s3 = s1" by simp (iprover intro: evaln_nonstrict max.cobounded1 max.cobounded2) ultimately have "G\Norm s0 \l\ While(e) c\max n1 n2\ s3" apply - apply (rule evaln.Loop) apply (iprover intro: evaln_nonstrict intro: max.cobounded1) apply (auto intro: evaln_nonstrict intro: max.cobounded2) done then show ?case .. next case (Jmp s j) fix n have "G\Norm s \Jmp j\n\ (Some (Jump j), s)" by (rule evaln.Jmp) then show ?case .. next case (Throw s0 e a s1) then obtain n where "G\Norm s0 \e-\a\n\ s1" by (iprover) then have "G\Norm s0 \Throw e\n\ abupd (throw a) s1" by (rule evaln.Throw) then show ?case .. next case (Try s0 c1 s1 s2 catchC vn c2 s3) from Try.hyps obtain n1 where "G\Norm s0 \c1\n1\ s1" by (iprover) moreover note sxalloc = \G\s1 \sxalloc\ s2\ moreover from Try.hyps obtain n2 where "if G,s2\catch catchC then G\new_xcpt_var vn s2 \c2\n2\ s3 else s3 = s2" by fastforce ultimately have "G\Norm s0 \Try c1 Catch(catchC vn) c2\max n1 n2\ s3" by (auto intro!: evaln.Try max.cobounded1 max.cobounded2) then show ?case .. next case (Fin s0 c1 x1 s1 c2 s2 s3) from Fin obtain n1 n2 where "G\Norm s0 \c1\n1\ (x1, s1)" "G\Norm s1 \c2\n2\ s2" by iprover moreover note s3 = \s3 = (if \err. x1 = Some (Error err) then (x1, s1) else abupd (abrupt_if (x1 \ None) x1) s2)\ ultimately have "G\Norm s0 \c1 Finally c2\max n1 n2\ s3" by (blast intro: evaln.Fin dest: evaln_max2) then show ?case .. next case (Init C c s0 s3 s1 s2) note cls = \the (class G C) = c\ moreover from Init.hyps obtain n where "if inited C (globs s0) then s3 = Norm s0 else (G\Norm (init_class_obj G C s0) \(if C = Object then Skip else Init (super c))\n\ s1 \ G\set_lvars Map.empty s1 \init c\n\ s2 \ s3 = restore_lvars s1 s2)" by (auto intro: evaln_nonstrict max.cobounded1 max.cobounded2) ultimately have "G\Norm s0 \Init C\n\ s3" by (rule evaln.Init) then show ?case .. next case (NewC s0 C s1 a s2) then obtain n where "G\Norm s0 \Init C\n\ s1" by (iprover) with NewC have "G\Norm s0 \NewC C-\Addr a\n\ s2" by (iprover intro: evaln.NewC) then show ?case .. next case (NewA s0 T s1 e i s2 a s3) then obtain n1 n2 where "G\Norm s0 \init_comp_ty T\n1\ s1" "G\s1 \e-\i\n2\ s2" by (iprover) moreover note \G\abupd (check_neg i) s2 \halloc Arr T (the_Intg i)\a\ s3\ ultimately have "G\Norm s0 \New T[e]-\Addr a\max n1 n2\ s3" by (blast intro: evaln.NewA dest: evaln_max2) then show ?case .. next case (Cast s0 e v s1 s2 castT) then obtain n where "G\Norm s0 \e-\v\n\ s1" by (iprover) moreover note \s2 = abupd (raise_if (\ G,snd s1\v fits castT) ClassCast) s1\ ultimately have "G\Norm s0 \Cast castT e-\v\n\ s2" by (rule evaln.Cast) then show ?case .. next case (Inst s0 e v s1 b T) then obtain n where "G\Norm s0 \e-\v\n\ s1" by (iprover) moreover note \b = (v \ Null \ G,snd s1\v fits RefT T)\ ultimately have "G\Norm s0 \e InstOf T-\Bool b\n\ s1" by (rule evaln.Inst) then show ?case .. next case (Lit s v) fix n have "G\Norm s \Lit v-\v\n\ Norm s" by (rule evaln.Lit) then show ?case .. next case (UnOp s0 e v s1 unop) then obtain n where "G\Norm s0 \e-\v\n\ s1" by (iprover) hence "G\Norm s0 \UnOp unop e-\eval_unop unop v\n\ s1" by (rule evaln.UnOp) then show ?case .. next case (BinOp s0 e1 v1 s1 binop e2 v2 s2) then obtain n1 n2 where "G\Norm s0 \e1-\v1\n1\ s1" "G\s1 \(if need_second_arg binop v1 then In1l e2 else In1r Skip)\\n2\ (In1 v2, s2)" by (iprover) hence "G\Norm s0 \BinOp binop e1 e2-\(eval_binop binop v1 v2)\max n1 n2 \ s2" by (blast intro!: evaln.BinOp dest: evaln_max2) then show ?case .. next case (Super s ) fix n have "G\Norm s \Super-\val_this s\n\ Norm s" by (rule evaln.Super) then show ?case .. next case (Acc s0 va v f s1) then obtain n where "G\Norm s0 \va=\(v, f)\n\ s1" by (iprover) then have "G\Norm s0 \Acc va-\v\n\ s1" by (rule evaln.Acc) then show ?case .. next case (Ass s0 var w f s1 e v s2) then obtain n1 n2 where "G\Norm s0 \var=\(w, f)\n1\ s1" "G\s1 \e-\v\n2\ s2" by (iprover) then have "G\Norm s0 \var:=e-\v\max n1 n2\ assign f v s2" by (blast intro: evaln.Ass dest: evaln_max2) then show ?case .. next case (Cond s0 e0 b s1 e1 e2 v s2) then obtain n1 n2 where "G\Norm s0 \e0-\b\n1\ s1" "G\s1 \(if the_Bool b then e1 else e2)-\v\n2\ s2" by (iprover) then have "G\Norm s0 \e0 ? e1 : e2-\v\max n1 n2\ s2" by (blast intro: evaln.Cond dest: evaln_max2) then show ?case .. next case (Call s0 e a' s1 args vs s2 invDeclC mode statT mn pTs' s3 s3' accC' v s4) then obtain n1 n2 where "G\Norm s0 \e-\a'\n1\ s1" "G\s1 \args\\vs\n2\ s2" by iprover moreover note \invDeclC = invocation_declclass G mode (store s2) a' statT \name=mn,parTs=pTs'\\ moreover note \s3 = init_lvars G invDeclC \name=mn,parTs=pTs'\ mode a' vs s2\ moreover note \s3'=check_method_access G accC' statT mode \name=mn,parTs=pTs'\ a' s3\ moreover from Call.hyps obtain m where "G\s3' \Methd invDeclC \name=mn, parTs=pTs'\-\v\m\ s4" by iprover ultimately have "G\Norm s0 \{accC',statT,mode}e\mn( {pTs'}args)-\v\max n1 (max n2 m)\ (set_lvars (locals (store s2))) s4" by (auto intro!: evaln.Call max.cobounded1 le_max3I1 le_max3I2) thus ?case .. next case (Methd s0 D sig v s1) then obtain n where "G\Norm s0 \body G D sig-\v\n\ s1" by iprover then have "G\Norm s0 \Methd D sig-\v\Suc n\ s1" by (rule evaln.Methd) then show ?case .. next case (Body s0 D s1 c s2 s3) from Body.hyps obtain n1 n2 where evaln_init: "G\Norm s0 \Init D\n1\ s1" and evaln_c: "G\s1 \c\n2\ s2" by (iprover) moreover note \s3 = (if \l. fst s2 = Some (Jump (Break l)) \ fst s2 = Some (Jump (Cont l)) then abupd (\x. Some (Error CrossMethodJump)) s2 else s2)\ ultimately have "G\Norm s0 \Body D c-\the (locals (store s2) Result)\max n1 n2 \ abupd (absorb Ret) s3" by (iprover intro: evaln.Body dest: evaln_max2) then show ?case .. next case (LVar s vn ) obtain n where "G\Norm s \LVar vn=\lvar vn s\n\ Norm s" by (iprover intro: evaln.LVar) then show ?case .. next case (FVar s0 statDeclC s1 e a s2 v s2' stat fn s3 accC) then obtain n1 n2 where "G\Norm s0 \Init statDeclC\n1\ s1" "G\s1 \e-\a\n2\ s2" by iprover moreover note \s3 = check_field_access G accC statDeclC fn stat a s2'\ and \(v, s2') = fvar statDeclC stat fn a s2\ ultimately have "G\Norm s0 \{accC,statDeclC,stat}e..fn=\v\max n1 n2\ s3" by (iprover intro: evaln.FVar dest: evaln_max2) then show ?case .. next case (AVar s0 e1 a s1 e2 i s2 v s2') then obtain n1 n2 where "G\Norm s0 \e1-\a\n1\ s1" "G\s1 \e2-\i\n2\ s2" by iprover moreover note \(v, s2') = avar G i a s2\ ultimately have "G\Norm s0 \e1.[e2]=\v\max n1 n2\ s2'" by (blast intro!: evaln.AVar dest: evaln_max2) then show ?case .. next case (Nil s0) show ?case by (iprover intro: evaln.Nil) next case (Cons s0 e v s1 es vs s2) then obtain n1 n2 where "G\Norm s0 \e-\v\n1\ s1" "G\s1 \es\\vs\n2\ s2" by iprover then have "G\Norm s0 \e # es\\v # vs\max n1 n2\ s2" by (blast intro!: evaln.Cons dest: evaln_max2) then show ?case .. qed end diff --git a/src/HOL/Bali/WellType.thy b/src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy +++ b/src/HOL/Bali/WellType.thy @@ -1,686 +1,686 @@ (* Title: HOL/Bali/WellType.thy Author: David von Oheimb *) subsection \Well-typedness of Java programs\ theory WellType imports DeclConcepts begin text \ improvements over Java Specification 1.0: \begin{itemize} \item methods of Object can be called upon references of interface or array type \end{itemize} simplifications: \begin{itemize} \item the type rules include all static checks on statements and expressions, e.g. definedness of names (of parameters, locals, fields, methods) \end{itemize} design issues: \begin{itemize} \item unified type judgment for statements, variables, expressions, expression lists \item statements are typed like expressions with dummy type Void \item the typing rules take an extra argument that is capable of determining the dynamic type of objects. Therefore, they can be used for both checking static types and determining runtime types in transition semantics. \end{itemize} \ type_synonym lenv = "(lname, ty) table" \ \local variables, including This and Result\ record env = prg:: "prog" \ \program\ cls:: "qtname" \ \current package and class name\ lcl:: "lenv" \ \local environment\ translations (type) "lenv" <= (type) "(lname, ty) table" (type) "lenv" <= (type) "lname \ ty option" (type) "env" <= (type) "\prg::prog,cls::qtname,lcl::lenv\" (type) "env" <= (type) "\prg::prog,cls::qtname,lcl::lenv,\::'a\" abbreviation pkg :: "env \ pname" \ \select the current package from an environment\ where "pkg e == pid (cls e)" subsubsection "Static overloading: maximally specific methods " type_synonym emhead = "ref_ty \ mhead" \ \Some mnemotic selectors for emhead\ definition "declrefT" :: "emhead \ ref_ty" where "declrefT = fst" definition "mhd" :: "emhead \ mhead" where "mhd \ snd" lemma declrefT_simp[simp]:"declrefT (r,m) = r" by (simp add: declrefT_def) lemma mhd_simp[simp]:"mhd (r,m) = m" by (simp add: mhd_def) lemma static_mhd_simp[simp]: "static (mhd m) = is_static m" by (cases m) (simp add: member_is_static_simp mhd_def) lemma mhd_resTy_simp [simp]: "resTy (mhd m) = resTy m" by (cases m) simp lemma mhd_is_static_simp [simp]: "is_static (mhd m) = is_static m" by (cases m) simp lemma mhd_accmodi_simp [simp]: "accmodi (mhd m) = accmodi m" by (cases m) simp definition cmheads :: "prog \ qtname \ qtname \ sig \ emhead set" where "cmheads G S C = (\sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) ` set_option (accmethd G S C sig))" definition Objectmheads :: "prog \ qtname \ sig \ emhead set" where "Objectmheads G S = (\sig. (\(Cls,mthd). (ClassT Cls,(mhead mthd))) ` set_option (filter_tab (\sig m. accmodi m \ Private) (accmethd G S Object) sig))" definition accObjectmheads :: "prog \ qtname \ ref_ty \ sig \ emhead set" where "accObjectmheads G S T = (if G\RefT T accessible_in (pid S) then Objectmheads G S else (\sig. {}))" primrec mheads :: "prog \ qtname \ ref_ty \ sig \ emhead set" where "mheads G S NullT = (\sig. {})" | "mheads G S (IfaceT I) = (\sig. (\(I,h).(IfaceT I,h)) ` accimethds G (pid S) I sig \ accObjectmheads G S (IfaceT I) sig)" | "mheads G S (ClassT C) = cmheads G S C" | "mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)" definition \ \applicable methods, cf. 15.11.2.1\ appl_methds :: "prog \ qtname \ ref_ty \ sig \ (emhead \ ty list) set" where "appl_methds G S rt = (\ sig. {(mh,pTs') |mh pTs'. mh \ mheads G S rt \name=name sig,parTs=pTs'\ \ G\(parTs sig)[\]pTs'})" definition \ \more specific methods, cf. 15.11.2.2\ more_spec :: "prog \ emhead \ ty list \ emhead \ ty list \ bool" where "more_spec G = (\(mh,pTs). \(mh',pTs'). G\pTs[\]pTs')" (*more_spec G \\((d,h),pTs). \((d',h'),pTs'). G\RefT d\RefT d'\G\pTs[\]pTs'*) definition \ \maximally specific methods, cf. 15.11.2.2\ max_spec :: "prog \ qtname \ ref_ty \ sig \ (emhead \ ty list) set" where "max_spec G S rt sig = {m. m \appl_methds G S rt sig \ (\m'\appl_methds G S rt sig. more_spec G m' m \ m'=m)}" lemma max_spec2appl_meths: "x \ max_spec G S T sig \ x \ appl_methds G S T sig" by (auto simp: max_spec_def) lemma appl_methsD: "(mh,pTs')\appl_methds G S T \name=mn,parTs=pTs\ \ mh \ mheads G S T \name=mn,parTs=pTs'\ \ G\pTs[\]pTs'" by (auto simp: appl_methds_def) lemma max_spec2mheads: "max_spec G S rt \name=mn,parTs=pTs\ = insert (mh, pTs') A \ mh \ mheads G S rt \name=mn,parTs=pTs'\ \ G\pTs[\]pTs'" apply (auto dest: equalityD2 subsetD max_spec2appl_meths appl_methsD) done definition empty_dt :: "dyn_ty" where "empty_dt = (\a. None)" definition invmode :: "('a::type)member_scheme \ expr \ inv_mode" where "invmode m e = (if is_static m then Static else if e=Super then SuperM else IntVir)" lemma invmode_nonstatic [simp]: "invmode \access=a,static=False,\=x\ (Acc (LVar e)) = IntVir" apply (unfold invmode_def) apply (simp (no_asm) add: member_is_static_simp) done lemma invmode_Static_eq [simp]: "(invmode m e = Static) = is_static m" apply (unfold invmode_def) apply (simp (no_asm)) done lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\(is_static m) \ e\Super)" apply (unfold invmode_def) apply (simp (no_asm)) done lemma Null_staticD: "a'=Null \ (is_static m) \ invmode m e = IntVir \ a' \ Null" apply (clarsimp simp add: invmode_IntVir_eq) done subsubsection "Typing for unary operations" primrec unop_type :: "unop \ prim_ty" where "unop_type UPlus = Integer" | "unop_type UMinus = Integer" | "unop_type UBitNot = Integer" | "unop_type UNot = Boolean" primrec wt_unop :: "unop \ ty \ bool" where "wt_unop UPlus t = (t = PrimT Integer)" | "wt_unop UMinus t = (t = PrimT Integer)" | "wt_unop UBitNot t = (t = PrimT Integer)" | "wt_unop UNot t = (t = PrimT Boolean)" subsubsection "Typing for binary operations" primrec binop_type :: "binop \ prim_ty" where "binop_type Mul = Integer" | "binop_type Div = Integer" | "binop_type Mod = Integer" | "binop_type Plus = Integer" | "binop_type Minus = Integer" | "binop_type LShift = Integer" | "binop_type RShift = Integer" | "binop_type RShiftU = Integer" | "binop_type Less = Boolean" | "binop_type Le = Boolean" | "binop_type Greater = Boolean" | "binop_type Ge = Boolean" | "binop_type Eq = Boolean" | "binop_type Neq = Boolean" | "binop_type BitAnd = Integer" | "binop_type And = Boolean" | "binop_type BitXor = Integer" | "binop_type Xor = Boolean" | "binop_type BitOr = Integer" | "binop_type Or = Boolean" | "binop_type CondAnd = Boolean" | "binop_type CondOr = Boolean" primrec wt_binop :: "prog \ binop \ ty \ ty \ bool" where "wt_binop G Mul t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Div t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Mod t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Plus t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Minus t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G LShift t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G RShift t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G RShiftU t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Less t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Le t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Greater t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Ge t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Eq t1 t2 = (G\t1\t2 \ G\t2\t1)" | "wt_binop G Neq t1 t2 = (G\t1\t2 \ G\t2\t1)" | "wt_binop G BitAnd t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G And t1 t2 = ((t1 = PrimT Boolean) \ (t2 = PrimT Boolean))" | "wt_binop G BitXor t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Xor t1 t2 = ((t1 = PrimT Boolean) \ (t2 = PrimT Boolean))" | "wt_binop G BitOr t1 t2 = ((t1 = PrimT Integer) \ (t2 = PrimT Integer))" | "wt_binop G Or t1 t2 = ((t1 = PrimT Boolean) \ (t2 = PrimT Boolean))" | "wt_binop G CondAnd t1 t2 = ((t1 = PrimT Boolean) \ (t2 = PrimT Boolean))" | "wt_binop G CondOr t1 t2 = ((t1 = PrimT Boolean) \ (t2 = PrimT Boolean))" subsubsection "Typing for terms" type_synonym tys = "ty + ty list" translations (type) "tys" <= (type) "ty + ty list" inductive wt :: "env \ dyn_ty \ [term,tys] \ bool" ("_,_\_\_" [51,51,51,51] 50) and wt_stmt :: "env \ dyn_ty \ stmt \ bool" ("_,_\_\\" [51,51,51] 50) and ty_expr :: "env \ dyn_ty \ [expr ,ty ] \ bool" ("_,_\_\-_" [51,51,51,51] 50) and ty_var :: "env \ dyn_ty \ [var ,ty ] \ bool" ("_,_\_\=_" [51,51,51,51] 50) and ty_exprs :: "env \ dyn_ty \ [expr list, ty list] \ bool" ("_,_\_\\_" [51,51,51,51] 50) where "E,dt\s\\ \ E,dt\In1r s\Inl (PrimT Void)" | "E,dt\e\-T \ E,dt\In1l e\Inl T" | "E,dt\e\=T \ E,dt\In2 e\Inl T" | "E,dt\e\\T \ E,dt\In3 e\Inr T" \ \well-typed statements\ | Skip: "E,dt\Skip\\" | Expr: "\E,dt\e\-T\ \ E,dt\Expr e\\" \ \cf. 14.6\ | Lab: "E,dt\c\\ \ E,dt\l\ c\\" | Comp: "\E,dt\c1\\; E,dt\c2\\\ \ E,dt\c1;; c2\\" \ \cf. 14.8\ | If: "\E,dt\e\-PrimT Boolean; E,dt\c1\\; E,dt\c2\\\ \ E,dt\If(e) c1 Else c2\\" \ \cf. 14.10\ | Loop: "\E,dt\e\-PrimT Boolean; E,dt\c\\\ \ E,dt\l\ While(e) c\\" \ \cf. 14.13, 14.15, 14.16\ | Jmp: "E,dt\Jmp jump\\" \ \cf. 14.16\ | Throw: "\E,dt\e\-Class tn; prg E\tn\\<^sub>C SXcpt Throwable\ \ E,dt\Throw e\\" \ \cf. 14.18\ | Try: "\E,dt\c1\\; prg E\tn\\<^sub>C SXcpt Throwable; lcl E (VName vn)=None; E \lcl := (lcl E)(VName vn\Class tn)\,dt\c2\\\ \ E,dt\Try c1 Catch(tn vn) c2\\" \ \cf. 14.18\ | Fin: "\E,dt\c1\\; E,dt\c2\\\ \ E,dt\c1 Finally c2\\" | Init: "\is_class (prg E) C\ \ E,dt\Init C\\" \ \\<^term>\Init\ is created on the fly during evaluation (see Eval.thy). The class isn't necessarily accessible from the points \<^term>\Init\ is called. Therefor we only demand \<^term>\is_class\ and not \<^term>\is_acc_class\ here.\ \ \well-typed expressions\ \ \cf. 15.8\ | NewC: "\is_acc_class (prg E) (pkg E) C\ \ E,dt\NewC C\-Class C" \ \cf. 15.9\ | NewA: "\is_acc_type (prg E) (pkg E) T; E,dt\i\-PrimT Integer\ \ E,dt\New T[i]\-T.[]" \ \cf. 15.15\ | Cast: "\E,dt\e\-T; is_acc_type (prg E) (pkg E) T'; prg E\T\? T'\ \ E,dt\Cast T' e\-T'" \ \cf. 15.19.2\ | Inst: "\E,dt\e\-RefT T; is_acc_type (prg E) (pkg E) (RefT T'); prg E\RefT T\? RefT T'\ \ E,dt\e InstOf T'\-PrimT Boolean" \ \cf. 15.7.1\ | Lit: "\typeof dt x = Some T\ \ E,dt\Lit x\-T" | UnOp: "\E,dt\e\-Te; wt_unop unop Te; T=PrimT (unop_type unop)\ \ E,dt\UnOp unop e\-T" | BinOp: "\E,dt\e1\-T1; E,dt\e2\-T2; wt_binop (prg E) binop T1 T2; T=PrimT (binop_type binop)\ \ E,dt\BinOp binop e1 e2\-T" \ \cf. 15.10.2, 15.11.1\ | Super: "\lcl E This = Some (Class C); C \ Object; class (prg E) C = Some c\ \ E,dt\Super\-Class (super c)" \ \cf. 15.13.1, 15.10.1, 15.12\ | Acc: "\E,dt\va\=T\ \ E,dt\Acc va\-T" \ \cf. 15.25, 15.25.1\ | Ass: "\E,dt\va\=T; va \ LVar This; E,dt\v \-T'; prg E\T'\T\ \ E,dt\va:=v\-T'" \ \cf. 15.24\ | Cond: "\E,dt\e0\-PrimT Boolean; E,dt\e1\-T1; E,dt\e2\-T2; prg E\T1\T2 \ T = T2 \ prg E\T2\T1 \ T = T1\ \ E,dt\e0 ? e1 : e2\-T" \ \cf. 15.11.1, 15.11.2, 15.11.3\ | Call: "\E,dt\e\-RefT statT; E,dt\ps\\pTs; max_spec (prg E) (cls E) statT \name=mn,parTs=pTs\ = {((statDeclT,m),pTs')} \ \ E,dt\{cls E,statT,invmode m e}e\mn({pTs'}ps)\-(resTy m)" | Methd: "\is_class (prg E) C; methd (prg E) C sig = Some m; E,dt\Body (declclass m) (stmt (mbody (mthd m)))\-T\ \ E,dt\Methd C sig\-T" \ \The class \<^term>\C\ is the dynamic class of the method call (cf. Eval.thy). It hasn't got to be directly accessible from the current package \<^term>\(pkg E)\. Only the static class must be accessible (enshured indirectly by \<^term>\Call\). Note that l is just a dummy value. It is only used in the smallstep semantics. To proof typesafety directly for the smallstep semantics we would have to assume conformance of l here!\ | Body: "\is_class (prg E) D; E,dt\blk\\; (lcl E) Result = Some T; is_type (prg E) T\ \ E,dt\Body D blk\-T" \ \The class \<^term>\D\ implementing the method must not directly be accessible from the current package \<^term>\(pkg E)\, but can also be indirectly accessible due to inheritance (enshured in \<^term>\Call\) The result type hasn't got to be accessible in Java! (If it is not accessible you can only assign it to Object). For dummy value l see rule \<^term>\Methd\.\ \ \well-typed variables\ \ \cf. 15.13.1\ | LVar: "\lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\ \ E,dt\LVar vn\=T" \ \cf. 15.10.1\ | FVar: "\E,dt\e\-Class C; accfield (prg E) (cls E) C fn = Some (statDeclC,f)\ \ E,dt\{cls E,statDeclC,is_static f}e..fn\=(type f)" \ \cf. 15.12\ | AVar: "\E,dt\e\-T.[]; E,dt\i\-PrimT Integer\ \ E,dt\e.[i]\=T" \ \well-typed expression lists\ \ \cf. 15.11.???\ | Nil: "E,dt\[]\\[]" \ \cf. 15.11.???\ | Cons: "\E,dt\e \-T; E,dt\es\\Ts\ \ E,dt\e#es\\T#Ts" (* for purely static typing *) abbreviation wt_syntax :: "env \ [term,tys] \ bool" ("_\_\_" [51,51,51] 50) where "E\t\T == E,empty_dt\t\ T" abbreviation wt_stmt_syntax :: "env \ stmt \ bool" ("_\_\\" [51,51 ] 50) where "E\s\\ == E\In1r s \ Inl (PrimT Void)" abbreviation ty_expr_syntax :: "env \ [expr, ty] \ bool" ("_\_\-_" [51,51,51] 50) where "E\e\-T == E\In1l e \ Inl T" abbreviation ty_var_syntax :: "env \ [var, ty] \ bool" ("_\_\=_" [51,51,51] 50) where "E\e\=T == E\In2 e \ Inl T" abbreviation ty_exprs_syntax :: "env \ [expr list, ty list] \ bool" ("_\_\\_" [51,51,51] 50) where "E\e\\T == E\In3 e \ Inr T" notation (ASCII) wt_syntax ("_|-_::_" [51,51,51] 50) and wt_stmt_syntax ("_|-_:<>" [51,51 ] 50) and ty_expr_syntax ("_|-_:-_" [51,51,51] 50) and ty_var_syntax ("_|-_:=_" [51,51,51] 50) and ty_exprs_syntax ("_|-_:#_" [51,51,51] 50) declare not_None_eq [simp del] declare if_split [split del] if_split_asm [split del] declare split_paired_All [simp del] split_paired_Ex [simp del] setup \map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\ inductive_cases wt_elim_cases [cases set]: "E,dt\In2 (LVar vn) \T" "E,dt\In2 ({accC,statDeclC,s}e..fn)\T" "E,dt\In2 (e.[i]) \T" "E,dt\In1l (NewC C) \T" "E,dt\In1l (New T'[i]) \T" "E,dt\In1l (Cast T' e) \T" "E,dt\In1l (e InstOf T') \T" "E,dt\In1l (Lit x) \T" "E,dt\In1l (UnOp unop e) \T" "E,dt\In1l (BinOp binop e1 e2) \T" "E,dt\In1l (Super) \T" "E,dt\In1l (Acc va) \T" "E,dt\In1l (Ass va v) \T" "E,dt\In1l (e0 ? e1 : e2) \T" "E,dt\In1l ({accC,statT,mode}e\mn({pT'}p))\T" "E,dt\In1l (Methd C sig) \T" "E,dt\In1l (Body D blk) \T" "E,dt\In3 ([]) \Ts" "E,dt\In3 (e#es) \Ts" "E,dt\In1r Skip \x" "E,dt\In1r (Expr e) \x" "E,dt\In1r (c1;; c2) \x" "E,dt\In1r (l\ c) \x" "E,dt\In1r (If(e) c1 Else c2) \x" "E,dt\In1r (l\ While(e) c) \x" "E,dt\In1r (Jmp jump) \x" "E,dt\In1r (Throw e) \x" "E,dt\In1r (Try c1 Catch(tn vn) c2)\x" "E,dt\In1r (c1 Finally c2) \x" "E,dt\In1r (Init C) \x" declare not_None_eq [simp] declare if_split [split] if_split_asm [split] declare split_paired_All [simp] split_paired_Ex [simp] setup \map_theory_simpset (fn ctxt => ctxt addloop ("split_all_tac", split_all_tac))\ lemma is_acc_class_is_accessible: "is_acc_class G P C \ G\(Class C) accessible_in P" by (auto simp add: is_acc_class_def) lemma is_acc_iface_is_iface: "is_acc_iface G P I \ is_iface G I" by (auto simp add: is_acc_iface_def) lemma is_acc_iface_Iface_is_accessible: "is_acc_iface G P I \ G\(Iface I) accessible_in P" by (auto simp add: is_acc_iface_def) lemma is_acc_type_is_type: "is_acc_type G P T \ is_type G T" by (auto simp add: is_acc_type_def) lemma is_acc_iface_is_accessible: "is_acc_type G P T \ G\T accessible_in P" by (auto simp add: is_acc_type_def) lemma wt_Methd_is_methd: "E\In1l (Methd C sig)\T \ is_methd (prg E) C sig" apply (erule_tac wt_elim_cases) apply clarsimp apply (erule is_methdI, assumption) done text \Special versions of some typing rules, better suited to pattern match the conclusion (no selectors in the conclusion) \ lemma wt_Call: "\E,dt\e\-RefT statT; E,dt\ps\\pTs; max_spec (prg E) (cls E) statT \name=mn,parTs=pTs\ = {((statDeclC,m),pTs')};rT=(resTy m);accC=cls E; mode = invmode m e\ \ E,dt\{accC,statT,mode}e\mn({pTs'}ps)\-rT" by (auto elim: wt.Call) lemma invocationTypeExpr_noClassD: "\ E\e\-RefT statT\ \ (\ statC. statT \ ClassT statC) \ invmode m e \ SuperM" proof - assume wt: "E\e\-RefT statT" show ?thesis proof (cases "e=Super") case True with wt obtain "C" where "statT = ClassT C" by (blast elim: wt_elim_cases) then show ?thesis by blast next case False then show ?thesis by (auto simp add: invmode_def) qed qed lemma wt_Super: "\lcl E This = Some (Class C); C \ Object; class (prg E) C = Some c; D=super c\ \ E,dt\Super\-Class D" by (auto elim: wt.Super) lemma wt_FVar: "\E,dt\e\-Class C; accfield (prg E) (cls E) C fn = Some (statDeclC,f); sf=is_static f; fT=(type f); accC=cls E\ \ E,dt\{accC,statDeclC,sf}e..fn\=fT" by (auto dest: wt.FVar) lemma wt_init [iff]: "E,dt\Init C\\ = is_class (prg E) C" by (auto elim: wt_elim_cases intro: "wt.Init") declare wt.Skip [iff] lemma wt_StatRef: "is_acc_type (prg E) (pkg E) (RefT rt) \ E\StatRef rt\-RefT rt" apply (rule wt.Cast) apply (rule wt.Lit) apply (simp (no_asm)) apply (simp (no_asm_simp)) apply (rule cast.widen) apply (simp (no_asm)) done lemma wt_Inj_elim: "\E. E,dt\t\U \ case t of In1 ec \ (case ec of Inl e \ \T. U=Inl T | Inr s \ U=Inl (PrimT Void)) | In2 e \ (\T. U=Inl T) | In3 e \ (\T. U=Inr T)" apply (erule wt.induct) apply auto done \ \In the special syntax to distinguish the typing judgements for expressions, statements, variables and expression lists the kind of term corresponds to the kind of type in the end e.g. An statement (injection \<^term>\In3\ into terms, always has type void (injection \<^term>\Inl\ into the generalised types. The following simplification procedures establish these kinds of correlation.\ lemma wt_expr_eq: "E,dt\In1l t\U = (\T. U=Inl T \ E,dt\t\-T)" by (auto, frule wt_Inj_elim, auto) lemma wt_var_eq: "E,dt\In2 t\U = (\T. U=Inl T \ E,dt\t\=T)" by (auto, frule wt_Inj_elim, auto) lemma wt_exprs_eq: "E,dt\In3 t\U = (\Ts. U=Inr Ts \ E,dt\t\\Ts)" by (auto, frule wt_Inj_elim, auto) lemma wt_stmt_eq: "E,dt\In1r t\U = (U=Inl(PrimT Void)\E,dt\t\\)" by (auto, frule wt_Inj_elim, auto, frule wt_Inj_elim, auto) simproc_setup wt_expr ("E,dt\In1l t\U") = \ - fn _ => fn _ => fn ct => - (case Thm.term_of ct of - (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE - | _ => SOME (mk_meta_eq @{thm wt_expr_eq}))\ - -simproc_setup wt_var ("E,dt\In2 t\U") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE - | _ => SOME (mk_meta_eq @{thm wt_var_eq}))\ + | _ => SOME (mk_meta_eq @{thm wt_expr_eq}))))\ -simproc_setup wt_exprs ("E,dt\In3 t\U") = \ - fn _ => fn _ => fn ct => +simproc_setup wt_var ("E,dt\In2 t\U") = \ + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE - | _ => SOME (mk_meta_eq @{thm wt_exprs_eq}))\ + | _ => SOME (mk_meta_eq @{thm wt_var_eq}))))\ + +simproc_setup wt_exprs ("E,dt\In3 t\U") = \ + K (K (fn ct => + (case Thm.term_of ct of + (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE + | _ => SOME (mk_meta_eq @{thm wt_exprs_eq}))))\ simproc_setup wt_stmt ("E,dt\In1r t\U") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of (_ $ _ $ _ $ _ $ (Const _ $ _)) => NONE - | _ => SOME (mk_meta_eq @{thm wt_stmt_eq}))\ + | _ => SOME (mk_meta_eq @{thm wt_stmt_eq}))))\ lemma wt_elim_BinOp: "\E,dt\In1l (BinOp binop e1 e2)\T; \T1 T2 T3. \E,dt\e1\-T1; E,dt\e2\-T2; wt_binop (prg E) binop T1 T2; E,dt\(if b then In1l e2 else In1r Skip)\T3; T = Inl (PrimT (binop_type binop))\ \ P\ \ P" apply (erule wt_elim_cases) apply (cases b) apply auto done lemma Inj_eq_lemma [simp]: "(\T. (\T'. T = Inj T' \ P T') \ Q T) = (\T'. P T' \ Q (Inj T'))" by auto (* unused *) lemma single_valued_tys_lemma [rule_format (no_asm)]: "\S T. G\S\T \ G\T\S \ S = T \ E,dt\t\T \ G = prg E \ (\T'. E,dt\t\T' \ T = T')" apply (cases "E", erule wt.induct) apply (safe del: disjE) apply (simp_all (no_asm_use) split del: if_split_asm) apply (safe del: disjE) (* 17 subgoals *) apply (tactic \ALLGOALS (fn i => if i = 11 then EVERY' [Rule_Insts.thin_tac \<^context> "E,dt\e0\-PrimT Boolean" [(\<^binding>\E\, NONE, NoSyn)], Rule_Insts.thin_tac \<^context> "E,dt\e1\-T1" [(\<^binding>\E\, NONE, NoSyn), (\<^binding>\T1\, NONE, NoSyn)], Rule_Insts.thin_tac \<^context> "E,dt\e2\-T2" [(\<^binding>\E\, NONE, NoSyn), (\<^binding>\T2\, NONE, NoSyn)]] i else Rule_Insts.thin_tac \<^context> "All P" [(\<^binding>\P\, NONE, NoSyn)] i)\) (*apply (safe del: disjE elim!: wt_elim_cases)*) apply (tactic \ALLGOALS (eresolve_tac \<^context> @{thms wt_elim_cases})\) apply (simp_all (no_asm_use) split del: if_split_asm) apply (erule_tac [12] V = "All P" for P in thin_rl) (* Call *) apply (blast del: equalityCE dest: sym [THEN trans])+ done (* unused *) lemma single_valued_tys: "ws_prog (prg E) \ single_valued {(t,T). E,dt\t\T}" apply (unfold single_valued_def) apply clarsimp apply (rule single_valued_tys_lemma) apply (auto intro!: widen_antisym) done lemma typeof_empty_is_type: "typeof (\a. None) v = Some T \ is_type G T" by (induct v) auto (* unused *) lemma typeof_is_type: "(\a. v \ Addr a) \ \T. typeof dt v = Some T \ is_type G T" by (induct v) auto end diff --git a/src/HOL/Boolean_Algebras.thy b/src/HOL/Boolean_Algebras.thy --- a/src/HOL/Boolean_Algebras.thy +++ b/src/HOL/Boolean_Algebras.thy @@ -1,573 +1,573 @@ (* Title: HOL/Boolean_Algebras.thy Author: Brian Huffman Author: Florian Haftmann *) section \Boolean Algebras\ theory Boolean_Algebras imports Lattices begin subsection \Abstract boolean algebra\ locale abstract_boolean_algebra = conj: abel_semigroup \(\<^bold>\)\ + disj: abel_semigroup \(\<^bold>\)\ for conj :: \'a \ 'a \ 'a\ (infixr \\<^bold>\\ 70) and disj :: \'a \ 'a \ 'a\ (infixr \\<^bold>\\ 65) + fixes compl :: \'a \ 'a\ (\\<^bold>- _\ [81] 80) and zero :: \'a\ (\\<^bold>0\) and one :: \'a\ (\\<^bold>1\) assumes conj_disj_distrib: \x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)\ and disj_conj_distrib: \x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)\ and conj_one_right: \x \<^bold>\ \<^bold>1 = x\ and disj_zero_right: \x \<^bold>\ \<^bold>0 = x\ and conj_cancel_right [simp]: \x \<^bold>\ \<^bold>- x = \<^bold>0\ and disj_cancel_right [simp]: \x \<^bold>\ \<^bold>- x = \<^bold>1\ begin sublocale conj: semilattice_neutr \(\<^bold>\)\ \\<^bold>1\ proof show "x \<^bold>\ \<^bold>1 = x" for x by (fact conj_one_right) show "x \<^bold>\ x = x" for x proof - have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>0" by (simp add: disj_zero_right) also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" by simp also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" by (simp only: conj_disj_distrib) also have "\ = x \<^bold>\ \<^bold>1" by simp also have "\ = x" by (simp add: conj_one_right) finally show ?thesis . qed qed sublocale disj: semilattice_neutr \(\<^bold>\)\ \\<^bold>0\ proof show "x \<^bold>\ \<^bold>0 = x" for x by (fact disj_zero_right) show "x \<^bold>\ x = x" for x proof - have "x \<^bold>\ x = (x \<^bold>\ x) \<^bold>\ \<^bold>1" by simp also have "\ = (x \<^bold>\ x) \<^bold>\ (x \<^bold>\ \<^bold>- x)" by simp also have "\ = x \<^bold>\ (x \<^bold>\ \<^bold>- x)" by (simp only: disj_conj_distrib) also have "\ = x \<^bold>\ \<^bold>0" by simp also have "\ = x" by (simp add: disj_zero_right) finally show ?thesis . qed qed subsubsection \Complement\ lemma complement_unique: assumes 1: "a \<^bold>\ x = \<^bold>0" assumes 2: "a \<^bold>\ x = \<^bold>1" assumes 3: "a \<^bold>\ y = \<^bold>0" assumes 4: "a \<^bold>\ y = \<^bold>1" shows "x = y" proof - from 1 3 have "(a \<^bold>\ x) \<^bold>\ (x \<^bold>\ y) = (a \<^bold>\ y) \<^bold>\ (x \<^bold>\ y)" by simp then have "(x \<^bold>\ a) \<^bold>\ (x \<^bold>\ y) = (y \<^bold>\ a) \<^bold>\ (y \<^bold>\ x)" by (simp add: ac_simps) then have "x \<^bold>\ (a \<^bold>\ y) = y \<^bold>\ (a \<^bold>\ x)" by (simp add: conj_disj_distrib) with 2 4 have "x \<^bold>\ \<^bold>1 = y \<^bold>\ \<^bold>1" by simp then show "x = y" by simp qed lemma compl_unique: "x \<^bold>\ y = \<^bold>0 \ x \<^bold>\ y = \<^bold>1 \ \<^bold>- x = y" by (rule complement_unique [OF conj_cancel_right disj_cancel_right]) lemma double_compl [simp]: "\<^bold>- (\<^bold>- x) = x" proof (rule compl_unique) show "\<^bold>- x \<^bold>\ x = \<^bold>0" by (simp only: conj_cancel_right conj.commute) show "\<^bold>- x \<^bold>\ x = \<^bold>1" by (simp only: disj_cancel_right disj.commute) qed lemma compl_eq_compl_iff [simp]: \\<^bold>- x = \<^bold>- y \ x = y\ (is \?P \ ?Q\) proof assume \?Q\ then show ?P by simp next assume \?P\ then have \\<^bold>- (\<^bold>- x) = \<^bold>- (\<^bold>- y)\ by simp then show ?Q by simp qed subsubsection \Conjunction\ lemma conj_zero_right [simp]: "x \<^bold>\ \<^bold>0 = \<^bold>0" using conj.left_idem conj_cancel_right by fastforce lemma compl_one [simp]: "\<^bold>- \<^bold>1 = \<^bold>0" by (rule compl_unique [OF conj_zero_right disj_zero_right]) lemma conj_zero_left [simp]: "\<^bold>0 \<^bold>\ x = \<^bold>0" by (subst conj.commute) (rule conj_zero_right) lemma conj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>0" by (subst conj.commute) (rule conj_cancel_right) lemma conj_disj_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" by (simp only: conj.commute conj_disj_distrib) lemmas conj_disj_distribs = conj_disj_distrib conj_disj_distrib2 subsubsection \Disjunction\ context begin interpretation dual: abstract_boolean_algebra \(\<^bold>\)\ \(\<^bold>\)\ compl \\<^bold>1\ \\<^bold>0\ apply standard apply (rule disj_conj_distrib) apply (rule conj_disj_distrib) apply simp_all done lemma disj_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>1" by (fact dual.conj_zero_right) lemma compl_zero [simp]: "\<^bold>- \<^bold>0 = \<^bold>1" by (fact dual.compl_one) lemma disj_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>1" by (fact dual.conj_zero_left) lemma disj_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>1" by (fact dual.conj_cancel_left) lemma disj_conj_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" by (fact dual.conj_disj_distrib2) lemmas disj_conj_distribs = disj_conj_distrib disj_conj_distrib2 end subsubsection \De Morgan's Laws\ lemma de_Morgan_conj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" proof (rule compl_unique) have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = ((x \<^bold>\ y) \<^bold>\ \<^bold>- x) \<^bold>\ ((x \<^bold>\ y) \<^bold>\ \<^bold>- y)" by (rule conj_disj_distrib) also have "\ = (y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (x \<^bold>\ (y \<^bold>\ \<^bold>- y))" by (simp only: ac_simps) finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>0" by (simp only: conj_cancel_right conj_zero_right disj_zero_right) next have "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = (x \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)) \<^bold>\ (y \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y))" by (rule disj_conj_distrib2) also have "\ = (\<^bold>- y \<^bold>\ (x \<^bold>\ \<^bold>- x)) \<^bold>\ (\<^bold>- x \<^bold>\ (y \<^bold>\ \<^bold>- y))" by (simp only: ac_simps) finally show "(x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y) = \<^bold>1" by (simp only: disj_cancel_right disj_one_right conj_one_right) qed context begin interpretation dual: abstract_boolean_algebra \(\<^bold>\)\ \(\<^bold>\)\ compl \\<^bold>1\ \\<^bold>0\ apply standard apply (rule disj_conj_distrib) apply (rule conj_disj_distrib) apply simp_all done lemma de_Morgan_disj [simp]: "\<^bold>- (x \<^bold>\ y) = \<^bold>- x \<^bold>\ \<^bold>- y" by (fact dual.de_Morgan_conj) end end subsection \Symmetric Difference\ locale abstract_boolean_algebra_sym_diff = abstract_boolean_algebra + fixes xor :: \'a \ 'a \ 'a\ (infixr \\<^bold>\\ 65) assumes xor_def : \x \<^bold>\ y = (x \<^bold>\ \<^bold>- y) \<^bold>\ (\<^bold>- x \<^bold>\ y)\ begin sublocale xor: comm_monoid xor \\<^bold>0\ proof fix x y z :: 'a let ?t = "(x \<^bold>\ y \<^bold>\ z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y \<^bold>\ z)" have "?t \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ y \<^bold>\ \<^bold>- y) = ?t \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- y) \<^bold>\ (x \<^bold>\ z \<^bold>\ \<^bold>- z)" by (simp only: conj_cancel_right conj_zero_right) then show "(x \<^bold>\ y) \<^bold>\ z = x \<^bold>\ (y \<^bold>\ z)" by (simp only: xor_def de_Morgan_disj de_Morgan_conj double_compl) (simp only: conj_disj_distribs conj_ac ac_simps) show "x \<^bold>\ y = y \<^bold>\ x" by (simp only: xor_def ac_simps) show "x \<^bold>\ \<^bold>0 = x" by (simp add: xor_def) qed lemma xor_def2: \x \<^bold>\ y = (x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)\ proof - note xor_def [of x y] also have \x \<^bold>\ \<^bold>- y \<^bold>\ \<^bold>- x \<^bold>\ y = ((x \<^bold>\ \<^bold>- x) \<^bold>\ (\<^bold>- y \<^bold>\ \<^bold>- x)) \<^bold>\ (x \<^bold>\ y) \<^bold>\ (\<^bold>- y \<^bold>\ y)\ by (simp add: ac_simps disj_conj_distribs) also have \\ = (x \<^bold>\ y) \<^bold>\ (\<^bold>- x \<^bold>\ \<^bold>- y)\ by (simp add: ac_simps) finally show ?thesis . qed lemma xor_one_right [simp]: "x \<^bold>\ \<^bold>1 = \<^bold>- x" by (simp only: xor_def compl_one conj_zero_right conj_one_right disj.left_neutral) lemma xor_one_left [simp]: "\<^bold>1 \<^bold>\ x = \<^bold>- x" using xor_one_right [of x] by (simp add: ac_simps) lemma xor_self [simp]: "x \<^bold>\ x = \<^bold>0" by (simp only: xor_def conj_cancel_right conj_cancel_left disj_zero_right) lemma xor_left_self [simp]: "x \<^bold>\ (x \<^bold>\ y) = y" by (simp only: xor.assoc [symmetric] xor_self xor.left_neutral) lemma xor_compl_left [simp]: "\<^bold>- x \<^bold>\ y = \<^bold>- (x \<^bold>\ y)" by (simp add: ac_simps flip: xor_one_left) lemma xor_compl_right [simp]: "x \<^bold>\ \<^bold>- y = \<^bold>- (x \<^bold>\ y)" using xor.commute xor_compl_left by auto lemma xor_cancel_right [simp]: "x \<^bold>\ \<^bold>- x = \<^bold>1" by (simp only: xor_compl_right xor_self compl_zero) lemma xor_cancel_left [simp]: "\<^bold>- x \<^bold>\ x = \<^bold>1" by (simp only: xor_compl_left xor_self compl_zero) lemma conj_xor_distrib: "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" proof - have *: "(x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z) = (y \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (z \<^bold>\ x \<^bold>\ \<^bold>- x) \<^bold>\ (x \<^bold>\ y \<^bold>\ \<^bold>- z) \<^bold>\ (x \<^bold>\ \<^bold>- y \<^bold>\ z)" by (simp only: conj_cancel_right conj_zero_right disj.left_neutral) then show "x \<^bold>\ (y \<^bold>\ z) = (x \<^bold>\ y) \<^bold>\ (x \<^bold>\ z)" by (simp (no_asm_use) only: xor_def de_Morgan_disj de_Morgan_conj double_compl conj_disj_distribs ac_simps) qed lemma conj_xor_distrib2: "(y \<^bold>\ z) \<^bold>\ x = (y \<^bold>\ x) \<^bold>\ (z \<^bold>\ x)" by (simp add: conj.commute conj_xor_distrib) lemmas conj_xor_distribs = conj_xor_distrib conj_xor_distrib2 end subsection \Type classes\ class boolean_algebra = distrib_lattice + bounded_lattice + minus + uminus + assumes inf_compl_bot: \x \ - x = \\ and sup_compl_top: \x \ - x = \\ assumes diff_eq: \x - y = x \ - y\ begin sublocale boolean_algebra: abstract_boolean_algebra \(\)\ \(\)\ uminus \ \ apply standard apply (rule inf_sup_distrib1) apply (rule sup_inf_distrib1) apply (simp_all add: ac_simps inf_compl_bot sup_compl_top) done lemma compl_inf_bot: "- x \ x = \" by (fact boolean_algebra.conj_cancel_left) lemma compl_sup_top: "- x \ x = \" by (fact boolean_algebra.disj_cancel_left) lemma compl_unique: assumes "x \ y = \" and "x \ y = \" shows "- x = y" using assms by (rule boolean_algebra.compl_unique) lemma double_compl: "- (- x) = x" by (fact boolean_algebra.double_compl) lemma compl_eq_compl_iff: "- x = - y \ x = y" by (fact boolean_algebra.compl_eq_compl_iff) lemma compl_bot_eq: "- \ = \" by (fact boolean_algebra.compl_zero) lemma compl_top_eq: "- \ = \" by (fact boolean_algebra.compl_one) lemma compl_inf: "- (x \ y) = - x \ - y" by (fact boolean_algebra.de_Morgan_conj) lemma compl_sup: "- (x \ y) = - x \ - y" by (fact boolean_algebra.de_Morgan_disj) lemma compl_mono: assumes "x \ y" shows "- y \ - x" proof - from assms have "x \ y = y" by (simp only: le_iff_sup) then have "- (x \ y) = - y" by simp then have "- x \ - y = - y" by simp then have "- y \ - x = - y" by (simp only: inf_commute) then show ?thesis by (simp only: le_iff_inf) qed lemma compl_le_compl_iff [simp]: "- x \ - y \ y \ x" by (auto dest: compl_mono) lemma compl_le_swap1: assumes "y \ - x" shows "x \ -y" proof - from assms have "- (- x) \ - y" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed lemma compl_le_swap2: assumes "- y \ x" shows "- x \ y" proof - from assms have "- x \ - (- y)" by (simp only: compl_le_compl_iff) then show ?thesis by simp qed lemma compl_less_compl_iff [simp]: "- x < - y \ y < x" by (auto simp add: less_le) lemma compl_less_swap1: assumes "y < - x" shows "x < - y" proof - from assms have "- (- x) < - y" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma compl_less_swap2: assumes "- y < x" shows "- x < y" proof - from assms have "- x < - (- y)" by (simp only: compl_less_compl_iff) then show ?thesis by simp qed lemma sup_cancel_left1: \x \ a \ (- x \ b) = \\ by (simp add: ac_simps) lemma sup_cancel_left2: \- x \ a \ (x \ b) = \\ by (simp add: ac_simps) lemma inf_cancel_left1: \x \ a \ (- x \ b) = \\ by (simp add: ac_simps) lemma inf_cancel_left2: \- x \ a \ (x \ b) = \\ by (simp add: ac_simps) lemma sup_compl_top_left1 [simp]: \- x \ (x \ y) = \\ by (simp add: sup_assoc [symmetric]) lemma sup_compl_top_left2 [simp]: \x \ (- x \ y) = \\ using sup_compl_top_left1 [of "- x" y] by simp lemma inf_compl_bot_left1 [simp]: \- x \ (x \ y) = \\ by (simp add: inf_assoc [symmetric]) lemma inf_compl_bot_left2 [simp]: \x \ (- x \ y) = \\ using inf_compl_bot_left1 [of "- x" y] by simp lemma inf_compl_bot_right [simp]: \x \ (y \ - x) = \\ by (subst inf_left_commute) simp end subsection \Lattice on \<^typ>\bool\\ instantiation bool :: boolean_algebra begin definition bool_Compl_def [simp]: "uminus = Not" definition bool_diff_def [simp]: "A - B \ A \ \ B" definition [simp]: "P \ Q \ P \ Q" definition [simp]: "P \ Q \ P \ Q" instance by standard auto end lemma sup_boolI1: "P \ P \ Q" by simp lemma sup_boolI2: "Q \ P \ Q" by simp lemma sup_boolE: "P \ Q \ (P \ R) \ (Q \ R) \ R" by auto instance "fun" :: (type, boolean_algebra) boolean_algebra by standard (rule ext, simp_all add: inf_compl_bot sup_compl_top diff_eq)+ subsection \Lattice on unary and binary predicates\ lemma inf1I: "A x \ B x \ (A \ B) x" by (simp add: inf_fun_def) lemma inf2I: "A x y \ B x y \ (A \ B) x y" by (simp add: inf_fun_def) lemma inf1E: "(A \ B) x \ (A x \ B x \ P) \ P" by (simp add: inf_fun_def) lemma inf2E: "(A \ B) x y \ (A x y \ B x y \ P) \ P" by (simp add: inf_fun_def) lemma inf1D1: "(A \ B) x \ A x" by (rule inf1E) lemma inf2D1: "(A \ B) x y \ A x y" by (rule inf2E) lemma inf1D2: "(A \ B) x \ B x" by (rule inf1E) lemma inf2D2: "(A \ B) x y \ B x y" by (rule inf2E) lemma sup1I1: "A x \ (A \ B) x" by (simp add: sup_fun_def) lemma sup2I1: "A x y \ (A \ B) x y" by (simp add: sup_fun_def) lemma sup1I2: "B x \ (A \ B) x" by (simp add: sup_fun_def) lemma sup2I2: "B x y \ (A \ B) x y" by (simp add: sup_fun_def) lemma sup1E: "(A \ B) x \ (A x \ P) \ (B x \ P) \ P" by (simp add: sup_fun_def) iprover lemma sup2E: "(A \ B) x y \ (A x y \ P) \ (B x y \ P) \ P" by (simp add: sup_fun_def) iprover text \ \<^medskip> Classical introduction rule: no commitment to \A\ vs \B\.\ lemma sup1CI: "(\ B x \ A x) \ (A \ B) x" by (auto simp add: sup_fun_def) lemma sup2CI: "(\ B x y \ A x y) \ (A \ B) x y" by (auto simp add: sup_fun_def) subsection \Simproc setup\ locale boolean_algebra_cancel begin lemma sup1: "(A::'a::semilattice_sup) \ sup k a \ sup A b \ sup k (sup a b)" by (simp only: ac_simps) lemma sup2: "(B::'a::semilattice_sup) \ sup k b \ sup a B \ sup k (sup a b)" by (simp only: ac_simps) lemma sup0: "(a::'a::bounded_semilattice_sup_bot) \ sup a bot" by simp lemma inf1: "(A::'a::semilattice_inf) \ inf k a \ inf A b \ inf k (inf a b)" by (simp only: ac_simps) lemma inf2: "(B::'a::semilattice_inf) \ inf k b \ inf a B \ inf k (inf a b)" by (simp only: ac_simps) lemma inf0: "(a::'a::bounded_semilattice_inf_top) \ inf a top" by simp end ML_file \Tools/boolean_algebra_cancel.ML\ simproc_setup boolean_algebra_cancel_sup ("sup a b::'a::boolean_algebra") = - \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_sup_conv\ + \K (K (try Boolean_Algebra_Cancel.cancel_sup_conv))\ simproc_setup boolean_algebra_cancel_inf ("inf a b::'a::boolean_algebra") = - \fn phi => fn ss => try Boolean_Algebra_Cancel.cancel_inf_conv\ + \K (K (try Boolean_Algebra_Cancel.cancel_inf_conv))\ context boolean_algebra begin lemma shunt1: "(x \ y \ z) \ (x \ -y \ z)" proof assume "x \ y \ z" hence "-y \ (x \ y) \ -y \ z" using sup.mono by blast hence "-y \ x \ -y \ z" by (simp add: sup_inf_distrib1) thus "x \ -y \ z" by simp next assume "x \ -y \ z" hence "x \ y \ (-y \ z) \ y" using inf_mono by auto thus "x \ y \ z" using inf.boundedE inf_sup_distrib2 by auto qed lemma shunt2: "(x \ -y \ z) \ (x \ y \ z)" by (simp add: shunt1) lemma inf_shunt: "(x \ y = \) \ (x \ - y)" by (simp add: order.eq_iff shunt1) lemma sup_shunt: "(x \ y = \) \ (- x \ y)" using inf_shunt [of \- x\ \- y\, symmetric] by (simp flip: compl_sup compl_top_eq) lemma diff_shunt_var: "(x - y = \) \ (x \ y)" by (simp add: diff_eq inf_shunt) lemma sup_neg_inf: \p \ q \ r \ p \ -q \ r\ (is \?P \ ?Q\) proof assume ?P then have \p \ - q \ (q \ r) \ - q\ by (rule inf_mono) simp then show ?Q by (simp add: inf_sup_distrib2) next assume ?Q then have \p \ - q \ q \ r \ q\ by (rule sup_mono) simp then show ?P by (simp add: sup_inf_distrib ac_simps) qed end end diff --git a/src/HOL/Data_Structures/Less_False.thy b/src/HOL/Data_Structures/Less_False.thy --- a/src/HOL/Data_Structures/Less_False.thy +++ b/src/HOL/Data_Structures/Less_False.thy @@ -1,31 +1,31 @@ (* Author: Tobias Nipkow *) section \Improved Simproc for $<$\ theory Less_False imports Main begin -simproc_setup less_False ("(x::'a::order) < y") = \fn _ => fn ctxt => fn ct => +simproc_setup less_False ("(x::'a::order) < y") = \K (fn ctxt => fn ct => let fun prp t thm = Thm.full_prop_of thm aconv t; val eq_False_if_not = @{thm eq_False} RS iffD2 fun prove_less_False ((less as Const(_,T)) $ r $ s) = let val prems = Simplifier.prems_of ctxt; val le = Const (\<^const_name>\less_eq\, T); val t = HOLogic.mk_Trueprop(le $ s $ r); in case find_first (prp t) prems of NONE => let val t = HOLogic.mk_Trueprop(less $ s $ r) in case find_first (prp t) prems of NONE => NONE | SOME thm => SOME(mk_meta_eq((thm RS @{thm less_not_sym}) RS eq_False_if_not)) end | SOME thm => NONE end; - in prove_less_False (Thm.term_of ct) end + in prove_less_False (Thm.term_of ct) end) \ end diff --git a/src/HOL/Decision_Procs/Commutative_Ring.thy b/src/HOL/Decision_Procs/Commutative_Ring.thy --- a/src/HOL/Decision_Procs/Commutative_Ring.thy +++ b/src/HOL/Decision_Procs/Commutative_Ring.thy @@ -1,963 +1,962 @@ (* Title: HOL/Decision_Procs/Commutative_Ring.thy Author: Bernhard Haeupler, Stefan Berghofer, and Amine Chaieb Proving equalities in commutative rings done "right" in Isabelle/HOL. *) section \Proving equalities in commutative rings\ theory Commutative_Ring imports Conversions Algebra_Aux "HOL-Library.Code_Target_Numeral" begin text \Syntax of multivariate polynomials (pol) and polynomial expressions.\ datatype pol = Pc int | Pinj nat pol | PX pol nat pol datatype polex = Var nat | Const int | Add polex polex | Sub polex polex | Mul polex polex | Pow polex nat | Neg polex text \Interpretation functions for the shadow syntax.\ context cring begin definition in_carrier :: "'a list \ bool" where "in_carrier xs \ (\x\set xs. x \ carrier R)" lemma in_carrier_Nil: "in_carrier []" by (simp add: in_carrier_def) lemma in_carrier_Cons: "x \ carrier R \ in_carrier xs \ in_carrier (x # xs)" by (simp add: in_carrier_def) lemma drop_in_carrier [simp]: "in_carrier xs \ in_carrier (drop n xs)" using set_drop_subset [of n xs] by (auto simp add: in_carrier_def) primrec head :: "'a list \ 'a" where "head [] = \" | "head (x # xs) = x" lemma head_closed [simp]: "in_carrier xs \ head xs \ carrier R" by (cases xs) (simp_all add: in_carrier_def) primrec Ipol :: "'a list \ pol \ 'a" where "Ipol l (Pc c) = \c\" | "Ipol l (Pinj i P) = Ipol (drop i l) P" | "Ipol l (PX P x Q) = Ipol l P \ head l [^] x \ Ipol (drop 1 l) Q" lemma Ipol_Pc: "Ipol l (Pc 0) = \" "Ipol l (Pc 1) = \" "Ipol l (Pc (numeral n)) = \numeral n\" "Ipol l (Pc (- numeral n)) = \ \numeral n\" by simp_all lemma Ipol_closed [simp]: "in_carrier l \ Ipol l p \ carrier R" by (induct p arbitrary: l) simp_all primrec Ipolex :: "'a list \ polex \ 'a" where "Ipolex l (Var n) = head (drop n l)" | "Ipolex l (Const i) = \i\" | "Ipolex l (Add P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Sub P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Mul P Q) = Ipolex l P \ Ipolex l Q" | "Ipolex l (Pow p n) = Ipolex l p [^] n" | "Ipolex l (Neg P) = \ Ipolex l P" lemma Ipolex_Const: "Ipolex l (Const 0) = \" "Ipolex l (Const 1) = \" "Ipolex l (Const (numeral n)) = \numeral n\" by simp_all end text \Create polynomial normalized polynomials given normalized inputs.\ definition mkPinj :: "nat \ pol \ pol" where "mkPinj x P = (case P of Pc c \ Pc c | Pinj y P \ Pinj (x + y) P | PX p1 y p2 \ Pinj x P)" definition mkPX :: "pol \ nat \ pol \ pol" where "mkPX P i Q = (case P of Pc c \ if c = 0 then mkPinj 1 Q else PX P i Q | Pinj j R \ PX P i Q | PX P2 i2 Q2 \ if Q2 = Pc 0 then PX P2 (i + i2) Q else PX P i Q)" text \Defining the basic ring operations on normalized polynomials\ function add :: "pol \ pol \ pol" (infixl "\+\" 65) where "Pc a \+\ Pc b = Pc (a + b)" | "Pc c \+\ Pinj i P = Pinj i (P \+\ Pc c)" | "Pinj i P \+\ Pc c = Pinj i (P \+\ Pc c)" | "Pc c \+\ PX P i Q = PX P i (Q \+\ Pc c)" | "PX P i Q \+\ Pc c = PX P i (Q \+\ Pc c)" | "Pinj x P \+\ Pinj y Q = (if x = y then mkPinj x (P \+\ Q) else (if x > y then mkPinj y (Pinj (x - y) P \+\ Q) else mkPinj x (Pinj (y - x) Q \+\ P)))" | "Pinj x P \+\ PX Q y R = (if x = 0 then P \+\ PX Q y R else (if x = 1 then PX Q y (R \+\ P) else PX Q y (R \+\ Pinj (x - 1) P)))" | "PX P x R \+\ Pinj y Q = (if y = 0 then PX P x R \+\ Q else (if y = 1 then PX P x (R \+\ Q) else PX P x (R \+\ Pinj (y - 1) Q)))" | "PX P1 x P2 \+\ PX Q1 y Q2 = (if x = y then mkPX (P1 \+\ Q1) x (P2 \+\ Q2) else (if x > y then mkPX (PX P1 (x - y) (Pc 0) \+\ Q1) y (P2 \+\ Q2) else mkPX (PX Q1 (y - x) (Pc 0) \+\ P1) x (P2 \+\ Q2)))" by pat_completeness auto termination by (relation "measure (\(x, y). size x + size y)") auto function mul :: "pol \ pol \ pol" (infixl "\*\" 70) where "Pc a \*\ Pc b = Pc (a * b)" | "Pc c \*\ Pinj i P = (if c = 0 then Pc 0 else mkPinj i (P \*\ Pc c))" | "Pinj i P \*\ Pc c = (if c = 0 then Pc 0 else mkPinj i (P \*\ Pc c))" | "Pc c \*\ PX P i Q = (if c = 0 then Pc 0 else mkPX (P \*\ Pc c) i (Q \*\ Pc c))" | "PX P i Q \*\ Pc c = (if c = 0 then Pc 0 else mkPX (P \*\ Pc c) i (Q \*\ Pc c))" | "Pinj x P \*\ Pinj y Q = (if x = y then mkPinj x (P \*\ Q) else (if x > y then mkPinj y (Pinj (x - y) P \*\ Q) else mkPinj x (Pinj (y - x) Q \*\ P)))" | "Pinj x P \*\ PX Q y R = (if x = 0 then P \*\ PX Q y R else (if x = 1 then mkPX (Pinj x P \*\ Q) y (R \*\ P) else mkPX (Pinj x P \*\ Q) y (R \*\ Pinj (x - 1) P)))" | "PX P x R \*\ Pinj y Q = (if y = 0 then PX P x R \*\ Q else (if y = 1 then mkPX (Pinj y Q \*\ P) x (R \*\ Q) else mkPX (Pinj y Q \*\ P) x (R \*\ Pinj (y - 1) Q)))" | "PX P1 x P2 \*\ PX Q1 y Q2 = mkPX (P1 \*\ Q1) (x + y) (P2 \*\ Q2) \+\ (mkPX (P1 \*\ mkPinj 1 Q2) x (Pc 0) \+\ (mkPX (Q1 \*\ mkPinj 1 P2) y (Pc 0)))" by pat_completeness auto termination by (relation "measure (\(x, y). size x + size y)") (auto simp add: mkPinj_def split: pol.split) text \Negation\ primrec neg :: "pol \ pol" where "neg (Pc c) = Pc (- c)" | "neg (Pinj i P) = Pinj i (neg P)" | "neg (PX P x Q) = PX (neg P) x (neg Q)" text \Subtraction\ definition sub :: "pol \ pol \ pol" (infixl "\-\" 65) where "sub P Q = P \+\ neg Q" text \Square for Fast Exponentiation\ primrec sqr :: "pol \ pol" where "sqr (Pc c) = Pc (c * c)" | "sqr (Pinj i P) = mkPinj i (sqr P)" | "sqr (PX A x B) = mkPX (sqr A) (x + x) (sqr B) \+\ mkPX (Pc 2 \*\ A \*\ mkPinj 1 B) x (Pc 0)" text \Fast Exponentiation\ fun pow :: "nat \ pol \ pol" where pow_if [simp del]: "pow n P = (if n = 0 then Pc 1 else if even n then pow (n div 2) (sqr P) else P \*\ pow (n div 2) (sqr P))" lemma pow_simps [simp]: "pow 0 P = Pc 1" "pow (2 * n) P = pow n (sqr P)" "pow (Suc (2 * n)) P = P \*\ pow n (sqr P)" by (simp_all add: pow_if) lemma even_pow: "even n \ pow n P = pow (n div 2) (sqr P)" by (erule evenE) simp lemma odd_pow: "odd n \ pow n P = P \*\ pow (n div 2) (sqr P)" by (erule oddE) simp text \Normalization of polynomial expressions\ primrec norm :: "polex \ pol" where "norm (Var n) = (if n = 0 then PX (Pc 1) 1 (Pc 0) else Pinj n (PX (Pc 1) 1 (Pc 0)))" | "norm (Const i) = Pc i" | "norm (Add P Q) = norm P \+\ norm Q" | "norm (Sub P Q) = norm P \-\ norm Q" | "norm (Mul P Q) = norm P \*\ norm Q" | "norm (Pow P n) = pow n (norm P)" | "norm (Neg P) = neg (norm P)" context cring begin text \mkPinj preserve semantics\ lemma mkPinj_ci: "Ipol l (mkPinj a B) = Ipol l (Pinj a B)" by (induct B) (auto simp add: mkPinj_def algebra_simps) text \mkPX preserves semantics\ lemma mkPX_ci: "in_carrier l \ Ipol l (mkPX A b C) = Ipol l (PX A b C)" by (cases A) (auto simp add: mkPX_def mkPinj_ci nat_pow_mult [symmetric] m_ac) text \Correctness theorems for the implemented operations\ text \Negation\ lemma neg_ci: "in_carrier l \ Ipol l (neg P) = \ (Ipol l P)" by (induct P arbitrary: l) (auto simp add: minus_add l_minus) text \Addition\ lemma add_ci: "in_carrier l \ Ipol l (P \+\ Q) = Ipol l P \ Ipol l Q" proof (induct P Q arbitrary: l rule: add.induct) case (6 x P y Q) consider "x < y" | "x = y" | "x > y" by arith then show ?case proof cases case 1 with 6 show ?thesis by (simp add: mkPinj_ci a_ac) next case 2 with 6 show ?thesis by (simp add: mkPinj_ci) next case 3 with 6 show ?thesis by (simp add: mkPinj_ci) qed next case (7 x P Q y R) consider "x = 0" | "x = 1" | "x > 1" by arith then show ?case proof cases case 1 with 7 show ?thesis by simp next case 2 with 7 show ?thesis by (simp add: a_ac) next case 3 with 7 show ?thesis by (cases x) (simp_all add: a_ac) qed next case (8 P x R y Q) then show ?case by (simp add: a_ac) next case (9 P1 x P2 Q1 y Q2) consider "x = y" | d where "d + x = y" | d where "d + y = x" by atomize_elim arith then show ?case proof cases case 1 with 9 show ?thesis by (simp add: mkPX_ci r_distr a_ac m_ac) next case 2 with 9 show ?thesis by (auto simp add: mkPX_ci nat_pow_mult [symmetric] r_distr a_ac m_ac) next case 3 with 9 show ?thesis by (auto simp add: nat_pow_mult [symmetric] mkPX_ci r_distr a_ac m_ac) qed qed (auto simp add: a_ac m_ac) text \Multiplication\ lemma mul_ci: "in_carrier l \ Ipol l (P \*\ Q) = Ipol l P \ Ipol l Q" by (induct P Q arbitrary: l rule: mul.induct) (simp_all add: mkPX_ci mkPinj_ci a_ac m_ac r_distr add_ci nat_pow_mult [symmetric]) text \Subtraction\ lemma sub_ci: "in_carrier l \ Ipol l (P \-\ Q) = Ipol l P \ Ipol l Q" by (simp add: add_ci neg_ci sub_def minus_eq) text \Square\ lemma sqr_ci: "in_carrier ls \ Ipol ls (sqr P) = Ipol ls P \ Ipol ls P" by (induct P arbitrary: ls) (simp_all add: add_ci mkPinj_ci mkPX_ci mul_ci l_distr r_distr a_ac m_ac nat_pow_mult [symmetric] of_int_2) text \Power\ lemma pow_ci: "in_carrier ls \ Ipol ls (pow n P) = Ipol ls P [^] n" proof (induct n arbitrary: P rule: less_induct) case (less k) consider "k = 0" | "k > 0" by arith then show ?case proof cases case 1 then show ?thesis by simp next case 2 then have "k div 2 < k" by arith with less have *: "Ipol ls (pow (k div 2) (sqr P)) = Ipol ls (sqr P) [^] (k div 2)" by simp show ?thesis proof (cases "even k") case True with * \in_carrier ls\ show ?thesis by (simp add: even_pow sqr_ci nat_pow_distrib nat_pow_mult mult_2 [symmetric] even_two_times_div_two) next case False with * \in_carrier ls\ show ?thesis by (simp add: odd_pow mul_ci sqr_ci nat_pow_distrib nat_pow_mult mult_2 [symmetric] trans [OF nat_pow_Suc m_comm, symmetric]) qed qed qed text \Normalization preserves semantics\ lemma norm_ci: "in_carrier l \ Ipolex l Pe = Ipol l (norm Pe)" by (induct Pe) (simp_all add: add_ci sub_ci mul_ci neg_ci pow_ci) text \Reflection lemma: Key to the (incomplete) decision procedure\ lemma norm_eq: assumes "in_carrier l" and "norm P1 = norm P2" shows "Ipolex l P1 = Ipolex l P2" proof - from assms have "Ipol l (norm P1) = Ipol l (norm P2)" by simp with assms show ?thesis by (simp only: norm_ci) qed end text \Monomials\ datatype mon = Mc int | Minj nat mon | MX nat mon primrec (in cring) Imon :: "'a list \ mon \ 'a" where "Imon l (Mc c) = \c\" | "Imon l (Minj i M) = Imon (drop i l) M" | "Imon l (MX x M) = Imon (drop 1 l) M \ head l [^] x" lemma (in cring) Imon_closed [simp]: "in_carrier l \ Imon l m \ carrier R" by (induct m arbitrary: l) simp_all definition mkMinj :: "nat \ mon \ mon" where "mkMinj i M = (case M of Mc c \ Mc c | Minj j M \ Minj (i + j) M | _ \ Minj i M)" definition Minj_pred :: "nat \ mon \ mon" where "Minj_pred i M = (if i = 1 then M else mkMinj (i - 1) M)" primrec mkMX :: "nat \ mon \ mon" where "mkMX i (Mc c) = MX i (Mc c)" | "mkMX i (Minj j M) = (if j = 0 then mkMX i M else MX i (Minj_pred j M))" | "mkMX i (MX j M) = MX (i + j) M" lemma (in cring) mkMinj_correct: "Imon l (mkMinj i M) = Imon l (Minj i M)" by (simp add: mkMinj_def add.commute split: mon.split) lemma (in cring) Minj_pred_correct: "0 < i \ Imon (drop 1 l) (Minj_pred i M) = Imon l (Minj i M)" by (simp add: Minj_pred_def mkMinj_correct) lemma (in cring) mkMX_correct: "in_carrier l \ Imon l (mkMX i M) = Imon l M \ head l [^] i" by (induct M) (simp_all add: Minj_pred_correct [simplified] nat_pow_mult [symmetric] m_ac split: mon.split) fun cfactor :: "pol \ int \ pol \ pol" where "cfactor (Pc c') c = (Pc (c' mod c), Pc (c' div c))" | "cfactor (Pinj i P) c = (let (R, S) = cfactor P c in (mkPinj i R, mkPinj i S))" | "cfactor (PX P i Q) c = (let (R1, S1) = cfactor P c; (R2, S2) = cfactor Q c in (mkPX R1 i R2, mkPX S1 i S2))" lemma (in cring) cfactor_correct: "in_carrier l \ Ipol l P = Ipol l (fst (cfactor P c)) \ \c\ \ Ipol l (snd (cfactor P c))" proof (induct P c arbitrary: l rule: cfactor.induct) case (1 c' c) show ?case by (simp add: of_int_mult [symmetric] of_int_add [symmetric] del: of_int_mult) next case (2 i P c) then show ?case by (simp add: mkPinj_ci split_beta) next case (3 P i Q c) from 3(1) 3(2) [OF refl prod.collapse] 3(3) show ?case by (simp add: mkPX_ci r_distr a_ac m_ac split_beta) qed fun mfactor :: "pol \ mon \ pol \ pol" where "mfactor P (Mc c) = (if c = 1 then (Pc 0, P) else cfactor P c)" | "mfactor (Pc d) M = (Pc d, Pc 0)" | "mfactor (Pinj i P) (Minj j M) = (if i = j then let (R, S) = mfactor P M in (mkPinj i R, mkPinj i S) else if i < j then let (R, S) = mfactor P (Minj (j - i) M) in (mkPinj i R, mkPinj i S) else (Pinj i P, Pc 0))" | "mfactor (Pinj i P) (MX j M) = (Pinj i P, Pc 0)" | "mfactor (PX P i Q) (Minj j M) = (if j = 0 then mfactor (PX P i Q) M else let (R1, S1) = mfactor P (Minj j M); (R2, S2) = mfactor Q (Minj_pred j M) in (mkPX R1 i R2, mkPX S1 i S2))" | "mfactor (PX P i Q) (MX j M) = (if i = j then let (R, S) = mfactor P (mkMinj 1 M) in (mkPX R i Q, S) else if i < j then let (R, S) = mfactor P (MX (j - i) M) in (mkPX R i Q, S) else let (R, S) = mfactor P (mkMinj 1 M) in (mkPX R i Q, mkPX S (i - j) (Pc 0)))" lemmas mfactor_induct = mfactor.induct [case_names Mc Pc_Minj Pc_MX Pinj_Minj Pinj_MX PX_Minj PX_MX] lemma (in cring) mfactor_correct: "in_carrier l \ Ipol l P = Ipol l (fst (mfactor P M)) \ Imon l M \ Ipol l (snd (mfactor P M))" proof (induct P M arbitrary: l rule: mfactor_induct) case (Mc P c) then show ?case by (simp add: cfactor_correct) next case (Pc_Minj d i M) then show ?case by simp next case (Pc_MX d i M) then show ?case by simp next case (Pinj_Minj i P j M) then show ?case by (simp add: mkPinj_ci split_beta) next case (Pinj_MX i P j M) then show ?case by simp next case (PX_Minj P i Q j M) from PX_Minj(1,2) PX_Minj(3) [OF _ refl prod.collapse] PX_Minj(4) show ?case by (simp add: mkPX_ci Minj_pred_correct [simplified] r_distr a_ac m_ac split_beta) next case (PX_MX P i Q j M) from \in_carrier l\ have eq1: "(Imon (drop (Suc 0) l) M \ head l [^] (j - i)) \ Ipol l (snd (mfactor P (MX (j - i) M))) \ head l [^] i = Imon (drop (Suc 0) l) M \ Ipol l (snd (mfactor P (MX (j - i) M))) \ (head l [^] (j - i) \ head l [^] i)" by (simp add: m_ac) from \in_carrier l\ have eq2: "(Imon (drop (Suc 0) l) M \ head l [^] j) \ (Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \ head l [^] (i - j)) = Imon (drop (Suc 0) l) M \ Ipol l (snd (mfactor P (mkMinj (Suc 0) M))) \ (head l [^] (i - j) \ head l [^] j)" by (simp add: m_ac) from PX_MX \in_carrier l\ show ?case by (simp add: mkPX_ci mkMinj_correct l_distr eq1 eq2 split_beta nat_pow_mult) (simp add: a_ac m_ac) qed primrec mon_of_pol :: "pol \ mon option" where "mon_of_pol (Pc c) = Some (Mc c)" | "mon_of_pol (Pinj i P) = (case mon_of_pol P of None \ None | Some M \ Some (mkMinj i M))" | "mon_of_pol (PX P i Q) = (if Q = Pc 0 then (case mon_of_pol P of None \ None | Some M \ Some (mkMX i M)) else None)" lemma (in cring) mon_of_pol_correct: assumes "in_carrier l" and "mon_of_pol P = Some M" shows "Ipol l P = Imon l M" using assms proof (induct P arbitrary: M l) case (PX P1 i P2) from PX(1,3,4) show ?case by (auto simp add: mkMinj_correct mkMX_correct split: if_split_asm option.split_asm) qed (auto simp add: mkMinj_correct split: option.split_asm) fun (in cring) Ipolex_polex_list :: "'a list \ (polex \ polex) list \ bool" where "Ipolex_polex_list l [] = True" | "Ipolex_polex_list l ((P, Q) # pps) = ((Ipolex l P = Ipolex l Q) \ Ipolex_polex_list l pps)" fun (in cring) Imon_pol_list :: "'a list \ (mon \ pol) list \ bool" where "Imon_pol_list l [] = True" | "Imon_pol_list l ((M, P) # mps) = ((Imon l M = Ipol l P) \ Imon_pol_list l mps)" fun mk_monpol_list :: "(polex \ polex) list \ (mon \ pol) list" where "mk_monpol_list [] = []" | "mk_monpol_list ((P, Q) # pps) = (case mon_of_pol (norm P) of None \ mk_monpol_list pps | Some M \ (M, norm Q) # mk_monpol_list pps)" lemma (in cring) mk_monpol_list_correct: "in_carrier l \ Ipolex_polex_list l pps \ Imon_pol_list l (mk_monpol_list pps)" by (induct pps rule: mk_monpol_list.induct) (auto split: option.split simp add: norm_ci [symmetric] mon_of_pol_correct [symmetric]) definition ponesubst :: "pol \ mon \ pol \ pol option" where "ponesubst P1 M P2 = (let (Q, R) = mfactor P1 M in (case R of Pc c \ if c = 0 then None else Some (add Q (mul P2 R)) | _ \ Some (add Q (mul P2 R))))" fun pnsubst1 :: "pol \ mon \ pol \ nat \ pol" where "pnsubst1 P1 M P2 n = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ if n = 0 then P3 else pnsubst1 P3 M P2 (n - 1))" lemma pnsubst1_0 [simp]: "pnsubst1 P1 M P2 0 = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ P3)" by (simp split: option.split) lemma pnsubst1_Suc [simp]: "pnsubst1 P1 M P2 (Suc n) = (case ponesubst P1 M P2 of None \ P1 | Some P3 \ pnsubst1 P3 M P2 n)" by (simp split: option.split) declare pnsubst1.simps [simp del] definition pnsubst :: "pol \ mon \ pol \ nat \ pol option" where "pnsubst P1 M P2 n = (case ponesubst P1 M P2 of None \ None | Some P3 \ Some (pnsubst1 P3 M P2 n))" fun psubstl1 :: "pol \ (mon \ pol) list \ nat \ pol" where "psubstl1 P1 [] n = P1" | "psubstl1 P1 ((M, P2) # mps) n = psubstl1 (pnsubst1 P1 M P2 n) mps n" fun psubstl :: "pol \ (mon \ pol) list \ nat \ pol option" where "psubstl P1 [] n = None" | "psubstl P1 ((M, P2) # mps) n = (case pnsubst P1 M P2 n of None \ psubstl P1 mps n | Some P3 \ Some (psubstl1 P3 mps n))" fun pnsubstl :: "pol \ (mon \ pol) list \ nat \ nat \ pol" where "pnsubstl P1 mps m n = (case psubstl P1 mps n of None \ P1 | Some P3 \ if m = 0 then P3 else pnsubstl P3 mps (m - 1) n)" lemma pnsubstl_0 [simp]: "pnsubstl P1 mps 0 n = (case psubstl P1 mps n of None \ P1 | Some P3 \ P3)" by (simp split: option.split) lemma pnsubstl_Suc [simp]: "pnsubstl P1 mps (Suc m) n = (case psubstl P1 mps n of None \ P1 | Some P3 \ pnsubstl P3 mps m n)" by (simp split: option.split) declare pnsubstl.simps [simp del] lemma (in cring) ponesubst_correct: "in_carrier l \ ponesubst P1 M P2 = Some P3 \ Imon l M = Ipol l P2 \ Ipol l P1 = Ipol l P3" by (auto simp add: ponesubst_def split_beta mfactor_correct [of l P1 M] add_ci mul_ci split: pol.split_asm if_split_asm) lemma (in cring) pnsubst1_correct: "in_carrier l \ Imon l M = Ipol l P2 \ Ipol l (pnsubst1 P1 M P2 n) = Ipol l P1" by (induct n arbitrary: P1) (simp_all add: ponesubst_correct split: option.split) lemma (in cring) pnsubst_correct: "in_carrier l \ pnsubst P1 M P2 n = Some P3 \ Imon l M = Ipol l P2 \ Ipol l P1 = Ipol l P3" by (auto simp add: pnsubst_def pnsubst1_correct ponesubst_correct split: option.split_asm) lemma (in cring) psubstl1_correct: "in_carrier l \ Imon_pol_list l mps \ Ipol l (psubstl1 P1 mps n) = Ipol l P1" by (induct P1 mps n rule: psubstl1.induct) (simp_all add: pnsubst1_correct) lemma (in cring) psubstl_correct: "in_carrier l \ psubstl P1 mps n = Some P2 \ Imon_pol_list l mps \ Ipol l P1 = Ipol l P2" by (induct P1 mps n rule: psubstl.induct) (auto simp add: psubstl1_correct pnsubst_correct split: option.split_asm) lemma (in cring) pnsubstl_correct: "in_carrier l \ Imon_pol_list l mps \ Ipol l (pnsubstl P1 mps m n) = Ipol l P1" by (induct m arbitrary: P1) (simp_all add: psubstl_correct split: option.split) lemma (in cring) norm_subst_correct: "in_carrier l \ Ipolex_polex_list l pps \ Ipolex l P = Ipol l (pnsubstl (norm P) (mk_monpol_list pps) m n)" by (simp add: pnsubstl_correct mk_monpol_list_correct norm_ci) lemma in_carrier_trivial: "cring_class.in_carrier l" by (induct l) (simp_all add: cring_class.in_carrier_def carrier_class) ML \ val term_of_nat = HOLogic.mk_number \<^Type>\nat\ o @{code integer_of_nat}; val term_of_int = HOLogic.mk_number \<^Type>\int\ o @{code integer_of_int}; fun term_of_pol (@{code Pc} k) = \<^Const>\Pc\ $ term_of_int k | term_of_pol (@{code Pinj} (n, p)) = \<^Const>\Pinj\ $ term_of_nat n $ term_of_pol p | term_of_pol (@{code PX} (p, n, q)) = \<^Const>\PX\ $ term_of_pol p $ term_of_nat n $ term_of_pol q; local fun pol (ctxt, ct, t) = \<^instantiate>\x = ct and y = \Thm.cterm_of ctxt t\ in cterm \x \ y\ for x y :: pol\; val (_, raw_pol_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\pnsubstl\, pol))); fun pol_oracle ctxt ct t = raw_pol_oracle (ctxt, ct, t); in val cv = @{computation_conv pol terms: pnsubstl mk_monpol_list norm nat_of_integer Code_Target_Nat.natural "0::nat" "1::nat" "2::nat" "3::nat" "0::int" "1::int" "2::int" "3::int" "-1::int" datatypes: polex "(polex * polex) list" int integer num} (fn ctxt => fn p => fn ct => pol_oracle ctxt ct (term_of_pol p)) end \ ML \ signature RING_TAC = sig structure Ring_Simps: sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end val insert_rules: ((term * 'a) * (term * 'a) -> bool) -> (term * 'a) -> (term * 'a) Net.net -> (term * 'a) Net.net val eq_ring_simps: (term * (thm list * thm list * thm list * thm list * thm * thm)) * (term * (thm list * thm list * thm list * thm list * thm * thm)) -> bool val ring_tac: bool -> thm list -> Proof.context -> int -> tactic val get_matching_rules: Proof.context -> (term * 'a) Net.net -> term -> 'a option val norm: thm -> thm val mk_in_carrier: Proof.context -> term -> thm list -> (string * typ) list -> thm val mk_ring: typ -> term end structure Ring_Tac : RING_TAC = struct fun eq_ring_simps ((t, (ths1, ths2, ths3, ths4, th5, th)), (t', (ths1', ths2', ths3', ths4', th5', th'))) = t aconv t' andalso eq_list Thm.eq_thm (ths1, ths1') andalso eq_list Thm.eq_thm (ths2, ths2') andalso eq_list Thm.eq_thm (ths3, ths3') andalso eq_list Thm.eq_thm (ths4, ths4') andalso Thm.eq_thm (th5, th5') andalso Thm.eq_thm (th, th'); structure Ring_Simps = Generic_Data (struct type T = (term * (thm list * thm list * thm list * thm list * thm * thm)) Net.net val empty = Net.empty val merge = Net.merge eq_ring_simps end); fun get_matching_rules ctxt net t = get_first (fn (p, x) => if Pattern.matches (Proof_Context.theory_of ctxt) (p, t) then SOME x else NONE) (Net.match_term net t); fun insert_rules eq (t, x) = Net.insert_term eq (t, (t, x)); fun norm thm = thm COMP_INCR asm_rl; fun get_ring_simps ctxt optcT t = (case get_matching_rules ctxt (Ring_Simps.get (Context.Proof ctxt)) t of SOME (ths1, ths2, ths3, ths4, th5, th) => let val tr = Thm.transfer' ctxt #> (case optcT of NONE => I | SOME cT => inst [cT] [] #> norm) in (map tr ths1, map tr ths2, map tr ths3, map tr ths4, tr th5, tr th) end | NONE => error "get_ring_simps: lookup failed"); fun ring_struct \<^Const_>\Ring.ring.add _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.a_minus _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Group.monoid.mult _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.a_inv _ _ for R _\ = SOME R | ring_struct \<^Const_>\Group.pow _ _ _ for R _ _\ = SOME R | ring_struct \<^Const_>\Ring.ring.zero _ _ for R\ = SOME R | ring_struct \<^Const_>\Group.monoid.one _ _ for R\ = SOME R | ring_struct \<^Const_>\Algebra_Aux.of_integer _ _ for R _\ = SOME R | ring_struct _ = NONE; fun reif_polex vs \<^Const_>\Ring.ring.add _ _ for _ a b\ = \<^Const>\Add for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Ring.a_minus _ _ for _ a b\ = \<^Const>\Sub for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Group.monoid.mult _ _ for _ a b\ = \<^Const>\Mul for \reif_polex vs a\ \reif_polex vs b\\ | reif_polex vs \<^Const_>\Ring.a_inv _ _ for _ a\ = \<^Const>\Neg for \reif_polex vs a\\ | reif_polex vs \<^Const_>\Group.pow _ _ _ for _ a n\ = \<^Const>\Pow for \reif_polex vs a\ n\ | reif_polex vs (Free x) = \<^Const>\Var for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_polex _ \<^Const_>\Ring.ring.zero _ _ for _\ = \<^term>\Const 0\ | reif_polex _ \<^Const_>\Group.monoid.one _ _ for _\ = \<^term>\Const 1\ | reif_polex _ \<^Const_>\Algebra_Aux.of_integer _ _ for _ n\ = \<^Const>\Const for n\ | reif_polex _ _ = error "reif_polex: bad expression"; fun reif_polex' vs \<^Const_>\plus _ for a b\ = \<^Const>\Add for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\minus _ for a b\ = \<^Const>\Sub for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\times _ for a b\ = \<^Const>\Mul for \reif_polex' vs a\ \reif_polex' vs b\\ | reif_polex' vs \<^Const_>\uminus _ for a\ = \<^Const>\Neg for \reif_polex' vs a\\ | reif_polex' vs \<^Const_>\power _ for a n\ = \<^Const>\Pow for \reif_polex' vs a\ n\ | reif_polex' vs (Free x) = \<^Const>\Var for \HOLogic.mk_number \<^Type>\nat\ (find_index (equal x) vs)\\ | reif_polex' _ \<^Const_>\numeral _ for b\ = \<^Const>\Const for \<^Const>\numeral \<^Type>\int\ for b\\ | reif_polex' _ \<^Const_>\zero_class.zero _\ = \<^term>\Const 0\ | reif_polex' _ \<^Const_>\one_class.one _\ = \<^term>\Const 1\ | reif_polex' _ t = error "reif_polex: bad expression"; fun head_conv (_, _, _, _, head_simp, _) ys = (case strip_app ys of (\<^const_name>\Cons\, [y, xs]) => inst [] [y, xs] head_simp); fun Ipol_conv (rls as ([Ipol_simps_1, Ipol_simps_2, Ipol_simps_3, Ipol_simps_4, Ipol_simps_5, Ipol_simps_6, Ipol_simps_7], _, _, _, _, _)) = let val a = type_of_eqn Ipol_simps_1; val drop_conv_a = drop_conv a; fun conv l p = (case strip_app p of (\<^const_name>\Pc\, [c]) => (case strip_numeral c of (\<^const_name>\zero_class.zero\, _) => inst [] [l] Ipol_simps_4 | (\<^const_name>\one_class.one\, _) => inst [] [l] Ipol_simps_5 | (\<^const_name>\numeral\, [m]) => inst [] [l, m] Ipol_simps_6 | (\<^const_name>\uminus\, [m]) => inst [] [l, m] Ipol_simps_7 | _ => inst [] [l, c] Ipol_simps_1) | (\<^const_name>\Pinj\, [i, P]) => transitive' (inst [] [l, i, P] Ipol_simps_2) (cong2' conv (args2 drop_conv_a) Thm.reflexive) | (\<^const_name>\PX\, [P, x, Q]) => transitive' (inst [] [l, P, x, Q] Ipol_simps_3) (cong2 (cong2 (args2 conv) (cong2 (args1 (head_conv rls)) Thm.reflexive)) (cong2' conv (args2 drop_conv_a) Thm.reflexive))) in conv end; fun Ipolex_conv (rls as (_, [Ipolex_Var, Ipolex_Const, Ipolex_Add, Ipolex_Sub, Ipolex_Mul, Ipolex_Pow, Ipolex_Neg, Ipolex_Const_0, Ipolex_Const_1, Ipolex_Const_numeral], _, _, _, _)) = let val a = type_of_eqn Ipolex_Var; val drop_conv_a = drop_conv a; fun conv l r = (case strip_app r of (\<^const_name>\Var\, [n]) => transitive' (inst [] [l, n] Ipolex_Var) (cong1' (head_conv rls) (args2 drop_conv_a)) | (\<^const_name>\Const\, [i]) => (case strip_app i of (\<^const_name>\zero_class.zero\, _) => inst [] [l] Ipolex_Const_0 | (\<^const_name>\one_class.one\, _) => inst [] [l] Ipolex_Const_1 | (\<^const_name>\numeral\, [m]) => inst [] [l, m] Ipolex_Const_numeral | _ => inst [] [l, i] Ipolex_Const) | (\<^const_name>\Add\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Add) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Sub\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Sub) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Mul\, [P, Q]) => transitive' (inst [] [l, P, Q] Ipolex_Mul) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\Pow\, [P, n]) => transitive' (inst [] [l, P, n] Ipolex_Pow) (cong2 (args2 conv) Thm.reflexive) | (\<^const_name>\Neg\, [P]) => transitive' (inst [] [l, P] Ipolex_Neg) (cong1 (args2 conv))) in conv end; fun Ipolex_polex_list_conv (rls as (_, _, [Ipolex_polex_list_Nil, Ipolex_polex_list_Cons], _, _, _)) l pps = (case strip_app pps of (\<^const_name>\Nil\, []) => inst [] [l] Ipolex_polex_list_Nil | (\<^const_name>\Cons\, [p, pps']) => (case strip_app p of (\<^const_name>\Pair\, [P, Q]) => transitive' (inst [] [l, P, Q, pps'] Ipolex_polex_list_Cons) (cong2 (cong2 (args2 (Ipolex_conv rls)) (args2 (Ipolex_conv rls))) (args2 (Ipolex_polex_list_conv rls))))); fun dest_conj th = let val th1 = th RS @{thm conjunct1}; val th2 = th RS @{thm conjunct2} in dest_conj th1 @ dest_conj th2 end handle THM _ => [th]; fun mk_in_carrier ctxt R prems xs = let val (_, _, _, [in_carrier_Nil, in_carrier_Cons], _, _) = get_ring_simps ctxt NONE R; val props = map fst (Facts.props (Proof_Context.facts_of ctxt)) @ maps dest_conj prems; val ths = map (fn p as (x, _) => (case find_first ((fn \<^Const_>\Trueprop for \<^Const_>\Set.member _ for \Free (y, _)\ \<^Const_>\carrier _ _ for S\\\ => x = y andalso R aconv S | _ => false) o Thm.prop_of) props of SOME th => th | NONE => error ("Variable " ^ Syntax.string_of_term ctxt (Free p) ^ " not in carrier"))) xs in fold_rev (fn th1 => fn th2 => [th1, th2] MRS in_carrier_Cons) ths in_carrier_Nil end; fun mk_ring T = \<^Const>\cring_class_ops T\; val iterations = \<^cterm>\1000::nat\; val Trueprop_cong = Thm.combination (Thm.reflexive \<^cterm>\Trueprop\); fun commutative_ring_conv ctxt prems eqs ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; val eqs' = map (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of) eqs; val xs = filter (equal T o snd) (rev (fold Term.add_frees (map fst eqs' @ map snd eqs' @ [Thm.term_of ct]) [])); val (R, optcT, prem', reif) = (case ring_struct (Thm.term_of ct) of SOME R => (R, NONE, mk_in_carrier ctxt R prems xs, reif_polex xs) | NONE => (mk_ring T, SOME cT, @{thm in_carrier_trivial}, reif_polex' xs)); val rls as (_, _, _, _, _, norm_subst_correct) = get_ring_simps ctxt optcT R; val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs)); val ceqs = Thm.cterm_of ctxt (HOLogic.mk_list \<^typ>\polex \ polex\ (map (HOLogic.mk_prod o apply2 reif) eqs')); val cp = Thm.cterm_of ctxt (reif (Thm.term_of ct)); val prem = Thm.equal_elim (Trueprop_cong (Thm.symmetric (Ipolex_polex_list_conv rls cxs ceqs))) (fold_rev (fn th1 => fn th2 => [th1, th2] MRS @{thm conjI}) eqs @{thm TrueI}); in Thm.transitive (Thm.symmetric (Ipolex_conv rls cxs cp)) (transitive' ([prem', prem] MRS inst [] [cxs, ceqs, cp, iterations, iterations] norm_subst_correct) (cong2' (Ipol_conv rls) Thm.reflexive (cv ctxt))) end; fun ring_tac in_prems thms ctxt = tactic_of_conv (fn ct => (if in_prems then Conv.prems_conv else Conv.concl_conv) (Logic.count_prems (Thm.term_of ct)) (Conv.arg_conv (Conv.binop_conv (commutative_ring_conv ctxt [] thms))) ct) THEN' TRY o (assume_tac ctxt ORELSE' resolve_tac ctxt [@{thm refl}]); end \ context cring begin -local_setup \ -Local_Theory.declaration {syntax = false, pervasive = false} - (fn phi => Ring_Tac.Ring_Simps.map (Ring_Tac.insert_rules Ring_Tac.eq_ring_simps +declaration \fn phi => + Ring_Tac.Ring_Simps.map (Ring_Tac.insert_rules Ring_Tac.eq_ring_simps (Morphism.term phi \<^term>\R\, (Morphism.fact phi @{thms Ipol.simps [meta] Ipol_Pc [meta]}, Morphism.fact phi @{thms Ipolex.simps [meta] Ipolex_Const [meta]}, Morphism.fact phi @{thms Ipolex_polex_list.simps [meta]}, Morphism.fact phi @{thms in_carrier_Nil in_carrier_Cons}, singleton (Morphism.fact phi) @{thm head.simps(2) [meta]}, - singleton (Morphism.fact phi) @{thm norm_subst_correct [meta]})))) + singleton (Morphism.fact phi) @{thm norm_subst_correct [meta]}))) \ end method_setup ring = \ Scan.lift (Args.mode "prems") -- Attrib.thms >> (SIMPLE_METHOD' oo uncurry Ring_Tac.ring_tac) \ "simplify equations involving rings" end diff --git a/src/HOL/Decision_Procs/Conversions.thy b/src/HOL/Decision_Procs/Conversions.thy --- a/src/HOL/Decision_Procs/Conversions.thy +++ b/src/HOL/Decision_Procs/Conversions.thy @@ -1,839 +1,839 @@ (* Title: HOL/Decision_Procs/Conversions.thy Author: Stefan Berghofer *) theory Conversions imports Main begin ML \ fun tactic_of_conv cv i st = if i > Thm.nprems_of st then Seq.empty else Seq.single (Conv.gconv_rule cv i st); fun binop_conv cv cv' = Conv.combination_conv (Conv.arg_conv cv) cv'; \ ML \ fun err s ct = error (s ^ ": " ^ Syntax.string_of_term_global (Thm.theory_of_cterm ct) (Thm.term_of ct)); \ attribute_setup meta = - \Scan.succeed (fn (ctxt, th) => (NONE, SOME (mk_meta_eq th)))\ + \Scan.succeed (Thm.rule_attribute [] (K mk_meta_eq))\ \convert equality to meta equality\ ML \ fun strip_app ct = ct |> Drule.strip_comb |>> Thm.term_of |>> dest_Const |>> fst; fun inst cTs cts th = Thm.instantiate' (map SOME cTs) (map SOME cts) th; fun transitive' eq eq' = Thm.transitive eq (eq' (Thm.rhs_of eq)); fun type_of_eqn eqn = Thm.ctyp_of_cterm (Thm.dest_arg1 (Thm.cprop_of eqn)); fun cong1 conv ct = Thm.combination (Thm.reflexive (Thm.dest_fun ct)) (conv (Thm.dest_arg ct)); fun cong1' conv' conv ct = let val eqn = conv (Thm.dest_arg ct) in Thm.transitive (Thm.combination (Thm.reflexive (Thm.dest_fun ct)) eqn) (conv' (Thm.rhs_of eqn)) end; fun cong2 conv1 conv2 ct = Thm.combination (Thm.combination (Thm.reflexive (Thm.dest_fun2 ct)) (conv1 (Thm.dest_arg1 ct))) (conv2 (Thm.dest_arg ct)); fun cong2' conv conv1 conv2 ct = let val eqn1 = conv1 (Thm.dest_arg1 ct); val eqn2 = conv2 (Thm.dest_arg ct) in Thm.transitive (Thm.combination (Thm.combination (Thm.reflexive (Thm.dest_fun2 ct)) eqn1) eqn2) (conv (Thm.rhs_of eqn1) (Thm.rhs_of eqn2)) end; fun cong2'' conv eqn1 eqn2 = let val eqn3 = conv (Thm.rhs_of eqn1) (Thm.rhs_of eqn2) in Thm.transitive (Thm.combination (Thm.combination (Thm.reflexive (Thm.dest_fun2 (Thm.lhs_of eqn3))) eqn1) eqn2) eqn3 end; fun args1 conv ct = conv (Thm.dest_arg ct); fun args2 conv ct = conv (Thm.dest_arg1 ct) (Thm.dest_arg ct); \ ML \ fun strip_numeral ct = (case strip_app ct of (\<^const_name>\uminus\, [n]) => (case strip_app n of (\<^const_name>\numeral\, [b]) => (\<^const_name>\uminus\, [b]) | _ => ("", [])) | x => x); \ lemma nat_minus1_eq: "nat (- 1) = 0" by simp ML \ fun nat_conv i = (case strip_app i of (\<^const_name>\zero_class.zero\, []) => @{thm nat_0 [meta]} | (\<^const_name>\one_class.one\, []) => @{thm nat_one_as_int [meta, symmetric]} | (\<^const_name>\numeral\, [b]) => inst [] [b] @{thm nat_numeral [meta]} | (\<^const_name>\uminus\, [b]) => (case strip_app b of (\<^const_name>\one_class.one\, []) => @{thm nat_minus1_eq [meta]} | (\<^const_name>\numeral\, [b']) => inst [] [b'] @{thm nat_neg_numeral [meta]})); \ ML \ fun add_num_conv b b' = (case (strip_app b, strip_app b') of ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.One\, [])) => @{thm add_num_simps(1) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit0\, [n])) => inst [] [n] @{thm add_num_simps(2) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit1\, [n])) => transitive' (inst [] [n] @{thm add_num_simps(3) [meta]}) (cong1 (args2 add_num_conv)) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm add_num_simps(4) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit0\, [n])) => transitive' (inst [] [m, n] @{thm add_num_simps(5) [meta]}) (cong1 (args2 add_num_conv)) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit1\, [n])) => transitive' (inst [] [m, n] @{thm add_num_simps(6) [meta]}) (cong1 (args2 add_num_conv)) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.One\, [])) => transitive' (inst [] [m] @{thm add_num_simps(7) [meta]}) (cong1 (args2 add_num_conv)) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit0\, [n])) => transitive' (inst [] [m, n] @{thm add_num_simps(8) [meta]}) (cong1 (args2 add_num_conv)) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit1\, [n])) => transitive' (inst [] [m, n] @{thm add_num_simps(9) [meta]}) (cong1 (cong2' add_num_conv (args2 add_num_conv) Thm.reflexive))); \ ML \ fun BitM_conv m = (case strip_app m of (\<^const_name>\Num.One\, []) => @{thm BitM.simps(1) [meta]} | (\<^const_name>\Num.Bit0\, [n]) => transitive' (inst [] [n] @{thm BitM.simps(2) [meta]}) (cong1 (args1 BitM_conv)) | (\<^const_name>\Num.Bit1\, [n]) => inst [] [n] @{thm BitM.simps(3) [meta]}); \ lemma dbl_neg_numeral: "Num.dbl (- Num.numeral k) = - Num.numeral (Num.Bit0 k)" by simp ML \ fun dbl_conv a = let val dbl_neg_numeral_a = inst [a] [] @{thm dbl_neg_numeral [meta]}; val dbl_0_a = inst [a] [] @{thm dbl_simps(2) [meta]}; val dbl_numeral_a = inst [a] [] @{thm dbl_simps(5) [meta]} in fn n => case strip_numeral n of (\<^const_name>\zero_class.zero\, []) => dbl_0_a | (\<^const_name>\numeral\, [k]) => inst [] [k] dbl_numeral_a | (\<^const_name>\uminus\, [k]) => inst [] [k] dbl_neg_numeral_a end; \ lemma dbl_inc_neg_numeral: "Num.dbl_inc (- Num.numeral k) = - Num.numeral (Num.BitM k)" by simp ML \ fun dbl_inc_conv a = let val dbl_inc_neg_numeral_a = inst [a] [] @{thm dbl_inc_neg_numeral [meta]}; val dbl_inc_0_a = inst [a] [] @{thm dbl_inc_simps(2) [folded numeral_One, meta]}; val dbl_inc_numeral_a = inst [a] [] @{thm dbl_inc_simps(5) [meta]}; in fn n => case strip_numeral n of (\<^const_name>\zero_class.zero\, []) => dbl_inc_0_a | (\<^const_name>\numeral\, [k]) => inst [] [k] dbl_inc_numeral_a | (\<^const_name>\uminus\, [k]) => transitive' (inst [] [k] dbl_inc_neg_numeral_a) (cong1 (cong1 (args1 BitM_conv))) end; \ lemma dbl_dec_neg_numeral: "Num.dbl_dec (- Num.numeral k) = - Num.numeral (Num.Bit1 k)" by simp ML \ fun dbl_dec_conv a = let val dbl_dec_neg_numeral_a = inst [a] [] @{thm dbl_dec_neg_numeral [meta]}; val dbl_dec_0_a = inst [a] [] @{thm dbl_dec_simps(2) [folded numeral_One, meta]}; val dbl_dec_numeral_a = inst [a] [] @{thm dbl_dec_simps(5) [meta]}; in fn n => case strip_numeral n of (\<^const_name>\zero_class.zero\, []) => dbl_dec_0_a | (\<^const_name>\uminus\, [k]) => inst [] [k] dbl_dec_neg_numeral_a | (\<^const_name>\numeral\, [k]) => transitive' (inst [] [k] dbl_dec_numeral_a) (cong1 (args1 BitM_conv)) end; \ ML \ fun sub_conv a = let val [sub_One_One, sub_One_Bit0, sub_One_Bit1, sub_Bit0_One, sub_Bit1_One, sub_Bit0_Bit0, sub_Bit0_Bit1, sub_Bit1_Bit0, sub_Bit1_Bit1] = map (inst [a] []) @{thms sub_num_simps [meta]}; val dbl_conv_a = dbl_conv a; val dbl_inc_conv_a = dbl_inc_conv a; val dbl_dec_conv_a = dbl_dec_conv a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.One\, [])) => sub_One_One | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit0\, [l])) => transitive' (inst [] [l] sub_One_Bit0) (cong1 (cong1 (args1 BitM_conv))) | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit1\, [l])) => inst [] [l] sub_One_Bit1 | ((\<^const_name>\Num.Bit0\, [k]), (\<^const_name>\Num.One\, [])) => transitive' (inst [] [k] sub_Bit0_One) (cong1 (args1 BitM_conv)) | ((\<^const_name>\Num.Bit1\, [k]), (\<^const_name>\Num.One\, [])) => inst [] [k] sub_Bit1_One | ((\<^const_name>\Num.Bit0\, [k]), (\<^const_name>\Num.Bit0\, [l])) => transitive' (inst [] [k, l] sub_Bit0_Bit0) (cong1' dbl_conv_a (args2 conv)) | ((\<^const_name>\Num.Bit0\, [k]), (\<^const_name>\Num.Bit1\, [l])) => transitive' (inst [] [k, l] sub_Bit0_Bit1) (cong1' dbl_dec_conv_a (args2 conv)) | ((\<^const_name>\Num.Bit1\, [k]), (\<^const_name>\Num.Bit0\, [l])) => transitive' (inst [] [k, l] sub_Bit1_Bit0) (cong1' dbl_inc_conv_a (args2 conv)) | ((\<^const_name>\Num.Bit1\, [k]), (\<^const_name>\Num.Bit1\, [l])) => transitive' (inst [] [k, l] sub_Bit1_Bit1) (cong1' dbl_conv_a (args2 conv))) in conv end; \ ML \ fun expand1 a = let val numeral_1_eq_1_a = inst [a] [] @{thm numeral_One [meta, symmetric]} in fn n => case Thm.term_of n of \<^Const_>\one_class.one _\ => numeral_1_eq_1_a | \<^Const_>\uminus _ for \<^Const_>\one_class.one _\\ => Thm.combination (Thm.reflexive (Thm.dest_fun n)) numeral_1_eq_1_a | \<^Const_>\zero_class.zero _\ => Thm.reflexive n | \<^Const_>\numeral _ for _\ => Thm.reflexive n | \<^Const_>\uminus _ for \<^Const_>\numeral _ for _\\ => Thm.reflexive n | _ => err "expand1" n end; fun norm1_eq a = let val numeral_1_eq_1_a = inst [a] [] @{thm numeral_One [meta]} in fn eq => case Thm.term_of (Thm.rhs_of eq) of \<^Const_>\Num.numeral _ for \<^Const_>\Num.One\\ => Thm.transitive eq numeral_1_eq_1_a | \<^Const_>\uminus _ for \<^Const_>\Num.numeral _ for \<^Const_>\Num.One\\\ => Thm.transitive eq (Thm.combination (Thm.reflexive (Thm.dest_fun (Thm.rhs_of eq))) numeral_1_eq_1_a) | _ => eq end; \ ML \ fun plus_conv f a = let val add_0_a = inst [a] [] @{thm add_0 [meta]}; val add_0_right_a = inst [a] [] @{thm add_0_right [meta]}; val numeral_plus_numeral_a = inst [a] [] @{thm numeral_plus_numeral [meta]}; val expand1_a = expand1 a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), _) => inst [] [n] add_0_a | (_, (\<^const_name>\zero_class.zero\, [])) => inst [] [m] add_0_right_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => transitive' (inst [] [m, n] numeral_plus_numeral_a) (cong1 (args2 add_num_conv)) | _ => cong2'' (f conv) (expand1_a m) (expand1_a n)) in f conv end; val nat_plus_conv = plus_conv I \<^ctyp>\nat\; \ lemma neg_numeral_plus_neg_numeral: "- Num.numeral m + - Num.numeral n = (- Num.numeral (m + n) ::'a::neg_numeral)" by simp ML \ fun plus_neg_conv a = let val numeral_plus_neg_numeral_a = inst [a] [] @{thm add_neg_numeral_simps(1) [meta]}; val neg_numeral_plus_numeral_a = inst [a] [] @{thm add_neg_numeral_simps(2) [meta]}; val neg_numeral_plus_neg_numeral_a = inst [a] [] @{thm neg_numeral_plus_neg_numeral [meta]}; val sub_conv_a = sub_conv a; in fn conv => fn m => fn n => case (strip_numeral m, strip_numeral n) of ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\uminus\, [n])) => Thm.transitive (inst [] [m, n] numeral_plus_neg_numeral_a) (sub_conv_a m n) | ((\<^const_name>\uminus\, [m]), (\<^const_name>\Num.numeral\, [n])) => Thm.transitive (inst [] [m, n] neg_numeral_plus_numeral_a) (sub_conv_a n m) | ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => transitive' (inst [] [m, n] neg_numeral_plus_neg_numeral_a) (cong1 (cong1 (args2 add_num_conv))) | _ => conv m n end; fun plus_conv' a = norm1_eq a oo plus_conv (plus_neg_conv a) a; val int_plus_conv = plus_conv' \<^ctyp>\int\; \ lemma minus_one: "- 1 = - 1" by simp lemma minus_numeral: "- numeral b = - numeral b" by simp ML \ fun uminus_conv a = let val minus_zero_a = inst [a] [] @{thm minus_zero [meta]}; val minus_one_a = inst [a] [] @{thm minus_one [meta]}; val minus_numeral_a = inst [a] [] @{thm minus_numeral [meta]}; val minus_minus_a = inst [a] [] @{thm minus_minus [meta]} in fn n => case strip_app n of (\<^const_name>\zero_class.zero\, []) => minus_zero_a | (\<^const_name>\one_class.one\, []) => minus_one_a | (\<^const_name>\Num.numeral\, [m]) => inst [] [m] minus_numeral_a | (\<^const_name>\uminus\, [m]) => inst [] [m] minus_minus_a end; val int_neg_conv = uminus_conv \<^ctyp>\int\; \ ML \ fun minus_conv a = let val [numeral_minus_numeral_a, numeral_minus_neg_numeral_a, neg_numeral_minus_numeral_a, neg_numeral_minus_neg_numeral_a] = map (inst [a] []) @{thms diff_numeral_simps [meta]}; val diff_0_a = inst [a] [] @{thm diff_0 [meta]}; val diff_0_right_a = inst [a] [] @{thm diff_0_right [meta]}; val sub_conv_a = sub_conv a; val uminus_conv_a = uminus_conv a; val expand1_a = expand1 a; val norm1_eq_a = norm1_eq a; fun conv m n = (case (strip_numeral m, strip_numeral n) of ((\<^const_name>\zero_class.zero\, []), _) => Thm.transitive (inst [] [n] diff_0_a) (uminus_conv_a n) | (_, (\<^const_name>\zero_class.zero\, [])) => inst [] [m] diff_0_right_a | ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\Num.numeral\, [n])) => Thm.transitive (inst [] [m, n] numeral_minus_numeral_a) (sub_conv_a m n) | ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\uminus\, [n])) => transitive' (inst [] [m, n] numeral_minus_neg_numeral_a) (cong1 (args2 add_num_conv)) | ((\<^const_name>\uminus\, [m]), (\<^const_name>\Num.numeral\, [n])) => transitive' (inst [] [m, n] neg_numeral_minus_numeral_a) (cong1 (cong1 (args2 add_num_conv))) | ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => Thm.transitive (inst [] [m, n] neg_numeral_minus_neg_numeral_a) (sub_conv_a n m) | _ => cong2'' conv (expand1_a m) (expand1_a n)) in norm1_eq_a oo conv end; val int_minus_conv = minus_conv \<^ctyp>\int\; \ ML \ val int_numeral = Thm.apply \<^cterm>\numeral :: num \ int\; val nat_minus_refl = Thm.reflexive \<^cterm>\minus :: nat \ nat \ nat\; val expand1_nat = expand1 \<^ctyp>\nat\; fun nat_minus_conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), _) => inst [] [n] @{thm diff_0_eq_0 [meta]} | (_, (\<^const_name>\zero_class.zero\, [])) => inst [] [m] @{thm minus_nat.diff_0 [meta]} | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => transitive' (inst [] [m, n] @{thm diff_nat_numeral [meta]}) (cong1' nat_conv (args2 int_minus_conv)) | _ => cong2'' nat_minus_conv (expand1_nat m) (expand1_nat n)); \ ML \ fun mult_num_conv m n = (case (strip_app m, strip_app n) of (_, (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm mult_num_simps(1) [meta]} | ((\<^const_name>\Num.One\, []), _) => inst [] [n] @{thm mult_num_simps(2) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit0\, [n])) => transitive' (inst [] [m, n] @{thm mult_num_simps(3) [meta]}) (cong1 (cong1 (args2 mult_num_conv))) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit1\, [n'])) => transitive' (inst [] [m, n'] @{thm mult_num_simps(4) [meta]}) (cong1 (args2 mult_num_conv)) | ((\<^const_name>\Num.Bit1\, [m']), (\<^const_name>\Num.Bit0\, [n])) => transitive' (inst [] [m', n] @{thm mult_num_simps(5) [meta]}) (cong1 (args2 mult_num_conv)) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit1\, [n])) => transitive' (inst [] [m, n] @{thm mult_num_simps(6) [meta]}) (cong1 (cong2' add_num_conv (args2 add_num_conv) (cong1 (args2 mult_num_conv))))); \ ML \ fun mult_conv f a = let val mult_zero_left_a = inst [a] [] @{thm mult_zero_left [meta]}; val mult_zero_right_a = inst [a] [] @{thm mult_zero_right [meta]}; val numeral_times_numeral_a = inst [a] [] @{thm numeral_times_numeral [meta]}; val expand1_a = expand1 a; val norm1_eq_a = norm1_eq a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), _) => inst [] [n] mult_zero_left_a | (_, (\<^const_name>\zero_class.zero\, [])) => inst [] [m] mult_zero_right_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => transitive' (inst [] [m, n] numeral_times_numeral_a) (cong1 (args2 mult_num_conv)) | _ => cong2'' (f conv) (expand1_a m) (expand1_a n)) in norm1_eq_a oo f conv end; val nat_mult_conv = mult_conv I \<^ctyp>\nat\; \ ML \ fun mult_neg_conv a = let val [neg_numeral_times_neg_numeral_a, neg_numeral_times_numeral_a, numeral_times_neg_numeral_a] = map (inst [a] []) @{thms mult_neg_numeral_simps [meta]}; in fn conv => fn m => fn n => case (strip_numeral m, strip_numeral n) of ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => transitive' (inst [] [m, n] neg_numeral_times_neg_numeral_a) (cong1 (args2 mult_num_conv)) | ((\<^const_name>\uminus\, [m]), (\<^const_name>\numeral\, [n])) => transitive' (inst [] [m, n] neg_numeral_times_numeral_a) (cong1 (cong1 (args2 mult_num_conv))) | ((\<^const_name>\numeral\, [m]), (\<^const_name>\uminus\, [n])) => transitive' (inst [] [m, n] numeral_times_neg_numeral_a) (cong1 (cong1 (args2 mult_num_conv))) | _ => conv m n end; fun mult_conv' a = mult_conv (mult_neg_conv a) a; val int_mult_conv = mult_conv' \<^ctyp>\int\; \ ML \ fun eq_num_conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.One\, [])) => @{thm eq_num_simps(1) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit0\, [n])) => inst [] [n] @{thm eq_num_simps(2) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit1\, [n])) => inst [] [n] @{thm eq_num_simps(3) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm eq_num_simps(4) [meta]} | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm eq_num_simps(5) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit0\, [n])) => Thm.transitive (inst [] [m, n] @{thm eq_num_simps(6) [meta]}) (eq_num_conv m n) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit1\, [n])) => inst [] [m, n] @{thm eq_num_simps(7) [meta]} | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit0\, [n])) => inst [] [m, n] @{thm eq_num_simps(8) [meta]} | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit1\, [n])) => Thm.transitive (inst [] [m, n] @{thm eq_num_simps(9) [meta]}) (eq_num_conv m n)); \ ML \ fun eq_conv f a = let val zero_eq_zero_a = inst [a] [] @{thm refl [of 0, THEN Eq_TrueI]}; val zero_neq_numeral_a = inst [a] [] @{thm zero_neq_numeral [THEN Eq_FalseI]}; val numeral_neq_zero_a = inst [a] [] @{thm numeral_neq_zero [THEN Eq_FalseI]}; val numeral_eq_iff_a = inst [a] [] @{thm numeral_eq_iff [meta]}; val expand1_a = expand1 a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\zero_class.zero\, [])) => zero_eq_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\numeral\, [n])) => inst [] [n] zero_neq_numeral_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] numeral_neq_zero_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => Thm.transitive (inst [] [m, n] numeral_eq_iff_a) (eq_num_conv m n) | _ => cong2'' (f conv) (expand1_a m) (expand1_a n)) in f conv end; val nat_eq_conv = eq_conv I \<^ctyp>\nat\; \ ML \ fun eq_neg_conv a = let val neg_numeral_neq_zero_a = inst [a] [] @{thm neg_numeral_neq_zero [THEN Eq_FalseI]}; val zero_neq_neg_numeral_a = inst [a] [] @{thm zero_neq_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_neq_numeral_a = inst [a] [] @{thm neg_numeral_neq_numeral [THEN Eq_FalseI]}; val numeral_neq_neg_numeral_a = inst [a] [] @{thm numeral_neq_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_eq_iff_a = inst [a] [] @{thm neg_numeral_eq_iff [meta]} in fn conv => fn m => fn n => case (strip_numeral m, strip_numeral n) of ((\<^const_name>\uminus\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] neg_numeral_neq_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\uminus\, [n])) => inst [] [n] zero_neq_neg_numeral_a | ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\uminus\, [n])) => inst [] [m, n] numeral_neq_neg_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\Num.numeral\, [n])) => inst [] [m, n] neg_numeral_neq_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => Thm.transitive (inst [] [m, n] neg_numeral_eq_iff_a) (eq_num_conv m n) | _ => conv m n end; fun eq_conv' a = eq_conv (eq_neg_conv a) a; val int_eq_conv = eq_conv' \<^ctyp>\int\; \ ML \ fun le_num_conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\Num.One\, []), _) => inst [] [n] @{thm le_num_simps(1) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm le_num_simps(2) [meta]} | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm le_num_simps(3) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit0\, [n])) => Thm.transitive (inst [] [m, n] @{thm le_num_simps(4) [meta]}) (le_num_conv m n) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit1\, [n])) => Thm.transitive (inst [] [m, n] @{thm le_num_simps(5) [meta]}) (le_num_conv m n) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit1\, [n])) => Thm.transitive (inst [] [m, n] @{thm le_num_simps(6) [meta]}) (le_num_conv m n) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit0\, [n])) => Thm.transitive (inst [] [m, n] @{thm le_num_simps(7) [meta]}) (less_num_conv m n)) and less_num_conv m n = (case (strip_app m, strip_app n) of (_, (\<^const_name>\Num.One\, [])) => inst [] [m] @{thm less_num_simps(1) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit0\, [n])) => inst [] [n] @{thm less_num_simps(2) [meta]} | ((\<^const_name>\Num.One\, []), (\<^const_name>\Num.Bit1\, [n])) => inst [] [n] @{thm less_num_simps(3) [meta]} | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit0\, [n])) => Thm.transitive (inst [] [m, n] @{thm less_num_simps(4) [meta]}) (less_num_conv m n) | ((\<^const_name>\Num.Bit0\, [m]), (\<^const_name>\Num.Bit1\, [n])) => Thm.transitive (inst [] [m, n] @{thm less_num_simps(5) [meta]}) (le_num_conv m n) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit1\, [n])) => Thm.transitive (inst [] [m, n] @{thm less_num_simps(6) [meta]}) (less_num_conv m n) | ((\<^const_name>\Num.Bit1\, [m]), (\<^const_name>\Num.Bit0\, [n])) => Thm.transitive (inst [] [m, n] @{thm less_num_simps(7) [meta]}) (less_num_conv m n)); \ ML \ fun le_conv f a = let val zero_le_zero_a = inst [a] [] @{thm order_refl [of 0, THEN Eq_TrueI]}; val zero_le_numeral_a = inst [a] [] @{thm zero_le_numeral [THEN Eq_TrueI]}; val not_numeral_le_zero_a = inst [a] [] @{thm not_numeral_le_zero [THEN Eq_FalseI]}; val numeral_le_iff_a = inst [a] [] @{thm numeral_le_iff [meta]}; val expand1_a = expand1 a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\zero_class.zero\, [])) => zero_le_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\numeral\, [n])) => inst [] [n] zero_le_numeral_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] not_numeral_le_zero_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => Thm.transitive (inst [] [m, n] numeral_le_iff_a) (le_num_conv m n) | _ => cong2'' (f conv) (expand1_a m) (expand1_a n)) in f conv end; val nat_le_conv = le_conv I \<^ctyp>\nat\; \ ML \ fun le_neg_conv a = let val neg_numeral_le_zero_a = inst [a] [] @{thm neg_numeral_le_zero [THEN Eq_TrueI]}; val not_zero_le_neg_numeral_a = inst [a] [] @{thm not_zero_le_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_le_numeral_a = inst [a] [] @{thm neg_numeral_le_numeral [THEN Eq_TrueI]}; val not_numeral_le_neg_numeral_a = inst [a] [] @{thm not_numeral_le_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_le_iff_a = inst [a] [] @{thm neg_numeral_le_iff [meta]} in fn conv => fn m => fn n => case (strip_numeral m, strip_numeral n) of ((\<^const_name>\uminus\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] neg_numeral_le_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\uminus\, [n])) => inst [] [n] not_zero_le_neg_numeral_a | ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\uminus\, [n])) => inst [] [m, n] not_numeral_le_neg_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\Num.numeral\, [n])) => inst [] [m, n] neg_numeral_le_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => Thm.transitive (inst [] [m, n] neg_numeral_le_iff_a) (le_num_conv n m) | _ => conv m n end; fun le_conv' a = le_conv (le_neg_conv a) a; val int_le_conv = le_conv' \<^ctyp>\int\; \ ML \ fun less_conv f a = let val not_zero_less_zero_a = inst [a] [] @{thm less_irrefl [of 0, THEN Eq_FalseI]}; val zero_less_numeral_a = inst [a] [] @{thm zero_less_numeral [THEN Eq_TrueI]}; val not_numeral_less_zero_a = inst [a] [] @{thm not_numeral_less_zero [THEN Eq_FalseI]}; val numeral_less_iff_a = inst [a] [] @{thm numeral_less_iff [meta]}; val expand1_a = expand1 a; fun conv m n = (case (strip_app m, strip_app n) of ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\zero_class.zero\, [])) => not_zero_less_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\numeral\, [n])) => inst [] [n] zero_less_numeral_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] not_numeral_less_zero_a | ((\<^const_name>\numeral\, [m]), (\<^const_name>\numeral\, [n])) => Thm.transitive (inst [] [m, n] numeral_less_iff_a) (less_num_conv m n) | _ => cong2'' (f conv) (expand1_a m) (expand1_a n)) in f conv end; val nat_less_conv = less_conv I \<^ctyp>\nat\; \ ML \ fun less_neg_conv a = let val neg_numeral_less_zero_a = inst [a] [] @{thm neg_numeral_less_zero [THEN Eq_TrueI]}; val not_zero_less_neg_numeral_a = inst [a] [] @{thm not_zero_less_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_less_numeral_a = inst [a] [] @{thm neg_numeral_less_numeral [THEN Eq_TrueI]}; val not_numeral_less_neg_numeral_a = inst [a] [] @{thm not_numeral_less_neg_numeral [THEN Eq_FalseI]}; val neg_numeral_less_iff_a = inst [a] [] @{thm neg_numeral_less_iff [meta]} in fn conv => fn m => fn n => case (strip_numeral m, strip_numeral n) of ((\<^const_name>\uminus\, [m]), (\<^const_name>\zero_class.zero\, [])) => inst [] [m] neg_numeral_less_zero_a | ((\<^const_name>\zero_class.zero\, []), (\<^const_name>\uminus\, [n])) => inst [] [n] not_zero_less_neg_numeral_a | ((\<^const_name>\Num.numeral\, [m]), (\<^const_name>\uminus\, [n])) => inst [] [m, n] not_numeral_less_neg_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\Num.numeral\, [n])) => inst [] [m, n] neg_numeral_less_numeral_a | ((\<^const_name>\uminus\, [m]), (\<^const_name>\uminus\, [n])) => Thm.transitive (inst [] [m, n] neg_numeral_less_iff_a) (less_num_conv n m) | _ => conv m n end; fun less_conv' a = less_conv (less_neg_conv a) a; val int_less_conv = less_conv' \<^ctyp>\int\; \ ML \ fun If_conv a = let val if_True = inst [a] [] @{thm if_True [meta]}; val if_False = inst [a] [] @{thm if_False [meta]} in fn p => fn x => fn y => fn ct => case strip_app ct of (\<^const_name>\If\, [cb, cx, cy]) => let val p_eq = p cb val eq = Thm.combination (Thm.reflexive (Thm.dest_fun (Thm.dest_fun2 ct))) p_eq in case Thm.term_of (Thm.rhs_of p_eq) of \<^Const_>\True\ => let val x_eq = x cx; val cx = Thm.rhs_of x_eq; in Thm.transitive (Thm.combination (Thm.combination eq x_eq) (Thm.reflexive cy)) (inst [] [cx, cy] if_True) end | \<^Const_>\False\ => let val y_eq = y cy; val cy = Thm.rhs_of y_eq; in Thm.transitive (Thm.combination (Thm.combination eq (Thm.reflexive cx)) y_eq) (inst [] [cx, cy] if_False) end | _ => err "If_conv" (Thm.rhs_of p_eq) end end; \ ML \ fun drop_conv a = let val drop_0_a = inst [a] [] @{thm drop_0 [meta]}; val drop_Cons_a = inst [a] [] @{thm drop_Cons' [meta]}; val If_conv_a = If_conv (type_of_eqn drop_0_a); fun conv n ys = (case Thm.term_of n of \<^Const_>\zero_class.zero _\ => inst [] [ys] drop_0_a | _ => (case strip_app ys of (\<^const_name>\Cons\, [x, xs]) => transitive' (inst [] [n, x, xs] drop_Cons_a) (If_conv_a (args2 nat_eq_conv) Thm.reflexive (cong2' conv (args2 nat_minus_conv) Thm.reflexive)))) in conv end; \ ML \ fun nth_conv a = let val nth_Cons_a = inst [a] [] @{thm nth_Cons' [meta]}; val If_conv_a = If_conv a; fun conv ys n = (case strip_app ys of (\<^const_name>\Cons\, [x, xs]) => transitive' (inst [] [x, xs, n] nth_Cons_a) (If_conv_a (args2 nat_eq_conv) Thm.reflexive (cong2' conv Thm.reflexive (args2 nat_minus_conv)))) in conv end; \ end diff --git a/src/HOL/Decision_Procs/Reflective_Field.thy b/src/HOL/Decision_Procs/Reflective_Field.thy --- a/src/HOL/Decision_Procs/Reflective_Field.thy +++ b/src/HOL/Decision_Procs/Reflective_Field.thy @@ -1,930 +1,929 @@ (* Title: HOL/Decision_Procs/Reflective_Field.thy Author: Stefan Berghofer Reducing equalities in fields to equalities in rings. *) theory Reflective_Field imports Commutative_Ring begin datatype fexpr = FCnst int | FVar nat | FAdd fexpr fexpr | FSub fexpr fexpr | FMul fexpr fexpr | FNeg fexpr | FDiv fexpr fexpr | FPow fexpr nat fun (in field) nth_el :: "'a list \ nat \ 'a" where "nth_el [] n = \" | "nth_el (x # xs) 0 = x" | "nth_el (x # xs) (Suc n) = nth_el xs n" lemma (in field) nth_el_Cons: "nth_el (x # xs) n = (if n = 0 then x else nth_el xs (n - 1))" by (cases n) simp_all lemma (in field) nth_el_closed [simp]: "in_carrier xs \ nth_el xs n \ carrier R" by (induct xs n rule: nth_el.induct) (simp_all add: in_carrier_def) primrec (in field) feval :: "'a list \ fexpr \ 'a" where "feval xs (FCnst c) = \c\" | "feval xs (FVar n) = nth_el xs n" | "feval xs (FAdd a b) = feval xs a \ feval xs b" | "feval xs (FSub a b) = feval xs a \ feval xs b" | "feval xs (FMul a b) = feval xs a \ feval xs b" | "feval xs (FNeg a) = \ feval xs a" | "feval xs (FDiv a b) = feval xs a \ feval xs b" | "feval xs (FPow a n) = feval xs a [^] n" lemma (in field) feval_Cnst: "feval xs (FCnst 0) = \" "feval xs (FCnst 1) = \" "feval xs (FCnst (numeral n)) = \numeral n\" by simp_all datatype pexpr = PExpr1 pexpr1 | PExpr2 pexpr2 and pexpr1 = PCnst int | PVar nat | PAdd pexpr pexpr | PSub pexpr pexpr | PNeg pexpr and pexpr2 = PMul pexpr pexpr | PPow pexpr nat lemma pexpr_cases [case_names PCnst PVar PAdd PSub PNeg PMul PPow]: assumes "\c. e = PExpr1 (PCnst c) \ P" "\n. e = PExpr1 (PVar n) \ P" "\e1 e2. e = PExpr1 (PAdd e1 e2) \ P" "\e1 e2. e = PExpr1 (PSub e1 e2) \ P" "\e'. e = PExpr1 (PNeg e') \ P" "\e1 e2. e = PExpr2 (PMul e1 e2) \ P" "\e' n. e = PExpr2 (PPow e' n) \ P" shows P proof (cases e) case (PExpr1 e') then show ?thesis apply (cases e') apply simp_all apply (erule assms)+ done next case (PExpr2 e') then show ?thesis apply (cases e') apply simp_all apply (erule assms)+ done qed lemmas pexpr_cases2 = pexpr_cases [case_product pexpr_cases] fun (in field) peval :: "'a list \ pexpr \ 'a" where "peval xs (PExpr1 (PCnst c)) = \c\" | "peval xs (PExpr1 (PVar n)) = nth_el xs n" | "peval xs (PExpr1 (PAdd a b)) = peval xs a \ peval xs b" | "peval xs (PExpr1 (PSub a b)) = peval xs a \ peval xs b" | "peval xs (PExpr1 (PNeg a)) = \ peval xs a" | "peval xs (PExpr2 (PMul a b)) = peval xs a \ peval xs b" | "peval xs (PExpr2 (PPow a n)) = peval xs a [^] n" lemma (in field) peval_Cnst: "peval xs (PExpr1 (PCnst 0)) = \" "peval xs (PExpr1 (PCnst 1)) = \" "peval xs (PExpr1 (PCnst (numeral n))) = \numeral n\" "peval xs (PExpr1 (PCnst (- numeral n))) = \ \numeral n\" by simp_all lemma (in field) peval_closed [simp]: "in_carrier xs \ peval xs e \ carrier R" "in_carrier xs \ peval xs (PExpr1 e1) \ carrier R" "in_carrier xs \ peval xs (PExpr2 e2) \ carrier R" by (induct e and e1 and e2) simp_all definition npepow :: "pexpr \ nat \ pexpr" where "npepow e n = (if n = 0 then PExpr1 (PCnst 1) else if n = 1 then e else (case e of PExpr1 (PCnst c) \ PExpr1 (PCnst (c ^ n)) | _ \ PExpr2 (PPow e n)))" lemma (in field) npepow_correct: "in_carrier xs \ peval xs (npepow e n) = peval xs (PExpr2 (PPow e n))" by (cases e rule: pexpr_cases) (simp_all add: npepow_def) fun npemul :: "pexpr \ pexpr \ pexpr" where "npemul x y = (case x of PExpr1 (PCnst c) \ if c = 0 then x else if c = 1 then y else (case y of PExpr1 (PCnst d) \ PExpr1 (PCnst (c * d)) | _ \ PExpr2 (PMul x y)) | PExpr2 (PPow e1 n) \ (case y of PExpr2 (PPow e2 m) \ if n = m then npepow (npemul e1 e2) n else PExpr2 (PMul x y) | PExpr1 (PCnst d) \ if d = 0 then y else if d = 1 then x else PExpr2 (PMul x y) | _ \ PExpr2 (PMul x y)) | _ \ (case y of PExpr1 (PCnst d) \ if d = 0 then y else if d = 1 then x else PExpr2 (PMul x y) | _ \ PExpr2 (PMul x y)))" lemma (in field) npemul_correct: "in_carrier xs \ peval xs (npemul e1 e2) = peval xs (PExpr2 (PMul e1 e2))" proof (induct e1 e2 rule: npemul.induct) case (1 x y) then show ?case proof (cases x y rule: pexpr_cases2) case (PPow_PPow e n e' m) then show ?thesis by (simp del: npemul.simps add: 1 npepow_correct nat_pow_distrib npemul.simps [of "PExpr2 (PPow e n)" "PExpr2 (PPow e' m)"]) qed simp_all qed declare npemul.simps [simp del] definition npeadd :: "pexpr \ pexpr \ pexpr" where "npeadd x y = (case x of PExpr1 (PCnst c) \ if c = 0 then y else (case y of PExpr1 (PCnst d) \ PExpr1 (PCnst (c + d)) | _ \ PExpr1 (PAdd x y)) | _ \ (case y of PExpr1 (PCnst d) \ if d = 0 then x else PExpr1 (PAdd x y) | _ \ PExpr1 (PAdd x y)))" lemma (in field) npeadd_correct: "in_carrier xs \ peval xs (npeadd e1 e2) = peval xs (PExpr1 (PAdd e1 e2))" by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npeadd_def) definition npesub :: "pexpr \ pexpr \ pexpr" where "npesub x y = (case y of PExpr1 (PCnst d) \ if d = 0 then x else (case x of PExpr1 (PCnst c) \ PExpr1 (PCnst (c - d)) | _ \ PExpr1 (PSub x y)) | _ \ (case x of PExpr1 (PCnst c) \ if c = 0 then PExpr1 (PNeg y) else PExpr1 (PSub x y) | _ \ PExpr1 (PSub x y)))" lemma (in field) npesub_correct: "in_carrier xs \ peval xs (npesub e1 e2) = peval xs (PExpr1 (PSub e1 e2))" by (cases e1 e2 rule: pexpr_cases2) (simp_all add: npesub_def) definition npeneg :: "pexpr \ pexpr" where "npeneg e = (case e of PExpr1 (PCnst c) \ PExpr1 (PCnst (- c)) | _ \ PExpr1 (PNeg e))" lemma (in field) npeneg_correct: "peval xs (npeneg e) = peval xs (PExpr1 (PNeg e))" by (cases e rule: pexpr_cases) (simp_all add: npeneg_def) lemma option_pair_cases [case_names None Some]: obtains (None) "x = None" | (Some) p q where "x = Some (p, q)" proof (cases x) case None then show ?thesis by (rule that) next case (Some r) then show ?thesis by (cases r, simp) (rule that) qed fun isin :: "pexpr \ nat \ pexpr \ nat \ (nat \ pexpr) option" where "isin e n (PExpr2 (PMul e1 e2)) m = (case isin e n e1 m of Some (k, e3) \ if k = 0 then Some (0, npemul e3 (npepow e2 m)) else (case isin e k e2 m of Some (l, e4) \ Some (l, npemul e3 e4) | None \ Some (k, npemul e3 (npepow e2 m))) | None \ (case isin e n e2 m of Some (k, e3) \ Some (k, npemul (npepow e1 m) e3) | None \ None))" | "isin e n (PExpr2 (PPow e' k)) m = (if k = 0 then None else isin e n e' (k * m))" | "isin (PExpr1 e) n (PExpr1 e') m = (if e = e' then if n \ m then Some (n - m, PExpr1 (PCnst 1)) else Some (0, npepow (PExpr1 e) (m - n)) else None)" | "isin (PExpr2 e) n (PExpr1 e') m = None" lemma (in field) isin_correct: assumes "in_carrier xs" and "isin e n e' m = Some (p, e'')" shows "peval xs (PExpr2 (PPow e' m)) = peval xs (PExpr2 (PMul (PExpr2 (PPow e (n - p))) e''))" and "p \ n" using assms by (induct e n e' m arbitrary: p e'' rule: isin.induct) (force simp add: nat_pow_distrib nat_pow_pow nat_pow_mult m_ac npemul_correct npepow_correct split: option.split_asm prod.split_asm if_split_asm)+ lemma (in field) isin_correct': "in_carrier xs \ isin e n e' 1 = Some (p, e'') \ peval xs e' = peval xs e [^] (n - p) \ peval xs e''" "in_carrier xs \ isin e n e' 1 = Some (p, e'') \ p \ n" using isin_correct [where m = 1] by simp_all fun split_aux :: "pexpr \ nat \ pexpr \ pexpr \ pexpr \ pexpr" where "split_aux (PExpr2 (PMul e1 e2)) n e = (let (left1, common1, right1) = split_aux e1 n e; (left2, common2, right2) = split_aux e2 n right1 in (npemul left1 left2, npemul common1 common2, right2))" | "split_aux (PExpr2 (PPow e' m)) n e = (if m = 0 then (PExpr1 (PCnst 1), PExpr1 (PCnst 1), e) else split_aux e' (m * n) e)" | "split_aux (PExpr1 e') n e = (case isin (PExpr1 e') n e 1 of Some (m, e'') \ (if m = 0 then (PExpr1 (PCnst 1), npepow (PExpr1 e') n, e'') else (npepow (PExpr1 e') m, npepow (PExpr1 e') (n - m), e'')) | None \ (npepow (PExpr1 e') n, PExpr1 (PCnst 1), e))" hide_const Left Right (* FIXME !? *) abbreviation Left :: "pexpr \ pexpr \ pexpr" where "Left e1 e2 \ fst (split_aux e1 (Suc 0) e2)" abbreviation Common :: "pexpr \ pexpr \ pexpr" where "Common e1 e2 \ fst (snd (split_aux e1 (Suc 0) e2))" abbreviation Right :: "pexpr \ pexpr \ pexpr" where "Right e1 e2 \ snd (snd (split_aux e1 (Suc 0) e2))" lemma split_aux_induct [case_names 1 2 3]: assumes I1: "\e1 e2 n e. P e1 n e \ P e2 n (snd (snd (split_aux e1 n e))) \ P (PExpr2 (PMul e1 e2)) n e" and I2: "\e' m n e. (m \ 0 \ P e' (m * n) e) \ P (PExpr2 (PPow e' m)) n e" and I3: "\e' n e. P (PExpr1 e') n e" shows "P x y z" proof (induct x y z rule: split_aux.induct) case 1 from 1(1) 1(2) [OF refl prod.collapse prod.collapse] show ?case by (rule I1) next case 2 then show ?case by (rule I2) next case 3 then show ?case by (rule I3) qed lemma (in field) split_aux_correct: "in_carrier xs \ peval xs (PExpr2 (PPow e\<^sub>1 n)) = peval xs (PExpr2 (PMul (fst (split_aux e\<^sub>1 n e\<^sub>2)) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))" "in_carrier xs \ peval xs e\<^sub>2 = peval xs (PExpr2 (PMul (snd (snd (split_aux e\<^sub>1 n e\<^sub>2))) (fst (snd (split_aux e\<^sub>1 n e\<^sub>2)))))" by (induct e\<^sub>1 n e\<^sub>2 rule: split_aux_induct) (auto simp add: split_beta nat_pow_distrib nat_pow_pow nat_pow_mult m_ac npemul_correct npepow_correct isin_correct' split: option.split) lemma (in field) split_aux_correct': "in_carrier xs \ peval xs e\<^sub>1 = peval xs (Left e\<^sub>1 e\<^sub>2) \ peval xs (Common e\<^sub>1 e\<^sub>2)" "in_carrier xs \ peval xs e\<^sub>2 = peval xs (Right e\<^sub>1 e\<^sub>2) \ peval xs (Common e\<^sub>1 e\<^sub>2)" using split_aux_correct [where n = 1] by simp_all fun fnorm :: "fexpr \ pexpr \ pexpr \ pexpr list" where "fnorm (FCnst c) = (PExpr1 (PCnst c), PExpr1 (PCnst 1), [])" | "fnorm (FVar n) = (PExpr1 (PVar n), PExpr1 (PCnst 1), [])" | "fnorm (FAdd e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left, common, right) = split_aux xd 1 yd in (npeadd (npemul xn right) (npemul yn left), npemul left (npemul right common), List.union xc yc))" | "fnorm (FSub e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left, common, right) = split_aux xd 1 yd in (npesub (npemul xn right) (npemul yn left), npemul left (npemul right common), List.union xc yc))" | "fnorm (FMul e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left1, common1, right1) = split_aux xn 1 yd; (left2, common2, right2) = split_aux yn 1 xd in (npemul left1 left2, npemul right2 right1, List.union xc yc))" | "fnorm (FNeg e) = (let (n, d, c) = fnorm e in (npeneg n, d, c))" | "fnorm (FDiv e1 e2) = (let (xn, xd, xc) = fnorm e1; (yn, yd, yc) = fnorm e2; (left1, common1, right1) = split_aux xn 1 yn; (left2, common2, right2) = split_aux xd 1 yd in (npemul left1 right2, npemul left2 right1, List.insert yn (List.union xc yc)))" | "fnorm (FPow e m) = (let (n, d, c) = fnorm e in (npepow n m, npepow d m, c))" abbreviation Num :: "fexpr \ pexpr" where "Num e \ fst (fnorm e)" abbreviation Denom :: "fexpr \ pexpr" where "Denom e \ fst (snd (fnorm e))" abbreviation Cond :: "fexpr \ pexpr list" where "Cond e \ snd (snd (fnorm e))" primrec (in field) nonzero :: "'a list \ pexpr list \ bool" where "nonzero xs [] \ True" | "nonzero xs (p # ps) \ peval xs p \ \ \ nonzero xs ps" lemma (in field) nonzero_singleton: "nonzero xs [p] = (peval xs p \ \)" by simp lemma (in field) nonzero_append: "nonzero xs (ps @ qs) = (nonzero xs ps \ nonzero xs qs)" by (induct ps) simp_all lemma (in field) nonzero_idempotent: "p \ set ps \ (peval xs p \ \ \ nonzero xs ps) = nonzero xs ps" by (induct ps) auto lemma (in field) nonzero_insert: "nonzero xs (List.insert p ps) = (peval xs p \ \ \ nonzero xs ps)" by (simp add: List.insert_def nonzero_idempotent) lemma (in field) nonzero_union: "nonzero xs (List.union ps qs) = (nonzero xs ps \ nonzero xs qs)" by (induct ps rule: rev_induct) (auto simp add: List.union_def nonzero_insert nonzero_append) lemma (in field) fnorm_correct: assumes "in_carrier xs" and "nonzero xs (Cond e)" shows "feval xs e = peval xs (Num e) \ peval xs (Denom e)" and "peval xs (Denom e) \ \" using assms proof (induct e) case (FCnst c) { case 1 show ?case by simp next case 2 show ?case by simp } next case (FVar n) { case 1 then show ?case by simp next case 2 show ?case by simp } next case (FAdd e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left = "peval xs (Left (Denom e1) (Denom e2))" let ?common = "peval xs (Common (Denom e1) (Denom e2))" let ?right = "peval xs (Right (Denom e1) (Denom e2))" from 1 FAdd have "feval xs (FAdd e1 e2) = (?common \ (peval xs (Num e1) \ ?right \ peval xs (Num e2) \ ?left)) \ (?common \ (?left \ (?right \ ?common)))" by (simp add: split_beta split nonzero_union add_frac_eq r_distr m_ac) also from 1 FAdd have "\ = peval xs (Num (FAdd e1 e2)) \ peval xs (Denom (FAdd e1 e2))" by (simp add: split_beta split nonzero_union npeadd_correct npemul_correct integral_iff) finally show ?case . next case 2 with FAdd show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FSub e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left = "peval xs (Left (Denom e1) (Denom e2))" let ?common = "peval xs (Common (Denom e1) (Denom e2))" let ?right = "peval xs (Right (Denom e1) (Denom e2))" from 1 FSub have "feval xs (FSub e1 e2) = (?common \ (peval xs (Num e1) \ ?right \ peval xs (Num e2) \ ?left)) \ (?common \ (?left \ (?right \ ?common)))" by (simp add: split_beta split nonzero_union diff_frac_eq r_diff_distr m_ac) also from 1 FSub have "\ = peval xs (Num (FSub e1 e2)) \ peval xs (Denom (FSub e1 e2))" by (simp add: split_beta split nonzero_union npesub_correct npemul_correct integral_iff) finally show ?case . next case 2 with FSub show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FMul e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Denom e2"] split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e2" and e\<^sub>2 = "Denom e1"] { case 1 let ?left\<^sub>1 = "peval xs (Left (Num e1) (Denom e2))" let ?common\<^sub>1 = "peval xs (Common (Num e1) (Denom e2))" let ?right\<^sub>1 = "peval xs (Right (Num e1) (Denom e2))" let ?left\<^sub>2 = "peval xs (Left (Num e2) (Denom e1))" let ?common\<^sub>2 = "peval xs (Common (Num e2) (Denom e1))" let ?right\<^sub>2 = "peval xs (Right (Num e2) (Denom e1))" from 1 FMul have "feval xs (FMul e1 e2) = ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>1 \ ?left\<^sub>2)) \ ((?common\<^sub>1 \ ?common\<^sub>2) \ (?right\<^sub>2 \ ?right\<^sub>1))" by (simp add: split_beta split nonzero_union nonzero_divide_divide_eq_left m_ac) also from 1 FMul have "\ = peval xs (Num (FMul e1 e2)) \ peval xs (Denom (FMul e1 e2))" by (simp add: split_beta split nonzero_union npemul_correct integral_iff) finally show ?case . next case 2 with FMul show ?case by (simp add: split_beta split nonzero_union npemul_correct integral_iff) } next case (FNeg e) { case 1 with FNeg show ?case by (simp add: split_beta npeneg_correct) next case 2 with FNeg show ?case by (simp add: split_beta) } next case (FDiv e1 e2) note split = split_aux_correct' [where xs=xs and e\<^sub>1 = "Num e1" and e\<^sub>2 = "Num e2"] split_aux_correct' [where xs=xs and e\<^sub>1 = "Denom e1" and e\<^sub>2 = "Denom e2"] { case 1 let ?left\<^sub>1 = "peval xs (Left (Num e1) (Num e2))" let ?common\<^sub>1 = "peval xs (Common (Num e1) (Num e2))" let ?right\<^sub>1 = "peval xs (Right (Num e1) (Num e2))" let ?left\<^sub>2 = "peval xs (Left (Denom e1) (Denom e2))" let ?common\<^sub>2 = "peval xs (Common (Denom e1) (Denom e2))" let ?right\<^sub>2 = "peval xs (Right (Denom e1) (Denom e2))" from 1 FDiv have "feval xs (FDiv e1 e2) = ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>1 \ ?right\<^sub>2)) \ ((?common\<^sub>1 \ ?common\<^sub>2) \ (?left\<^sub>2 \ ?right\<^sub>1))" by (simp add: split_beta split nonzero_union nonzero_insert nonzero_divide_divide_eq m_ac) also from 1 FDiv have "\ = peval xs (Num (FDiv e1 e2)) \ peval xs (Denom (FDiv e1 e2))" by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff) finally show ?case . next case 2 with FDiv show ?case by (simp add: split_beta split nonzero_union nonzero_insert npemul_correct integral_iff) } next case (FPow e n) { case 1 with FPow show ?case by (simp add: split_beta nonzero_power_divide npepow_correct) next case 2 with FPow show ?case by (simp add: split_beta npepow_correct) } qed lemma (in field) feval_eq0: assumes "in_carrier xs" and "fnorm e = (n, d, c)" and "nonzero xs c" and "peval xs n = \" shows "feval xs e = \" using assms fnorm_correct [of xs e] by simp lemma (in field) fexpr_in_carrier: assumes "in_carrier xs" and "nonzero xs (Cond e)" shows "feval xs e \ carrier R" using assms proof (induct e) case (FDiv e1 e2) then have "feval xs e1 \ carrier R" "feval xs e2 \ carrier R" "peval xs (Num e2) \ \" "nonzero xs (Cond e2)" by (simp_all add: nonzero_union nonzero_insert split: prod.split_asm) from \in_carrier xs\ \nonzero xs (Cond e2)\ have "feval xs e2 = peval xs (Num e2) \ peval xs (Denom e2)" by (rule fnorm_correct) moreover from \in_carrier xs\ \nonzero xs (Cond e2)\ have "peval xs (Denom e2) \ \" by (rule fnorm_correct) ultimately have "feval xs e2 \ \" using \peval xs (Num e2) \ \\ \in_carrier xs\ by (simp add: divide_eq_0_iff) with \feval xs e1 \ carrier R\ \feval xs e2 \ carrier R\ show ?case by simp qed (simp_all add: nonzero_union split: prod.split_asm) lemma (in field) feval_eq: assumes "in_carrier xs" and "fnorm (FSub e e') = (n, d, c)" and "nonzero xs c" shows "(feval xs e = feval xs e') = (peval xs n = \)" proof - from assms have "nonzero xs (Cond e)" "nonzero xs (Cond e')" by (auto simp add: nonzero_union split: prod.split_asm) with assms fnorm_correct [of xs "FSub e e'"] have "feval xs e \ feval xs e' = peval xs n \ peval xs d" "peval xs d \ \" by simp_all show ?thesis proof assume "feval xs e = feval xs e'" with \feval xs e \ feval xs e' = peval xs n \ peval xs d\ \in_carrier xs\ \nonzero xs (Cond e')\ have "peval xs n \ peval xs d = \" by (simp add: fexpr_in_carrier minus_eq r_neg) with \peval xs d \ \\ \in_carrier xs\ show "peval xs n = \" by (simp add: divide_eq_0_iff) next assume "peval xs n = \" with \feval xs e \ feval xs e' = peval xs n \ peval xs d\ \peval xs d \ \\ \nonzero xs (Cond e)\ \nonzero xs (Cond e')\ \in_carrier xs\ show "feval xs e = feval xs e'" by (simp add: eq_diff0 fexpr_in_carrier) qed qed ML \ val term_of_nat = HOLogic.mk_number \<^Type>\nat\ o @{code integer_of_nat}; val term_of_int = HOLogic.mk_number \<^Type>\int\ o @{code integer_of_int}; fun term_of_pexpr (@{code PExpr1} x) = \<^Const>\PExpr1\ $ term_of_pexpr1 x | term_of_pexpr (@{code PExpr2} x) = \<^Const>\PExpr2\ $ term_of_pexpr2 x and term_of_pexpr1 (@{code PCnst} k) = \<^Const>\PCnst\ $ term_of_int k | term_of_pexpr1 (@{code PVar} n) = \<^Const>\PVar\ $ term_of_nat n | term_of_pexpr1 (@{code PAdd} (x, y)) = \<^Const>\PAdd\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr1 (@{code PSub} (x, y)) = \<^Const>\PSub\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr1 (@{code PNeg} x) = \<^Const>\PNeg\ $ term_of_pexpr x and term_of_pexpr2 (@{code PMul} (x, y)) = \<^Const>\PMul\ $ term_of_pexpr x $ term_of_pexpr y | term_of_pexpr2 (@{code PPow} (x, n)) = \<^Const>\PPow\ $ term_of_pexpr x $ term_of_nat n fun term_of_result (x, (y, zs)) = HOLogic.mk_prod (term_of_pexpr x, HOLogic.mk_prod (term_of_pexpr y, HOLogic.mk_list \<^Type>\pexpr\ (map term_of_pexpr zs))); local fun fnorm (ctxt, ct, t) = \<^instantiate>\x = ct and y = \Thm.cterm_of ctxt t\ in cterm \x \ y\ for x y :: \pexpr \ pexpr \ pexpr list\\; val (_, raw_fnorm_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\fnorm\, fnorm))); fun fnorm_oracle ctxt ct t = raw_fnorm_oracle (ctxt, ct, t); in val cv = @{computation_conv "pexpr \ pexpr \ pexpr list" terms: fnorm nat_of_integer Code_Target_Nat.natural "0::nat" "1::nat" "2::nat" "3::nat" "0::int" "1::int" "2::int" "3::int" "-1::int" datatypes: fexpr int integer num} (fn ctxt => fn result => fn ct => fnorm_oracle ctxt ct (term_of_result result)) end \ ML \ signature FIELD_TAC = sig structure Field_Simps: sig type T val get: Context.generic -> T val put: T -> Context.generic -> Context.generic val map: (T -> T) -> Context.generic -> Context.generic end val eq_field_simps: (term * (thm list * thm list * thm list * thm * thm)) * (term * (thm list * thm list * thm list * thm * thm)) -> bool val field_tac: bool -> Proof.context -> int -> tactic end structure Field_Tac : FIELD_TAC = struct open Ring_Tac; fun field_struct \<^Const_>\Ring.ring.add _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Ring.a_minus _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Group.monoid.mult _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Ring.a_inv _ _ for R _\ = SOME R | field_struct \<^Const_>\Group.pow _ _ _ for R _ _\ = SOME R | field_struct \<^Const_>\Algebra_Aux.m_div _ _for R _ _\ = SOME R | field_struct \<^Const_>\Ring.ring.zero _ _ for R\ = SOME R | field_struct \<^Const_>\Group.monoid.one _ _ for R\ = SOME R | field_struct \<^Const_>\Algebra_Aux.of_integer _ _ for R _\ = SOME R | field_struct _ = NONE; fun reif_fexpr vs \<^Const_>\Ring.ring.add _ _ for _ a b\ = \<^Const>\FAdd for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Ring.a_minus _ _ for _ a b\ = \<^Const>\FSub for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Group.monoid.mult _ _ for _ a b\ = \<^Const>\FMul for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs \<^Const_>\Ring.a_inv _ _ for _ a\ = \<^Const>\FNeg for \reif_fexpr vs a\\ | reif_fexpr vs \<^Const>\Group.pow _ _ _ for _ a n\ = \<^Const>\FPow for \reif_fexpr vs a\ n\ | reif_fexpr vs \<^Const_>\Algebra_Aux.m_div _ _ for _ a b\ = \<^Const>\FDiv for \reif_fexpr vs a\ \reif_fexpr vs b\\ | reif_fexpr vs (Free x) = \<^Const>\FVar for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_fexpr vs \<^Const_>\Ring.ring.zero _ _ for _\ = \<^term>\FCnst 0\ | reif_fexpr vs \<^Const_>\Group.monoid.one _ _ for _\ = \<^term>\FCnst 1\ | reif_fexpr vs \<^Const_>\Algebra_Aux.of_integer _ _ for _ n\ = \<^Const>\FCnst for n\ | reif_fexpr _ _ = error "reif_fexpr: bad expression"; fun reif_fexpr' vs \<^Const_>\plus _ for a b\ = \<^Const>\FAdd for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\minus _ for a b\ = \<^Const>\FSub for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\times _ for a b\ = \<^Const>\FMul for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs \<^Const_>\uminus _ for a\ = \<^Const>\FNeg for \reif_fexpr' vs a\\ | reif_fexpr' vs \<^Const_>\power _ for a n\ = \<^Const>\FPow for \reif_fexpr' vs a\ n\ | reif_fexpr' vs \<^Const_>\divide _ for a b\ = \<^Const>\FDiv for \reif_fexpr' vs a\ \reif_fexpr' vs b\\ | reif_fexpr' vs (Free x) = \<^Const>\FVar for \HOLogic.mk_number HOLogic.natT (find_index (equal x) vs)\\ | reif_fexpr' vs \<^Const_>\zero_class.zero _\ = \<^term>\FCnst 0\ | reif_fexpr' vs \<^Const_>\one_class.one _\ = \<^term>\FCnst 1\ | reif_fexpr' vs \<^Const_>\numeral _ for b\ = \<^Const>\FCnst for \<^Const>\numeral \<^Type>\int\ for b\\ | reif_fexpr' _ _ = error "reif_fexpr: bad expression"; fun eq_field_simps ((t, (ths1, ths2, ths3, th4, th)), (t', (ths1', ths2', ths3', th4', th'))) = t aconv t' andalso eq_list Thm.eq_thm (ths1, ths1') andalso eq_list Thm.eq_thm (ths2, ths2') andalso eq_list Thm.eq_thm (ths3, ths3') andalso Thm.eq_thm (th4, th4') andalso Thm.eq_thm (th, th'); structure Field_Simps = Generic_Data (struct type T = (term * (thm list * thm list * thm list * thm * thm)) Net.net val empty = Net.empty val merge = Net.merge eq_field_simps end); fun get_field_simps ctxt optcT t = (case get_matching_rules ctxt (Field_Simps.get (Context.Proof ctxt)) t of SOME (ths1, ths2, ths3, th4, th) => let val tr = Thm.transfer' ctxt #> (case optcT of NONE => I | SOME cT => inst [cT] [] #> norm) in (map tr ths1, map tr ths2, map tr ths3, tr th4, tr th) end | NONE => error "get_field_simps: lookup failed"); fun nth_el_conv (_, _, _, nth_el_Cons, _) = let val a = type_of_eqn nth_el_Cons; val If_conv_a = If_conv a; fun conv ys n = (case strip_app ys of (\<^const_name>\Cons\, [x, xs]) => transitive' (inst [] [x, xs, n] nth_el_Cons) (If_conv_a (args2 nat_eq_conv) Thm.reflexive (cong2' conv Thm.reflexive (args2 nat_minus_conv)))) in conv end; fun feval_conv (rls as ([feval_simps_1, feval_simps_2, feval_simps_3, feval_simps_4, feval_simps_5, feval_simps_6, feval_simps_7, feval_simps_8, feval_simps_9, feval_simps_10, feval_simps_11], _, _, _, _)) = let val nth_el_conv' = nth_el_conv rls; fun conv xs x = (case strip_app x of (\<^const_name>\FCnst\, [c]) => (case strip_app c of (\<^const_name>\zero_class.zero\, _) => inst [] [xs] feval_simps_9 | (\<^const_name>\one_class.one\, _) => inst [] [xs] feval_simps_10 | (\<^const_name>\numeral\, [n]) => inst [] [xs, n] feval_simps_11 | _ => inst [] [xs, c] feval_simps_1) | (\<^const_name>\FVar\, [n]) => transitive' (inst [] [xs, n] feval_simps_2) (args2 nth_el_conv') | (\<^const_name>\FAdd\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_3) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FSub\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_4) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FMul\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_5) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FNeg\, [a]) => transitive' (inst [] [xs, a] feval_simps_6) (cong1 (args2 conv)) | (\<^const_name>\FDiv\, [a, b]) => transitive' (inst [] [xs, a, b] feval_simps_7) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\FPow\, [a, n]) => transitive' (inst [] [xs, a, n] feval_simps_8) (cong2 (args2 conv) Thm.reflexive)) in conv end; fun peval_conv (rls as (_, [peval_simps_1, peval_simps_2, peval_simps_3, peval_simps_4, peval_simps_5, peval_simps_6, peval_simps_7, peval_simps_8, peval_simps_9, peval_simps_10, peval_simps_11], _, _, _)) = let val nth_el_conv' = nth_el_conv rls; fun conv xs x = (case strip_app x of (\<^const_name>\PExpr1\, [e]) => (case strip_app e of (\<^const_name>\PCnst\, [c]) => (case strip_numeral c of (\<^const_name>\zero_class.zero\, _) => inst [] [xs] peval_simps_8 | (\<^const_name>\one_class.one\, _) => inst [] [xs] peval_simps_9 | (\<^const_name>\numeral\, [n]) => inst [] [xs, n] peval_simps_10 | (\<^const_name>\uminus\, [n]) => inst [] [xs, n] peval_simps_11 | _ => inst [] [xs, c] peval_simps_1) | (\<^const_name>\PVar\, [n]) => transitive' (inst [] [xs, n] peval_simps_2) (args2 nth_el_conv') | (\<^const_name>\PAdd\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_3) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PSub\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_4) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PNeg\, [a]) => transitive' (inst [] [xs, a] peval_simps_5) (cong1 (args2 conv))) | (\<^const_name>\PExpr2\, [e]) => (case strip_app e of (\<^const_name>\PMul\, [a, b]) => transitive' (inst [] [xs, a, b] peval_simps_6) (cong2 (args2 conv) (args2 conv)) | (\<^const_name>\PPow\, [a, n]) => transitive' (inst [] [xs, a, n] peval_simps_7) (cong2 (args2 conv) Thm.reflexive))) in conv end; fun nonzero_conv (rls as (_, _, [nonzero_Nil, nonzero_Cons, nonzero_singleton], _, _)) = let val peval_conv' = peval_conv rls; fun conv xs qs = (case strip_app qs of (\<^const_name>\Nil\, []) => inst [] [xs] nonzero_Nil | (\<^const_name>\Cons\, [p, ps]) => (case Thm.term_of ps of \<^Const_>\Nil _\ => transitive' (inst [] [xs, p] nonzero_singleton) (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) | _ => transitive' (inst [] [xs, p, ps] nonzero_Cons) (cong2 (cong1 (cong2 (args2 peval_conv') Thm.reflexive)) (args2 conv)))) in conv end; fun field_tac in_prem ctxt = SUBGOAL (fn (g, i) => let val (prems, concl) = Logic.strip_horn g; fun find_eq s = (case s of (_ $ \<^Const_>\HOL.eq T for t u\) => (case (field_struct t, field_struct u) of (SOME R, _) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr) | (_, SOME R) => SOME ((t, u), R, T, NONE, mk_in_carrier ctxt R [], reif_fexpr) | _ => if Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\field\) then SOME ((t, u), mk_ring T, T, SOME T, K @{thm in_carrier_trivial}, reif_fexpr') else NONE) | _ => NONE); val ((t, u), R, T, optT, mkic, reif) = (case get_first find_eq (if in_prem then prems else [concl]) of SOME q => q | NONE => error "cannot determine field"); val rls as (_, _, _, _, feval_eq) = get_field_simps ctxt (Option.map (Thm.ctyp_of ctxt) optT) R; val xs = [] |> Term.add_frees t |> Term.add_frees u |> filter (equal T o snd); val cxs = Thm.cterm_of ctxt (HOLogic.mk_list T (map Free xs)); val ce = Thm.cterm_of ctxt (reif xs t); val ce' = Thm.cterm_of ctxt (reif xs u); val fnorm = cv ctxt \<^instantiate>\e = ce and e' = ce' in cterm \fnorm (FSub e e')\\; val (_, [n, dc]) = strip_app (Thm.rhs_of fnorm); val (_, [_, c]) = strip_app dc; val th = Conv.fconv_rule (Conv.concl_conv 1 (Conv.arg_conv (binop_conv (binop_conv (K (feval_conv rls cxs ce)) (K (feval_conv rls cxs ce'))) (Conv.arg1_conv (K (peval_conv rls cxs n)))))) ([mkic xs, HOLogic.mk_obj_eq fnorm, HOLogic.mk_obj_eq (nonzero_conv rls cxs c) RS @{thm iffD2}] MRS feval_eq); val th' = Drule.rotate_prems 1 (th RS (if in_prem then @{thm iffD1} else @{thm iffD2})); in if in_prem then dresolve_tac ctxt [th'] 1 THEN defer_tac 1 else resolve_tac ctxt [th'] 1 end); end \ context field begin -local_setup \ -Local_Theory.declaration {syntax = false, pervasive = false} - (fn phi => Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps +declaration \fn phi => + Field_Tac.Field_Simps.map (Ring_Tac.insert_rules Field_Tac.eq_field_simps (Morphism.term phi \<^term>\R\, (Morphism.fact phi @{thms feval.simps [meta] feval_Cnst [meta]}, Morphism.fact phi @{thms peval.simps [meta] peval_Cnst [meta]}, Morphism.fact phi @{thms nonzero.simps [meta] nonzero_singleton [meta]}, singleton (Morphism.fact phi) @{thm nth_el_Cons [meta]}, - singleton (Morphism.fact phi) @{thm feval_eq})))) + singleton (Morphism.fact phi) @{thm feval_eq}))) \ end method_setup field = \ Scan.lift (Args.mode "prems") -- Attrib.thms >> (fn (in_prem, thms) => fn ctxt => SIMPLE_METHOD' (Field_Tac.field_tac in_prem ctxt THEN' Ring_Tac.ring_tac in_prem thms ctxt)) \ "reduce equations over fields to equations over rings" end diff --git a/src/HOL/Eisbach/Tests.thy b/src/HOL/Eisbach/Tests.thy --- a/src/HOL/Eisbach/Tests.thy +++ b/src/HOL/Eisbach/Tests.thy @@ -1,606 +1,604 @@ (* Title: HOL/Eisbach/Tests.thy Author: Daniel Matichuk, NICTA/UNSW *) section \Miscellaneous Eisbach tests\ theory Tests imports Main Eisbach_Tools begin subsection \Named Theorems Tests\ named_theorems foo method foo declares foo = (rule foo) lemma assumes A [foo]: A shows A apply foo done method abs_used for P = (match (P) in "\a. ?Q" \ fail \ _ \ -) subsection \Match Tests\ notepad begin have dup: "\A. A \ A \ A" by simp fix A y have "(\x. A x) \ A y" apply (rule dup, match premises in Y: "\B. P B" for P \ \match (P) in A \ \print_fact Y, rule Y\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match (P) in A \ \print_fact Y, rule Y\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match conclusion in "P y" for y \ \print_fact Y, print_term y, rule Y[where B=y]\\) apply (rule dup, match premises in Y: "\B :: 'a. P B" for P \ \match conclusion in "P z" for z \ \print_fact Y, print_term y, rule Y[where B=z]\\) apply (rule dup, match conclusion in "P y" for P \ \match premises in Y: "\z. P z" \ \print_fact Y, rule Y[where z=y]\\) apply (match premises in Y: "\z :: 'a. P z" for P \ \match conclusion in "P y" \ \print_fact Y, rule Y[where z=y]\\) done assume X: "\x. A x" "A y" have "A y" apply (match X in Y:"\B. A B" and Y':"B ?x" for B \ \print_fact Y[where B=y], print_term B\) apply (match X in Y:"B ?x" and Y':"B ?x" for B \ \print_fact Y, print_term B\) apply (match X in Y:"B x" and Y':"B x" for B x \ \print_fact Y, print_term B, print_term x\) apply (insert X) apply (match premises in Y:"\B. A B" and Y':"B y" for B and y :: 'a \ \print_fact Y[where B=y], print_term B\) apply (match premises in Y:"B ?x" and Y':"B ?x" for B \ \print_fact Y, print_term B\) apply (match premises in Y:"B x" and Y':"B x" for B x \ \print_fact Y, print_term B\) apply (match conclusion in "P x" and "P y" for P x \ \print_term P, print_term x\) apply assumption done { fix B x y assume X: "\x y. B x x y" have "B x x y" by (match X in Y:"\y. B y y z" for z \ \rule Y[where y=x]\) fix A B have "(\x y. A (B x) y) \ A (B x) y" by (match premises in Y: "\xx. ?H (B xx)" \ \rule Y\) } (* match focusing retains prems *) fix B x have "(\x. A x) \ (\z. B z) \ A y \ B x" apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y': "\z :: 'b. B z" \ \print_fact Y, print_fact Y', rule Y'[where z=x]\\) done (*Attributes *) fix C have "(\x :: 'a. A x) \ (\z. B z) \ A y \ B x \ B x \ A y" apply (intro conjI) apply (match premises in Y: "\z :: 'a. A z" and Y'[intro]:"\z :: 'b. B z" \ fastforce) apply (match premises in Y: "\z :: 'a. A z" \ \match premises in Y'[intro]:"\z :: 'b. B z" \ fastforce\) apply (match premises in Y[thin]: "\z :: 'a. A z" \ \(match premises in Y':"\z :: 'a. A z" \ \print_fact Y,fail\ \ _ \ \print_fact Y\)\) (*apply (match premises in Y: "\z :: 'b. B z" \ \(match premises in Y'[thin]:"\z :: 'b. B z" \ \(match premises in Y':"\z :: 'a. A z" \ fail \ Y': _ \ -)\)\)*) apply assumption done fix A B C D have "\uu'' uu''' uu uu'. (\x :: 'a. A uu' x) \ D uu y \ (\z. B uu z) \ C uu y \ (\z y. C uu z) \ B uu x \ B uu x \ C uu y" apply (match premises in Y[thin]: "\z :: 'a. A ?zz' z" and Y'[thin]: "\rr :: 'b. B ?zz rr" \ \print_fact Y, print_fact Y', intro conjI, rule Y', insert Y', insert Y'[where rr=x]\) apply (match premises in Y:"B ?u ?x" \ \rule Y\) apply (insert TrueI) apply (match premises in Y'[thin]: "\ff. B uu ff" for uu \ \insert Y', drule meta_spec[where x=x]\) apply assumption done (* Multi-matches. As many facts as match are bound. *) fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ \intro conjI, (rule Y)+\) apply (match premises in Y[thin]: "\z :: 'a. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done fix A B C x have "(\x :: 'a. A x) \ (\y :: 'a. B y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y[thin]: "\z. ?A z" (multi) \ \intro conjI, (rule Y)+\) apply (match premises in Y[thin]: "\z. ?A z" (multi) \ fail \ "C y" \ -) (* multi-match must bind something *) apply (match premises in Y: "C y" \ \rule Y\) done fix A B C P Q and x :: 'a and y :: 'a have "(\x y :: 'a. A x y \ Q) \ (\a b. B (a :: 'a) (b :: 'a) \ Q) \ (\x y. C (x :: 'a) (y :: 'a) \ P) \ A y x \ B y x" by (match premises in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule conjI, rule Y[where z=x,THEN conjunct1], rule Y[THEN conjunct1]\) (*We may use for-fixes in multi-matches too. All bound facts must agree on the fixed term *) fix A B C x have "(\y :: 'a. B y \ C y) \ (\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match (P) in B \ fail \ "\a. B" \ fail \ _ \ -, intro conjI, (rule Y[THEN conjunct1])\) apply (rule dup) apply (match premises in Y':"\x :: 'a. ?U x \ Q x" and Y: "\x :: 'a. Q x \ ?U x" (multi) for Q \ \insert Y[THEN conjunct1]\) apply assumption (* Previous match requires that Q is consistent *) apply (match premises in Y: "\z :: 'a. ?A z \ False" (multi) \ \print_fact Y, fail\ \ "C y" \ \print_term C\) (* multi-match must bind something *) apply (match premises in Y: "\x. B x \ C x" \ \intro conjI Y[THEN conjunct1]\) apply (match premises in Y: "C ?x" \ \rule Y\) done (* All bindings must be tried for a particular theorem. However all combinations are NOT explored. *) fix B A C assume asms:"\a b. B (a :: 'a) (b :: 'a) \ Q" "\x :: 'a. A x x \ Q" "\a b. C (a :: 'a) (b :: 'a) \ Q" have "B y x \ C x y \ B x y \ C y x \ A x x" apply (intro conjI) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where z=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where a=x,THEN conjunct1]\) apply (match asms in Y: "\z a. ?A (z :: 'a) (a :: 'a) \ R" (multi) for R \ \rule Y[where z=x,THEN conjunct1]\) apply (match asms in Y: "\z a. A (z :: 'a) (a :: 'a) \ R" for R \ fail \ _ \ -) apply (rule asms[THEN conjunct1]) done (* Attributes *) fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match Y[THEN conjunct1] in Y':"?H" (multi) \ \intro conjI,rule Y'\\) apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match Y[THEN conjunct2] in Y':"?H" (multi) \ \rule Y'\\) apply assumption done (* Removed feature for now *) (* fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match prems in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match \K @{thms Y TrueI}\ in Y':"?H" (multi) \ \rule conjI; (rule Y')?\\) apply (match prems in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \match \K [@{thm Y}]\ in Y':"?H" (multi) \ \rule Y'\\) done *) (* Testing THEN_ALL_NEW within match *) fix A B C x have "(\x :: 'a. A x \ B x) \ (\y :: 'a. A y \ C y) \ (\y :: 'a. B y \ C y) \ C y \ (A x \ B y \ C y)" apply (match premises in Y: "\x :: 'a. P x \ ?U x" (multi) for P \ \intro conjI ; ((rule Y[THEN conjunct1])?); rule Y[THEN conjunct2] \) done (* Cut tests *) fix A B C D have "D \ C \ A \ B \ A \ C \ D \ True \ C" by (((match premises in I: "P \ Q" (cut) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\)?), simp) have "D \ C \ A \ B \ A \ C \ D \ True \ C" by (match premises in I: "P \ Q" (cut 2) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\) have "A \ B \ A \ C \ C" by (((match premises in I: "P \ Q" (cut) and I': "P \ ?U" for P Q \ \rule mp [OF I' I[THEN conjunct1]]\)?, simp) | simp) fix f x y have "f x y \ f x y" by (match conclusion in "f x y" for f x y \ \print_term f\) fix A B C assume X: "A \ B" "A \ C" C have "A \ B \ C" by (match X in H: "A \ ?H" (multi, cut) \ \match H in "A \ C" and "A \ B" \ fail\ | simp add: X) (* Thinning an inner focus *) (* Thinning should persist within a match, even when on an external premise *) fix A have "(\x. A x \ B) \ B \ C \ C" apply (match premises in H:"\x. A x \ B" \ \match premises in H'[thin]: "\x. A x \ B" \ \match premises in H'':"\x. A x \ B" \ fail \ _ \ -\ ,match premises in H'':"\x. A x \ B" \ fail \ _ \ -\) apply (match premises in H:"\x. A x \ B" \ fail \ H':_ \ \rule H'[THEN conjunct2]\) done (* Local premises *) (* Only match premises which actually existed in the goal we just focused.*) fix A assume asms: "C \ D" have "B \ C \ C" by (match premises in _ \ \insert asms, match premises (local) in "B \ C" \ fail \ H:"C \ D" \ \rule H[THEN conjunct1]\\) end (* Testing inner focusing. This fails if we don't smash flex-flex pairs produced by retrofitting. This needs to be done more carefully to avoid smashing legitimate pairs.*) schematic_goal "?A x \ A x" apply (match conclusion in "H" for H \ \match conclusion in Y for Y \ \print_term Y\\) apply assumption done (* Ensure short-circuit after first match failure *) lemma assumes A: "P \ Q" shows "P" by ((match A in "P \ Q" \ fail \ "?H" \ -) | simp add: A) lemma assumes A: "D \ C" "A \ B" "A \ B" shows "A" apply ((match A in U: "P \ Q" (cut) and "P' \ Q'" for P Q P' Q' \ \simp add: U\ \ "?H" \ -) | -) apply (simp add: A) done subsection \Uses Tests\ ML \ fun test_internal_fact ctxt factnm = (case \<^try>\Proof_Context.get_thms ctxt factnm\ of NONE => () | SOME _ => error "Found internal fact"); \ method uses_test\<^sub>1 uses uses_test\<^sub>1_uses = (rule uses_test\<^sub>1_uses) lemma assumes A shows A by (uses_test\<^sub>1 uses_test\<^sub>1_uses: assms) ML \test_internal_fact \<^context> "uses_test\<^sub>1_uses"\ ML \test_internal_fact \<^context> "Tests.uses_test\<^sub>1_uses"\ ML \test_internal_fact \<^context> "Tests.uses_test\<^sub>1.uses_test\<^sub>1_uses"\ subsection \Basic fact passing\ method find_fact for x y :: bool uses facts1 facts2 = (match facts1 in U: "x" \ \insert U, match facts2 in U: "y" \ \insert U\\) lemma assumes A: A and B: B shows "A \ B" apply (find_fact "A" "B" facts1: A facts2: B) apply (rule conjI; assumption) done subsection \Testing term and fact passing in recursion\ method recursion_example for x :: bool uses facts = (match (x) in "A \ B" for A B \ \(recursion_example A facts: facts, recursion_example B facts: facts)\ \ "?H" \ \match facts in U: "x" \ \insert U\\) lemma assumes asms: "A" "B" "C" "D" shows "(A \ B) \ (C \ D)" apply (recursion_example "(A \ B) \ (C \ D)" facts: asms) apply simp done (* uses facts are not accumulated *) method recursion_example' for A :: bool and B :: bool uses facts = (match facts in H: "A" and H': "B" \ \recursion_example' "A" "B" facts: H TrueI\ \ "A" and "True" \ \recursion_example' "A" "B" facts: TrueI\ \ "True" \ - \ "PROP ?P" \ fail) lemma assumes asms: "A" "B" shows "True" apply (recursion_example' "A" "B" facts: asms) apply simp done (*Method.sections in existing method*) method my_simp\<^sub>1 uses my_simp\<^sub>1_facts = (simp add: my_simp\<^sub>1_facts) lemma assumes A shows A by (my_simp\<^sub>1 my_simp\<^sub>1_facts: assms) (*Method.sections via Eisbach argument parser*) method uses_test\<^sub>2 uses uses_test\<^sub>2_uses = (uses_test\<^sub>1 uses_test\<^sub>1_uses: uses_test\<^sub>2_uses) lemma assumes A shows A by (uses_test\<^sub>2 uses_test\<^sub>2_uses: assms) subsection \Declaration Tests\ named_theorems declare_facts\<^sub>1 method declares_test\<^sub>1 declares declare_facts\<^sub>1 = (rule declare_facts\<^sub>1) lemma assumes A shows A by (declares_test\<^sub>1 declare_facts\<^sub>1: assms) lemma assumes A[declare_facts\<^sub>1]: A shows A by declares_test\<^sub>1 subsection \Rule Instantiation Tests\ method my_allE\<^sub>1 for x :: 'a and P :: "'a \ bool" = (erule allE [where x = x and P = P]) lemma "\x. Q x \ Q x" by (my_allE\<^sub>1 x Q) method my_allE\<^sub>2 for x :: 'a and P :: "'a \ bool" = (erule allE [of P x]) lemma "\x. Q x \ Q x" by (my_allE\<^sub>2 x Q) method my_allE\<^sub>3 for x :: 'a and P :: "'a \ bool" = (match allE [where 'a = 'a] in X: "\(x :: 'a) P R. \x. P x \ (P x \ R) \ R" \ \erule X [where x = x and P = P]\) lemma "\x. Q x \ Q x" by (my_allE\<^sub>3 x Q) method my_allE\<^sub>4 for x :: 'a and P :: "'a \ bool" = (match allE [where 'a = 'a] in X: "\(x :: 'a) P R. \x. P x \ (P x \ R) \ R" \ \erule X [of x P]\) lemma "\x. Q x \ Q x" by (my_allE\<^sub>4 x Q) subsection \Polymorphism test\ axiomatization foo' :: "'a \ 'b \ 'c \ bool" axiomatization where foo'_ax1: "foo' x y z \ z \ y" axiomatization where foo'_ax2: "foo' x y y \ x \ z" axiomatization where foo'_ax3: "foo' (x :: int) y y \ y \ y" lemmas my_thms = foo'_ax1 foo'_ax2 foo'_ax3 definition first_id where "first_id x = x" lemmas my_thms' = my_thms[of "first_id x" for x] method print_conclusion = (match conclusion in concl for concl \ \print_term concl\) lemma assumes foo: "\x (y :: bool). foo' (A x) B (A x)" shows "\z. A z \ B" apply (match conclusion in "f x y" for f y and x :: "'d :: type" \ \ match my_thms' in R:"\(x :: 'f :: type). ?P (first_id x) \ ?R" and R':"\(x :: 'f :: type). ?P' (first_id x) \ ?R'" \ \ match (x) in "q :: 'f" for q \ \ rule R[of q,simplified first_id_def], print_conclusion, rule foo \\\) done subsection \Unchecked rule instantiation, with the possibility of runtime errors\ named_theorems my_thms_named declare foo'_ax3[my_thms_named] method foo_method3 declares my_thms_named = (match my_thms_named[of (unchecked) z for z] in R:"PROP ?H" \ \rule R\) notepad begin (*FIXME: Shouldn't need unchecked keyword here. See Tests_Failing.thy *) fix A B x have "foo' x B A \ A \ B" by (match my_thms[of (unchecked) z for z] in R:"PROP ?H" \ \rule R\) fix A B x note foo'_ax1[my_thms_named] have "foo' x B A \ A \ B" by (match my_thms_named[where x=z for z] in R:"PROP ?H" \ \rule R\) fix A B x note foo'_ax1[my_thms_named] foo'_ax2[my_thms_named] foo'_ax3[my_thms_named] have "foo' x B A \ A \ B" by foo_method3 end ML \ structure Data = Generic_Data ( type T = thm list; val empty: T = []; fun merge data : T = Thm.merge_thms data; ); \ local_setup \Local_Theory.add_thms_dynamic (\<^binding>\test_dyn\, Data.get)\ setup \Context.theory_map (Data.put @{thms TrueI})\ method dynamic_thms_test = (rule test_dyn) locale foo = fixes A assumes A : "A" begin -local_setup - \Local_Theory.declaration {pervasive = false, syntax = false} - (fn phi => Data.put (Morphism.fact phi @{thms A}))\ +declaration \fn phi => Data.put (Morphism.fact phi @{thms A})\ lemma A by dynamic_thms_test end notepad begin fix A x assume X: "\x. A x" have "A x" by (match X in H[of x]:"\x. A x" \ \print_fact H,match H in "A x" \ \rule H\\) fix A x B assume X: "\x :: bool. A x \ B" "\x. A x" assume Y: "A B" have "B \ B \ B \ B \ B \ B" apply (intro conjI) apply (match X in H[OF X(2)]:"\x. A x \ B" \ \print_fact H,rule H\) apply (match X in H':"\x. A x" and H[OF H']:"\x. A x \ B" \ \print_fact H',print_fact H,rule H\) apply (match X in H[of Q]:"\x. A x \ ?R" and "?P \ Q" for Q \ \print_fact H,rule H, rule Y\) apply (match X in H[of Q,OF Y]:"\x. A x \ ?R" and "?P \ Q" for Q \ \print_fact H,rule H\) apply (match X in H[OF Y,intro]:"\x. A x \ ?R" \ \print_fact H,fastforce\) apply (match X in H[intro]:"\x. A x \ ?R" \ \rule H[where x=B], rule Y\) done fix x :: "prop" and A assume X: "TERM x" assume Y: "\x :: prop. A x" have "A TERM x" apply (match X in "PROP y" for y \ \rule Y[where x="PROP y"]\) done end subsection \Proper context for method parameters\ method add_simp methods m uses f = (match f in H[simp]:_ \ m) method add_my_thms methods m uses f = (match f in H[my_thms_named]:_ \ m) method rule_my_thms = (rule my_thms_named) method rule_my_thms' declares my_thms_named = (rule my_thms_named) lemma assumes A: A and B: B shows "(A \ B) \ A \ A \ A" apply (intro conjI) apply (add_simp \add_simp simp f: B\ f: A) apply (add_my_thms rule_my_thms f:A) apply (add_my_thms rule_my_thms' f:A) apply (add_my_thms \rule my_thms_named\ f:A) done subsection \Shallow parser tests\ method all_args for A B methods m1 m2 uses f1 f2 declares my_thms_named = fail lemma True by (all_args True False - fail f1: TrueI f2: TrueI my_thms_named: TrueI | rule TrueI) subsection \Method name internalization test\ method test2 = (simp) method simp = fail lemma "A \ A" by test2 subsection \Dynamic facts\ named_theorems my_thms_named' method foo_method1 for x = (match my_thms_named' [of (unchecked) x] in R: "PROP ?H" \ \rule R\) lemma assumes A [my_thms_named']: "\x. A x" shows "A y" by (foo_method1 y) subsection \Eisbach method invocation from ML\ method test_method for x y uses r = (print_term x, print_term y, rule r) method_setup test_method' = \ Args.term -- Args.term -- (Scan.lift (Args.$$$ "rule" -- Args.colon) |-- Attrib.thms) >> (fn ((x, y), r) => fn ctxt => Method_Closure.apply_method ctxt \<^method>\test_method\ [x, y] [r] [] ctxt) \ lemma fixes a b :: nat assumes "a = b" shows "b = a" apply (test_method a b)? apply (test_method' a b rule: refl)? apply (test_method' a b rule: assms [symmetric])? done subsection \Eisbach methods in locales\ locale my_locale1 = fixes A assumes A: A begin method apply_A = (match conclusion in "A" \ \match A in U:"A" \ \print_term A, print_fact A, rule U\\) end locale my_locale2 = fixes B assumes B: B begin interpretation my_locale1 B by (unfold_locales; rule B) lemma B by apply_A end context fixes C assumes C: C begin interpretation my_locale1 C by (unfold_locales; rule C) lemma C by apply_A end context begin interpretation my_locale1 "True \ True" by (unfold_locales; blast) lemma "True \ True" by apply_A end locale locale_poly = fixes P assumes P: "\x :: 'a. P x" begin method solve_P for z :: 'a = (rule P[where x = z]) end context begin interpretation locale_poly "\x:: nat. 0 \ x" by (unfold_locales; blast) lemma "0 \ (n :: nat)" by (solve_P n) end subsection \Mutual recursion via higher-order methods\ experiment begin method inner_method methods passed_method = (rule conjI; passed_method) method outer_method = (inner_method \outer_method\ | assumption) lemma "Q \ R \ P \ (Q \ R) \ P" by outer_method end end diff --git a/src/HOL/Eisbach/method_closure.ML b/src/HOL/Eisbach/method_closure.ML --- a/src/HOL/Eisbach/method_closure.ML +++ b/src/HOL/Eisbach/method_closure.ML @@ -1,251 +1,252 @@ (* Title: HOL/Eisbach/method_closure.ML Author: Daniel Matichuk, NICTA/UNSW Facilities for treating method syntax as a closure, with abstraction over terms, facts and other methods. The 'method' command allows to define new proof methods by combining existing ones with their usual syntax. *) signature METHOD_CLOSURE = sig val apply_method: Proof.context -> string -> term list -> thm list list -> (Proof.context -> Method.method) list -> Proof.context -> thm list -> context_tactic val method: binding -> (binding * typ option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory val method_cmd: binding -> (binding * string option * mixfix) list -> binding list -> binding list -> binding list -> Token.src -> local_theory -> string * local_theory end; structure Method_Closure: METHOD_CLOSURE = struct (* auxiliary data for method definition *) structure Method_Definition = Proof_Data ( type T = (Proof.context -> Method.method) Symtab.table * (*dynamic methods*) (term list -> Proof.context -> Method.method) Symtab.table (*recursive methods*); fun init _ : T = (Symtab.empty, Symtab.empty); ); fun lookup_dynamic_method ctxt full_name = (case Symtab.lookup (#1 (Method_Definition.get ctxt)) full_name of SOME m => m ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val update_dynamic_method = Method_Definition.map o apfst o Symtab.update; fun get_recursive_method full_name ts ctxt = (case Symtab.lookup (#2 (Method_Definition.get ctxt)) full_name of SOME m => m ts ctxt | NONE => error ("Illegal use of internal Eisbach method: " ^ quote full_name)); val put_recursive_method = Method_Definition.map o apsnd o Symtab.update; (* stored method closures *) type closure = {vars: term list, named_thms: string list, methods: string list, body: Method.text}; structure Data = Generic_Data ( type T = closure Symtab.table; val empty: T = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun get_closure ctxt name = (case Symtab.lookup (Data.get (Context.Proof ctxt)) name of SOME closure => closure | NONE => error ("Unknown Eisbach method: " ^ quote name)); fun put_closure binding (closure: closure) lthy = let val name = Local_Theory.full_name lthy binding; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map - (Symtab.update (name, - {vars = map (Morphism.term phi) (#vars closure), - named_thms = #named_thms closure, - methods = #methods closure, - body = (Method.map_source o map) (Token.transform phi) (#body closure)}))) + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding} + (fn phi => + Data.map + (Symtab.update (name, + {vars = map (Morphism.term phi) (#vars closure), + named_thms = #named_thms closure, + methods = #methods closure, + body = (Method.map_source o map) (Token.transform phi) (#body closure)}))) end; (* instantiate and evaluate method text *) fun method_instantiate vars body ts ctxt = let val thy = Proof_Context.theory_of ctxt; val subst = fold (Pattern.match thy) (vars ~~ ts) (Vartab.empty, Vartab.empty); val morphism = Morphism.term_morphism "method_instantiate" (Envir.subst_term subst); val body' = (Method.map_source o map) (Token.transform morphism) body; in Method.evaluate_runtime body' ctxt end; (** apply method closure **) fun recursive_method full_name vars body ts = let val m = method_instantiate vars body in put_recursive_method (full_name, m) #> m ts end; fun apply_method ctxt method_name terms facts methods = let fun declare_facts (name :: names) (fact :: facts) = fold (Context.proof_map o Named_Theorems.add_thm name) fact #> declare_facts names facts | declare_facts _ [] = I | declare_facts [] (_ :: _) = error ("Excessive facts for method " ^ quote method_name); val {vars, named_thms, methods = method_args, body} = get_closure ctxt method_name; in declare_facts named_thms facts #> fold update_dynamic_method (method_args ~~ methods) #> recursive_method method_name vars body terms end; (** define method closure **) local fun setup_local_method binding lthy = let val full_name = Local_Theory.full_name lthy binding; fun dynamic_method ctxt = lookup_dynamic_method ctxt full_name; in lthy |> update_dynamic_method (full_name, K Method.fail) |> Method.local_setup binding (Scan.succeed dynamic_method) "(internal)" end; fun check_named_thm ctxt binding = let val bname = Binding.name_of binding; val pos = Binding.pos_of binding; val full_name = Named_Theorems.check ctxt (bname, pos); val parser: Method.modifier parser = Args.$$$ bname -- Args.colon >> K {init = I, attribute = Named_Theorems.add full_name, pos = pos}; in (full_name, parser) end; fun parse_term_args args = Args.context :|-- (fn ctxt => let val ctxt' = Proof_Context.set_mode (Proof_Context.mode_schematic) ctxt; fun parse T = (if T = propT then Syntax.parse_prop ctxt' else Syntax.parse_term ctxt') #> Type.constraint (Type_Infer.paramify_vars T); fun do_parse' T = Parse_Tools.name_term >> Parse_Tools.parse_val_cases (parse T); fun do_parse (Var (_, T)) = do_parse' T | do_parse (Free (_, T)) = do_parse' T | do_parse t = error ("Unexpected method parameter: " ^ Syntax.string_of_term ctxt' t); fun rep [] x = Scan.succeed [] x | rep (t :: ts) x = (do_parse t ::: rep ts) x; fun check ts = let val (ts, fs) = split_list ts; val ts' = Syntax.check_terms ctxt' ts |> Variable.polymorphic ctxt'; val _ = ListPair.app (fn (f, t) => f t) (fs, ts'); in ts' end; in Scan.lift (rep args) >> check end); fun parse_method_args method_args = let fun bind_method (name, text) ctxt = let val method = Method.evaluate_runtime text; val inner_update = method o update_dynamic_method (name, K (method ctxt)); in update_dynamic_method (name, inner_update) ctxt end; fun rep [] x = Scan.succeed [] x | rep (m :: ms) x = ((Method.text_closure >> pair m) ::: rep ms) x; in rep method_args >> fold bind_method end; fun gen_method add_fixes name vars uses declares methods source lthy = let val (uses_internal, lthy1) = lthy |> Proof_Context.concealed |> Local_Theory.begin_nested |-> Proof_Context.private_scope |> Local_Theory.map_background_naming (Name_Space.add_path (Binding.name_of name)) |> Config.put Method.old_section_parser true |> fold setup_local_method methods |> fold_map (fn b => Named_Theorems.declare b "") uses; val (term_args, lthy2) = lthy1 |> add_fixes vars |-> fold_map Proof_Context.inferred_param |>> map Free; val (named_thms, modifiers) = map (check_named_thm lthy2) (declares @ uses) |> split_list; val method_args = map (Local_Theory.full_name lthy2) methods; fun parser args meth = apfst (Config.put_generic Method.old_section_parser true) #> (parse_term_args args -- parse_method_args method_args --| (Scan.depend (fn context => Scan.succeed (fold Named_Theorems.clear uses_internal context, ())) -- Method.sections modifiers)) >> (fn (ts, decl) => meth ts o decl); val full_name = Local_Theory.full_name lthy name; val lthy3 = lthy2 |> Method.local_setup (Binding.make (Binding.name_of name, Position.none)) (parser term_args (get_recursive_method full_name)) "(internal)" |> put_recursive_method (full_name, fn _ => fn _ => Method.fail); val (text, src) = Method.read_closure (Config.put Proof_Context.dynamic_facts_dummy true lthy3) source; val morphism = Variable.export_morphism lthy3 (lthy |> Proof_Context.transfer (Proof_Context.theory_of lthy3) |> fold Token.declare_maxidx src |> Variable.declare_maxidx (Variable.maxidx_of lthy3)); val text' = (Method.map_source o map) (Token.transform morphism) text; val term_args' = map (Morphism.term morphism) term_args; in lthy3 |> Local_Theory.end_nested |> Proof_Context.restore_naming lthy |> put_closure name {vars = term_args', named_thms = named_thms, methods = method_args, body = text'} |> Method.local_setup name (Args.context :|-- (fn ctxt => let val {body, vars, ...} = get_closure ctxt full_name in parser vars (recursive_method full_name vars body) end)) "" |> pair full_name end; in val method = gen_method Proof_Context.add_fixes; val method_cmd = gen_method Proof_Context.add_fixes_cmd; end; val _ = Outer_Syntax.local_theory \<^command_keyword>\method\ "Eisbach method definition" (Parse.binding -- Parse.for_fixes -- ((Scan.optional (\<^keyword>\methods\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- (Scan.optional (\<^keyword>\uses\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) [])) -- (Scan.optional (\<^keyword>\declares\ |-- Parse.!!! (Scan.repeat1 Parse.binding)) []) -- Parse.!!! (\<^keyword>\=\ |-- Parse.args1 (K true)) >> (fn ((((name, vars), (methods, uses)), declares), src) => #2 o method_cmd name vars uses declares methods src)); end; diff --git a/src/HOL/Enum.thy b/src/HOL/Enum.thy --- a/src/HOL/Enum.thy +++ b/src/HOL/Enum.thy @@ -1,1160 +1,1160 @@ (* Author: Florian Haftmann, TU Muenchen *) section \Finite types as explicit enumerations\ theory Enum imports Map Groups_List begin subsection \Class \enum\\ class enum = fixes enum :: "'a list" fixes enum_all :: "('a \ bool) \ bool" fixes enum_ex :: "('a \ bool) \ bool" assumes UNIV_enum: "UNIV = set enum" and enum_distinct: "distinct enum" assumes enum_all_UNIV: "enum_all P \ Ball UNIV P" assumes enum_ex_UNIV: "enum_ex P \ Bex UNIV P" \ \tailored towards simple instantiation\ begin subclass finite proof qed (simp add: UNIV_enum) lemma enum_UNIV: "set enum = UNIV" by (simp only: UNIV_enum) lemma in_enum: "x \ set enum" by (simp add: enum_UNIV) lemma enum_eq_I: assumes "\x. x \ set xs" shows "set enum = set xs" proof - from assms UNIV_eq_I have "UNIV = set xs" by auto with enum_UNIV show ?thesis by simp qed lemma card_UNIV_length_enum: "card (UNIV :: 'a set) = length enum" by (simp add: UNIV_enum distinct_card enum_distinct) lemma enum_all [simp]: "enum_all = HOL.All" by (simp add: fun_eq_iff enum_all_UNIV) lemma enum_ex [simp]: "enum_ex = HOL.Ex" by (simp add: fun_eq_iff enum_ex_UNIV) end subsection \Implementations using \<^class>\enum\\ subsubsection \Unbounded operations and quantifiers\ lemma Collect_code [code]: "Collect P = set (filter P enum)" by (simp add: enum_UNIV) lemma vimage_code [code]: "f -` B = set (filter (\x. f x \ B) enum_class.enum)" unfolding vimage_def Collect_code .. definition card_UNIV :: "'a itself \ nat" where [code del]: "card_UNIV TYPE('a) = card (UNIV :: 'a set)" lemma [code]: "card_UNIV TYPE('a :: enum) = card (set (Enum.enum :: 'a list))" by (simp only: card_UNIV_def enum_UNIV) lemma all_code [code]: "(\x. P x) \ enum_all P" by simp lemma exists_code [code]: "(\x. P x) \ enum_ex P" by simp lemma exists1_code [code]: "(\!x. P x) \ list_ex1 P enum" by (auto simp add: list_ex1_iff enum_UNIV) subsubsection \An executable choice operator\ definition [code del]: "enum_the = The" lemma [code]: "The P = (case filter P enum of [x] \ x | _ \ enum_the P)" proof - { fix a assume filter_enum: "filter P enum = [a]" have "The P = a" proof (rule the_equality) fix x assume "P x" show "x = a" proof (rule ccontr) assume "x \ a" from filter_enum obtain us vs where enum_eq: "enum = us @ [a] @ vs" and "\ x \ set us. \ P x" and "\ x \ set vs. \ P x" and "P a" by (auto simp add: filter_eq_Cons_iff) (simp only: filter_empty_conv[symmetric]) with \P x\ in_enum[of x, unfolded enum_eq] \x \ a\ show "False" by auto qed next from filter_enum show "P a" by (auto simp add: filter_eq_Cons_iff) qed } from this show ?thesis unfolding enum_the_def by (auto split: list.split) qed declare [[code abort: enum_the]] code_printing constant enum_the \ (Eval) "(fn '_ => raise Match)" subsubsection \Equality and order on functions\ instantiation "fun" :: (enum, equal) equal begin definition "HOL.equal f g \ (\x \ set enum. f x = g x)" instance proof qed (simp_all add: equal_fun_def fun_eq_iff enum_UNIV) end lemma [code]: "HOL.equal f g \ enum_all (%x. f x = g x)" by (auto simp add: equal fun_eq_iff) lemma [code nbe]: "HOL.equal (f :: _ \ _) f \ True" by (fact equal_refl) lemma order_fun [code]: fixes f g :: "'a::enum \ 'b::order" shows "f \ g \ enum_all (\x. f x \ g x)" and "f < g \ f \ g \ enum_ex (\x. f x \ g x)" by (simp_all add: fun_eq_iff le_fun_def order_less_le) subsubsection \Operations on relations\ lemma [code]: "Id = image (\x. (x, x)) (set Enum.enum)" by (auto intro: imageI in_enum) lemma tranclp_unfold [code]: "tranclp r a b \ (a, b) \ trancl {(x, y). r x y}" by (simp add: trancl_def) lemma rtranclp_rtrancl_eq [code]: "rtranclp r x y \ (x, y) \ rtrancl {(x, y). r x y}" by (simp add: rtrancl_def) lemma max_ext_eq [code]: "max_ext R = {(X, Y). finite X \ finite Y \ Y \ {} \ (\x. x \ X \ (\xa \ Y. (x, xa) \ R))}" by (auto simp add: max_ext.simps) lemma max_extp_eq [code]: "max_extp r x y \ (x, y) \ max_ext {(x, y). r x y}" by (simp add: max_ext_def) lemma mlex_eq [code]: "f <*mlex*> R = {(x, y). f x < f y \ (f x \ f y \ (x, y) \ R)}" by (auto simp add: mlex_prod_def) subsubsection \Bounded accessible part\ primrec bacc :: "('a \ 'a) set \ nat \ 'a set" where "bacc r 0 = {x. \ y. (y, x) \ r}" | "bacc r (Suc n) = (bacc r n \ {x. \y. (y, x) \ r \ y \ bacc r n})" lemma bacc_subseteq_acc: "bacc r n \ Wellfounded.acc r" by (induct n) (auto intro: acc.intros) lemma bacc_mono: "n \ m \ bacc r n \ bacc r m" by (induct rule: dec_induct) auto lemma bacc_upper_bound: "bacc (r :: ('a \ 'a) set) (card (UNIV :: 'a::finite set)) = (\n. bacc r n)" proof - have "mono (bacc r)" unfolding mono_def by (simp add: bacc_mono) moreover have "\n. bacc r n = bacc r (Suc n) \ bacc r (Suc n) = bacc r (Suc (Suc n))" by auto moreover have "finite (range (bacc r))" by auto ultimately show ?thesis by (intro finite_mono_strict_prefix_implies_finite_fixpoint) (auto intro: finite_mono_remains_stable_implies_strict_prefix) qed lemma acc_subseteq_bacc: assumes "finite r" shows "Wellfounded.acc r \ (\n. bacc r n)" proof fix x assume "x \ Wellfounded.acc r" then have "\n. x \ bacc r n" proof (induct x arbitrary: rule: acc.induct) case (accI x) then have "\y. \ n. (y, x) \ r \ y \ bacc r n" by simp from choice[OF this] obtain n where n: "\y. (y, x) \ r \ y \ bacc r (n y)" .. obtain n where "\y. (y, x) \ r \ y \ bacc r n" proof fix y assume y: "(y, x) \ r" with n have "y \ bacc r (n y)" by auto moreover have "n y <= Max ((\(y, x). n y) ` r)" using y \finite r\ by (auto intro!: Max_ge) note bacc_mono[OF this, of r] ultimately show "y \ bacc r (Max ((\(y, x). n y) ` r))" by auto qed then show ?case by (auto simp add: Let_def intro!: exI[of _ "Suc n"]) qed then show "x \ (\n. bacc r n)" by auto qed lemma acc_bacc_eq: fixes A :: "('a :: finite \ 'a) set" assumes "finite A" shows "Wellfounded.acc A = bacc A (card (UNIV :: 'a set))" using assms by (metis acc_subseteq_bacc bacc_subseteq_acc bacc_upper_bound order_eq_iff) lemma [code]: fixes xs :: "('a::finite \ 'a) list" shows "Wellfounded.acc (set xs) = bacc (set xs) (card_UNIV TYPE('a))" by (simp add: card_UNIV_def acc_bacc_eq) subsection \Default instances for \<^class>\enum\\ lemma map_of_zip_enum_is_Some: assumes "length ys = length (enum :: 'a::enum list)" shows "\y. map_of (zip (enum :: 'a::enum list) ys) x = Some y" proof - from assms have "x \ set (enum :: 'a::enum list) \ (\y. map_of (zip (enum :: 'a::enum list) ys) x = Some y)" by (auto intro!: map_of_zip_is_Some) then show ?thesis using enum_UNIV by auto qed lemma map_of_zip_enum_inject: fixes xs ys :: "'b::enum list" assumes length: "length xs = length (enum :: 'a::enum list)" "length ys = length (enum :: 'a::enum list)" and map_of: "the \ map_of (zip (enum :: 'a::enum list) xs) = the \ map_of (zip (enum :: 'a::enum list) ys)" shows "xs = ys" proof - have "map_of (zip (enum :: 'a list) xs) = map_of (zip (enum :: 'a list) ys)" proof fix x :: 'a from length map_of_zip_enum_is_Some obtain y1 y2 where "map_of (zip (enum :: 'a list) xs) x = Some y1" and "map_of (zip (enum :: 'a list) ys) x = Some y2" by blast moreover from map_of have "the (map_of (zip (enum :: 'a::enum list) xs) x) = the (map_of (zip (enum :: 'a::enum list) ys) x)" by (auto dest: fun_cong) ultimately show "map_of (zip (enum :: 'a::enum list) xs) x = map_of (zip (enum :: 'a::enum list) ys) x" by simp qed with length enum_distinct show "xs = ys" by (rule map_of_zip_inject) qed definition all_n_lists :: "(('a :: enum) list \ bool) \ nat \ bool" where "all_n_lists P n \ (\xs \ set (List.n_lists n enum). P xs)" lemma [code]: "all_n_lists P n \ (if n = 0 then P [] else enum_all (%x. all_n_lists (%xs. P (x # xs)) (n - 1)))" unfolding all_n_lists_def enum_all by (cases n) (auto simp add: enum_UNIV) definition ex_n_lists :: "(('a :: enum) list \ bool) \ nat \ bool" where "ex_n_lists P n \ (\xs \ set (List.n_lists n enum). P xs)" lemma [code]: "ex_n_lists P n \ (if n = 0 then P [] else enum_ex (%x. ex_n_lists (%xs. P (x # xs)) (n - 1)))" unfolding ex_n_lists_def enum_ex by (cases n) (auto simp add: enum_UNIV) instantiation "fun" :: (enum, enum) enum begin definition "enum = map (\ys. the \ map_of (zip (enum::'a list) ys)) (List.n_lists (length (enum::'a::enum list)) enum)" definition "enum_all P = all_n_lists (\bs. P (the \ map_of (zip enum bs))) (length (enum :: 'a list))" definition "enum_ex P = ex_n_lists (\bs. P (the \ map_of (zip enum bs))) (length (enum :: 'a list))" instance proof show "UNIV = set (enum :: ('a \ 'b) list)" proof (rule UNIV_eq_I) fix f :: "'a \ 'b" have "f = the \ map_of (zip (enum :: 'a::enum list) (map f enum))" by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) then show "f \ set enum" by (auto simp add: enum_fun_def set_n_lists intro: in_enum) qed next from map_of_zip_enum_inject show "distinct (enum :: ('a \ 'b) list)" by (auto intro!: inj_onI simp add: enum_fun_def distinct_map distinct_n_lists enum_distinct set_n_lists) next fix P show "enum_all (P :: ('a \ 'b) \ bool) = Ball UNIV P" proof assume "enum_all P" show "Ball UNIV P" proof fix f :: "'a \ 'b" have f: "f = the \ map_of (zip (enum :: 'a::enum list) (map f enum))" by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) from \enum_all P\ have "P (the \ map_of (zip enum (map f enum)))" unfolding enum_all_fun_def all_n_lists_def apply (simp add: set_n_lists) apply (erule_tac x="map f enum" in allE) apply (auto intro!: in_enum) done from this f show "P f" by auto qed next assume "Ball UNIV P" from this show "enum_all P" unfolding enum_all_fun_def all_n_lists_def by auto qed next fix P show "enum_ex (P :: ('a \ 'b) \ bool) = Bex UNIV P" proof assume "enum_ex P" from this show "Bex UNIV P" unfolding enum_ex_fun_def ex_n_lists_def by auto next assume "Bex UNIV P" from this obtain f where "P f" .. have f: "f = the \ map_of (zip (enum :: 'a::enum list) (map f enum))" by (auto simp add: map_of_zip_map fun_eq_iff intro: in_enum) from \P f\ this have "P (the \ map_of (zip (enum :: 'a::enum list) (map f enum)))" by auto from this show "enum_ex P" unfolding enum_ex_fun_def ex_n_lists_def apply (auto simp add: set_n_lists) apply (rule_tac x="map f enum" in exI) apply (auto intro!: in_enum) done qed qed end lemma enum_fun_code [code]: "enum = (let enum_a = (enum :: 'a::{enum, equal} list) in map (\ys. the \ map_of (zip enum_a ys)) (List.n_lists (length enum_a) enum))" by (simp add: enum_fun_def Let_def) lemma enum_all_fun_code [code]: "enum_all P = (let enum_a = (enum :: 'a::{enum, equal} list) in all_n_lists (\bs. P (the \ map_of (zip enum_a bs))) (length enum_a))" by (simp only: enum_all_fun_def Let_def) lemma enum_ex_fun_code [code]: "enum_ex P = (let enum_a = (enum :: 'a::{enum, equal} list) in ex_n_lists (\bs. P (the \ map_of (zip enum_a bs))) (length enum_a))" by (simp only: enum_ex_fun_def Let_def) instantiation set :: (enum) enum begin definition "enum = map set (subseqs enum)" definition "enum_all P \ (\A\set enum. P (A::'a set))" definition "enum_ex P \ (\A\set enum. P (A::'a set))" instance proof qed (simp_all add: enum_set_def enum_all_set_def enum_ex_set_def subseqs_powset distinct_set_subseqs enum_distinct enum_UNIV) end instantiation unit :: enum begin definition "enum = [()]" definition "enum_all P = P ()" definition "enum_ex P = P ()" instance proof qed (auto simp add: enum_unit_def enum_all_unit_def enum_ex_unit_def) end instantiation bool :: enum begin definition "enum = [False, True]" definition "enum_all P \ P False \ P True" definition "enum_ex P \ P False \ P True" instance proof qed (simp_all only: enum_bool_def enum_all_bool_def enum_ex_bool_def UNIV_bool, simp_all) end instantiation prod :: (enum, enum) enum begin definition "enum = List.product enum enum" definition "enum_all P = enum_all (%x. enum_all (%y. P (x, y)))" definition "enum_ex P = enum_ex (%x. enum_ex (%y. P (x, y)))" instance by standard (simp_all add: enum_prod_def distinct_product enum_UNIV enum_distinct enum_all_prod_def enum_ex_prod_def) end instantiation sum :: (enum, enum) enum begin definition "enum = map Inl enum @ map Inr enum" definition "enum_all P \ enum_all (\x. P (Inl x)) \ enum_all (\x. P (Inr x))" definition "enum_ex P \ enum_ex (\x. P (Inl x)) \ enum_ex (\x. P (Inr x))" instance proof qed (simp_all only: enum_sum_def enum_all_sum_def enum_ex_sum_def UNIV_sum, auto simp add: enum_UNIV distinct_map enum_distinct) end instantiation option :: (enum) enum begin definition "enum = None # map Some enum" definition "enum_all P \ P None \ enum_all (\x. P (Some x))" definition "enum_ex P \ P None \ enum_ex (\x. P (Some x))" instance proof qed (simp_all only: enum_option_def enum_all_option_def enum_ex_option_def UNIV_option_conv, auto simp add: distinct_map enum_UNIV enum_distinct) end subsection \Small finite types\ text \We define small finite types for use in Quickcheck\ datatype (plugins only: code "quickcheck" extraction) finite_1 = a\<^sub>1 notation (output) a\<^sub>1 ("a\<^sub>1") lemma UNIV_finite_1: "UNIV = {a\<^sub>1}" by (auto intro: finite_1.exhaust) instantiation finite_1 :: enum begin definition "enum = [a\<^sub>1]" definition "enum_all P = P a\<^sub>1" definition "enum_ex P = P a\<^sub>1" instance proof qed (simp_all only: enum_finite_1_def enum_all_finite_1_def enum_ex_finite_1_def UNIV_finite_1, simp_all) end instantiation finite_1 :: linorder begin definition less_finite_1 :: "finite_1 \ finite_1 \ bool" where "x < (y :: finite_1) \ False" definition less_eq_finite_1 :: "finite_1 \ finite_1 \ bool" where "x \ (y :: finite_1) \ True" instance apply (intro_classes) apply (auto simp add: less_finite_1_def less_eq_finite_1_def) apply (metis (full_types) finite_1.exhaust) done end instance finite_1 :: "{dense_linorder, wellorder}" by intro_classes (simp_all add: less_finite_1_def) instantiation finite_1 :: complete_lattice begin definition [simp]: "Inf = (\_. a\<^sub>1)" definition [simp]: "Sup = (\_. a\<^sub>1)" definition [simp]: "bot = a\<^sub>1" definition [simp]: "top = a\<^sub>1" definition [simp]: "inf = (\_ _. a\<^sub>1)" definition [simp]: "sup = (\_ _. a\<^sub>1)" instance by intro_classes(simp_all add: less_eq_finite_1_def) end instance finite_1 :: complete_distrib_lattice by standard simp_all instance finite_1 :: complete_linorder .. lemma finite_1_eq: "x = a\<^sub>1" by(cases x) simp simproc_setup finite_1_eq ("x::finite_1") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => (case Thm.term_of ct of Const (\<^const_name>\a\<^sub>1\, _) => NONE - | _ => SOME (mk_meta_eq @{thm finite_1_eq})) + | _ => SOME (mk_meta_eq @{thm finite_1_eq})))) \ instantiation finite_1 :: complete_boolean_algebra begin definition [simp]: "(-) = (\_ _. a\<^sub>1)" definition [simp]: "uminus = (\_. a\<^sub>1)" instance by intro_classes simp_all end instantiation finite_1 :: "{linordered_ring_strict, linordered_comm_semiring_strict, ordered_comm_ring, ordered_cancel_comm_monoid_diff, comm_monoid_mult, ordered_ring_abs, one, modulo, sgn, inverse}" begin definition [simp]: "Groups.zero = a\<^sub>1" definition [simp]: "Groups.one = a\<^sub>1" definition [simp]: "(+) = (\_ _. a\<^sub>1)" definition [simp]: "(*) = (\_ _. a\<^sub>1)" definition [simp]: "(mod) = (\_ _. a\<^sub>1)" definition [simp]: "abs = (\_. a\<^sub>1)" definition [simp]: "sgn = (\_. a\<^sub>1)" definition [simp]: "inverse = (\_. a\<^sub>1)" definition [simp]: "divide = (\_ _. a\<^sub>1)" instance by intro_classes(simp_all add: less_finite_1_def) end declare [[simproc del: finite_1_eq]] hide_const (open) a\<^sub>1 datatype (plugins only: code "quickcheck" extraction) finite_2 = a\<^sub>1 | a\<^sub>2 notation (output) a\<^sub>1 ("a\<^sub>1") notation (output) a\<^sub>2 ("a\<^sub>2") lemma UNIV_finite_2: "UNIV = {a\<^sub>1, a\<^sub>2}" by (auto intro: finite_2.exhaust) instantiation finite_2 :: enum begin definition "enum = [a\<^sub>1, a\<^sub>2]" definition "enum_all P \ P a\<^sub>1 \ P a\<^sub>2" definition "enum_ex P \ P a\<^sub>1 \ P a\<^sub>2" instance proof qed (simp_all only: enum_finite_2_def enum_all_finite_2_def enum_ex_finite_2_def UNIV_finite_2, simp_all) end instantiation finite_2 :: linorder begin definition less_finite_2 :: "finite_2 \ finite_2 \ bool" where "x < y \ x = a\<^sub>1 \ y = a\<^sub>2" definition less_eq_finite_2 :: "finite_2 \ finite_2 \ bool" where "x \ y \ x = y \ x < (y :: finite_2)" instance apply (intro_classes) apply (auto simp add: less_finite_2_def less_eq_finite_2_def) apply (metis finite_2.nchotomy)+ done end instance finite_2 :: wellorder by(rule wf_wellorderI)(simp add: less_finite_2_def, intro_classes) instantiation finite_2 :: complete_lattice begin definition "\A = (if a\<^sub>1 \ A then a\<^sub>1 else a\<^sub>2)" definition "\A = (if a\<^sub>2 \ A then a\<^sub>2 else a\<^sub>1)" definition [simp]: "bot = a\<^sub>1" definition [simp]: "top = a\<^sub>2" definition "x \ y = (if x = a\<^sub>1 \ y = a\<^sub>1 then a\<^sub>1 else a\<^sub>2)" definition "x \ y = (if x = a\<^sub>2 \ y = a\<^sub>2 then a\<^sub>2 else a\<^sub>1)" lemma neq_finite_2_a\<^sub>1_iff [simp]: "x \ a\<^sub>1 \ x = a\<^sub>2" by(cases x) simp_all lemma neq_finite_2_a\<^sub>1_iff' [simp]: "a\<^sub>1 \ x \ x = a\<^sub>2" by(cases x) simp_all lemma neq_finite_2_a\<^sub>2_iff [simp]: "x \ a\<^sub>2 \ x = a\<^sub>1" by(cases x) simp_all lemma neq_finite_2_a\<^sub>2_iff' [simp]: "a\<^sub>2 \ x \ x = a\<^sub>1" by(cases x) simp_all instance proof fix x :: finite_2 and A assume "x \ A" then show "\A \ x" "x \ \A" by(cases x; auto simp add: less_eq_finite_2_def less_finite_2_def Inf_finite_2_def Sup_finite_2_def)+ qed(auto simp add: less_eq_finite_2_def less_finite_2_def inf_finite_2_def sup_finite_2_def Inf_finite_2_def Sup_finite_2_def) end instance finite_2 :: complete_linorder .. instance finite_2 :: complete_distrib_lattice .. instantiation finite_2 :: "{field, idom_abs_sgn, idom_modulo}" begin definition [simp]: "0 = a\<^sub>1" definition [simp]: "1 = a\<^sub>2" definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \ a\<^sub>1 | (a\<^sub>2, a\<^sub>2) \ a\<^sub>1 | _ \ a\<^sub>2)" definition "uminus = (\x :: finite_2. x)" definition "(-) = ((+) :: finite_2 \ _)" definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \ a\<^sub>2 | _ \ a\<^sub>1)" definition "inverse = (\x :: finite_2. x)" definition "divide = ((*) :: finite_2 \ _)" definition "x mod y = (case (x, y) of (a\<^sub>2, a\<^sub>1) \ a\<^sub>2 | _ \ a\<^sub>1)" definition "abs = (\x :: finite_2. x)" definition "sgn = (\x :: finite_2. x)" instance by standard (subproofs \simp_all add: plus_finite_2_def uminus_finite_2_def minus_finite_2_def times_finite_2_def inverse_finite_2_def divide_finite_2_def modulo_finite_2_def abs_finite_2_def sgn_finite_2_def split: finite_2.splits\) end lemma two_finite_2 [simp]: "2 = a\<^sub>1" by (simp add: numeral.simps plus_finite_2_def) lemma dvd_finite_2_unfold: "x dvd y \ x = a\<^sub>2 \ y = a\<^sub>1" by (auto simp add: dvd_def times_finite_2_def split: finite_2.splits) instantiation finite_2 :: "{normalization_semidom, unique_euclidean_semiring}" begin definition [simp]: "normalize = (id :: finite_2 \ _)" definition [simp]: "unit_factor = (id :: finite_2 \ _)" definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \ 0 | a\<^sub>2 \ 1)" definition [simp]: "division_segment (x :: finite_2) = 1" instance by standard (subproofs \auto simp add: divide_finite_2_def times_finite_2_def dvd_finite_2_unfold split: finite_2.splits\) end hide_const (open) a\<^sub>1 a\<^sub>2 datatype (plugins only: code "quickcheck" extraction) finite_3 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 notation (output) a\<^sub>1 ("a\<^sub>1") notation (output) a\<^sub>2 ("a\<^sub>2") notation (output) a\<^sub>3 ("a\<^sub>3") lemma UNIV_finite_3: "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3}" by (auto intro: finite_3.exhaust) instantiation finite_3 :: enum begin definition "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3]" definition "enum_all P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3" definition "enum_ex P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3" instance proof qed (simp_all only: enum_finite_3_def enum_all_finite_3_def enum_ex_finite_3_def UNIV_finite_3, simp_all) end lemma finite_3_not_eq_unfold: "x \ a\<^sub>1 \ x \ {a\<^sub>2, a\<^sub>3}" "x \ a\<^sub>2 \ x \ {a\<^sub>1, a\<^sub>3}" "x \ a\<^sub>3 \ x \ {a\<^sub>1, a\<^sub>2}" by (cases x; simp)+ instantiation finite_3 :: linorder begin definition less_finite_3 :: "finite_3 \ finite_3 \ bool" where "x < y = (case x of a\<^sub>1 \ y \ a\<^sub>1 | a\<^sub>2 \ y = a\<^sub>3 | a\<^sub>3 \ False)" definition less_eq_finite_3 :: "finite_3 \ finite_3 \ bool" where "x \ y \ x = y \ x < (y :: finite_3)" instance proof (intro_classes) qed (auto simp add: less_finite_3_def less_eq_finite_3_def split: finite_3.split_asm) end instance finite_3 :: wellorder proof(rule wf_wellorderI) have "inv_image less_than (case_finite_3 0 1 2) = {(x, y). x < y}" by(auto simp add: less_finite_3_def split: finite_3.splits) from this[symmetric] show "wf \" by simp qed intro_classes class finite_lattice = finite + lattice + Inf + Sup + bot + top + assumes Inf_finite_empty: "Inf {} = Sup UNIV" assumes Inf_finite_insert: "Inf (insert a A) = a \ Inf A" assumes Sup_finite_empty: "Sup {} = Inf UNIV" assumes Sup_finite_insert: "Sup (insert a A) = a \ Sup A" assumes bot_finite_def: "bot = Inf UNIV" assumes top_finite_def: "top = Sup UNIV" begin subclass complete_lattice proof fix x A show "x \ A \ \A \ x" by (metis Set.set_insert abel_semigroup.commute local.Inf_finite_insert local.inf.abel_semigroup_axioms local.inf.left_idem local.inf.orderI) show "x \ A \ x \ \A" by (metis Set.set_insert insert_absorb2 local.Sup_finite_insert local.sup.absorb_iff2) next fix A z have "\ UNIV = z \ \UNIV" by (subst Sup_finite_insert [symmetric], simp add: insert_UNIV) from this have [simp]: "z \ \UNIV" using local.le_iff_sup by auto have "(\ x. x \ A \ z \ x) \ z \ \A" by (rule finite_induct [of A "\ A . (\ x. x \ A \ z \ x) \ z \ \A"]) (simp_all add: Inf_finite_empty Inf_finite_insert) from this show "(\x. x \ A \ z \ x) \ z \ \A" by simp have "\ UNIV = z \ \UNIV" by (subst Inf_finite_insert [symmetric], simp add: insert_UNIV) from this have [simp]: "\UNIV \ z" by (simp add: local.inf.absorb_iff2) have "(\ x. x \ A \ x \ z) \ \A \ z" by (rule finite_induct [of A "\ A . (\ x. x \ A \ x \ z) \ \A \ z" ], simp_all add: Sup_finite_empty Sup_finite_insert) from this show " (\x. x \ A \ x \ z) \ \A \ z" by blast next show "\{} = \" by (simp add: Inf_finite_empty top_finite_def) show " \{} = \" by (simp add: Sup_finite_empty bot_finite_def) qed end class finite_distrib_lattice = finite_lattice + distrib_lattice begin lemma finite_inf_Sup: "a \ (Sup A) = Sup {a \ b | b . b \ A}" proof (rule finite_induct [of A "\ A . a \ (Sup A) = Sup {a \ b | b . b \ A}"], simp_all) fix x::"'a" fix F assume "x \ F" assume [simp]: "a \ \F = \{a \ b |b. b \ F}" have [simp]: " insert (a \ x) {a \ b |b. b \ F} = {a \ b |b. b = x \ b \ F}" by blast have "a \ (x \ \F) = a \ x \ a \ \F" by (simp add: inf_sup_distrib1) also have "... = a \ x \ \{a \ b |b. b \ F}" by simp also have "... = \{a \ b |b. b = x \ b \ F}" by (unfold Sup_insert[THEN sym], simp) finally show "a \ (x \ \F) = \{a \ b |b. b = x \ b \ F}" by simp qed lemma finite_Inf_Sup: "\(Sup ` A) \ \(Inf ` {f ` A |f. \Y\A. f Y \ Y})" proof (rule finite_induct [of A "\A. \(Sup ` A) \ \(Inf ` {f ` A |f. \Y\A. f Y \ Y})"], simp_all add: finite_UnionD) fix x::"'a set" fix F assume "x \ F" have [simp]: "{\x \ b |b . b \ Inf ` {f ` F |f. \Y\F. f Y \ Y} } = {\x \ (Inf (f ` F)) |f . (\Y\F. f Y \ Y)}" by auto define fa where "fa = (\ (b::'a) f Y . (if Y = x then b else f Y))" have "\f b. \Y\F. f Y \ Y \ b \ x \ insert b (f ` (F \ {Y. Y \ x})) = insert (fa b f x) (fa b f ` F) \ fa b f x \ x \ (\Y\F. fa b f Y \ Y)" by (auto simp add: fa_def) from this have B: "\f b. \Y\F. f Y \ Y \ b \ x \ fa b f ` ({x} \ F) \ {insert (f x) (f ` F) |f. f x \ x \ (\Y\F. f Y \ Y)}" by blast have [simp]: "\f b. \Y\F. f Y \ Y \ b \ x \ b \ (\x\F. f x) \ \(Inf ` {insert (f x) (f ` F) |f. f x \ x \ (\Y\F. f Y \ Y)})" using B apply (rule SUP_upper2) using \x \ F\ apply (simp_all add: fa_def Inf_union_distrib) apply (simp add: image_mono Inf_superset_mono inf.coboundedI2) done assume "\(Sup ` F) \ \(Inf ` {f ` F |f. \Y\F. f Y \ Y})" from this have "\x \ \(Sup ` F) \ \x \ \(Inf ` {f ` F |f. \Y\F. f Y \ Y})" using inf.coboundedI2 by auto also have "... = Sup {\x \ (Inf (f ` F)) |f . (\Y\F. f Y \ Y)}" by (simp add: finite_inf_Sup) also have "... = Sup {Sup {Inf (f ` F) \ b | b . b \ x} |f . (\Y\F. f Y \ Y)}" by (subst inf_commute) (simp add: finite_inf_Sup) also have "... \ \(Inf ` {insert (f x) (f ` F) |f. f x \ x \ (\Y\F. f Y \ Y)})" apply (rule Sup_least, clarsimp)+ apply (subst inf_commute, simp) done finally show "\x \ \(Sup ` F) \ \(Inf ` {insert (f x) (f ` F) |f. f x \ x \ (\Y\F. f Y \ Y)})" by simp qed subclass complete_distrib_lattice by (standard, rule finite_Inf_Sup) end instantiation finite_3 :: finite_lattice begin definition "\A = (if a\<^sub>1 \ A then a\<^sub>1 else if a\<^sub>2 \ A then a\<^sub>2 else a\<^sub>3)" definition "\A = (if a\<^sub>3 \ A then a\<^sub>3 else if a\<^sub>2 \ A then a\<^sub>2 else a\<^sub>1)" definition [simp]: "bot = a\<^sub>1" definition [simp]: "top = a\<^sub>3" definition [simp]: "inf = (min :: finite_3 \ _)" definition [simp]: "sup = (max :: finite_3 \ _)" instance proof qed (auto simp add: Inf_finite_3_def Sup_finite_3_def max_def min_def less_eq_finite_3_def less_finite_3_def split: finite_3.split) end instance finite_3 :: complete_lattice .. instance finite_3 :: finite_distrib_lattice proof qed (auto simp add: min_def max_def) instance finite_3 :: complete_distrib_lattice .. instance finite_3 :: complete_linorder .. instantiation finite_3 :: "{field, idom_abs_sgn, idom_modulo}" begin definition [simp]: "0 = a\<^sub>1" definition [simp]: "1 = a\<^sub>2" definition "x + y = (case (x, y) of (a\<^sub>1, a\<^sub>1) \ a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \ a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \ a\<^sub>1 | (a\<^sub>1, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>2, a\<^sub>1) \ a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \ a\<^sub>2 | _ \ a\<^sub>3)" definition "- x = (case x of a\<^sub>1 \ a\<^sub>1 | a\<^sub>2 \ a\<^sub>3 | a\<^sub>3 \ a\<^sub>2)" definition "x - y = x + (- y :: finite_3)" definition "x * y = (case (x, y) of (a\<^sub>2, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>3, a\<^sub>3) \ a\<^sub>2 | (a\<^sub>2, a\<^sub>3) \ a\<^sub>3 | (a\<^sub>3, a\<^sub>2) \ a\<^sub>3 | _ \ a\<^sub>1)" definition "inverse = (\x :: finite_3. x)" definition "x div y = x * inverse (y :: finite_3)" definition "x mod y = (case y of a\<^sub>1 \ x | _ \ a\<^sub>1)" definition "abs = (\x. case x of a\<^sub>3 \ a\<^sub>2 | _ \ x)" definition "sgn = (\x :: finite_3. x)" instance by standard (subproofs \simp_all add: plus_finite_3_def uminus_finite_3_def minus_finite_3_def times_finite_3_def inverse_finite_3_def divide_finite_3_def modulo_finite_3_def abs_finite_3_def sgn_finite_3_def less_finite_3_def split: finite_3.splits\) end lemma two_finite_3 [simp]: "2 = a\<^sub>3" by (simp add: numeral.simps plus_finite_3_def) lemma dvd_finite_3_unfold: "x dvd y \ x = a\<^sub>2 \ x = a\<^sub>3 \ y = a\<^sub>1" by (cases x) (auto simp add: dvd_def times_finite_3_def split: finite_3.splits) instantiation finite_3 :: "{normalization_semidom, unique_euclidean_semiring}" begin definition [simp]: "normalize x = (case x of a\<^sub>3 \ a\<^sub>2 | _ \ x)" definition [simp]: "unit_factor = (id :: finite_3 \ _)" definition [simp]: "euclidean_size x = (case x of a\<^sub>1 \ 0 | _ \ 1)" definition [simp]: "division_segment (x :: finite_3) = 1" instance proof fix x :: finite_3 assume "x \ 0" then show "is_unit (unit_factor x)" by (cases x) (simp_all add: dvd_finite_3_unfold) qed (subproofs \auto simp add: divide_finite_3_def times_finite_3_def dvd_finite_3_unfold inverse_finite_3_def plus_finite_3_def split: finite_3.splits\) end hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 datatype (plugins only: code "quickcheck" extraction) finite_4 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 notation (output) a\<^sub>1 ("a\<^sub>1") notation (output) a\<^sub>2 ("a\<^sub>2") notation (output) a\<^sub>3 ("a\<^sub>3") notation (output) a\<^sub>4 ("a\<^sub>4") lemma UNIV_finite_4: "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4}" by (auto intro: finite_4.exhaust) instantiation finite_4 :: enum begin definition "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4]" definition "enum_all P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3 \ P a\<^sub>4" definition "enum_ex P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3 \ P a\<^sub>4" instance proof qed (simp_all only: enum_finite_4_def enum_all_finite_4_def enum_ex_finite_4_def UNIV_finite_4, simp_all) end instantiation finite_4 :: finite_distrib_lattice begin text \\<^term>\a\<^sub>1\ $<$ \<^term>\a\<^sub>2\,\<^term>\a\<^sub>3\ $<$ \<^term>\a\<^sub>4\, but \<^term>\a\<^sub>2\ and \<^term>\a\<^sub>3\ are incomparable.\ definition "x < y \ (case (x, y) of (a\<^sub>1, a\<^sub>1) \ False | (a\<^sub>1, _) \ True | (a\<^sub>2, a\<^sub>4) \ True | (a\<^sub>3, a\<^sub>4) \ True | _ \ False)" definition "x \ y \ (case (x, y) of (a\<^sub>1, _) \ True | (a\<^sub>2, a\<^sub>2) \ True | (a\<^sub>2, a\<^sub>4) \ True | (a\<^sub>3, a\<^sub>3) \ True | (a\<^sub>3, a\<^sub>4) \ True | (a\<^sub>4, a\<^sub>4) \ True | _ \ False)" definition "\A = (if a\<^sub>1 \ A \ a\<^sub>2 \ A \ a\<^sub>3 \ A then a\<^sub>1 else if a\<^sub>2 \ A then a\<^sub>2 else if a\<^sub>3 \ A then a\<^sub>3 else a\<^sub>4)" definition "\A = (if a\<^sub>4 \ A \ a\<^sub>2 \ A \ a\<^sub>3 \ A then a\<^sub>4 else if a\<^sub>2 \ A then a\<^sub>2 else if a\<^sub>3 \ A then a\<^sub>3 else a\<^sub>1)" definition [simp]: "bot = a\<^sub>1" definition [simp]: "top = a\<^sub>4" definition "x \ y = (case (x, y) of (a\<^sub>1, _) \ a\<^sub>1 | (_, a\<^sub>1) \ a\<^sub>1 | (a\<^sub>2, a\<^sub>3) \ a\<^sub>1 | (a\<^sub>3, a\<^sub>2) \ a\<^sub>1 | (a\<^sub>2, _) \ a\<^sub>2 | (_, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>3, _) \ a\<^sub>3 | (_, a\<^sub>3) \ a\<^sub>3 | _ \ a\<^sub>4)" definition "x \ y = (case (x, y) of (a\<^sub>4, _) \ a\<^sub>4 | (_, a\<^sub>4) \ a\<^sub>4 | (a\<^sub>2, a\<^sub>3) \ a\<^sub>4 | (a\<^sub>3, a\<^sub>2) \ a\<^sub>4 | (a\<^sub>2, _) \ a\<^sub>2 | (_, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>3, _) \ a\<^sub>3 | (_, a\<^sub>3) \ a\<^sub>3 | _ \ a\<^sub>1)" instance by standard (subproofs \auto simp add: less_finite_4_def less_eq_finite_4_def Inf_finite_4_def Sup_finite_4_def inf_finite_4_def sup_finite_4_def split: finite_4.splits\) end instance finite_4 :: complete_lattice .. instance finite_4 :: complete_distrib_lattice .. instantiation finite_4 :: complete_boolean_algebra begin definition "- x = (case x of a\<^sub>1 \ a\<^sub>4 | a\<^sub>2 \ a\<^sub>3 | a\<^sub>3 \ a\<^sub>2 | a\<^sub>4 \ a\<^sub>1)" definition "x - y = x \ - (y :: finite_4)" instance by standard (subproofs \simp_all add: inf_finite_4_def sup_finite_4_def uminus_finite_4_def minus_finite_4_def split: finite_4.splits\) end hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 datatype (plugins only: code "quickcheck" extraction) finite_5 = a\<^sub>1 | a\<^sub>2 | a\<^sub>3 | a\<^sub>4 | a\<^sub>5 notation (output) a\<^sub>1 ("a\<^sub>1") notation (output) a\<^sub>2 ("a\<^sub>2") notation (output) a\<^sub>3 ("a\<^sub>3") notation (output) a\<^sub>4 ("a\<^sub>4") notation (output) a\<^sub>5 ("a\<^sub>5") lemma UNIV_finite_5: "UNIV = {a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5}" by (auto intro: finite_5.exhaust) instantiation finite_5 :: enum begin definition "enum = [a\<^sub>1, a\<^sub>2, a\<^sub>3, a\<^sub>4, a\<^sub>5]" definition "enum_all P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3 \ P a\<^sub>4 \ P a\<^sub>5" definition "enum_ex P \ P a\<^sub>1 \ P a\<^sub>2 \ P a\<^sub>3 \ P a\<^sub>4 \ P a\<^sub>5" instance proof qed (simp_all only: enum_finite_5_def enum_all_finite_5_def enum_ex_finite_5_def UNIV_finite_5, simp_all) end instantiation finite_5 :: finite_lattice begin text \The non-distributive pentagon lattice $N_5$\ definition "x < y \ (case (x, y) of (a\<^sub>1, a\<^sub>1) \ False | (a\<^sub>1, _) \ True | (a\<^sub>2, a\<^sub>3) \ True | (a\<^sub>2, a\<^sub>5) \ True | (a\<^sub>3, a\<^sub>5) \ True | (a\<^sub>4, a\<^sub>5) \ True | _ \ False)" definition "x \ y \ (case (x, y) of (a\<^sub>1, _) \ True | (a\<^sub>2, a\<^sub>2) \ True | (a\<^sub>2, a\<^sub>3) \ True | (a\<^sub>2, a\<^sub>5) \ True | (a\<^sub>3, a\<^sub>3) \ True | (a\<^sub>3, a\<^sub>5) \ True | (a\<^sub>4, a\<^sub>4) \ True | (a\<^sub>4, a\<^sub>5) \ True | (a\<^sub>5, a\<^sub>5) \ True | _ \ False)" definition "\A = (if a\<^sub>1 \ A \ a\<^sub>4 \ A \ (a\<^sub>2 \ A \ a\<^sub>3 \ A) then a\<^sub>1 else if a\<^sub>2 \ A then a\<^sub>2 else if a\<^sub>3 \ A then a\<^sub>3 else if a\<^sub>4 \ A then a\<^sub>4 else a\<^sub>5)" definition "\A = (if a\<^sub>5 \ A \ a\<^sub>4 \ A \ (a\<^sub>2 \ A \ a\<^sub>3 \ A) then a\<^sub>5 else if a\<^sub>3 \ A then a\<^sub>3 else if a\<^sub>2 \ A then a\<^sub>2 else if a\<^sub>4 \ A then a\<^sub>4 else a\<^sub>1)" definition [simp]: "bot = a\<^sub>1" definition [simp]: "top = a\<^sub>5" definition "x \ y = (case (x, y) of (a\<^sub>1, _) \ a\<^sub>1 | (_, a\<^sub>1) \ a\<^sub>1 | (a\<^sub>2, a\<^sub>4) \ a\<^sub>1 | (a\<^sub>4, a\<^sub>2) \ a\<^sub>1 | (a\<^sub>3, a\<^sub>4) \ a\<^sub>1 | (a\<^sub>4, a\<^sub>3) \ a\<^sub>1 | (a\<^sub>2, _) \ a\<^sub>2 | (_, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>3, _) \ a\<^sub>3 | (_, a\<^sub>3) \ a\<^sub>3 | (a\<^sub>4, _) \ a\<^sub>4 | (_, a\<^sub>4) \ a\<^sub>4 | _ \ a\<^sub>5)" definition "x \ y = (case (x, y) of (a\<^sub>5, _) \ a\<^sub>5 | (_, a\<^sub>5) \ a\<^sub>5 | (a\<^sub>2, a\<^sub>4) \ a\<^sub>5 | (a\<^sub>4, a\<^sub>2) \ a\<^sub>5 | (a\<^sub>3, a\<^sub>4) \ a\<^sub>5 | (a\<^sub>4, a\<^sub>3) \ a\<^sub>5 | (a\<^sub>3, _) \ a\<^sub>3 | (_, a\<^sub>3) \ a\<^sub>3 | (a\<^sub>2, _) \ a\<^sub>2 | (_, a\<^sub>2) \ a\<^sub>2 | (a\<^sub>4, _) \ a\<^sub>4 | (_, a\<^sub>4) \ a\<^sub>4 | _ \ a\<^sub>1)" instance by standard (subproofs \auto simp add: less_eq_finite_5_def less_finite_5_def inf_finite_5_def sup_finite_5_def Inf_finite_5_def Sup_finite_5_def split: finite_5.splits if_split_asm\) end instance finite_5 :: complete_lattice .. hide_const (open) a\<^sub>1 a\<^sub>2 a\<^sub>3 a\<^sub>4 a\<^sub>5 subsection \Closing up\ hide_type (open) finite_1 finite_2 finite_3 finite_4 finite_5 hide_const (open) enum enum_all enum_ex all_n_lists ex_n_lists ntrancl end diff --git a/src/HOL/Finite_Set.thy b/src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy +++ b/src/HOL/Finite_Set.thy @@ -1,3044 +1,3044 @@ (* Title: HOL/Finite_Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad Author: Andrei Popescu *) section \Finite sets\ theory Finite_Set imports Product_Type Sum_Type Fields Relation begin subsection \Predicate for finite sets\ context notes [[inductive_internals]] begin inductive finite :: "'a set \ bool" where emptyI [simp, intro!]: "finite {}" | insertI [simp, intro!]: "finite A \ finite (insert a A)" end simproc_setup finite_Collect ("finite (Collect P)") = \K Set_Comprehension_Pointfree.simproc\ declare [[simproc del: finite_Collect]] lemma finite_induct [case_names empty insert, induct set: finite]: \ \Discharging \x \ F\ entails extra work.\ assumes "finite F" assumes "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P F" using \finite F\ proof induct show "P {}" by fact next fix x F assume F: "finite F" and P: "P F" show "P (insert x F)" proof cases assume "x \ F" then have "insert x F = F" by (rule insert_absorb) with P show ?thesis by (simp only:) next assume "x \ F" from F this P show ?thesis by (rule insert) qed qed lemma infinite_finite_induct [case_names infinite empty insert]: assumes infinite: "\A. \ finite A \ P A" and empty: "P {}" and insert: "\x F. finite F \ x \ F \ P F \ P (insert x F)" shows "P A" proof (cases "finite A") case False with infinite show ?thesis . next case True then show ?thesis by (induct A) (fact empty insert)+ qed subsubsection \Choice principles\ lemma ex_new_if_finite: \ \does not depend on def of finite at all\ assumes "\ finite (UNIV :: 'a set)" and "finite A" shows "\a::'a. a \ A" proof - from assms have "A \ UNIV" by blast then show ?thesis by blast qed text \A finite choice principle. Does not need the SOME choice operator.\ lemma finite_set_choice: "finite A \ \x\A. \y. P x y \ \f. \x\A. P x (f x)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert a A) then obtain f b where f: "\x\A. P x (f x)" and ab: "P a b" by auto show ?case (is "\f. ?P f") proof show "?P (\x. if x = a then b else f x)" using f ab by auto qed qed subsubsection \Finite sets are the images of initial segments of natural numbers\ lemma finite_imp_nat_seg_image_inj_on: assumes "finite A" shows "\(n::nat) f. A = f ` {i. i < n} \ inj_on f {i. i < n}" using assms proof induct case empty show ?case proof show "\f. {} = f ` {i::nat. i < 0} \ inj_on f {i. i < 0}" by simp qed next case (insert a A) have notinA: "a \ A" by fact from insert.hyps obtain n f where "A = f ` {i::nat. i < n}" "inj_on f {i. i < n}" by blast then have "insert a A = f(n:=a) ` {i. i < Suc n}" and "inj_on (f(n:=a)) {i. i < Suc n}" using notinA by (auto simp add: image_def Ball_def inj_on_def less_Suc_eq) then show ?case by blast qed lemma nat_seg_image_imp_finite: "A = f ` {i::nat. i < n} \ finite A" proof (induct n arbitrary: A) case 0 then show ?case by simp next case (Suc n) let ?B = "f ` {i. i < n}" have finB: "finite ?B" by (rule Suc.hyps[OF refl]) show ?case proof (cases "\k (\n f. A = f ` {i::nat. i < n})" by (blast intro: nat_seg_image_imp_finite dest: finite_imp_nat_seg_image_inj_on) lemma finite_imp_inj_to_nat_seg: assumes "finite A" shows "\f n. f ` A = {i::nat. i < n} \ inj_on f A" proof - from finite_imp_nat_seg_image_inj_on [OF \finite A\] obtain f and n :: nat where bij: "bij_betw f {i. i ?f ` A = {i. i k}" by (simp add: le_eq_less_or_eq Collect_disj_eq) subsection \Finiteness and common set operations\ lemma rev_finite_subset: "finite B \ A \ B \ finite A" proof (induct arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert x F A) have A: "A \ insert x F" and r: "A - {x} \ F \ finite (A - {x})" by fact+ show "finite A" proof cases assume x: "x \ A" with A have "A - {x} \ F" by (simp add: subset_insert_iff) with r have "finite (A - {x})" . then have "finite (insert x (A - {x}))" .. also have "insert x (A - {x}) = A" using x by (rule insert_Diff) finally show ?thesis . next show ?thesis when "A \ F" using that by fact assume "x \ A" with A show "A \ F" by (simp add: subset_insert_iff) qed qed lemma finite_subset: "A \ B \ finite B \ finite A" by (rule rev_finite_subset) -simproc_setup finite ("finite A") = \fn _ => +simproc_setup finite ("finite A") = \ let val finite_subset = @{thm finite_subset} val Eq_TrueI = @{thm Eq_TrueI} fun is_subset A th = case Thm.prop_of th of (_ $ (Const (\<^const_name>\less_eq\, Type (\<^type_name>\fun\, [Type (\<^type_name>\set\, _), _])) $ A' $ B)) => if A aconv A' then SOME(B,th) else NONE | _ => NONE; fun is_finite th = case Thm.prop_of th of (_ $ (Const (\<^const_name>\finite\, _) $ A)) => SOME(A,th) | _ => NONE; fun comb (A,sub_th) (A',fin_th) ths = if A aconv A' then (sub_th,fin_th) :: ths else ths - fun proc ss ct = + fun proc ctxt ct = (let val _ $ A = Thm.term_of ct - val prems = Simplifier.prems_of ss + val prems = Simplifier.prems_of ctxt val fins = map_filter is_finite prems val subsets = map_filter (is_subset A) prems in case fold_product comb subsets fins [] of (sub_th,fin_th) :: _ => SOME((fin_th RS (sub_th RS finite_subset)) RS Eq_TrueI) | _ => NONE end) -in proc end +in K proc end \ (* Needs to be used with care *) declare [[simproc del: finite]] lemma finite_UnI: assumes "finite F" and "finite G" shows "finite (F \ G)" using assms by induct simp_all lemma finite_Un [iff]: "finite (F \ G) \ finite F \ finite G" by (blast intro: finite_UnI finite_subset [of _ "F \ G"]) lemma finite_insert [simp]: "finite (insert a A) \ finite A" proof - have "finite {a} \ finite A \ finite A" by simp then have "finite ({a} \ A) \ finite A" by (simp only: finite_Un) then show ?thesis by simp qed lemma finite_Int [simp, intro]: "finite F \ finite G \ finite (F \ G)" by (blast intro: finite_subset) lemma finite_Collect_conjI [simp, intro]: "finite {x. P x} \ finite {x. Q x} \ finite {x. P x \ Q x}" by (simp add: Collect_conj_eq) lemma finite_Collect_disjI [simp]: "finite {x. P x \ Q x} \ finite {x. P x} \ finite {x. Q x}" by (simp add: Collect_disj_eq) lemma finite_Diff [simp, intro]: "finite A \ finite (A - B)" by (rule finite_subset, rule Diff_subset) lemma finite_Diff2 [simp]: assumes "finite B" shows "finite (A - B) \ finite A" proof - have "finite A \ finite ((A - B) \ (A \ B))" by (simp add: Un_Diff_Int) also have "\ \ finite (A - B)" using \finite B\ by simp finally show ?thesis .. qed lemma finite_Diff_insert [iff]: "finite (A - insert a B) \ finite (A - B)" proof - have "finite (A - B) \ finite (A - B - {a})" by simp moreover have "A - insert a B = A - B - {a}" by auto ultimately show ?thesis by simp qed lemma finite_compl [simp]: "finite (A :: 'a set) \ finite (- A) \ finite (UNIV :: 'a set)" by (simp add: Compl_eq_Diff_UNIV) lemma finite_Collect_not [simp]: "finite {x :: 'a. P x} \ finite {x. \ P x} \ finite (UNIV :: 'a set)" by (simp add: Collect_neg_eq) lemma finite_Union [simp, intro]: "finite A \ (\M. M \ A \ finite M) \ finite (\A)" by (induct rule: finite_induct) simp_all lemma finite_UN_I [intro]: "finite A \ (\a. a \ A \ finite (B a)) \ finite (\a\A. B a)" by (induct rule: finite_induct) simp_all lemma finite_UN [simp]: "finite A \ finite (\(B ` A)) \ (\x\A. finite (B x))" by (blast intro: finite_subset) lemma finite_Inter [intro]: "\A\M. finite A \ finite (\M)" by (blast intro: Inter_lower finite_subset) lemma finite_INT [intro]: "\x\I. finite (A x) \ finite (\x\I. A x)" by (blast intro: INT_lower finite_subset) lemma finite_imageI [simp, intro]: "finite F \ finite (h ` F)" by (induct rule: finite_induct) simp_all lemma finite_image_set [simp]: "finite {x. P x} \ finite {f x |x. P x}" by (simp add: image_Collect [symmetric]) lemma finite_image_set2: "finite {x. P x} \ finite {y. Q y} \ finite {f x y |x y. P x \ Q y}" by (rule finite_subset [where B = "\x \ {x. P x}. \y \ {y. Q y}. {f x y}"]) auto lemma finite_imageD: assumes "finite (f ` A)" and "inj_on f A" shows "finite A" using assms proof (induct "f ` A" arbitrary: A) case empty then show ?case by simp next case (insert x B) then have B_A: "insert x B = f ` A" by simp then obtain y where "x = f y" and "y \ A" by blast from B_A \x \ B\ have "B = f ` A - {x}" by blast with B_A \x \ B\ \x = f y\ \inj_on f A\ \y \ A\ have "B = f ` (A - {y})" by (simp add: inj_on_image_set_diff) moreover from \inj_on f A\ have "inj_on f (A - {y})" by (rule inj_on_diff) ultimately have "finite (A - {y})" by (rule insert.hyps) then show "finite A" by simp qed lemma finite_image_iff: "inj_on f A \ finite (f ` A) \ finite A" using finite_imageD by blast lemma finite_surj: "finite A \ B \ f ` A \ finite B" by (erule finite_subset) (rule finite_imageI) lemma finite_range_imageI: "finite (range g) \ finite (range (\x. f (g x)))" by (drule finite_imageI) (simp add: range_composition) lemma finite_subset_image: assumes "finite B" shows "B \ f ` A \ \C\A. finite C \ B = f ` C" using assms proof induct case empty then show ?case by simp next case insert then show ?case by (clarsimp simp del: image_insert simp add: image_insert [symmetric]) blast qed lemma all_subset_image: "(\B. B \ f ` A \ P B) \ (\B. B \ A \ P(f ` B))" by (safe elim!: subset_imageE) (use image_mono in \blast+\) (* slow *) lemma all_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and P: "\B. finite B \ B \ A \ P (f ` B)" show "P B" using finite_subset_image [OF B] P by blast qed blast lemma ex_finite_subset_image: "(\B. finite B \ B \ f ` A \ P B) \ (\B. finite B \ B \ A \ P (f ` B))" proof safe fix B :: "'a set" assume B: "finite B" "B \ f ` A" and "P B" show "\B. finite B \ B \ A \ P (f ` B)" using finite_subset_image [OF B] \P B\ by blast qed blast lemma finite_vimage_IntI: "finite F \ inj_on h A \ finite (h -` F \ A)" proof (induct rule: finite_induct) case (insert x F) then show ?case by (simp add: vimage_insert [of h x F] finite_subset [OF inj_on_vimage_singleton] Int_Un_distrib2) qed simp lemma finite_finite_vimage_IntI: assumes "finite F" and "\y. y \ F \ finite ((h -` {y}) \ A)" shows "finite (h -` F \ A)" proof - have *: "h -` F \ A = (\ y\F. (h -` {y}) \ A)" by blast show ?thesis by (simp only: * assms finite_UN_I) qed lemma finite_vimageI: "finite F \ inj h \ finite (h -` F)" using finite_vimage_IntI[of F h UNIV] by auto lemma finite_vimageD': "finite (f -` A) \ A \ range f \ finite A" by (auto simp add: subset_image_iff intro: finite_subset[rotated]) lemma finite_vimageD: "finite (h -` F) \ surj h \ finite F" by (auto dest: finite_vimageD') lemma finite_vimage_iff: "bij h \ finite (h -` F) \ finite F" unfolding bij_def by (auto elim: finite_vimageD finite_vimageI) lemma finite_inverse_image_gen: assumes "finite A" "inj_on f D" shows "finite {j\D. f j \ A}" using finite_vimage_IntI [OF assms] by (simp add: Collect_conj_eq inf_commute vimage_def) lemma finite_inverse_image: assumes "finite A" "inj f" shows "finite {j. f j \ A}" using finite_inverse_image_gen [OF assms] by simp lemma finite_Collect_bex [simp]: assumes "finite A" shows "finite {x. \y\A. Q x y} \ (\y\A. finite {x. Q x y})" proof - have "{x. \y\A. Q x y} = (\y\A. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Collect_bounded_ex [simp]: assumes "finite {y. P y}" shows "finite {x. \y. P y \ Q x y} \ (\y. P y \ finite {x. Q x y})" proof - have "{x. \y. P y \ Q x y} = (\y\{y. P y}. {x. Q x y})" by auto with assms show ?thesis by simp qed lemma finite_Plus: "finite A \ finite B \ finite (A <+> B)" by (simp add: Plus_def) lemma finite_PlusD: fixes A :: "'a set" and B :: "'b set" assumes fin: "finite (A <+> B)" shows "finite A" "finite B" proof - have "Inl ` A \ A <+> B" by auto then have "finite (Inl ` A :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite A" by (rule finite_imageD) (auto intro: inj_onI) next have "Inr ` B \ A <+> B" by auto then have "finite (Inr ` B :: ('a + 'b) set)" using fin by (rule finite_subset) then show "finite B" by (rule finite_imageD) (auto intro: inj_onI) qed lemma finite_Plus_iff [simp]: "finite (A <+> B) \ finite A \ finite B" by (auto intro: finite_PlusD finite_Plus) lemma finite_Plus_UNIV_iff [simp]: "finite (UNIV :: ('a + 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" by (subst UNIV_Plus_UNIV [symmetric]) (rule finite_Plus_iff) lemma finite_SigmaI [simp, intro]: "finite A \ (\a. a\A \ finite (B a)) \ finite (SIGMA a:A. B a)" unfolding Sigma_def by blast lemma finite_SigmaI2: assumes "finite {x\A. B x \ {}}" and "\a. a \ A \ finite (B a)" shows "finite (Sigma A B)" proof - from assms have "finite (Sigma {x\A. B x \ {}} B)" by auto also have "Sigma {x:A. B x \ {}} B = Sigma A B" by auto finally show ?thesis . qed lemma finite_cartesian_product: "finite A \ finite B \ finite (A \ B)" by (rule finite_SigmaI) lemma finite_Prod_UNIV: "finite (UNIV :: 'a set) \ finite (UNIV :: 'b set) \ finite (UNIV :: ('a \ 'b) set)" by (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product) lemma finite_cartesian_productD1: assumes "finite (A \ B)" and "B \ {}" shows "finite A" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "fst ` (A \ B) = fst ` f ` {i::nat. i < n}" by simp with \B \ {}\ have "A = (fst \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. A = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_productD2: assumes "finite (A \ B)" and "A \ {}" shows "finite B" proof - from assms obtain n f where "A \ B = f ` {i::nat. i < n}" by (auto simp add: finite_conv_nat_seg_image) then have "snd ` (A \ B) = snd ` f ` {i::nat. i < n}" by simp with \A \ {}\ have "B = (snd \ f) ` {i::nat. i < n}" by (simp add: image_comp) then have "\n f. B = f ` {i::nat. i < n}" by blast then show ?thesis by (auto simp add: finite_conv_nat_seg_image) qed lemma finite_cartesian_product_iff: "finite (A \ B) \ (A = {} \ B = {} \ (finite A \ finite B))" by (auto dest: finite_cartesian_productD1 finite_cartesian_productD2 finite_cartesian_product) lemma finite_prod: "finite (UNIV :: ('a \ 'b) set) \ finite (UNIV :: 'a set) \ finite (UNIV :: 'b set)" using finite_cartesian_product_iff[of UNIV UNIV] by simp lemma finite_Pow_iff [iff]: "finite (Pow A) \ finite A" proof assume "finite (Pow A)" then have "finite ((\x. {x}) ` A)" by (blast intro: finite_subset) (* somewhat slow *) then show "finite A" by (rule finite_imageD [unfolded inj_on_def]) simp next assume "finite A" then show "finite (Pow A)" by induct (simp_all add: Pow_insert) qed corollary finite_Collect_subsets [simp, intro]: "finite A \ finite {B. B \ A}" by (simp add: Pow_def [symmetric]) lemma finite_set: "finite (UNIV :: 'a set set) \ finite (UNIV :: 'a set)" by (simp only: finite_Pow_iff Pow_UNIV[symmetric]) lemma finite_UnionD: "finite (\A) \ finite A" by (blast intro: finite_subset [OF subset_Pow_Union]) lemma finite_bind: assumes "finite S" assumes "\x \ S. finite (f x)" shows "finite (Set.bind S f)" using assms by (simp add: bind_UNION) lemma finite_filter [simp]: "finite S \ finite (Set.filter P S)" unfolding Set.filter_def by simp lemma finite_set_of_finite_funs: assumes "finite A" "finite B" shows "finite {f. \x. (x \ A \ f x \ B) \ (x \ A \ f x = d)}" (is "finite ?S") proof - let ?F = "\f. {(a,b). a \ A \ b = f a}" have "?F ` ?S \ Pow(A \ B)" by auto from finite_subset[OF this] assms have 1: "finite (?F ` ?S)" by simp have 2: "inj_on ?F ?S" by (fastforce simp add: inj_on_def set_eq_iff fun_eq_iff) (* somewhat slow *) show ?thesis by (rule finite_imageD [OF 1 2]) qed lemma not_finite_existsD: assumes "\ finite {a. P a}" shows "\a. P a" proof (rule classical) assume "\ ?thesis" with assms show ?thesis by auto qed lemma finite_converse [iff]: "finite (r\) \ finite r" unfolding converse_def conversep_iff using [[simproc add: finite_Collect]] by (auto elim: finite_imageD simp: inj_on_def) lemma finite_Domain: "finite r \ finite (Domain r)" by (induct set: finite) auto lemma finite_Range: "finite r \ finite (Range r)" by (induct set: finite) auto lemma finite_Field: "finite r \ finite (Field r)" by (simp add: Field_def finite_Domain finite_Range) lemma finite_Image[simp]: "finite R \ finite (R `` A)" by(rule finite_subset[OF _ finite_Range]) auto subsection \Further induction rules on finite sets\ lemma finite_ne_induct [case_names singleton insert, consumes 2]: assumes "finite F" and "F \ {}" assumes "\x. P {x}" and "\x F. finite F \ F \ {} \ x \ F \ P F \ P (insert x F)" shows "P F" using assms proof induct case empty then show ?case by simp next case (insert x F) then show ?case by cases auto qed lemma finite_subset_induct [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. finite F \ a \ A \ a \ F \ P F \ P (insert a F)" shows "P F" using \finite F\ \F \ A\ proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact qed qed lemma finite_empty_induct: assumes "finite A" and "P A" and remove: "\a A. finite A \ a \ A \ P A \ P (A - {a})" shows "P {}" proof - have "P (A - B)" if "B \ A" for B :: "'a set" proof - from \finite A\ that have "finite B" by (rule rev_finite_subset) from this \B \ A\ show "P (A - B)" proof induct case empty from \P A\ show ?case by simp next case (insert b B) have "P (A - B - {b})" proof (rule remove) from \finite A\ show "finite (A - B)" by induct auto from insert show "b \ A - B" by simp from insert show "P (A - B)" by simp qed also have "A - B - {b} = A - insert b B" by (rule Diff_insert [symmetric]) finally show ?case . qed qed then have "P (A - A)" by blast then show ?thesis by simp qed lemma finite_update_induct [consumes 1, case_names const update]: assumes finite: "finite {a. f a \ c}" and const: "P (\a. c)" and update: "\a b f. finite {a. f a \ c} \ f a = c \ b \ c \ P f \ P (f(a := b))" shows "P f" using finite proof (induct "{a. f a \ c}" arbitrary: f) case empty with const show ?case by simp next case (insert a A) then have "A = {a'. (f(a := c)) a' \ c}" and "f a \ c" by auto with \finite A\ have "finite {a'. (f(a := c)) a' \ c}" by simp have "(f(a := c)) a = c" by simp from insert \A = {a'. (f(a := c)) a' \ c}\ have "P (f(a := c))" by simp with \finite {a'. (f(a := c)) a' \ c}\ \(f(a := c)) a = c\ \f a \ c\ have "P ((f(a := c))(a := f a))" by (rule update) then show ?case by simp qed lemma finite_subset_induct' [consumes 2, case_names empty insert]: assumes "finite F" and "F \ A" and empty: "P {}" and insert: "\a F. \finite F; a \ A; F \ A; a \ F; P F \ \ P (insert a F)" shows "P F" using assms(1,2) proof induct show "P {}" by fact next fix x F assume "finite F" and "x \ F" and P: "F \ A \ P F" and i: "insert x F \ A" show "P (insert x F)" proof (rule insert) from i show "x \ A" by blast from i have "F \ A" by blast with P show "P F" . show "finite F" by fact show "x \ F" by fact show "F \ A" by fact qed qed subsection \Class \finite\\ class finite = assumes finite_UNIV: "finite (UNIV :: 'a set)" begin lemma finite [simp]: "finite (A :: 'a set)" by (rule subset_UNIV finite_UNIV finite_subset)+ lemma finite_code [code]: "finite (A :: 'a set) \ True" by simp end instance prod :: (finite, finite) finite by standard (simp only: UNIV_Times_UNIV [symmetric] finite_cartesian_product finite) lemma inj_graph: "inj (\f. {(x, y). y = f x})" by (rule inj_onI) (auto simp add: set_eq_iff fun_eq_iff) instance "fun" :: (finite, finite) finite proof show "finite (UNIV :: ('a \ 'b) set)" proof (rule finite_imageD) let ?graph = "\f::'a \ 'b. {(x, y). y = f x}" have "range ?graph \ Pow UNIV" by simp moreover have "finite (Pow (UNIV :: ('a * 'b) set))" by (simp only: finite_Pow_iff finite) ultimately show "finite (range ?graph)" by (rule finite_subset) show "inj ?graph" by (rule inj_graph) qed qed instance bool :: finite by standard (simp add: UNIV_bool) instance set :: (finite) finite by standard (simp only: Pow_UNIV [symmetric] finite_Pow_iff finite) instance unit :: finite by standard (simp add: UNIV_unit) instance sum :: (finite, finite) finite by standard (simp only: UNIV_Plus_UNIV [symmetric] finite_Plus finite) subsection \A basic fold functional for finite sets\ text \ The intended behaviour is \fold f z {x\<^sub>1, \, x\<^sub>n} = f x\<^sub>1 (\ (f x\<^sub>n z)\)\ if \f\ is ``left-commutative''. The commutativity requirement is relativised to the carrier set \S\: \ locale comp_fun_commute_on = fixes S :: "'a set" fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute_on: "x \ S \ y \ S \ f y \ f x = f x \ f y" begin lemma fun_left_comm: "x \ S \ y \ S \ f y (f x z) = f x (f y z)" using comp_fun_commute_on by (simp add: fun_eq_iff) lemma commute_left_comp: "x \ S \ y \ S \ f y \ (f x \ g) = f x \ (f y \ g)" by (simp add: o_assoc comp_fun_commute_on) end inductive fold_graph :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b \ bool" for f :: "'a \ 'b \ 'b" and z :: 'b where emptyI [intro]: "fold_graph f z {} z" | insertI [intro]: "x \ A \ fold_graph f z A y \ fold_graph f z (insert x A) (f x y)" inductive_cases empty_fold_graphE [elim!]: "fold_graph f z {} x" lemma fold_graph_closed_lemma: "fold_graph f z A x \ x \ B" if "fold_graph g z A x" "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using that(1-3) proof (induction rule: fold_graph.induct) case (insertI x A y) have "fold_graph f z A y" "y \ B" unfolding atomize_conj by (rule insertI.IH) (auto intro: insertI.prems) then have "g x y \ B" and f_eq: "f x y = g x y" by (auto simp: insertI.prems) moreover have "fold_graph f z (insert x A) (f x y)" by (rule fold_graph.insertI; fact) ultimately show ?case by (simp add: f_eq) qed (auto intro!: that) lemma fold_graph_closed_eq: "fold_graph f z A = fold_graph g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" using fold_graph_closed_lemma[of f z A _ B g] fold_graph_closed_lemma[of g z A _ B f] that by auto definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a set \ 'b" where "fold f z A = (if finite A then (THE y. fold_graph f z A y) else z)" lemma fold_closed_eq: "fold f z A = fold g z A" if "\a b. a \ A \ b \ B \ f a b = g a b" "\a b. a \ A \ b \ B \ g a b \ B" "z \ B" unfolding Finite_Set.fold_def by (subst fold_graph_closed_eq[where B=B and g=g]) (auto simp: that) text \ A tempting alternative for the definition is \<^term>\if finite A then THE y. fold_graph f z A y else e\. It allows the removal of finiteness assumptions from the theorems \fold_comm\, \fold_reindex\ and \fold_distrib\. The proofs become ugly. It is not worth the effort. (???) \ lemma finite_imp_fold_graph: "finite A \ \x. fold_graph f z A x" by (induct rule: finite_induct) auto subsubsection \From \<^const>\fold_graph\ to \<^term>\fold\\ context comp_fun_commute_on begin lemma fold_graph_finite: assumes "fold_graph f z A y" shows "finite A" using assms by induct simp_all lemma fold_graph_insertE_aux: assumes "A \ S" assumes "fold_graph f z A y" "a \ A" shows "\y'. y = f a y' \ fold_graph f z (A - {a}) y'" using assms(2-,1) proof (induct set: fold_graph) case emptyI then show ?case by simp next case (insertI x A y) show ?case proof (cases "x = a") case True with insertI show ?thesis by auto next case False then obtain y' where y: "y = f a y'" and y': "fold_graph f z (A - {a}) y'" using insertI by auto from insertI have "x \ S" "a \ S" by auto then have "f x y = f a (f x y')" unfolding y by (intro fun_left_comm; simp) moreover have "fold_graph f z (insert x A - {a}) (f x y')" using y' and \x \ a\ and \x \ A\ by (simp add: insert_Diff_if fold_graph.insertI) ultimately show ?thesis by fast qed qed lemma fold_graph_insertE: assumes "insert x A \ S" assumes "fold_graph f z (insert x A) v" and "x \ A" obtains y where "v = f x y" and "fold_graph f z A y" using assms by (auto dest: fold_graph_insertE_aux[OF \insert x A \ S\ _ insertI1]) lemma fold_graph_determ: assumes "A \ S" assumes "fold_graph f z A x" "fold_graph f z A y" shows "y = x" using assms(2-,1) proof (induct arbitrary: y set: fold_graph) case emptyI then show ?case by fast next case (insertI x A y v) from \insert x A \ S\ and \fold_graph f z (insert x A) v\ and \x \ A\ obtain y' where "v = f x y'" and "fold_graph f z A y'" by (rule fold_graph_insertE) from \fold_graph f z A y'\ insertI have "y' = y" by simp with \v = f x y'\ show "v = f x y" by simp qed lemma fold_equality: "A \ S \ fold_graph f z A y \ fold f z A = y" by (cases "finite A") (auto simp add: fold_def intro: fold_graph_determ dest: fold_graph_finite) lemma fold_graph_fold: assumes "A \ S" assumes "finite A" shows "fold_graph f z A (fold f z A)" proof - from \finite A\ have "\x. fold_graph f z A x" by (rule finite_imp_fold_graph) moreover note fold_graph_determ[OF \A \ S\] ultimately have "\!x. fold_graph f z A x" by (rule ex_ex1I) then have "fold_graph f z A (The (fold_graph f z A))" by (rule theI') with assms show ?thesis by (simp add: fold_def) qed text \The base case for \fold\:\ lemma (in -) fold_infinite [simp]: "\ finite A \ fold f z A = z" by (auto simp: fold_def) lemma (in -) fold_empty [simp]: "fold f z {} = z" by (auto simp: fold_def) text \The various recursion equations for \<^const>\fold\:\ lemma fold_insert [simp]: assumes "insert x A \ S" assumes "finite A" and "x \ A" shows "fold f z (insert x A) = f x (fold f z A)" proof (rule fold_equality[OF \insert x A \ S\]) fix z from \insert x A \ S\ \finite A\ have "fold_graph f z A (fold f z A)" by (blast intro: fold_graph_fold) with \x \ A\ have "fold_graph f z (insert x A) (f x (fold f z A))" by (rule fold_graph.insertI) then show "fold_graph f z (insert x A) (f x (fold f z A))" by simp qed declare (in -) empty_fold_graphE [rule del] fold_graph.intros [rule del] \ \No more proofs involve these.\ lemma fold_fun_left_comm: assumes "insert x A \ S" "finite A" shows "f x (fold f z A) = fold f (f x z) A" using assms(2,1) proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert y F) then have "fold f (f x z) (insert y F) = f y (fold f (f x z) F)" by simp also have "\ = f x (f y (fold f z F))" using insert by (simp add: fun_left_comm[where ?y=x]) also have "\ = f x (fold f z (insert y F))" proof - from insert have "insert y F \ S" by simp from fold_insert[OF this] insert show ?thesis by simp qed finally show ?case .. qed lemma fold_insert2: "insert x A \ S \ finite A \ x \ A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) lemma fold_rec: assumes "A \ S" assumes "finite A" and "x \ A" shows "fold f z A = f x (fold f z (A - {x}))" proof - have A: "A = insert x (A - {x})" using \x \ A\ by blast then have "fold f z A = fold f z (insert x (A - {x}))" by simp also have "\ = f x (fold f z (A - {x}))" by (rule fold_insert) (use assms in \auto\) finally show ?thesis . qed lemma fold_insert_remove: assumes "insert x A \ S" assumes "finite A" shows "fold f z (insert x A) = f x (fold f z (A - {x}))" proof - from \finite A\ have "finite (insert x A)" by auto moreover have "x \ insert x A" by auto ultimately have "fold f z (insert x A) = f x (fold f z (insert x A - {x}))" using \insert x A \ S\ by (blast intro: fold_rec) then show ?thesis by simp qed lemma fold_set_union_disj: assumes "A \ S" "B \ S" assumes "finite A" "finite B" "A \ B = {}" shows "Finite_Set.fold f z (A \ B) = Finite_Set.fold f (Finite_Set.fold f z A) B" using \finite B\ assms(1,2,3,5) proof induct case (insert x F) have "fold f z (A \ insert x F) = f x (fold f (fold f z A) F)" using insert by auto also have "\ = fold f (fold f z A) (insert x F)" using insert by (blast intro: fold_insert[symmetric]) finally show ?case . qed simp end text \Other properties of \<^const>\fold\:\ lemma fold_graph_image: assumes "inj_on g A" shows "fold_graph f z (g ` A) = fold_graph (f \ g) z A" proof fix w show "fold_graph f z (g ` A) w = fold_graph (f o g) z A w" proof assume "fold_graph f z (g ` A) w" then show "fold_graph (f \ g) z A w" using assms proof (induct "g ` A" w arbitrary: A) case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r B) from \inj_on g B\ \x \ A\ \insert x A = image g B\ obtain x' A' where "x' \ A'" and [simp]: "B = insert x' A'" "x = g x'" "A = g ` A'" by (rule inj_img_insertE) from insertI.prems have "fold_graph (f \ g) z A' r" by (auto intro: insertI.hyps) with \x' \ A'\ have "fold_graph (f \ g) z (insert x' A') ((f \ g) x' r)" by (rule fold_graph.insertI) then show ?case by simp qed next assume "fold_graph (f \ g) z A w" then show "fold_graph f z (g ` A) w" using assms proof induct case emptyI then show ?case by (auto intro: fold_graph.emptyI) next case (insertI x A r) from \x \ A\ insertI.prems have "g x \ g ` A" by auto moreover from insertI have "fold_graph f z (g ` A) r" by simp ultimately have "fold_graph f z (insert (g x) (g ` A)) (f (g x) r)" by (rule fold_graph.insertI) then show ?case by simp qed qed qed lemma fold_image: assumes "inj_on g A" shows "fold f z (g ` A) = fold (f \ g) z A" proof (cases "finite A") case False with assms show ?thesis by (auto dest: finite_imageD simp add: fold_def) next case True then show ?thesis by (auto simp add: fold_def fold_graph_image[OF assms]) qed lemma fold_cong: assumes "comp_fun_commute_on S f" "comp_fun_commute_on S g" and "A \ S" "finite A" and cong: "\x. x \ A \ f x = g x" and "s = t" and "A = B" shows "fold f s A = fold g t B" proof - have "fold f s A = fold g s A" using \finite A\ \A \ S\ cong proof (induct A) case empty then show ?case by simp next case insert interpret f: comp_fun_commute_on S f by (fact \comp_fun_commute_on S f\) interpret g: comp_fun_commute_on S g by (fact \comp_fun_commute_on S g\) from insert show ?case by simp qed with assms show ?thesis by simp qed text \A simplified version for idempotent functions:\ locale comp_fun_idem_on = comp_fun_commute_on + assumes comp_fun_idem_on: "x \ S \ f x \ f x = f x" begin lemma fun_left_idem: "x \ S \ f x (f x z) = f x z" using comp_fun_idem_on by (simp add: fun_eq_iff) lemma fold_insert_idem: assumes "insert x A \ S" assumes fin: "finite A" shows "fold f z (insert x A) = f x (fold f z A)" proof cases assume "x \ A" then obtain B where "A = insert x B" and "x \ B" by (rule set_insert) then show ?thesis using assms by (simp add: comp_fun_idem_on fun_left_idem) next assume "x \ A" then show ?thesis using assms by auto qed declare fold_insert [simp del] fold_insert_idem [simp] lemma fold_insert_idem2: "insert x A \ S \ finite A \ fold f z (insert x A) = fold f (f x z) A" by (simp add: fold_fun_left_comm) end subsubsection \Liftings to \comp_fun_commute_on\ etc.\ lemma (in comp_fun_commute_on) comp_comp_fun_commute_on: "range g \ S \ comp_fun_commute_on R (f \ g)" by standard (force intro: comp_fun_commute_on) lemma (in comp_fun_idem_on) comp_comp_fun_idem_on: assumes "range g \ S" shows "comp_fun_idem_on R (f \ g)" proof interpret f_g: comp_fun_commute_on R "f o g" by (fact comp_comp_fun_commute_on[OF \range g \ S\]) show "x \ R \ y \ R \ (f \ g) y \ (f \ g) x = (f \ g) x \ (f \ g) y" for x y by (fact f_g.comp_fun_commute_on) qed (use \range g \ S\ in \force intro: comp_fun_idem_on\) lemma (in comp_fun_commute_on) comp_fun_commute_on_funpow: "comp_fun_commute_on S (\x. f x ^^ g x)" proof fix x y assume "x \ S" "y \ S" show "f y ^^ g y \ f x ^^ g x = f x ^^ g x \ f y ^^ g y" proof (cases "x = y") case True then show ?thesis by simp next case False show ?thesis proof (induct "g x" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) have hyp1: "f y ^^ g y \ f x = f x \ f y ^^ g y" proof (induct "g y" arbitrary: g) case 0 then show ?case by simp next case (Suc n g) define h where "h z = g z - 1" for z with Suc have "n = h y" by simp with Suc have hyp: "f y ^^ h y \ f x = f x \ f y ^^ h y" by auto from Suc h_def have "g y = Suc (h y)" by simp with \x \ S\ \y \ S\ show ?case by (simp add: comp_assoc hyp) (simp add: o_assoc comp_fun_commute_on) qed define h where "h z = (if z = x then g x - 1 else g z)" for z with Suc have "n = h x" by simp with Suc have "f y ^^ h y \ f x ^^ h x = f x ^^ h x \ f y ^^ h y" by auto with False h_def have hyp2: "f y ^^ g y \ f x ^^ h x = f x ^^ h x \ f y ^^ g y" by simp from Suc h_def have "g x = Suc (h x)" by simp then show ?case by (simp del: funpow.simps add: funpow_Suc_right o_assoc hyp2) (simp add: comp_assoc hyp1) qed qed qed subsubsection \\<^term>\UNIV\ as carrier set\ locale comp_fun_commute = fixes f :: "'a \ 'b \ 'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma (in -) comp_fun_commute_def': "comp_fun_commute f = comp_fun_commute_on UNIV f" unfolding comp_fun_commute_def comp_fun_commute_on_def by blast text \ We abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale comp_fun_commute_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "comp_fun_commute_on UNIV f" by standard (simp add: comp_fun_commute) qed simp_all end lemma (in comp_fun_commute) comp_comp_fun_commute: "comp_fun_commute (f o g)" unfolding comp_fun_commute_def' by (fact comp_comp_fun_commute_on) lemma (in comp_fun_commute) comp_fun_commute_funpow: "comp_fun_commute (\x. f x ^^ g x)" unfolding comp_fun_commute_def' by (fact comp_fun_commute_on_funpow) locale comp_fun_idem = comp_fun_commute + assumes comp_fun_idem: "f x o f x = f x" begin lemma (in -) comp_fun_idem_def': "comp_fun_idem f = comp_fun_idem_on UNIV f" unfolding comp_fun_idem_on_def comp_fun_idem_def comp_fun_commute_def' unfolding comp_fun_idem_axioms_def comp_fun_idem_on_axioms_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale comp_fun_idem_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "comp_fun_idem_on UNIV f" by standard (simp_all add: comp_fun_idem comp_fun_commute) qed simp_all end lemma (in comp_fun_idem) comp_comp_fun_idem: "comp_fun_idem (f o g)" unfolding comp_fun_idem_def' by (fact comp_comp_fun_idem_on) subsubsection \Expressing set operations via \<^const>\fold\\ lemma comp_fun_commute_const: "comp_fun_commute (\_. f)" by standard (rule refl) lemma comp_fun_idem_insert: "comp_fun_idem insert" by standard auto lemma comp_fun_idem_remove: "comp_fun_idem Set.remove" by standard auto lemma (in semilattice_inf) comp_fun_idem_inf: "comp_fun_idem inf" by standard (auto simp add: inf_left_commute) lemma (in semilattice_sup) comp_fun_idem_sup: "comp_fun_idem sup" by standard (auto simp add: sup_left_commute) lemma union_fold_insert: assumes "finite A" shows "A \ B = fold insert B A" proof - interpret comp_fun_idem insert by (fact comp_fun_idem_insert) from \finite A\ show ?thesis by (induct A arbitrary: B) simp_all qed lemma minus_fold_remove: assumes "finite A" shows "B - A = fold Set.remove B A" proof - interpret comp_fun_idem Set.remove by (fact comp_fun_idem_remove) from \finite A\ have "fold Set.remove B A = B - A" by (induct A arbitrary: B) auto (* slow *) then show ?thesis .. qed lemma comp_fun_commute_filter_fold: "comp_fun_commute (\x A'. if P x then Set.insert x A' else A')" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff) qed lemma Set_filter_fold: assumes "finite A" shows "Set.filter P A = fold (\x A'. if P x then Set.insert x A' else A') {} A" using assms proof - interpret commute_insert: comp_fun_commute "(\x A'. if P x then Set.insert x A' else A')" by (fact comp_fun_commute_filter_fold) from \finite A\ show ?thesis by induct (auto simp add: Set.filter_def) qed lemma inter_Set_filter: assumes "finite B" shows "A \ B = Set.filter (\x. x \ A) B" using assms by induct (auto simp: Set.filter_def) lemma image_fold_insert: assumes "finite A" shows "image f A = fold (\k A. Set.insert (f k) A) {} A" proof - interpret comp_fun_commute "\k A. Set.insert (f k) A" by standard auto show ?thesis using assms by (induct A) auto qed lemma Ball_fold: assumes "finite A" shows "Ball A P = fold (\k s. s \ P k) True A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma Bex_fold: assumes "finite A" shows "Bex A P = fold (\k s. s \ P k) False A" proof - interpret comp_fun_commute "\k s. s \ P k" by standard auto show ?thesis using assms by (induct A) auto qed lemma comp_fun_commute_Pow_fold: "comp_fun_commute (\x A. A \ Set.insert x ` A)" by (clarsimp simp: fun_eq_iff comp_fun_commute_def) blast lemma Pow_fold: assumes "finite A" shows "Pow A = fold (\x A. A \ Set.insert x ` A) {{}} A" proof - interpret comp_fun_commute "\x A. A \ Set.insert x ` A" by (rule comp_fun_commute_Pow_fold) show ?thesis using assms by (induct A) (auto simp: Pow_insert) qed lemma fold_union_pair: assumes "finite B" shows "(\y\B. {(x, y)}) \ A = fold (\y. Set.insert (x, y)) A B" proof - interpret comp_fun_commute "\y. Set.insert (x, y)" by standard auto show ?thesis using assms by (induct arbitrary: A) simp_all qed lemma comp_fun_commute_product_fold: "finite B \ comp_fun_commute (\x z. fold (\y. Set.insert (x, y)) z B)" by standard (auto simp: fold_union_pair [symmetric]) lemma product_fold: assumes "finite A" "finite B" shows "A \ B = fold (\x z. fold (\y. Set.insert (x, y)) z B) {} A" proof - interpret commute_product: comp_fun_commute "(\x z. fold (\y. Set.insert (x, y)) z B)" by (fact comp_fun_commute_product_fold[OF \finite B\]) from assms show ?thesis unfolding Sigma_def by (induct A) (simp_all add: fold_union_pair) qed context complete_lattice begin lemma inf_Inf_fold_inf: assumes "finite A" shows "inf (Inf A) B = fold inf B A" proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: inf_commute fun_eq_iff) qed lemma sup_Sup_fold_sup: assumes "finite A" shows "sup (Sup A) B = fold sup B A" proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) from \finite A\ fold_fun_left_comm show ?thesis by (induct A arbitrary: B) (simp_all add: sup_commute fun_eq_iff) qed lemma Inf_fold_inf: "finite A \ Inf A = fold inf top A" using inf_Inf_fold_inf [of A top] by (simp add: inf_absorb2) lemma Sup_fold_sup: "finite A \ Sup A = fold sup bot A" using sup_Sup_fold_sup [of A bot] by (simp add: sup_absorb2) lemma inf_INF_fold_inf: assumes "finite A" shows "inf B (\(f ` A)) = fold (inf \ f) B A" (is "?inf = ?fold") proof - interpret comp_fun_idem inf by (fact comp_fun_idem_inf) interpret comp_fun_idem "inf \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?inf" by (induct A arbitrary: B) (simp_all add: inf_left_commute) then show ?thesis .. qed lemma sup_SUP_fold_sup: assumes "finite A" shows "sup B (\(f ` A)) = fold (sup \ f) B A" (is "?sup = ?fold") proof - interpret comp_fun_idem sup by (fact comp_fun_idem_sup) interpret comp_fun_idem "sup \ f" by (fact comp_comp_fun_idem) from \finite A\ have "?fold = ?sup" by (induct A arbitrary: B) (simp_all add: sup_left_commute) then show ?thesis .. qed lemma INF_fold_inf: "finite A \ \(f ` A) = fold (inf \ f) top A" using inf_INF_fold_inf [of A top] by simp lemma SUP_fold_sup: "finite A \ \(f ` A) = fold (sup \ f) bot A" using sup_SUP_fold_sup [of A bot] by simp lemma finite_Inf_in: assumes "finite A" "A\{}" and inf: "\x y. \x \ A; y \ A\ \ inf x y \ A" shows "Inf A \ A" proof - have "Inf B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use inf in \force+\) then show ?thesis by (simp add: assms) qed lemma finite_Sup_in: assumes "finite A" "A\{}" and sup: "\x y. \x \ A; y \ A\ \ sup x y \ A" shows "Sup A \ A" proof - have "Sup B \ A" if "B \ A" "B\{}" for B using finite_subset [OF \B \ A\ \finite A\] that by (induction B) (use sup in \force+\) then show ?thesis by (simp add: assms) qed end subsubsection \Expressing relation operations via \<^const>\fold\\ lemma Id_on_fold: assumes "finite A" shows "Id_on A = Finite_Set.fold (\x. Set.insert (Pair x x)) {} A" proof - interpret comp_fun_commute "\x. Set.insert (Pair x x)" by standard auto from assms show ?thesis unfolding Id_on_def by (induct A) simp_all qed lemma comp_fun_commute_Image_fold: "comp_fun_commute (\(x,y) A. if x \ S then Set.insert y A else A)" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show ?thesis by standard (auto simp: fun_eq_iff comp_fun_commute split: prod.split) qed lemma Image_fold: assumes "finite R" shows "R `` S = Finite_Set.fold (\(x,y) A. if x \ S then Set.insert y A else A) {} R" proof - interpret comp_fun_commute "(\(x,y) A. if x \ S then Set.insert y A else A)" by (rule comp_fun_commute_Image_fold) have *: "\x F. Set.insert x F `` S = (if fst x \ S then Set.insert (snd x) (F `` S) else (F `` S))" by (force intro: rev_ImageI) show ?thesis using assms by (induct R) (auto simp: * ) qed lemma insert_relcomp_union_fold: assumes "finite S" shows "{x} O S \ X = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') X S" proof - interpret comp_fun_commute "\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A'" proof - interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert) show "comp_fun_commute (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A')" by standard (auto simp add: fun_eq_iff split: prod.split) qed have *: "{x} O S = {(x', z). x' = fst x \ (snd x, z) \ S}" by (auto simp: relcomp_unfold intro!: exI) show ?thesis unfolding * using \finite S\ by (induct S) (auto split: prod.split) qed lemma insert_relcomp_fold: assumes "finite S" shows "Set.insert x R O S = Finite_Set.fold (\(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') (R O S) S" proof - have "Set.insert x R O S = ({x} O S) \ (R O S)" by auto then show ?thesis by (auto simp: insert_relcomp_union_fold [OF assms]) qed lemma comp_fun_commute_relcomp_fold: assumes "finite S" shows "comp_fun_commute (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S)" proof - have *: "\a b A. Finite_Set.fold (\(w, z) A'. if b = w then Set.insert (a, z) A' else A') A S = {(a,b)} O S \ A" by (auto simp: insert_relcomp_union_fold[OF assms] cong: if_cong) show ?thesis by standard (auto simp: * ) qed lemma relcomp_fold: assumes "finite R" "finite S" shows "R O S = Finite_Set.fold (\(x,y) A. Finite_Set.fold (\(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S) {} R" proof - interpret commute_relcomp_fold: comp_fun_commute "(\(x, y) A. Finite_Set.fold (\(w, z) A'. if y = w then insert (x, z) A' else A') A S)" by (fact comp_fun_commute_relcomp_fold[OF \finite S\]) from assms show ?thesis by (induct R) (auto simp: comp_fun_commute_relcomp_fold insert_relcomp_fold cong: if_cong) qed subsection \Locales as mini-packages for fold operations\ subsubsection \The natural case\ locale folding_on = fixes S :: "'a set" fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute_on: "x \ S \ y \ S \ f y o f x = f x o f y" begin interpretation fold?: comp_fun_commute_on S f by standard (simp add: comp_fun_commute_on) definition F :: "'a set \ 'b" where eq_fold: "F A = Finite_Set.fold f z A" lemma empty [simp]: "F {} = z" by (simp add: eq_fold) lemma infinite [simp]: "\ finite A \ F A = z" by (simp add: eq_fold) lemma insert [simp]: assumes "insert x A \ S" and "finite A" and "x \ A" shows "F (insert x A) = f x (F A)" proof - from fold_insert assms have "Finite_Set.fold f z (insert x A) = f x (Finite_Set.fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed lemma remove: assumes "A \ S" and "finite A" and "x \ A" shows "F A = f x (F (A - {x}))" proof - from \x \ A\ obtain B where A: "A = insert x B" and "x \ B" by (auto dest: mk_disjoint_insert) moreover from \finite A\ A have "finite B" by simp ultimately show ?thesis using \A \ S\ by auto qed lemma insert_remove: assumes "insert x A \ S" and "finite A" shows "F (insert x A) = f x (F (A - {x}))" using assms by (cases "x \ A") (simp_all add: remove insert_absorb) end subsubsection \With idempotency\ locale folding_idem_on = folding_on + assumes comp_fun_idem_on: "x \ S \ y \ S \ f x \ f x = f x" begin declare insert [simp del] interpretation fold?: comp_fun_idem_on S f by standard (simp_all add: comp_fun_commute_on comp_fun_idem_on) lemma insert_idem [simp]: assumes "insert x A \ S" and "finite A" shows "F (insert x A) = f x (F A)" proof - from fold_insert_idem assms have "fold f z (insert x A) = f x (fold f z A)" by simp with \finite A\ show ?thesis by (simp add: eq_fold fun_eq_iff) qed end subsubsection \\<^term>\UNIV\ as the carrier set\ locale folding = fixes f :: "'a \ 'b \ 'b" and z :: "'b" assumes comp_fun_commute: "f y \ f x = f x \ f y" begin lemma (in -) folding_def': "folding f = folding_on UNIV f" unfolding folding_def folding_on_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale folding_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "folding_on UNIV f" by standard (simp add: comp_fun_commute) qed simp_all end locale folding_idem = folding + assumes comp_fun_idem: "f x \ f x = f x" begin lemma (in -) folding_idem_def': "folding_idem f = folding_idem_on UNIV f" unfolding folding_idem_def folding_def' folding_idem_on_def unfolding folding_idem_axioms_def folding_idem_on_axioms_def by blast text \ Again, we abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the carrier set to \<^term>\UNIV\. \ sublocale folding_idem_on UNIV f rewrites "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "folding_idem_on UNIV f" by standard (simp add: comp_fun_idem) qed simp_all end subsection \Finite cardinality\ text \ The traditional definition \<^prop>\card A \ LEAST n. \f. A = {f i |i. i < n}\ is ugly to work with. But now that we have \<^const>\fold\ things are easy: \ global_interpretation card: folding "\_. Suc" 0 defines card = "folding_on.F (\_. Suc) 0" by standard (rule refl) lemma card_insert_disjoint: "finite A \ x \ A \ card (insert x A) = Suc (card A)" by (fact card.insert) lemma card_insert_if: "finite A \ card (insert x A) = (if x \ A then card A else Suc (card A))" by auto (simp add: card.insert_remove card.remove) lemma card_ge_0_finite: "card A > 0 \ finite A" by (rule ccontr) simp lemma card_0_eq [simp]: "finite A \ card A = 0 \ A = {}" by (auto dest: mk_disjoint_insert) lemma finite_UNIV_card_ge_0: "finite (UNIV :: 'a set) \ card (UNIV :: 'a set) > 0" by (rule ccontr) simp lemma card_eq_0_iff: "card A = 0 \ A = {} \ \ finite A" by auto lemma card_range_greater_zero: "finite (range f) \ card (range f) > 0" by (rule ccontr) (simp add: card_eq_0_iff) lemma card_gt_0_iff: "0 < card A \ A \ {} \ finite A" by (simp add: neq0_conv [symmetric] card_eq_0_iff) lemma card_Suc_Diff1: assumes "finite A" "x \ A" shows "Suc (card (A - {x})) = card A" proof - have "Suc (card (A - {x})) = card (insert x (A - {x}))" using assms by (simp add: card.insert_remove) also have "... = card A" using assms by (simp add: card_insert_if) finally show ?thesis . qed lemma card_insert_le_m1: assumes "n > 0" "card y \ n - 1" shows "card (insert x y) \ n" using assms by (cases "finite y") (auto simp: card_insert_if) lemma card_Diff_singleton: assumes "x \ A" shows "card (A - {x}) = card A - 1" proof (cases "finite A") case True with assms show ?thesis by (simp add: card_Suc_Diff1 [symmetric]) qed auto lemma card_Diff_singleton_if: "card (A - {x}) = (if x \ A then card A - 1 else card A)" by (simp add: card_Diff_singleton) lemma card_Diff_insert[simp]: assumes "a \ A" and "a \ B" shows "card (A - insert a B) = card (A - B) - 1" proof - have "A - insert a B = (A - B) - {a}" using assms by blast then show ?thesis using assms by (simp add: card_Diff_singleton) qed lemma card_insert_le: "card A \ card (insert x A)" proof (cases "finite A") case True then show ?thesis by (simp add: card_insert_if) qed auto lemma card_Collect_less_nat[simp]: "card {i::nat. i < n} = n" by (induct n) (simp_all add:less_Suc_eq Collect_disj_eq) lemma card_Collect_le_nat[simp]: "card {i::nat. i \ n} = Suc n" using card_Collect_less_nat[of "Suc n"] by (simp add: less_Suc_eq_le) lemma card_mono: assumes "finite B" and "A \ B" shows "card A \ card B" proof - from assms have "finite A" by (auto intro: finite_subset) then show ?thesis using assms proof (induct A arbitrary: B) case empty then show ?case by simp next case (insert x A) then have "x \ B" by simp from insert have "A \ B - {x}" and "finite (B - {x})" by auto with insert.hyps have "card A \ card (B - {x})" by auto with \finite A\ \x \ A\ \finite B\ \x \ B\ show ?case by simp (simp only: card.remove) qed qed lemma card_seteq: assumes "finite B" and A: "A \ B" "card B \ card A" shows "A = B" using assms proof (induction arbitrary: A rule: finite_induct) case (insert b B) then have A: "finite A" "A - {b} \ B" by force+ then have "card B \ card (A - {b})" using insert by (auto simp add: card_Diff_singleton_if) then have "A - {b} = B" using A insert.IH by auto then show ?case using insert.hyps insert.prems by auto qed auto lemma psubset_card_mono: "finite B \ A < B \ card A < card B" using card_seteq [of B A] by (auto simp add: psubset_eq) lemma card_Un_Int: assumes "finite A" "finite B" shows "card A + card B = card (A \ B) + card (A \ B)" using assms proof (induct A) case empty then show ?case by simp next case insert then show ?case by (auto simp add: insert_absorb Int_insert_left) qed lemma card_Un_disjoint: "finite A \ finite B \ A \ B = {} \ card (A \ B) = card A + card B" using card_Un_Int [of A B] by simp lemma card_Un_disjnt: "\finite A; finite B; disjnt A B\ \ card (A \ B) = card A + card B" by (simp add: card_Un_disjoint disjnt_def) lemma card_Un_le: "card (A \ B) \ card A + card B" proof (cases "finite A \ finite B") case True then show ?thesis using le_iff_add card_Un_Int [of A B] by auto qed auto lemma card_Diff_subset: assumes "finite B" and "B \ A" shows "card (A - B) = card A - card B" using assms proof (cases "finite A") case False with assms show ?thesis by simp next case True with assms show ?thesis by (induct B arbitrary: A) simp_all qed lemma card_Diff_subset_Int: assumes "finite (A \ B)" shows "card (A - B) = card A - card (A \ B)" proof - have "A - B = A - A \ B" by auto with assms show ?thesis by (simp add: card_Diff_subset) qed lemma diff_card_le_card_Diff: assumes "finite B" shows "card A - card B \ card (A - B)" proof - have "card A - card B \ card A - card (A \ B)" using card_mono[OF assms Int_lower2, of A] by arith also have "\ = card (A - B)" using assms by (simp add: card_Diff_subset_Int) finally show ?thesis . qed lemma card_le_sym_Diff: assumes "finite A" "finite B" "card A \ card B" shows "card(A - B) \ card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ \ card B - card (A \ B)" using assms(3) by linarith also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_less_sym_Diff: assumes "finite A" "finite B" "card A < card B" shows "card(A - B) < card(B - A)" proof - have "card(A - B) = card A - card (A \ B)" using assms(1,2) by(simp add: card_Diff_subset_Int) also have "\ < card B - card (A \ B)" using assms(1,3) by (simp add: card_mono diff_less_mono) also have "\ = card(B - A)" using assms(1,2) by(simp add: card_Diff_subset_Int Int_commute) finally show ?thesis . qed lemma card_Diff1_less_iff: "card (A - {x}) < card A \ finite A \ x \ A" proof (cases "finite A \ x \ A") case True then show ?thesis by (auto simp: card_gt_0_iff intro: diff_less) qed auto lemma card_Diff1_less: "finite A \ x \ A \ card (A - {x}) < card A" unfolding card_Diff1_less_iff by auto lemma card_Diff2_less: assumes "finite A" "x \ A" "y \ A" shows "card (A - {x} - {y}) < card A" proof (cases "x = y") case True with assms show ?thesis by (simp add: card_Diff1_less del: card_Diff_insert) next case False then have "card (A - {x} - {y}) < card (A - {x})" "card (A - {x}) < card A" using assms by (intro card_Diff1_less; simp)+ then show ?thesis by (blast intro: less_trans) qed lemma card_Diff1_le: "card (A - {x}) \ card A" proof (cases "finite A") case True then show ?thesis by (cases "x \ A") (simp_all add: card_Diff1_less less_imp_le) qed auto lemma card_psubset: "finite B \ A \ B \ card A < card B \ A < B" by (erule psubsetI) blast lemma card_le_inj: assumes fA: "finite A" and fB: "finite B" and c: "card A \ card B" shows "\f. f ` A \ B \ inj_on f A" using fA fB c proof (induct arbitrary: B rule: finite_induct) case empty then show ?case by simp next case (insert x s t) then show ?case proof (induct rule: finite_induct [OF insert.prems(1)]) case 1 then show ?case by simp next case (2 y t) from "2.prems"(1,2,5) "2.hyps"(1,2) have cst: "card s \ card t" by simp from "2.prems"(3) [OF "2.hyps"(1) cst] obtain f where *: "f ` s \ t" "inj_on f s" by blast let ?g = "(\a. if a = x then y else f a)" have "?g ` insert x s \ insert y t \ inj_on ?g (insert x s)" using * "2.prems"(2) "2.hyps"(2) unfolding inj_on_def by auto then show ?case by (rule exI[where ?x="?g"]) qed qed lemma card_subset_eq: assumes fB: "finite B" and AB: "A \ B" and c: "card A = card B" shows "A = B" proof - from fB AB have fA: "finite A" by (auto intro: finite_subset) from fA fB have fBA: "finite (B - A)" by auto have e: "A \ (B - A) = {}" by blast have eq: "A \ (B - A) = B" using AB by blast from card_Un_disjoint[OF fA fBA e, unfolded eq c] have "card (B - A) = 0" by arith then have "B - A = {}" unfolding card_eq_0_iff using fA fB by simp with AB show "A = B" by blast qed lemma insert_partition: "x \ F \ \c1 \ insert x F. \c2 \ insert x F. c1 \ c2 \ c1 \ c2 = {} \ x \ \F = {}" by auto lemma finite_psubset_induct [consumes 1, case_names psubset]: assumes finite: "finite A" and major: "\A. finite A \ (\B. B \ A \ P B) \ P A" shows "P A" using finite proof (induct A taking: card rule: measure_induct_rule) case (less A) have fin: "finite A" by fact have ih: "card B < card A \ finite B \ P B" for B by fact have "P B" if "B \ A" for B proof - from that have "card B < card A" using psubset_card_mono fin by blast moreover from that have "B \ A" by auto then have "finite B" using fin finite_subset by blast ultimately show ?thesis using ih by simp qed with fin show "P A" using major by blast qed lemma finite_induct_select [consumes 1, case_names empty select]: assumes "finite S" and "P {}" and select: "\T. T \ S \ P T \ \s\S - T. P (insert s T)" shows "P S" proof - have "0 \ card S" by simp then have "\T \ S. card T = card S \ P T" proof (induct rule: dec_induct) case base with \P {}\ show ?case by (intro exI[of _ "{}"]) auto next case (step n) then obtain T where T: "T \ S" "card T = n" "P T" by auto with \n < card S\ have "T \ S" "P T" by auto with select[of T] obtain s where "s \ S" "s \ T" "P (insert s T)" by auto with step(2) T \finite S\ show ?case by (intro exI[of _ "insert s T"]) (auto dest: finite_subset) qed with \finite S\ show "P S" by (auto dest: card_subset_eq) qed lemma remove_induct [case_names empty infinite remove]: assumes empty: "P ({} :: 'a set)" and infinite: "\ finite B \ P B" and remove: "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" shows "P B" proof (cases "finite B") case False then show ?thesis by (rule infinite) next case True define A where "A = B" with True have "finite A" "A \ B" by simp_all then show "P A" proof (induct "card A" arbitrary: A) case 0 then have "A = {}" by auto with empty show ?case by simp next case (Suc n A) from \A \ B\ and \finite B\ have "finite A" by (rule finite_subset) moreover from Suc.hyps have "A \ {}" by auto moreover note \A \ B\ moreover have "P (A - {x})" if x: "x \ A" for x using x Suc.prems \Suc n = card A\ by (intro Suc) auto ultimately show ?case by (rule remove) qed qed lemma finite_remove_induct [consumes 1, case_names empty remove]: fixes P :: "'a set \ bool" assumes "finite B" and "P {}" and "\A. finite A \ A \ {} \ A \ B \ (\x. x \ A \ P (A - {x})) \ P A" defines "B' \ B" shows "P B'" by (induct B' rule: remove_induct) (simp_all add: assms) text \Main cardinality theorem.\ lemma card_partition [rule_format]: "finite C \ finite (\C) \ (\c\C. card c = k) \ (\c1 \ C. \c2 \ C. c1 \ c2 \ c1 \ c2 = {}) \ k * card C = card (\C)" proof (induct rule: finite_induct) case empty then show ?case by simp next case (insert x F) then show ?case by (simp add: card_Un_disjoint insert_partition finite_subset [of _ "\(insert _ _)"]) qed lemma card_eq_UNIV_imp_eq_UNIV: assumes fin: "finite (UNIV :: 'a set)" and card: "card A = card (UNIV :: 'a set)" shows "A = (UNIV :: 'a set)" proof show "A \ UNIV" by simp show "UNIV \ A" proof show "x \ A" for x proof (rule ccontr) assume "x \ A" then have "A \ UNIV" by auto with fin have "card A < card (UNIV :: 'a set)" by (fact psubset_card_mono) with card show False by simp qed qed qed text \The form of a finite set of given cardinality\ lemma card_eq_SucD: assumes "card A = Suc k" shows "\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {})" proof - have fin: "finite A" using assms by (auto intro: ccontr) moreover have "card A \ 0" using assms by auto ultimately obtain b where b: "b \ A" by auto show ?thesis proof (intro exI conjI) show "A = insert b (A - {b})" using b by blast show "b \ A - {b}" by blast show "card (A - {b}) = k" and "k = 0 \ A - {b} = {}" using assms b fin by (fastforce dest: mk_disjoint_insert)+ qed qed lemma card_Suc_eq: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ (k = 0 \ B = {}))" by (auto simp: card_insert_if card_gt_0_iff elim!: card_eq_SucD) lemma card_Suc_eq_finite: "card A = Suc k \ (\b B. A = insert b B \ b \ B \ card B = k \ finite B)" unfolding card_Suc_eq using card_gt_0_iff by fastforce lemma card_1_singletonE: assumes "card A = 1" obtains x where "A = {x}" using assms by (auto simp: card_Suc_eq) lemma is_singleton_altdef: "is_singleton A \ card A = 1" unfolding is_singleton_def by (auto elim!: card_1_singletonE is_singletonE simp del: One_nat_def) lemma card_1_singleton_iff: "card A = Suc 0 \ (\x. A = {x})" by (simp add: card_Suc_eq) lemma card_le_Suc0_iff_eq: assumes "finite A" shows "card A \ Suc 0 \ (\a1 \ A. \a2 \ A. a1 = a2)" (is "?C = ?A") proof assume ?C thus ?A using assms by (auto simp: le_Suc_eq dest: card_eq_SucD) next assume ?A show ?C proof cases assume "A = {}" thus ?C using \?A\ by simp next assume "A \ {}" then obtain a where "A = {a}" using \?A\ by blast thus ?C by simp qed qed lemma card_le_Suc_iff: "Suc n \ card A = (\a B. A = insert a B \ a \ B \ n \ card B \ finite B)" proof (cases "finite A") case True then show ?thesis by (fastforce simp: card_Suc_eq less_eq_nat.simps split: nat.splits) qed auto lemma finite_fun_UNIVD2: assumes fin: "finite (UNIV :: ('a \ 'b) set)" shows "finite (UNIV :: 'b set)" proof - from fin have "finite (range (\f :: 'a \ 'b. f arbitrary))" for arbitrary by (rule finite_imageI) moreover have "UNIV = range (\f :: 'a \ 'b. f arbitrary)" for arbitrary by (rule UNIV_eq_I) auto ultimately show "finite (UNIV :: 'b set)" by simp qed lemma card_UNIV_unit [simp]: "card (UNIV :: unit set) = 1" unfolding UNIV_unit by simp lemma infinite_arbitrarily_large: assumes "\ finite A" shows "\B. finite B \ card B = n \ B \ A" proof (induction n) case 0 show ?case by (intro exI[of _ "{}"]) auto next case (Suc n) then obtain B where B: "finite B \ card B = n \ B \ A" .. with \\ finite A\ have "A \ B" by auto with B have "B \ A" by auto then have "\x. x \ A - B" by (elim psubset_imp_ex_mem) then obtain x where x: "x \ A - B" .. with B have "finite (insert x B) \ card (insert x B) = Suc n \ insert x B \ A" by auto then show "\B. finite B \ card B = Suc n \ B \ A" .. qed text \Sometimes, to prove that a set is finite, it is convenient to work with finite subsets and to show that their cardinalities are uniformly bounded. This possibility is formalized in the next criterion.\ lemma finite_if_finite_subsets_card_bdd: assumes "\G. G \ F \ finite G \ card G \ C" shows "finite F \ card F \ C" proof (cases "finite F") case False obtain n::nat where n: "n > max C 0" by auto obtain G where G: "G \ F" "card G = n" using infinite_arbitrarily_large[OF False] by auto hence "finite G" using \n > max C 0\ using card.infinite gr_implies_not0 by blast hence False using assms G n not_less by auto thus ?thesis .. next case True thus ?thesis using assms[of F] by auto qed lemma obtain_subset_with_card_n: assumes "n \ card S" obtains T where "T \ S" "card T = n" "finite T" proof - obtain n' where "card S = n + n'" using le_Suc_ex[OF assms] by blast with that show thesis proof (induct n' arbitrary: S) case 0 thus ?case by (cases "finite S") auto next case Suc thus ?case by (auto simp add: card_Suc_eq) qed qed lemma exists_subset_between: assumes "card A \ n" "n \ card C" "A \ C" "finite C" shows "\B. A \ B \ B \ C \ card B = n" using assms proof (induct n arbitrary: A C) case 0 thus ?case using finite_subset[of A C] by (intro exI[of _ "{}"], auto) next case (Suc n A C) show ?case proof (cases "A = {}") case True from obtain_subset_with_card_n[OF Suc(3)] obtain B where "B \ C" "card B = Suc n" by blast thus ?thesis unfolding True by blast next case False then obtain a where a: "a \ A" by auto let ?A = "A - {a}" let ?C = "C - {a}" have 1: "card ?A \ n" using Suc(2-) a using finite_subset by fastforce have 2: "card ?C \ n" using Suc(2-) a by auto from Suc(1)[OF 1 2 _ finite_subset[OF _ Suc(5)]] Suc(2-) obtain B where "?A \ B" "B \ ?C" "card B = n" by blast thus ?thesis using a Suc(2-) by (intro exI[of _ "insert a B"], auto intro!: card_insert_disjoint finite_subset[of B C]) qed qed subsubsection \Cardinality of image\ lemma card_image_le: "finite A \ card (f ` A) \ card A" by (induct rule: finite_induct) (simp_all add: le_SucI card_insert_if) lemma card_image: "inj_on f A \ card (f ` A) = card A" proof (induct A rule: infinite_finite_induct) case (infinite A) then have "\ finite (f ` A)" by (auto dest: finite_imageD) with infinite show ?case by simp qed simp_all lemma bij_betw_same_card: "bij_betw f A B \ card A = card B" by (auto simp: card_image bij_betw_def) lemma endo_inj_surj: "finite A \ f ` A \ A \ inj_on f A \ f ` A = A" by (simp add: card_seteq card_image) lemma eq_card_imp_inj_on: assumes "finite A" "card(f ` A) = card A" shows "inj_on f A" using assms proof (induct rule:finite_induct) case empty show ?case by simp next case (insert x A) then show ?case using card_image_le [of A f] by (simp add: card_insert_if split: if_splits) qed lemma inj_on_iff_eq_card: "finite A \ inj_on f A \ card (f ` A) = card A" by (blast intro: card_image eq_card_imp_inj_on) lemma card_inj_on_le: assumes "inj_on f A" "f ` A \ B" "finite B" shows "card A \ card B" proof - have "finite A" using assms by (blast intro: finite_imageD dest: finite_subset) then show ?thesis using assms by (force intro: card_mono simp: card_image [symmetric]) qed lemma inj_on_iff_card_le: "\ finite A; finite B \ \ (\f. inj_on f A \ f ` A \ B) = (card A \ card B)" using card_inj_on_le[of _ A B] card_le_inj[of A B] by blast lemma surj_card_le: "finite A \ B \ f ` A \ card B \ card A" by (blast intro: card_image_le card_mono le_trans) lemma card_bij_eq: "inj_on f A \ f ` A \ B \ inj_on g B \ g ` B \ A \ finite A \ finite B \ card A = card B" by (auto intro: le_antisym card_inj_on_le) lemma bij_betw_finite: "bij_betw f A B \ finite A \ finite B" unfolding bij_betw_def using finite_imageD [of f A] by auto lemma inj_on_finite: "inj_on f A \ f ` A \ B \ finite B \ finite A" using finite_imageD finite_subset by blast lemma card_vimage_inj_on_le: assumes "inj_on f D" "finite A" shows "card (f-`A \ D) \ card A" proof (rule card_inj_on_le) show "inj_on f (f -` A \ D)" by (blast intro: assms inj_on_subset) qed (use assms in auto) lemma card_vimage_inj: "inj f \ A \ range f \ card (f -` A) = card A" by (auto 4 3 simp: subset_image_iff inj_vimage_image_eq intro: card_image[symmetric, OF subset_inj_on]) lemma card_inverse[simp]: "card (R\) = card R" proof - have *: "\R. prod.swap ` R = R\" by auto { assume "\finite R" hence ?thesis by auto } moreover { assume "finite R" with card_image_le[of R prod.swap] card_image_le[of "R\" prod.swap] have ?thesis by (auto simp: * ) } ultimately show ?thesis by blast qed subsubsection \Pigeonhole Principles\ lemma pigeonhole: "card A > card (f ` A) \ \ inj_on f A " by (auto dest: card_image less_irrefl_nat) lemma pigeonhole_infinite: assumes "\ finite A" and "finite (f`A)" shows "\a0\A. \ finite {a\A. f a = f a0}" using assms(2,1) proof (induct "f`A" arbitrary: A rule: finite_induct) case empty then show ?case by simp next case (insert b F) show ?case proof (cases "finite {a\A. f a = b}") case True with \\ finite A\ have "\ finite (A - {a\A. f a = b})" by simp also have "A - {a\A. f a = b} = {a\A. f a \ b}" by blast finally have "\ finite {a\A. f a \ b}" . from insert(3)[OF _ this] insert(2,4) show ?thesis by simp (blast intro: rev_finite_subset) next case False then have "{a \ A. f a = b} \ {}" by force with False show ?thesis by blast qed qed lemma pigeonhole_infinite_rel: assumes "\ finite A" and "finite B" and "\a\A. \b\B. R a b" shows "\b\B. \ finite {a:A. R a b}" proof - let ?F = "\a. {b\B. R a b}" from finite_Pow_iff[THEN iffD2, OF \finite B\] have "finite (?F ` A)" by (blast intro: rev_finite_subset) from pigeonhole_infinite [where f = ?F, OF assms(1) this] obtain a0 where "a0 \ A" and infinite: "\ finite {a\A. ?F a = ?F a0}" .. obtain b0 where "b0 \ B" and "R a0 b0" using \a0 \ A\ assms(3) by blast have "finite {a\A. ?F a = ?F a0}" if "finite {a\A. R a b0}" using \b0 \ B\ \R a0 b0\ that by (blast intro: rev_finite_subset) with infinite \b0 \ B\ show ?thesis by blast qed subsubsection \Cardinality of sums\ lemma card_Plus: assumes "finite A" "finite B" shows "card (A <+> B) = card A + card B" proof - have "Inl`A \ Inr`B = {}" by fast with assms show ?thesis by (simp add: Plus_def card_Un_disjoint card_image) qed lemma card_Plus_conv_if: "card (A <+> B) = (if finite A \ finite B then card A + card B else 0)" by (auto simp add: card_Plus) text \Relates to equivalence classes. Based on a theorem of F. Kammüller.\ lemma dvd_partition: assumes f: "finite (\C)" and "\c\C. k dvd card c" "\c1\C. \c2\C. c1 \ c2 \ c1 \ c2 = {}" shows "k dvd card (\C)" proof - have "finite C" by (rule finite_UnionD [OF f]) then show ?thesis using assms proof (induct rule: finite_induct) case empty show ?case by simp next case (insert c C) then have "c \ \C = {}" by auto with insert show ?case by (simp add: card_Un_disjoint) qed qed subsection \Minimal and maximal elements of finite sets\ context begin qualified lemma assumes "finite A" and "asymp_on A R" and "transp_on A R" and "\x \ A. P x" shows bex_min_element_with_property: "\x \ A. P x \ (\y \ A. R y x \ \ P y)" and bex_max_element_with_property: "\x \ A. P x \ (\y \ A. R x y \ \ P y)" unfolding atomize_conj using assms proof (induction A rule: finite_induct) case empty hence False by simp_all thus ?case .. next case (insert x F) from insert.prems have "asymp_on F R" using asymp_on_subset by blast from insert.prems have "transp_on F R" using transp_on_subset by blast show ?case proof (cases "P x") case True show ?thesis proof (cases "\a\F. P a") case True with insert.IH obtain min max where "min \ F" and "P min" and "\z \ F. R z min \ \ P z" "max \ F" and "P max" and "\z \ F. R max z \ \ P z" using \asymp_on F R\ \transp_on F R\ by auto show ?thesis proof (rule conjI) show "\y \ insert x F. P y \ (\z \ insert x F. R y z \ \ P z)" proof (cases "R max x") case True show ?thesis proof (intro bexI conjI ballI impI) show "x \ insert x F" by simp next show "P x" using \P x\ by simp next fix z assume "z \ insert x F" and "R x z" hence "z = x \ z \ F" by simp thus "\ P z" proof (rule disjE) assume "z = x" hence "R x x" using \R x z\ by simp moreover have "\ R x x" using \asymp_on (insert x F) R\[THEN irreflp_on_if_asymp_on, THEN irreflp_onD] by simp ultimately have False by simp thus ?thesis .. next assume "z \ F" moreover have "R max z" using \R max x\ \R x z\ using \transp_on (insert x F) R\[THEN transp_onD, of max x z] using \max \ F\ \z \ F\ by simp ultimately show ?thesis using \\z \ F. R max z \ \ P z\ by simp qed qed next case False show ?thesis proof (intro bexI conjI ballI impI) show "max \ insert x F" using \max \ F\ by simp next show "P max" using \P max\ by simp next fix z assume "z \ insert x F" and "R max z" hence "z = x \ z \ F" by simp thus "\ P z" proof (rule disjE) assume "z = x" hence False using \\ R max x\ \R max z\ by simp thus ?thesis .. next assume "z \ F" thus ?thesis using \R max z\ \\z\F. R max z \ \ P z\ by simp qed qed qed next show "\y \ insert x F. P y \ (\z \ insert x F. R z y \ \ P z)" proof (cases "R x min") case True show ?thesis proof (intro bexI conjI ballI impI) show "x \ insert x F" by simp next show "P x" using \P x\ by simp next fix z assume "z \ insert x F" and "R z x" hence "z = x \ z \ F" by simp thus "\ P z" proof (rule disjE) assume "z = x" hence "R x x" using \R z x\ by simp moreover have "\ R x x" using \asymp_on (insert x F) R\[THEN irreflp_on_if_asymp_on, THEN irreflp_onD] by simp ultimately have False by simp thus ?thesis .. next assume "z \ F" moreover have "R z min" using \R z x\ \R x min\ using \transp_on (insert x F) R\[THEN transp_onD, of z x min] using \min \ F\ \z \ F\ by simp ultimately show ?thesis using \\z \ F. R z min \ \ P z\ by simp qed qed next case False show ?thesis proof (intro bexI conjI ballI impI) show "min \ insert x F" using \min \ F\ by simp next show "P min" using \P min\ by simp next fix z assume "z \ insert x F" and "R z min" hence "z = x \ z \ F" by simp thus "\ P z" proof (rule disjE) assume "z = x" hence False using \\ R x min\ \R z min\ by simp thus ?thesis .. next assume "z \ F" thus ?thesis using \R z min\ \\z\F. R z min \ \ P z\ by simp qed qed qed qed next case False then show ?thesis using \\a\insert x F. P a\ using \asymp_on (insert x F) R\[THEN asymp_onD, of x] insert_iff[of _ x F] by blast qed next case False with insert.prems have "\x \ F. P x" by simp with insert.IH have "\y \ F. P y \ (\z\F. R z y \ \ P z)" "\y \ F. P y \ (\z\F. R y z \ \ P z)" using \asymp_on F R\ \transp_on F R\ by auto thus ?thesis using False by auto qed qed qualified lemma assumes "finite A" and "asymp_on A R" and "transp_on A R" and "A \ {}" shows bex_min_element: "\m \ A. \x \ A. x \ m \ \ R x m" and bex_max_element: "\m \ A. \x \ A. x \ m \ \ R m x" using \A \ {}\ bex_min_element_with_property[OF assms(1,2,3), of "\_. True", simplified] bex_max_element_with_property[OF assms(1,2,3), of "\_. True", simplified] by blast+ end text \The following alternative form might sometimes be easier to work with.\ lemma is_min_element_in_set_iff: "asymp_on A R \ (\y \ A. y \ x \ \ R y x) \ (\y. R y x \ y \ A)" by (auto dest: asymp_onD) lemma is_max_element_in_set_iff: "asymp_on A R \ (\y \ A. y \ x \ \ R x y) \ (\y. R x y \ y \ A)" by (auto dest: asymp_onD) context begin qualified lemma assumes "finite A" and "A \ {}" and "transp_on A R" and "totalp_on A R" shows bex_least_element: "\l \ A. \x \ A. x \ l \ R l x" and bex_greatest_element: "\g \ A. \x \ A. x \ g \ R x g" unfolding atomize_conj using assms proof (induction A rule: finite_induct) case empty hence False by simp thus ?case .. next case (insert a A') from insert.prems(2) have transp_on_A': "transp_on A' R" by (auto intro: transp_onI dest: transp_onD) from insert.prems(3) have totalp_on_a_A'_raw: "\y \ A'. a \ y \ R a y \ R y a" and totalp_on_A': "totalp_on A' R" by (simp_all add: totalp_on_def) show ?case proof (cases "A' = {}") case True thus ?thesis by simp next case False then obtain least greatest where "least \ A'" and least_of_A': "\x\A'. x \ least \ R least x" and "greatest \ A'" and greatest_of_A': "\x\A'. x \ greatest \ R x greatest" using insert.IH[OF _ transp_on_A' totalp_on_A'] by auto show ?thesis proof (rule conjI) show "\l\insert a A'. \x\insert a A'. x \ l \ R l x" proof (cases "R a least") case True show ?thesis proof (intro bexI ballI impI) show "a \ insert a A'" by simp next fix x show "\x. x \ insert a A' \ x \ a \ R a x" using True \least \ A'\ least_of_A' using insert.prems(2)[THEN transp_onD, of a least] by auto qed next case False show ?thesis proof (intro bexI ballI impI) show "least \ insert a A'" using \least \ A'\ by simp next fix x show "x \ insert a A' \ x \ least \ R least x" using False \least \ A'\ least_of_A' totalp_on_a_A'_raw by (cases "x = a") auto qed qed next show "\g \ insert a A'. \x \ insert a A'. x \ g \ R x g" proof (cases "R greatest a") case True show ?thesis proof (intro bexI ballI impI) show "a \ insert a A'" by simp next fix x show "\x. x \ insert a A' \ x \ a \ R x a" using True \greatest \ A'\ greatest_of_A' using insert.prems(2)[THEN transp_onD, of _ greatest a] by auto qed next case False show ?thesis proof (intro bexI ballI impI) show "greatest \ insert a A'" using \greatest \ A'\ by simp next fix x show "x \ insert a A' \ x \ greatest \ R x greatest" using False \greatest \ A'\ greatest_of_A' totalp_on_a_A'_raw by (cases "x = a") auto qed qed qed qed qed end subsubsection \Finite orders\ context order begin lemma finite_has_maximal: assumes "finite A" and "A \ {}" shows "\ m \ A. \ b \ A. m \ b \ m = b" proof - obtain m where "m \ A" and m_is_max: "\x\A. x \ m \ \ m < x" using Finite_Set.bex_max_element[OF \finite A\ _ _ \A \ {}\, of "(<)"] by auto moreover have "\b \ A. m \ b \ m = b" using m_is_max by (auto simp: le_less) ultimately show ?thesis by auto qed lemma finite_has_maximal2: "\ finite A; a \ A \ \ \ m \ A. a \ m \ (\ b \ A. m \ b \ m = b)" using finite_has_maximal[of "{b \ A. a \ b}"] by fastforce lemma finite_has_minimal: assumes "finite A" and "A \ {}" shows "\ m \ A. \ b \ A. b \ m \ m = b" proof - obtain m where "m \ A" and m_is_min: "\x\A. x \ m \ \ x < m" using Finite_Set.bex_min_element[OF \finite A\ _ _ \A \ {}\, of "(<)"] by auto moreover have "\b \ A. b \ m \ m = b" using m_is_min by (auto simp: le_less) ultimately show ?thesis by auto qed lemma finite_has_minimal2: "\ finite A; a \ A \ \ \ m \ A. m \ a \ (\ b \ A. b \ m \ m = b)" using finite_has_minimal[of "{b \ A. b \ a}"] by fastforce end subsubsection \Relating injectivity and surjectivity\ lemma finite_surj_inj: assumes "finite A" "A \ f ` A" shows "inj_on f A" proof - have "f ` A = A" by (rule card_seteq [THEN sym]) (auto simp add: assms card_image_le) then show ?thesis using assms by (simp add: eq_card_imp_inj_on) qed lemma finite_UNIV_surj_inj: "finite(UNIV:: 'a set) \ surj f \ inj f" for f :: "'a \ 'a" by (blast intro: finite_surj_inj subset_UNIV) lemma finite_UNIV_inj_surj: "finite(UNIV:: 'a set) \ inj f \ surj f" for f :: "'a \ 'a" by (fastforce simp:surj_def dest!: endo_inj_surj) lemma surjective_iff_injective_gen: assumes fS: "finite S" and fT: "finite T" and c: "card S = card T" and ST: "f ` S \ T" shows "(\y \ T. \x \ S. f x = y) \ inj_on f S" (is "?lhs \ ?rhs") proof assume h: "?lhs" { fix x y assume x: "x \ S" assume y: "y \ S" assume f: "f x = f y" from x fS have S0: "card S \ 0" by auto have "x = y" proof (rule ccontr) assume xy: "\ ?thesis" have th: "card S \ card (f ` (S - {y}))" unfolding c proof (rule card_mono) show "finite (f ` (S - {y}))" by (simp add: fS) have "\x \ y; x \ S; z \ S; f x = f y\ \ \x \ S. x \ y \ f z = f x" for z by (cases "z = y \ z = x") auto then show "T \ f ` (S - {y})" using h xy x y f by fastforce qed also have " \ \ card (S - {y})" by (simp add: card_image_le fS) also have "\ \ card S - 1" using y fS by simp finally show False using S0 by arith qed } then show ?rhs unfolding inj_on_def by blast next assume h: ?rhs have "f ` S = T" by (simp add: ST c card_image card_subset_eq fT h) then show ?lhs by blast qed hide_const (open) Finite_Set.fold subsection \Infinite Sets\ text \ Some elementary facts about infinite sets, mostly by Stephan Merz. Beware! Because "infinite" merely abbreviates a negation, these lemmas may not work well with \blast\. \ abbreviation infinite :: "'a set \ bool" where "infinite S \ \ finite S" text \ Infinite sets are non-empty, and if we remove some elements from an infinite set, the result is still infinite. \ lemma infinite_UNIV_nat [iff]: "infinite (UNIV :: nat set)" proof assume "finite (UNIV :: nat set)" with finite_UNIV_inj_surj [of Suc] show False by simp (blast dest: Suc_neq_Zero surjD) qed lemma infinite_UNIV_char_0: "infinite (UNIV :: 'a::semiring_char_0 set)" proof assume "finite (UNIV :: 'a set)" with subset_UNIV have "finite (range of_nat :: 'a set)" by (rule finite_subset) moreover have "inj (of_nat :: nat \ 'a)" by (simp add: inj_on_def) ultimately have "finite (UNIV :: nat set)" by (rule finite_imageD) then show False by simp qed lemma infinite_imp_nonempty: "infinite S \ S \ {}" by auto lemma infinite_remove: "infinite S \ infinite (S - {a})" by simp lemma Diff_infinite_finite: assumes "finite T" "infinite S" shows "infinite (S - T)" using \finite T\ proof induct from \infinite S\ show "infinite (S - {})" by auto next fix T x assume ih: "infinite (S - T)" have "S - (insert x T) = (S - T) - {x}" by (rule Diff_insert) with ih show "infinite (S - (insert x T))" by (simp add: infinite_remove) qed lemma Un_infinite: "infinite S \ infinite (S \ T)" by simp lemma infinite_Un: "infinite (S \ T) \ infinite S \ infinite T" by simp lemma infinite_super: assumes "S \ T" and "infinite S" shows "infinite T" proof assume "finite T" with \S \ T\ have "finite S" by (simp add: finite_subset) with \infinite S\ show False by simp qed proposition infinite_coinduct [consumes 1, case_names infinite]: assumes "X A" and step: "\A. X A \ \x\A. X (A - {x}) \ infinite (A - {x})" shows "infinite A" proof assume "finite A" then show False using \X A\ proof (induction rule: finite_psubset_induct) case (psubset A) then obtain x where "x \ A" "X (A - {x}) \ infinite (A - {x})" using local.step psubset.prems by blast then have "X (A - {x})" using psubset.hyps by blast show False proof (rule psubset.IH [where B = "A - {x}"]) show "A - {x} \ A" using \x \ A\ by blast qed fact qed qed text \ For any function with infinite domain and finite range there is some element that is the image of infinitely many domain elements. In particular, any infinite sequence of elements from a finite set contains some element that occurs infinitely often. \ lemma inf_img_fin_dom': assumes img: "finite (f ` A)" and dom: "infinite A" shows "\y \ f ` A. infinite (f -` {y} \ A)" proof (rule ccontr) have "A \ (\y\f ` A. f -` {y} \ A)" by auto moreover assume "\ ?thesis" with img have "finite (\y\f ` A. f -` {y} \ A)" by blast ultimately have "finite A" by (rule finite_subset) with dom show False by contradiction qed lemma inf_img_fin_domE': assumes "finite (f ` A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y} \ A)" using assms by (blast dest: inf_img_fin_dom') lemma inf_img_fin_dom: assumes img: "finite (f`A)" and dom: "infinite A" shows "\y \ f`A. infinite (f -` {y})" using inf_img_fin_dom'[OF assms] by auto lemma inf_img_fin_domE: assumes "finite (f`A)" and "infinite A" obtains y where "y \ f`A" and "infinite (f -` {y})" using assms by (blast dest: inf_img_fin_dom) proposition finite_image_absD: "finite (abs ` S) \ finite S" for S :: "'a::linordered_ring set" by (rule ccontr) (auto simp: abs_eq_iff vimage_def dest: inf_img_fin_dom) subsection \The finite powerset operator\ definition Fpow :: "'a set \ 'a set set" where "Fpow A \ {X. X \ A \ finite X}" lemma Fpow_mono: "A \ B \ Fpow A \ Fpow B" unfolding Fpow_def by auto lemma empty_in_Fpow: "{} \ Fpow A" unfolding Fpow_def by auto lemma Fpow_not_empty: "Fpow A \ {}" using empty_in_Fpow by blast lemma Fpow_subset_Pow: "Fpow A \ Pow A" unfolding Fpow_def by auto lemma Fpow_Pow_finite: "Fpow A = Pow A Int {A. finite A}" unfolding Fpow_def Pow_def by blast lemma inj_on_image_Fpow: assumes "inj_on f A" shows "inj_on (image f) (Fpow A)" using assms Fpow_subset_Pow[of A] subset_inj_on[of "image f" "Pow A"] inj_on_image_Pow by blast lemma image_Fpow_mono: assumes "f ` A \ B" shows "(image f) ` (Fpow A) \ Fpow B" using assms by(unfold Fpow_def, auto) end diff --git a/src/HOL/Fun.thy b/src/HOL/Fun.thy --- a/src/HOL/Fun.thy +++ b/src/HOL/Fun.thy @@ -1,1410 +1,1410 @@ (* Title: HOL/Fun.thy Author: Tobias Nipkow, Cambridge University Computer Laboratory Author: Andrei Popescu, TU Muenchen Copyright 1994, 2012 *) section \Notions about functions\ theory Fun imports Set keywords "functor" :: thy_goal_defn begin lemma apply_inverse: "f x = u \ (\x. P x \ g (f x) = x) \ P x \ x = g u" by auto text \Uniqueness, so NOT the axiom of choice.\ lemma uniq_choice: "\x. \!y. Q x y \ \f. \x. Q x (f x)" by (force intro: theI') lemma b_uniq_choice: "\x\S. \!y. Q x y \ \f. \x\S. Q x (f x)" by (force intro: theI') subsection \The Identity Function \id\\ definition id :: "'a \ 'a" where "id = (\x. x)" lemma id_apply [simp]: "id x = x" by (simp add: id_def) lemma image_id [simp]: "image id = id" by (simp add: id_def fun_eq_iff) lemma vimage_id [simp]: "vimage id = id" by (simp add: id_def fun_eq_iff) lemma eq_id_iff: "(\x. f x = x) \ f = id" by auto code_printing constant id \ (Haskell) "id" subsection \The Composition Operator \f \ g\\ definition comp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "\" 55) where "f \ g = (\x. f (g x))" notation (ASCII) comp (infixl "o" 55) lemma comp_apply [simp]: "(f \ g) x = f (g x)" by (simp add: comp_def) lemma comp_assoc: "(f \ g) \ h = f \ (g \ h)" by (simp add: fun_eq_iff) lemma id_comp [simp]: "id \ g = g" by (simp add: fun_eq_iff) lemma comp_id [simp]: "f \ id = f" by (simp add: fun_eq_iff) lemma comp_eq_dest: "a \ b = c \ d \ a (b v) = c (d v)" by (simp add: fun_eq_iff) lemma comp_eq_elim: "a \ b = c \ d \ ((\v. a (b v) = c (d v)) \ R) \ R" by (simp add: fun_eq_iff) lemma comp_eq_dest_lhs: "a \ b = c \ a (b v) = c v" by clarsimp lemma comp_eq_id_dest: "a \ b = id \ c \ a (b v) = c v" by clarsimp lemma image_comp: "f ` (g ` r) = (f \ g) ` r" by auto lemma vimage_comp: "f -` (g -` x) = (g \ f) -` x" by auto lemma image_eq_imp_comp: "f ` A = g ` B \ (h \ f) ` A = (h \ g) ` B" by (auto simp: comp_def elim!: equalityE) lemma image_bind: "f ` (Set.bind A g) = Set.bind A ((`) f \ g)" by (auto simp add: Set.bind_def) lemma bind_image: "Set.bind (f ` A) g = Set.bind A (g \ f)" by (auto simp add: Set.bind_def) lemma (in group_add) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) lemma (in boolean_algebra) minus_comp_minus [simp]: "uminus \ uminus = id" by (simp add: fun_eq_iff) code_printing constant comp \ (SML) infixl 5 "o" and (Haskell) infixr 9 "." subsection \The Forward Composition Operator \fcomp\\ definition fcomp :: "('a \ 'b) \ ('b \ 'c) \ 'a \ 'c" (infixl "\>" 60) where "f \> g = (\x. g (f x))" lemma fcomp_apply [simp]: "(f \> g) x = g (f x)" by (simp add: fcomp_def) lemma fcomp_assoc: "(f \> g) \> h = f \> (g \> h)" by (simp add: fcomp_def) lemma id_fcomp [simp]: "id \> g = g" by (simp add: fcomp_def) lemma fcomp_id [simp]: "f \> id = f" by (simp add: fcomp_def) lemma fcomp_comp: "fcomp f g = comp g f" by (simp add: ext) code_printing constant fcomp \ (Eval) infixl 1 "#>" no_notation fcomp (infixl "\>" 60) subsection \Mapping functions\ definition map_fun :: "('c \ 'a) \ ('b \ 'd) \ ('a \ 'b) \ 'c \ 'd" where "map_fun f g h = g \ h \ f" lemma map_fun_apply [simp]: "map_fun f g h x = g (h (f x))" by (simp add: map_fun_def) subsection \Injectivity and Bijectivity\ definition inj_on :: "('a \ 'b) \ 'a set \ bool" \ \injective\ where "inj_on f A \ (\x\A. \y\A. f x = f y \ x = y)" definition bij_betw :: "('a \ 'b) \ 'a set \ 'b set \ bool" \ \bijective\ where "bij_betw f A B \ inj_on f A \ f ` A = B" text \ A common special case: functions injective, surjective or bijective over the entire domain type. \ abbreviation inj :: "('a \ 'b) \ bool" where "inj f \ inj_on f UNIV" abbreviation surj :: "('a \ 'b) \ bool" where "surj f \ range f = UNIV" translations \ \The negated case:\ "\ CONST surj f" \ "CONST range f \ CONST UNIV" abbreviation bij :: "('a \ 'b) \ bool" where "bij f \ bij_betw f UNIV UNIV" lemma inj_def: "inj f \ (\x y. f x = f y \ x = y)" unfolding inj_on_def by blast lemma injI: "(\x y. f x = f y \ x = y) \ inj f" unfolding inj_def by blast theorem range_ex1_eq: "inj f \ b \ range f \ (\!x. b = f x)" unfolding inj_def by blast lemma injD: "inj f \ f x = f y \ x = y" by (simp add: inj_def) lemma inj_on_eq_iff: "inj_on f A \ x \ A \ y \ A \ f x = f y \ x = y" by (auto simp: inj_on_def) lemma inj_on_cong: "(\a. a \ A \ f a = g a) \ inj_on f A \ inj_on g A" by (auto simp: inj_on_def) lemma image_strict_mono: "inj_on f B \ A \ B \ f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_compose: "inj f \ inj g \ inj (f \ g)" by (simp add: inj_def) lemma inj_fun: "inj f \ inj (\x y. f x)" by (simp add: inj_def fun_eq_iff) lemma inj_eq: "inj f \ f x = f y \ x = y" by (simp add: inj_on_eq_iff) lemma inj_on_iff_Uniq: "inj_on f A \ (\x\A. \\<^sub>\\<^sub>1y. y\A \ f x = f y)" by (auto simp: Uniq_def inj_on_def) lemma inj_on_id[simp]: "inj_on id A" by (simp add: inj_on_def) lemma inj_on_id2[simp]: "inj_on (\x. x) A" by (simp add: inj_on_def) lemma inj_on_Int: "inj_on f A \ inj_on f B \ inj_on f (A \ B)" unfolding inj_on_def by blast lemma surj_id: "surj id" by simp lemma bij_id[simp]: "bij id" by (simp add: bij_betw_def) lemma bij_uminus: "bij (uminus :: 'a \ 'a::group_add)" unfolding bij_betw_def inj_on_def by (force intro: minus_minus [symmetric]) lemma bij_betwE: "bij_betw f A B \ \a\A. f a \ B" unfolding bij_betw_def by auto lemma inj_onI [intro?]: "(\x y. x \ A \ y \ A \ f x = f y \ x = y) \ inj_on f A" by (simp add: inj_on_def) lemma inj_on_inverseI: "(\x. x \ A \ g (f x) = x) \ inj_on f A" by (auto dest: arg_cong [of concl: g] simp add: inj_on_def) lemma inj_onD: "inj_on f A \ f x = f y \ x \ A \ y \ A \ x = y" unfolding inj_on_def by blast lemma inj_on_subset: assumes "inj_on f A" and "B \ A" shows "inj_on f B" proof (rule inj_onI) fix a b assume "a \ B" and "b \ B" with assms have "a \ A" and "b \ A" by auto moreover assume "f a = f b" ultimately show "a = b" using assms by (auto dest: inj_onD) qed lemma comp_inj_on: "inj_on f A \ inj_on g (f ` A) \ inj_on (g \ f) A" by (simp add: comp_def inj_on_def) lemma inj_on_imageI: "inj_on (g \ f) A \ inj_on g (f ` A)" by (auto simp add: inj_on_def) lemma inj_on_image_iff: "\x\A. \y\A. g (f x) = g (f y) \ g x = g y \ inj_on f A \ inj_on g (f ` A) \ inj_on g A" unfolding inj_on_def by blast lemma inj_on_contraD: "inj_on f A \ x \ y \ x \ A \ y \ A \ f x \ f y" unfolding inj_on_def by blast lemma inj_singleton [simp]: "inj_on (\x. {x}) A" by (simp add: inj_on_def) lemma inj_on_empty[iff]: "inj_on f {}" by (simp add: inj_on_def) lemma subset_inj_on: "inj_on f B \ A \ B \ inj_on f A" unfolding inj_on_def by blast lemma inj_on_Un: "inj_on f (A \ B) \ inj_on f A \ inj_on f B \ f ` (A - B) \ f ` (B - A) = {}" unfolding inj_on_def by (blast intro: sym) lemma inj_on_insert [iff]: "inj_on f (insert a A) \ inj_on f A \ f a \ f ` (A - {a})" unfolding inj_on_def by (blast intro: sym) lemma inj_on_diff: "inj_on f A \ inj_on f (A - B)" unfolding inj_on_def by blast lemma comp_inj_on_iff: "inj_on f A \ inj_on f' (f ` A) \ inj_on (f' \ f) A" by (auto simp: comp_inj_on inj_on_def) lemma inj_on_imageI2: "inj_on (f' \ f) A \ inj_on f A" by (auto simp: comp_inj_on inj_on_def) lemma inj_img_insertE: assumes "inj_on f A" assumes "x \ B" and "insert x B = f ` A" obtains x' A' where "x' \ A'" and "A = insert x' A'" and "x = f x'" and "B = f ` A'" proof - from assms have "x \ f ` A" by auto then obtain x' where *: "x' \ A" "x = f x'" by auto then have A: "A = insert x' (A - {x'})" by auto with assms * have B: "B = f ` (A - {x'})" by (auto dest: inj_on_contraD) have "x' \ A - {x'}" by simp from this A \x = f x'\ B show ?thesis .. qed lemma linorder_inj_onI: fixes A :: "'a::order set" assumes ne: "\x y. \x < y; x\A; y\A\ \ f x \ f y" and lin: "\x y. \x\A; y\A\ \ x\y \ y\x" shows "inj_on f A" proof (rule inj_onI) fix x y assume eq: "f x = f y" and "x\A" "y\A" then show "x = y" using lin [of x y] ne by (force simp: dual_order.order_iff_strict) qed lemma linorder_inj_onI': fixes A :: "'a :: linorder set" assumes "\i j. i \ A \ j \ A \ i < j \ f i \ f j" shows "inj_on f A" by (intro linorder_inj_onI) (auto simp add: assms) lemma linorder_injI: assumes "\x y::'a::linorder. x < y \ f x \ f y" shows "inj f" \ \Courtesy of Stephan Merz\ using assms by (simp add: linorder_inj_onI') lemma inj_on_image_Pow: "inj_on f A \inj_on (image f) (Pow A)" unfolding Pow_def inj_on_def by blast lemma bij_betw_image_Pow: "bij_betw f A B \ bij_betw (image f) (Pow A) (Pow B)" by (auto simp add: bij_betw_def inj_on_image_Pow image_Pow_surj) lemma surj_def: "surj f \ (\y. \x. y = f x)" by auto lemma surjI: assumes "\x. g (f x) = x" shows "surj g" using assms [symmetric] by auto lemma surjD: "surj f \ \x. y = f x" by (simp add: surj_def) lemma surjE: "surj f \ (\x. y = f x \ C) \ C" by (simp add: surj_def) blast lemma comp_surj: "surj f \ surj g \ surj (g \ f)" using image_comp [of g f UNIV] by simp lemma bij_betw_imageI: "inj_on f A \ f ` A = B \ bij_betw f A B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj_on: "bij_betw f A B \ f ` A = B" unfolding bij_betw_def by clarify lemma bij_betw_imp_surj: "bij_betw f A UNIV \ surj f" unfolding bij_betw_def by auto lemma bij_betw_empty1: "bij_betw f {} A \ A = {}" unfolding bij_betw_def by blast lemma bij_betw_empty2: "bij_betw f A {} \ A = {}" unfolding bij_betw_def by blast lemma inj_on_imp_bij_betw: "inj_on f A \ bij_betw f A (f ` A)" unfolding bij_betw_def by simp lemma bij_betw_DiffI: assumes "bij_betw f A B" "bij_betw f C D" "C \ A" "D \ B" shows "bij_betw f (A - C) (B - D)" using assms unfolding bij_betw_def inj_on_def by auto lemma bij_betw_singleton_iff [simp]: "bij_betw f {x} {y} \ f x = y" by (auto simp: bij_betw_def) lemma bij_betw_singletonI [intro]: "f x = y \ bij_betw f {x} {y}" by auto lemma bij_betw_apply: "\bij_betw f A B; a \ A\ \ f a \ B" unfolding bij_betw_def by auto lemma bij_def: "bij f \ inj f \ surj f" by (rule bij_betw_def) lemma bijI: "inj f \ surj f \ bij f" by (rule bij_betw_imageI) lemma bij_is_inj: "bij f \ inj f" by (simp add: bij_def) lemma bij_is_surj: "bij f \ surj f" by (simp add: bij_def) lemma bij_betw_imp_inj_on: "bij_betw f A B \ inj_on f A" by (simp add: bij_betw_def) lemma bij_betw_trans: "bij_betw f A B \ bij_betw g B C \ bij_betw (g \ f) A C" by (auto simp add:bij_betw_def comp_inj_on) lemma bij_comp: "bij f \ bij g \ bij (g \ f)" by (rule bij_betw_trans) lemma bij_betw_comp_iff: "bij_betw f A A' \ bij_betw f' A' A'' \ bij_betw (f' \ f) A A''" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_comp_iff2: assumes bij: "bij_betw f' A' A''" and img: "f ` A \ A'" shows "bij_betw f A A' \ bij_betw (f' \ f) A A''" (is "?L \ ?R") proof assume "?L" then show "?R" using assms by (auto simp add: bij_betw_comp_iff) next assume *: "?R" have "inj_on (f' \ f) A \ inj_on f A" using inj_on_imageI2 by blast moreover have "A' \ f ` A" proof fix a' assume **: "a' \ A'" with bij have "f' a' \ A''" unfolding bij_betw_def by auto with * obtain a where 1: "a \ A \ f' (f a) = f' a'" unfolding bij_betw_def by force with img have "f a \ A'" by auto with bij ** 1 have "f a = a'" unfolding bij_betw_def inj_on_def by auto with 1 show "a' \ f ` A" by auto qed ultimately show "?L" using img * by (auto simp add: bij_betw_def) qed lemma bij_betw_inv: assumes "bij_betw f A B" shows "\g. bij_betw g B A" proof - have i: "inj_on f A" and s: "f ` A = B" using assms by (auto simp: bij_betw_def) let ?P = "\b a. a \ A \ f a = b" let ?g = "\b. The (?P b)" have g: "?g b = a" if P: "?P b a" for a b proof - from that s have ex1: "\a. ?P b a" by blast then have uex1: "\!a. ?P b a" by (blast dest:inj_onD[OF i]) then show ?thesis using the1_equality[OF uex1, OF P] P by simp qed have "inj_on ?g B" proof (rule inj_onI) fix x y assume "x \ B" "y \ B" "?g x = ?g y" from s \x \ B\ obtain a1 where a1: "?P x a1" by blast from s \y \ B\ obtain a2 where a2: "?P y a2" by blast from g [OF a1] a1 g [OF a2] a2 \?g x = ?g y\ show "x = y" by simp qed moreover have "?g ` B = A" proof safe fix b assume "b \ B" with s obtain a where P: "?P b a" by blast with g[OF P] show "?g b \ A" by auto next fix a assume "a \ A" with s obtain b where P: "?P b a" by blast with s have "b \ B" by blast with g[OF P] have "\b\B. a = ?g b" by blast then show "a \ ?g ` B" by auto qed ultimately show ?thesis by (auto simp: bij_betw_def) qed lemma bij_betw_cong: "(\a. a \ A \ f a = g a) \ bij_betw f A A' = bij_betw g A A'" unfolding bij_betw_def inj_on_def by safe force+ (* somewhat slow *) lemma bij_betw_id[intro, simp]: "bij_betw id A A" unfolding bij_betw_def id_def by auto lemma bij_betw_id_iff: "bij_betw id A B \ A = B" by (auto simp add: bij_betw_def) lemma bij_betw_combine: "bij_betw f A B \ bij_betw f C D \ B \ D = {} \ bij_betw f (A \ C) (B \ D)" unfolding bij_betw_def inj_on_Un image_Un by auto lemma bij_betw_subset: "bij_betw f A A' \ B \ A \ f ` B = B' \ bij_betw f B B'" by (auto simp add: bij_betw_def inj_on_def) lemma bij_betw_ball: "bij_betw f A B \ (\b \ B. phi b) = (\a \ A. phi (f a))" unfolding bij_betw_def inj_on_def by blast lemma bij_pointE: assumes "bij f" obtains x where "y = f x" and "\x'. y = f x' \ x' = x" proof - from assms have "inj f" by (rule bij_is_inj) moreover from assms have "surj f" by (rule bij_is_surj) then have "y \ range f" by simp ultimately have "\!x. y = f x" by (simp add: range_ex1_eq) with that show thesis by blast qed lemma bij_iff: \<^marker>\contributor \Amine Chaieb\\ \bij f \ (\x. \!y. f y = x)\ (is \?P \ ?Q\) proof assume ?P then have \inj f\ \surj f\ by (simp_all add: bij_def) show ?Q proof fix y from \surj f\ obtain x where \y = f x\ by (auto simp add: surj_def) with \inj f\ show \\!x. f x = y\ by (auto simp add: inj_def) qed next assume ?Q then have \inj f\ by (auto simp add: inj_def) moreover have \\x. y = f x\ for y proof - from \?Q\ obtain x where \f x = y\ by blast then have \y = f x\ by simp then show ?thesis .. qed then have \surj f\ by (auto simp add: surj_def) ultimately show ?P by (rule bijI) qed lemma bij_betw_partition: \bij_betw f A B\ if \bij_betw f (A \ C) (B \ D)\ \bij_betw f C D\ \A \ C = {}\ \B \ D = {}\ proof - from that have \inj_on f (A \ C)\ \inj_on f C\ \f ` (A \ C) = B \ D\ \f ` C = D\ by (simp_all add: bij_betw_def) then have \inj_on f A\ and \f ` (A - C) \ f ` (C - A) = {}\ by (simp_all add: inj_on_Un) with \A \ C = {}\ have \f ` A \ f ` C = {}\ by auto with \f ` (A \ C) = B \ D\ \f ` C = D\ \B \ D = {}\ have \f ` A = B\ by blast with \inj_on f A\ show ?thesis by (simp add: bij_betw_def) qed lemma surj_image_vimage_eq: "surj f \ f ` (f -` A) = A" by simp lemma surj_vimage_empty: assumes "surj f" shows "f -` A = {} \ A = {}" using surj_image_vimage_eq [OF \surj f\, of A] by (intro iffI) fastforce+ lemma inj_vimage_image_eq: "inj f \ f -` (f ` A) = A" unfolding inj_def by blast lemma vimage_subsetD: "surj f \ f -` B \ A \ B \ f ` A" by (blast intro: sym) lemma vimage_subsetI: "inj f \ B \ f ` A \ f -` B \ A" unfolding inj_def by blast lemma vimage_subset_eq: "bij f \ f -` B \ A \ B \ f ` A" unfolding bij_def by (blast del: subsetI intro: vimage_subsetI vimage_subsetD) lemma inj_on_image_eq_iff: "inj_on f C \ A \ C \ B \ C \ f ` A = f ` B \ A = B" by (fastforce simp: inj_on_def) lemma inj_on_Un_image_eq_iff: "inj_on f (A \ B) \ f ` A = f ` B \ A = B" by (erule inj_on_image_eq_iff) simp_all lemma inj_on_image_Int: "inj_on f C \ A \ C \ B \ C \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_on_def by blast lemma inj_on_image_set_diff: "inj_on f C \ A - B \ C \ B \ C \ f ` (A - B) = f ` A - f ` B" unfolding inj_on_def by blast lemma image_Int: "inj f \ f ` (A \ B) = f ` A \ f ` B" unfolding inj_def by blast lemma image_set_diff: "inj f \ f ` (A - B) = f ` A - f ` B" unfolding inj_def by blast lemma inj_on_image_mem_iff: "inj_on f B \ a \ B \ A \ B \ f a \ f ` A \ a \ A" by (auto simp: inj_on_def) lemma inj_image_mem_iff: "inj f \ f a \ f ` A \ a \ A" by (blast dest: injD) lemma inj_image_subset_iff: "inj f \ f ` A \ f ` B \ A \ B" by (blast dest: injD) lemma inj_image_eq_iff: "inj f \ f ` A = f ` B \ A = B" by (blast dest: injD) lemma surj_Compl_image_subset: "surj f \ - (f ` A) \ f ` (- A)" by auto lemma inj_image_Compl_subset: "inj f \ f ` (- A) \ - (f ` A)" by (auto simp: inj_def) lemma bij_image_Compl_eq: "bij f \ f ` (- A) = - (f ` A)" by (simp add: bij_def inj_image_Compl_subset surj_Compl_image_subset equalityI) lemma inj_vimage_singleton: "inj f \ f -` {a} \ {THE x. f x = a}" \ \The inverse image of a singleton under an injective function is included in a singleton.\ by (simp add: inj_def) (blast intro: the_equality [symmetric]) lemma inj_on_vimage_singleton: "inj_on f A \ f -` {a} \ A \ {THE x. x \ A \ f x = a}" by (auto simp add: inj_on_def intro: the_equality [symmetric]) lemma bij_betw_byWitness: assumes left: "\a \ A. f' (f a) = a" and right: "\a' \ A'. f (f' a') = a'" and "f ` A \ A'" and img2: "f' ` A' \ A" shows "bij_betw f A A'" using assms unfolding bij_betw_def inj_on_def proof safe fix a b assume "a \ A" "b \ A" with left have "a = f' (f a) \ b = f' (f b)" by simp moreover assume "f a = f b" ultimately show "a = b" by simp next fix a' assume *: "a' \ A'" with img2 have "f' a' \ A" by blast moreover from * right have "a' = f (f' a')" by simp ultimately show "a' \ f ` A" by blast qed corollary notIn_Un_bij_betw: assumes "b \ A" and "f b \ A'" and "bij_betw f A A'" shows "bij_betw f (A \ {b}) (A' \ {f b})" proof - have "bij_betw f {b} {f b}" unfolding bij_betw_def inj_on_def by simp with assms show ?thesis using bij_betw_combine[of f A A' "{b}" "{f b}"] by blast qed lemma notIn_Un_bij_betw3: assumes "b \ A" and "f b \ A'" shows "bij_betw f A A' = bij_betw f (A \ {b}) (A' \ {f b})" proof assume "bij_betw f A A'" then show "bij_betw f (A \ {b}) (A' \ {f b})" using assms notIn_Un_bij_betw [of b A f A'] by blast next assume *: "bij_betw f (A \ {b}) (A' \ {f b})" have "f ` A = A'" proof safe fix a assume **: "a \ A" then have "f a \ A' \ {f b}" using * unfolding bij_betw_def by blast moreover have False if "f a = f b" proof - have "a = b" using * ** that unfolding bij_betw_def inj_on_def by blast with \b \ A\ ** show ?thesis by blast qed ultimately show "f a \ A'" by blast next fix a' assume **: "a' \ A'" then have "a' \ f ` (A \ {b})" using * by (auto simp add: bij_betw_def) then obtain a where 1: "a \ A \ {b} \ f a = a'" by blast moreover have False if "a = b" using 1 ** \f b \ A'\ that by blast ultimately have "a \ A" by blast with 1 show "a' \ f ` A" by blast qed then show "bij_betw f A A'" using * bij_betw_subset[of f "A \ {b}" _ A] by blast qed lemma inj_on_disjoint_Un: assumes "inj_on f A" and "inj_on g B" and "f ` A \ g ` B = {}" shows "inj_on (\x. if x \ A then f x else g x) (A \ B)" using assms by (simp add: inj_on_def disjoint_iff) (blast) lemma bij_betw_disjoint_Un: assumes "bij_betw f A C" and "bij_betw g B D" and "A \ B = {}" and "C \ D = {}" shows "bij_betw (\x. if x \ A then f x else g x) (A \ B) (C \ D)" using assms by (auto simp: inj_on_disjoint_Un bij_betw_def) lemma involuntory_imp_bij: \bij f\ if \\x. f (f x) = x\ proof (rule bijI) from that show \surj f\ by (rule surjI) show \inj f\ proof (rule injI) fix x y assume \f x = f y\ then have \f (f x) = f (f y)\ by simp then show \x = y\ by (simp add: that) qed qed subsubsection \Inj/surj/bij of Algebraic Operations\ context cancel_semigroup_add begin lemma inj_on_add [simp]: "inj_on ((+) a) A" by (rule inj_onI) simp lemma inj_on_add' [simp]: "inj_on (\b. b + a) A" by (rule inj_onI) simp lemma bij_betw_add [simp]: "bij_betw ((+) a) A B \ (+) a ` A = B" by (simp add: bij_betw_def) end context group_add begin lemma diff_left_imp_eq: "a - b = a - c \ b = c" unfolding add_uminus_conv_diff[symmetric] by(drule local.add_left_imp_eq) simp lemma inj_uminus[simp, intro]: "inj_on uminus A" by (auto intro!: inj_onI) lemma surj_uminus[simp]: "surj uminus" using surjI minus_minus by blast lemma surj_plus [simp]: "surj ((+) a)" proof (standard, simp, standard, simp) fix x have "x = a + (-a + x)" by (simp add: add.assoc) thus "x \ range ((+) a)" by blast qed lemma surj_plus_right [simp]: "surj (\b. b+a)" proof (standard, simp, standard, simp) fix b show "b \ range (\b. b+a)" using diff_add_cancel[of b a, symmetric] by blast qed lemma inj_on_diff_left [simp]: \inj_on ((-) a) A\ by (auto intro: inj_onI dest!: diff_left_imp_eq) lemma inj_on_diff_right [simp]: \inj_on (\b. b - a) A\ by (auto intro: inj_onI simp add: algebra_simps) lemma surj_diff [simp]: "surj ((-) a)" proof (standard, simp, standard, simp) fix x have "x = a - (- x + a)" by (simp add: algebra_simps) thus "x \ range ((-) a)" by blast qed lemma surj_diff_right [simp]: "surj (\x. x - a)" proof (standard, simp, standard, simp) fix x have "x = x + a - a" by simp thus "x \ range (\x. x - a)" by fast qed lemma shows bij_plus: "bij ((+) a)" and bij_plus_right: "bij (\x. x + a)" and bij_uminus: "bij uminus" and bij_diff: "bij ((-) a)" and bij_diff_right: "bij (\x. x - a)" by(simp_all add: bij_def) lemma translation_subtract_Compl: "(\x. x - a) ` (- t) = - ((\x. x - a) ` t)" by(rule bij_image_Compl_eq) (auto simp add: bij_def surj_def inj_def diff_eq_eq intro!: add_diff_cancel[symmetric]) lemma translation_diff: "(+) a ` (s - t) = ((+) a ` s) - ((+) a ` t)" by auto lemma translation_subtract_diff: "(\x. x - a) ` (s - t) = ((\x. x - a) ` s) - ((\x. x - a) ` t)" by(rule image_set_diff)(simp add: inj_on_def diff_eq_eq) lemma translation_Int: "(+) a ` (s \ t) = ((+) a ` s) \ ((+) a ` t)" by auto lemma translation_subtract_Int: "(\x. x - a) ` (s \ t) = ((\x. x - a) ` s) \ ((\x. x - a) ` t)" by(rule image_Int)(simp add: inj_on_def diff_eq_eq) end (* TODO: prove in group_add *) context ab_group_add begin lemma translation_Compl: "(+) a ` (- t) = - ((+) a ` t)" proof (rule set_eqI) fix b show "b \ (+) a ` (- t) \ b \ - (+) a ` t" by (auto simp: image_iff algebra_simps intro!: bexI [of _ "b - a"]) qed end subsection \Function Updating\ definition fun_upd :: "('a \ 'b) \ 'a \ 'b \ ('a \ 'b)" where "fun_upd f a b = (\x. if x = a then b else f x)" nonterminal updbinds and updbind syntax "_updbind" :: "'a \ 'a \ updbind" ("(2_ :=/ _)") "" :: "updbind \ updbinds" ("_") "_updbinds":: "updbind \ updbinds \ updbinds" ("_,/ _") "_Update" :: "'a \ updbinds \ 'a" ("_/'((_)')" [1000, 0] 900) translations "_Update f (_updbinds b bs)" \ "_Update (_Update f b) bs" "f(x:=y)" \ "CONST fun_upd f x y" (* Hint: to define the sum of two functions (or maps), use case_sum. A nice infix syntax could be defined by notation case_sum (infixr "'(+')"80) *) lemma fun_upd_idem_iff: "f(x:=y) = f \ f x = y" unfolding fun_upd_def apply safe apply (erule subst) apply auto done lemma fun_upd_idem: "f x = y \ f(x := y) = f" by (simp only: fun_upd_idem_iff) lemma fun_upd_triv [iff]: "f(x := f x) = f" by (simp only: fun_upd_idem) lemma fun_upd_apply [simp]: "(f(x := y)) z = (if z = x then y else f z)" by (simp add: fun_upd_def) (* fun_upd_apply supersedes these two, but they are useful if fun_upd_apply is intentionally removed from the simpset *) lemma fun_upd_same: "(f(x := y)) x = y" by simp lemma fun_upd_other: "z \ x \ (f(x := y)) z = f z" by simp lemma fun_upd_upd [simp]: "f(x := y, x := z) = f(x := z)" by (simp add: fun_eq_iff) lemma fun_upd_twist: "a \ c \ (m(a := b))(c := d) = (m(c := d))(a := b)" by auto lemma inj_on_fun_updI: "inj_on f A \ y \ f ` A \ inj_on (f(x := y)) A" by (auto simp: inj_on_def) lemma fun_upd_image: "f(x := y) ` A = (if x \ A then insert y (f ` (A - {x})) else f ` A)" by auto lemma fun_upd_comp: "f \ (g(x := y)) = (f \ g)(x := f y)" by auto lemma fun_upd_eqD: "f(x := y) = g(x := z) \ y = z" by (simp add: fun_eq_iff split: if_split_asm) subsection \\override_on\\ definition override_on :: "('a \ 'b) \ ('a \ 'b) \ 'a set \ 'a \ 'b" where "override_on f g A = (\a. if a \ A then g a else f a)" lemma override_on_emptyset[simp]: "override_on f g {} = f" by (simp add: override_on_def) lemma override_on_apply_notin[simp]: "a \ A \ (override_on f g A) a = f a" by (simp add: override_on_def) lemma override_on_apply_in[simp]: "a \ A \ (override_on f g A) a = g a" by (simp add: override_on_def) lemma override_on_insert: "override_on f g (insert x X) = (override_on f g X)(x:=g x)" by (simp add: override_on_def fun_eq_iff) lemma override_on_insert': "override_on f g (insert x X) = (override_on (f(x:=g x)) g X)" by (simp add: override_on_def fun_eq_iff) subsection \Inversion of injective functions\ definition the_inv_into :: "'a set \ ('a \ 'b) \ ('b \ 'a)" where "the_inv_into A f = (\x. THE y. y \ A \ f y = x)" lemma the_inv_into_f_f: "inj_on f A \ x \ A \ the_inv_into A f (f x) = x" unfolding the_inv_into_def inj_on_def by blast lemma f_the_inv_into_f: "inj_on f A \ y \ f ` A \ f (the_inv_into A f y) = y" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma f_the_inv_into_f_bij_betw: "bij_betw f A B \ (bij_betw f A B \ x \ B) \ f (the_inv_into A f x) = x" unfolding bij_betw_def by (blast intro: f_the_inv_into_f) lemma the_inv_into_into: "inj_on f A \ x \ f ` A \ A \ B \ the_inv_into A f x \ B" unfolding the_inv_into_def by (rule the1I2; blast dest: inj_onD) lemma the_inv_into_onto [simp]: "inj_on f A \ the_inv_into A f ` (f ` A) = A" by (fast intro: the_inv_into_into the_inv_into_f_f [symmetric]) lemma the_inv_into_f_eq: "inj_on f A \ f x = y \ x \ A \ the_inv_into A f y = x" by (force simp add: the_inv_into_f_f) lemma the_inv_into_comp: "inj_on f (g ` A) \ inj_on g A \ x \ f ` g ` A \ the_inv_into A (f \ g) x = (the_inv_into A g \ the_inv_into (g ` A) f) x" apply (rule the_inv_into_f_eq) apply (fast intro: comp_inj_on) apply (simp add: f_the_inv_into_f the_inv_into_into) apply (simp add: the_inv_into_into) done lemma inj_on_the_inv_into: "inj_on f A \ inj_on (the_inv_into A f) (f ` A)" by (auto intro: inj_onI simp: the_inv_into_f_f) lemma bij_betw_the_inv_into: "bij_betw f A B \ bij_betw (the_inv_into A f) B A" by (auto simp add: bij_betw_def inj_on_the_inv_into the_inv_into_into) lemma bij_betw_iff_bijections: "bij_betw f A B \ (\g. (\x \ A. f x \ B \ g(f x) = x) \ (\y \ B. g y \ A \ f(g y) = y))" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" by (auto simp: bij_betw_def f_the_inv_into_f the_inv_into_f_f the_inv_into_into exI[where ?x="the_inv_into A f"]) next show "?rhs \ ?lhs" by (force intro: bij_betw_byWitness) qed abbreviation the_inv :: "('a \ 'b) \ ('b \ 'a)" where "the_inv f \ the_inv_into UNIV f" lemma the_inv_f_f: "the_inv f (f x) = x" if "inj f" using that UNIV_I by (rule the_inv_into_f_f) subsection \Monotonicity\ definition monotone_on :: "'a set \ ('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone_on A orda ordb f \ (\x\A. \y\A. orda x y \ ordb (f x) (f y))" abbreviation monotone :: "('a \ 'a \ bool) \ ('b \ 'b \ bool) \ ('a \ 'b) \ bool" where "monotone \ monotone_on UNIV" lemma monotone_def[no_atp]: "monotone orda ordb f \ (\x y. orda x y \ ordb (f x) (f y))" by (simp add: monotone_on_def) text \Lemma @{thm [source] monotone_def} is provided for backward compatibility.\ lemma monotone_onI: "(\x y. x \ A \ y \ A \ orda x y \ ordb (f x) (f y)) \ monotone_on A orda ordb f" by (simp add: monotone_on_def) lemma monotoneI[intro?]: "(\x y. orda x y \ ordb (f x) (f y)) \ monotone orda ordb f" by (rule monotone_onI) lemma monotone_onD: "monotone_on A orda ordb f \ x \ A \ y \ A \ orda x y \ ordb (f x) (f y)" by (simp add: monotone_on_def) lemma monotoneD[dest?]: "monotone orda ordb f \ orda x y \ ordb (f x) (f y)" by (rule monotone_onD[of UNIV, simplified]) lemma monotone_on_subset: "monotone_on A orda ordb f \ B \ A \ monotone_on B orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_empty[simp]: "monotone_on {} orda ordb f" by (auto intro: monotone_onI dest: monotone_onD) lemma monotone_on_o: assumes mono_f: "monotone_on A orda ordb f" and mono_g: "monotone_on B ordc orda g" and "g ` B \ A" shows "monotone_on B ordc ordb (f \ g)" proof (rule monotone_onI) fix x y assume "x \ B" and "y \ B" and "ordc x y" hence "orda (g x) (g y)" by (rule mono_g[THEN monotone_onD]) moreover from \g ` B \ A\ \x \ B\ \y \ B\ have "g x \ A" and "g y \ A" unfolding image_subset_iff by simp_all ultimately show "ordb ((f \ g) x) ((f \ g) y)" using mono_f[THEN monotone_onD] by simp qed subsubsection \Specializations For @{class ord} Type Class And More\ context ord begin abbreviation mono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "mono_on A \ monotone_on A (\) (\)" abbreviation strict_mono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "strict_mono_on A \ monotone_on A (<) (<)" abbreviation antimono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "antimono_on A \ monotone_on A (\) (\)" abbreviation strict_antimono_on :: "'a set \ ('a \ 'b :: ord) \ bool" where "strict_antimono_on A \ monotone_on A (<) (>)" lemma mono_on_def[no_atp]: "mono_on A f \ (\r s. r \ A \ s \ A \ r \ s \ f r \ f s)" by (auto simp add: monotone_on_def) lemma strict_mono_on_def[no_atp]: "strict_mono_on A f \ (\r s. r \ A \ s \ A \ r < s \ f r < f s)" by (auto simp add: monotone_on_def) text \Lemmas @{thm [source] mono_on_def} and @{thm [source] strict_mono_on_def} are provided for backward compatibility.\ lemma mono_onI: "(\r s. r \ A \ s \ A \ r \ s \ f r \ f s) \ mono_on A f" by (rule monotone_onI) lemma strict_mono_onI: "(\r s. r \ A \ s \ A \ r < s \ f r < f s) \ strict_mono_on A f" by (rule monotone_onI) lemma mono_onD: "\mono_on A f; r \ A; s \ A; r \ s\ \ f r \ f s" by (rule monotone_onD) lemma strict_mono_onD: "\strict_mono_on A f; r \ A; s \ A; r < s\ \ f r < f s" by (rule monotone_onD) lemma mono_on_subset: "mono_on A f \ B \ A \ mono_on B f" by (rule monotone_on_subset) end lemma mono_on_greaterD: assumes "mono_on A g" "x \ A" "y \ A" "g x > (g (y::_::linorder) :: _ :: linorder)" shows "x > y" proof (rule ccontr) assume "\x > y" hence "x \ y" by (simp add: not_less) from assms(1-3) and this have "g x \ g y" by (rule mono_onD) with assms(4) show False by simp qed context order begin abbreviation mono :: "('a \ 'b::order) \ bool" where "mono \ mono_on UNIV" abbreviation strict_mono :: "('a \ 'b::order) \ bool" where "strict_mono \ strict_mono_on UNIV" abbreviation antimono :: "('a \ 'b::order) \ bool" where "antimono \ monotone (\) (\x y. y \ x)" lemma mono_def[no_atp]: "mono f \ (\x y. x \ y \ f x \ f y)" by (simp add: monotone_on_def) lemma strict_mono_def[no_atp]: "strict_mono f \ (\x y. x < y \ f x < f y)" by (simp add: monotone_on_def) lemma antimono_def[no_atp]: "antimono f \ (\x y. x \ y \ f x \ f y)" by (simp add: monotone_on_def) text \Lemmas @{thm [source] mono_def}, @{thm [source] strict_mono_def}, and @{thm [source] antimono_def} are provided for backward compatibility.\ lemma monoI [intro?]: "(\x y. x \ y \ f x \ f y) \ mono f" by (rule monotoneI) lemma strict_monoI [intro?]: "(\x y. x < y \ f x < f y) \ strict_mono f" by (rule monotoneI) lemma antimonoI [intro?]: "(\x y. x \ y \ f x \ f y) \ antimono f" by (rule monotoneI) lemma monoD [dest?]: "mono f \ x \ y \ f x \ f y" by (rule monotoneD) lemma strict_monoD [dest?]: "strict_mono f \ x < y \ f x < f y" by (rule monotoneD) lemma antimonoD [dest?]: "antimono f \ x \ y \ f x \ f y" by (rule monotoneD) lemma monoE: assumes "mono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: mono_def) qed lemma antimonoE: fixes f :: "'a \ 'b::order" assumes "antimono f" assumes "x \ y" obtains "f x \ f y" proof from assms show "f x \ f y" by (simp add: antimono_def) qed lemma mono_imp_mono_on: "mono f \ mono_on A f" by (rule monotone_on_subset[OF _ subset_UNIV]) lemma strict_mono_mono [dest?]: assumes "strict_mono f" shows "mono f" proof (rule monoI) fix x y assume "x \ y" show "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with \x \ y\ have "x < y" by simp with assms strict_monoD have "f x < f y" by auto then show ?thesis by simp qed qed end context linorder begin lemma mono_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x \ y" proof show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma mono_strict_invE: fixes f :: "'a \ 'b::order" assumes "mono f" assumes "f x < f y" obtains "x < y" proof show "x < y" proof (rule ccontr) assume "\ x < y" then have "y \ x" by simp with \mono f\ obtain "f y \ f x" by (rule monoE) with \f x < f y\ show False by simp qed qed lemma strict_mono_eq: assumes "strict_mono f" shows "f x = f y \ x = y" proof assume "f x = f y" show "x = y" proof (cases x y rule: linorder_cases) case less with assms strict_monoD have "f x < f y" by auto with \f x = f y\ show ?thesis by simp next case equal then show ?thesis . next case greater with assms strict_monoD have "f y < f x" by auto with \f x = f y\ show ?thesis by simp qed qed simp lemma strict_mono_less_eq: assumes "strict_mono f" shows "f x \ f y \ x \ y" proof assume "x \ y" with assms strict_mono_mono monoD show "f x \ f y" by auto next assume "f x \ f y" show "x \ y" proof (rule ccontr) assume "\ x \ y" then have "y < x" by simp with assms strict_monoD have "f y < f x" by auto with \f x \ f y\ show False by simp qed qed lemma strict_mono_less: assumes "strict_mono f" shows "f x < f y \ x < y" using assms by (auto simp add: less_le Orderings.less_le strict_mono_eq strict_mono_less_eq) end lemma strict_mono_inv: fixes f :: "('a::linorder) \ ('b::linorder)" assumes "strict_mono f" and "surj f" and inv: "\x. g (f x) = x" shows "strict_mono g" proof fix x y :: 'b assume "x < y" from \surj f\ obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast with \x < y\ and \strict_mono f\ have "x' < y'" by (simp add: strict_mono_less) with inv show "g x < g y" by simp qed lemma strict_mono_on_imp_inj_on: assumes "strict_mono_on A (f :: (_ :: linorder) \ (_ :: preorder))" shows "inj_on f A" proof (rule inj_onI) fix x y assume "x \ A" "y \ A" "f x = f y" thus "x = y" by (cases x y rule: linorder_cases) (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x]) qed lemma strict_mono_on_leD: assumes "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder)" "x \ A" "y \ A" "x \ y" shows "f x \ f y" proof (cases "x = y") case True then show ?thesis by simp next case False with assms have "f x < f y" using strict_mono_onD[OF assms(1)] by simp then show ?thesis by (rule less_imp_le) qed lemma strict_mono_on_eqD: fixes f :: "(_ :: linorder) \ (_ :: preorder)" assumes "strict_mono_on A f" "f x = f y" "x \ A" "y \ A" shows "y = x" using assms by (cases rule: linorder_cases) (auto dest: strict_mono_onD) lemma strict_mono_on_imp_mono_on: "strict_mono_on A (f :: (_ :: linorder) \ _ :: preorder) \ mono_on A f" by (rule mono_onI, rule strict_mono_on_leD) lemma mono_imp_strict_mono: fixes f :: "'a::order \ 'b::order" shows "\mono_on S f; inj_on f S\ \ strict_mono_on S f" by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) lemma strict_mono_iff_mono: fixes f :: "'a::linorder \ 'b::order" shows "strict_mono_on S f \ mono_on S f \ inj_on f S" proof show "strict_mono_on S f \ mono_on S f \ inj_on f S" by (simp add: strict_mono_on_imp_inj_on strict_mono_on_imp_mono_on) qed (auto intro: mono_imp_strict_mono) lemma antimono_imp_strict_antimono: fixes f :: "'a::order \ 'b::order" shows "\antimono_on S f; inj_on f S\ \ strict_antimono_on S f" by (auto simp add: monotone_on_def order_less_le inj_on_eq_iff) lemma strict_antimono_iff_antimono: fixes f :: "'a::linorder \ 'b::order" shows "strict_antimono_on S f \ antimono_on S f \ inj_on f S" proof show "strict_antimono_on S f \ antimono_on S f \ inj_on f S" by (force simp add: monotone_on_def intro: linorder_inj_onI) qed (auto intro: antimono_imp_strict_antimono) lemma mono_compose: "mono Q \ mono (\i x. Q i (f x))" unfolding mono_def le_fun_def by auto lemma mono_add: fixes a :: "'a::ordered_ab_semigroup_add" shows "mono ((+) a)" by (simp add: add_left_mono monoI) lemma (in semilattice_inf) mono_inf: "mono f \ f (A \ B) \ f A \ f B" for f :: "'a \ 'b::semilattice_inf" by (auto simp add: mono_def intro: Lattices.inf_greatest) lemma (in semilattice_sup) mono_sup: "mono f \ f A \ f B \ f (A \ B)" for f :: "'a \ 'b::semilattice_sup" by (auto simp add: mono_def intro: Lattices.sup_least) lemma (in linorder) min_of_mono: "mono f \ min (f m) (f n) = f (min m n)" by (auto simp: mono_def Orderings.min_def min_def intro: Orderings.antisym) lemma (in linorder) max_of_mono: "mono f \ max (f m) (f n) = f (max m n)" by (auto simp: mono_def Orderings.max_def max_def intro: Orderings.antisym) lemma (in linorder) max_of_antimono: "antimono f \ max (f x) (f y) = f (min x y)" and min_of_antimono: "antimono f \ min (f x) (f y) = f (max x y)" by (auto simp: antimono_def Orderings.max_def max_def Orderings.min_def min_def intro!: antisym) lemma (in linorder) strict_mono_imp_inj_on: "strict_mono f \ inj_on f A" by (auto intro!: inj_onI dest: strict_mono_eq) lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" by (fact mono_inf) lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" by (fact mono_sup) subsubsection \Least value operator\ lemma Least_mono: "mono f \ \x\S. \y\S. x \ y \ (LEAST y. y \ f ` S) = f (LEAST x. x \ S)" for f :: "'a::order \ 'b::order" \ \Courtesy of Stephan Merz\ apply clarify apply (erule_tac P = "\x. x \ S" in LeastI2_order) apply fast apply (rule LeastI2_order) apply (auto elim: monoD intro!: order_antisym) done subsection \Setup\ subsubsection \Proof tools\ text \Simplify terms of the form \f(\,x:=y,\,x:=z,\)\ to \f(\,x:=z,\)\\ -simproc_setup fun_upd2 ("f(v := w, x := y)") = \fn _ => +simproc_setup fun_upd2 ("f(v := w, x := y)") = \ let fun gen_fun_upd NONE T _ _ = NONE | gen_fun_upd (SOME f) T x y = SOME (Const (\<^const_name>\fun_upd\, T) $ f $ x $ y) fun dest_fun_T1 (Type (_, T :: Ts)) = T fun find_double (t as Const (\<^const_name>\fun_upd\,T) $ f $ x $ y) = let fun find (Const (\<^const_name>\fun_upd\,T) $ g $ v $ w) = if v aconv x then SOME g else gen_fun_upd (find g) T v w | find t = NONE in (dest_fun_T1 T, gen_fun_upd (find f) T x y) end val ss = simpset_of \<^context> fun proc ctxt ct = let val t = Thm.term_of ct in (case find_double t of (T, NONE) => NONE | (T, SOME rhs) => SOME (Goal.prove ctxt [] [] (Logic.mk_equals (t, rhs)) (fn _ => resolve_tac ctxt [eq_reflection] 1 THEN resolve_tac ctxt @{thms ext} 1 THEN simp_tac (put_simpset ss ctxt) 1))) end - in proc end + in K proc end \ subsubsection \Functorial structure of types\ ML_file \Tools/functor.ML\ functor map_fun: map_fun by (simp_all add: fun_eq_iff) functor vimage by (simp_all add: fun_eq_iff vimage_comp) text \Legacy theorem names\ lemmas o_def = comp_def lemmas o_apply = comp_apply lemmas o_assoc = comp_assoc [symmetric] lemmas id_o = id_comp lemmas o_id = comp_id lemmas o_eq_dest = comp_eq_dest lemmas o_eq_elim = comp_eq_elim lemmas o_eq_dest_lhs = comp_eq_dest_lhs lemmas o_eq_id_dest = comp_eq_id_dest end diff --git a/src/HOL/Groups.thy b/src/HOL/Groups.thy --- a/src/HOL/Groups.thy +++ b/src/HOL/Groups.thy @@ -1,1495 +1,1495 @@ (* Title: HOL/Groups.thy Author: Gertrud Bauer Author: Steven Obua Author: Lawrence C Paulson Author: Markus Wenzel Author: Jeremy Avigad *) section \Groups, also combined with orderings\ theory Groups imports Orderings begin subsection \Dynamic facts\ named_theorems ac_simps "associativity and commutativity simplification rules" and algebra_simps "algebra simplification rules for rings" and algebra_split_simps "algebra simplification rules for rings, with potential goal splitting" and field_simps "algebra simplification rules for fields" and field_split_simps "algebra simplification rules for fields, with potential goal splitting" text \ The rewrites accumulated in \algebra_simps\ deal with the classical algebraic structures of groups, rings and family. They simplify terms by multiplying everything out (in case of a ring) and bringing sums and products into a canonical form (by ordered rewriting). As a result it decides group and ring equalities but also helps with inequalities. Of course it also works for fields, but it knows nothing about multiplicative inverses or division. This is catered for by \field_simps\. Facts in \field_simps\ multiply with denominators in (in)equations if they can be proved to be non-zero (for equations) or positive/negative (for inequalities). Can be too aggressive and is therefore separate from the more benign \algebra_simps\. Collections \algebra_split_simps\ and \field_split_simps\ correspond to \algebra_simps\ and \field_simps\ but contain more aggresive rules that may lead to goal splitting. \ subsection \Abstract structures\ text \ These locales provide basic structures for interpretation into bigger structures; extensions require careful thinking, otherwise undesired effects may occur due to interpretation. \ locale semigroup = fixes f :: "'a \ 'a \ 'a" (infixl "\<^bold>*" 70) assumes assoc [ac_simps]: "a \<^bold>* b \<^bold>* c = a \<^bold>* (b \<^bold>* c)" locale abel_semigroup = semigroup + assumes commute [ac_simps]: "a \<^bold>* b = b \<^bold>* a" begin lemma left_commute [ac_simps]: "b \<^bold>* (a \<^bold>* c) = a \<^bold>* (b \<^bold>* c)" proof - have "(b \<^bold>* a) \<^bold>* c = (a \<^bold>* b) \<^bold>* c" by (simp only: commute) then show ?thesis by (simp only: assoc) qed end locale monoid = semigroup + fixes z :: 'a ("\<^bold>1") assumes left_neutral [simp]: "\<^bold>1 \<^bold>* a = a" assumes right_neutral [simp]: "a \<^bold>* \<^bold>1 = a" locale comm_monoid = abel_semigroup + fixes z :: 'a ("\<^bold>1") assumes comm_neutral: "a \<^bold>* \<^bold>1 = a" begin sublocale monoid by standard (simp_all add: commute comm_neutral) end locale group = semigroup + fixes z :: 'a ("\<^bold>1") fixes inverse :: "'a \ 'a" assumes group_left_neutral: "\<^bold>1 \<^bold>* a = a" assumes left_inverse [simp]: "inverse a \<^bold>* a = \<^bold>1" begin lemma left_cancel: "a \<^bold>* b = a \<^bold>* c \ b = c" proof assume "a \<^bold>* b = a \<^bold>* c" then have "inverse a \<^bold>* (a \<^bold>* b) = inverse a \<^bold>* (a \<^bold>* c)" by simp then have "(inverse a \<^bold>* a) \<^bold>* b = (inverse a \<^bold>* a) \<^bold>* c" by (simp only: assoc) then show "b = c" by (simp add: group_left_neutral) qed simp sublocale monoid proof fix a have "inverse a \<^bold>* a = \<^bold>1" by simp then have "inverse a \<^bold>* (a \<^bold>* \<^bold>1) = inverse a \<^bold>* a" by (simp add: group_left_neutral assoc [symmetric]) with left_cancel show "a \<^bold>* \<^bold>1 = a" by (simp only: left_cancel) qed (fact group_left_neutral) lemma inverse_unique: assumes "a \<^bold>* b = \<^bold>1" shows "inverse a = b" proof - from assms have "inverse a \<^bold>* (a \<^bold>* b) = inverse a" by simp then show ?thesis by (simp add: assoc [symmetric]) qed lemma inverse_neutral [simp]: "inverse \<^bold>1 = \<^bold>1" by (rule inverse_unique) simp lemma inverse_inverse [simp]: "inverse (inverse a) = a" by (rule inverse_unique) simp lemma right_inverse [simp]: "a \<^bold>* inverse a = \<^bold>1" proof - have "a \<^bold>* inverse a = inverse (inverse a) \<^bold>* inverse a" by simp also have "\ = \<^bold>1" by (rule left_inverse) then show ?thesis by simp qed lemma inverse_distrib_swap: "inverse (a \<^bold>* b) = inverse b \<^bold>* inverse a" proof (rule inverse_unique) have "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = a \<^bold>* (b \<^bold>* inverse b) \<^bold>* inverse a" by (simp only: assoc) also have "\ = \<^bold>1" by simp finally show "a \<^bold>* b \<^bold>* (inverse b \<^bold>* inverse a) = \<^bold>1" . qed lemma right_cancel: "b \<^bold>* a = c \<^bold>* a \ b = c" proof assume "b \<^bold>* a = c \<^bold>* a" then have "b \<^bold>* a \<^bold>* inverse a= c \<^bold>* a \<^bold>* inverse a" by simp then show "b = c" by (simp add: assoc) qed simp end subsection \Generic operations\ class zero = fixes zero :: 'a ("0") class one = fixes one :: 'a ("1") hide_const (open) zero one lemma Let_0 [simp]: "Let 0 f = f 0" unfolding Let_def .. lemma Let_1 [simp]: "Let 1 f = f 1" unfolding Let_def .. setup \ Reorient_Proc.add (fn Const(\<^const_name>\Groups.zero\, _) => true | Const(\<^const_name>\Groups.one\, _) => true | _ => false) \ -simproc_setup reorient_zero ("0 = x") = Reorient_Proc.proc -simproc_setup reorient_one ("1 = x") = Reorient_Proc.proc +simproc_setup reorient_zero ("0 = x") = \K Reorient_Proc.proc\ +simproc_setup reorient_one ("1 = x") = \K Reorient_Proc.proc\ typed_print_translation \ let fun tr' c = (c, fn ctxt => fn T => fn ts => if null ts andalso Printer.type_emphasis ctxt T then Syntax.const \<^syntax_const>\_constrain\ $ Syntax.const c $ Syntax_Phases.term_of_typ ctxt T else raise Match); in map tr' [\<^const_syntax>\Groups.one\, \<^const_syntax>\Groups.zero\] end \ \ \show types that are presumably too general\ class plus = fixes plus :: "'a \ 'a \ 'a" (infixl "+" 65) class minus = fixes minus :: "'a \ 'a \ 'a" (infixl "-" 65) class uminus = fixes uminus :: "'a \ 'a" ("- _" [81] 80) class times = fixes times :: "'a \ 'a \ 'a" (infixl "*" 70) subsection \Semigroups and Monoids\ class semigroup_add = plus + assumes add_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a + b) + c = a + (b + c)" begin sublocale add: semigroup plus by standard (fact add_assoc) end hide_fact add_assoc class ab_semigroup_add = semigroup_add + assumes add_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a + b = b + a" begin sublocale add: abel_semigroup plus by standard (fact add_commute) declare add.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] lemmas add_ac = add.assoc add.commute add.left_commute end hide_fact add_commute lemmas add_ac = add.assoc add.commute add.left_commute class semigroup_mult = times + assumes mult_assoc [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a * b) * c = a * (b * c)" begin sublocale mult: semigroup times by standard (fact mult_assoc) end hide_fact mult_assoc class ab_semigroup_mult = semigroup_mult + assumes mult_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a * b = b * a" begin sublocale mult: abel_semigroup times by standard (fact mult_commute) declare mult.left_commute [algebra_simps, algebra_split_simps, field_simps, field_split_simps] lemmas mult_ac = mult.assoc mult.commute mult.left_commute end hide_fact mult_commute lemmas mult_ac = mult.assoc mult.commute mult.left_commute class monoid_add = zero + semigroup_add + assumes add_0_left: "0 + a = a" and add_0_right: "a + 0 = a" begin sublocale add: monoid plus 0 by standard (fact add_0_left add_0_right)+ end lemma zero_reorient: "0 = x \ x = 0" by (fact eq_commute) class comm_monoid_add = zero + ab_semigroup_add + assumes add_0: "0 + a = a" begin subclass monoid_add by standard (simp_all add: add_0 add.commute [of _ 0]) sublocale add: comm_monoid plus 0 by standard (simp add: ac_simps) end class monoid_mult = one + semigroup_mult + assumes mult_1_left: "1 * a = a" and mult_1_right: "a * 1 = a" begin sublocale mult: monoid times 1 by standard (fact mult_1_left mult_1_right)+ end lemma one_reorient: "1 = x \ x = 1" by (fact eq_commute) class comm_monoid_mult = one + ab_semigroup_mult + assumes mult_1: "1 * a = a" begin subclass monoid_mult by standard (simp_all add: mult_1 mult.commute [of _ 1]) sublocale mult: comm_monoid times 1 by standard (simp add: ac_simps) end class cancel_semigroup_add = semigroup_add + assumes add_left_imp_eq: "a + b = a + c \ b = c" assumes add_right_imp_eq: "b + a = c + a \ b = c" begin lemma add_left_cancel [simp]: "a + b = a + c \ b = c" by (blast dest: add_left_imp_eq) lemma add_right_cancel [simp]: "b + a = c + a \ b = c" by (blast dest: add_right_imp_eq) end class cancel_ab_semigroup_add = ab_semigroup_add + minus + assumes add_diff_cancel_left' [simp]: "(a + b) - a = b" assumes diff_diff_add [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b - c = a - (b + c)" begin lemma add_diff_cancel_right' [simp]: "(a + b) - b = a" using add_diff_cancel_left' [of b a] by (simp add: ac_simps) subclass cancel_semigroup_add proof fix a b c :: 'a assume "a + b = a + c" then have "a + b - a = a + c - a" by simp then show "b = c" by simp next fix a b c :: 'a assume "b + a = c + a" then have "b + a - a = c + a - a" by simp then show "b = c" by simp qed lemma add_diff_cancel_left [simp]: "(c + a) - (c + b) = a - b" unfolding diff_diff_add [symmetric] by simp lemma add_diff_cancel_right [simp]: "(a + c) - (b + c) = a - b" using add_diff_cancel_left [symmetric] by (simp add: ac_simps) lemma diff_right_commute: "a - c - b = a - b - c" by (simp add: diff_diff_add add.commute) end class cancel_comm_monoid_add = cancel_ab_semigroup_add + comm_monoid_add begin lemma diff_zero [simp]: "a - 0 = a" using add_diff_cancel_right' [of a 0] by simp lemma diff_cancel [simp]: "a - a = 0" proof - have "(a + 0) - (a + 0) = 0" by (simp only: add_diff_cancel_left diff_zero) then show ?thesis by simp qed lemma add_implies_diff: assumes "c + b = a" shows "c = a - b" proof - from assms have "(b + c) - (b + 0) = a - b" by (simp add: add.commute) then show "c = a - b" by simp qed lemma add_cancel_right_right [simp]: "a = a + b \ b = 0" (is "?P \ ?Q") proof assume ?Q then show ?P by simp next assume ?P then have "a - a = a + b - a" by simp then show ?Q by simp qed lemma add_cancel_right_left [simp]: "a = b + a \ b = 0" using add_cancel_right_right [of a b] by (simp add: ac_simps) lemma add_cancel_left_right [simp]: "a + b = a \ b = 0" by (auto dest: sym) lemma add_cancel_left_left [simp]: "b + a = a \ b = 0" by (auto dest: sym) end class comm_monoid_diff = cancel_comm_monoid_add + assumes zero_diff [simp]: "0 - a = 0" begin lemma diff_add_zero [simp]: "a - (a + b) = 0" proof - have "a - (a + b) = (a + 0) - (a + b)" by simp also have "\ = 0" by (simp only: add_diff_cancel_left zero_diff) finally show ?thesis . qed end subsection \Groups\ class group_add = minus + uminus + monoid_add + assumes left_minus: "- a + a = 0" assumes add_uminus_conv_diff [simp]: "a + (- b) = a - b" begin lemma diff_conv_add_uminus: "a - b = a + (- b)" by simp sublocale add: group plus 0 uminus by standard (simp_all add: left_minus) lemma minus_unique: "a + b = 0 \ - a = b" by (fact add.inverse_unique) lemma minus_zero: "- 0 = 0" by (fact add.inverse_neutral) lemma minus_minus: "- (- a) = a" by (fact add.inverse_inverse) lemma right_minus: "a + - a = 0" by (fact add.right_inverse) lemma diff_self [simp]: "a - a = 0" using right_minus [of a] by simp subclass cancel_semigroup_add by standard (simp_all add: add.left_cancel add.right_cancel) lemma minus_add_cancel [simp]: "- a + (a + b) = b" by (simp add: add.assoc [symmetric]) lemma add_minus_cancel [simp]: "a + (- a + b) = b" by (simp add: add.assoc [symmetric]) lemma diff_add_cancel [simp]: "a - b + b = a" by (simp only: diff_conv_add_uminus add.assoc) simp lemma add_diff_cancel [simp]: "a + b - b = a" by (simp only: diff_conv_add_uminus add.assoc) simp lemma minus_add: "- (a + b) = - b + - a" by (fact add.inverse_distrib_swap) lemma right_minus_eq [simp]: "a - b = 0 \ a = b" proof assume "a - b = 0" have "a = (a - b) + b" by (simp add: add.assoc) also have "\ = b" using \a - b = 0\ by simp finally show "a = b" . next assume "a = b" then show "a - b = 0" by simp qed lemma eq_iff_diff_eq_0: "a = b \ a - b = 0" by (fact right_minus_eq [symmetric]) lemma diff_0 [simp]: "0 - a = - a" by (simp only: diff_conv_add_uminus add_0_left) lemma diff_0_right [simp]: "a - 0 = a" by (simp only: diff_conv_add_uminus minus_zero add_0_right) lemma diff_minus_eq_add [simp]: "a - - b = a + b" by (simp only: diff_conv_add_uminus minus_minus) lemma neg_equal_iff_equal [simp]: "- a = - b \ a = b" proof assume "- a = - b" then have "- (- a) = - (- b)" by simp then show "a = b" by simp next assume "a = b" then show "- a = - b" by simp qed lemma neg_equal_0_iff_equal [simp]: "- a = 0 \ a = 0" by (subst neg_equal_iff_equal [symmetric]) simp lemma neg_0_equal_iff_equal [simp]: "0 = - a \ 0 = a" by (subst neg_equal_iff_equal [symmetric]) simp text \The next two equations can make the simplifier loop!\ lemma equation_minus_iff: "a = - b \ b = - a" proof - have "- (- a) = - b \ - a = b" by (rule neg_equal_iff_equal) then show ?thesis by (simp add: eq_commute) qed lemma minus_equation_iff: "- a = b \ - b = a" proof - have "- a = - (- b) \ a = -b" by (rule neg_equal_iff_equal) then show ?thesis by (simp add: eq_commute) qed lemma eq_neg_iff_add_eq_0: "a = - b \ a + b = 0" proof assume "a = - b" then show "a + b = 0" by simp next assume "a + b = 0" moreover have "a + (b + - b) = (a + b) + - b" by (simp only: add.assoc) ultimately show "a = - b" by simp qed lemma add_eq_0_iff2: "a + b = 0 \ a = - b" by (fact eq_neg_iff_add_eq_0 [symmetric]) lemma neg_eq_iff_add_eq_0: "- a = b \ a + b = 0" by (auto simp add: add_eq_0_iff2) lemma add_eq_0_iff: "a + b = 0 \ b = - a" by (auto simp add: neg_eq_iff_add_eq_0 [symmetric]) lemma minus_diff_eq [simp]: "- (a - b) = b - a" by (simp only: neg_eq_iff_add_eq_0 diff_conv_add_uminus add.assoc minus_add_cancel) simp lemma add_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a + (b - c) = (a + b) - c" by (simp only: diff_conv_add_uminus add.assoc) lemma diff_add_eq_diff_diff_swap: "a - (b + c) = a - c - b" by (simp only: diff_conv_add_uminus add.assoc minus_add) lemma diff_eq_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b = c \ a = c + b" by auto lemma eq_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a = c - b \ a + b = c" by auto lemma diff_diff_eq2 [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - (b - c) = (a + c) - b" by (simp only: diff_conv_add_uminus add.assoc) simp lemma diff_eq_diff_eq: "a - b = c - d \ a = b \ c = d" by (simp only: eq_iff_diff_eq_0 [of a b] eq_iff_diff_eq_0 [of c d]) end class ab_group_add = minus + uminus + comm_monoid_add + assumes ab_left_minus: "- a + a = 0" assumes ab_diff_conv_add_uminus: "a - b = a + (- b)" begin subclass group_add by standard (simp_all add: ab_left_minus ab_diff_conv_add_uminus) subclass cancel_comm_monoid_add proof fix a b c :: 'a have "b + a - a = b" by simp then show "a + b - a = b" by (simp add: ac_simps) show "a - b - c = a - (b + c)" by (simp add: algebra_simps) qed lemma uminus_add_conv_diff [simp]: "- a + b = b - a" by (simp add: add.commute) lemma minus_add_distrib [simp]: "- (a + b) = - a + - b" by (simp add: algebra_simps) lemma diff_add_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "(a - b) + c = (a + c) - b" by (simp add: algebra_simps) lemma minus_diff_commute: "- b - a = - a - b" by (simp only: diff_conv_add_uminus add.commute) end subsection \(Partially) Ordered Groups\ text \ The theory of partially ordered groups is taken from the books: \<^item> \<^emph>\Lattice Theory\ by Garret Birkhoff, American Mathematical Society, 1979 \<^item> \<^emph>\Partially Ordered Algebraic Systems\, Pergamon Press, 1963 Most of the used notions can also be looked up in \<^item> \<^url>\http://www.mathworld.com\ by Eric Weisstein et. al. \<^item> \<^emph>\Algebra I\ by van der Waerden, Springer \ class ordered_ab_semigroup_add = order + ab_semigroup_add + assumes add_left_mono: "a \ b \ c + a \ c + b" begin lemma add_right_mono: "a \ b \ a + c \ b + c" by (simp add: add.commute [of _ c] add_left_mono) text \non-strict, in both arguments\ lemma add_mono: "a \ b \ c \ d \ a + c \ b + d" by (simp add: add.commute add_left_mono add_right_mono [THEN order_trans]) end text \Strict monotonicity in both arguments\ class strict_ordered_ab_semigroup_add = ordered_ab_semigroup_add + assumes add_strict_mono: "a < b \ c < d \ a + c < b + d" class ordered_cancel_ab_semigroup_add = ordered_ab_semigroup_add + cancel_ab_semigroup_add begin lemma add_strict_left_mono: "a < b \ c + a < c + b" by (auto simp add: less_le add_left_mono) lemma add_strict_right_mono: "a < b \ a + c < b + c" by (simp add: add.commute [of _ c] add_strict_left_mono) subclass strict_ordered_ab_semigroup_add proof show "\a b c d. \a < b; c < d\ \ a + c < b + d" by (iprover intro: add_strict_left_mono add_strict_right_mono less_trans) qed lemma add_less_le_mono: "a < b \ c \ d \ a + c < b + d" by (iprover intro: add_left_mono add_strict_right_mono less_le_trans) lemma add_le_less_mono: "a \ b \ c < d \ a + c < b + d" by (iprover intro: add_strict_left_mono add_right_mono less_le_trans) end class ordered_ab_semigroup_add_imp_le = ordered_cancel_ab_semigroup_add + assumes add_le_imp_le_left: "c + a \ c + b \ a \ b" begin lemma add_less_imp_less_left: assumes less: "c + a < c + b" shows "a < b" proof - from less have le: "c + a \ c + b" by (simp add: order_le_less) have "a \ b" using add_le_imp_le_left [OF le] . moreover have "a \ b" proof (rule ccontr) assume "\ ?thesis" then have "a = b" by simp then have "c + a = c + b" by simp with less show "False" by simp qed ultimately show "a < b" by (simp add: order_le_less) qed lemma add_less_imp_less_right: "a + c < b + c \ a < b" by (rule add_less_imp_less_left [of c]) (simp add: add.commute) lemma add_less_cancel_left [simp]: "c + a < c + b \ a < b" by (blast intro: add_less_imp_less_left add_strict_left_mono) lemma add_less_cancel_right [simp]: "a + c < b + c \ a < b" by (blast intro: add_less_imp_less_right add_strict_right_mono) lemma add_le_cancel_left [simp]: "c + a \ c + b \ a \ b" by (auto simp: dest: add_le_imp_le_left add_left_mono) lemma add_le_cancel_right [simp]: "a + c \ b + c \ a \ b" by (simp add: add.commute [of a c] add.commute [of b c]) lemma add_le_imp_le_right: "a + c \ b + c \ a \ b" by simp lemma max_add_distrib_left: "max x y + z = max (x + z) (y + z)" unfolding max_def by auto lemma min_add_distrib_left: "min x y + z = min (x + z) (y + z)" unfolding min_def by auto lemma max_add_distrib_right: "x + max y z = max (x + y) (x + z)" unfolding max_def by auto lemma min_add_distrib_right: "x + min y z = min (x + y) (x + z)" unfolding min_def by auto end subsection \Support for reasoning about signs\ class ordered_comm_monoid_add = comm_monoid_add + ordered_ab_semigroup_add begin lemma add_nonneg_nonneg [simp]: "0 \ a \ 0 \ b \ 0 \ a + b" using add_mono[of 0 a 0 b] by simp lemma add_nonpos_nonpos: "a \ 0 \ b \ 0 \ a + b \ 0" using add_mono[of a 0 b 0] by simp lemma add_nonneg_eq_0_iff: "0 \ x \ 0 \ y \ x + y = 0 \ x = 0 \ y = 0" using add_left_mono[of 0 y x] add_right_mono[of 0 x y] by auto lemma add_nonpos_eq_0_iff: "x \ 0 \ y \ 0 \ x + y = 0 \ x = 0 \ y = 0" using add_left_mono[of y 0 x] add_right_mono[of x 0 y] by auto lemma add_increasing: "0 \ a \ b \ c \ b \ a + c" using add_mono [of 0 a b c] by simp lemma add_increasing2: "0 \ c \ b \ a \ b \ a + c" by (simp add: add_increasing add.commute [of a]) lemma add_decreasing: "a \ 0 \ c \ b \ a + c \ b" using add_mono [of a 0 c b] by simp lemma add_decreasing2: "c \ 0 \ a \ b \ a + c \ b" using add_mono[of a b c 0] by simp lemma add_pos_nonneg: "0 < a \ 0 \ b \ 0 < a + b" using less_le_trans[of 0 a "a + b"] by (simp add: add_increasing2) lemma add_pos_pos: "0 < a \ 0 < b \ 0 < a + b" by (intro add_pos_nonneg less_imp_le) lemma add_nonneg_pos: "0 \ a \ 0 < b \ 0 < a + b" using add_pos_nonneg[of b a] by (simp add: add_commute) lemma add_neg_nonpos: "a < 0 \ b \ 0 \ a + b < 0" using le_less_trans[of "a + b" a 0] by (simp add: add_decreasing2) lemma add_neg_neg: "a < 0 \ b < 0 \ a + b < 0" by (intro add_neg_nonpos less_imp_le) lemma add_nonpos_neg: "a \ 0 \ b < 0 \ a + b < 0" using add_neg_nonpos[of b a] by (simp add: add_commute) lemmas add_sign_intros = add_pos_nonneg add_pos_pos add_nonneg_pos add_nonneg_nonneg add_neg_nonpos add_neg_neg add_nonpos_neg add_nonpos_nonpos end class strict_ordered_comm_monoid_add = comm_monoid_add + strict_ordered_ab_semigroup_add begin lemma pos_add_strict: "0 < a \ b < c \ b < a + c" using add_strict_mono [of 0 a b c] by simp end class ordered_cancel_comm_monoid_add = ordered_comm_monoid_add + cancel_ab_semigroup_add begin subclass ordered_cancel_ab_semigroup_add .. subclass strict_ordered_comm_monoid_add .. lemma add_strict_increasing: "0 < a \ b \ c \ b < a + c" using add_less_le_mono [of 0 a b c] by simp lemma add_strict_increasing2: "0 \ a \ b < c \ b < a + c" using add_le_less_mono [of 0 a b c] by simp end class ordered_ab_semigroup_monoid_add_imp_le = monoid_add + ordered_ab_semigroup_add_imp_le begin lemma add_less_same_cancel1 [simp]: "b + a < b \ a < 0" using add_less_cancel_left [of _ _ 0] by simp lemma add_less_same_cancel2 [simp]: "a + b < b \ a < 0" using add_less_cancel_right [of _ _ 0] by simp lemma less_add_same_cancel1 [simp]: "a < a + b \ 0 < b" using add_less_cancel_left [of _ 0] by simp lemma less_add_same_cancel2 [simp]: "a < b + a \ 0 < b" using add_less_cancel_right [of 0] by simp lemma add_le_same_cancel1 [simp]: "b + a \ b \ a \ 0" using add_le_cancel_left [of _ _ 0] by simp lemma add_le_same_cancel2 [simp]: "a + b \ b \ a \ 0" using add_le_cancel_right [of _ _ 0] by simp lemma le_add_same_cancel1 [simp]: "a \ a + b \ 0 \ b" using add_le_cancel_left [of _ 0] by simp lemma le_add_same_cancel2 [simp]: "a \ b + a \ 0 \ b" using add_le_cancel_right [of 0] by simp subclass cancel_comm_monoid_add by standard auto subclass ordered_cancel_comm_monoid_add by standard end class ordered_ab_group_add = ab_group_add + ordered_ab_semigroup_add begin subclass ordered_cancel_ab_semigroup_add .. subclass ordered_ab_semigroup_monoid_add_imp_le proof fix a b c :: 'a assume "c + a \ c + b" then have "(-c) + (c + a) \ (-c) + (c + b)" by (rule add_left_mono) then have "((-c) + c) + a \ ((-c) + c) + b" by (simp only: add.assoc) then show "a \ b" by simp qed lemma max_diff_distrib_left: "max x y - z = max (x - z) (y - z)" using max_add_distrib_left [of x y "- z"] by simp lemma min_diff_distrib_left: "min x y - z = min (x - z) (y - z)" using min_add_distrib_left [of x y "- z"] by simp lemma le_imp_neg_le: assumes "a \ b" shows "- b \ - a" proof - from assms have "- a + a \ - a + b" by (rule add_left_mono) then have "0 \ - a + b" by simp then have "0 + (- b) \ (- a + b) + (- b)" by (rule add_right_mono) then show ?thesis by (simp add: algebra_simps) qed lemma neg_le_iff_le [simp]: "- b \ - a \ a \ b" proof assume "- b \ - a" then have "- (- a) \ - (- b)" by (rule le_imp_neg_le) then show "a \ b" by simp next assume "a \ b" then show "- b \ - a" by (rule le_imp_neg_le) qed lemma neg_le_0_iff_le [simp]: "- a \ 0 \ 0 \ a" by (subst neg_le_iff_le [symmetric]) simp lemma neg_0_le_iff_le [simp]: "0 \ - a \ a \ 0" by (subst neg_le_iff_le [symmetric]) simp lemma neg_less_iff_less [simp]: "- b < - a \ a < b" by (auto simp add: less_le) lemma neg_less_0_iff_less [simp]: "- a < 0 \ 0 < a" by (subst neg_less_iff_less [symmetric]) simp lemma neg_0_less_iff_less [simp]: "0 < - a \ a < 0" by (subst neg_less_iff_less [symmetric]) simp text \The next several equations can make the simplifier loop!\ lemma less_minus_iff: "a < - b \ b < - a" proof - have "- (- a) < - b \ b < - a" by (rule neg_less_iff_less) then show ?thesis by simp qed lemma minus_less_iff: "- a < b \ - b < a" proof - have "- a < - (- b) \ - b < a" by (rule neg_less_iff_less) then show ?thesis by simp qed lemma le_minus_iff: "a \ - b \ b \ - a" by (auto simp: order.order_iff_strict less_minus_iff) lemma minus_le_iff: "- a \ b \ - b \ a" by (auto simp add: le_less minus_less_iff) lemma diff_less_0_iff_less [simp]: "a - b < 0 \ a < b" proof - have "a - b < 0 \ a + (- b) < b + (- b)" by simp also have "\ \ a < b" by (simp only: add_less_cancel_right) finally show ?thesis . qed lemmas less_iff_diff_less_0 = diff_less_0_iff_less [symmetric] lemma diff_less_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b < c \ a < c + b" proof (subst less_iff_diff_less_0 [of a]) show "(a - b < c) = (a - (c + b) < 0)" by (simp add: algebra_simps less_iff_diff_less_0 [of _ c]) qed lemma less_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a < c - b \ a + b < c" proof (subst less_iff_diff_less_0 [of "a + b"]) show "(a < c - b) = (a + b - c < 0)" by (simp add: algebra_simps less_iff_diff_less_0 [of a]) qed lemma diff_gt_0_iff_gt [simp]: "a - b > 0 \ a > b" by (simp add: less_diff_eq) lemma diff_le_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a - b \ c \ a \ c + b" by (auto simp add: le_less diff_less_eq ) lemma le_diff_eq [algebra_simps, algebra_split_simps, field_simps, field_split_simps]: "a \ c - b \ a + b \ c" by (auto simp add: le_less less_diff_eq) lemma diff_le_0_iff_le [simp]: "a - b \ 0 \ a \ b" by (simp add: algebra_simps) lemmas le_iff_diff_le_0 = diff_le_0_iff_le [symmetric] lemma diff_ge_0_iff_ge [simp]: "a - b \ 0 \ a \ b" by (simp add: le_diff_eq) lemma diff_eq_diff_less: "a - b = c - d \ a < b \ c < d" by (auto simp only: less_iff_diff_less_0 [of a b] less_iff_diff_less_0 [of c d]) lemma diff_eq_diff_less_eq: "a - b = c - d \ a \ b \ c \ d" by (auto simp only: le_iff_diff_le_0 [of a b] le_iff_diff_le_0 [of c d]) lemma diff_mono: "a \ b \ d \ c \ a - c \ b - d" by (simp add: field_simps add_mono) lemma diff_left_mono: "b \ a \ c - a \ c - b" by (simp add: field_simps) lemma diff_right_mono: "a \ b \ a - c \ b - c" by (simp add: field_simps) lemma diff_strict_mono: "a < b \ d < c \ a - c < b - d" by (simp add: field_simps add_strict_mono) lemma diff_strict_left_mono: "b < a \ c - a < c - b" by (simp add: field_simps) lemma diff_strict_right_mono: "a < b \ a - c < b - c" by (simp add: field_simps) end locale group_cancel begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma sub1: "(A::'a::ab_group_add) \ k + a \ A - b \ k + (a - b)" by (simp only: add_diff_eq) lemma sub2: "(B::'a::ab_group_add) \ k + b \ a - B \ - k + (a - b)" by (simp only: minus_add diff_conv_add_uminus ac_simps) lemma neg1: "(A::'a::ab_group_add) \ k + a \ - A \ - k + - a" by (simp only: minus_add_distrib) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/group_cancel.ML\ simproc_setup group_cancel_add ("a + b::'a::ab_group_add") = \fn phi => fn ss => try Group_Cancel.cancel_add_conv\ simproc_setup group_cancel_diff ("a - b::'a::ab_group_add") = \fn phi => fn ss => try Group_Cancel.cancel_diff_conv\ simproc_setup group_cancel_eq ("a = (b::'a::ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_eq_conv\ simproc_setup group_cancel_le ("a \ (b::'a::ordered_ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_le_conv\ simproc_setup group_cancel_less ("a < (b::'a::ordered_ab_group_add)") = \fn phi => fn ss => try Group_Cancel.cancel_less_conv\ class linordered_ab_semigroup_add = linorder + ordered_ab_semigroup_add class linordered_cancel_ab_semigroup_add = linorder + ordered_cancel_ab_semigroup_add begin subclass linordered_ab_semigroup_add .. subclass ordered_ab_semigroup_add_imp_le proof fix a b c :: 'a assume le1: "c + a \ c + b" show "a \ b" proof (rule ccontr) assume *: "\ ?thesis" then have "b \ a" by (simp add: linorder_not_le) then have "c + b \ c + a" by (rule add_left_mono) then have "c + a = c + b" using le1 by (iprover intro: order.antisym) then have "a = b" by simp with * show False by (simp add: linorder_not_le [symmetric]) qed qed end class linordered_ab_group_add = linorder + ordered_ab_group_add begin subclass linordered_cancel_ab_semigroup_add .. lemma equal_neg_zero [simp]: "a = - a \ a = 0" proof assume "a = 0" then show "a = - a" by simp next assume A: "a = - a" show "a = 0" proof (cases "0 \ a") case True with A have "0 \ - a" by auto with le_minus_iff have "a \ 0" by simp with True show ?thesis by (auto intro: order_trans) next case False then have B: "a \ 0" by auto with A have "- a \ 0" by auto with B show ?thesis by (auto intro: order_trans) qed qed lemma neg_equal_zero [simp]: "- a = a \ a = 0" by (auto dest: sym) lemma neg_less_eq_nonneg [simp]: "- a \ a \ 0 \ a" proof assume *: "- a \ a" show "0 \ a" proof (rule classical) assume "\ ?thesis" then have "a < 0" by auto with * have "- a < 0" by (rule le_less_trans) then show ?thesis by auto qed next assume *: "0 \ a" then have "- a \ 0" by (simp add: minus_le_iff) from this * show "- a \ a" by (rule order_trans) qed lemma neg_less_pos [simp]: "- a < a \ 0 < a" by (auto simp add: less_le) lemma less_eq_neg_nonpos [simp]: "a \ - a \ a \ 0" using neg_less_eq_nonneg [of "- a"] by simp lemma less_neg_neg [simp]: "a < - a \ a < 0" using neg_less_pos [of "- a"] by simp lemma double_zero [simp]: "a + a = 0 \ a = 0" proof assume "a + a = 0" then have a: "- a = a" by (rule minus_unique) then show "a = 0" by (simp only: neg_equal_zero) next assume "a = 0" then show "a + a = 0" by simp qed lemma double_zero_sym [simp]: "0 = a + a \ a = 0" using double_zero [of a] by (simp only: eq_commute) lemma zero_less_double_add_iff_zero_less_single_add [simp]: "0 < a + a \ 0 < a" proof assume "0 < a + a" then have "0 - a < a" by (simp only: diff_less_eq) then have "- a < a" by simp then show "0 < a" by simp next assume "0 < a" with this have "0 + 0 < a + a" by (rule add_strict_mono) then show "0 < a + a" by simp qed lemma zero_le_double_add_iff_zero_le_single_add [simp]: "0 \ a + a \ 0 \ a" by (auto simp add: le_less) lemma double_add_less_zero_iff_single_add_less_zero [simp]: "a + a < 0 \ a < 0" proof - have "\ a + a < 0 \ \ a < 0" by (simp add: not_less) then show ?thesis by simp qed lemma double_add_le_zero_iff_single_add_le_zero [simp]: "a + a \ 0 \ a \ 0" proof - have "\ a + a \ 0 \ \ a \ 0" by (simp add: not_le) then show ?thesis by simp qed lemma minus_max_eq_min: "- max x y = min (- x) (- y)" by (auto simp add: max_def min_def) lemma minus_min_eq_max: "- min x y = max (- x) (- y)" by (auto simp add: max_def min_def) end class abs = fixes abs :: "'a \ 'a" ("\_\") class sgn = fixes sgn :: "'a \ 'a" class ordered_ab_group_add_abs = ordered_ab_group_add + abs + assumes abs_ge_zero [simp]: "\a\ \ 0" and abs_ge_self: "a \ \a\" and abs_leI: "a \ b \ - a \ b \ \a\ \ b" and abs_minus_cancel [simp]: "\-a\ = \a\" and abs_triangle_ineq: "\a + b\ \ \a\ + \b\" begin lemma abs_minus_le_zero: "- \a\ \ 0" unfolding neg_le_0_iff_le by simp lemma abs_of_nonneg [simp]: assumes nonneg: "0 \ a" shows "\a\ = a" proof (rule order.antisym) show "a \ \a\" by (rule abs_ge_self) from nonneg le_imp_neg_le have "- a \ 0" by simp from this nonneg have "- a \ a" by (rule order_trans) then show "\a\ \ a" by (auto intro: abs_leI) qed lemma abs_idempotent [simp]: "\\a\\ = \a\" by (rule order.antisym) (auto intro!: abs_ge_self abs_leI order_trans [of "- \a\" 0 "\a\"]) lemma abs_eq_0 [simp]: "\a\ = 0 \ a = 0" proof - have "\a\ = 0 \ a = 0" proof (rule order.antisym) assume zero: "\a\ = 0" with abs_ge_self show "a \ 0" by auto from zero have "\-a\ = 0" by simp with abs_ge_self [of "- a"] have "- a \ 0" by auto with neg_le_0_iff_le show "0 \ a" by auto qed then show ?thesis by auto qed lemma abs_zero [simp]: "\0\ = 0" by simp lemma abs_0_eq [simp]: "0 = \a\ \ a = 0" proof - have "0 = \a\ \ \a\ = 0" by (simp only: eq_ac) then show ?thesis by simp qed lemma abs_le_zero_iff [simp]: "\a\ \ 0 \ a = 0" proof assume "\a\ \ 0" then have "\a\ = 0" by (rule order.antisym) simp then show "a = 0" by simp next assume "a = 0" then show "\a\ \ 0" by simp qed lemma abs_le_self_iff [simp]: "\a\ \ a \ 0 \ a" proof - have "0 \ \a\" using abs_ge_zero by blast then have "\a\ \ a \ 0 \ a" using order.trans by blast then show ?thesis using abs_of_nonneg eq_refl by blast qed lemma zero_less_abs_iff [simp]: "0 < \a\ \ a \ 0" by (simp add: less_le) lemma abs_not_less_zero [simp]: "\ \a\ < 0" proof - have "x \ y \ \ y < x" for x y by auto then show ?thesis by simp qed lemma abs_ge_minus_self: "- a \ \a\" proof - have "- a \ \-a\" by (rule abs_ge_self) then show ?thesis by simp qed lemma abs_minus_commute: "\a - b\ = \b - a\" proof - have "\a - b\ = \- (a - b)\" by (simp only: abs_minus_cancel) also have "\ = \b - a\" by simp finally show ?thesis . qed lemma abs_of_pos: "0 < a \ \a\ = a" by (rule abs_of_nonneg) (rule less_imp_le) lemma abs_of_nonpos [simp]: assumes "a \ 0" shows "\a\ = - a" proof - let ?b = "- a" have "- ?b \ 0 \ \- ?b\ = - (- ?b)" unfolding abs_minus_cancel [of ?b] unfolding neg_le_0_iff_le [of ?b] unfolding minus_minus by (erule abs_of_nonneg) then show ?thesis using assms by auto qed lemma abs_of_neg: "a < 0 \ \a\ = - a" by (rule abs_of_nonpos) (rule less_imp_le) lemma abs_le_D1: "\a\ \ b \ a \ b" using abs_ge_self by (blast intro: order_trans) lemma abs_le_D2: "\a\ \ b \ - a \ b" using abs_le_D1 [of "- a"] by simp lemma abs_le_iff: "\a\ \ b \ a \ b \ - a \ b" by (blast intro: abs_leI dest: abs_le_D1 abs_le_D2) lemma abs_triangle_ineq2: "\a\ - \b\ \ \a - b\" proof - have "\a\ = \b + (a - b)\" by (simp add: algebra_simps) then have "\a\ \ \b\ + \a - b\" by (simp add: abs_triangle_ineq) then show ?thesis by (simp add: algebra_simps) qed lemma abs_triangle_ineq2_sym: "\a\ - \b\ \ \b - a\" by (simp only: abs_minus_commute [of b] abs_triangle_ineq2) lemma abs_triangle_ineq3: "\\a\ - \b\\ \ \a - b\" by (simp add: abs_le_iff abs_triangle_ineq2 abs_triangle_ineq2_sym) lemma abs_triangle_ineq4: "\a - b\ \ \a\ + \b\" proof - have "\a - b\ = \a + - b\" by (simp add: algebra_simps) also have "\ \ \a\ + \- b\" by (rule abs_triangle_ineq) finally show ?thesis by simp qed lemma abs_diff_triangle_ineq: "\a + b - (c + d)\ \ \a - c\ + \b - d\" proof - have "\a + b - (c + d)\ = \(a - c) + (b - d)\" by (simp add: algebra_simps) also have "\ \ \a - c\ + \b - d\" by (rule abs_triangle_ineq) finally show ?thesis . qed lemma abs_add_abs [simp]: "\\a\ + \b\\ = \a\ + \b\" (is "?L = ?R") proof (rule order.antisym) show "?L \ ?R" by (rule abs_ge_self) have "?L \ \\a\\ + \\b\\" by (rule abs_triangle_ineq) also have "\ = ?R" by simp finally show "?L \ ?R" . qed end lemma dense_eq0_I: fixes x::"'a::{dense_linorder,ordered_ab_group_add_abs}" assumes "\e. 0 < e \ \x\ \ e" shows "x = 0" proof (cases "\x\ = 0") case False then have "\x\ > 0" by simp then obtain z where "0 < z" "z < \x\" using dense by force then show ?thesis using assms by (simp flip: not_less) qed auto hide_fact (open) ab_diff_conv_add_uminus add_0 mult_1 ab_left_minus lemmas add_0 = add_0_left (* FIXME duplicate *) lemmas mult_1 = mult_1_left (* FIXME duplicate *) lemmas ab_left_minus = left_minus (* FIXME duplicate *) lemmas diff_diff_eq = diff_diff_add (* FIXME duplicate *) subsection \Canonically ordered monoids\ text \Canonically ordered monoids are never groups.\ class canonically_ordered_monoid_add = comm_monoid_add + order + assumes le_iff_add: "a \ b \ (\c. b = a + c)" begin lemma zero_le[simp]: "0 \ x" by (auto simp: le_iff_add) lemma le_zero_eq[simp]: "n \ 0 \ n = 0" by (auto intro: order.antisym) lemma not_less_zero[simp]: "\ n < 0" by (auto simp: less_le) lemma zero_less_iff_neq_zero: "0 < n \ n \ 0" by (auto simp: less_le) text \This theorem is useful with \blast\\ lemma gr_zeroI: "(n = 0 \ False) \ 0 < n" by (rule zero_less_iff_neq_zero[THEN iffD2]) iprover lemma not_gr_zero[simp]: "\ 0 < n \ n = 0" by (simp add: zero_less_iff_neq_zero) subclass ordered_comm_monoid_add proof qed (auto simp: le_iff_add add_ac) lemma gr_implies_not_zero: "m < n \ n \ 0" by auto lemma add_eq_0_iff_both_eq_0[simp]: "x + y = 0 \ x = 0 \ y = 0" by (intro add_nonneg_eq_0_iff zero_le) lemma zero_eq_add_iff_both_eq_0[simp]: "0 = x + y \ x = 0 \ y = 0" using add_eq_0_iff_both_eq_0[of x y] unfolding eq_commute[of 0] . lemma less_eqE: assumes \a \ b\ obtains c where \b = a + c\ using assms by (auto simp add: le_iff_add) lemma lessE: assumes \a < b\ obtains c where \b = a + c\ and \c \ 0\ proof - from assms have \a \ b\ \a \ b\ by simp_all from \a \ b\ obtain c where \b = a + c\ by (rule less_eqE) moreover have \c \ 0\ using \a \ b\ \b = a + c\ by auto ultimately show ?thesis by (rule that) qed lemmas zero_order = zero_le le_zero_eq not_less_zero zero_less_iff_neq_zero not_gr_zero \ \This should be attributed with \[iff]\, but then \blast\ fails in \Set\.\ end class ordered_cancel_comm_monoid_diff = canonically_ordered_monoid_add + comm_monoid_diff + ordered_ab_semigroup_add_imp_le begin context fixes a b :: 'a assumes le: "a \ b" begin lemma add_diff_inverse: "a + (b - a) = b" using le by (auto simp add: le_iff_add) lemma add_diff_assoc: "c + (b - a) = c + b - a" using le by (auto simp add: le_iff_add add.left_commute [of c]) lemma add_diff_assoc2: "b - a + c = b + c - a" using le by (auto simp add: le_iff_add add.assoc) lemma diff_add_assoc: "c + b - a = c + (b - a)" using le by (simp add: add.commute add_diff_assoc) lemma diff_add_assoc2: "b + c - a = b - a + c" using le by (simp add: add.commute add_diff_assoc) lemma diff_diff_right: "c - (b - a) = c + a - b" by (simp add: add_diff_inverse add_diff_cancel_left [of a c "b - a", symmetric] add.commute) lemma diff_add: "b - a + a = b" by (simp add: add.commute add_diff_inverse) lemma le_add_diff: "c \ b + c - a" by (auto simp add: add.commute diff_add_assoc2 le_iff_add) lemma le_imp_diff_is_add: "a \ b \ b - a = c \ b = c + a" by (auto simp add: add.commute add_diff_inverse) lemma le_diff_conv2: "c \ b - a \ c + a \ b" (is "?P \ ?Q") proof assume ?P then have "c + a \ b - a + a" by (rule add_right_mono) then show ?Q by (simp add: add_diff_inverse add.commute) next assume ?Q then have "a + c \ a + (b - a)" by (simp add: add_diff_inverse add.commute) then show ?P by simp qed end end subsection \Tools setup\ lemma add_mono_thms_linordered_semiring: fixes i j k :: "'a::ordered_ab_semigroup_add" shows "i \ j \ k \ l \ i + k \ j + l" and "i = j \ k \ l \ i + k \ j + l" and "i \ j \ k = l \ i + k \ j + l" and "i = j \ k = l \ i + k = j + l" by (rule add_mono, clarify+)+ lemma add_mono_thms_linordered_field: fixes i j k :: "'a::ordered_cancel_ab_semigroup_add" shows "i < j \ k = l \ i + k < j + l" and "i = j \ k < l \ i + k < j + l" and "i < j \ k \ l \ i + k < j + l" and "i \ j \ k < l \ i + k < j + l" and "i < j \ k < l \ i + k < j + l" by (auto intro: add_strict_right_mono add_strict_left_mono add_less_le_mono add_le_less_mono add_strict_mono) code_identifier code_module Groups \ (SML) Arith and (OCaml) Arith and (Haskell) Arith end diff --git a/src/HOL/HOL.thy b/src/HOL/HOL.thy --- a/src/HOL/HOL.thy +++ b/src/HOL/HOL.thy @@ -1,2177 +1,2177 @@ (* Title: HOL/HOL.thy Author: Tobias Nipkow, Markus Wenzel, and Larry Paulson *) section \The basis of Higher-Order Logic\ theory HOL imports Pure Tools.Code_Generator keywords "try" "solve_direct" "quickcheck" "print_coercions" "print_claset" "print_induct_rules" :: diag and "quickcheck_params" :: thy_decl abbrevs "?<" = "\\<^sub>\\<^sub>1" begin ML_file \~~/src/Tools/misc_legacy.ML\ ML_file \~~/src/Tools/try.ML\ ML_file \~~/src/Tools/quickcheck.ML\ ML_file \~~/src/Tools/solve_direct.ML\ ML_file \~~/src/Tools/IsaPlanner/zipper.ML\ ML_file \~~/src/Tools/IsaPlanner/isand.ML\ ML_file \~~/src/Tools/IsaPlanner/rw_inst.ML\ ML_file \~~/src/Provers/hypsubst.ML\ ML_file \~~/src/Provers/splitter.ML\ ML_file \~~/src/Provers/classical.ML\ ML_file \~~/src/Provers/blast.ML\ ML_file \~~/src/Provers/clasimp.ML\ ML_file \~~/src/Tools/eqsubst.ML\ ML_file \~~/src/Provers/quantifier1.ML\ ML_file \~~/src/Tools/atomize_elim.ML\ ML_file \~~/src/Tools/cong_tac.ML\ ML_file \~~/src/Tools/intuitionistic.ML\ setup \Intuitionistic.method_setup \<^binding>\iprover\\ ML_file \~~/src/Tools/project_rule.ML\ ML_file \~~/src/Tools/subtyping.ML\ ML_file \~~/src/Tools/case_product.ML\ ML \Plugin_Name.declare_setup \<^binding>\extraction\\ ML \ Plugin_Name.declare_setup \<^binding>\quickcheck_random\; Plugin_Name.declare_setup \<^binding>\quickcheck_exhaustive\; Plugin_Name.declare_setup \<^binding>\quickcheck_bounded_forall\; Plugin_Name.declare_setup \<^binding>\quickcheck_full_exhaustive\; Plugin_Name.declare_setup \<^binding>\quickcheck_narrowing\; \ ML \ Plugin_Name.define_setup \<^binding>\quickcheck\ [\<^plugin>\quickcheck_exhaustive\, \<^plugin>\quickcheck_random\, \<^plugin>\quickcheck_bounded_forall\, \<^plugin>\quickcheck_full_exhaustive\, \<^plugin>\quickcheck_narrowing\] \ subsection \Primitive logic\ text \ The definition of the logic is based on Mike Gordon's technical report \<^cite>\"Gordon-TR68"\ that describes the first implementation of HOL. However, there are a number of differences. In particular, we start with the definite description operator and introduce Hilbert's \\\ operator only much later. Moreover, axiom \(P \ Q) \ (Q \ P) \ (P = Q)\ is derived from the other axioms. The fact that this axiom is derivable was first noticed by Bruno Barras (for Mike Gordon's line of HOL systems) and later independently by Alexander Maletzky (for Isabelle/HOL). \ subsubsection \Core syntax\ setup \Axclass.class_axiomatization (\<^binding>\type\, [])\ default_sort type setup \Object_Logic.add_base_sort \<^sort>\type\\ setup \Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc [])\ axiomatization where fun_arity: "OFCLASS('a \ 'b, type_class)" instance "fun" :: (type, type) type by (rule fun_arity) axiomatization where itself_arity: "OFCLASS('a itself, type_class)" instance itself :: (type) type by (rule itself_arity) typedecl bool judgment Trueprop :: "bool \ prop" ("(_)" 5) axiomatization implies :: "[bool, bool] \ bool" (infixr "\" 25) and eq :: "['a, 'a] \ bool" and The :: "('a \ bool) \ 'a" notation (input) eq (infixl "=" 50) notation (output) eq (infix "=" 50) text \The input syntax for \eq\ is more permissive than the output syntax because of the large amount of material that relies on infixl.\ subsubsection \Defined connectives and quantifiers\ definition True :: bool where "True \ ((\x::bool. x) = (\x. x))" definition All :: "('a \ bool) \ bool" (binder "\" 10) where "All P \ (P = (\x. True))" definition Ex :: "('a \ bool) \ bool" (binder "\" 10) where "Ex P \ \Q. (\x. P x \ Q) \ Q" definition False :: bool where "False \ (\P. P)" definition Not :: "bool \ bool" ("\ _" [40] 40) where not_def: "\ P \ P \ False" definition conj :: "[bool, bool] \ bool" (infixr "\" 35) where and_def: "P \ Q \ \R. (P \ Q \ R) \ R" definition disj :: "[bool, bool] \ bool" (infixr "\" 30) where or_def: "P \ Q \ \R. (P \ R) \ (Q \ R) \ R" definition Uniq :: "('a \ bool) \ bool" where "Uniq P \ (\x y. P x \ P y \ y = x)" definition Ex1 :: "('a \ bool) \ bool" where "Ex1 P \ \x. P x \ (\y. P y \ y = x)" subsubsection \Additional concrete syntax\ syntax (ASCII) "_Uniq" :: "pttrn \ bool \ bool" ("(4?< _./ _)" [0, 10] 10) syntax "_Uniq" :: "pttrn \ bool \ bool" ("(2\\<^sub>\\<^sub>1 _./ _)" [0, 10] 10) translations "\\<^sub>\\<^sub>1x. P" \ "CONST Uniq (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\Uniq\ \<^syntax_const>\_Uniq\] \ \ \to avoid eta-contraction of body\ syntax (ASCII) "_Ex1" :: "pttrn \ bool \ bool" ("(3EX! _./ _)" [0, 10] 10) syntax (input) "_Ex1" :: "pttrn \ bool \ bool" ("(3?! _./ _)" [0, 10] 10) syntax "_Ex1" :: "pttrn \ bool \ bool" ("(3\!_./ _)" [0, 10] 10) translations "\!x. P" \ "CONST Ex1 (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs_tr' \<^const_syntax>\Ex1\ \<^syntax_const>\_Ex1\] \ \ \to avoid eta-contraction of body\ syntax "_Not_Ex" :: "idts \ bool \ bool" ("(3\_./ _)" [0, 10] 10) "_Not_Ex1" :: "pttrn \ bool \ bool" ("(3\!_./ _)" [0, 10] 10) translations "\x. P" \ "\ (\x. P)" "\!x. P" \ "\ (\!x. P)" abbreviation not_equal :: "['a, 'a] \ bool" (infix "\" 50) where "x \ y \ \ (x = y)" notation (ASCII) Not ("~ _" [40] 40) and conj (infixr "&" 35) and disj (infixr "|" 30) and implies (infixr "-->" 25) and not_equal (infix "~=" 50) abbreviation (iff) iff :: "[bool, bool] \ bool" (infixr "\" 25) where "A \ B \ A = B" syntax "_The" :: "[pttrn, bool] \ 'a" ("(3THE _./ _)" [0, 10] 10) translations "THE x. P" \ "CONST The (\x. P)" print_translation \ [(\<^const_syntax>\The\, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_The\ $ x $ t end)] \ \ \To avoid eta-contraction of body\ nonterminal letbinds and letbind syntax "_bind" :: "[pttrn, 'a] \ letbind" ("(2_ =/ _)" 10) "" :: "letbind \ letbinds" ("_") "_binds" :: "[letbind, letbinds] \ letbinds" ("_;/ _") "_Let" :: "[letbinds, 'a] \ 'a" ("(let (_)/ in (_))" [0, 10] 10) nonterminal case_syn and cases_syn syntax "_case_syntax" :: "['a, cases_syn] \ 'b" ("(case _ of/ _)" 10) "_case1" :: "['a, 'b] \ case_syn" ("(2_ \/ _)" 10) "" :: "case_syn \ cases_syn" ("_") "_case2" :: "[case_syn, cases_syn] \ cases_syn" ("_/ | _") syntax (ASCII) "_case1" :: "['a, 'b] \ case_syn" ("(2_ =>/ _)" 10) notation (ASCII) All (binder "ALL " 10) and Ex (binder "EX " 10) notation (input) All (binder "! " 10) and Ex (binder "? " 10) subsubsection \Axioms and basic definitions\ axiomatization where refl: "t = (t::'a)" and subst: "s = t \ P s \ P t" and ext: "(\x::'a. (f x ::'b) = g x) \ (\x. f x) = (\x. g x)" \ \Extensionality is built into the meta-logic, and this rule expresses a related property. It is an eta-expanded version of the traditional rule, and similar to the ABS rule of HOL\ and the_eq_trivial: "(THE x. x = a) = (a::'a)" axiomatization where impI: "(P \ Q) \ P \ Q" and mp: "\P \ Q; P\ \ Q" and True_or_False: "(P = True) \ (P = False)" definition If :: "bool \ 'a \ 'a \ 'a" ("(if (_)/ then (_)/ else (_))" [0, 0, 10] 10) where "If P x y \ (THE z::'a. (P = True \ z = x) \ (P = False \ z = y))" definition Let :: "'a \ ('a \ 'b) \ 'b" where "Let s f \ f s" translations "_Let (_binds b bs) e" \ "_Let b (_Let bs e)" "let x = a in e" \ "CONST Let a (\x. e)" axiomatization undefined :: 'a class default = fixes default :: 'a subsection \Fundamental rules\ subsubsection \Equality\ lemma sym: "s = t \ t = s" by (erule subst) (rule refl) lemma ssubst: "t = s \ P s \ P t" by (drule sym) (erule subst) lemma trans: "\r = s; s = t\ \ r = t" by (erule subst) lemma trans_sym [Pure.elim?]: "r = s \ t = s \ r = t" by (rule trans [OF _ sym]) lemma meta_eq_to_obj_eq: assumes "A \ B" shows "A = B" unfolding assms by (rule refl) text \Useful with \erule\ for proving equalities from known equalities.\ (* a = b | | c = d *) lemma box_equals: "\a = b; a = c; b = d\ \ c = d" by (iprover intro: sym trans) text \For calculational reasoning:\ lemma forw_subst: "a = b \ P b \ P a" by (rule ssubst) lemma back_subst: "P a \ a = b \ P b" by (rule subst) subsubsection \Congruence rules for application\ text \Similar to \AP_THM\ in Gordon's HOL.\ lemma fun_cong: "(f :: 'a \ 'b) = g \ f x = g x" by (iprover intro: refl elim: subst) text \Similar to \AP_TERM\ in Gordon's HOL and FOL's \subst_context\.\ lemma arg_cong: "x = y \ f x = f y" by (iprover intro: refl elim: subst) lemma arg_cong2: "\a = b; c = d\ \ f a c = f b d" by (iprover intro: refl elim: subst) lemma cong: "\f = g; (x::'a) = y\ \ f x = g y" by (iprover intro: refl elim: subst) ML \fun cong_tac ctxt = Cong_Tac.cong_tac ctxt @{thm cong}\ subsubsection \Equality of booleans -- iff\ lemma iffD2: "\P = Q; Q\ \ P" by (erule ssubst) lemma rev_iffD2: "\Q; P = Q\ \ P" by (erule iffD2) lemma iffD1: "Q = P \ Q \ P" by (drule sym) (rule iffD2) lemma rev_iffD1: "Q \ Q = P \ P" by (drule sym) (rule rev_iffD2) lemma iffE: assumes major: "P = Q" and minor: "\P \ Q; Q \ P\ \ R" shows R by (iprover intro: minor impI major [THEN iffD2] major [THEN iffD1]) subsubsection \True (1)\ lemma TrueI: True unfolding True_def by (rule refl) lemma eqTrueE: "P = True \ P" by (erule iffD2) (rule TrueI) subsubsection \Universal quantifier (1)\ lemma spec: "\x::'a. P x \ P x" unfolding All_def by (iprover intro: eqTrueE fun_cong) lemma allE: assumes major: "\x. P x" and minor: "P x \ R" shows R by (iprover intro: minor major [THEN spec]) lemma all_dupE: assumes major: "\x. P x" and minor: "\P x; \x. P x\ \ R" shows R by (iprover intro: minor major major [THEN spec]) subsubsection \False\ text \ Depends upon \spec\; it is impossible to do propositional logic before quantifiers! \ lemma FalseE: "False \ P" unfolding False_def by (erule spec) lemma False_neq_True: "False = True \ P" by (erule eqTrueE [THEN FalseE]) subsubsection \Negation\ lemma notI: assumes "P \ False" shows "\ P" unfolding not_def by (iprover intro: impI assms) lemma False_not_True: "False \ True" by (iprover intro: notI elim: False_neq_True) lemma True_not_False: "True \ False" by (iprover intro: notI dest: sym elim: False_neq_True) lemma notE: "\\ P; P\ \ R" unfolding not_def by (iprover intro: mp [THEN FalseE]) subsubsection \Implication\ lemma impE: assumes "P \ Q" P "Q \ R" shows R by (iprover intro: assms mp) text \Reduces \Q\ to \P \ Q\, allowing substitution in \P\.\ lemma rev_mp: "\P; P \ Q\ \ Q" by (rule mp) lemma contrapos_nn: assumes major: "\ Q" and minor: "P \ Q" shows "\ P" by (iprover intro: notI minor major [THEN notE]) text \Not used at all, but we already have the other 3 combinations.\ lemma contrapos_pn: assumes major: "Q" and minor: "P \ \ Q" shows "\ P" by (iprover intro: notI minor major notE) lemma not_sym: "t \ s \ s \ t" by (erule contrapos_nn) (erule sym) lemma eq_neq_eq_imp_neq: "\x = a; a \ b; b = y\ \ x \ y" by (erule subst, erule ssubst, assumption) subsubsection \Disjunction (1)\ lemma disjE: assumes major: "P \ Q" and minorP: "P \ R" and minorQ: "Q \ R" shows R by (iprover intro: minorP minorQ impI major [unfolded or_def, THEN spec, THEN mp, THEN mp]) subsubsection \Derivation of \iffI\\ text \In an intuitionistic version of HOL \iffI\ needs to be an axiom.\ lemma iffI: assumes "P \ Q" and "Q \ P" shows "P = Q" proof (rule disjE[OF True_or_False[of P]]) assume 1: "P = True" note Q = assms(1)[OF eqTrueE[OF this]] from 1 show ?thesis proof (rule ssubst) from True_or_False[of Q] show "True = Q" proof (rule disjE) assume "Q = True" thus ?thesis by(rule sym) next assume "Q = False" with Q have False by (rule rev_iffD1) thus ?thesis by (rule FalseE) qed qed next assume 2: "P = False" thus ?thesis proof (rule ssubst) from True_or_False[of Q] show "False = Q" proof (rule disjE) assume "Q = True" from 2 assms(2)[OF eqTrueE[OF this]] have False by (rule iffD1) thus ?thesis by (rule FalseE) next assume "Q = False" thus ?thesis by(rule sym) qed qed qed subsubsection \True (2)\ lemma eqTrueI: "P \ P = True" by (iprover intro: iffI TrueI) subsubsection \Universal quantifier (2)\ lemma allI: assumes "\x::'a. P x" shows "\x. P x" unfolding All_def by (iprover intro: ext eqTrueI assms) subsubsection \Existential quantifier\ lemma exI: "P x \ \x::'a. P x" unfolding Ex_def by (iprover intro: allI allE impI mp) lemma exE: assumes major: "\x::'a. P x" and minor: "\x. P x \ Q" shows "Q" by (rule major [unfolded Ex_def, THEN spec, THEN mp]) (iprover intro: impI [THEN allI] minor) subsubsection \Conjunction\ lemma conjI: "\P; Q\ \ P \ Q" unfolding and_def by (iprover intro: impI [THEN allI] mp) lemma conjunct1: "\P \ Q\ \ P" unfolding and_def by (iprover intro: impI dest: spec mp) lemma conjunct2: "\P \ Q\ \ Q" unfolding and_def by (iprover intro: impI dest: spec mp) lemma conjE: assumes major: "P \ Q" and minor: "\P; Q\ \ R" shows R proof (rule minor) show P by (rule major [THEN conjunct1]) show Q by (rule major [THEN conjunct2]) qed lemma context_conjI: assumes P "P \ Q" shows "P \ Q" by (iprover intro: conjI assms) subsubsection \Disjunction (2)\ lemma disjI1: "P \ P \ Q" unfolding or_def by (iprover intro: allI impI mp) lemma disjI2: "Q \ P \ Q" unfolding or_def by (iprover intro: allI impI mp) subsubsection \Classical logic\ lemma classical: assumes "\ P \ P" shows P proof (rule True_or_False [THEN disjE]) show P if "P = True" using that by (iprover intro: eqTrueE) show P if "P = False" proof (intro notI assms) assume P with that show False by (iprover elim: subst) qed qed lemmas ccontr = FalseE [THEN classical] text \\notE\ with premises exchanged; it discharges \\ R\ so that it can be used to make elimination rules.\ lemma rev_notE: assumes premp: P and premnot: "\ R \ \ P" shows R by (iprover intro: ccontr notE [OF premnot premp]) text \Double negation law.\ lemma notnotD: "\\ P \ P" by (iprover intro: ccontr notE ) lemma contrapos_pp: assumes p1: Q and p2: "\ P \ \ Q" shows P by (iprover intro: classical p1 p2 notE) subsubsection \Unique existence\ lemma Uniq_I [intro?]: assumes "\x y. \P x; P y\ \ y = x" shows "Uniq P" unfolding Uniq_def by (iprover intro: assms allI impI) lemma Uniq_D [dest?]: "\Uniq P; P a; P b\ \ a=b" unfolding Uniq_def by (iprover dest: spec mp) lemma ex1I: assumes "P a" "\x. P x \ x = a" shows "\!x. P x" unfolding Ex1_def by (iprover intro: assms exI conjI allI impI) text \Sometimes easier to use: the premises have no shared variables. Safe!\ lemma ex_ex1I: assumes ex_prem: "\x. P x" and eq: "\x y. \P x; P y\ \ x = y" shows "\!x. P x" by (iprover intro: ex_prem [THEN exE] ex1I eq) lemma ex1E: assumes major: "\!x. P x" and minor: "\x. \P x; \y. P y \ y = x\ \ R" shows R proof (rule major [unfolded Ex1_def, THEN exE]) show "\x. P x \ (\y. P y \ y = x) \ R" by (iprover intro: minor elim: conjE) qed lemma ex1_implies_ex: "\!x. P x \ \x. P x" by (iprover intro: exI elim: ex1E) subsubsection \Classical intro rules for disjunction and existential quantifiers\ lemma disjCI: assumes "\ Q \ P" shows "P \ Q" by (rule classical) (iprover intro: assms disjI1 disjI2 notI elim: notE) lemma excluded_middle: "\ P \ P" by (iprover intro: disjCI) text \ case distinction as a natural deduction rule. Note that \\ P\ is the second case, not the first. \ lemma case_split [case_names True False]: assumes "P \ Q" "\ P \ Q" shows Q using excluded_middle [of P] by (iprover intro: assms elim: disjE) text \Classical implies (\\\) elimination.\ lemma impCE: assumes major: "P \ Q" and minor: "\ P \ R" "Q \ R" shows R using excluded_middle [of P] by (iprover intro: minor major [THEN mp] elim: disjE)+ text \ This version of \\\ elimination works on \Q\ before \P\. It works best for those cases in which \P\ holds "almost everywhere". Can't install as default: would break old proofs. \ lemma impCE': assumes major: "P \ Q" and minor: "Q \ R" "\ P \ R" shows R using assms by (elim impCE) text \Classical \\\ elimination.\ lemma iffCE: assumes major: "P = Q" and minor: "\P; Q\ \ R" "\\ P; \ Q\ \ R" shows R by (rule major [THEN iffE]) (iprover intro: minor elim: impCE notE) lemma exCI: assumes "\x. \ P x \ P a" shows "\x. P x" by (rule ccontr) (iprover intro: assms exI allI notI notE [of "\x. P x"]) subsubsection \Intuitionistic Reasoning\ lemma impE': assumes 1: "P \ Q" and 2: "Q \ R" and 3: "P \ Q \ P" shows R proof - from 3 and 1 have P . with 1 have Q by (rule impE) with 2 show R . qed lemma allE': assumes 1: "\x. P x" and 2: "P x \ \x. P x \ Q" shows Q proof - from 1 have "P x" by (rule spec) from this and 1 show Q by (rule 2) qed lemma notE': assumes 1: "\ P" and 2: "\ P \ P" shows R proof - from 2 and 1 have P . with 1 show R by (rule notE) qed lemma TrueE: "True \ P \ P" . lemma notFalseE: "\ False \ P \ P" . lemmas [Pure.elim!] = disjE iffE FalseE conjE exE TrueE notFalseE and [Pure.intro!] = iffI conjI impI TrueI notI allI refl and [Pure.elim 2] = allE notE' impE' and [Pure.intro] = exI disjI2 disjI1 lemmas [trans] = trans and [sym] = sym not_sym and [Pure.elim?] = iffD1 iffD2 impE subsubsection \Atomizing meta-level connectives\ axiomatization where eq_reflection: "x = y \ x \ y" \ \admissible axiom\ lemma atomize_all [atomize]: "(\x. P x) \ Trueprop (\x. P x)" proof assume "\x. P x" then show "\x. P x" .. next assume "\x. P x" then show "\x. P x" by (rule allE) qed lemma atomize_imp [atomize]: "(A \ B) \ Trueprop (A \ B)" proof assume r: "A \ B" show "A \ B" by (rule impI) (rule r) next assume "A \ B" and A then show B by (rule mp) qed lemma atomize_not: "(A \ False) \ Trueprop (\ A)" proof assume r: "A \ False" show "\ A" by (rule notI) (rule r) next assume "\ A" and A then show False by (rule notE) qed lemma atomize_eq [atomize, code]: "(x \ y) \ Trueprop (x = y)" proof assume "x \ y" show "x = y" by (unfold \x \ y\) (rule refl) next assume "x = y" then show "x \ y" by (rule eq_reflection) qed lemma atomize_conj [atomize]: "(A &&& B) \ Trueprop (A \ B)" proof assume conj: "A &&& B" show "A \ B" proof (rule conjI) from conj show A by (rule conjunctionD1) from conj show B by (rule conjunctionD2) qed next assume conj: "A \ B" show "A &&& B" proof - from conj show A .. from conj show B .. qed qed lemmas [symmetric, rulify] = atomize_all atomize_imp and [symmetric, defn] = atomize_all atomize_imp atomize_eq subsubsection \Atomizing elimination rules\ lemma atomize_exL[atomize_elim]: "(\x. P x \ Q) \ ((\x. P x) \ Q)" by (rule equal_intr_rule) iprover+ lemma atomize_conjL[atomize_elim]: "(A \ B \ C) \ (A \ B \ C)" by (rule equal_intr_rule) iprover+ lemma atomize_disjL[atomize_elim]: "((A \ C) \ (B \ C) \ C) \ ((A \ B \ C) \ C)" by (rule equal_intr_rule) iprover+ lemma atomize_elimL[atomize_elim]: "(\B. (A \ B) \ B) \ Trueprop A" .. subsection \Package setup\ ML_file \Tools/hologic.ML\ ML_file \Tools/rewrite_hol_proof.ML\ setup \Proofterm.set_preproc (Proof_Rewrite_Rules.standard_preproc Rewrite_HOL_Proof.rews)\ subsubsection \Sledgehammer setup\ text \ Theorems blacklisted to Sledgehammer. These theorems typically produce clauses that are prolific (match too many equality or membership literals) and relate to seldom-used facts. Some duplicate other rules. \ named_theorems no_atp "theorems that should be filtered out by Sledgehammer" subsubsection \Classical Reasoner setup\ lemma imp_elim: "P \ Q \ (\ R \ P) \ (Q \ R) \ R" by (rule classical) iprover lemma swap: "\ P \ (\ R \ P) \ R" by (rule classical) iprover lemma thin_refl: "\x = x; PROP W\ \ PROP W" . ML \ structure Hypsubst = Hypsubst ( val dest_eq = HOLogic.dest_eq val dest_Trueprop = HOLogic.dest_Trueprop val dest_imp = HOLogic.dest_imp val eq_reflection = @{thm eq_reflection} val rev_eq_reflection = @{thm meta_eq_to_obj_eq} val imp_intr = @{thm impI} val rev_mp = @{thm rev_mp} val subst = @{thm subst} val sym = @{thm sym} val thin_refl = @{thm thin_refl}; ); open Hypsubst; structure Classical = Classical ( val imp_elim = @{thm imp_elim} val not_elim = @{thm notE} val swap = @{thm swap} val classical = @{thm classical} val sizef = Drule.size_of_thm val hyp_subst_tacs = [Hypsubst.hyp_subst_tac] ); structure Basic_Classical: BASIC_CLASSICAL = Classical; open Basic_Classical; \ setup \ (*prevent substitution on bool*) let fun non_bool_eq (\<^const_name>\HOL.eq\, Type (_, [T, _])) = T <> \<^typ>\bool\ | non_bool_eq _ = false; fun hyp_subst_tac' ctxt = SUBGOAL (fn (goal, i) => if Term.exists_Const non_bool_eq goal then Hypsubst.hyp_subst_tac ctxt i else no_tac); in Context_Rules.addSWrapper (fn ctxt => fn tac => hyp_subst_tac' ctxt ORELSE' tac) end \ declare iffI [intro!] and notI [intro!] and impI [intro!] and disjCI [intro!] and conjI [intro!] and TrueI [intro!] and refl [intro!] declare iffCE [elim!] and FalseE [elim!] and impCE [elim!] and disjE [elim!] and conjE [elim!] declare ex_ex1I [intro!] and allI [intro!] and exI [intro] declare exE [elim!] allE [elim] ML \val HOL_cs = claset_of \<^context>\ lemma contrapos_np: "\ Q \ (\ P \ Q) \ P" by (erule swap) declare ex_ex1I [rule del, intro! 2] and ex1I [intro] declare ext [intro] lemmas [intro?] = ext and [elim?] = ex1_implies_ex text \Better than \ex1E\ for classical reasoner: needs no quantifier duplication!\ lemma alt_ex1E [elim!]: assumes major: "\!x. P x" and minor: "\x. \P x; \y y'. P y \ P y' \ y = y'\ \ R" shows R proof (rule ex1E [OF major minor]) show "\y y'. P y \ P y' \ y = y'" if "P x" and \
: "\y. P y \ y = x" for x using \P x\ \
\
by fast qed assumption text \And again using Uniq\ lemma alt_ex1E': assumes "\!x. P x" "\x. \P x; \\<^sub>\\<^sub>1x. P x\ \ R" shows R using assms unfolding Uniq_def by fast lemma ex1_iff_ex_Uniq: "(\!x. P x) \ (\x. P x) \ (\\<^sub>\\<^sub>1x. P x)" unfolding Uniq_def by fast ML \ structure Blast = Blast ( structure Classical = Classical val Trueprop_const = dest_Const \<^Const>\Trueprop\ val equality_name = \<^const_name>\HOL.eq\ val not_name = \<^const_name>\Not\ val notE = @{thm notE} val ccontr = @{thm ccontr} val hyp_subst_tac = Hypsubst.blast_hyp_subst_tac ); val blast_tac = Blast.blast_tac; \ subsubsection \THE: definite description operator\ lemma the_equality [intro]: assumes "P a" and "\x. P x \ x = a" shows "(THE x. P x) = a" by (blast intro: assms trans [OF arg_cong [where f=The] the_eq_trivial]) lemma theI: assumes "P a" and "\x. P x \ x = a" shows "P (THE x. P x)" by (iprover intro: assms the_equality [THEN ssubst]) lemma theI': "\!x. P x \ P (THE x. P x)" by (blast intro: theI) text \Easier to apply than \theI\: only one occurrence of \P\.\ lemma theI2: assumes "P a" "\x. P x \ x = a" "\x. P x \ Q x" shows "Q (THE x. P x)" by (iprover intro: assms theI) lemma the1I2: assumes "\!x. P x" "\x. P x \ Q x" shows "Q (THE x. P x)" by (iprover intro: assms(2) theI2[where P=P and Q=Q] ex1E[OF assms(1)] elim: allE impE) lemma the1_equality [elim?]: "\\!x. P x; P a\ \ (THE x. P x) = a" by blast lemma the1_equality': "\\\<^sub>\\<^sub>1x. P x; P a\ \ (THE x. P x) = a" unfolding Uniq_def by blast lemma the_sym_eq_trivial: "(THE y. x = y) = x" by blast subsubsection \Simplifier\ lemma eta_contract_eq: "(\s. f s) = f" .. lemma subst_all: \(\x. x = a \ PROP P x) \ PROP P a\ \(\x. a = x \ PROP P x) \ PROP P a\ proof - show \(\x. x = a \ PROP P x) \ PROP P a\ proof (rule equal_intr_rule) assume *: \\x. x = a \ PROP P x\ show \PROP P a\ by (rule *) (rule refl) next fix x assume \PROP P a\ and \x = a\ from \x = a\ have \x \ a\ by (rule eq_reflection) with \PROP P a\ show \PROP P x\ by simp qed show \(\x. a = x \ PROP P x) \ PROP P a\ proof (rule equal_intr_rule) assume *: \\x. a = x \ PROP P x\ show \PROP P a\ by (rule *) (rule refl) next fix x assume \PROP P a\ and \a = x\ from \a = x\ have \a \ x\ by (rule eq_reflection) with \PROP P a\ show \PROP P x\ by simp qed qed lemma simp_thms: shows not_not: "(\ \ P) = P" and Not_eq_iff: "((\ P) = (\ Q)) = (P = Q)" and "(P \ Q) = (P = (\ Q))" "(P \ \P) = True" "(\ P \ P) = True" "(x = x) = True" and not_True_eq_False [code]: "(\ True) = False" and not_False_eq_True [code]: "(\ False) = True" and "(\ P) \ P" "P \ (\ P)" "(True = P) = P" and eq_True: "(P = True) = P" and "(False = P) = (\ P)" and eq_False: "(P = False) = (\ P)" and "(True \ P) = P" "(False \ P) = True" "(P \ True) = True" "(P \ P) = True" "(P \ False) = (\ P)" "(P \ \ P) = (\ P)" "(P \ True) = P" "(True \ P) = P" "(P \ False) = False" "(False \ P) = False" "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" "(P \ \ P) = False" "(\ P \ P) = False" "(P \ True) = True" "(True \ P) = True" "(P \ False) = P" "(False \ P) = P" "(P \ P) = P" "(P \ (P \ Q)) = (P \ Q)" and "(\x. P) = P" "(\x. P) = P" "\x. x = t" "\x. t = x" and "\P. (\x. x = t \ P x) = P t" "\P. (\x. t = x \ P x) = P t" "\P. (\x. x = t \ P x) = P t" "\P. (\x. t = x \ P x) = P t" "(\x. x \ t) = False" "(\x. t \ x) = False" by (blast, blast, blast, blast, blast, iprover+) lemma disj_absorb: "A \ A \ A" by blast lemma disj_left_absorb: "A \ (A \ B) \ A \ B" by blast lemma conj_absorb: "A \ A \ A" by blast lemma conj_left_absorb: "A \ (A \ B) \ A \ B" by blast lemma eq_ac: shows eq_commute: "a = b \ b = a" and iff_left_commute: "(P \ (Q \ R)) \ (Q \ (P \ R))" and iff_assoc: "((P \ Q) \ R) \ (P \ (Q \ R))" by (iprover, blast+) lemma neq_commute: "a \ b \ b \ a" by iprover lemma conj_comms: shows conj_commute: "P \ Q \ Q \ P" and conj_left_commute: "P \ (Q \ R) \ Q \ (P \ R)" by iprover+ lemma conj_assoc: "(P \ Q) \ R \ P \ (Q \ R)" by iprover lemmas conj_ac = conj_commute conj_left_commute conj_assoc lemma disj_comms: shows disj_commute: "P \ Q \ Q \ P" and disj_left_commute: "P \ (Q \ R) \ Q \ (P \ R)" by iprover+ lemma disj_assoc: "(P \ Q) \ R \ P \ (Q \ R)" by iprover lemmas disj_ac = disj_commute disj_left_commute disj_assoc lemma conj_disj_distribL: "P \ (Q \ R) \ P \ Q \ P \ R" by iprover lemma conj_disj_distribR: "(P \ Q) \ R \ P \ R \ Q \ R" by iprover lemma disj_conj_distribL: "P \ (Q \ R) \ (P \ Q) \ (P \ R)" by iprover lemma disj_conj_distribR: "(P \ Q) \ R \ (P \ R) \ (Q \ R)" by iprover lemma imp_conjR: "(P \ (Q \ R)) = ((P \ Q) \ (P \ R))" by iprover lemma imp_conjL: "((P \ Q) \ R) = (P \ (Q \ R))" by iprover lemma imp_disjL: "((P \ Q) \ R) = ((P \ R) \ (Q \ R))" by iprover text \These two are specialized, but \imp_disj_not1\ is useful in \Auth/Yahalom\.\ lemma imp_disj_not1: "(P \ Q \ R) \ (\ Q \ P \ R)" by blast lemma imp_disj_not2: "(P \ Q \ R) \ (\ R \ P \ Q)" by blast lemma imp_disj1: "((P \ Q) \ R) \ (P \ Q \ R)" by blast lemma imp_disj2: "(Q \ (P \ R)) \ (P \ Q \ R)" by blast lemma imp_cong: "(P = P') \ (P' \ (Q = Q')) \ ((P \ Q) \ (P' \ Q'))" by iprover lemma de_Morgan_disj: "\ (P \ Q) \ \ P \ \ Q" by iprover lemma de_Morgan_conj: "\ (P \ Q) \ \ P \ \ Q" by blast lemma not_imp: "\ (P \ Q) \ P \ \ Q" by blast lemma not_iff: "P \ Q \ (P \ \ Q)" by blast lemma disj_not1: "\ P \ Q \ (P \ Q)" by blast lemma disj_not2: "P \ \ Q \ (Q \ P)" by blast \ \changes orientation :-(\ lemma imp_conv_disj: "(P \ Q) \ (\ P) \ Q" by blast lemma disj_imp: "P \ Q \ \ P \ Q" by blast lemma iff_conv_conj_imp: "(P \ Q) \ (P \ Q) \ (Q \ P)" by iprover lemma cases_simp: "(P \ Q) \ (\ P \ Q) \ Q" \ \Avoids duplication of subgoals after \if_split\, when the true and false\ \ \cases boil down to the same thing.\ by blast lemma not_all: "\ (\x. P x) \ (\x. \ P x)" by blast lemma imp_all: "((\x. P x) \ Q) \ (\x. P x \ Q)" by blast lemma not_ex: "\ (\x. P x) \ (\x. \ P x)" by iprover lemma imp_ex: "((\x. P x) \ Q) \ (\x. P x \ Q)" by iprover lemma all_not_ex: "(\x. P x) \ \ (\x. \ P x)" by blast declare All_def [no_atp] lemma ex_disj_distrib: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma all_conj_distrib: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover text \ \<^medskip> The \\\ congruence rule: not included by default! May slow rewrite proofs down by as much as 50\%\ lemma conj_cong: "P = P' \ (P' \ Q = Q') \ (P \ Q) = (P' \ Q')" by iprover lemma rev_conj_cong: "Q = Q' \ (Q' \ P = P') \ (P \ Q) = (P' \ Q')" by iprover text \The \|\ congruence rule: not included by default!\ lemma disj_cong: "P = P' \ (\ P' \ Q = Q') \ (P \ Q) = (P' \ Q')" by blast text \\<^medskip> if-then-else rules\ lemma if_True [code]: "(if True then x else y) = x" unfolding If_def by blast lemma if_False [code]: "(if False then x else y) = y" unfolding If_def by blast lemma if_P: "P \ (if P then x else y) = x" unfolding If_def by blast lemma if_not_P: "\ P \ (if P then x else y) = y" unfolding If_def by blast lemma if_split: "P (if Q then x else y) = ((Q \ P x) \ (\ Q \ P y))" proof (rule case_split [of Q]) show ?thesis if Q using that by (simplesubst if_P) blast+ show ?thesis if "\ Q" using that by (simplesubst if_not_P) blast+ qed lemma if_split_asm: "P (if Q then x else y) = (\ ((Q \ \ P x) \ (\ Q \ \ P y)))" by (simplesubst if_split) blast lemmas if_splits [no_atp] = if_split if_split_asm lemma if_cancel: "(if c then x else x) = x" by (simplesubst if_split) blast lemma if_eq_cancel: "(if x = y then y else x) = x" by (simplesubst if_split) blast lemma if_bool_eq_conj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" \ \This form is useful for expanding \if\s on the RIGHT of the \\\ symbol.\ by (rule if_split) lemma if_bool_eq_disj: "(if P then Q else R) = ((P \ Q) \ (\ P \ R))" \ \And this form is useful for expanding \if\s on the LEFT.\ by (simplesubst if_split) blast lemma Eq_TrueI: "P \ P \ True" unfolding atomize_eq by iprover lemma Eq_FalseI: "\ P \ P \ False" unfolding atomize_eq by iprover text \\<^medskip> let rules for simproc\ lemma Let_folded: "f x \ g x \ Let x f \ Let x g" by (unfold Let_def) lemma Let_unfold: "f x \ g \ Let x f \ g" by (unfold Let_def) text \ The following copy of the implication operator is useful for fine-tuning congruence rules. It instructs the simplifier to simplify its premise. \ definition simp_implies :: "prop \ prop \ prop" (infixr "=simp=>" 1) where "simp_implies \ (\)" lemma simp_impliesI: assumes PQ: "(PROP P \ PROP Q)" shows "PROP P =simp=> PROP Q" unfolding simp_implies_def by (iprover intro: PQ) lemma simp_impliesE: assumes PQ: "PROP P =simp=> PROP Q" and P: "PROP P" and QR: "PROP Q \ PROP R" shows "PROP R" by (iprover intro: QR P PQ [unfolded simp_implies_def]) lemma simp_implies_cong: assumes PP' :"PROP P \ PROP P'" and P'QQ': "PROP P' \ (PROP Q \ PROP Q')" shows "(PROP P =simp=> PROP Q) \ (PROP P' =simp=> PROP Q')" unfolding simp_implies_def proof (rule equal_intr_rule) assume PQ: "PROP P \ PROP Q" and P': "PROP P'" from PP' [symmetric] and P' have "PROP P" by (rule equal_elim_rule1) then have "PROP Q" by (rule PQ) with P'QQ' [OF P'] show "PROP Q'" by (rule equal_elim_rule1) next assume P'Q': "PROP P' \ PROP Q'" and P: "PROP P" from PP' and P have P': "PROP P'" by (rule equal_elim_rule1) then have "PROP Q'" by (rule P'Q') with P'QQ' [OF P', symmetric] show "PROP Q" by (rule equal_elim_rule1) qed lemma uncurry: assumes "P \ Q \ R" shows "P \ Q \ R" using assms by blast lemma iff_allI: assumes "\x. P x = Q x" shows "(\x. P x) = (\x. Q x)" using assms by blast lemma iff_exI: assumes "\x. P x = Q x" shows "(\x. P x) = (\x. Q x)" using assms by blast lemma all_comm: "(\x y. P x y) = (\y x. P x y)" by blast lemma ex_comm: "(\x y. P x y) = (\y x. P x y)" by blast ML_file \Tools/simpdata.ML\ ML \open Simpdata\ setup \ map_theory_simpset (put_simpset HOL_basic_ss) #> Simplifier.method_setup Splitter.split_modifiers \ simproc_setup defined_Ex ("\x. P x") = \K Quantifier1.rearrange_Ex\ simproc_setup defined_All ("\x. P x") = \K Quantifier1.rearrange_All\ simproc_setup defined_all("\x. PROP P x") = \K Quantifier1.rearrange_all\ text \Simproc for proving \(y = x) \ False\ from premise \\ (x = y)\:\ -simproc_setup neq ("x = y") = \fn _ => +simproc_setup neq ("x = y") = \ let val neq_to_EQ_False = @{thm not_sym} RS @{thm Eq_FalseI}; fun is_neq eq lhs rhs thm = (case Thm.prop_of thm of _ $ (Not $ (eq' $ l' $ r')) => Not = HOLogic.Not andalso eq' = eq andalso r' aconv lhs andalso l' aconv rhs | _ => false); fun proc ss ct = (case Thm.term_of ct of eq $ lhs $ rhs => (case find_first (is_neq eq lhs rhs) (Simplifier.prems_of ss) of SOME thm => SOME (thm RS neq_to_EQ_False) | NONE => NONE) | _ => NONE); - in proc end + in K proc end \ simproc_setup let_simp ("Let x f") = \ let fun count_loose (Bound i) k = if i >= k then 1 else 0 | count_loose (s $ t) k = count_loose s k + count_loose t k | count_loose (Abs (_, _, t)) k = count_loose t (k + 1) | count_loose _ _ = 0; fun is_trivial_let (Const (\<^const_name>\Let\, _) $ x $ t) = (case t of Abs (_, _, t') => count_loose t' 0 <= 1 | _ => true); in - fn _ => fn ctxt => fn ct => + K (fn ctxt => fn ct => if is_trivial_let (Thm.term_of ct) then SOME @{thm Let_def} (*no or one ocurrence of bound variable*) else let (*Norbert Schirmer's case*) val t = Thm.term_of ct; val (t', ctxt') = yield_singleton (Variable.import_terms false) t ctxt; in Option.map (hd o Variable.export ctxt' ctxt o single) (case t' of Const (\<^const_name>\Let\,_) $ x $ f => (* x and f are already in normal form *) if is_Free x orelse is_Bound x orelse is_Const x then SOME @{thm Let_def} else let val n = case f of (Abs (x, _, _)) => x | _ => "x"; val cx = Thm.cterm_of ctxt x; val xT = Thm.typ_of_cterm cx; val cf = Thm.cterm_of ctxt f; val fx_g = Simplifier.rewrite ctxt (Thm.apply cf cx); val (_ $ _ $ g) = Thm.prop_of fx_g; val g' = abstract_over (x, g); val abs_g'= Abs (n, xT, g'); in if g aconv g' then let val rl = infer_instantiate ctxt [(("f", 0), cf), (("x", 0), cx)] @{thm Let_unfold}; in SOME (rl OF [fx_g]) end else if (Envir.beta_eta_contract f) aconv (Envir.beta_eta_contract abs_g') then NONE (*avoid identity conversion*) else let val g'x = abs_g' $ x; val g_g'x = Thm.symmetric (Thm.beta_conversion false (Thm.cterm_of ctxt g'x)); val rl = @{thm Let_folded} |> infer_instantiate ctxt [(("f", 0), Thm.cterm_of ctxt f), (("x", 0), cx), (("g", 0), Thm.cterm_of ctxt abs_g')]; in SOME (rl OF [Thm.transitive fx_g g_g'x]) end end | _ => NONE) - end + end) end \ lemma True_implies_equals: "(True \ PROP P) \ PROP P" proof assume "True \ PROP P" from this [OF TrueI] show "PROP P" . next assume "PROP P" then show "PROP P" . qed lemma implies_True_equals: "(PROP P \ True) \ Trueprop True" by standard (intro TrueI) lemma False_implies_equals: "(False \ P) \ Trueprop True" by standard simp_all (* It seems that making this a simp rule is slower than using the simproc below *) lemma implies_False_swap: "(False \ PROP P \ PROP Q) \ (PROP P \ False \ PROP Q)" by (rule swap_prems_eq) ML \ fun eliminate_false_implies ct = let val (prems, concl) = Logic.strip_horn (Thm.term_of ct) fun go n = if n > 1 then Conv.rewr_conv @{thm Pure.swap_prems_eq} then_conv Conv.arg_conv (go (n - 1)) then_conv Conv.rewr_conv @{thm HOL.implies_True_equals} else Conv.rewr_conv @{thm HOL.False_implies_equals} in case concl of Const (@{const_name HOL.Trueprop}, _) $ _ => SOME (go (length prems) ct) | _ => NONE end \ simproc_setup eliminate_false_implies ("False \ PROP P") = \K (K eliminate_false_implies)\ lemma ex_simps: "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" \ \Miniscoping: pushing in existential quantifiers.\ by (iprover | blast)+ lemma all_simps: "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" "\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" "\P Q. (\x. P \ Q x) = (P \ (\x. Q x))" \ \Miniscoping: pushing in universal quantifiers.\ by (iprover | blast)+ lemmas [simp] = triv_forall_equality \ \prunes params\ True_implies_equals implies_True_equals \ \prune \True\ in asms\ False_implies_equals \ \prune \False\ in asms\ if_True if_False if_cancel if_eq_cancel imp_disjL \ \In general it seems wrong to add distributive laws by default: they might cause exponential blow-up. But \imp_disjL\ has been in for a while and cannot be removed without affecting existing proofs. Moreover, rewriting by \(P \ Q \ R) = ((P \ R) \ (Q \ R))\ might be justified on the grounds that it allows simplification of \R\ in the two cases.\ conj_assoc disj_assoc de_Morgan_conj de_Morgan_disj imp_disj1 imp_disj2 not_imp disj_not1 not_all not_ex cases_simp the_eq_trivial the_sym_eq_trivial ex_simps all_simps simp_thms subst_all lemmas [cong] = imp_cong simp_implies_cong lemmas [split] = if_split ML \val HOL_ss = simpset_of \<^context>\ text \Simplifies \x\ assuming \c\ and \y\ assuming \\ c\.\ lemma if_cong: assumes "b = c" and "c \ x = u" and "\ c \ y = v" shows "(if b then x else y) = (if c then u else v)" using assms by simp text \Prevents simplification of \x\ and \y\: faster and allows the execution of functional programs.\ lemma if_weak_cong [cong]: assumes "b = c" shows "(if b then x else y) = (if c then x else y)" using assms by (rule arg_cong) text \Prevents simplification of t: much faster\ lemma let_weak_cong: assumes "a = b" shows "(let x = a in t x) = (let x = b in t x)" using assms by (rule arg_cong) text \To tidy up the result of a simproc. Only the RHS will be simplified.\ lemma eq_cong2: assumes "u = u'" shows "(t \ u) \ (t \ u')" using assms by simp lemma if_distrib: "f (if c then x else y) = (if c then f x else f y)" by simp lemma if_distribR: "(if b then f else g) x = (if b then f x else g x)" by simp lemma all_if_distrib: "(\x. if x = a then P x else Q x) \ P a \ (\x. x\a \ Q x)" by auto lemma ex_if_distrib: "(\x. if x = a then P x else Q x) \ P a \ (\x. x\a \ Q x)" by auto lemma if_if_eq_conj: "(if P then if Q then x else y else y) = (if P \ Q then x else y)" by simp text \As a simplification rule, it replaces all function equalities by first-order equalities.\ lemma fun_eq_iff: "f = g \ (\x. f x = g x)" by auto subsubsection \Generic cases and induction\ text \Rule projections:\ ML \ structure Project_Rule = Project_Rule ( val conjunct1 = @{thm conjunct1} val conjunct2 = @{thm conjunct2} val mp = @{thm mp} ); \ context begin qualified definition "induct_forall P \ \x. P x" qualified definition "induct_implies A B \ A \ B" qualified definition "induct_equal x y \ x = y" qualified definition "induct_conj A B \ A \ B" qualified definition "induct_true \ True" qualified definition "induct_false \ False" lemma induct_forall_eq: "(\x. P x) \ Trueprop (induct_forall (\x. P x))" by (unfold atomize_all induct_forall_def) lemma induct_implies_eq: "(A \ B) \ Trueprop (induct_implies A B)" by (unfold atomize_imp induct_implies_def) lemma induct_equal_eq: "(x \ y) \ Trueprop (induct_equal x y)" by (unfold atomize_eq induct_equal_def) lemma induct_conj_eq: "(A &&& B) \ Trueprop (induct_conj A B)" by (unfold atomize_conj induct_conj_def) lemmas induct_atomize' = induct_forall_eq induct_implies_eq induct_conj_eq lemmas induct_atomize = induct_atomize' induct_equal_eq lemmas induct_rulify' [symmetric] = induct_atomize' lemmas induct_rulify [symmetric] = induct_atomize lemmas induct_rulify_fallback = induct_forall_def induct_implies_def induct_equal_def induct_conj_def induct_true_def induct_false_def lemma induct_forall_conj: "induct_forall (\x. induct_conj (A x) (B x)) = induct_conj (induct_forall A) (induct_forall B)" by (unfold induct_forall_def induct_conj_def) iprover lemma induct_implies_conj: "induct_implies C (induct_conj A B) = induct_conj (induct_implies C A) (induct_implies C B)" by (unfold induct_implies_def induct_conj_def) iprover lemma induct_conj_curry: "(induct_conj A B \ PROP C) \ (A \ B \ PROP C)" proof assume r: "induct_conj A B \ PROP C" assume ab: A B show "PROP C" by (rule r) (simp add: induct_conj_def ab) next assume r: "A \ B \ PROP C" assume ab: "induct_conj A B" show "PROP C" by (rule r) (simp_all add: ab [unfolded induct_conj_def]) qed lemmas induct_conj = induct_forall_conj induct_implies_conj induct_conj_curry lemma induct_trueI: "induct_true" by (simp add: induct_true_def) text \Method setup.\ ML_file \~~/src/Tools/induct.ML\ ML \ structure Induct = Induct ( val cases_default = @{thm case_split} val atomize = @{thms induct_atomize} val rulify = @{thms induct_rulify'} val rulify_fallback = @{thms induct_rulify_fallback} val equal_def = @{thm induct_equal_def} fun dest_def (Const (\<^const_name>\induct_equal\, _) $ t $ u) = SOME (t, u) | dest_def _ = NONE fun trivial_tac ctxt = match_tac ctxt @{thms induct_trueI} ) \ ML_file \~~/src/Tools/induction.ML\ declaration \ fn _ => Induct.map_simpset (fn ss => ss addsimprocs [Simplifier.make_simproc \<^context> "swap_induct_false" {lhss = [\<^term>\induct_false \ PROP P \ PROP Q\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (P as _ $ \<^Const_>\induct_false\) $ (_ $ Q $ _) => if P <> Q then SOME Drule.swap_prems_eq else NONE | _ => NONE)}, Simplifier.make_simproc \<^context> "induct_equal_conj_curry" {lhss = [\<^term>\induct_conj P Q \ PROP R\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of _ $ (_ $ P) $ _ => let fun is_conj \<^Const_>\induct_conj for P Q\ = is_conj P andalso is_conj Q | is_conj \<^Const_>\induct_equal _ for _ _\ = true | is_conj \<^Const_>\induct_true\ = true | is_conj \<^Const_>\induct_false\ = true | is_conj _ = false in if is_conj P then SOME @{thm induct_conj_curry} else NONE end | _ => NONE)}] |> Simplifier.set_mksimps (fn ctxt => Simpdata.mksimps Simpdata.mksimps_pairs ctxt #> map (rewrite_rule ctxt (map Thm.symmetric @{thms induct_rulify_fallback})))) \ text \Pre-simplification of induction and cases rules\ lemma [induct_simp]: "(\x. induct_equal x t \ PROP P x) \ PROP P t" unfolding induct_equal_def proof assume r: "\x. x = t \ PROP P x" show "PROP P t" by (rule r [OF refl]) next fix x assume "PROP P t" "x = t" then show "PROP P x" by simp qed lemma [induct_simp]: "(\x. induct_equal t x \ PROP P x) \ PROP P t" unfolding induct_equal_def proof assume r: "\x. t = x \ PROP P x" show "PROP P t" by (rule r [OF refl]) next fix x assume "PROP P t" "t = x" then show "PROP P x" by simp qed lemma [induct_simp]: "(induct_false \ P) \ Trueprop induct_true" unfolding induct_false_def induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "(induct_true \ PROP P) \ PROP P" unfolding induct_true_def proof assume "True \ PROP P" then show "PROP P" using TrueI . next assume "PROP P" then show "PROP P" . qed lemma [induct_simp]: "(PROP P \ induct_true) \ Trueprop induct_true" unfolding induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "(\x::'a::{}. induct_true) \ Trueprop induct_true" unfolding induct_true_def by (iprover intro: equal_intr_rule) lemma [induct_simp]: "induct_implies induct_true P \ P" by (simp add: induct_implies_def induct_true_def) lemma [induct_simp]: "x = x \ True" by (rule simp_thms) end ML_file \~~/src/Tools/induct_tacs.ML\ subsubsection \Coherent logic\ ML_file \~~/src/Tools/coherent.ML\ ML \ structure Coherent = Coherent ( val atomize_elimL = @{thm atomize_elimL}; val atomize_exL = @{thm atomize_exL}; val atomize_conjL = @{thm atomize_conjL}; val atomize_disjL = @{thm atomize_disjL}; val operator_names = [\<^const_name>\HOL.disj\, \<^const_name>\HOL.conj\, \<^const_name>\Ex\]; ); \ subsubsection \Reorienting equalities\ ML \ signature REORIENT_PROC = sig val add : (term -> bool) -> theory -> theory - val proc : morphism -> Proof.context -> cterm -> thm option + val proc : Proof.context -> cterm -> thm option end; structure Reorient_Proc : REORIENT_PROC = struct structure Data = Theory_Data ( type T = ((term -> bool) * stamp) list; val empty = []; fun merge data : T = Library.merge (eq_snd (op =)) data; ); fun add m = Data.map (cons (m, stamp ())); fun matches thy t = exists (fn (m, _) => m t) (Data.get thy); val meta_reorient = @{thm eq_commute [THEN eq_reflection]}; - fun proc phi ctxt ct = + fun proc ctxt ct = let val thy = Proof_Context.theory_of ctxt; in case Thm.term_of ct of (_ $ t $ u) => if matches thy u then NONE else SOME meta_reorient | _ => NONE end; end; \ subsection \Other simple lemmas and lemma duplicates\ lemma eq_iff_swap: "(x = y \ P) \ (y = x \ P)" by blast lemma all_cong1: "(\x. P x = P' x) \ (\x. P x) = (\x. P' x)" by auto lemma ex_cong1: "(\x. P x = P' x) \ (\x. P x) = (\x. P' x)" by auto lemma all_cong: "(\x. Q x \ P x = P' x) \ (\x. Q x \ P x) = (\x. Q x \ P' x)" by auto lemma ex_cong: "(\x. Q x \ P x = P' x) \ (\x. Q x \ P x) = (\x. Q x \ P' x)" by auto lemma ex1_eq [iff]: "\!x. x = t" "\!x. t = x" by blast+ lemma choice_eq: "(\x. \!y. P x y) = (\!f. \x. P x (f x))" (is "?lhs = ?rhs") proof (intro iffI allI) assume L: ?lhs then have \
: "\x. P x (THE y. P x y)" by (best intro: theI') show ?rhs by (rule ex1I) (use L \
in \fast+\) next fix x assume R: ?rhs then obtain f where f: "\x. P x (f x)" and f1: "\y. (\x. P x (y x)) \ y = f" by (blast elim: ex1E) show "\!y. P x y" proof (rule ex1I) show "P x (f x)" using f by blast show "y = f x" if "P x y" for y proof - have "P z (if z = x then y else f z)" for z using f that by (auto split: if_split) with f1 [of "\z. if z = x then y else f z"] f show ?thesis by (auto simp add: split: if_split_asm dest: fun_cong) qed qed qed lemmas eq_sym_conv = eq_commute lemma nnf_simps: "(\ (P \ Q)) = (\ P \ \ Q)" "(\ (P \ Q)) = (\ P \ \ Q)" "(P \ Q) = (\ P \ Q)" "(P = Q) = ((P \ Q) \ (\ P \ \ Q))" "(\ (P = Q)) = ((P \ \ Q) \ (\ P \ Q))" "(\ \ P) = P" by blast+ subsection \Basic ML bindings\ ML \ val FalseE = @{thm FalseE} val Let_def = @{thm Let_def} val TrueI = @{thm TrueI} val allE = @{thm allE} val allI = @{thm allI} val all_dupE = @{thm all_dupE} val arg_cong = @{thm arg_cong} val box_equals = @{thm box_equals} val ccontr = @{thm ccontr} val classical = @{thm classical} val conjE = @{thm conjE} val conjI = @{thm conjI} val conjunct1 = @{thm conjunct1} val conjunct2 = @{thm conjunct2} val disjCI = @{thm disjCI} val disjE = @{thm disjE} val disjI1 = @{thm disjI1} val disjI2 = @{thm disjI2} val eq_reflection = @{thm eq_reflection} val ex1E = @{thm ex1E} val ex1I = @{thm ex1I} val ex1_implies_ex = @{thm ex1_implies_ex} val exE = @{thm exE} val exI = @{thm exI} val excluded_middle = @{thm excluded_middle} val ext = @{thm ext} val fun_cong = @{thm fun_cong} val iffD1 = @{thm iffD1} val iffD2 = @{thm iffD2} val iffI = @{thm iffI} val impE = @{thm impE} val impI = @{thm impI} val meta_eq_to_obj_eq = @{thm meta_eq_to_obj_eq} val mp = @{thm mp} val notE = @{thm notE} val notI = @{thm notI} val not_all = @{thm not_all} val not_ex = @{thm not_ex} val not_iff = @{thm not_iff} val not_not = @{thm not_not} val not_sym = @{thm not_sym} val refl = @{thm refl} val rev_mp = @{thm rev_mp} val spec = @{thm spec} val ssubst = @{thm ssubst} val subst = @{thm subst} val sym = @{thm sym} val trans = @{thm trans} \ locale cnf begin lemma clause2raw_notE: "\P; \P\ \ False" by auto lemma clause2raw_not_disj: "\\ P; \ Q\ \ \ (P \ Q)" by auto lemma clause2raw_not_not: "P \ \\ P" by auto lemma iff_refl: "(P::bool) = P" by auto lemma iff_trans: "[| (P::bool) = Q; Q = R |] ==> P = R" by auto lemma conj_cong: "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma disj_cong: "[| P = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma make_nnf_imp: "[| (\P) = P'; Q = Q' |] ==> (P \ Q) = (P' \ Q')" by auto lemma make_nnf_iff: "[| P = P'; (\P) = NP; Q = Q'; (\Q) = NQ |] ==> (P = Q) = ((P' \ NQ) \ (NP \ Q'))" by auto lemma make_nnf_not_false: "(\False) = True" by auto lemma make_nnf_not_true: "(\True) = False" by auto lemma make_nnf_not_conj: "[| (\P) = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_disj: "[| (\P) = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_imp: "[| P = P'; (\Q) = Q' |] ==> (\(P \ Q)) = (P' \ Q')" by auto lemma make_nnf_not_iff: "[| P = P'; (\P) = NP; Q = Q'; (\Q) = NQ |] ==> (\(P = Q)) = ((P' \ Q') \ (NP \ NQ))" by auto lemma make_nnf_not_not: "P = P' ==> (\\P) = P'" by auto lemma simp_TF_conj_True_l: "[| P = True; Q = Q' |] ==> (P \ Q) = Q'" by auto lemma simp_TF_conj_True_r: "[| P = P'; Q = True |] ==> (P \ Q) = P'" by auto lemma simp_TF_conj_False_l: "P = False ==> (P \ Q) = False" by auto lemma simp_TF_conj_False_r: "Q = False ==> (P \ Q) = False" by auto lemma simp_TF_disj_True_l: "P = True ==> (P \ Q) = True" by auto lemma simp_TF_disj_True_r: "Q = True ==> (P \ Q) = True" by auto lemma simp_TF_disj_False_l: "[| P = False; Q = Q' |] ==> (P \ Q) = Q'" by auto lemma simp_TF_disj_False_r: "[| P = P'; Q = False |] ==> (P \ Q) = P'" by auto lemma make_cnf_disj_conj_l: "[| (P \ R) = PR; (Q \ R) = QR |] ==> ((P \ Q) \ R) = (PR \ QR)" by auto lemma make_cnf_disj_conj_r: "[| (P \ Q) = PQ; (P \ R) = PR |] ==> (P \ (Q \ R)) = (PQ \ PR)" by auto lemma make_cnfx_disj_ex_l: "((\(x::bool). P x) \ Q) = (\x. P x \ Q)" by auto lemma make_cnfx_disj_ex_r: "(P \ (\(x::bool). Q x)) = (\x. P \ Q x)" by auto lemma make_cnfx_newlit: "(P \ Q) = (\x. (P \ x) \ (Q \ \x))" by auto lemma make_cnfx_ex_cong: "(\(x::bool). P x = Q x) \ (\x. P x) = (\x. Q x)" by auto lemma weakening_thm: "[| P; Q |] ==> Q" by auto lemma cnftac_eq_imp: "[| P = Q; P |] ==> Q" by auto end ML_file \Tools/cnf.ML\ section \\NO_MATCH\ simproc\ text \ The simplification procedure can be used to avoid simplification of terms of a certain form. \ definition NO_MATCH :: "'a \ 'b \ bool" where "NO_MATCH pat val \ True" lemma NO_MATCH_cong[cong]: "NO_MATCH pat val = NO_MATCH pat val" by (rule refl) declare [[coercion_args NO_MATCH - -]] -simproc_setup NO_MATCH ("NO_MATCH pat val") = \fn _ => fn ctxt => fn ct => +simproc_setup NO_MATCH ("NO_MATCH pat val") = \K (fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt val dest_binop = Term.dest_comb #> apfst (Term.dest_comb #> snd) val m = Pattern.matches thy (dest_binop (Thm.term_of ct)) - in if m then NONE else SOME @{thm NO_MATCH_def} end + in if m then NONE else SOME @{thm NO_MATCH_def} end) \ text \ This setup ensures that a rewrite rule of the form \<^term>\NO_MATCH pat val \ t\ is only applied, if the pattern \pat\ does not match the value \val\. \ text\ Tagging a premise of a simp rule with ASSUMPTION forces the simplifier not to simplify the argument and to solve it by an assumption. \ definition ASSUMPTION :: "bool \ bool" where "ASSUMPTION A \ A" lemma ASSUMPTION_cong[cong]: "ASSUMPTION A = ASSUMPTION A" by (rule refl) lemma ASSUMPTION_I: "A \ ASSUMPTION A" by (simp add: ASSUMPTION_def) lemma ASSUMPTION_D: "ASSUMPTION A \ A" by (simp add: ASSUMPTION_def) setup \ let val asm_sol = mk_solver "ASSUMPTION" (fn ctxt => resolve_tac ctxt [@{thm ASSUMPTION_I}] THEN' resolve_tac ctxt (Simplifier.prems_of ctxt)) in map_theory_simpset (fn ctxt => Simplifier.addSolver (ctxt,asm_sol)) end \ subsection \Code generator setup\ subsubsection \Generic code generator preprocessor setup\ lemma conj_left_cong: "P \ Q \ P \ R \ Q \ R" by (fact arg_cong) lemma disj_left_cong: "P \ Q \ P \ R \ Q \ R" by (fact arg_cong) setup \ Code_Preproc.map_pre (put_simpset HOL_basic_ss) #> Code_Preproc.map_post (put_simpset HOL_basic_ss) #> Code_Simp.map_ss (put_simpset HOL_basic_ss #> Simplifier.add_cong @{thm conj_left_cong} #> Simplifier.add_cong @{thm disj_left_cong}) \ subsubsection \Equality\ class equal = fixes equal :: "'a \ 'a \ bool" assumes equal_eq: "equal x y \ x = y" begin lemma equal: "equal = (=)" by (rule ext equal_eq)+ lemma equal_refl: "equal x x \ True" unfolding equal by (rule iffI TrueI refl)+ lemma eq_equal: "(=) \ equal" by (rule eq_reflection) (rule ext, rule ext, rule sym, rule equal_eq) end declare eq_equal [symmetric, code_post] declare eq_equal [code] setup \ Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs [Simplifier.make_simproc \<^context> "equal" {lhss = [\<^term>\HOL.eq\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (_, Type (\<^type_name>\fun\, [Type _, _])) => SOME @{thm eq_equal} | _ => NONE)}]) \ subsubsection \Generic code generator foundation\ text \Datatype \<^typ>\bool\\ code_datatype True False lemma [code]: shows "False \ P \ False" and "True \ P \ P" and "P \ False \ False" and "P \ True \ P" by simp_all lemma [code]: shows "False \ P \ P" and "True \ P \ True" and "P \ False \ P" and "P \ True \ True" by simp_all lemma [code]: shows "(False \ P) \ True" and "(True \ P) \ P" and "(P \ False) \ \ P" and "(P \ True) \ True" by simp_all text \More about \<^typ>\prop\\ lemma [code nbe]: shows "(True \ PROP Q) \ PROP Q" and "(PROP Q \ True) \ Trueprop True" and "(P \ R) \ Trueprop (P \ R)" by (auto intro!: equal_intr_rule) lemma Trueprop_code [code]: "Trueprop True \ Code_Generator.holds" by (auto intro!: equal_intr_rule holds) declare Trueprop_code [symmetric, code_post] text \Equality\ declare simp_thms(6) [code nbe] instantiation itself :: (type) equal begin definition equal_itself :: "'a itself \ 'a itself \ bool" where "equal_itself x y \ x = y" instance by standard (fact equal_itself_def) end lemma equal_itself_code [code]: "equal TYPE('a) TYPE('a) \ True" by (simp add: equal) setup \Sign.add_const_constraint (\<^const_name>\equal\, SOME \<^typ>\'a::type \ 'a \ bool\)\ lemma equal_alias_cert: "OFCLASS('a, equal_class) \ (((=) :: 'a \ 'a \ bool) \ equal)" (is "?ofclass \ ?equal") proof assume "PROP ?ofclass" show "PROP ?equal" by (tactic \ALLGOALS (resolve_tac \<^context> [Thm.unconstrainT @{thm eq_equal}])\) (fact \PROP ?ofclass\) next assume "PROP ?equal" show "PROP ?ofclass" proof qed (simp add: \PROP ?equal\) qed setup \Sign.add_const_constraint (\<^const_name>\equal\, SOME \<^typ>\'a::equal \ 'a \ bool\)\ setup \Nbe.add_const_alias @{thm equal_alias_cert}\ text \Cases\ lemma Let_case_cert: assumes "CASE \ (\x. Let x f)" shows "CASE x \ f x" using assms by simp_all setup \ Code.declare_case_global @{thm Let_case_cert} #> Code.declare_undefined_global \<^const_name>\undefined\ \ declare [[code abort: undefined]] subsubsection \Generic code generator target languages\ text \type \<^typ>\bool\\ code_printing type_constructor bool \ (SML) "bool" and (OCaml) "bool" and (Haskell) "Bool" and (Scala) "Boolean" | constant True \ (SML) "true" and (OCaml) "true" and (Haskell) "True" and (Scala) "true" | constant False \ (SML) "false" and (OCaml) "false" and (Haskell) "False" and (Scala) "false" code_reserved SML bool true false code_reserved OCaml bool code_reserved Scala Boolean code_printing constant Not \ (SML) "not" and (OCaml) "not" and (Haskell) "not" and (Scala) "'! _" | constant HOL.conj \ (SML) infixl 1 "andalso" and (OCaml) infixl 3 "&&" and (Haskell) infixr 3 "&&" and (Scala) infixl 3 "&&" | constant HOL.disj \ (SML) infixl 0 "orelse" and (OCaml) infixl 2 "||" and (Haskell) infixl 2 "||" and (Scala) infixl 1 "||" | constant HOL.implies \ (SML) "!(if (_)/ then (_)/ else true)" and (OCaml) "!(if (_)/ then (_)/ else true)" and (Haskell) "!(if (_)/ then (_)/ else True)" and (Scala) "!(if ((_))/ (_)/ else true)" | constant If \ (SML) "!(if (_)/ then (_)/ else (_))" and (OCaml) "!(if (_)/ then (_)/ else (_))" and (Haskell) "!(if (_)/ then (_)/ else (_))" and (Scala) "!(if ((_))/ (_)/ else (_))" code_reserved SML not code_reserved OCaml not code_identifier code_module Pure \ (SML) HOL and (OCaml) HOL and (Haskell) HOL and (Scala) HOL text \Using built-in Haskell equality.\ code_printing type_class equal \ (Haskell) "Eq" | constant HOL.equal \ (Haskell) infix 4 "==" | constant HOL.eq \ (Haskell) infix 4 "==" text \\undefined\\ code_printing constant undefined \ (SML) "!(raise/ Fail/ \"undefined\")" and (OCaml) "failwith/ \"undefined\"" and (Haskell) "error/ \"undefined\"" and (Scala) "!sys.error(\"undefined\")" subsubsection \Evaluation and normalization by evaluation\ method_setup eval = \ let fun eval_tac ctxt = let val conv = Code_Runtime.dynamic_holds_conv in CONVERSION (Conv.params_conv ~1 (Conv.concl_conv ~1 o conv) ctxt) THEN' resolve_tac ctxt [TrueI] end in Scan.succeed (SIMPLE_METHOD' o eval_tac) end \ "solve goal by evaluation" method_setup normalization = \ Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o (CONVERSION (Nbe.dynamic_conv ctxt) THEN_ALL_NEW (TRY o resolve_tac ctxt [TrueI])))) \ "solve goal by normalization" subsection \Counterexample Search Units\ subsubsection \Quickcheck\ quickcheck_params [size = 5, iterations = 50] subsubsection \Nitpick setup\ named_theorems nitpick_unfold "alternative definitions of constants as needed by Nitpick" and nitpick_simp "equational specification of constants as needed by Nitpick" and nitpick_psimp "partial equational specification of constants as needed by Nitpick" and nitpick_choice_spec "choice specification of constants as needed by Nitpick" declare if_bool_eq_conj [nitpick_unfold, no_atp] and if_bool_eq_disj [no_atp] subsection \Preprocessing for the predicate compiler\ named_theorems code_pred_def "alternative definitions of constants for the Predicate Compiler" and code_pred_inline "inlining definitions for the Predicate Compiler" and code_pred_simp "simplification rules for the optimisations in the Predicate Compiler" subsection \Legacy tactics and ML bindings\ ML \ (* combination of (spec RS spec RS ...(j times) ... spec RS mp) *) local fun wrong_prem (Const (\<^const_name>\All\, _) $ Abs (_, _, t)) = wrong_prem t | wrong_prem (Bound _) = true | wrong_prem _ = false; val filter_right = filter (not o wrong_prem o HOLogic.dest_Trueprop o hd o Thm.prems_of); fun smp i = funpow i (fn m => filter_right ([spec] RL m)) [mp]; in fun smp_tac ctxt j = EVERY' [dresolve_tac ctxt (smp j), assume_tac ctxt]; end; local val nnf_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms nnf_simps}); in fun nnf_conv ctxt = Simplifier.rewrite (put_simpset nnf_ss ctxt); end \ hide_const (open) eq equal end diff --git a/src/HOL/HOLCF/Cfun.thy b/src/HOL/HOLCF/Cfun.thy --- a/src/HOL/HOLCF/Cfun.thy +++ b/src/HOL/HOLCF/Cfun.thy @@ -1,523 +1,523 @@ (* Title: HOL/HOLCF/Cfun.thy Author: Franz Regensburger Author: Brian Huffman *) section \The type of continuous functions\ theory Cfun imports Cpodef Fun_Cpo Product_Cpo begin default_sort cpo subsection \Definition of continuous function type\ definition "cfun = {f::'a \ 'b. cont f}" cpodef ('a, 'b) cfun ("(_ \/ _)" [1, 0] 0) = "cfun :: ('a \ 'b) set" by (auto simp: cfun_def intro: cont_const adm_cont) type_notation (ASCII) cfun (infixr "->" 0) notation (ASCII) Rep_cfun ("(_$/_)" [999,1000] 999) notation Rep_cfun ("(_\/_)" [999,1000] 999) subsection \Syntax for continuous lambda abstraction\ syntax "_cabs" :: "[logic, logic] \ logic" parse_translation \ (* rewrite (_cabs x t) => (Abs_cfun (%x. t)) *) [Syntax_Trans.mk_binder_tr (\<^syntax_const>\_cabs\, \<^const_syntax>\Abs_cfun\)] \ print_translation \ [(\<^const_syntax>\Abs_cfun\, fn _ => fn [Abs abs] => let val (x, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_cabs\ $ x $ t end)] \ \ \To avoid eta-contraction of body\ text \Syntax for nested abstractions\ syntax (ASCII) "_Lambda" :: "[cargs, logic] \ logic" ("(3LAM _./ _)" [1000, 10] 10) syntax "_Lambda" :: "[cargs, logic] \ logic" ("(3\ _./ _)" [1000, 10] 10) parse_ast_translation \ (* rewrite (LAM x y z. t) => (_cabs x (_cabs y (_cabs z t))) *) (* cf. Syntax.lambda_ast_tr from src/Pure/Syntax/syn_trans.ML *) let fun Lambda_ast_tr [pats, body] = Ast.fold_ast_p \<^syntax_const>\_cabs\ (Ast.unfold_ast \<^syntax_const>\_cargs\ (Ast.strip_positions pats), body) | Lambda_ast_tr asts = raise Ast.AST ("Lambda_ast_tr", asts); in [(\<^syntax_const>\_Lambda\, K Lambda_ast_tr)] end \ print_ast_translation \ (* rewrite (_cabs x (_cabs y (_cabs z t))) => (LAM x y z. t) *) (* cf. Syntax.abs_ast_tr' from src/Pure/Syntax/syn_trans.ML *) let fun cabs_ast_tr' asts = (case Ast.unfold_ast_p \<^syntax_const>\_cabs\ (Ast.Appl (Ast.Constant \<^syntax_const>\_cabs\ :: asts)) of ([], _) => raise Ast.AST ("cabs_ast_tr'", asts) | (xs, body) => Ast.Appl [Ast.Constant \<^syntax_const>\_Lambda\, Ast.fold_ast \<^syntax_const>\_cargs\ xs, body]); in [(\<^syntax_const>\_cabs\, K cabs_ast_tr')] end \ text \Dummy patterns for continuous abstraction\ translations "\ _. t" \ "CONST Abs_cfun (\_. t)" subsection \Continuous function space is pointed\ lemma bottom_cfun: "\ \ cfun" by (simp add: cfun_def inst_fun_pcpo) instance cfun :: (cpo, discrete_cpo) discrete_cpo by intro_classes (simp add: below_cfun_def Rep_cfun_inject) instance cfun :: (cpo, pcpo) pcpo by (rule typedef_pcpo [OF type_definition_cfun below_cfun_def bottom_cfun]) lemmas Rep_cfun_strict = typedef_Rep_strict [OF type_definition_cfun below_cfun_def bottom_cfun] lemmas Abs_cfun_strict = typedef_Abs_strict [OF type_definition_cfun below_cfun_def bottom_cfun] text \function application is strict in its first argument\ lemma Rep_cfun_strict1 [simp]: "\\x = \" by (simp add: Rep_cfun_strict) lemma LAM_strict [simp]: "(\ x. \) = \" by (simp add: inst_fun_pcpo [symmetric] Abs_cfun_strict) text \for compatibility with old HOLCF-Version\ lemma inst_cfun_pcpo: "\ = (\ x. \)" by simp subsection \Basic properties of continuous functions\ text \Beta-equality for continuous functions\ lemma Abs_cfun_inverse2: "cont f \ Rep_cfun (Abs_cfun f) = f" by (simp add: Abs_cfun_inverse cfun_def) lemma beta_cfun: "cont f \ (\ x. f x)\u = f u" by (simp add: Abs_cfun_inverse2) subsubsection \Beta-reduction simproc\ text \ Given the term \<^term>\(\ x. f x)\y\, the procedure tries to construct the theorem \<^term>\(\ x. f x)\y \ f y\. If this theorem cannot be completely solved by the cont2cont rules, then the procedure returns the ordinary conditional \beta_cfun\ rule. The simproc does not solve any more goals that would be solved by using \beta_cfun\ as a simp rule. The advantage of the simproc is that it can avoid deeply-nested calls to the simplifier that would otherwise be caused by large continuity side conditions. Update: The simproc now uses rule \Abs_cfun_inverse2\ instead of \beta_cfun\, to avoid problems with eta-contraction. \ simproc_setup beta_cfun_proc ("Rep_cfun (Abs_cfun f)") = \ - fn phi => fn ctxt => fn ct => + K (fn ctxt => fn ct => let val f = #2 (Thm.dest_comb (#2 (Thm.dest_comb ct))); val [T, U] = Thm.dest_ctyp (Thm.ctyp_of_cterm f); val tr = Thm.instantiate' [SOME T, SOME U] [SOME f] (mk_meta_eq @{thm Abs_cfun_inverse2}); val rules = Named_Theorems.get ctxt \<^named_theorems>\cont2cont\; val tac = SOLVED' (REPEAT_ALL_NEW (match_tac ctxt (rev rules))); - in SOME (perhaps (SINGLE (tac 1)) tr) end + in SOME (perhaps (SINGLE (tac 1)) tr) end) \ text \Eta-equality for continuous functions\ lemma eta_cfun: "(\ x. f\x) = f" by (rule Rep_cfun_inverse) text \Extensionality for continuous functions\ lemma cfun_eq_iff: "f = g \ (\x. f\x = g\x)" by (simp add: Rep_cfun_inject [symmetric] fun_eq_iff) lemma cfun_eqI: "(\x. f\x = g\x) \ f = g" by (simp add: cfun_eq_iff) text \Extensionality wrt. ordering for continuous functions\ lemma cfun_below_iff: "f \ g \ (\x. f\x \ g\x)" by (simp add: below_cfun_def fun_below_iff) lemma cfun_belowI: "(\x. f\x \ g\x) \ f \ g" by (simp add: cfun_below_iff) text \Congruence for continuous function application\ lemma cfun_cong: "f = g \ x = y \ f\x = g\y" by simp lemma cfun_fun_cong: "f = g \ f\x = g\x" by simp lemma cfun_arg_cong: "x = y \ f\x = f\y" by simp subsection \Continuity of application\ lemma cont_Rep_cfun1: "cont (\f. f\x)" by (rule cont_Rep_cfun [OF cont_id, THEN cont2cont_fun]) lemma cont_Rep_cfun2: "cont (\x. f\x)" using Rep_cfun [where x = f] by (simp add: cfun_def) lemmas monofun_Rep_cfun = cont_Rep_cfun [THEN cont2mono] lemmas monofun_Rep_cfun1 = cont_Rep_cfun1 [THEN cont2mono] lemmas monofun_Rep_cfun2 = cont_Rep_cfun2 [THEN cont2mono] text \contlub, cont properties of \<^term>\Rep_cfun\ in each argument\ lemma contlub_cfun_arg: "chain Y \ f\(\i. Y i) = (\i. f\(Y i))" by (rule cont_Rep_cfun2 [THEN cont2contlubE]) lemma contlub_cfun_fun: "chain F \ (\i. F i)\x = (\i. F i\x)" by (rule cont_Rep_cfun1 [THEN cont2contlubE]) text \monotonicity of application\ lemma monofun_cfun_fun: "f \ g \ f\x \ g\x" by (simp add: cfun_below_iff) lemma monofun_cfun_arg: "x \ y \ f\x \ f\y" by (rule monofun_Rep_cfun2 [THEN monofunE]) lemma monofun_cfun: "f \ g \ x \ y \ f\x \ g\y" by (rule below_trans [OF monofun_cfun_fun monofun_cfun_arg]) text \ch2ch - rules for the type \<^typ>\'a \ 'b\\ lemma chain_monofun: "chain Y \ chain (\i. f\(Y i))" by (erule monofun_Rep_cfun2 [THEN ch2ch_monofun]) lemma ch2ch_Rep_cfunR: "chain Y \ chain (\i. f\(Y i))" by (rule monofun_Rep_cfun2 [THEN ch2ch_monofun]) lemma ch2ch_Rep_cfunL: "chain F \ chain (\i. (F i)\x)" by (rule monofun_Rep_cfun1 [THEN ch2ch_monofun]) lemma ch2ch_Rep_cfun [simp]: "chain F \ chain Y \ chain (\i. (F i)\(Y i))" by (simp add: chain_def monofun_cfun) lemma ch2ch_LAM [simp]: "(\x. chain (\i. S i x)) \ (\i. cont (\x. S i x)) \ chain (\i. \ x. S i x)" by (simp add: chain_def cfun_below_iff) text \contlub, cont properties of \<^term>\Rep_cfun\ in both arguments\ lemma lub_APP: "chain F \ chain Y \ (\i. F i\(Y i)) = (\i. F i)\(\i. Y i)" by (simp add: contlub_cfun_fun contlub_cfun_arg diag_lub) lemma lub_LAM: assumes "\x. chain (\i. F i x)" and "\i. cont (\x. F i x)" shows "(\i. \ x. F i x) = (\ x. \i. F i x)" using assms by (simp add: lub_cfun lub_fun ch2ch_lambda) lemmas lub_distribs = lub_APP lub_LAM text \strictness\ lemma strictI: "f\x = \ \ f\\ = \" apply (rule bottomI) apply (erule subst) apply (rule minimal [THEN monofun_cfun_arg]) done text \type \<^typ>\'a \ 'b\ is chain complete\ lemma lub_cfun: "chain F \ (\i. F i) = (\ x. \i. F i\x)" by (simp add: lub_cfun lub_fun ch2ch_lambda) subsection \Continuity simplification procedure\ text \cont2cont lemma for \<^term>\Rep_cfun\\ lemma cont2cont_APP [simp, cont2cont]: assumes f: "cont (\x. f x)" assumes t: "cont (\x. t x)" shows "cont (\x. (f x)\(t x))" proof - from cont_Rep_cfun1 f have "cont (\x. (f x)\y)" for y by (rule cont_compose) with t cont_Rep_cfun2 show "cont (\x. (f x)\(t x))" by (rule cont_apply) qed text \ Two specific lemmas for the combination of LCF and HOL terms. These lemmas are needed in theories that use types like \<^typ>\'a \ 'b \ 'c\. \ lemma cont_APP_app [simp]: "cont f \ cont g \ cont (\x. ((f x)\(g x)) s)" by (rule cont2cont_APP [THEN cont2cont_fun]) lemma cont_APP_app_app [simp]: "cont f \ cont g \ cont (\x. ((f x)\(g x)) s t)" by (rule cont_APP_app [THEN cont2cont_fun]) text \cont2mono Lemma for \<^term>\\x. LAM y. c1(x)(y)\\ lemma cont2mono_LAM: "\\x. cont (\y. f x y); \y. monofun (\x. f x y)\ \ monofun (\x. \ y. f x y)" by (simp add: monofun_def cfun_below_iff) text \cont2cont Lemma for \<^term>\\x. LAM y. f x y\\ text \ Not suitable as a cont2cont rule, because on nested lambdas it causes exponential blow-up in the number of subgoals. \ lemma cont2cont_LAM: assumes f1: "\x. cont (\y. f x y)" assumes f2: "\y. cont (\x. f x y)" shows "cont (\x. \ y. f x y)" proof (rule cont_Abs_cfun) from f1 show "f x \ cfun" for x by (simp add: cfun_def) from f2 show "cont f" by (rule cont2cont_lambda) qed text \ This version does work as a cont2cont rule, since it has only a single subgoal. \ lemma cont2cont_LAM' [simp, cont2cont]: fixes f :: "'a::cpo \ 'b::cpo \ 'c::cpo" assumes f: "cont (\p. f (fst p) (snd p))" shows "cont (\x. \ y. f x y)" using assms by (simp add: cont2cont_LAM prod_cont_iff) lemma cont2cont_LAM_discrete [simp, cont2cont]: "(\y::'a::discrete_cpo. cont (\x. f x y)) \ cont (\x. \ y. f x y)" by (simp add: cont2cont_LAM) subsection \Miscellaneous\ text \Monotonicity of \<^term>\Abs_cfun\\ lemma monofun_LAM: "cont f \ cont g \ (\x. f x \ g x) \ (\ x. f x) \ (\ x. g x)" by (simp add: cfun_below_iff) text \some lemmata for functions with flat/chfin domain/range types\ lemma chfin_Rep_cfunR: "chain Y \ \s. \n. (LUB i. Y i)\s = Y n\s" for Y :: "nat \ 'a::cpo \ 'b::chfin" apply (rule allI) apply (subst contlub_cfun_fun) apply assumption apply (fast intro!: lub_eqI chfin lub_finch2 chfin2finch ch2ch_Rep_cfunL) done lemma adm_chfindom: "adm (\(u::'a::cpo \ 'b::chfin). P(u\s))" by (rule adm_subst, simp, rule adm_chfin) subsection \Continuous injection-retraction pairs\ text \Continuous retractions are strict.\ lemma retraction_strict: "\x. f\(g\x) = x \ f\\ = \" apply (rule bottomI) apply (drule_tac x="\" in spec) apply (erule subst) apply (rule monofun_cfun_arg) apply (rule minimal) done lemma injection_eq: "\x. f\(g\x) = x \ (g\x = g\y) = (x = y)" apply (rule iffI) apply (drule_tac f=f in cfun_arg_cong) apply simp apply simp done lemma injection_below: "\x. f\(g\x) = x \ (g\x \ g\y) = (x \ y)" apply (rule iffI) apply (drule_tac f=f in monofun_cfun_arg) apply simp apply (erule monofun_cfun_arg) done lemma injection_defined_rev: "\x. f\(g\x) = x \ g\z = \ \ z = \" apply (drule_tac f=f in cfun_arg_cong) apply (simp add: retraction_strict) done lemma injection_defined: "\x. f\(g\x) = x \ z \ \ \ g\z \ \" by (erule contrapos_nn, rule injection_defined_rev) text \a result about functions with flat codomain\ lemma flat_eqI: "x \ y \ x \ \ \ x = y" for x y :: "'a::flat" by (drule ax_flat) simp lemma flat_codom: "f\x = c \ f\\ = \ \ (\z. f\z = c)" for c :: "'b::flat" apply (cases "f\x = \") apply (rule disjI1) apply (rule bottomI) apply (erule_tac t="\" in subst) apply (rule minimal [THEN monofun_cfun_arg]) apply clarify apply (rule_tac a = "f\\" in refl [THEN box_equals]) apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI]) apply (erule minimal [THEN monofun_cfun_arg, THEN flat_eqI]) done subsection \Identity and composition\ definition ID :: "'a \ 'a" where "ID = (\ x. x)" definition cfcomp :: "('b \ 'c) \ ('a \ 'b) \ 'a \ 'c" where oo_def: "cfcomp = (\ f g x. f\(g\x))" abbreviation cfcomp_syn :: "['b \ 'c, 'a \ 'b] \ 'a \ 'c" (infixr "oo" 100) where "f oo g == cfcomp\f\g" lemma ID1 [simp]: "ID\x = x" by (simp add: ID_def) lemma cfcomp1: "(f oo g) = (\ x. f\(g\x))" by (simp add: oo_def) lemma cfcomp2 [simp]: "(f oo g)\x = f\(g\x)" by (simp add: cfcomp1) lemma cfcomp_LAM: "cont g \ f oo (\ x. g x) = (\ x. f\(g x))" by (simp add: cfcomp1) lemma cfcomp_strict [simp]: "\ oo f = \" by (simp add: cfun_eq_iff) text \ Show that interpretation of (pcpo, \_\_\) is a category. \<^item> The class of objects is interpretation of syntactical class pcpo. \<^item> The class of arrows between objects \<^typ>\'a\ and \<^typ>\'b\ is interpret. of \<^typ>\'a \ 'b\. \<^item> The identity arrow is interpretation of \<^term>\ID\. \<^item> The composition of f and g is interpretation of \oo\. \ lemma ID2 [simp]: "f oo ID = f" by (rule cfun_eqI, simp) lemma ID3 [simp]: "ID oo f = f" by (rule cfun_eqI) simp lemma assoc_oo: "f oo (g oo h) = (f oo g) oo h" by (rule cfun_eqI) simp subsection \Strictified functions\ default_sort pcpo definition seq :: "'a \ 'b \ 'b" where "seq = (\ x. if x = \ then \ else ID)" lemma cont2cont_if_bottom [cont2cont, simp]: assumes f: "cont (\x. f x)" and g: "cont (\x. g x)" shows "cont (\x. if f x = \ then \ else g x)" proof (rule cont_apply [OF f]) show "cont (\y. if y = \ then \ else g x)" for x unfolding cont_def is_lub_def is_ub_def ball_simps by (simp add: lub_eq_bottom_iff) show "cont (\x. if y = \ then \ else g x)" for y by (simp add: g) qed lemma seq_conv_if: "seq\x = (if x = \ then \ else ID)" by (simp add: seq_def) lemma seq_simps [simp]: "seq\\ = \" "seq\x\\ = \" "x \ \ \ seq\x = ID" by (simp_all add: seq_conv_if) definition strictify :: "('a \ 'b) \ 'a \ 'b" where "strictify = (\ f x. seq\x\(f\x))" lemma strictify_conv_if: "strictify\f\x = (if x = \ then \ else f\x)" by (simp add: strictify_def) lemma strictify1 [simp]: "strictify\f\\ = \" by (simp add: strictify_conv_if) lemma strictify2 [simp]: "x \ \ \ strictify\f\x = f\x" by (simp add: strictify_conv_if) subsection \Continuity of let-bindings\ lemma cont2cont_Let: assumes f: "cont (\x. f x)" assumes g1: "\y. cont (\x. g x y)" assumes g2: "\x. cont (\y. g x y)" shows "cont (\x. let y = f x in g x y)" unfolding Let_def using f g2 g1 by (rule cont_apply) lemma cont2cont_Let' [simp, cont2cont]: assumes f: "cont (\x. f x)" assumes g: "cont (\p. g (fst p) (snd p))" shows "cont (\x. let y = f x in g x y)" using f proof (rule cont2cont_Let) from g show "cont (\y. g x y)" for x by (simp add: prod_cont_iff) from g show "cont (\x. g x y)" for y by (simp add: prod_cont_iff) qed text \The simple version (suggested by Joachim Breitner) is needed if the type of the defined term is not a cpo.\ lemma cont2cont_Let_simple [simp, cont2cont]: assumes "\y. cont (\x. g x y)" shows "cont (\x. let y = t in g x y)" unfolding Let_def using assms . end diff --git a/src/HOL/HOLCF/Pcpo.thy b/src/HOL/HOLCF/Pcpo.thy --- a/src/HOL/HOLCF/Pcpo.thy +++ b/src/HOL/HOLCF/Pcpo.thy @@ -1,258 +1,258 @@ (* Title: HOL/HOLCF/Pcpo.thy Author: Franz Regensburger *) section \Classes cpo and pcpo\ theory Pcpo imports Porder begin subsection \Complete partial orders\ text \The class cpo of chain complete partial orders\ class cpo = po + assumes cpo: "chain S \ \x. range S <<| x" begin text \in cpo's everthing equal to THE lub has lub properties for every chain\ lemma cpo_lubI: "chain S \ range S <<| (\i. S i)" by (fast dest: cpo elim: is_lub_lub) lemma thelubE: "\chain S; (\i. S i) = l\ \ range S <<| l" by (blast dest: cpo intro: is_lub_lub) text \Properties of the lub\ lemma is_ub_thelub: "chain S \ S x \ (\i. S i)" by (blast dest: cpo intro: is_lub_lub [THEN is_lub_rangeD1]) lemma is_lub_thelub: "\chain S; range S <| x\ \ (\i. S i) \ x" by (blast dest: cpo intro: is_lub_lub [THEN is_lubD2]) lemma lub_below_iff: "chain S \ (\i. S i) \ x \ (\i. S i \ x)" by (simp add: is_lub_below_iff [OF cpo_lubI] is_ub_def) lemma lub_below: "\chain S; \i. S i \ x\ \ (\i. S i) \ x" by (simp add: lub_below_iff) lemma below_lub: "\chain S; x \ S i\ \ x \ (\i. S i)" by (erule below_trans, erule is_ub_thelub) lemma lub_range_mono: "\range X \ range Y; chain Y; chain X\ \ (\i. X i) \ (\i. Y i)" apply (erule lub_below) apply (subgoal_tac "\j. X i = Y j") apply clarsimp apply (erule is_ub_thelub) apply auto done lemma lub_range_shift: "chain Y \ (\i. Y (i + j)) = (\i. Y i)" apply (rule below_antisym) apply (rule lub_range_mono) apply fast apply assumption apply (erule chain_shift) apply (rule lub_below) apply assumption apply (rule_tac i="i" in below_lub) apply (erule chain_shift) apply (erule chain_mono) apply (rule le_add1) done lemma maxinch_is_thelub: "chain Y \ max_in_chain i Y = ((\i. Y i) = Y i)" apply (rule iffI) apply (fast intro!: lub_eqI lub_finch1) apply (unfold max_in_chain_def) apply (safe intro!: below_antisym) apply (fast elim!: chain_mono) apply (drule sym) apply (force elim!: is_ub_thelub) done text \the \\\ relation between two chains is preserved by their lubs\ lemma lub_mono: "\chain X; chain Y; \i. X i \ Y i\ \ (\i. X i) \ (\i. Y i)" by (fast elim: lub_below below_lub) text \the = relation between two chains is preserved by their lubs\ lemma lub_eq: "(\i. X i = Y i) \ (\i. X i) = (\i. Y i)" by simp lemma ch2ch_lub: assumes 1: "\j. chain (\i. Y i j)" assumes 2: "\i. chain (\j. Y i j)" shows "chain (\i. \j. Y i j)" apply (rule chainI) apply (rule lub_mono [OF 2 2]) apply (rule chainE [OF 1]) done lemma diag_lub: assumes 1: "\j. chain (\i. Y i j)" assumes 2: "\i. chain (\j. Y i j)" shows "(\i. \j. Y i j) = (\i. Y i i)" proof (rule below_antisym) have 3: "chain (\i. Y i i)" apply (rule chainI) apply (rule below_trans) apply (rule chainE [OF 1]) apply (rule chainE [OF 2]) done have 4: "chain (\i. \j. Y i j)" by (rule ch2ch_lub [OF 1 2]) show "(\i. \j. Y i j) \ (\i. Y i i)" apply (rule lub_below [OF 4]) apply (rule lub_below [OF 2]) apply (rule below_lub [OF 3]) apply (rule below_trans) apply (rule chain_mono [OF 1 max.cobounded1]) apply (rule chain_mono [OF 2 max.cobounded2]) done show "(\i. Y i i) \ (\i. \j. Y i j)" apply (rule lub_mono [OF 3 4]) apply (rule is_ub_thelub [OF 2]) done qed lemma ex_lub: assumes 1: "\j. chain (\i. Y i j)" assumes 2: "\i. chain (\j. Y i j)" shows "(\i. \j. Y i j) = (\j. \i. Y i j)" by (simp add: diag_lub 1 2) end subsection \Pointed cpos\ text \The class pcpo of pointed cpos\ class pcpo = cpo + assumes least: "\x. \y. x \ y" begin definition bottom :: "'a" ("\") where "bottom = (THE x. \y. x \ y)" lemma minimal [iff]: "\ \ x" unfolding bottom_def apply (rule the1I2) apply (rule ex_ex1I) apply (rule least) apply (blast intro: below_antisym) apply simp done end text \Old "UU" syntax:\ syntax UU :: logic translations "UU" \ "CONST bottom" text \Simproc to rewrite \<^term>\\ = x\ to \<^term>\x = \\.\ setup \Reorient_Proc.add (fn \<^Const_>\bottom _\ => true | _ => false)\ -simproc_setup reorient_bottom ("\ = x") = Reorient_Proc.proc +simproc_setup reorient_bottom ("\ = x") = \K Reorient_Proc.proc\ text \useful lemmas about \<^term>\\\\ lemma below_bottom_iff [simp]: "x \ \ \ x = \" by (simp add: po_eq_conv) lemma eq_bottom_iff: "x = \ \ x \ \" by simp lemma bottomI: "x \ \ \ x = \" by (subst eq_bottom_iff) lemma lub_eq_bottom_iff: "chain Y \ (\i. Y i) = \ \ (\i. Y i = \)" by (simp only: eq_bottom_iff lub_below_iff) subsection \Chain-finite and flat cpos\ text \further useful classes for HOLCF domains\ class chfin = po + assumes chfin: "chain Y \ \n. max_in_chain n Y" begin subclass cpo apply standard apply (frule chfin) apply (blast intro: lub_finch1) done lemma chfin2finch: "chain Y \ finite_chain Y" by (simp add: chfin finite_chain_def) end class flat = pcpo + assumes ax_flat: "x \ y \ x = \ \ x = y" begin subclass chfin proof fix Y assume *: "chain Y" show "\n. max_in_chain n Y" apply (unfold max_in_chain_def) apply (cases "\i. Y i = \") apply simp apply simp apply (erule exE) apply (rule_tac x="i" in exI) apply clarify using * apply (blast dest: chain_mono ax_flat) done qed lemma flat_below_iff: "x \ y \ x = \ \ x = y" by (safe dest!: ax_flat) lemma flat_eq: "a \ \ \ a \ b = (a = b)" by (safe dest!: ax_flat) end subsection \Discrete cpos\ class discrete_cpo = below + assumes discrete_cpo [simp]: "x \ y \ x = y" begin subclass po by standard simp_all text \In a discrete cpo, every chain is constant\ lemma discrete_chain_const: assumes S: "chain S" shows "\x. S = (\i. x)" proof (intro exI ext) fix i :: nat from S le0 have "S 0 \ S i" by (rule chain_mono) then have "S 0 = S i" by simp then show "S i = S 0" by (rule sym) qed subclass chfin proof fix S :: "nat \ 'a" assume S: "chain S" then have "\x. S = (\i. x)" by (rule discrete_chain_const) then have "max_in_chain 0 S" by (auto simp: max_in_chain_def) then show "\i. max_in_chain i S" .. qed end end diff --git a/src/HOL/HOLCF/Tools/fixrec.ML b/src/HOL/HOLCF/Tools/fixrec.ML --- a/src/HOL/HOLCF/Tools/fixrec.ML +++ b/src/HOL/HOLCF/Tools/fixrec.ML @@ -1,403 +1,403 @@ (* Title: HOL/HOLCF/Tools/fixrec.ML Author: Amber Telfer and Brian Huffman Recursive function definition package for HOLCF. *) signature FIXREC = sig val add_fixrec: (binding * typ option * mixfix) list -> (bool * (Attrib.binding * term)) list -> local_theory -> local_theory val add_fixrec_cmd: (binding * string option * mixfix) list -> (bool * (Attrib.binding * string)) list -> local_theory -> local_theory val add_matchers: (string * string) list -> theory -> theory val fixrec_simp_tac: Proof.context -> int -> tactic end structure Fixrec : FIXREC = struct open HOLCF_Library infixr 6 ->> infix -->> infix 9 ` val def_cont_fix_eq = @{thm def_cont_fix_eq} val def_cont_fix_ind = @{thm def_cont_fix_ind} fun fixrec_err s = error ("fixrec definition error:\n" ^ s) (*************************************************************************) (***************************** building types ****************************) (*************************************************************************) local fun binder_cfun \<^Type>\cfun T U\ = T :: binder_cfun U | binder_cfun \<^Type>\fun T U\ = T :: binder_cfun U | binder_cfun _ = [] fun body_cfun \<^Type>\cfun _ U\ = body_cfun U | body_cfun \<^Type>\fun _ U\ = body_cfun U | body_cfun T = T in fun matcherT (T, U) = body_cfun T ->> (binder_cfun T -->> U) ->> U end (*************************************************************************) (***************************** building terms ****************************) (*************************************************************************) val mk_trp = HOLogic.mk_Trueprop (* splits a cterm into the right and lefthand sides of equality *) fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t) (* similar to Thm.head_of, but for continuous application *) fun chead_of \<^Const_>\Rep_cfun _ _ for f _\ = chead_of f | chead_of u = u infix 1 === val (op ===) = HOLogic.mk_eq fun mk_mplus (t, u) = let val T = dest_matchT (Term.fastype_of t) in \<^Const>\Fixrec.mplus T\ ` t ` u end fun mk_run t = let val mT = Term.fastype_of t val T = dest_matchT mT val run = \<^Const>\Fixrec.run T\ in case t of \<^Const_>\Rep_cfun _ _\ $ \<^Const_>\Fixrec.succeed _\ $ u => u | _ => run ` t end (*************************************************************************) (************* fixed-point definitions and unfolding theorems ************) (*************************************************************************) structure FixrecUnfoldData = Generic_Data ( type T = thm Symtab.table val empty = Symtab.empty fun merge data : T = Symtab.merge (K true) data ) local fun name_of (Const (n, _)) = n | name_of (Free (n, _)) = n | name_of t = raise TERM ("Fixrec.add_unfold: lhs not a constant", [t]) val lhs_name = name_of o head_of o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of in val add_unfold : attribute = Thm.declaration_attribute (fn th => FixrecUnfoldData.map (Symtab.insert (K true) (lhs_name th, th))) end fun add_fixdefs (fixes : ((binding * typ) * mixfix) list) (spec : (Attrib.binding * term) list) (lthy : local_theory) = let val thy = Proof_Context.theory_of lthy val names = map (Binding.name_of o fst o fst) fixes val all_names = space_implode "_" names val (lhss, rhss) = ListPair.unzip (map (dest_eqs o snd) spec) val functional = lambda_tuple lhss (mk_tuple rhss) val fixpoint = mk_fix (mk_cabs functional) val cont_thm = let val prop = mk_trp (mk_cont functional) fun err _ = error ( "Continuity proof failed please check that cont2cont rules\n" ^ "or simp rules are configured for all non-HOLCF constants.\n" ^ "The error occurred for the goal statement:\n" ^ Syntax.string_of_term lthy prop) val rules = Named_Theorems.get lthy \<^named_theorems>\cont2cont\ val fast_tac = SOLVED' (REPEAT_ALL_NEW (match_tac lthy (rev rules))) val slow_tac = SOLVED' (simp_tac lthy) val tac = fast_tac 1 ORELSE slow_tac 1 ORELSE err in Goal.prove lthy [] [] prop (K tac) end fun one_def (Free(n,_)) r = let val b = Long_Name.base_name n in ((Binding.name (Thm.def_name b), []), r) end | one_def _ _ = fixrec_err "fixdefs: lhs not of correct form" fun defs [] _ = [] | defs (l::[]) r = [one_def l r] | defs (l::ls) r = one_def l (mk_fst r) :: defs ls (mk_snd r) val fixdefs = defs lhss fixpoint val (fixdef_thms : (term * (string * thm)) list, lthy) = lthy |> fold_map Local_Theory.define (map (apfst fst) fixes ~~ fixdefs) fun pair_equalI (thm1, thm2) = @{thm Pair_equalI} OF [thm1, thm2] val tuple_fixdef_thm = foldr1 pair_equalI (map (snd o snd) fixdef_thms) val P = Var (("P", 0), map Term.fastype_of lhss ---> \<^Type>\bool\) val predicate = lambda_tuple lhss (list_comb (P, lhss)) val tuple_induct_thm = (def_cont_fix_ind OF [tuple_fixdef_thm, cont_thm]) |> Thm.instantiate' [] [SOME (Thm.global_cterm_of thy predicate)] |> Local_Defs.unfold lthy @{thms split_paired_all split_conv split_strict} val tuple_unfold_thm = (def_cont_fix_eq OF [tuple_fixdef_thm, cont_thm]) |> Local_Defs.unfold lthy @{thms split_conv} fun unfolds [] _ = [] | unfolds (n::[]) thm = [(n, thm)] | unfolds (n::ns) thm = let val thmL = thm RS @{thm Pair_eqD1} val thmR = thm RS @{thm Pair_eqD2} in (n, thmL) :: unfolds ns thmR end val unfold_thms = unfolds names tuple_unfold_thm val induct_note : Attrib.binding * Thm.thm list = let val thm_name = Binding.qualify true all_names (Binding.name "induct") in ((thm_name, []), [tuple_induct_thm]) end fun unfold_note (name, thm) : Attrib.binding * Thm.thm list = let val thm_name = Binding.qualify true name (Binding.name "unfold") - val src = Attrib.internal (K add_unfold) + val src = Attrib.internal \<^here> (K add_unfold) in ((thm_name, [src]), [thm]) end val (_, lthy) = lthy |> fold_map Local_Theory.note (induct_note :: map unfold_note unfold_thms) in (lthy, names, fixdef_thms, map snd unfold_thms) end (*************************************************************************) (*********** monadic notation and pattern matching compilation ***********) (*************************************************************************) structure FixrecMatchData = Theory_Data ( type T = string Symtab.table val empty = Symtab.empty fun merge data = Symtab.merge (K true) data ) (* associate match functions with pattern constants *) fun add_matchers ms = FixrecMatchData.map (fold Symtab.update ms) fun taken_names (t : term) : bstring list = let fun taken (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | taken (Free(a,_) , bs) = insert (op =) a bs | taken (f $ u , bs) = taken (f, taken (u, bs)) | taken (Abs(a,_,t), bs) = taken (t, insert (op =) a bs) | taken (_ , bs) = bs in taken (t, []) end (* builds a monadic term for matching a pattern *) (* returns (rhs, free variable, used varnames) *) fun compile_pat match_name pat rhs taken = let fun comp_pat p rhs taken = if is_Free p then (rhs, p, taken) else comp_con (fastype_of p) p rhs [] taken (* compiles a monadic term for a constructor pattern *) and comp_con T p rhs vs taken = case p of \<^Const_>\Rep_cfun _ _ for f x\ => let val (rhs', v, taken') = comp_pat x rhs taken in comp_con T f rhs' (v::vs) taken' end | f $ x => let val (rhs', v, taken') = comp_pat x rhs taken in comp_con T f rhs' (v::vs) taken' end | Const (c, cT) => let val n = singleton (Name.variant_list taken) "v" val v = Free(n, T) val m = Const(match_name c, matcherT (cT, fastype_of rhs)) val k = big_lambdas vs rhs in (m`v`k, v, n::taken) end | _ => raise TERM ("fixrec: invalid pattern ", [p]) in comp_pat pat rhs taken end (* builds a monadic term for matching a function definition pattern *) (* returns (constant, (vars, matcher)) *) fun compile_lhs match_name pat rhs vs taken = case pat of \<^Const_>\Rep_cfun _ _ for f x\ => let val (rhs', v, taken') = compile_pat match_name x rhs taken in compile_lhs match_name f rhs' (v::vs) taken' end | Free(_,_) => (pat, (vs, rhs)) | Const(_,_) => (pat, (vs, rhs)) | _ => fixrec_err ("invalid function pattern: " ^ ML_Syntax.print_term pat) fun strip_alls t = (case try Logic.dest_all_global t of SOME (_, u) => strip_alls u | NONE => t) fun compile_eq match_name eq = let val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq)) in compile_lhs match_name lhs (mk_succeed rhs) [] (taken_names eq) end (* this is the pattern-matching compiler function *) fun compile_eqs match_name eqs = let val (consts, matchers) = ListPair.unzip (map (compile_eq match_name) eqs) val const = case distinct (op =) consts of [n] => n | [] => fixrec_err "no defining equations for function" | _ => fixrec_err "all equations in block must define the same function" val vars = case distinct (op = o apply2 length) (map fst matchers) of [vars] => vars | _ => fixrec_err "all equations in block must have the same arity" (* rename so all matchers use same free variables *) fun rename (vs, t) = Term.subst_free (filter_out (op =) (vs ~~ vars)) t val rhs = big_lambdas vars (mk_run (foldr1 mk_mplus (map rename matchers))) in mk_trp (const === rhs) end (*************************************************************************) (********************** Proving associated theorems **********************) (*************************************************************************) fun eta_tac i = CONVERSION Thm.eta_conversion i fun fixrec_simp_tac ctxt = let val tab = FixrecUnfoldData.get (Context.Proof ctxt) val concl = HOLogic.dest_Trueprop o Logic.strip_imp_concl o strip_alls fun tac (t, i) = let val (c, _) = (dest_Const o head_of o chead_of o fst o HOLogic.dest_eq o concl) t val unfold_thm = the (Symtab.lookup tab c) val rule = unfold_thm RS @{thm ssubst_lhs} in CHANGED (resolve_tac ctxt [rule] i THEN eta_tac i THEN asm_simp_tac ctxt i) end in SUBGOAL (fn ti => the_default no_tac (try tac ti)) end (* proves a block of pattern matching equations as theorems, using unfold *) fun make_simps ctxt (unfold_thm, eqns : (Attrib.binding * term) list) = let val rule = unfold_thm RS @{thm ssubst_lhs} val tac = resolve_tac ctxt [rule] 1 THEN eta_tac 1 THEN asm_simp_tac ctxt 1 fun prove_term t = Goal.prove ctxt [] [] t (K tac) fun prove_eqn (bind, eqn_t) = (bind, prove_term eqn_t) in map prove_eqn eqns end (*************************************************************************) (************************* Main fixrec function **************************) (*************************************************************************) local (* code adapted from HOL/Tools/Datatype/primrec.ML *) fun gen_fixrec prep_spec (raw_fixes : (binding * 'a option * mixfix) list) (raw_spec' : (bool * (Attrib.binding * 'b)) list) (lthy : local_theory) = let val (skips, raw_spec) = ListPair.unzip raw_spec' val (fixes : ((binding * typ) * mixfix) list, spec : (Attrib.binding * term) list) = fst (prep_spec raw_fixes (map (fn s => (s, [], [])) raw_spec) lthy) val names = map (Binding.name_of o fst o fst) fixes fun check_head name = member (op =) names name orelse fixrec_err ("Illegal equation head. Expected " ^ commas_quote names) val chead_of_spec = chead_of o fst o dest_eqs o Logic.strip_imp_concl o strip_alls o snd fun name_of (Free (n, _)) = tap check_head n | name_of _ = fixrec_err ("unknown term") val all_names = map (name_of o chead_of_spec) spec fun block_of_name n = map_filter (fn (m,eq) => if m = n then SOME eq else NONE) (all_names ~~ (spec ~~ skips)) val blocks = map block_of_name names val matcher_tab = FixrecMatchData.get (Proof_Context.theory_of lthy) fun match_name c = case Symtab.lookup matcher_tab c of SOME m => m | NONE => fixrec_err ("unknown pattern constructor: " ^ c) val matches = map (compile_eqs match_name) (map (map (snd o fst)) blocks) val spec' = map (pair Binding.empty_atts) matches val (lthy, _, _, unfold_thms) = add_fixdefs fixes spec' lthy val blocks' = map (map fst o filter_out snd) blocks val simps : (Attrib.binding * thm) list list = map (make_simps lthy) (unfold_thms ~~ blocks') fun mk_bind n : Attrib.binding = (Binding.qualify true n (Binding.name "simps"), @{attributes [simp]}) val simps1 : (Attrib.binding * thm list) list = map (fn (n,xs) => (mk_bind n, map snd xs)) (names ~~ simps) val simps2 : (Attrib.binding * thm list) list = map (apsnd (fn thm => [thm])) (flat simps) val (_, lthy) = lthy |> fold_map Local_Theory.note (simps1 @ simps2) in lthy end in val add_fixrec = gen_fixrec Specification.check_multi_specs val add_fixrec_cmd = gen_fixrec Specification.read_multi_specs end (* local *) (*************************************************************************) (******************************** Parsers ********************************) (*************************************************************************) val opt_thm_name' : (bool * Attrib.binding) parser = \<^keyword>\(\ -- \<^keyword>\unchecked\ -- \<^keyword>\)\ >> K (true, Binding.empty_atts) || Parse_Spec.opt_thm_name ":" >> pair false val spec' : (bool * (Attrib.binding * string)) parser = opt_thm_name' -- Parse.prop >> (fn ((a, b), c) => (a, (b, c))) val multi_specs' : (bool * (Attrib.binding * string)) list parser = let val unexpected = Scan.ahead (Parse.name || \<^keyword>\[\ || \<^keyword>\(\) in Parse.enum1 "|" (spec' --| Scan.option (unexpected -- Parse.!!! \<^keyword>\|\)) end val _ = Outer_Syntax.local_theory \<^command_keyword>\fixrec\ "define recursive functions (HOLCF)" (Parse.vars -- (Parse.where_ |-- Parse.!!! multi_specs') >> (fn (vars, specs) => add_fixrec_cmd vars specs)) end diff --git a/src/HOL/Library/Extended_Nat.thy b/src/HOL/Library/Extended_Nat.thy --- a/src/HOL/Library/Extended_Nat.thy +++ b/src/HOL/Library/Extended_Nat.thy @@ -1,707 +1,707 @@ (* Title: HOL/Library/Extended_Nat.thy Author: David von Oheimb, TU Muenchen; Florian Haftmann, TU Muenchen Contributions: David Trachtenherz, TU Muenchen *) section \Extended natural numbers (i.e. with infinity)\ theory Extended_Nat imports Main Countable Order_Continuity begin class infinity = fixes infinity :: "'a" ("\") context fixes f :: "nat \ 'a::{canonically_ordered_monoid_add, linorder_topology, complete_linorder}" begin lemma sums_SUP[simp, intro]: "f sums (SUP n. \iiType definition\ text \ We extend the standard natural numbers by a special value indicating infinity. \ typedef enat = "UNIV :: nat option set" .. text \TODO: introduce enat as coinductive datatype, enat is just \<^const>\of_nat\\ definition enat :: "nat \ enat" where "enat n = Abs_enat (Some n)" instantiation enat :: infinity begin definition "\ = Abs_enat None" instance .. end instance enat :: countable proof show "\to_nat::enat \ nat. inj to_nat" by (rule exI[of _ "to_nat \ Rep_enat"]) (simp add: inj_on_def Rep_enat_inject) qed old_rep_datatype enat "\ :: enat" proof - fix P i assume "\j. P (enat j)" "P \" then show "P i" proof induct case (Abs_enat y) then show ?case by (cases y rule: option.exhaust) (auto simp: enat_def infinity_enat_def) qed qed (auto simp add: enat_def infinity_enat_def Abs_enat_inject) declare [[coercion "enat::nat\enat"]] lemmas enat2_cases = enat.exhaust[case_product enat.exhaust] lemmas enat3_cases = enat.exhaust[case_product enat.exhaust enat.exhaust] lemma not_infinity_eq [iff]: "(x \ \) = (\i. x = enat i)" by (cases x) auto lemma not_enat_eq [iff]: "(\y. x \ enat y) = (x = \)" by (cases x) auto lemma enat_ex_split: "(\c::enat. P c) \ P \ \ (\c::nat. P c)" by (metis enat.exhaust) primrec the_enat :: "enat \ nat" where "the_enat (enat n) = n" subsection \Constructors and numbers\ instantiation enat :: zero_neq_one begin definition "0 = enat 0" definition "1 = enat 1" instance proof qed (simp add: zero_enat_def one_enat_def) end definition eSuc :: "enat \ enat" where "eSuc i = (case i of enat n \ enat (Suc n) | \ \ \)" lemma enat_0 [code_post]: "enat 0 = 0" by (simp add: zero_enat_def) lemma enat_1 [code_post]: "enat 1 = 1" by (simp add: one_enat_def) lemma enat_0_iff: "enat x = 0 \ x = 0" "0 = enat x \ x = 0" by (auto simp add: zero_enat_def) lemma enat_1_iff: "enat x = 1 \ x = 1" "1 = enat x \ x = 1" by (auto simp add: one_enat_def) lemma one_eSuc: "1 = eSuc 0" by (simp add: zero_enat_def one_enat_def eSuc_def) lemma infinity_ne_i0 [simp]: "(\::enat) \ 0" by (simp add: zero_enat_def) lemma i0_ne_infinity [simp]: "0 \ (\::enat)" by (simp add: zero_enat_def) lemma zero_one_enat_neq: "\ 0 = (1::enat)" "\ 1 = (0::enat)" unfolding zero_enat_def one_enat_def by simp_all lemma infinity_ne_i1 [simp]: "(\::enat) \ 1" by (simp add: one_enat_def) lemma i1_ne_infinity [simp]: "1 \ (\::enat)" by (simp add: one_enat_def) lemma eSuc_enat: "eSuc (enat n) = enat (Suc n)" by (simp add: eSuc_def) lemma eSuc_infinity [simp]: "eSuc \ = \" by (simp add: eSuc_def) lemma eSuc_ne_0 [simp]: "eSuc n \ 0" by (simp add: eSuc_def zero_enat_def split: enat.splits) lemma zero_ne_eSuc [simp]: "0 \ eSuc n" by (rule eSuc_ne_0 [symmetric]) lemma eSuc_inject [simp]: "eSuc m = eSuc n \ m = n" by (simp add: eSuc_def split: enat.splits) lemma eSuc_enat_iff: "eSuc x = enat y \ (\n. y = Suc n \ x = enat n)" by (cases y) (auto simp: enat_0 eSuc_enat[symmetric]) lemma enat_eSuc_iff: "enat y = eSuc x \ (\n. y = Suc n \ enat n = x)" by (cases y) (auto simp: enat_0 eSuc_enat[symmetric]) subsection \Addition\ instantiation enat :: comm_monoid_add begin definition [nitpick_simp]: "m + n = (case m of \ \ \ | enat m \ (case n of \ \ \ | enat n \ enat (m + n)))" lemma plus_enat_simps [simp, code]: fixes q :: enat shows "enat m + enat n = enat (m + n)" and "\ + q = \" and "q + \ = \" by (simp_all add: plus_enat_def split: enat.splits) instance proof fix n m q :: enat show "n + m + q = n + (m + q)" by (cases n m q rule: enat3_cases) auto show "n + m = m + n" by (cases n m rule: enat2_cases) auto show "0 + n = n" by (cases n) (simp_all add: zero_enat_def) qed end lemma eSuc_plus_1: "eSuc n = n + 1" by (cases n) (simp_all add: eSuc_enat one_enat_def) lemma plus_1_eSuc: "1 + q = eSuc q" "q + 1 = eSuc q" by (simp_all add: eSuc_plus_1 ac_simps) lemma iadd_Suc: "eSuc m + n = eSuc (m + n)" by (simp_all add: eSuc_plus_1 ac_simps) lemma iadd_Suc_right: "m + eSuc n = eSuc (m + n)" by (simp only: add.commute[of m] iadd_Suc) subsection \Multiplication\ instantiation enat :: "{comm_semiring_1, semiring_no_zero_divisors}" begin definition times_enat_def [nitpick_simp]: "m * n = (case m of \ \ if n = 0 then 0 else \ | enat m \ (case n of \ \ if m = 0 then 0 else \ | enat n \ enat (m * n)))" lemma times_enat_simps [simp, code]: "enat m * enat n = enat (m * n)" "\ * \ = (\::enat)" "\ * enat n = (if n = 0 then 0 else \)" "enat m * \ = (if m = 0 then 0 else \)" unfolding times_enat_def zero_enat_def by (simp_all split: enat.split) instance proof fix a b c :: enat show "(a * b) * c = a * (b * c)" unfolding times_enat_def zero_enat_def by (simp split: enat.split) show comm: "a * b = b * a" unfolding times_enat_def zero_enat_def by (simp split: enat.split) show "1 * a = a" unfolding times_enat_def zero_enat_def one_enat_def by (simp split: enat.split) show distr: "(a + b) * c = a * c + b * c" unfolding times_enat_def zero_enat_def by (simp split: enat.split add: distrib_right) show "0 * a = 0" unfolding times_enat_def zero_enat_def by (simp split: enat.split) show "a * 0 = 0" unfolding times_enat_def zero_enat_def by (simp split: enat.split) show "a * (b + c) = a * b + a * c" by (cases a b c rule: enat3_cases) (auto simp: times_enat_def zero_enat_def distrib_left) show "a \ 0 \ b \ 0 \ a * b \ 0" by (cases a b rule: enat2_cases) (auto simp: times_enat_def zero_enat_def) qed end lemma mult_eSuc: "eSuc m * n = n + m * n" unfolding eSuc_plus_1 by (simp add: algebra_simps) lemma mult_eSuc_right: "m * eSuc n = m + m * n" unfolding eSuc_plus_1 by (simp add: algebra_simps) lemma of_nat_eq_enat: "of_nat n = enat n" apply (induct n) apply (simp add: enat_0) apply (simp add: plus_1_eSuc eSuc_enat) done instance enat :: semiring_char_0 proof have "inj enat" by (rule injI) simp then show "inj (\n. of_nat n :: enat)" by (simp add: of_nat_eq_enat) qed lemma imult_is_infinity: "((a::enat) * b = \) = (a = \ \ b \ 0 \ b = \ \ a \ 0)" by (auto simp add: times_enat_def zero_enat_def split: enat.split) subsection \Numerals\ lemma numeral_eq_enat: "numeral k = enat (numeral k)" using of_nat_eq_enat [of "numeral k"] by simp lemma enat_numeral [code_abbrev]: "enat (numeral k) = numeral k" using numeral_eq_enat .. lemma infinity_ne_numeral [simp]: "(\::enat) \ numeral k" by (simp add: numeral_eq_enat) lemma numeral_ne_infinity [simp]: "numeral k \ (\::enat)" by (simp add: numeral_eq_enat) lemma eSuc_numeral [simp]: "eSuc (numeral k) = numeral (k + Num.One)" by (simp only: eSuc_plus_1 numeral_plus_one) subsection \Subtraction\ instantiation enat :: minus begin definition diff_enat_def: "a - b = (case a of (enat x) \ (case b of (enat y) \ enat (x - y) | \ \ 0) | \ \ \)" instance .. end lemma idiff_enat_enat [simp, code]: "enat a - enat b = enat (a - b)" by (simp add: diff_enat_def) lemma idiff_infinity [simp, code]: "\ - n = (\::enat)" by (simp add: diff_enat_def) lemma idiff_infinity_right [simp, code]: "enat a - \ = 0" by (simp add: diff_enat_def) lemma idiff_0 [simp]: "(0::enat) - n = 0" by (cases n, simp_all add: zero_enat_def) lemmas idiff_enat_0 [simp] = idiff_0 [unfolded zero_enat_def] lemma idiff_0_right [simp]: "(n::enat) - 0 = n" by (cases n) (simp_all add: zero_enat_def) lemmas idiff_enat_0_right [simp] = idiff_0_right [unfolded zero_enat_def] lemma idiff_self [simp]: "n \ \ \ (n::enat) - n = 0" by (auto simp: zero_enat_def) lemma eSuc_minus_eSuc [simp]: "eSuc n - eSuc m = n - m" by (simp add: eSuc_def split: enat.split) lemma eSuc_minus_1 [simp]: "eSuc n - 1 = n" by (simp add: one_enat_def flip: eSuc_enat zero_enat_def) (*lemmas idiff_self_eq_0_enat = idiff_self_eq_0[unfolded zero_enat_def]*) subsection \Ordering\ instantiation enat :: linordered_ab_semigroup_add begin definition [nitpick_simp]: "m \ n = (case n of enat n1 \ (case m of enat m1 \ m1 \ n1 | \ \ False) | \ \ True)" definition [nitpick_simp]: "m < n = (case m of enat m1 \ (case n of enat n1 \ m1 < n1 | \ \ True) | \ \ False)" lemma enat_ord_simps [simp]: "enat m \ enat n \ m \ n" "enat m < enat n \ m < n" "q \ (\::enat)" "q < (\::enat) \ q \ \" "(\::enat) \ q \ q = \" "(\::enat) < q \ False" by (simp_all add: less_eq_enat_def less_enat_def split: enat.splits) lemma numeral_le_enat_iff[simp]: shows "numeral m \ enat n \ numeral m \ n" by (auto simp: numeral_eq_enat) lemma numeral_less_enat_iff[simp]: shows "numeral m < enat n \ numeral m < n" by (auto simp: numeral_eq_enat) lemma enat_ord_code [code]: "enat m \ enat n \ m \ n" "enat m < enat n \ m < n" "q \ (\::enat) \ True" "enat m < \ \ True" "\ \ enat n \ False" "(\::enat) < q \ False" by simp_all instance by standard (auto simp add: less_eq_enat_def less_enat_def plus_enat_def split: enat.splits) end instance enat :: dioid proof fix a b :: enat show "(a \ b) = (\c. b = a + c)" by (cases a b rule: enat2_cases) (auto simp: le_iff_add enat_ex_split) qed instance enat :: "{linordered_nonzero_semiring, strict_ordered_comm_monoid_add}" proof fix a b c :: enat show "a \ b \ 0 \ c \c * a \ c * b" unfolding times_enat_def less_eq_enat_def zero_enat_def by (simp split: enat.splits) show "a < b \ c < d \ a + c < b + d" for a b c d :: enat by (cases a b c d rule: enat2_cases[case_product enat2_cases]) auto show "a < b \ a + 1 < b + 1" by (metis add_right_mono eSuc_minus_1 eSuc_plus_1 less_le) qed (simp add: zero_enat_def one_enat_def) (* BH: These equations are already proven generally for any type in class linordered_semidom. However, enat is not in that class because it does not have the cancellation property. Would it be worthwhile to a generalize linordered_semidom to a new class that includes enat? *) lemma add_diff_assoc_enat: "z \ y \ x + (y - z) = x + y - (z::enat)" by(cases x)(auto simp add: diff_enat_def split: enat.split) lemma enat_ord_number [simp]: "(numeral m :: enat) \ numeral n \ (numeral m :: nat) \ numeral n" "(numeral m :: enat) < numeral n \ (numeral m :: nat) < numeral n" by (simp_all add: numeral_eq_enat) lemma infinity_ileE [elim!]: "\ \ enat m \ R" by (simp add: zero_enat_def less_eq_enat_def split: enat.splits) lemma infinity_ilessE [elim!]: "\ < enat m \ R" by simp lemma eSuc_ile_mono [simp]: "eSuc n \ eSuc m \ n \ m" by (simp add: eSuc_def less_eq_enat_def split: enat.splits) lemma eSuc_mono [simp]: "eSuc n < eSuc m \ n < m" by (simp add: eSuc_def less_enat_def split: enat.splits) lemma ile_eSuc [simp]: "n \ eSuc n" by (simp add: eSuc_def less_eq_enat_def split: enat.splits) lemma not_eSuc_ilei0 [simp]: "\ eSuc n \ 0" by (simp add: zero_enat_def eSuc_def less_eq_enat_def split: enat.splits) lemma i0_iless_eSuc [simp]: "0 < eSuc n" by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.splits) lemma iless_eSuc0[simp]: "(n < eSuc 0) = (n = 0)" by (simp add: zero_enat_def eSuc_def less_enat_def split: enat.split) lemma ileI1: "m < n \ eSuc m \ n" by (simp add: eSuc_def less_eq_enat_def less_enat_def split: enat.splits) lemma Suc_ile_eq: "enat (Suc m) \ n \ enat m < n" by (cases n) auto lemma iless_Suc_eq [simp]: "enat m < eSuc n \ enat m \ n" by (auto simp add: eSuc_def less_enat_def split: enat.splits) lemma imult_infinity: "(0::enat) < n \ \ * n = \" by (simp add: zero_enat_def less_enat_def split: enat.splits) lemma imult_infinity_right: "(0::enat) < n \ n * \ = \" by (simp add: zero_enat_def less_enat_def split: enat.splits) lemma enat_0_less_mult_iff: "(0 < (m::enat) * n) = (0 < m \ 0 < n)" by (simp only: zero_less_iff_neq_zero mult_eq_0_iff, simp) lemma mono_eSuc: "mono eSuc" by (simp add: mono_def) lemma min_enat_simps [simp]: "min (enat m) (enat n) = enat (min m n)" "min q 0 = 0" "min 0 q = 0" "min q (\::enat) = q" "min (\::enat) q = q" by (auto simp add: min_def) lemma max_enat_simps [simp]: "max (enat m) (enat n) = enat (max m n)" "max q 0 = q" "max 0 q = q" "max q \ = (\::enat)" "max \ q = (\::enat)" by (simp_all add: max_def) lemma enat_ile: "n \ enat m \ \k. n = enat k" by (cases n) simp_all lemma enat_iless: "n < enat m \ \k. n = enat k" by (cases n) simp_all lemma iadd_le_enat_iff: "x + y \ enat n \ (\y' x'. x = enat x' \ y = enat y' \ x' + y' \ n)" by(cases x y rule: enat.exhaust[case_product enat.exhaust]) simp_all lemma chain_incr: "\i. \j. Y i < Y j \ \j. enat k < Y j" apply (induct_tac k) apply (simp (no_asm) only: enat_0) apply (fast intro: le_less_trans [OF zero_le]) apply (erule exE) apply (drule spec) apply (erule exE) apply (drule ileI1) apply (rule eSuc_enat [THEN subst]) apply (rule exI) apply (erule (1) le_less_trans) done lemma eSuc_max: "eSuc (max x y) = max (eSuc x) (eSuc y)" by (simp add: eSuc_def split: enat.split) lemma eSuc_Max: assumes "finite A" "A \ {}" shows "eSuc (Max A) = Max (eSuc ` A)" using assms proof induction case (insert x A) thus ?case by(cases "A = {}")(simp_all add: eSuc_max) qed simp instantiation enat :: "{order_bot, order_top}" begin definition bot_enat :: enat where "bot_enat = 0" definition top_enat :: enat where "top_enat = \" instance by standard (simp_all add: bot_enat_def top_enat_def) end lemma finite_enat_bounded: assumes le_fin: "\y. y \ A \ y \ enat n" shows "finite A" proof (rule finite_subset) show "finite (enat ` {..n})" by blast have "A \ {..enat n}" using le_fin by fastforce also have "\ \ enat ` {..n}" apply (rule subsetI) subgoal for x by (cases x) auto done finally show "A \ enat ` {..n}" . qed subsection \Cancellation simprocs\ lemma add_diff_cancel_enat[simp]: "x \ \ \ x + y - x = (y::enat)" by (metis add.commute add.right_neutral add_diff_assoc_enat idiff_self order_refl) lemma enat_add_left_cancel: "a + b = a + c \ a = (\::enat) \ b = c" unfolding plus_enat_def by (simp split: enat.split) lemma enat_add_left_cancel_le: "a + b \ a + c \ a = (\::enat) \ b \ c" unfolding plus_enat_def by (simp split: enat.split) lemma enat_add_left_cancel_less: "a + b < a + c \ a \ (\::enat) \ b < c" unfolding plus_enat_def by (simp split: enat.split) lemma plus_eq_infty_iff_enat: "(m::enat) + n = \ \ m=\ \ n=\" using enat_add_left_cancel by fastforce ML \ structure Cancel_Enat_Common = struct (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) fun find_first_t _ _ [] = raise TERM("find_first_t", []) | find_first_t past u (t::terms) = if u aconv t then (rev past @ terms) else find_first_t (t::past) u terms fun dest_summing (Const (\<^const_name>\Groups.plus\, _) $ t $ u, ts) = dest_summing (t, dest_summing (u, ts)) | dest_summing (t, ts) = t :: ts val mk_sum = Arith_Data.long_mk_sum fun dest_sum t = dest_summing (t, []) val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac val norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms ac_simps add_0_left add_0_right}) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) fun simplify_meta_eq ctxt cancel_th th = Arith_Data.simplify_meta_eq [] ctxt ([th, cancel_th] MRS trans) fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end structure Eq_Enat_Cancel = ExtractCommonTermFun (open Cancel_Enat_Common val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ \<^typ>\enat\ fun simp_conv _ _ = SOME @{thm enat_add_left_cancel} ) structure Le_Enat_Cancel = ExtractCommonTermFun (open Cancel_Enat_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ \<^typ>\enat\ fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_le} ) structure Less_Enat_Cancel = ExtractCommonTermFun (open Cancel_Enat_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ \<^typ>\enat\ fun simp_conv _ _ = SOME @{thm enat_add_left_cancel_less} ) \ simproc_setup enat_eq_cancel ("(l::enat) + m = n" | "(l::enat) = m + n") = - \fn phi => fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Eq_Enat_Cancel.proc ctxt (Thm.term_of ct))\ simproc_setup enat_le_cancel ("(l::enat) + m \ n" | "(l::enat) \ m + n") = - \fn phi => fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Le_Enat_Cancel.proc ctxt (Thm.term_of ct))\ simproc_setup enat_less_cancel ("(l::enat) + m < n" | "(l::enat) < m + n") = - \fn phi => fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Less_Enat_Cancel.proc ctxt (Thm.term_of ct))\ text \TODO: add regression tests for these simprocs\ text \TODO: add simprocs for combining and cancelling numerals\ subsection \Well-ordering\ lemma less_enatE: "[| n < enat m; !!k. n = enat k ==> k < m ==> P |] ==> P" by (induct n) auto lemma less_infinityE: "[| n < \; !!k. n = enat k ==> P |] ==> P" by (induct n) auto lemma enat_less_induct: assumes prem: "\n. \m::enat. m < n \ P m \ P n" shows "P n" proof - have P_enat: "\k. P (enat k)" apply (rule nat_less_induct) apply (rule prem, clarify) apply (erule less_enatE, simp) done show ?thesis proof (induct n) fix nat show "P (enat nat)" by (rule P_enat) next show "P \" apply (rule prem, clarify) apply (erule less_infinityE) apply (simp add: P_enat) done qed qed instance enat :: wellorder proof fix P and n assume hyp: "(\n::enat. (\m::enat. m < n \ P m) \ P n)" show "P n" by (blast intro: enat_less_induct hyp) qed subsection \Complete Lattice\ instantiation enat :: complete_lattice begin definition inf_enat :: "enat \ enat \ enat" where "inf_enat = min" definition sup_enat :: "enat \ enat \ enat" where "sup_enat = max" definition Inf_enat :: "enat set \ enat" where "Inf_enat A = (if A = {} then \ else (LEAST x. x \ A))" definition Sup_enat :: "enat set \ enat" where "Sup_enat A = (if A = {} then 0 else if finite A then Max A else \)" instance proof fix x :: "enat" and A :: "enat set" { assume "x \ A" then show "Inf A \ x" unfolding Inf_enat_def by (auto intro: Least_le) } { assume "\y. y \ A \ x \ y" then show "x \ Inf A" unfolding Inf_enat_def by (cases "A = {}") (auto intro: LeastI2_ex) } { assume "x \ A" then show "x \ Sup A" unfolding Sup_enat_def by (cases "finite A") auto } { assume "\y. y \ A \ y \ x" then show "Sup A \ x" unfolding Sup_enat_def using finite_enat_bounded by auto } qed (simp_all add: inf_enat_def sup_enat_def bot_enat_def top_enat_def Inf_enat_def Sup_enat_def) end instance enat :: complete_linorder .. lemma eSuc_Sup: "A \ {} \ eSuc (Sup A) = Sup (eSuc ` A)" by(auto simp add: Sup_enat_def eSuc_Max inj_on_def dest: finite_imageD) lemma sup_continuous_eSuc: "sup_continuous f \ sup_continuous (\x. eSuc (f x))" using eSuc_Sup [of "_ ` UNIV"] by (auto simp: sup_continuous_def image_comp) subsection \Traditional theorem names\ lemmas enat_defs = zero_enat_def one_enat_def eSuc_def plus_enat_def less_eq_enat_def less_enat_def lemma iadd_is_0: "(m + n = (0::enat)) = (m = 0 \ n = 0)" by (rule add_eq_0_iff_both_eq_0) lemma i0_lb : "(0::enat) \ n" by (rule zero_le) lemma ile0_eq: "n \ (0::enat) \ n = 0" by (rule le_zero_eq) lemma not_iless0: "\ n < (0::enat)" by (rule not_less_zero) lemma i0_less[simp]: "(0::enat) < n \ n \ 0" by (rule zero_less_iff_neq_zero) lemma imult_is_0: "((m::enat) * n = 0) = (m = 0 \ n = 0)" by (rule mult_eq_0_iff) end diff --git a/src/HOL/Library/Extended_Nonnegative_Real.thy b/src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy @@ -1,2047 +1,2047 @@ (* Title: HOL/Library/Extended_Nonnegative_Real.thy Author: Johannes Hölzl *) section \The type of non-negative extended real numbers\ theory Extended_Nonnegative_Real imports Extended_Real Indicator_Function begin lemma ereal_ineq_diff_add: assumes "b \ (-\::ereal)" "a \ b" shows "a = b + (a-b)" by (metis add.commute assms ereal_eq_minus_iff ereal_minus_le_iff ereal_plus_eq_PInfty) lemma Limsup_const_add: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Limsup F (\x. c + f x) = c + Limsup F f" by (rule Limsup_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma Liminf_const_add: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Liminf F (\x. c + f x) = c + Liminf F f" by (rule Liminf_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma Liminf_add_const: fixes c :: "'a::{complete_linorder, linorder_topology, topological_monoid_add, ordered_ab_semigroup_add}" shows "F \ bot \ Liminf F (\x. f x + c) = Liminf F f + c" by (rule Liminf_compose_continuous_mono) (auto intro!: monoI add_mono continuous_on_add continuous_on_id continuous_on_const) lemma sums_offset: fixes f g :: "nat \ 'a :: {t2_space, topological_comm_monoid_add}" assumes "(\n. f (n + i)) sums l" shows "f sums (l + (\jk. (\nj l + (\jjj=i..j=0..j=i..j\(\n. n + i)`{0..jnjk. (\n l + (\j 'a :: {t2_space, topological_comm_monoid_add}" shows "summable (\j. f (j + i)) \ suminf f = (\j. f (j + i)) + (\jz::real. 0 < z \ z < 1 \ P z) \ eventually P (at_left 1)" by (subst eventually_at_left[of 0]) (auto intro: exI[of _ 0]) lemma mult_eq_1: fixes a b :: "'a :: {ordered_semiring, comm_monoid_mult}" shows "0 \ a \ a \ 1 \ b \ 1 \ a * b = 1 \ (a = 1 \ b = 1)" by (metis mult.left_neutral eq_iff mult.commute mult_right_mono) lemma ereal_add_diff_cancel: fixes a b :: ereal shows "\b\ \ \ \ (a + b) - b = a" by (cases a b rule: ereal2_cases) auto lemma add_top: fixes x :: "'a::{order_top, ordered_comm_monoid_add}" shows "0 \ x \ x + top = top" by (intro top_le add_increasing order_refl) lemma top_add: fixes x :: "'a::{order_top, ordered_comm_monoid_add}" shows "0 \ x \ top + x = top" by (intro top_le add_increasing2 order_refl) lemma le_lfp: "mono f \ x \ lfp f \ f x \ lfp f" by (subst lfp_unfold) (auto dest: monoD) lemma lfp_transfer: assumes \: "sup_continuous \" and f: "sup_continuous f" and mg: "mono g" assumes bot: "\ bot \ lfp g" and eq: "\x. x \ lfp f \ \ (f x) = g (\ x)" shows "\ (lfp f) = lfp g" proof (rule antisym) note mf = sup_continuous_mono[OF f] have f_le_lfp: "(f ^^ i) bot \ lfp f" for i by (induction i) (auto intro: le_lfp mf) have "\ ((f ^^ i) bot) \ lfp g" for i by (induction i) (auto simp: bot eq f_le_lfp intro!: le_lfp mg) then show "\ (lfp f) \ lfp g" unfolding sup_continuous_lfp[OF f] by (subst \[THEN sup_continuousD]) (auto intro!: mono_funpow sup_continuous_mono[OF f] SUP_least) show "lfp g \ \ (lfp f)" by (rule lfp_lowerbound) (simp add: eq[symmetric] lfp_fixpoint[OF mf]) qed lemma sup_continuous_applyD: "sup_continuous f \ sup_continuous (\x. f x h)" using sup_continuous_apply[THEN sup_continuous_compose] . lemma sup_continuous_SUP[order_continuous_intros]: fixes M :: "_ \ _ \ 'a::complete_lattice" assumes M: "\i. i \ I \ sup_continuous (M i)" shows "sup_continuous (SUP i\I. M i)" unfolding sup_continuous_def by (auto simp add: sup_continuousD [OF M] image_comp intro: SUP_commute) lemma sup_continuous_apply_SUP[order_continuous_intros]: fixes M :: "_ \ _ \ 'a::complete_lattice" shows "(\i. i \ I \ sup_continuous (M i)) \ sup_continuous (\x. SUP i\I. M i x)" unfolding SUP_apply[symmetric] by (rule sup_continuous_SUP) lemma sup_continuous_lfp'[order_continuous_intros]: assumes 1: "sup_continuous f" assumes 2: "\g. sup_continuous g \ sup_continuous (f g)" shows "sup_continuous (lfp f)" proof - have "sup_continuous ((f ^^ i) bot)" for i proof (induction i) case (Suc i) then show ?case by (auto intro!: 2) qed (simp add: bot_fun_def sup_continuous_const) then show ?thesis unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) qed lemma sup_continuous_lfp''[order_continuous_intros]: assumes 1: "\s. sup_continuous (f s)" assumes 2: "\g. sup_continuous g \ sup_continuous (\s. f s (g s))" shows "sup_continuous (\x. lfp (f x))" proof - have "sup_continuous (\x. (f x ^^ i) bot)" for i proof (induction i) case (Suc i) then show ?case by (auto intro!: 2) qed (simp add: bot_fun_def sup_continuous_const) then show ?thesis unfolding sup_continuous_lfp[OF 1] by (intro order_continuous_intros) qed lemma mono_INF_fun: "(\x y. mono (F x y)) \ mono (\z x. INF y \ X x. F x y z :: 'a :: complete_lattice)" by (auto intro!: INF_mono[OF bexI] simp: le_fun_def mono_def) lemma continuous_on_cmult_ereal: "\c::ereal\ \ \ \ continuous_on A f \ continuous_on A (\x. c * f x)" using tendsto_cmult_ereal[of c f "f x" "at x within A" for x] by (auto simp: continuous_on_def simp del: tendsto_cmult_ereal) lemma real_of_nat_Sup: assumes "A \ {}" "bdd_above A" shows "of_nat (Sup A) = (SUP a\A. of_nat a :: real)" proof (intro antisym) show "(SUP a\A. of_nat a::real) \ of_nat (Sup A)" using assms by (intro cSUP_least of_nat_mono) (auto intro: cSup_upper) have "Sup A \ A" using assms by (auto simp: Sup_nat_def bdd_above_nat) then show "of_nat (Sup A) \ (SUP a\A. of_nat a::real)" by (intro cSUP_upper bdd_above_image_mono assms) (auto simp: mono_def) qed lemma (in complete_lattice) SUP_sup_const1: "I \ {} \ (SUP i\I. sup c (f i)) = sup c (SUP i\I. f i)" using SUP_sup_distrib[of "\_. c" I f] by simp lemma (in complete_lattice) SUP_sup_const2: "I \ {} \ (SUP i\I. sup (f i) c) = sup (SUP i\I. f i) c" using SUP_sup_distrib[of f I "\_. c"] by simp lemma one_less_of_natD: assumes "(1::'a::linordered_semidom) < of_nat n" shows "1 < n" by (cases n) (use assms in auto) subsection \Defining the extended non-negative reals\ text \Basic definitions and type class setup\ typedef ennreal = "{x :: ereal. 0 \ x}" morphisms enn2ereal e2ennreal' by auto definition "e2ennreal x = e2ennreal' (max 0 x)" lemma enn2ereal_range: "e2ennreal ` {0..} = UNIV" proof - have "\y\0. x = e2ennreal y" for x by (cases x) (auto simp: e2ennreal_def max_absorb2) then show ?thesis by (auto simp: image_iff Bex_def) qed lemma type_definition_ennreal': "type_definition enn2ereal e2ennreal {x. 0 \ x}" using type_definition_ennreal by (auto simp: type_definition_def e2ennreal_def max_absorb2) setup_lifting type_definition_ennreal' declare [[coercion e2ennreal]] instantiation ennreal :: complete_linorder begin lift_definition top_ennreal :: ennreal is top by (rule top_greatest) lift_definition bot_ennreal :: ennreal is 0 by (rule order_refl) lift_definition sup_ennreal :: "ennreal \ ennreal \ ennreal" is sup by (rule le_supI1) lift_definition inf_ennreal :: "ennreal \ ennreal \ ennreal" is inf by (rule le_infI) lift_definition Inf_ennreal :: "ennreal set \ ennreal" is "Inf" by (rule Inf_greatest) lift_definition Sup_ennreal :: "ennreal set \ ennreal" is "sup 0 \ Sup" by auto lift_definition less_eq_ennreal :: "ennreal \ ennreal \ bool" is "(\)" . lift_definition less_ennreal :: "ennreal \ ennreal \ bool" is "(<)" . instance by standard (transfer ; auto simp: Inf_lower Inf_greatest Sup_upper Sup_least le_max_iff_disj max.absorb1)+ end lemma pcr_ennreal_enn2ereal[simp]: "pcr_ennreal (enn2ereal x) x" by (simp add: ennreal.pcr_cr_eq cr_ennreal_def) lemma rel_fun_eq_pcr_ennreal: "rel_fun (=) pcr_ennreal f g \ f = enn2ereal \ g" by (auto simp: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) instantiation ennreal :: infinity begin definition infinity_ennreal :: ennreal where [simp]: "\ = (top::ennreal)" instance .. end instantiation ennreal :: "{semiring_1_no_zero_divisors, comm_semiring_1}" begin lift_definition one_ennreal :: ennreal is 1 by simp lift_definition zero_ennreal :: ennreal is 0 by simp lift_definition plus_ennreal :: "ennreal \ ennreal \ ennreal" is "(+)" by simp lift_definition times_ennreal :: "ennreal \ ennreal \ ennreal" is "(*)" by simp instance by standard (transfer; auto simp: field_simps ereal_right_distrib)+ end instantiation ennreal :: minus begin lift_definition minus_ennreal :: "ennreal \ ennreal \ ennreal" is "\a b. max 0 (a - b)" by simp instance .. end instance ennreal :: numeral .. instantiation ennreal :: inverse begin lift_definition inverse_ennreal :: "ennreal \ ennreal" is inverse by (rule inverse_ereal_ge0I) definition divide_ennreal :: "ennreal \ ennreal \ ennreal" where "x div y = x * inverse (y :: ennreal)" instance .. end lemma ennreal_zero_less_one: "0 < (1::ennreal)" \ \TODO: remove\ by transfer auto instance ennreal :: dioid proof (standard; transfer) fix a b :: ereal assume "0 \ a" "0 \ b" then show "(a \ b) = (\c\Collect ((\) 0). b = a + c)" unfolding ereal_ex_split Bex_def by (cases a b rule: ereal2_cases) (auto intro!: exI[of _ "real_of_ereal (b - a)"]) qed instance ennreal :: ordered_comm_semiring by standard (transfer ; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+ instance ennreal :: linordered_nonzero_semiring proof fix a b::ennreal show "a < b \ a + 1 < b + 1" by transfer (simp add: add_right_mono ereal_add_cancel_right less_le) qed (transfer; simp) instance ennreal :: strict_ordered_ab_semigroup_add proof fix a b c d :: ennreal show "a < b \ c < d \ a + c < b + d" by transfer (auto intro!: ereal_add_strict_mono) qed declare [[coercion "of_nat :: nat \ ennreal"]] lemma e2ennreal_neg: "x \ 0 \ e2ennreal x = 0" unfolding zero_ennreal_def e2ennreal_def by (simp add: max_absorb1) lemma e2ennreal_mono: "x \ y \ e2ennreal x \ e2ennreal y" by (cases "0 \ x" "0 \ y" rule: bool.exhaust[case_product bool.exhaust]) (auto simp: e2ennreal_neg less_eq_ennreal.abs_eq eq_onp_def) lemma enn2ereal_nonneg[simp]: "0 \ enn2ereal x" using ennreal.enn2ereal[of x] by simp lemma ereal_ennreal_cases: obtains b where "0 \ a" "a = enn2ereal b" | "a < 0" using e2ennreal'_inverse[of a, symmetric] by (cases "0 \ a") (auto intro: enn2ereal_nonneg) lemma rel_fun_liminf[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal liminf liminf" proof - have "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal (\x. sup 0 (liminf x)) liminf" unfolding liminf_SUP_INF[abs_def] by (transfer_prover_start, transfer_step+; simp) then show ?thesis apply (subst (asm) (2) rel_fun_def) apply (subst (2) rel_fun_def) apply (auto simp: comp_def max.absorb2 Liminf_bounded rel_fun_eq_pcr_ennreal) done qed lemma rel_fun_limsup[transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal limsup limsup" proof - have "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal (\x. INF n. sup 0 (SUP i\{n..}. x i)) limsup" unfolding limsup_INF_SUP[abs_def] by (transfer_prover_start, transfer_step+; simp) then show ?thesis unfolding limsup_INF_SUP[abs_def] apply (subst (asm) (2) rel_fun_def) apply (subst (2) rel_fun_def) apply (auto simp: comp_def max.absorb2 Sup_upper2 rel_fun_eq_pcr_ennreal) apply (subst (asm) max.absorb2) apply (auto intro: SUP_upper2) done qed lemma sum_enn2ereal[simp]: "(\i. i \ I \ 0 \ f i) \ (\i\I. enn2ereal (f i)) = enn2ereal (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg zero_ennreal.rep_eq plus_ennreal.rep_eq) lemma transfer_e2ennreal_sum [transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) (rel_fun (=) pcr_ennreal) sum sum" by (auto intro!: rel_funI simp: rel_fun_eq_pcr_ennreal comp_def) lemma enn2ereal_of_nat[simp]: "enn2ereal (of_nat n) = ereal n" by (induction n) (auto simp: zero_ennreal.rep_eq one_ennreal.rep_eq plus_ennreal.rep_eq) lemma enn2ereal_numeral[simp]: "enn2ereal (numeral a) = numeral a" by (metis enn2ereal_of_nat numeral_eq_ereal of_nat_numeral) lemma transfer_numeral[transfer_rule]: "pcr_ennreal (numeral a) (numeral a)" unfolding cr_ennreal_def pcr_ennreal_def by auto subsection \Cancellation simprocs\ lemma ennreal_add_left_cancel: "a + b = a + c \ a = (\::ennreal) \ b = c" unfolding infinity_ennreal_def by transfer (simp add: top_ereal_def ereal_add_cancel_left) lemma ennreal_add_left_cancel_le: "a + b \ a + c \ a = (\::ennreal) \ b \ c" unfolding infinity_ennreal_def by transfer (simp add: ereal_add_le_add_iff top_ereal_def disj_commute) lemma ereal_add_left_cancel_less: fixes a b c :: ereal shows "0 \ a \ 0 \ b \ a + b < a + c \ a \ \ \ b < c" by (cases a b c rule: ereal3_cases) auto lemma ennreal_add_left_cancel_less: "a + b < a + c \ a \ (\::ennreal) \ b < c" unfolding infinity_ennreal_def by transfer (simp add: top_ereal_def ereal_add_left_cancel_less) ML \ structure Cancel_Ennreal_Common = struct (* copied from src/HOL/Tools/nat_numeral_simprocs.ML *) fun find_first_t _ _ [] = raise TERM("find_first_t", []) | find_first_t past u (t::terms) = if u aconv t then (rev past @ terms) else find_first_t (t::past) u terms fun dest_summing (Const (\<^const_name>\Groups.plus\, _) $ t $ u, ts) = dest_summing (t, dest_summing (u, ts)) | dest_summing (t, ts) = t :: ts val mk_sum = Arith_Data.long_mk_sum fun dest_sum t = dest_summing (t, []) val find_first = find_first_t [] val trans_tac = Numeral_Simprocs.trans_tac val norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms ac_simps add_0_left add_0_right}) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) fun simplify_meta_eq ctxt cancel_th th = Arith_Data.simplify_meta_eq [] ctxt ([th, cancel_th] MRS trans) fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end structure Eq_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel} ) structure Le_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_le} ) structure Less_Ennreal_Cancel = ExtractCommonTermFun (open Cancel_Ennreal_Common val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ \<^typ>\ennreal\ fun simp_conv _ _ = SOME @{thm ennreal_add_left_cancel_less} ) \ simproc_setup ennreal_eq_cancel ("(l::ennreal) + m = n" | "(l::ennreal) = m + n") = - \fn phi => fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Eq_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\ simproc_setup ennreal_le_cancel ("(l::ennreal) + m \ n" | "(l::ennreal) \ m + n") = - \fn phi => fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Le_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\ simproc_setup ennreal_less_cancel ("(l::ennreal) + m < n" | "(l::ennreal) < m + n") = - \fn phi => fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => Less_Ennreal_Cancel.proc ctxt (Thm.term_of ct))\ subsection \Order with top\ lemma ennreal_zero_less_top[simp]: "0 < (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_one_less_top[simp]: "1 < (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_zero_neq_top[simp]: "0 \ (top::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_top_neq_zero[simp]: "(top::ennreal) \ 0" by transfer (simp add: top_ereal_def) lemma ennreal_top_neq_one[simp]: "top \ (1::ennreal)" by transfer (simp add: top_ereal_def one_ereal_def flip: ereal_max) lemma ennreal_one_neq_top[simp]: "1 \ (top::ennreal)" by transfer (simp add: top_ereal_def one_ereal_def flip: ereal_max) lemma ennreal_add_less_top[simp]: fixes a b :: ennreal shows "a + b < top \ a < top \ b < top" by transfer (auto simp: top_ereal_def) lemma ennreal_add_eq_top[simp]: fixes a b :: ennreal shows "a + b = top \ a = top \ b = top" by transfer (auto simp: top_ereal_def) lemma ennreal_sum_less_top[simp]: fixes f :: "'a \ ennreal" shows "finite I \ (\i\I. f i) < top \ (\i\I. f i < top)" by (induction I rule: finite_induct) auto lemma ennreal_sum_eq_top[simp]: fixes f :: "'a \ ennreal" shows "finite I \ (\i\I. f i) = top \ (\i\I. f i = top)" by (induction I rule: finite_induct) auto lemma ennreal_mult_eq_top_iff: fixes a b :: ennreal shows "a * b = top \ (a = top \ b \ 0) \ (b = top \ a \ 0)" by transfer (auto simp: top_ereal_def) lemma ennreal_top_eq_mult_iff: fixes a b :: ennreal shows "top = a * b \ (a = top \ b \ 0) \ (b = top \ a \ 0)" using ennreal_mult_eq_top_iff[of a b] by auto lemma ennreal_mult_less_top: fixes a b :: ennreal shows "a * b < top \ (a = 0 \ b = 0 \ (a < top \ b < top))" by transfer (auto simp add: top_ereal_def) lemma top_power_ennreal: "top ^ n = (if n = 0 then 1 else top :: ennreal)" by (induction n) (simp_all add: ennreal_mult_eq_top_iff) lemma ennreal_prod_eq_0[simp]: fixes f :: "'a \ ennreal" shows "(prod f A = 0) = (finite A \ (\i\A. f i = 0))" by (induction A rule: infinite_finite_induct) auto lemma ennreal_prod_eq_top: fixes f :: "'a \ ennreal" shows "(\i\I. f i) = top \ (finite I \ ((\i\I. f i \ 0) \ (\i\I. f i = top)))" by (induction I rule: infinite_finite_induct) (auto simp: ennreal_mult_eq_top_iff) lemma ennreal_top_mult: "top * a = (if a = 0 then 0 else top :: ennreal)" by (simp add: ennreal_mult_eq_top_iff) lemma ennreal_mult_top: "a * top = (if a = 0 then 0 else top :: ennreal)" by (simp add: ennreal_mult_eq_top_iff) lemma enn2ereal_eq_top_iff[simp]: "enn2ereal x = \ \ x = top" by transfer (simp add: top_ereal_def) lemma enn2ereal_top[simp]: "enn2ereal top = \" by transfer (simp add: top_ereal_def) lemma e2ennreal_infty[simp]: "e2ennreal \ = top" by (simp add: top_ennreal.abs_eq top_ereal_def) lemma ennreal_top_minus[simp]: "top - x = (top::ennreal)" by transfer (auto simp: top_ereal_def max_def) lemma minus_top_ennreal: "x - top = (if x = top then top else 0:: ennreal)" by transfer (use ereal_eq_minus_iff top_ereal_def in force) lemma bot_ennreal: "bot = (0::ennreal)" by transfer rule lemma ennreal_of_nat_neq_top[simp]: "of_nat i \ (top::ennreal)" by (induction i) auto lemma numeral_eq_of_nat: "(numeral a::ennreal) = of_nat (numeral a)" by simp lemma of_nat_less_top: "of_nat i < (top::ennreal)" using less_le_trans[of "of_nat i" "of_nat (Suc i)" "top::ennreal"] by simp lemma top_neq_numeral[simp]: "top \ (numeral i::ennreal)" using of_nat_less_top[of "numeral i"] by simp lemma ennreal_numeral_less_top[simp]: "numeral i < (top::ennreal)" using of_nat_less_top[of "numeral i"] by simp lemma ennreal_add_bot[simp]: "bot + x = (x::ennreal)" by transfer simp lemma add_top_right_ennreal [simp]: "x + top = (top :: ennreal)" by (cases x) auto lemma add_top_left_ennreal [simp]: "top + x = (top :: ennreal)" by (cases x) auto lemma ennreal_top_mult_left [simp]: "x \ 0 \ x * top = (top :: ennreal)" by (subst ennreal_mult_eq_top_iff) auto lemma ennreal_top_mult_right [simp]: "x \ 0 \ top * x = (top :: ennreal)" by (subst ennreal_mult_eq_top_iff) auto lemma power_top_ennreal [simp]: "n > 0 \ top ^ n = (top :: ennreal)" by (induction n) auto lemma power_eq_top_ennreal_iff: "x ^ n = top \ x = (top :: ennreal) \ n > 0" by (induction n) (auto simp: ennreal_mult_eq_top_iff) lemma ennreal_mult_le_mult_iff: "c \ 0 \ c \ top \ c * a \ c * b \ a \ (b :: ennreal)" including ennreal.lifting by (transfer, subst ereal_mult_le_mult_iff) (auto simp: top_ereal_def) lemma power_mono_ennreal: "x \ y \ x ^ n \ (y ^ n :: ennreal)" by (induction n) (auto intro!: mult_mono) instance ennreal :: semiring_char_0 proof (standard, safe intro!: linorder_injI) have *: "1 + of_nat k \ (0::ennreal)" for k using add_pos_nonneg[OF zero_less_one, of "of_nat k :: ennreal"] by auto fix x y :: nat assume "x < y" "of_nat x = (of_nat y::ennreal)" then show False by (auto simp add: less_iff_Suc_add *) qed subsection \Arithmetic\ lemma ennreal_minus_zero[simp]: "a - (0::ennreal) = a" by transfer (auto simp: max_def) lemma ennreal_add_diff_cancel_right[simp]: fixes x y z :: ennreal shows "y \ top \ (x + y) - y = x" by transfer (metis ereal_eq_minus_iff max_absorb2 not_MInfty_nonneg top_ereal_def) lemma ennreal_add_diff_cancel_left[simp]: fixes x y z :: ennreal shows "y \ top \ (y + x) - y = x" by (simp add: add.commute) lemma fixes a b :: ennreal shows "a - b = 0 \ a \ b" by transfer (metis ereal_diff_gr0 le_cases max.absorb2 not_less) lemma ennreal_minus_cancel: fixes a b c :: ennreal shows "c \ top \ a \ c \ b \ c \ c - a = c - b \ a = b" by (metis ennreal_add_diff_cancel_left ennreal_add_diff_cancel_right ennreal_add_eq_top less_eqE) lemma sup_const_add_ennreal: fixes a b c :: "ennreal" shows "sup (c + a) (c + b) = c + sup a b" by transfer (metis add_left_mono le_cases sup.absorb2 sup.orderE) lemma ennreal_diff_add_assoc: fixes a b c :: ennreal shows "a \ b \ c + b - a = c + (b - a)" by (metis add.left_commute ennreal_add_diff_cancel_left ennreal_add_eq_top ennreal_top_minus less_eqE) lemma mult_divide_eq_ennreal: fixes a b :: ennreal shows "b \ 0 \ b \ top \ (a * b) / b = a" unfolding divide_ennreal_def apply transfer by (metis abs_ereal_ge0 divide_ereal_def ereal_divide_eq ereal_times_divide_eq top_ereal_def) lemma divide_mult_eq: "a \ 0 \ a \ \ \ x * a / (b * a) = x / (b::ennreal)" unfolding divide_ennreal_def infinity_ennreal_def apply transfer subgoal for a b c apply (cases a b c rule: ereal3_cases) apply (auto simp: top_ereal_def) done done lemma ennreal_mult_divide_eq: fixes a b :: ennreal shows "b \ 0 \ b \ top \ (a * b) / b = a" by (fact mult_divide_eq_ennreal) lemma ennreal_add_diff_cancel: fixes a b :: ennreal shows "b \ \ \ (a + b) - b = a" by simp lemma ennreal_minus_eq_0: "a - b = 0 \ a \ (b::ennreal)" by transfer (metis ereal_diff_gr0 le_cases max.absorb2 not_less) lemma ennreal_mono_minus_cancel: fixes a b c :: ennreal shows "a - b \ a - c \ a < top \ b \ a \ c \ a \ c \ b" by transfer (auto simp add: max.absorb2 ereal_diff_positive top_ereal_def dest: ereal_mono_minus_cancel) lemma ennreal_mono_minus: fixes a b c :: ennreal shows "c \ b \ a - b \ a - c" by transfer (meson ereal_minus_mono max.mono order_refl) lemma ennreal_minus_pos_iff: fixes a b :: ennreal shows "a < top \ b < top \ 0 < a - b \ b < a" by transfer (use add.left_neutral ereal_minus_le_iff less_irrefl not_less in fastforce) lemma ennreal_inverse_top[simp]: "inverse top = (0::ennreal)" by transfer (simp add: top_ereal_def ereal_inverse_eq_0) lemma ennreal_inverse_zero[simp]: "inverse 0 = (top::ennreal)" by transfer (simp add: top_ereal_def ereal_inverse_eq_0) lemma ennreal_top_divide: "top / (x::ennreal) = (if x = top then 0 else top)" unfolding divide_ennreal_def by transfer (simp add: top_ereal_def ereal_inverse_eq_0 ereal_0_gt_inverse) lemma ennreal_zero_divide[simp]: "0 / (x::ennreal) = 0" by (simp add: divide_ennreal_def) lemma ennreal_divide_zero[simp]: "x / (0::ennreal) = (if x = 0 then 0 else top)" by (simp add: divide_ennreal_def ennreal_mult_top) lemma ennreal_divide_top[simp]: "x / (top::ennreal) = 0" by (simp add: divide_ennreal_def ennreal_top_mult) lemma ennreal_times_divide: "a * (b / c) = a * b / (c::ennreal)" unfolding divide_ennreal_def by transfer (simp add: divide_ereal_def[symmetric] ereal_times_divide_eq) lemma ennreal_zero_less_divide: "0 < a / b \ (0 < a \ b < (top::ennreal))" unfolding divide_ennreal_def by transfer (auto simp: ereal_zero_less_0_iff top_ereal_def ereal_0_gt_inverse) lemma add_divide_distrib_ennreal: "(a + b) / c = a / c + b / (c :: ennreal)" by (simp add: divide_ennreal_def ring_distribs) lemma divide_right_mono_ennreal: fixes a b c :: ennreal shows "a \ b \ a / c \ b / c" unfolding divide_ennreal_def by (intro mult_mono) auto lemma ennreal_mult_strict_right_mono: "(a::ennreal) < c \ 0 < b \ b < top \ a * b < c * b" by transfer (auto intro!: ereal_mult_strict_right_mono) lemma ennreal_indicator_less[simp]: "indicator A x \ (indicator B x::ennreal) \ (x \ A \ x \ B)" by (simp add: indicator_def not_le) lemma ennreal_inverse_positive: "0 < inverse x \ (x::ennreal) \ top" by transfer (simp add: ereal_0_gt_inverse top_ereal_def) lemma ennreal_inverse_mult': "((0 < b \ a < top) \ (0 < a \ b < top)) \ inverse (a * b::ennreal) = inverse a * inverse b" apply transfer subgoal for a b by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def) done lemma ennreal_inverse_mult: "a < top \ b < top \ inverse (a * b::ennreal) = inverse a * inverse b" apply transfer subgoal for a b by (cases a b rule: ereal2_cases) (auto simp: top_ereal_def) done lemma ennreal_inverse_1[simp]: "inverse (1::ennreal) = 1" by transfer simp lemma ennreal_inverse_eq_0_iff[simp]: "inverse (a::ennreal) = 0 \ a = top" by transfer (simp add: ereal_inverse_eq_0 top_ereal_def) lemma ennreal_inverse_eq_top_iff[simp]: "inverse (a::ennreal) = top \ a = 0" by transfer (simp add: top_ereal_def) lemma ennreal_divide_eq_0_iff[simp]: "(a::ennreal) / b = 0 \ (a = 0 \ b = top)" by (simp add: divide_ennreal_def) lemma ennreal_divide_eq_top_iff: "(a::ennreal) / b = top \ ((a \ 0 \ b = 0) \ (a = top \ b \ top))" by (auto simp add: divide_ennreal_def ennreal_mult_eq_top_iff) lemma one_divide_one_divide_ennreal[simp]: "1 / (1 / c) = (c::ennreal)" including ennreal.lifting unfolding divide_ennreal_def by transfer auto lemma ennreal_mult_left_cong: "((a::ennreal) \ 0 \ b = c) \ a * b = a * c" by (cases "a = 0") simp_all lemma ennreal_mult_right_cong: "((a::ennreal) \ 0 \ b = c) \ b * a = c * a" by (cases "a = 0") simp_all lemma ennreal_zero_less_mult_iff: "0 < a * b \ 0 < a \ 0 < (b::ennreal)" by transfer (auto simp add: ereal_zero_less_0_iff le_less) lemma less_diff_eq_ennreal: fixes a b c :: ennreal shows "b < top \ c < top \ a < b - c \ a + c < b" apply transfer subgoal for a b c by (cases a b c rule: ereal3_cases) (auto split: split_max) done lemma diff_add_cancel_ennreal: fixes a b :: ennreal shows "a \ b \ b - a + a = b" unfolding infinity_ennreal_def by transfer (metis (no_types) add.commute ereal_diff_positive ereal_ineq_diff_add max_def not_MInfty_nonneg) lemma ennreal_diff_self[simp]: "a \ top \ a - a = (0::ennreal)" by transfer (simp add: top_ereal_def) lemma ennreal_minus_mono: fixes a b c :: ennreal shows "a \ c \ d \ b \ a - b \ c - d" by transfer (meson ereal_minus_mono max.mono order_refl) lemma ennreal_minus_eq_top[simp]: "a - (b::ennreal) = top \ a = top" by (metis add_top diff_add_cancel_ennreal ennreal_mono_minus ennreal_top_minus zero_le) lemma ennreal_divide_self[simp]: "a \ 0 \ a < top \ a / a = (1::ennreal)" by (metis mult_1 mult_divide_eq_ennreal top.not_eq_extremum) subsection \Coercion from \<^typ>\real\ to \<^typ>\ennreal\\ lift_definition ennreal :: "real \ ennreal" is "sup 0 \ ereal" by simp declare [[coercion ennreal]] lemma ennreal_cong: "x = y \ ennreal x = ennreal y" by simp lemma ennreal_cases[cases type: ennreal]: fixes x :: ennreal obtains (real) r :: real where "0 \ r" "x = ennreal r" | (top) "x = top" apply transfer subgoal for x thesis by (cases x) (auto simp: max.absorb2 top_ereal_def) done lemmas ennreal2_cases = ennreal_cases[case_product ennreal_cases] lemmas ennreal3_cases = ennreal_cases[case_product ennreal2_cases] lemma ennreal_neq_top[simp]: "ennreal r \ top" by transfer (simp add: top_ereal_def zero_ereal_def flip: ereal_max) lemma top_neq_ennreal[simp]: "top \ ennreal r" using ennreal_neq_top[of r] by (auto simp del: ennreal_neq_top) lemma ennreal_less_top[simp]: "ennreal x < top" by transfer (simp add: top_ereal_def max_def) lemma ennreal_neg: "x \ 0 \ ennreal x = 0" by transfer (simp add: max.absorb1) lemma ennreal_inj[simp]: "0 \ a \ 0 \ b \ ennreal a = ennreal b \ a = b" by (transfer fixing: a b) (auto simp: max_absorb2) lemma ennreal_le_iff[simp]: "0 \ y \ ennreal x \ ennreal y \ x \ y" by (auto simp: ennreal_def zero_ereal_def less_eq_ennreal.abs_eq eq_onp_def split: split_max) lemma le_ennreal_iff: "0 \ r \ x \ ennreal r \ (\q\0. x = ennreal q \ q \ r)" by (cases x) (auto simp: top_unique) lemma ennreal_less_iff: "0 \ r \ ennreal r < ennreal q \ r < q" unfolding not_le[symmetric] by auto lemma ennreal_eq_zero_iff[simp]: "0 \ x \ ennreal x = 0 \ x = 0" by transfer (auto simp: max_absorb2) lemma ennreal_less_zero_iff[simp]: "0 < ennreal x \ 0 < x" by transfer (auto simp: max_def) lemma ennreal_lessI: "0 < q \ r < q \ ennreal r < ennreal q" by (cases "0 \ r") (auto simp: ennreal_less_iff ennreal_neg) lemma ennreal_leI: "x \ y \ ennreal x \ ennreal y" by (cases "0 \ y") (auto simp: ennreal_neg) lemma enn2ereal_ennreal[simp]: "0 \ x \ enn2ereal (ennreal x) = x" by transfer (simp add: max_absorb2) lemma e2ennreal_enn2ereal[simp]: "e2ennreal (enn2ereal x) = x" by (simp add: e2ennreal_def max_absorb2 ennreal.enn2ereal_inverse) lemma enn2ereal_e2ennreal: "x \ 0 \ enn2ereal (e2ennreal x) = x" by (metis e2ennreal_enn2ereal ereal_ennreal_cases not_le) lemma e2ennreal_ereal [simp]: "e2ennreal (ereal x) = ennreal x" by (metis e2ennreal_def enn2ereal_inverse ennreal.rep_eq sup_ereal_def) lemma ennreal_0[simp]: "ennreal 0 = 0" by (simp add: ennreal_def max.absorb1 zero_ennreal.abs_eq) lemma ennreal_1[simp]: "ennreal 1 = 1" by transfer (simp add: max_absorb2) lemma ennreal_eq_0_iff: "ennreal x = 0 \ x \ 0" by (cases "0 \ x") (auto simp: ennreal_neg) lemma ennreal_le_iff2: "ennreal x \ ennreal y \ ((0 \ y \ x \ y) \ (x \ 0 \ y \ 0))" by (cases "0 \ y") (auto simp: ennreal_eq_0_iff ennreal_neg) lemma ennreal_eq_1[simp]: "ennreal x = 1 \ x = 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma ennreal_le_1[simp]: "ennreal x \ 1 \ x \ 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma ennreal_ge_1[simp]: "ennreal x \ 1 \ x \ 1" by (cases "0 \ x") (auto simp: ennreal_neg simp flip: ennreal_1) lemma one_less_ennreal[simp]: "1 < ennreal x \ 1 < x" by (meson ennreal_le_1 linorder_not_le) lemma ennreal_plus[simp]: "0 \ a \ 0 \ b \ ennreal (a + b) = ennreal a + ennreal b" by (transfer fixing: a b) (auto simp: max_absorb2) lemma add_mono_ennreal: "x < ennreal y \ x' < ennreal y' \ x + x' < ennreal (y + y')" by (metis (full_types) add_strict_mono ennreal_less_zero_iff ennreal_plus less_le not_less zero_le) lemma sum_ennreal[simp]: "(\i. i \ I \ 0 \ f i) \ (\i\I. ennreal (f i)) = ennreal (sum f I)" by (induction I rule: infinite_finite_induct) (auto simp: sum_nonneg) lemma sum_list_ennreal[simp]: assumes "\x. x \ set xs \ f x \ 0" shows "sum_list (map (\x. ennreal (f x)) xs) = ennreal (sum_list (map f xs))" using assms proof (induction xs) case (Cons x xs) from Cons have "(\x\x # xs. ennreal (f x)) = ennreal (f x) + ennreal (sum_list (map f xs))" by simp also from Cons.prems have "\ = ennreal (f x + sum_list (map f xs))" by (intro ennreal_plus [symmetric] sum_list_nonneg) auto finally show ?case by simp qed simp_all lemma ennreal_of_nat_eq_real_of_nat: "of_nat i = ennreal (of_nat i)" by (induction i) simp_all lemma of_nat_le_ennreal_iff[simp]: "0 \ r \ of_nat i \ ennreal r \ of_nat i \ r" by (simp add: ennreal_of_nat_eq_real_of_nat) lemma ennreal_le_of_nat_iff[simp]: "ennreal r \ of_nat i \ r \ of_nat i" by (simp add: ennreal_of_nat_eq_real_of_nat) lemma ennreal_indicator: "ennreal (indicator A x) = indicator A x" by (auto split: split_indicator) lemma ennreal_numeral[simp]: "ennreal (numeral n) = numeral n" using ennreal_of_nat_eq_real_of_nat[of "numeral n"] by simp lemma ennreal_less_numeral_iff [simp]: "ennreal n < numeral w \ n < numeral w" by (metis ennreal_less_iff ennreal_numeral less_le not_less zero_less_numeral) lemma numeral_less_ennreal_iff [simp]: "numeral w < ennreal n \ numeral w < n" using ennreal_less_iff zero_le_numeral by fastforce lemma numeral_le_ennreal_iff [simp]: "numeral n \ ennreal m \ numeral n \ m" by (metis not_le ennreal_less_numeral_iff) lemma min_ennreal: "0 \ x \ 0 \ y \ min (ennreal x) (ennreal y) = ennreal (min x y)" by (auto split: split_min) lemma ennreal_half[simp]: "ennreal (1/2) = inverse 2" by transfer (simp add: max.absorb2) lemma ennreal_minus: "0 \ q \ ennreal r - ennreal q = ennreal (r - q)" by transfer (simp add: max.absorb2 zero_ereal_def flip: ereal_max) lemma ennreal_minus_top[simp]: "ennreal a - top = 0" by (simp add: minus_top_ennreal) lemma e2eenreal_enn2ereal_diff [simp]: "e2ennreal(enn2ereal x - enn2ereal y) = x - y" for x y by (cases x, cases y, auto simp add: ennreal_minus e2ennreal_neg) lemma ennreal_mult: "0 \ a \ 0 \ b \ ennreal (a * b) = ennreal a * ennreal b" by transfer (simp add: max_absorb2) lemma ennreal_mult': "0 \ a \ ennreal (a * b) = ennreal a * ennreal b" by (cases "0 \ b") (auto simp: ennreal_mult ennreal_neg mult_nonneg_nonpos) lemma indicator_mult_ennreal: "indicator A x * ennreal r = ennreal (indicator A x * r)" by (simp split: split_indicator) lemma ennreal_mult'': "0 \ b \ ennreal (a * b) = ennreal a * ennreal b" by (cases "0 \ a") (auto simp: ennreal_mult ennreal_neg mult_nonpos_nonneg) lemma numeral_mult_ennreal: "0 \ x \ numeral b * ennreal x = ennreal (numeral b * x)" by (simp add: ennreal_mult) lemma ennreal_power: "0 \ r \ ennreal r ^ n = ennreal (r ^ n)" by (induction n) (auto simp: ennreal_mult) lemma power_eq_top_ennreal: "x ^ n = top \ (n \ 0 \ (x::ennreal) = top)" by (cases x rule: ennreal_cases) (auto simp: ennreal_power top_power_ennreal) lemma inverse_ennreal: "0 < r \ inverse (ennreal r) = ennreal (inverse r)" by transfer (simp add: max.absorb2) lemma divide_ennreal: "0 \ r \ 0 < q \ ennreal r / ennreal q = ennreal (r / q)" by (simp add: divide_ennreal_def inverse_ennreal ennreal_mult[symmetric] inverse_eq_divide) lemma ennreal_inverse_power: "inverse (x ^ n :: ennreal) = inverse x ^ n" proof (cases x rule: ennreal_cases) case top with power_eq_top_ennreal[of x n] show ?thesis by (cases "n = 0") auto next case (real r) then show ?thesis proof (cases "x = 0") case False then show ?thesis by (smt (verit, best) ennreal_0 ennreal_power inverse_ennreal inverse_nonnegative_iff_nonnegative power_inverse real zero_less_power) qed (simp add: top_power_ennreal) qed lemma power_divide_distrib_ennreal [algebra_simps]: "(x / y) ^ n = x ^ n / (y ^ n :: ennreal)" by (simp add: divide_ennreal_def algebra_simps ennreal_inverse_power) lemma ennreal_divide_numeral: "0 \ x \ ennreal x / numeral b = ennreal (x / numeral b)" by (subst divide_ennreal[symmetric]) auto lemma prod_ennreal: "(\i. i \ A \ 0 \ f i) \ (\i\A. ennreal (f i)) = ennreal (prod f A)" by (induction A rule: infinite_finite_induct) (auto simp: ennreal_mult prod_nonneg) lemma prod_mono_ennreal: assumes "\x. x \ A \ f x \ (g x :: ennreal)" shows "prod f A \ prod g A" using assms by (induction A rule: infinite_finite_induct) (auto intro!: mult_mono) lemma mult_right_ennreal_cancel: "a * ennreal c = b * ennreal c \ (a = b \ c \ 0)" proof (cases "0 \ c") case True then show ?thesis by (metis ennreal_eq_0_iff ennreal_mult_right_cong ennreal_neq_top mult_divide_eq_ennreal) qed (use ennreal_neg in auto) lemma ennreal_le_epsilon: "(\e::real. y < top \ 0 < e \ x \ y + ennreal e) \ x \ y" apply (cases y rule: ennreal_cases) apply (cases x rule: ennreal_cases) apply (auto simp flip: ennreal_plus simp add: top_unique intro: zero_less_one field_le_epsilon) done lemma ennreal_rat_dense: fixes x y :: ennreal shows "x < y \ \r::rat. x < real_of_rat r \ real_of_rat r < y" proof transfer fix x y :: ereal assume xy: "0 \ x" "0 \ y" "x < y" moreover from ereal_dense3[OF \x < y\] obtain r where r: "x < ereal (real_of_rat r)" "ereal (real_of_rat r) < y" by auto then have "0 \ r" using le_less_trans[OF \0 \ x\ \x < ereal (real_of_rat r)\] by auto with r show "\r. x < (sup 0 \ ereal) (real_of_rat r) \ (sup 0 \ ereal) (real_of_rat r) < y" by (intro exI[of _ r]) (auto simp: max_absorb2) qed lemma ennreal_Ex_less_of_nat: "(x::ennreal) < top \ \n. x < of_nat n" by (cases x rule: ennreal_cases) (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_less_iff reals_Archimedean2) subsection \Coercion from \<^typ>\ennreal\ to \<^typ>\real\\ definition "enn2real x = real_of_ereal (enn2ereal x)" lemma enn2real_nonneg[simp]: "0 \ enn2real x" by (auto simp: enn2real_def intro!: real_of_ereal_pos enn2ereal_nonneg) lemma enn2real_mono: "a \ b \ b < top \ enn2real a \ enn2real b" by (auto simp add: enn2real_def less_eq_ennreal.rep_eq intro!: real_of_ereal_positive_mono enn2ereal_nonneg) lemma enn2real_of_nat[simp]: "enn2real (of_nat n) = n" by (auto simp: enn2real_def) lemma enn2real_ennreal[simp]: "0 \ r \ enn2real (ennreal r) = r" by (simp add: enn2real_def) lemma ennreal_enn2real[simp]: "r < top \ ennreal (enn2real r) = r" by (cases r rule: ennreal_cases) auto lemma real_of_ereal_enn2ereal[simp]: "real_of_ereal (enn2ereal x) = enn2real x" by (simp add: enn2real_def) lemma enn2real_top[simp]: "enn2real top = 0" unfolding enn2real_def top_ennreal.rep_eq top_ereal_def by simp lemma enn2real_0[simp]: "enn2real 0 = 0" unfolding enn2real_def zero_ennreal.rep_eq by simp lemma enn2real_1[simp]: "enn2real 1 = 1" unfolding enn2real_def one_ennreal.rep_eq by simp lemma enn2real_numeral[simp]: "enn2real (numeral n) = (numeral n)" unfolding enn2real_def by simp lemma enn2real_mult: "enn2real (a * b) = enn2real a * enn2real b" unfolding enn2real_def by (simp del: real_of_ereal_enn2ereal add: times_ennreal.rep_eq) lemma enn2real_leI: "0 \ B \ x \ ennreal B \ enn2real x \ B" by (cases x rule: ennreal_cases) (auto simp: top_unique) lemma enn2real_positive_iff: "0 < enn2real x \ (0 < x \ x < top)" by (cases x rule: ennreal_cases) auto lemma enn2real_eq_posreal_iff[simp]: "c > 0 \ enn2real x = c \ x = c" by (cases x) auto lemma ennreal_enn2real_if: "ennreal (enn2real r) = (if r = top then 0 else r)" by(auto intro!: ennreal_enn2real simp add: less_top) subsection \Coercion from \<^typ>\enat\ to \<^typ>\ennreal\\ definition ennreal_of_enat :: "enat \ ennreal" where "ennreal_of_enat n = (case n of \ \ top | enat n \ of_nat n)" declare [[coercion ennreal_of_enat]] declare [[coercion "of_nat :: nat \ ennreal"]] lemma ennreal_of_enat_infty[simp]: "ennreal_of_enat \ = \" by (simp add: ennreal_of_enat_def) lemma ennreal_of_enat_enat[simp]: "ennreal_of_enat (enat n) = of_nat n" by (simp add: ennreal_of_enat_def) lemma ennreal_of_enat_0[simp]: "ennreal_of_enat 0 = 0" using ennreal_of_enat_enat[of 0] unfolding enat_0 by simp lemma ennreal_of_enat_1[simp]: "ennreal_of_enat 1 = 1" using ennreal_of_enat_enat[of 1] unfolding enat_1 by simp lemma ennreal_top_neq_of_nat[simp]: "(top::ennreal) \ of_nat i" using ennreal_of_nat_neq_top[of i] by metis lemma ennreal_of_enat_inj[simp]: "ennreal_of_enat i = ennreal_of_enat j \ i = j" by (cases i j rule: enat.exhaust[case_product enat.exhaust]) auto lemma ennreal_of_enat_le_iff[simp]: "ennreal_of_enat m \ ennreal_of_enat n \ m \ n" by (auto simp: ennreal_of_enat_def top_unique split: enat.split) lemma of_nat_less_ennreal_of_nat[simp]: "of_nat n \ ennreal_of_enat x \ of_nat n \ x" by (cases x) (auto simp: of_nat_eq_enat) lemma ennreal_of_enat_Sup: "ennreal_of_enat (Sup X) = (SUP x\X. ennreal_of_enat x)" proof - have "ennreal_of_enat (Sup X) \ (SUP x \ X. ennreal_of_enat x)" unfolding Sup_enat_def proof (clarsimp, intro conjI impI) fix x assume "finite X" "X \ {}" then show "ennreal_of_enat (Max X) \ (SUP x \ X. ennreal_of_enat x)" by (intro SUP_upper Max_in) next assume "infinite X" "X \ {}" have "\y\X. r < ennreal_of_enat y" if r: "r < top" for r proof - obtain n where n: "r < of_nat n" using ennreal_Ex_less_of_nat[OF r] .. have "\ (X \ enat ` {.. n})" using \infinite X\ by (auto dest: finite_subset) then obtain x where x: "x \ X" "x \ enat ` {..n}" by blast then have "of_nat n \ x" by (cases x) (auto simp: of_nat_eq_enat) with x show ?thesis by (auto intro!: bexI[of _ x] less_le_trans[OF n]) qed then have "(SUP x \ X. ennreal_of_enat x) = top" by simp then show "top \ (SUP x \ X. ennreal_of_enat x)" unfolding top_unique by simp qed then show ?thesis by (auto intro!: antisym Sup_least intro: Sup_upper) qed lemma ennreal_of_enat_eSuc[simp]: "ennreal_of_enat (eSuc x) = 1 + ennreal_of_enat x" by (cases x) (auto simp: eSuc_enat) (* Contributed by Dominique Unruh *) lemma ennreal_of_enat_plus[simp]: \ennreal_of_enat (a+b) = ennreal_of_enat a + ennreal_of_enat b\ apply (induct a) apply (metis enat.exhaust ennreal_add_eq_top ennreal_of_enat_enat ennreal_of_enat_infty infinity_ennreal_def of_nat_add plus_enat_simps(1) plus_eq_infty_iff_enat) apply simp done (* Contributed by Dominique Unruh *) lemma sum_ennreal_of_enat[simp]: "(\i\I. ennreal_of_enat (f i)) = ennreal_of_enat (sum f I)" by (induct I rule: infinite_finite_induct) (auto simp: sum_nonneg) subsection \Topology on \<^typ>\ennreal\\ lemma enn2ereal_Iio: "enn2ereal -` {.. a then {..< e2ennreal a} else {})" using enn2ereal_nonneg by (cases a rule: ereal_ennreal_cases) (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 simp del: enn2ereal_nonneg intro: le_less_trans less_imp_le) lemma enn2ereal_Ioi: "enn2ereal -` {a <..} = (if 0 \ a then {e2ennreal a <..} else UNIV)" by (cases a rule: ereal_ennreal_cases) (auto simp add: vimage_def set_eq_iff ennreal.enn2ereal_inverse less_ennreal.rep_eq e2ennreal_def max_absorb2 intro: less_le_trans) instantiation ennreal :: linear_continuum_topology begin definition open_ennreal :: "ennreal set \ bool" where "(open :: ennreal set \ bool) = generate_topology (range lessThan \ range greaterThan)" instance proof show "\a b::ennreal. a \ b" using zero_neq_one by (intro exI) show "\x y::ennreal. x < y \ \z>x. z < y" proof transfer fix x y :: ereal assume *: "0 \ x" assume "x < y" from dense[OF this] obtain z where "x < z \ z < y" .. with * show "\z\Collect ((\) 0). x < z \ z < y" by (intro bexI[of _ z]) auto qed qed (rule open_ennreal_def) end lemma continuous_on_e2ennreal: "continuous_on A e2ennreal" proof (rule continuous_on_subset) show "continuous_on ({0..} \ {..0}) e2ennreal" proof (rule continuous_on_closed_Un) show "continuous_on {0 ..} e2ennreal" by (rule continuous_onI_mono) (auto simp add: less_eq_ennreal.abs_eq eq_onp_def enn2ereal_range) show "continuous_on {.. 0} e2ennreal" by (subst continuous_on_cong[OF refl, of _ _ "\_. 0"]) (auto simp add: e2ennreal_neg continuous_on_const) qed auto show "A \ {0..} \ {..0::ereal}" by auto qed lemma continuous_at_e2ennreal: "continuous (at x within A) e2ennreal" by (rule continuous_on_imp_continuous_within[OF continuous_on_e2ennreal, of _ UNIV]) auto lemma continuous_on_enn2ereal: "continuous_on UNIV enn2ereal" by (rule continuous_on_generate_topology[OF open_generated_order]) (auto simp add: enn2ereal_Iio enn2ereal_Ioi) lemma continuous_at_enn2ereal: "continuous (at x within A) enn2ereal" by (rule continuous_on_imp_continuous_within[OF continuous_on_enn2ereal]) auto lemma sup_continuous_e2ennreal[order_continuous_intros]: assumes f: "sup_continuous f" shows "sup_continuous (\x. e2ennreal (f x))" proof (rule sup_continuous_compose[OF _ f]) show "sup_continuous e2ennreal" by (simp add: continuous_at_e2ennreal continuous_at_left_imp_sup_continuous e2ennreal_mono mono_def) qed lemma sup_continuous_enn2ereal[order_continuous_intros]: assumes f: "sup_continuous f" shows "sup_continuous (\x. enn2ereal (f x))" proof (rule sup_continuous_compose[OF _ f]) show "sup_continuous enn2ereal" by (simp add: continuous_at_enn2ereal continuous_at_left_imp_sup_continuous less_eq_ennreal.rep_eq mono_def) qed lemma sup_continuous_mult_left_ennreal': fixes c :: "ennreal" shows "sup_continuous (\x. c * x)" unfolding sup_continuous_def by transfer (auto simp: SUP_ereal_mult_left max.absorb2 SUP_upper2) lemma sup_continuous_mult_left_ennreal[order_continuous_intros]: "sup_continuous f \ sup_continuous (\x. c * f x :: ennreal)" by (rule sup_continuous_compose[OF sup_continuous_mult_left_ennreal']) lemma sup_continuous_mult_right_ennreal[order_continuous_intros]: "sup_continuous f \ sup_continuous (\x. f x * c :: ennreal)" using sup_continuous_mult_left_ennreal[of f c] by (simp add: mult.commute) lemma sup_continuous_divide_ennreal[order_continuous_intros]: fixes f g :: "'a::complete_lattice \ ennreal" shows "sup_continuous f \ sup_continuous (\x. f x / c)" unfolding divide_ennreal_def by (rule sup_continuous_mult_right_ennreal) lemma transfer_enn2ereal_continuous_on [transfer_rule]: "rel_fun (=) (rel_fun (rel_fun (=) pcr_ennreal) (=)) continuous_on continuous_on" proof - have "continuous_on A f" if "continuous_on A (\x. enn2ereal (f x))" for A and f :: "'a \ ennreal" using continuous_on_compose2[OF continuous_on_e2ennreal[of "{0..}"] that] by (auto simp: ennreal.enn2ereal_inverse subset_eq e2ennreal_def max_absorb2) moreover have "continuous_on A (\x. enn2ereal (f x))" if "continuous_on A f" for A and f :: "'a \ ennreal" using continuous_on_compose2[OF continuous_on_enn2ereal that] by auto ultimately show ?thesis by (auto simp add: rel_fun_def ennreal.pcr_cr_eq cr_ennreal_def) qed lemma transfer_sup_continuous[transfer_rule]: "(rel_fun (rel_fun (=) pcr_ennreal) (=)) sup_continuous sup_continuous" proof (safe intro!: rel_funI dest!: rel_fun_eq_pcr_ennreal[THEN iffD1]) show "sup_continuous (enn2ereal \ f) \ sup_continuous f" for f :: "'a \ _" using sup_continuous_e2ennreal[of "enn2ereal \ f"] by simp show "sup_continuous f \ sup_continuous (enn2ereal \ f)" for f :: "'a \ _" using sup_continuous_enn2ereal[of f] by (simp add: comp_def) qed lemma continuous_on_ennreal[tendsto_intros]: "continuous_on A f \ continuous_on A (\x. ennreal (f x))" by transfer (auto intro!: continuous_on_max continuous_on_const continuous_on_ereal) lemma tendsto_ennrealD: assumes lim: "((\x. ennreal (f x)) \ ennreal x) F" assumes *: "\\<^sub>F x in F. 0 \ f x" and x: "0 \ x" shows "(f \ x) F" proof - have "((\x. enn2ereal (ennreal (f x))) \ enn2ereal (ennreal x)) F \ (f \ enn2ereal (ennreal x)) F" using "*" eventually_mono by (intro tendsto_cong) fastforce then show ?thesis using assms(1) continuous_at_enn2ereal isCont_tendsto_compose x by fastforce qed lemma tendsto_ennreal_iff [simp]: \((\x. ennreal (f x)) \ ennreal x) F \ (f \ x) F\ (is \?P \ ?Q\) if \\\<^sub>F x in F. 0 \ f x\ \0 \ x\ proof assume \?P\ then show \?Q\ using that by (rule tendsto_ennrealD) next assume \?Q\ have \continuous_on UNIV ereal\ using continuous_on_ereal [of _ id] by simp then have \continuous_on UNIV (e2ennreal \ ereal)\ by (rule continuous_on_compose) (simp_all add: continuous_on_e2ennreal) then have \((\x. (e2ennreal \ ereal) (f x)) \ (e2ennreal \ ereal) x) F\ using \?Q\ by (rule continuous_on_tendsto_compose) simp_all then show \?P\ by (simp flip: e2ennreal_ereal) qed lemma tendsto_enn2ereal_iff[simp]: "((\i. enn2ereal (f i)) \ enn2ereal x) F \ (f \ x) F" using continuous_on_enn2ereal[THEN continuous_on_tendsto_compose, of f x F] continuous_on_e2ennreal[THEN continuous_on_tendsto_compose, of "\x. enn2ereal (f x)" "enn2ereal x" F UNIV] by auto lemma ennreal_tendsto_0_iff: "(\n. f n \ 0) \ ((\n. ennreal (f n)) \ 0) \ (f \ 0)" by (metis (mono_tags) ennreal_0 eventuallyI order_refl tendsto_ennreal_iff) lemma continuous_on_add_ennreal: fixes f g :: "'a::topological_space \ ennreal" shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. f x + g x)" by (transfer fixing: A) (auto intro!: tendsto_add_ereal_nonneg simp: continuous_on_def) lemma continuous_on_inverse_ennreal[continuous_intros]: fixes f :: "'a::topological_space \ ennreal" shows "continuous_on A f \ continuous_on A (\x. inverse (f x))" proof (transfer fixing: A) show "pred_fun top ((\) 0) f \ continuous_on A (\x. inverse (f x))" if "continuous_on A f" for f :: "'a \ ereal" using continuous_on_compose2[OF continuous_on_inverse_ereal that] by (auto simp: subset_eq) qed instance ennreal :: topological_comm_monoid_add proof show "((\x. fst x + snd x) \ a + b) (nhds a \\<^sub>F nhds b)" for a b :: ennreal using continuous_on_add_ennreal[of UNIV fst snd] using tendsto_at_iff_tendsto_nhds[symmetric, of "\x::(ennreal \ ennreal). fst x + snd x"] by (auto simp: continuous_on_eq_continuous_at) (simp add: isCont_def nhds_prod[symmetric]) qed lemma sup_continuous_add_ennreal[order_continuous_intros]: fixes f g :: "'a::complete_lattice \ ennreal" shows "sup_continuous f \ sup_continuous g \ sup_continuous (\x. f x + g x)" by transfer (auto intro!: sup_continuous_add) lemma ennreal_suminf_lessD: "(\i. f i :: ennreal) < x \ f i < x" using le_less_trans[OF sum_le_suminf[OF summableI, of "{i}" f]] by simp lemma sums_ennreal[simp]: "(\i. 0 \ f i) \ 0 \ x \ (\i. ennreal (f i)) sums ennreal x \ f sums x" unfolding sums_def by (simp add: always_eventually sum_nonneg) lemma summable_suminf_not_top: "(\i. 0 \ f i) \ (\i. ennreal (f i)) \ top \ summable f" using summable_sums[OF summableI, of "\i. ennreal (f i)"] by (cases "\i. ennreal (f i)" rule: ennreal_cases) (auto simp: summable_def) lemma suminf_ennreal[simp]: "(\i. 0 \ f i) \ (\i. ennreal (f i)) \ top \ (\i. ennreal (f i)) = ennreal (\i. f i)" by (rule sums_unique[symmetric]) (simp add: summable_suminf_not_top suminf_nonneg summable_sums) lemma sums_enn2ereal[simp]: "(\i. enn2ereal (f i)) sums enn2ereal x \ f sums x" unfolding sums_def by (simp add: always_eventually sum_nonneg) lemma suminf_enn2ereal[simp]: "(\i. enn2ereal (f i)) = enn2ereal (suminf f)" by (rule sums_unique[symmetric]) (simp add: summable_sums) lemma transfer_e2ennreal_suminf [transfer_rule]: "rel_fun (rel_fun (=) pcr_ennreal) pcr_ennreal suminf suminf" by (auto simp: rel_funI rel_fun_eq_pcr_ennreal comp_def) lemma ennreal_suminf_cmult[simp]: "(\i. r * f i) = r * (\i. f i::ennreal)" by transfer (auto intro!: suminf_cmult_ereal) lemma ennreal_suminf_multc[simp]: "(\i. f i * r) = (\i. f i::ennreal) * r" using ennreal_suminf_cmult[of r f] by (simp add: ac_simps) lemma ennreal_suminf_divide[simp]: "(\i. f i / r) = (\i. f i::ennreal) / r" by (simp add: divide_ennreal_def) lemma ennreal_suminf_neq_top: "summable f \ (\i. 0 \ f i) \ (\i. ennreal (f i)) \ top" using sums_ennreal[of f "suminf f"] by (simp add: suminf_nonneg flip: sums_unique summable_sums_iff del: sums_ennreal) lemma suminf_ennreal_eq: "(\i. 0 \ f i) \ f sums x \ (\i. ennreal (f i)) = ennreal x" using suminf_nonneg[of f] sums_unique[of f x] by (intro sums_unique[symmetric]) (auto simp: summable_sums_iff) lemma ennreal_suminf_bound_add: fixes f :: "nat \ ennreal" shows "(\N. (\n x) \ suminf f + y \ x" by transfer (auto intro!: suminf_bound_add) lemma ennreal_suminf_SUP_eq_directed: fixes f :: "'a \ nat \ ennreal" assumes *: "\N i j. i \ I \ j \ I \ finite N \ \k\I. \n\N. f i n \ f k n \ f j n \ f k n" shows "(\n. SUP i\I. f i n) = (SUP i\I. \n. f i n)" proof cases assume "I \ {}" then obtain i where "i \ I" by auto from * show ?thesis by (transfer fixing: I) (auto simp: max_absorb2 SUP_upper2[OF \i \ I\] suminf_nonneg summable_ereal_pos \I \ {}\ intro!: suminf_SUP_eq_directed) qed (simp add: bot_ennreal) lemma INF_ennreal_add_const: fixes f g :: "nat \ ennreal" shows "(INF i. f i + c) = (INF i. f i) + c" using continuous_at_Inf_mono[of "\x. x + c" "f`UNIV"] using continuous_add[of "at_right (Inf (range f))", of "\x. x" "\x. c"] by (auto simp: mono_def image_comp) lemma INF_ennreal_const_add: fixes f g :: "nat \ ennreal" shows "(INF i. c + f i) = c + (INF i. f i)" using INF_ennreal_add_const[of f c] by (simp add: ac_simps) lemma SUP_mult_left_ennreal: "c * (SUP i\I. f i) = (SUP i\I. c * f i ::ennreal)" proof cases assume "I \ {}" then show ?thesis by transfer (auto simp add: SUP_ereal_mult_left max_absorb2 SUP_upper2) qed (simp add: bot_ennreal) lemma SUP_mult_right_ennreal: "(SUP i\I. f i) * c = (SUP i\I. f i * c ::ennreal)" using SUP_mult_left_ennreal by (simp add: mult.commute) lemma SUP_divide_ennreal: "(SUP i\I. f i) / c = (SUP i\I. f i / c ::ennreal)" using SUP_mult_right_ennreal by (simp add: divide_ennreal_def) lemma ennreal_SUP_of_nat_eq_top: "(SUP x. of_nat x :: ennreal) = top" proof (intro antisym top_greatest le_SUP_iff[THEN iffD2] allI impI) fix y :: ennreal assume "y < top" then obtain r where "y = ennreal r" by (cases y rule: ennreal_cases) auto then show "\i\UNIV. y < of_nat i" using reals_Archimedean2[of "max 1 r"] zero_less_one by (simp add: ennreal_Ex_less_of_nat) qed lemma ennreal_SUP_eq_top: fixes f :: "'a \ ennreal" assumes "\n. \i\I. of_nat n \ f i" shows "(SUP i \ I. f i) = top" proof - have "(SUP x. of_nat x :: ennreal) \ (SUP i \ I. f i)" using assms by (auto intro!: SUP_least intro: SUP_upper2) then show ?thesis by (auto simp: ennreal_SUP_of_nat_eq_top top_unique) qed lemma ennreal_INF_const_minus: fixes f :: "'a \ ennreal" shows "I \ {} \ (SUP x\I. c - f x) = c - (INF x\I. f x)" by (transfer fixing: I) (simp add: sup_max[symmetric] SUP_sup_const1 SUP_ereal_minus_right del: sup_ereal_def) lemma of_nat_Sup_ennreal: assumes "A \ {}" "bdd_above A" shows "of_nat (Sup A) = (SUP a\A. of_nat a :: ennreal)" proof (intro antisym) show "(SUP a\A. of_nat a::ennreal) \ of_nat (Sup A)" by (intro SUP_least of_nat_mono) (auto intro: cSup_upper assms) have "Sup A \ A" using assms by (auto simp: Sup_nat_def bdd_above_nat) then show "of_nat (Sup A) \ (SUP a\A. of_nat a::ennreal)" by (intro SUP_upper) qed lemma ennreal_tendsto_const_minus: fixes g :: "'a \ ennreal" assumes ae: "\\<^sub>F x in F. g x \ c" assumes g: "((\x. c - g x) \ 0) F" shows "(g \ c) F" proof (cases c rule: ennreal_cases) case top with tendsto_unique[OF _ g, of "top"] show ?thesis by (cases "F = bot") auto next case (real r) then have "\x. \q\0. g x \ c \ (g x = ennreal q \ q \ r)" by (auto simp: le_ennreal_iff) then obtain f where *: "0 \ f x" "g x = ennreal (f x)" "f x \ r" if "g x \ c" for x by metis from ae have ae2: "\\<^sub>F x in F. c - g x = ennreal (r - f x) \ f x \ r \ g x = ennreal (f x) \ 0 \ f x" proof eventually_elim fix x assume "g x \ c" with *[of x] \0 \ r\ show "c - g x = ennreal (r - f x) \ f x \ r \ g x = ennreal (f x) \ 0 \ f x" by (auto simp: real ennreal_minus) qed with g have "((\x. ennreal (r - f x)) \ ennreal 0) F" by (auto simp add: tendsto_cong eventually_conj_iff) with ae2 have "((\x. r - f x) \ 0) F" by (subst (asm) tendsto_ennreal_iff) (auto elim: eventually_mono) then have "(f \ r) F" by (rule Lim_transform2[OF tendsto_const]) with ae2 have "((\x. ennreal (f x)) \ ennreal r) F" by (subst tendsto_ennreal_iff) (auto elim: eventually_mono simp: real) with ae2 show ?thesis by (auto simp: real tendsto_cong eventually_conj_iff) qed lemma ennreal_SUP_add: fixes f g :: "nat \ ennreal" shows "incseq f \ incseq g \ (SUP i. f i + g i) = Sup (f ` UNIV) + Sup (g ` UNIV)" unfolding incseq_def le_fun_def by transfer (simp add: SUP_ereal_add incseq_def le_fun_def max_absorb2 SUP_upper2) lemma ennreal_SUP_sum: fixes f :: "'a \ nat \ ennreal" shows "(\i. i \ I \ incseq (f i)) \ (SUP n. \i\I. f i n) = (\i\I. SUP n. f i n)" unfolding incseq_def by transfer (simp add: SUP_ereal_sum incseq_def SUP_upper2 max_absorb2 sum_nonneg) lemma ennreal_liminf_minus: fixes f :: "nat \ ennreal" shows "(\n. f n \ c) \ liminf (\n. c - f n) = c - limsup f" apply transfer apply (simp add: ereal_diff_positive liminf_ereal_cminus) by (metis max.absorb2 ereal_diff_positive Limsup_bounded eventually_sequentiallyI) lemma ennreal_continuous_on_cmult: "(c::ennreal) < top \ continuous_on A f \ continuous_on A (\x. c * f x)" by (transfer fixing: A) (auto intro: continuous_on_cmult_ereal) lemma ennreal_tendsto_cmult: "(c::ennreal) < top \ (f \ x) F \ ((\x. c * f x) \ c * x) F" by (rule continuous_on_tendsto_compose[where g=f, OF ennreal_continuous_on_cmult, where s=UNIV]) (auto simp: continuous_on_id) lemma tendsto_ennrealI[intro, simp, tendsto_intros]: "(f \ x) F \ ((\x. ennreal (f x)) \ ennreal x) F" by (auto simp: ennreal_def intro!: continuous_on_tendsto_compose[OF continuous_on_e2ennreal[of UNIV]] tendsto_max) lemma tendsto_enn2erealI [tendsto_intros]: assumes "(f \ l) F" shows "((\i. enn2ereal(f i)) \ enn2ereal l) F" using tendsto_enn2ereal_iff assms by auto lemma tendsto_e2ennrealI [tendsto_intros]: assumes "(f \ l) F" shows "((\i. e2ennreal(f i)) \ e2ennreal l) F" proof - have *: "e2ennreal (max x 0) = e2ennreal x" for x by (simp add: e2ennreal_def max.commute) have "((\i. max (f i) 0) \ max l 0) F" apply (intro tendsto_intros) using assms by auto then have "((\i. enn2ereal(e2ennreal (max (f i) 0))) \ enn2ereal (e2ennreal (max l 0))) F" by (subst enn2ereal_e2ennreal, auto)+ then have "((\i. e2ennreal (max (f i) 0)) \ e2ennreal (max l 0)) F" using tendsto_enn2ereal_iff by auto then show ?thesis unfolding * by auto qed lemma ennreal_suminf_minus: fixes f g :: "nat \ ennreal" shows "(\i. g i \ f i) \ suminf f \ top \ suminf g \ top \ (\i. f i - g i) = suminf f - suminf g" by transfer (auto simp add: max.absorb2 ereal_diff_positive suminf_le_pos top_ereal_def intro!: suminf_ereal_minus) lemma ennreal_Sup_countable_SUP: "A \ {} \ \f::nat \ ennreal. incseq f \ range f \ A \ Sup A = (SUP i. f i)" unfolding incseq_def apply transfer subgoal for A using Sup_countable_SUP[of A] by (force simp add: incseq_def[symmetric] SUP_upper2 max.absorb2 image_subset_iff Sup_upper2 cong: conj_cong) done lemma ennreal_Inf_countable_INF: "A \ {} \ \f::nat \ ennreal. decseq f \ range f \ A \ Inf A = (INF i. f i)" unfolding decseq_def apply transfer subgoal for A using Inf_countable_INF[of A] apply (clarsimp simp flip: decseq_def) subgoal for f by (intro exI[of _ f]) auto done done lemma ennreal_SUP_countable_SUP: "A \ {} \ \f::nat \ ennreal. range f \ g`A \ Sup (g ` A) = Sup (f ` UNIV)" using ennreal_Sup_countable_SUP [of "g`A"] by auto lemma of_nat_tendsto_top_ennreal: "(\n::nat. of_nat n :: ennreal) \ top" using LIMSEQ_SUP[of "of_nat :: nat \ ennreal"] by (simp add: ennreal_SUP_of_nat_eq_top incseq_def) lemma SUP_sup_continuous_ennreal: fixes f :: "ennreal \ 'a::complete_lattice" assumes f: "sup_continuous f" and "I \ {}" shows "(SUP i\I. f (g i)) = f (SUP i\I. g i)" proof (rule antisym) show "(SUP i\I. f (g i)) \ f (SUP i\I. g i)" by (rule mono_SUP[OF sup_continuous_mono[OF f]]) from ennreal_Sup_countable_SUP[of "g`I"] \I \ {}\ obtain M :: "nat \ ennreal" where "incseq M" and M: "range M \ g ` I" and eq: "(SUP i \ I. g i) = (SUP i. M i)" by auto have "f (SUP i \ I. g i) = (SUP i \ range M. f i)" unfolding eq sup_continuousD[OF f \mono M\] by (simp add: image_comp) also have "\ \ (SUP i \ I. f (g i))" by (insert M, drule SUP_subset_mono) (auto simp add: image_comp) finally show "f (SUP i \ I. g i) \ (SUP i \ I. f (g i))" . qed lemma ennreal_suminf_SUP_eq: fixes f :: "nat \ nat \ ennreal" shows "(\i. incseq (\n. f n i)) \ (\i. SUP n. f n i) = (SUP n. \i. f n i)" apply (rule ennreal_suminf_SUP_eq_directed) subgoal for N n j by (auto simp: incseq_def intro!:exI[of _ "max n j"]) done lemma ennreal_SUP_add_left: fixes c :: ennreal shows "I \ {} \ (SUP i\I. f i + c) = (SUP i\I. f i) + c" apply transfer apply (simp add: SUP_ereal_add_left) by (metis SUP_upper all_not_in_conv ereal_le_add_mono1 max.absorb2 max.bounded_iff) lemma ennreal_SUP_const_minus: fixes f :: "'a \ ennreal" shows "I \ {} \ c < top \ (INF x\I. c - f x) = c - (SUP x\I. f x)" apply (transfer fixing: I) unfolding ex_in_conv[symmetric] apply (auto simp add: SUP_upper2 sup_absorb2 simp flip: sup_ereal_def) apply (subst INF_ereal_minus_right[symmetric]) apply (auto simp del: sup_ereal_def simp add: sup_INF) done (* Contributed by Dominique Unruh *) lemma isCont_ennreal[simp]: \isCont ennreal x\ apply (auto intro!: sequentially_imp_eventually_within simp: continuous_within tendsto_def) by (metis tendsto_def tendsto_ennrealI) (* Contributed by Dominique Unruh *) lemma isCont_ennreal_of_enat[simp]: \isCont ennreal_of_enat x\ proof - have continuous_at_open: \ \Copied lemma from \<^session>\HOL-Analysis\ to avoid dependency.\ "continuous (at x) f \ (\t. open t \ f x \ t --> (\s. open s \ x \ s \ (\x' \ s. (f x') \ t)))" for f :: \enat \ 'z::topological_space\ unfolding continuous_within_topological [of x UNIV f] unfolding imp_conjL by (intro all_cong imp_cong ex_cong conj_cong refl) auto show ?thesis proof (subst continuous_at_open, intro allI impI, cases \x = \\) case True fix t assume \open t \ ennreal_of_enat x \ t\ then have \\y<\. {y <.. \} \ t\ by (rule_tac open_left[where y=0]) (auto simp: True) then obtain y where \{y<..} \ t\ and \y \ \\ by fastforce from \y \ \\ obtain x' where x'y: \ennreal_of_enat x' > y\ and \x' \ \\ by (metis enat.simps(3) ennreal_Ex_less_of_nat ennreal_of_enat_enat infinity_ennreal_def top.not_eq_extremum) define s where \s = {x'<..}\ have \open s\ by (simp add: s_def) moreover have \x \ s\ by (simp add: \x' \ \\ s_def True) moreover have \ennreal_of_enat z \ t\ if \z \ s\ for z by (metis x'y \{y<..} \ t\ ennreal_of_enat_le_iff greaterThan_iff le_less_trans less_imp_le not_less s_def subsetD that) ultimately show \\s. open s \ x \ s \ (\z\s. ennreal_of_enat z \ t)\ by auto next case False fix t assume asm: \open t \ ennreal_of_enat x \ t\ define s where \s = {x}\ have \open s\ using False open_enat_iff s_def by blast moreover have \x \ s\ using s_def by auto moreover have \ennreal_of_enat z \ t\ if \z \ s\ for z using asm s_def that by blast ultimately show \\s. open s \ x \ s \ (\z\s. ennreal_of_enat z \ t)\ by auto qed qed subsection \Approximation lemmas\ lemma INF_approx_ennreal: fixes x::ennreal and e::real assumes "e > 0" assumes INF: "x = (INF i \ A. f i)" assumes "x \ \" shows "\i \ A. f i < x + e" proof - have "(INF i \ A. f i) < x + e" unfolding INF[symmetric] using \0 \x \ \\ by (cases x) auto then show ?thesis unfolding INF_less_iff . qed lemma SUP_approx_ennreal: fixes x::ennreal and e::real assumes "e > 0" "A \ {}" assumes SUP: "x = (SUP i \ A. f i)" assumes "x \ \" shows "\i \ A. x < f i + e" proof - have "x < x + e" using \0 \x \ \\ by (cases x) auto also have "x + e = (SUP i \ A. f i + e)" unfolding SUP ennreal_SUP_add_left[OF \A \ {}\] .. finally show ?thesis unfolding less_SUP_iff . qed lemma ennreal_approx_SUP: fixes x::ennreal assumes f_bound: "\i. i \ A \ f i \ x" assumes approx: "\e. (e::real) > 0 \ \i \ A. x \ f i + e" shows "x = (SUP i \ A. f i)" proof (rule antisym) show "x \ (SUP i\A. f i)" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" from approx[OF this] obtain i where "i \ A" and *: "x \ f i + ennreal e" by blast from * have "x \ f i + e" by simp also have "\ \ (SUP i\A. f i) + e" by (intro add_mono \i \ A\ SUP_upper order_refl) finally show "x \ (SUP i\A. f i) + e" . qed qed (intro SUP_least f_bound) lemma ennreal_approx_INF: fixes x::ennreal assumes f_bound: "\i. i \ A \ x \ f i" assumes approx: "\e. (e::real) > 0 \ \i \ A. f i \ x + e" shows "x = (INF i \ A. f i)" proof (rule antisym) show "(INF i\A. f i) \ x" proof (rule ennreal_le_epsilon) fix e :: real assume "0 < e" from approx[OF this] obtain i where "i\A" "f i \ x + ennreal e" by blast then have "(INF i\A. f i) \ f i" by (intro INF_lower) also have "\ \ x + e" by fact finally show "(INF i\A. f i) \ x + e" . qed qed (intro INF_greatest f_bound) lemma ennreal_approx_unit: "(\a::ennreal. 0 < a \ a < 1 \ a * z \ y) \ z \ y" apply (subst SUP_mult_right_ennreal[of "\x. x" "{0 <..< 1}" z, simplified]) apply (auto intro: SUP_least) done lemma suminf_ennreal2: "(\i. 0 \ f i) \ summable f \ (\i. ennreal (f i)) = ennreal (\i. f i)" using suminf_ennreal_eq by blast lemma less_top_ennreal: "x < top \ (\r\0. x = ennreal r)" by (cases x) auto lemma enn2real_less_iff[simp]: "x < top \ enn2real x < c \ x < c" using ennreal_less_iff less_top_ennreal by auto lemma enn2real_le_iff[simp]: "\x < top; c > 0\ \ enn2real x \ c \ x \ c" by (cases x) auto lemma enn2real_less: assumes "enn2real e < r" "e \ top" shows "e < ennreal r" using enn2real_less_iff assms top.not_eq_extremum by blast lemma enn2real_le: assumes "enn2real e \ r" "e \ top" shows "e \ ennreal r" by (metis assms enn2real_less ennreal_enn2real_if eq_iff less_le) lemma tendsto_top_iff_ennreal: fixes f :: "'a \ ennreal" shows "(f \ top) F \ (\l\0. eventually (\x. ennreal l < f x) F)" by (auto simp: less_top_ennreal order_tendsto_iff ) lemma ennreal_tendsto_top_eq_at_top: "((\z. ennreal (f z)) \ top) F \ (LIM z F. f z :> at_top)" unfolding filterlim_at_top_dense tendsto_top_iff_ennreal apply (auto simp: ennreal_less_iff) subgoal for y by (auto elim!: eventually_mono allE[of _ "max 0 y"]) done lemma tendsto_0_if_Limsup_eq_0_ennreal: fixes f :: "_ \ ennreal" shows "Limsup F f = 0 \ (f \ 0) F" using Liminf_le_Limsup[of F f] tendsto_iff_Liminf_eq_Limsup[of F f 0] by (cases "F = bot") auto lemma diff_le_self_ennreal[simp]: "a - b \ (a::ennreal)" by (cases a b rule: ennreal2_cases) (auto simp: ennreal_minus) lemma ennreal_ineq_diff_add: "b \ a \ a = b + (a - b::ennreal)" by transfer (auto simp: ereal_diff_positive max.absorb2 ereal_ineq_diff_add) lemma ennreal_mult_strict_left_mono: "(a::ennreal) < c \ 0 < b \ b < top \ b * a < b * c" by transfer (auto intro!: ereal_mult_strict_left_mono) lemma ennreal_between: "0 < e \ 0 < x \ x < top \ x - e < (x::ennreal)" by transfer (auto intro!: ereal_between) lemma minus_less_iff_ennreal: "b < top \ b \ a \ a - b < c \ a < c + (b::ennreal)" by transfer (auto simp: top_ereal_def ereal_minus_less le_less) lemma tendsto_zero_ennreal: assumes ev: "\r. 0 < r \ \\<^sub>F x in F. f x < ennreal r" shows "(f \ 0) F" proof (rule order_tendstoI) fix e::ennreal assume "e > 0" obtain e'::real where "e' > 0" "ennreal e' < e" using \0 < e\ dense[of 0 "if e = top then 1 else (enn2real e)"] by (cases e) (auto simp: ennreal_less_iff) from ev[OF \e' > 0\] show "\\<^sub>F x in F. f x < e" by eventually_elim (insert \ennreal e' < e\, auto) qed simp lifting_update ennreal.lifting lifting_forget ennreal.lifting subsection \\<^typ>\ennreal\ theorems\ lemma neq_top_trans: fixes x y :: ennreal shows "\ y \ top; x \ y \ \ x \ top" by (auto simp: top_unique) lemma diff_diff_ennreal: fixes a b :: ennreal shows "a \ b \ b \ \ \ b - (b - a) = a" by (cases a b rule: ennreal2_cases) (auto simp: ennreal_minus top_unique) lemma ennreal_less_one_iff[simp]: "ennreal x < 1 \ x < 1" by (cases "0 \ x") (auto simp: ennreal_neg ennreal_less_iff simp flip: ennreal_1) lemma SUP_const_minus_ennreal: fixes f :: "'a \ ennreal" shows "I \ {} \ (SUP x\I. c - f x) = c - (INF x\I. f x)" including ennreal.lifting by (transfer fixing: I) (simp add: SUP_sup_distrib[symmetric] SUP_ereal_minus_right flip: sup_ereal_def) lemma zero_minus_ennreal[simp]: "0 - (a::ennreal) = 0" including ennreal.lifting by transfer (simp split: split_max) lemma diff_diff_commute_ennreal: fixes a b c :: ennreal shows "a - b - c = a - c - b" by (cases a b c rule: ennreal3_cases) (simp_all add: ennreal_minus field_simps) lemma diff_gr0_ennreal: "b < (a::ennreal) \ 0 < a - b" including ennreal.lifting by transfer (auto simp: ereal_diff_gr0 ereal_diff_positive split: split_max) lemma divide_le_posI_ennreal: fixes x y z :: ennreal shows "x > 0 \ z \ x * y \ z / x \ y" by (cases x y z rule: ennreal3_cases) (auto simp: divide_ennreal ennreal_mult[symmetric] field_simps top_unique) lemma add_diff_eq_ennreal: fixes x y z :: ennreal shows "z \ y \ x + (y - z) = x + y - z" using ennreal_diff_add_assoc by auto lemma add_diff_inverse_ennreal: fixes x y :: ennreal shows "x \ y \ x + (y - x) = y" by (cases x) (simp_all add: top_unique add_diff_eq_ennreal) lemma add_diff_eq_iff_ennreal[simp]: fixes x y :: ennreal shows "x + (y - x) = y \ x \ y" proof assume *: "x + (y - x) = y" show "x \ y" by (subst *[symmetric]) simp qed (simp add: add_diff_inverse_ennreal) lemma add_diff_le_ennreal: "a + b - c \ a + (b - c::ennreal)" apply (cases a b c rule: ennreal3_cases) subgoal for a' b' c' by (cases "0 \ b' - c'") (simp_all add: ennreal_minus top_add ennreal_neg flip: ennreal_plus) apply (simp_all add: top_add flip: ennreal_plus) done lemma diff_eq_0_ennreal: "a < top \ a \ b \ a - b = (0::ennreal)" using ennreal_minus_pos_iff gr_zeroI not_less by blast lemma diff_diff_ennreal': fixes x y z :: ennreal shows "z \ y \ y - z \ x \ x - (y - z) = x + z - y" by (cases x; cases y; cases z) (auto simp add: top_add add_top minus_top_ennreal ennreal_minus top_unique simp flip: ennreal_plus) lemma diff_diff_ennreal'': fixes x y z :: ennreal shows "z \ y \ x - (y - z) = (if y - z \ x then x + z - y else 0)" by (cases x; cases y; cases z) (auto simp add: top_add add_top minus_top_ennreal ennreal_minus top_unique ennreal_neg simp flip: ennreal_plus) lemma power_less_top_ennreal: fixes x :: ennreal shows "x ^ n < top \ x < top \ n = 0" using power_eq_top_ennreal[of x n] by (auto simp: less_top) lemma ennreal_divide_times: "(a / b) * c = a * (c / b :: ennreal)" by (simp add: mult.commute ennreal_times_divide) lemma diff_less_top_ennreal: "a - b < top \ a < (top :: ennreal)" by (cases a; cases b) (auto simp: ennreal_minus) lemma divide_less_ennreal: "b \ 0 \ b < top \ a / b < c \ a < (c * b :: ennreal)" by (cases a; cases b; cases c) (auto simp: divide_ennreal ennreal_mult[symmetric] ennreal_less_iff field_simps ennreal_top_mult ennreal_top_divide) lemma one_less_numeral[simp]: "1 < (numeral n::ennreal) \ (num.One < n)" by (simp flip: ennreal_1 ennreal_numeral add: ennreal_less_iff) lemma divide_eq_1_ennreal: "a / b = (1::ennreal) \ (b \ top \ b \ 0 \ b = a)" by (cases a ; cases b; cases "b = 0") (auto simp: ennreal_top_divide divide_ennreal split: if_split_asm) lemma ennreal_mult_cancel_left: "(a * b = a * c) = (a = top \ b \ 0 \ c \ 0 \ a = 0 \ b = (c::ennreal))" by (cases a; cases b; cases c) (auto simp: ennreal_mult[symmetric] ennreal_mult_top ennreal_top_mult) lemma ennreal_minus_if: "ennreal a - ennreal b = ennreal (if 0 \ b then (if b \ a then a - b else 0) else a)" by (auto simp: ennreal_minus ennreal_neg) lemma ennreal_plus_if: "ennreal a + ennreal b = ennreal (if 0 \ a then (if 0 \ b then a + b else a) else b)" by (auto simp: ennreal_neg) lemma power_le_one_iff: "0 \ (a::real) \ a ^ n \ 1 \ (n = 0 \ a \ 1)" by (metis (mono_tags, opaque_lifting) le_less neq0_conv not_le one_le_power power_0 power_eq_imp_eq_base power_le_one zero_le_one) lemma ennreal_diff_le_mono_left: "a \ b \ a - c \ (b::ennreal)" using ennreal_mono_minus[of 0 c a, THEN order_trans, of b] by simp lemma ennreal_minus_le_iff: "a - b \ c \ (a \ b + (c::ennreal) \ (a = top \ b = top \ c = top))" by (cases a; cases b; cases c) (auto simp: top_unique top_add add_top ennreal_minus simp flip: ennreal_plus) lemma ennreal_le_minus_iff: "a \ b - c \ (a + c \ (b::ennreal) \ (a = 0 \ b \ c))" by (cases a; cases b; cases c) (auto simp: top_unique top_add add_top ennreal_minus ennreal_le_iff2 simp flip: ennreal_plus) lemma diff_add_eq_diff_diff_swap_ennreal: "x - (y + z :: ennreal) = x - y - z" by (cases x; cases y; cases z) (auto simp: ennreal_minus_if add_top top_add simp flip: ennreal_plus) lemma diff_add_assoc2_ennreal: "b \ a \ (a - b + c::ennreal) = a + c - b" by (cases a; cases b; cases c) (auto simp add: ennreal_minus_if ennreal_plus_if add_top top_add top_unique simp del: ennreal_plus) lemma diff_gt_0_iff_gt_ennreal: "0 < a - b \ (a = top \ b = top \ b < (a::ennreal))" by (cases a; cases b) (auto simp: ennreal_minus_if ennreal_less_iff) lemma diff_eq_0_iff_ennreal: "(a - b::ennreal) = 0 \ (a < top \ a \ b)" by (cases a) (auto simp: ennreal_minus_eq_0 diff_eq_0_ennreal) lemma add_diff_self_ennreal: "a + (b - a::ennreal) = (if a \ b then b else a)" by (auto simp: diff_eq_0_iff_ennreal less_top) lemma diff_add_self_ennreal: "(b - a + a::ennreal) = (if a \ b then b else a)" by (auto simp: diff_add_cancel_ennreal diff_eq_0_iff_ennreal less_top) lemma ennreal_minus_cancel_iff: fixes a b c :: ennreal shows "a - b = a - c \ (b = c \ (a \ b \ a \ c) \ a = top)" by (cases a; cases b; cases c) (auto simp: ennreal_minus_if) text \The next lemma is wrong for $a = top$, for $b = c = 1$ for instance.\ lemma ennreal_right_diff_distrib: fixes a b c :: ennreal assumes "a \ top" shows "a * (b - c) = a * b - a * c" apply (cases a; cases b; cases c) apply (use assms in \auto simp add: ennreal_mult_top ennreal_minus ennreal_mult' [symmetric]\) apply (simp add: algebra_simps) done lemma SUP_diff_ennreal: "c < top \ (SUP i\I. f i - c :: ennreal) = (SUP i\I. f i) - c" by (auto intro!: SUP_eqI ennreal_minus_mono SUP_least intro: SUP_upper simp: ennreal_minus_cancel_iff ennreal_minus_le_iff less_top[symmetric]) lemma ennreal_SUP_add_right: fixes c :: ennreal shows "I \ {} \ c + (SUP i\I. f i) = (SUP i\I. c + f i)" using ennreal_SUP_add_left[of I f c] by (simp add: add.commute) lemma SUP_add_directed_ennreal: fixes f g :: "_ \ ennreal" assumes directed: "\i j. i \ I \ j \ I \ \k\I. f i + g j \ f k + g k" shows "(SUP i\I. f i + g i) = (SUP i\I. f i) + (SUP i\I. g i)" proof (cases "I = {}") case False show ?thesis proof (rule antisym) show "(SUP i\I. f i + g i) \ (SUP i\I. f i) + (SUP i\I. g i)" by (rule SUP_least; intro add_mono SUP_upper) next have "(SUP i\I. f i) + (SUP i\I. g i) = (SUP i\I. f i + (SUP i\I. g i))" by (intro ennreal_SUP_add_left[symmetric] \I \ {}\) also have "\ = (SUP i\I. (SUP j\I. f i + g j))" using \I \ {}\ by (simp add: ennreal_SUP_add_right) also have "\ \ (SUP i\I. f i + g i)" using directed by (intro SUP_least) (blast intro: SUP_upper2) finally show "(SUP i\I. f i) + (SUP i\I. g i) \ (SUP i\I. f i + g i)" . qed qed (simp add: bot_ereal_def) lemma enn2real_eq_0_iff: "enn2real x = 0 \ x = 0 \ x = top" by (cases x) auto lemma continuous_on_diff_ennreal: "continuous_on A f \ continuous_on A g \ (\x. x \ A \ f x \ top) \ (\x. x \ A \ g x \ top) \ continuous_on A (\z. f z - g z::ennreal)" including ennreal.lifting proof (transfer fixing: A, simp add: top_ereal_def) fix f g :: "'a \ ereal" assume "\x. 0 \ f x" "\x. 0 \ g x" "continuous_on A f" "continuous_on A g" moreover assume "f x \ \" "g x \ \" if "x \ A" for x ultimately show "continuous_on A (\z. max 0 (f z - g z))" by (intro continuous_on_max continuous_on_const continuous_on_diff_ereal) auto qed lemma tendsto_diff_ennreal: "(f \ x) F \ (g \ y) F \ x \ top \ y \ top \ ((\z. f z - g z::ennreal) \ x - y) F" using continuous_on_tendsto_compose[where f="\x. fst x - snd x::ennreal" and s="{(x, y). x \ top \ y \ top}" and g="\x. (f x, g x)" and l="(x, y)" and F="F", OF continuous_on_diff_ennreal] by (auto simp: tendsto_Pair eventually_conj_iff less_top order_tendstoD continuous_on_fst continuous_on_snd continuous_on_id) declare lim_real_of_ereal [tendsto_intros] lemma tendsto_enn2real [tendsto_intros]: assumes "(u \ ennreal l) F" "l \ 0" shows "((\n. enn2real (u n)) \ l) F" unfolding enn2real_def by (metis assms enn2ereal_ennreal lim_real_of_ereal tendsto_enn2erealI) end diff --git a/src/HOL/Library/Multiset.thy b/src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy +++ b/src/HOL/Library/Multiset.thy @@ -1,4526 +1,4526 @@ (* Title: HOL/Library/Multiset.thy Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker Author: Andrei Popescu, TU Muenchen Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, TU Muenchen Author: Mathias Fleury, MPII Author: Martin Desharnais, MPI-INF Saarbruecken *) section \(Finite) Multisets\ theory Multiset imports Cancellation begin subsection \The type of multisets\ typedef 'a multiset = \{f :: 'a \ nat. finite {x. f x > 0}}\ morphisms count Abs_multiset proof show \(\x. 0::nat) \ {f. finite {x. f x > 0}}\ by simp qed setup_lifting type_definition_multiset lemma count_Abs_multiset: \count (Abs_multiset f) = f\ if \finite {x. f x > 0}\ by (rule Abs_multiset_inverse) (simp add: that) lemma multiset_eq_iff: "M = N \ (\a. count M a = count N a)" by (simp only: count_inject [symmetric] fun_eq_iff) lemma multiset_eqI: "(\x. count A x = count B x) \ A = B" using multiset_eq_iff by auto text \Preservation of the representing set \<^term>\multiset\.\ lemma diff_preserves_multiset: \finite {x. 0 < M x - N x}\ if \finite {x. 0 < M x}\ for M N :: \'a \ nat\ using that by (rule rev_finite_subset) auto lemma filter_preserves_multiset: \finite {x. 0 < (if P x then M x else 0)}\ if \finite {x. 0 < M x}\ for M N :: \'a \ nat\ using that by (rule rev_finite_subset) auto lemmas in_multiset = diff_preserves_multiset filter_preserves_multiset subsection \Representing multisets\ text \Multiset enumeration\ instantiation multiset :: (type) cancel_comm_monoid_add begin lift_definition zero_multiset :: \'a multiset\ is \\a. 0\ by simp abbreviation empty_mset :: \'a multiset\ (\{#}\) where \empty_mset \ 0\ lift_definition plus_multiset :: \'a multiset \ 'a multiset \ 'a multiset\ is \\M N a. M a + N a\ by simp lift_definition minus_multiset :: \'a multiset \ 'a multiset \ 'a multiset\ is \\M N a. M a - N a\ by (rule diff_preserves_multiset) instance by (standard; transfer) (simp_all add: fun_eq_iff) end context begin qualified definition is_empty :: "'a multiset \ bool" where [code_abbrev]: "is_empty A \ A = {#}" end lemma add_mset_in_multiset: \finite {x. 0 < (if x = a then Suc (M x) else M x)}\ if \finite {x. 0 < M x}\ using that by (simp add: flip: insert_Collect) lift_definition add_mset :: "'a \ 'a multiset \ 'a multiset" is "\a M b. if b = a then Suc (M b) else M b" by (rule add_mset_in_multiset) syntax "_multiset" :: "args \ 'a multiset" ("{#(_)#}") translations "{#x, xs#}" == "CONST add_mset x {#xs#}" "{#x#}" == "CONST add_mset x {#}" lemma count_empty [simp]: "count {#} a = 0" by (simp add: zero_multiset.rep_eq) lemma count_add_mset [simp]: "count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)" by (simp add: add_mset.rep_eq) lemma count_single: "count {#b#} a = (if b = a then 1 else 0)" by simp lemma add_mset_not_empty [simp]: \add_mset a A \ {#}\ and empty_not_add_mset [simp]: "{#} \ add_mset a A" by (auto simp: multiset_eq_iff) lemma add_mset_add_mset_same_iff [simp]: "add_mset a A = add_mset a B \ A = B" by (auto simp: multiset_eq_iff) lemma add_mset_commute: "add_mset x (add_mset y M) = add_mset y (add_mset x M)" by (auto simp: multiset_eq_iff) subsection \Basic operations\ subsubsection \Conversion to set and membership\ definition set_mset :: \'a multiset \ 'a set\ where \set_mset M = {x. count M x > 0}\ abbreviation member_mset :: \'a \ 'a multiset \ bool\ where \member_mset a M \ a \ set_mset M\ notation member_mset (\'(\#')\) and member_mset (\(_/ \# _)\ [50, 51] 50) notation (ASCII) member_mset (\'(:#')\) and member_mset (\(_/ :# _)\ [50, 51] 50) abbreviation not_member_mset :: \'a \ 'a multiset \ bool\ where \not_member_mset a M \ a \ set_mset M\ notation not_member_mset (\'(\#')\) and not_member_mset (\(_/ \# _)\ [50, 51] 50) notation (ASCII) not_member_mset (\'(~:#')\) and not_member_mset (\(_/ ~:# _)\ [50, 51] 50) context begin qualified abbreviation Ball :: "'a multiset \ ('a \ bool) \ bool" where "Ball M \ Set.Ball (set_mset M)" qualified abbreviation Bex :: "'a multiset \ ('a \ bool) \ bool" where "Bex M \ Set.Bex (set_mset M)" end syntax "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) syntax (ASCII) "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) translations "\x\#A. P" \ "CONST Multiset.Ball A (\x. P)" "\x\#A. P" \ "CONST Multiset.Bex A (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Multiset.Ball\ \<^syntax_const>\_MBall\, Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Multiset.Bex\ \<^syntax_const>\_MBex\] \ \ \to avoid eta-contraction of body\ lemma count_eq_zero_iff: "count M x = 0 \ x \# M" by (auto simp add: set_mset_def) lemma not_in_iff: "x \# M \ count M x = 0" by (auto simp add: count_eq_zero_iff) lemma count_greater_zero_iff [simp]: "count M x > 0 \ x \# M" by (auto simp add: set_mset_def) lemma count_inI: assumes "count M x = 0 \ False" shows "x \# M" proof (rule ccontr) assume "x \# M" with assms show False by (simp add: not_in_iff) qed lemma in_countE: assumes "x \# M" obtains n where "count M x = Suc n" proof - from assms have "count M x > 0" by simp then obtain n where "count M x = Suc n" using gr0_conv_Suc by blast with that show thesis . qed lemma count_greater_eq_Suc_zero_iff [simp]: "count M x \ Suc 0 \ x \# M" by (simp add: Suc_le_eq) lemma count_greater_eq_one_iff [simp]: "count M x \ 1 \ x \# M" by simp lemma set_mset_empty [simp]: "set_mset {#} = {}" by (simp add: set_mset_def) lemma set_mset_single: "set_mset {#b#} = {b}" by (simp add: set_mset_def) lemma set_mset_eq_empty_iff [simp]: "set_mset M = {} \ M = {#}" by (auto simp add: multiset_eq_iff count_eq_zero_iff) lemma finite_set_mset [iff]: "finite (set_mset M)" using count [of M] by simp lemma set_mset_add_mset_insert [simp]: \set_mset (add_mset a A) = insert a (set_mset A)\ by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits) lemma multiset_nonemptyE [elim]: assumes "A \ {#}" obtains x where "x \# A" proof - have "\x. x \# A" by (rule ccontr) (insert assms, auto) with that show ?thesis by blast qed subsubsection \Union\ lemma count_union [simp]: "count (M + N) a = count M a + count N a" by (simp add: plus_multiset.rep_eq) lemma set_mset_union [simp]: "set_mset (M + N) = set_mset M \ set_mset N" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp lemma union_mset_add_mset_left [simp]: "add_mset a A + B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma union_mset_add_mset_right [simp]: "A + add_mset a B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma add_mset_add_single: \add_mset a A = A + {#a#}\ by (subst union_mset_add_mset_right, subst add.comm_neutral) standard subsubsection \Difference\ instance multiset :: (type) comm_monoid_diff by standard (transfer; simp add: fun_eq_iff) lemma count_diff [simp]: "count (M - N) a = count M a - count N a" by (simp add: minus_multiset.rep_eq) lemma add_mset_diff_bothsides: \add_mset a M - add_mset a A = M - A\ by (auto simp: multiset_eq_iff) lemma in_diff_count: "a \# M - N \ count N a < count M a" by (simp add: set_mset_def) lemma count_in_diffI: assumes "\n. count N x = n + count M x \ False" shows "x \# M - N" proof (rule ccontr) assume "x \# M - N" then have "count N x = (count N x - count M x) + count M x" by (simp add: in_diff_count not_less) with assms show False by auto qed lemma in_diff_countE: assumes "x \# M - N" obtains n where "count M x = Suc n + count N x" proof - from assms have "count M x - count N x > 0" by (simp add: in_diff_count) then have "count M x > count N x" by simp then obtain n where "count M x = Suc n + count N x" using less_iff_Suc_add by auto with that show thesis . qed lemma in_diffD: assumes "a \# M - N" shows "a \# M" proof - have "0 \ count N a" by simp also from assms have "count N a < count M a" by (simp add: in_diff_count) finally show ?thesis by simp qed lemma set_mset_diff: "set_mset (M - N) = {a. count N a < count M a}" by (simp add: set_mset_def) lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" by rule (fact Groups.diff_zero, fact Groups.zero_diff) lemma diff_cancel: "A - A = {#}" by (fact Groups.diff_cancel) lemma diff_union_cancelR: "M + N - N = (M::'a multiset)" by (fact add_diff_cancel_right') lemma diff_union_cancelL: "N + M - N = (M::'a multiset)" by (fact add_diff_cancel_left') lemma diff_right_commute: fixes M N Q :: "'a multiset" shows "M - N - Q = M - Q - N" by (fact diff_right_commute) lemma diff_add: fixes M N Q :: "'a multiset" shows "M - (N + Q) = M - N - Q" by (rule sym) (fact diff_diff_add) lemma insert_DiffM [simp]: "x \# M \ add_mset x (M - {#x#}) = M" by (clarsimp simp: multiset_eq_iff) lemma insert_DiffM2: "x \# M \ (M - {#x#}) + {#x#} = M" by simp lemma diff_union_swap: "a \ b \ add_mset b (M - {#a#}) = add_mset b M - {#a#}" by (auto simp add: multiset_eq_iff) lemma diff_add_mset_swap [simp]: "b \# A \ add_mset b M - A = add_mset b (M - A)" by (auto simp add: multiset_eq_iff simp: not_in_iff) lemma diff_union_swap2 [simp]: "y \# M \ add_mset x M - {#y#} = add_mset x (M - {#y#})" by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM) lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)" by (rule diff_diff_add) lemma diff_union_single_conv: "a \# J \ I + J - {#a#} = I + (J - {#a#})" by (simp add: multiset_eq_iff Suc_le_eq) lemma mset_add [elim?]: assumes "a \# A" obtains B where "A = add_mset a B" proof - from assms have "A = add_mset a (A - {#a#})" by simp with that show thesis . qed lemma union_iff: "a \# A + B \ a \# A \ a \# B" by auto lemma count_minus_inter_lt_count_minus_inter_iff: "count (M2 - M1) y < count (M1 - M2) y \ y \# M1 - M2" by (meson count_greater_zero_iff gr_implies_not_zero in_diff_count leI order.strict_trans2 order_less_asym) lemma minus_inter_eq_minus_inter_iff: "(M1 - M2) = (M2 - M1) \ set_mset (M1 - M2) = set_mset (M2 - M1)" by (metis add.commute count_diff count_eq_zero_iff diff_add_zero in_diff_countE multiset_eq_iff) subsubsection \Min and Max\ abbreviation Min_mset :: "'a::linorder multiset \ 'a" where "Min_mset m \ Min (set_mset m)" abbreviation Max_mset :: "'a::linorder multiset \ 'a" where "Max_mset m \ Max (set_mset m)" subsubsection \Equality of multisets\ lemma single_eq_single [simp]: "{#a#} = {#b#} \ a = b" by (auto simp add: multiset_eq_iff) lemma union_eq_empty [iff]: "M + N = {#} \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma empty_eq_union [iff]: "{#} = M + N \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma multi_self_add_other_not_self [simp]: "M = add_mset x M \ False" by (auto simp add: multiset_eq_iff) lemma add_mset_remove_trivial [simp]: \add_mset x M - {#x#} = M\ by (auto simp: multiset_eq_iff) lemma diff_single_trivial: "\ x \# M \ M - {#x#} = M" by (auto simp add: multiset_eq_iff not_in_iff) lemma diff_single_eq_union: "x \# M \ M - {#x#} = N \ M = add_mset x N" by auto lemma union_single_eq_diff: "add_mset x M = N \ M = N - {#x#}" unfolding add_mset_add_single[of _ M] by (fact add_implies_diff) lemma union_single_eq_member: "add_mset x M = N \ x \# N" by auto lemma add_mset_remove_trivial_If: "add_mset a (N - {#a#}) = (if a \# N then N else add_mset a N)" by (simp add: diff_single_trivial) lemma add_mset_remove_trivial_eq: \N = add_mset a (N - {#a#}) \ a \# N\ by (auto simp: add_mset_remove_trivial_If) lemma union_is_single: "M + N = {#a#} \ M = {#a#} \ N = {#} \ M = {#} \ N = {#a#}" (is "?lhs = ?rhs") proof show ?lhs if ?rhs using that by auto show ?rhs if ?lhs by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that) qed lemma single_is_union: "{#a#} = M + N \ {#a#} = M \ N = {#} \ M = {#} \ {#a#} = N" by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single) lemma add_eq_conv_diff: "add_mset a M = add_mset b N \ M = N \ a = b \ M = add_mset b (N - {#a#}) \ N = add_mset a (M - {#b#})" (is "?lhs \ ?rhs") (* shorter: by (simp add: multiset_eq_iff) fastforce *) proof show ?lhs if ?rhs using that by (auto simp add: add_mset_commute[of a b]) show ?rhs if ?lhs proof (cases "a = b") case True with \?lhs\ show ?thesis by simp next case False from \?lhs\ have "a \# add_mset b N" by (rule union_single_eq_member) with False have "a \# N" by auto moreover from \?lhs\ have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff) moreover note False ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"]) qed qed lemma add_mset_eq_single [iff]: "add_mset b M = {#a#} \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma insert_noteq_member: assumes BC: "add_mset b B = add_mset c C" and bnotc: "b \ c" shows "c \# B" proof - have "c \# add_mset c C" by simp have nc: "\ c \# {#b#}" using bnotc by simp then have "c \# add_mset b B" using BC by simp then show "c \# B" using nc by simp qed lemma add_eq_conv_ex: "(add_mset a M = add_mset b N) = (M = N \ a = b \ (\K. M = add_mset b K \ N = add_mset a K))" by (auto simp add: add_eq_conv_diff) lemma multi_member_split: "x \# M \ \A. M = add_mset x A" by (rule exI [where x = "M - {#x#}"]) simp lemma multiset_add_sub_el_shuffle: assumes "c \# B" and "b \ c" shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}" proof - from \c \# B\ obtain A where B: "B = add_mset c A" by (blast dest: multi_member_split) have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp then have "add_mset b A = add_mset b (add_mset c A) - {#c#}" by (simp add: \b \ c\) then show ?thesis using B by simp qed lemma add_mset_eq_singleton_iff[iff]: "add_mset x M = {#y#} \ M = {#} \ x = y" by auto subsubsection \Pointwise ordering induced by count\ definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ (\a. count A a \ count B a)" definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ A \# B \ A \ B" abbreviation (input) supseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supseteq_mset A B \ B \# A" abbreviation (input) supset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supset_mset A B \ B \# A" notation (input) subseteq_mset (infix "\#" 50) and supseteq_mset (infix "\#" 50) notation (ASCII) subseteq_mset (infix "<=#" 50) and subset_mset (infix "<#" 50) and supseteq_mset (infix ">=#" 50) and supset_mset (infix ">#" 50) global_interpretation subset_mset: ordering \(\#)\ \(\#)\ by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym) interpretation subset_mset: ordered_ab_semigroup_add_imp_le \(+)\ \(-)\ \(\#)\ \(\#)\ by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(\#)" "(\#)" by standard \ \FIXME: avoid junk stemming from type class interpretation\ lemma mset_subset_eqI: "(\a. count A a \ count B a) \ A \# B" by (simp add: subseteq_mset_def) lemma mset_subset_eq_count: "A \# B \ count A a \ count B a" by (simp add: subseteq_mset_def) lemma mset_subset_eq_exists_conv: "(A::'a multiset) \# B \ (\C. B = A + C)" unfolding subseteq_mset_def apply (rule iffI) apply (rule exI [where x = "B - A"]) apply (auto intro: multiset_eq_iff [THEN iffD2]) done interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\#)" "(\#)" "(-)" by standard (simp, fact mset_subset_eq_exists_conv) \ \FIXME: avoid junk stemming from type class interpretation\ declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp] lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C \# B + C \ A \# B" by (fact subset_mset.add_le_cancel_right) lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) \# C + B \ A \# B" by (fact subset_mset.add_le_cancel_left) lemma mset_subset_eq_mono_add: "(A::'a multiset) \# B \ C \# D \ A + C \# B + D" by (fact subset_mset.add_mono) lemma mset_subset_eq_add_left: "(A::'a multiset) \# A + B" by simp lemma mset_subset_eq_add_right: "B \# (A::'a multiset) + B" by simp lemma single_subset_iff [simp]: "{#a#} \# M \ a \# M" by (auto simp add: subseteq_mset_def Suc_le_eq) lemma mset_subset_eq_single: "a \# B \ {#a#} \# B" by simp lemma mset_subset_eq_add_mset_cancel: \add_mset a A \# add_mset a B \ A \# B\ unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B] by (rule mset_subset_eq_mono_add_right_cancel) lemma multiset_diff_union_assoc: fixes A B C D :: "'a multiset" shows "C \# B \ A + B - C = A + (B - C)" by (fact subset_mset.diff_add_assoc) lemma mset_subset_eq_multiset_union_diff_commute: fixes A B C D :: "'a multiset" shows "B \# A \ A - B + C = A + C - B" by (fact subset_mset.add_diff_assoc2) lemma diff_subset_eq_self[simp]: "(M::'a multiset) - N \# M" by (simp add: subseteq_mset_def) lemma mset_subset_eqD: assumes "A \# B" and "x \# A" shows "x \# B" proof - from \x \# A\ have "count A x > 0" by simp also from \A \# B\ have "count A x \ count B x" by (simp add: subseteq_mset_def) finally show ?thesis by simp qed lemma mset_subsetD: "A \# B \ x \# A \ x \# B" by (auto intro: mset_subset_eqD [of A]) lemma set_mset_mono: "A \# B \ set_mset A \ set_mset B" by (metis mset_subset_eqD subsetI) lemma mset_subset_eq_insertD: "add_mset x A \# B \ x \# B \ A \# B" apply (rule conjI) apply (simp add: mset_subset_eqD) apply (clarsimp simp: subset_mset_def subseteq_mset_def) apply safe apply (erule_tac x = a in allE) apply (auto split: if_split_asm) done lemma mset_subset_insertD: "add_mset x A \# B \ x \# B \ A \# B" by (rule mset_subset_eq_insertD) simp lemma mset_subset_of_empty[simp]: "A \# {#} \ False" by (simp only: subset_mset.not_less_zero) lemma empty_subset_add_mset[simp]: "{#} \# add_mset x M" by (auto intro: subset_mset.gr_zeroI) lemma empty_le: "{#} \# A" by (fact subset_mset.zero_le) lemma insert_subset_eq_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" using le_diff_conv2 [of "Suc 0" "count B a" "count A a"] apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq) apply (rule ccontr) apply (auto simp add: not_in_iff) done lemma insert_union_subset_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" by (auto simp add: insert_subset_eq_iff subset_mset_def) lemma subset_eq_diff_conv: "A - C \# B \ A \# B + C" by (simp add: subseteq_mset_def le_diff_conv) lemma multi_psub_of_add_self [simp]: "A \# add_mset x A" by (auto simp: subset_mset_def subseteq_mset_def) lemma multi_psub_self: "A \# A = False" by simp lemma mset_subset_add_mset [simp]: "add_mset x N \# add_mset x M \ N \# M" unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M] by (fact subset_mset.add_less_cancel_right) lemma mset_subset_diff_self: "c \# B \ B - {#c#} \# B" by (auto simp: subset_mset_def elim: mset_add) lemma Diff_eq_empty_iff_mset: "A - B = {#} \ A \# B" by (auto simp: multiset_eq_iff subseteq_mset_def) lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \# {#b#} \ M = {#} \ a = b" proof assume A: "add_mset a M \# {#b#}" then have \a = b\ by (auto dest: mset_subset_eq_insertD) then show "M={#} \ a=b" using A by (simp add: mset_subset_eq_add_mset_cancel) qed simp subsubsection \Intersection and bounded union\ definition inter_mset :: \'a multiset \ 'a multiset \ 'a multiset\ (infixl \\#\ 70) where \A \# B = A - (A - B)\ lemma count_inter_mset [simp]: \count (A \# B) x = min (count A x) (count B x)\ by (simp add: inter_mset_def) (*global_interpretation subset_mset: semilattice_order \(\#)\ \(\#)\ \(\#)\ by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*) interpretation subset_mset: semilattice_inf \(\#)\ \(\#)\ \(\#)\ by standard (simp_all add: multiset_eq_iff subseteq_mset_def) \ \FIXME: avoid junk stemming from type class interpretation\ definition union_mset :: \'a multiset \ 'a multiset \ 'a multiset\ (infixl \\#\ 70) where \A \# B = A + (B - A)\ lemma count_union_mset [simp]: \count (A \# B) x = max (count A x) (count B x)\ by (simp add: union_mset_def) global_interpretation subset_mset: semilattice_neutr_order \(\#)\ \{#}\ \(\#)\ \(\#)\ apply standard apply (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def max_def) apply (auto simp add: le_antisym dest: sym) apply (metis nat_le_linear)+ done interpretation subset_mset: semilattice_sup \(\#)\ \(\#)\ \(\#)\ proof - have [simp]: "m \ n \ q \ n \ m + (q - m) \ n" for m n q :: nat by arith show "class.semilattice_sup (\#) (\#) (\#)" by standard (auto simp add: union_mset_def subseteq_mset_def) qed \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: bounded_lattice_bot "(\#)" "(\#)" "(\#)" "(\#)" "{#}" by standard auto \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Additional intersection facts\ lemma set_mset_inter [simp]: "set_mset (A \# B) = set_mset A \ set_mset B" by (simp only: set_mset_def) auto lemma diff_intersect_left_idem [simp]: "M - M \# N = M - N" by (simp add: multiset_eq_iff min_def) lemma diff_intersect_right_idem [simp]: "M - N \# M = M - N" by (simp add: multiset_eq_iff min_def) lemma multiset_inter_single[simp]: "a \ b \ {#a#} \# {#b#} = {#}" by (rule multiset_eqI) auto lemma multiset_union_diff_commute: assumes "B \# C = {#}" shows "A + B - C = A - C + B" proof (rule multiset_eqI) fix x from assms have "min (count B x) (count C x) = 0" by (auto simp add: multiset_eq_iff) then have "count B x = 0 \ count C x = 0" unfolding min_def by (auto split: if_splits) then show "count (A + B - C) x = count (A - C + B) x" by auto qed lemma disjunct_not_in: "A \# B = {#} \ (\a. a \# A \ a \# B)" (is "?P \ ?Q") proof assume ?P show ?Q proof fix a from \?P\ have "min (count A a) (count B a) = 0" by (simp add: multiset_eq_iff) then have "count A a = 0 \ count B a = 0" by (cases "count A a \ count B a") (simp_all add: min_def) then show "a \# A \ a \# B" by (simp add: not_in_iff) qed next assume ?Q show ?P proof (rule multiset_eqI) fix a from \?Q\ have "count A a = 0 \ count B a = 0" by (auto simp add: not_in_iff) then show "count (A \# B) a = count {#} a" by auto qed qed lemma inter_mset_empty_distrib_right: "A \# (B + C) = {#} \ A \# B = {#} \ A \# C = {#}" by (meson disjunct_not_in union_iff) lemma inter_mset_empty_distrib_left: "(A + B) \# C = {#} \ A \# C = {#} \ B \# C = {#}" by (meson disjunct_not_in union_iff) lemma add_mset_inter_add_mset [simp]: "add_mset a A \# add_mset a B = add_mset a (A \# B)" by (rule multiset_eqI) simp lemma add_mset_disjoint [simp]: "add_mset a A \# B = {#} \ a \# B \ A \# B = {#}" "{#} = add_mset a A \# B \ a \# B \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma disjoint_add_mset [simp]: "B \# add_mset a A = {#} \ a \# B \ B \# A = {#}" "{#} = A \# add_mset b B \ b \# A \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma inter_add_left1: "\ x \# N \ (add_mset x M) \# N = M \# N" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (auto simp add: multiset_eq_iff elim: mset_add) lemma inter_add_right1: "\ x \# N \ N \# (add_mset x M) = N \# M" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (auto simp add: multiset_eq_iff elim: mset_add) lemma disjunct_set_mset_diff: assumes "M \# N = {#}" shows "set_mset (M - N) = set_mset M" proof (rule set_eqI) fix a from assms have "a \# M \ a \# N" by (simp add: disjunct_not_in) then show "a \# M - N \ a \# M" by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff) qed lemma at_most_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M - {a}" using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff) lemma more_than_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M" proof (rule set_eqI) fix b have "Suc 0 < count M b \ count M b > 0" by arith then show "b \# M - {#a#} \ b \# M" using assms by (auto simp add: in_diff_count) qed lemma inter_iff: "a \# A \# B \ a \# A \ a \# B" by simp lemma inter_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff min_add_distrib_left) lemma inter_union_distrib_right: "C + A \# B = (C + A) \# (C + B)" using inter_union_distrib_left [of A B C] by (simp add: ac_simps) lemma inter_subset_eq_union: "A \# B \# A + B" by (auto simp add: subseteq_mset_def) subsubsection \Additional bounded union facts\ lemma set_mset_sup [simp]: \set_mset (A \# B) = set_mset A \ set_mset B\ by (simp only: set_mset_def) (auto simp add: less_max_iff_disj) lemma sup_union_left1 [simp]: "\ x \# N \ (add_mset x M) \# N = add_mset x (M \# N)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (simp add: multiset_eq_iff) lemma sup_union_right1 [simp]: "\ x \# N \ N \# (add_mset x M) = add_mset x (N \# M)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (simp add: multiset_eq_iff) lemma sup_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff max_add_distrib_left) lemma union_sup_distrib_right: "C + A \# B = (C + A) \# (C + B)" using sup_union_distrib_left [of A B C] by (simp add: ac_simps) lemma union_diff_inter_eq_sup: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma union_diff_sup_eq_inter: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma add_mset_union: \add_mset a A \# add_mset a B = add_mset a (A \# B)\ by (auto simp: multiset_eq_iff max_def) subsection \Replicate and repeat operations\ definition replicate_mset :: "nat \ 'a \ 'a multiset" where "replicate_mset n x = (add_mset x ^^ n) {#}" lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}" unfolding replicate_mset_def by simp lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)" unfolding replicate_mset_def by (induct n) (auto intro: add.commute) lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" unfolding replicate_mset_def by (induct n) auto lift_definition repeat_mset :: \nat \ 'a multiset \ 'a multiset\ is \\n M a. n * M a\ by simp lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a" by transfer rule lemma repeat_mset_0 [simp]: \repeat_mset 0 M = {#}\ by transfer simp lemma repeat_mset_Suc [simp]: \repeat_mset (Suc n) M = M + repeat_mset n M\ by transfer simp lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A" by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_diff_repeat_mset_distrib': \repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u\ by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_add_mult_distrib_mset: "repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k" by (auto simp: multiset_eq_iff add_mult_distrib) lemma repeat_mset_distrib: "repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A" by (auto simp: multiset_eq_iff Nat.add_mult_distrib) lemma repeat_mset_distrib2[simp]: "repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B" by (auto simp: multiset_eq_iff add_mult_distrib2) lemma repeat_mset_replicate_mset[simp]: "repeat_mset n {#a#} = replicate_mset n a" by (auto simp: multiset_eq_iff) lemma repeat_mset_distrib_add_mset[simp]: "repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A" by (auto simp: multiset_eq_iff) lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}" by transfer simp subsubsection \Simprocs\ lemma repeat_mset_iterate_add: \repeat_mset n M = iterate_add n M\ unfolding iterate_add_def by (induction n) auto lemma mset_subseteq_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" by (auto simp add: subseteq_mset_def nat_le_add_iff1) lemma mset_subseteq_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" by (auto simp add: subseteq_mset_def nat_le_add_iff2) lemma mset_subset_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add]) lemma mset_subset_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add]) ML_file \multiset_simprocs.ML\ lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: \NO_MATCH {#} M \ add_mset a M = {#a#} + M\ by simp declare repeat_mset_iterate_add[cancelation_simproc_pre] declare iterate_add_distrib[cancelation_simproc_pre] declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post] declare add_mset_not_empty[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] add_mset_not_empty[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] le_zero_eq[cancelation_simproc_eq_elim] simproc_setup mseteq_cancel ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" | "add_mset a m = n" | "m = add_mset a n" | "replicate_mset p a = n" | "m = replicate_mset p a" | "repeat_mset p m = n" | "m = repeat_mset p m") = - \fn phi => Cancel_Simprocs.eq_cancel\ + \K Cancel_Simprocs.eq_cancel\ simproc_setup msetsubset_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = - \fn phi => Multiset_Simprocs.subset_cancel_msets\ + \K Multiset_Simprocs.subset_cancel_msets\ simproc_setup msetsubset_eq_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = - \fn phi => Multiset_Simprocs.subseteq_cancel_msets\ + \K Multiset_Simprocs.subseteq_cancel_msets\ simproc_setup msetdiff_cancel ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" | "add_mset a m - n" | "m - add_mset a n" | "replicate_mset p r - n" | "m - replicate_mset p r" | "repeat_mset p m - n" | "m - repeat_mset p m") = - \fn phi => Cancel_Simprocs.diff_cancel\ + \K Cancel_Simprocs.diff_cancel\ subsubsection \Conditionally complete lattice\ instantiation multiset :: (type) Inf begin lift_definition Inf_multiset :: "'a multiset set \ 'a multiset" is "\A i. if A = {} then 0 else Inf ((\f. f i) ` A)" proof - fix A :: "('a \ nat) set" assume *: "\f. f \ A \ finite {x. 0 < f x}" show \finite {i. 0 < (if A = {} then 0 else INF f\A. f i)}\ proof (cases "A = {}") case False then obtain f where "f \ A" by blast hence "{i. Inf ((\f. f i) ` A) > 0} \ {i. f i > 0}" by (auto intro: less_le_trans[OF _ cInf_lower]) moreover from \f \ A\ * have "finite \" by simp ultimately have "finite {i. Inf ((\f. f i) ` A) > 0}" by (rule finite_subset) with False show ?thesis by simp qed simp_all qed instance .. end lemma Inf_multiset_empty: "Inf {} = {#}" by transfer simp_all lemma count_Inf_multiset_nonempty: "A \ {} \ count (Inf A) x = Inf ((\X. count X x) ` A)" by transfer simp_all instantiation multiset :: (type) Sup begin definition Sup_multiset :: "'a multiset set \ 'a multiset" where "Sup_multiset A = (if A \ {} \ subset_mset.bdd_above A then Abs_multiset (\i. Sup ((\X. count X i) ` A)) else {#})" lemma Sup_multiset_empty: "Sup {} = {#}" by (simp add: Sup_multiset_def) lemma Sup_multiset_unbounded: "\ subset_mset.bdd_above A \ Sup A = {#}" by (simp add: Sup_multiset_def) instance .. end lemma bdd_above_multiset_imp_bdd_above_count: assumes "subset_mset.bdd_above (A :: 'a multiset set)" shows "bdd_above ((\X. count X x) ` A)" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (meson subset_mset.bdd_above.E) hence "count X x \ count Y x" if "X \ A" for X using that by (auto intro: mset_subset_eq_count) thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto qed lemma bdd_above_multiset_imp_finite_support: assumes "A \ {}" "subset_mset.bdd_above (A :: 'a multiset set)" shows "finite (\X\A. {x. count X x > 0})" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (meson subset_mset.bdd_above.E) hence "count X x \ count Y x" if "X \ A" for X x using that by (auto intro: mset_subset_eq_count) hence "(\X\A. {x. count X x > 0}) \ {x. count Y x > 0}" by safe (erule less_le_trans) moreover have "finite \" by simp ultimately show ?thesis by (rule finite_subset) qed lemma Sup_multiset_in_multiset: \finite {i. 0 < (SUP M\A. count M i)}\ if \A \ {}\ \subset_mset.bdd_above A\ proof - have "{i. Sup ((\X. count X i) ` A) > 0} \ (\X\A. {i. 0 < count X i})" proof safe fix i assume pos: "(SUP X\A. count X i) > 0" show "i \ (\X\A. {i. 0 < count X i})" proof (rule ccontr) assume "i \ (\X\A. {i. 0 < count X i})" hence "\X\A. count X i \ 0" by (auto simp: count_eq_zero_iff) with that have "(SUP X\A. count X i) \ 0" by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto with pos show False by simp qed qed moreover from that have "finite \" by (rule bdd_above_multiset_imp_finite_support) ultimately show "finite {i. Sup ((\X. count X i) ` A) > 0}" by (rule finite_subset) qed lemma count_Sup_multiset_nonempty: \count (Sup A) x = (SUP X\A. count X x)\ if \A \ {}\ \subset_mset.bdd_above A\ using that by (simp add: Sup_multiset_def Sup_multiset_in_multiset count_Abs_multiset) interpretation subset_mset: conditionally_complete_lattice Inf Sup "(\#)" "(\#)" "(\#)" "(\#)" proof fix X :: "'a multiset" and A assume "X \ A" show "Inf A \# X" proof (rule mset_subset_eqI) fix x from \X \ A\ have "A \ {}" by auto hence "count (Inf A) x = (INF X\A. count X x)" by (simp add: count_Inf_multiset_nonempty) also from \X \ A\ have "\ \ count X x" by (intro cInf_lower) simp_all finally show "count (Inf A) x \ count X x" . qed next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and le: "\Y. Y \ A \ X \# Y" show "X \# Inf A" proof (rule mset_subset_eqI) fix x from nonempty have "count X x \ (INF X\A. count X x)" by (intro cInf_greatest) (auto intro: mset_subset_eq_count le) also from nonempty have "\ = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty) finally show "count X x \ count (Inf A) x" . qed next fix X :: "'a multiset" and A assume X: "X \ A" and bdd: "subset_mset.bdd_above A" show "X \# Sup A" proof (rule mset_subset_eqI) fix x from X have "A \ {}" by auto have "count X x \ (SUP X\A. count X x)" by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd) also from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "(SUP X\A. count X x) = count (Sup A) x" by simp finally show "count X x \ count (Sup A) x" . qed next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and ge: "\Y. Y \ A \ Y \# X" from ge have bdd: "subset_mset.bdd_above A" by blast show "Sup A \# X" proof (rule mset_subset_eqI) fix x from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "count (Sup A) x = (SUP X\A. count X x)" . also from nonempty have "\ \ count X x" by (intro cSup_least) (auto intro: mset_subset_eq_count ge) finally show "count (Sup A) x \ count X x" . qed qed \ \FIXME: avoid junk stemming from type class interpretation\ lemma set_mset_Inf: assumes "A \ {}" shows "set_mset (Inf A) = (\X\A. set_mset X)" proof safe fix x X assume "x \# Inf A" "X \ A" hence nonempty: "A \ {}" by (auto simp: Inf_multiset_empty) from \x \# Inf A\ have "{#x#} \# Inf A" by auto also from \X \ A\ have "\ \# X" by (rule subset_mset.cInf_lower) simp_all finally show "x \# X" by simp next fix x assume x: "x \ (\X\A. set_mset X)" hence "{#x#} \# X" if "X \ A" for X using that by auto from assms and this have "{#x#} \# Inf A" by (rule subset_mset.cInf_greatest) thus "x \# Inf A" by simp qed lemma in_Inf_multiset_iff: assumes "A \ {}" shows "x \# Inf A \ (\X\A. x \# X)" proof - from assms have "set_mset (Inf A) = (\X\A. set_mset X)" by (rule set_mset_Inf) also have "x \ \ \ (\X\A. x \# X)" by simp finally show ?thesis . qed lemma in_Inf_multisetD: "x \# Inf A \ X \ A \ x \# X" by (subst (asm) in_Inf_multiset_iff) auto lemma set_mset_Sup: assumes "subset_mset.bdd_above A" shows "set_mset (Sup A) = (\X\A. set_mset X)" proof safe fix x assume "x \# Sup A" hence nonempty: "A \ {}" by (auto simp: Sup_multiset_empty) show "x \ (\X\A. set_mset X)" proof (rule ccontr) assume x: "x \ (\X\A. set_mset X)" have "count X x \ count (Sup A) x" if "X \ A" for X x using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms) with x have "X \# Sup A - {#x#}" if "X \ A" for X using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff) hence "Sup A \# Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty) with \x \# Sup A\ show False by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff dest!: spec[of _ x]) qed next fix x X assume "x \ set_mset X" "X \ A" hence "{#x#} \# X" by auto also have "X \# Sup A" by (intro subset_mset.cSup_upper \X \ A\ assms) finally show "x \ set_mset (Sup A)" by simp qed lemma in_Sup_multiset_iff: assumes "subset_mset.bdd_above A" shows "x \# Sup A \ (\X\A. x \# X)" proof - from assms have "set_mset (Sup A) = (\X\A. set_mset X)" by (rule set_mset_Sup) also have "x \ \ \ (\X\A. x \# X)" by simp finally show ?thesis . qed lemma in_Sup_multisetD: assumes "x \# Sup A" shows "\X\A. x \# X" proof - have "subset_mset.bdd_above A" by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded) with assms show ?thesis by (simp add: in_Sup_multiset_iff) qed interpretation subset_mset: distrib_lattice "(\#)" "(\#)" "(\#)" "(\#)" proof fix A B C :: "'a multiset" show "A \# (B \# C) = A \# B \# (A \# C)" by (intro multiset_eqI) simp_all qed \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Filter (with comprehension syntax)\ text \Multiset comprehension\ lift_definition filter_mset :: "('a \ bool) \ 'a multiset \ 'a multiset" is "\P M. \x. if P x then M x else 0" by (rule filter_preserves_multiset) syntax (ASCII) "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ :# _./ _#})") syntax "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ \# _./ _#})") translations "{#x \# M. P#}" == "CONST filter_mset (\x. P) M" lemma count_filter_mset [simp]: "count (filter_mset P M) a = (if P a then count M a else 0)" by (simp add: filter_mset.rep_eq) lemma set_mset_filter [simp]: "set_mset (filter_mset P M) = {a \ set_mset M. P a}" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}" by (rule multiset_eqI) simp lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})" by (rule multiset_eqI) simp lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N" by (rule multiset_eqI) simp lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N" by (rule multiset_eqI) simp lemma filter_inter_mset [simp]: "filter_mset P (M \# N) = filter_mset P M \# filter_mset P N" by (rule multiset_eqI) simp lemma filter_sup_mset[simp]: "filter_mset P (A \# B) = filter_mset P A \# filter_mset P B" by (rule multiset_eqI) simp lemma filter_mset_add_mset [simp]: "filter_mset P (add_mset x A) = (if P x then add_mset x (filter_mset P A) else filter_mset P A)" by (auto simp: multiset_eq_iff) lemma multiset_filter_subset[simp]: "filter_mset f M \# M" by (simp add: mset_subset_eqI) lemma multiset_filter_mono: assumes "A \# B" shows "filter_mset f A \# filter_mset f B" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain C where B: "B = A + C" by auto show ?thesis unfolding B by auto qed lemma filter_mset_eq_conv: "filter_mset P M = N \ N \# M \ (\b\#N. P b) \ (\a\#M - N. \ P a)" (is "?P \ ?Q") proof assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count) next assume ?Q then obtain Q where M: "M = N + Q" by (auto simp add: mset_subset_eq_exists_conv) then have MN: "M - N = Q" by simp show ?P proof (rule multiset_eqI) fix a from \?Q\ MN have *: "\ P a \ a \# N" "P a \ a \# Q" by auto show "count (filter_mset P M) a = count N a" proof (cases "a \# M") case True with * show ?thesis by (simp add: not_in_iff M) next case False then have "count M a = 0" by (simp add: not_in_iff) with M show ?thesis by simp qed qed qed lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \# M. Q x \ P x#}" by (auto simp: multiset_eq_iff) lemma filter_mset_True[simp]: "{#y \# M. True#} = M" and filter_mset_False[simp]: "{#y \# M. False#} = {#}" by (auto simp: multiset_eq_iff) lemma filter_mset_cong0: assumes "\x. x \# M \ f x \ g x" shows "filter_mset f M = filter_mset g M" proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI) fix x show "count (filter_mset f M) x \ count (filter_mset g M) x" using assms by (cases "x \# M") (simp_all add: not_in_iff) next fix x show "count (filter_mset g M) x \ count (filter_mset f M) x" using assms by (cases "x \# M") (simp_all add: not_in_iff) qed lemma filter_mset_cong: assumes "M = M'" and "\x. x \# M' \ f x \ g x" shows "filter_mset f M = filter_mset g M'" unfolding \M = M'\ using assms by (auto intro: filter_mset_cong0) subsubsection \Size\ definition wcount where "wcount f M = (\x. count M x * Suc (f x))" lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" by (auto simp: wcount_def add_mult_distrib) lemma wcount_add_mset: "wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a" unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def) definition size_multiset :: "('a \ nat) \ 'a multiset \ nat" where "size_multiset f M = sum (wcount f M) (set_mset M)" lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] instantiation multiset :: (type) size begin definition size_multiset where size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\_. 0)" instance .. end lemmas size_multiset_overloaded_eq = size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] lemma size_multiset_empty [simp]: "size_multiset f {#} = 0" by (simp add: size_multiset_def) lemma size_empty [simp]: "size {#} = 0" by (simp add: size_multiset_overloaded_def) lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)" by (simp add: size_multiset_eq) lemma size_single: "size {#b#} = 1" by (simp add: size_multiset_overloaded_def size_multiset_single) lemma sum_wcount_Int: "finite A \ sum (wcount f N) (A \ set_mset N) = sum (wcount f N) A" by (induct rule: finite_induct) (simp_all add: Int_insert_left wcount_def count_eq_zero_iff) lemma size_multiset_union [simp]: "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) apply (subst Int_commute) apply (simp add: sum_wcount_Int) done lemma size_multiset_add_mset [simp]: "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M" unfolding add_mset_add_single[of _ M] size_multiset_union by (auto simp: size_multiset_single) lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)" by (simp add: size_multiset_overloaded_def wcount_add_mset) lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" by (auto simp add: size_multiset_overloaded_def) lemma size_multiset_eq_0_iff_empty [iff]: "size_multiset f M = 0 \ M = {#}" by (auto simp add: size_multiset_eq count_eq_zero_iff) lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" by (auto simp add: size_multiset_overloaded_def) lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a \# M" apply (unfold size_multiset_overloaded_eq) apply (drule sum_SucD) apply auto done lemma size_eq_Suc_imp_eq_union: assumes "size M = Suc n" shows "\a N. M = add_mset a N" proof - from assms obtain a where "a \# M" by (erule size_eq_Suc_imp_elem [THEN exE]) then have "M = add_mset a (M - {#a#})" by simp then show ?thesis by blast qed lemma size_mset_mono: fixes A B :: "'a multiset" assumes "A \# B" shows "size A \ size B" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain C where B: "B = A + C" by auto show ?thesis unfolding B by (induct C) auto qed lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \ size M" by (rule size_mset_mono[OF multiset_filter_subset]) lemma size_Diff_submset: "M \# M' \ size (M' - M) = size M' - size(M::'a multiset)" by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv) subsection \Induction and case splits\ theorem multiset_induct [case_names empty add, induct type: multiset]: assumes empty: "P {#}" assumes add: "\x M. P M \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case 0 thus "P M" by (simp add: empty) next case (Suc k) obtain N x where "M = add_mset x N" using \Suc k = size M\ [symmetric] using size_eq_Suc_imp_eq_union by fast with Suc add show "P M" by simp qed lemma multiset_induct_min[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Min_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Min_le finite_set_mset in_diffD) qed (simp add: empty) lemma multiset_induct_max[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Max_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Max_ge finite_set_mset in_diffD) qed (simp add: empty) lemma multi_nonempty_split: "M \ {#} \ \A a. M = add_mset a A" by (induct M) auto lemma multiset_cases [cases type]: obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N" by (induct M) simp_all lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" by (cases "B = {#}") (auto dest: multi_member_split) lemma union_filter_mset_complement[simp]: "\x. P x = (\ Q x) \ filter_mset P M + filter_mset Q M = M" by (subst multiset_eq_iff) auto lemma multiset_partition: "M = {#x \# M. P x#} + {#x \# M. \ P x#}" by simp lemma mset_subset_size: "A \# B \ size A < size B" proof (induct A arbitrary: B) case empty then show ?case using nonempty_has_size by auto next case (add x A) have "add_mset x A \# B" by (meson add.prems subset_mset_def) then show ?case by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff size_Diff_submset size_eq_0_iff_empty size_mset_mono subset_mset.le_iff_add subset_mset_def) qed lemma size_1_singleton_mset: "size M = 1 \ \a. M = {#a#}" by (cases M) auto subsubsection \Strong induction and subset induction for multisets\ text \Well-foundedness of strict subset relation\ lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \# N}" apply (rule wf_measure [THEN wf_subset, where f1=size]) apply (clarsimp simp: measure_def inv_image_def mset_subset_size) done lemma wfP_subset_mset[simp]: "wfP (\#)" by (rule wf_subset_mset_rel[to_pred]) lemma full_multiset_induct [case_names less]: assumes ih: "\B. \(A::'a multiset). A \# B \ P A \ P B" shows "P B" apply (rule wf_subset_mset_rel [THEN wf_induct]) apply (rule ih, auto) done lemma multi_subset_induct [consumes 2, case_names empty add]: assumes "F \# A" and empty: "P {#}" and insert: "\a F. a \# A \ P F \ P (add_mset a F)" shows "P F" proof - from \F \# A\ show ?thesis proof (induct F) show "P {#}" by fact next fix x F assume P: "F \# A \ P F" and i: "add_mset x F \# A" show "P (add_mset x F)" proof (rule insert) from i show "x \# A" by (auto dest: mset_subset_eq_insertD) from i have "F \# A" by (auto dest: mset_subset_eq_insertD) with P show "P F" . qed qed qed subsection \Least and greatest elements\ context begin qualified lemma assumes "M \ {#}" and "transp_on (set_mset M) R" and "totalp_on (set_mset M) R" shows bex_least_element: "(\l \# M. \x \# M. x \ l \ R l x)" and bex_greatest_element: "(\g \# M. \x \# M. x \ g \ R x g)" using assms by (auto intro: Finite_Set.bex_least_element Finite_Set.bex_greatest_element) end subsection \The fold combinator\ definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" where "fold_mset f s M = Finite_Set.fold (\x. f x ^^ count M x) s (set_mset M)" lemma fold_mset_empty [simp]: "fold_mset f s {#} = s" by (simp add: fold_mset_def) context comp_fun_commute begin lemma fold_mset_add_mset [simp]: "fold_mset f s (add_mset x M) = f x (fold_mset f s M)" proof - interpret mset: comp_fun_commute "\y. f y ^^ count M y" by (fact comp_fun_commute_funpow) interpret mset_union: comp_fun_commute "\y. f y ^^ count (add_mset x M) y" by (fact comp_fun_commute_funpow) show ?thesis proof (cases "x \ set_mset M") case False then have *: "count (add_mset x M) x = 1" by (simp add: not_in_iff) from False have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s (set_mset M) = Finite_Set.fold (\y. f y ^^ count M y) s (set_mset M)" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) with False * show ?thesis by (simp add: fold_mset_def del: count_add_mset) next case True define N where "N = set_mset M - {x}" from N_def True have *: "set_mset M = insert x N" "x \ N" "finite N" by auto then have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s N = Finite_Set.fold (\y. f y ^^ count M y) s N" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp qed qed corollary fold_mset_single: "fold_mset f s {#x#} = f x s" by simp lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" by (induct M) (simp_all add: fun_left_comm) lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" by (induct M) (simp_all add: fold_mset_fun_left_comm) lemma fold_mset_fusion: assumes "comp_fun_commute g" and *: "\x y. h (g x y) = f x (h y)" shows "h (fold_mset g w A) = fold_mset f (h w) A" proof - interpret comp_fun_commute g by (fact assms) from * show ?thesis by (induct A) auto qed end lemma union_fold_mset_add_mset: "A + B = fold_mset add_mset A B" proof - interpret comp_fun_commute add_mset by standard auto show ?thesis by (induction B) auto qed text \ A note on code generation: When defining some function containing a subterm \<^term>\fold_mset F\, code generation is not automatic. When interpreting locale \left_commutative\ with \F\, the would be code thms for \<^const>\fold_mset\ become thms like \<^term>\fold_mset F z {#} = z\ where \F\ is not a pattern but contains defined symbols, i.e.\ is not a code thm. Hence a separate constant with its own code thms needs to be introduced for \F\. See the image operator below. \ subsection \Image\ definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where "image_mset f = fold_mset (add_mset \ f) {#}" lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \ f)" by unfold_locales (simp add: fun_eq_iff) lemma image_mset_empty [simp]: "image_mset f {#} = {#}" by (simp add: image_mset_def) lemma image_mset_single: "image_mset f {#x#} = {#f x#}" by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def) lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" proof - interpret comp_fun_commute "add_mset \ f" by (fact comp_fun_commute_mset_image) show ?thesis by (induct N) (simp_all add: image_mset_def) qed corollary image_mset_add_mset [simp]: "image_mset f (add_mset a M) = add_mset (f a) (image_mset f M)" unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single) lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" by (induct M) simp_all lemma size_image_mset [simp]: "size (image_mset f M) = size M" by (induct M) simp_all lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" by (cases M) auto lemma image_mset_If: "image_mset (\x. if P x then f x else g x) A = image_mset f (filter_mset P A) + image_mset g (filter_mset (\x. \P x) A)" by (induction A) auto lemma image_mset_Diff: assumes "B \# A" shows "image_mset f (A - B) = image_mset f A - image_mset f B" proof - have "image_mset f (A - B + B) = image_mset f (A - B) + image_mset f B" by simp also from assms have "A - B + B = A" by (simp add: subset_mset.diff_add) finally show ?thesis by simp qed lemma count_image_mset: \count (image_mset f A) x = (\y\f -` {x} \ set_mset A. count A y)\ proof (induction A) case empty then show ?case by simp next case (add x A) moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y by simp ultimately show ?case by (auto simp: sum.distrib intro!: sum.mono_neutral_left) qed lemma count_image_mset': \count (image_mset f X) y = (\x | x \# X \ y = f x. count X x)\ by (auto simp add: count_image_mset simp flip: singleton_conv2 simp add: Collect_conj_eq ac_simps) lemma image_mset_subseteq_mono: "A \# B \ image_mset f A \# image_mset f B" by (metis image_mset_union subset_mset.le_iff_add) lemma image_mset_subset_mono: "M \# N \ image_mset f M \# image_mset f N" by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff image_mset_subseteq_mono subset_mset.less_le_not_le) syntax (ASCII) "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") syntax "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ \# _#})") translations "{#e. x \# M#}" \ "CONST image_mset (\x. e) M" syntax (ASCII) "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") syntax "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ \# _./ _#})") translations "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" text \ This allows to write not just filters like \<^term>\{#x\#M. x but also images like \<^term>\{#x+x. x\#M #}\ and @{term [source] "{#x+x|x\#M. x\{#x+x|x\#M. x. \ lemma in_image_mset: "y \# {#f x. x \# M#} \ y \ f ` set_mset M" by simp functor image_mset: image_mset proof - fix f g show "image_mset f \ image_mset g = image_mset (f \ g)" proof fix A show "(image_mset f \ image_mset g) A = image_mset (f \ g) A" by (induct A) simp_all qed show "image_mset id = id" proof fix A show "image_mset id A = id A" by (induct A) simp_all qed qed declare image_mset.id [simp] image_mset.identity [simp] lemma image_mset_id[simp]: "image_mset id x = x" unfolding id_def by auto lemma image_mset_cong: "(\x. x \# M \ f x = g x) \ {#f x. x \# M#} = {#g x. x \# M#}" by (induct M) auto lemma image_mset_cong_pair: "(\x y. (x, y) \# M \ f x y = g x y) \ {#f x y. (x, y) \# M#} = {#g x y. (x, y) \# M#}" by (metis image_mset_cong split_cong) lemma image_mset_const_eq: "{#c. a \# M#} = replicate_mset (size M) c" by (induct M) simp_all lemma image_mset_filter_mset_swap: "image_mset f (filter_mset (\x. P (f x)) M) = filter_mset P (image_mset f M)" by (induction M rule: multiset_induct) simp_all lemma image_mset_eq_plusD: "image_mset f A = B + C \ \B' C'. A = B' + C' \ B = image_mset f B' \ C = image_mset f C'" proof (induction A arbitrary: B C) case empty thus ?case by simp next case (add x A) show ?case proof (cases "f x \# B") case True with add.prems have "image_mset f A = (B - {#f x#}) + C" by (metis add_mset_remove_trivial image_mset_add_mset mset_subset_eq_single subset_mset.add_diff_assoc2) thus ?thesis using add.IH add.prems by force next case False with add.prems have "image_mset f A = B + (C - {#f x#})" by (metis diff_single_eq_union diff_union_single_conv image_mset_add_mset union_iff union_single_eq_member) then show ?thesis using add.IH add.prems by force qed qed lemma image_mset_eq_image_mset_plusD: assumes "image_mset f A = image_mset f B + C" and inj_f: "inj_on f (set_mset A \ set_mset B)" shows "\C'. A = B + C' \ C = image_mset f C'" using assms proof (induction A arbitrary: B C) case empty thus ?case by simp next case (add x A) show ?case proof (cases "x \# B") case True with add.prems have "image_mset f A = image_mset f (B - {#x#}) + C" by (smt (verit, del_insts) add.left_commute add_cancel_right_left diff_union_cancelL diff_union_single_conv image_mset_union union_mset_add_mset_left union_mset_add_mset_right) with add.IH have "\M3'. A = B - {#x#} + M3' \ image_mset f M3' = C" by (smt (verit, del_insts) True Un_insert_left Un_insert_right add.prems(2) inj_on_insert insert_DiffM set_mset_add_mset_insert) with True show ?thesis by auto next case False with add.prems(2) have "f x \# image_mset f B" by auto with add.prems(1) have "image_mset f A = image_mset f B + (C - {#f x#})" by (metis (no_types, lifting) diff_union_single_conv image_eqI image_mset_Diff image_mset_single mset_subset_eq_single set_image_mset union_iff union_single_eq_diff union_single_eq_member) with add.prems(2) add.IH have "\M3'. A = B + M3' \ C - {#f x#} = image_mset f M3'" by auto then show ?thesis by (metis add.prems(1) add_diff_cancel_left' image_mset_Diff mset_subset_eq_add_left union_mset_add_mset_right) qed qed lemma image_mset_eq_plus_image_msetD: "image_mset f A = B + image_mset f C \ inj_on f (set_mset A \ set_mset C) \ \B'. A = B' + C \ B = image_mset f B'" unfolding add.commute[of B] add.commute[of _ C] by (rule image_mset_eq_image_mset_plusD; assumption) subsection \Further conversions\ primrec mset :: "'a list \ 'a multiset" where "mset [] = {#}" | "mset (a # x) = add_mset a (mset x)" lemma in_multiset_in_set: "x \# mset xs \ x \ set xs" by (induct xs) simp_all lemma count_mset: "count (mset xs) x = length (filter (\y. x = y) xs)" by (induct xs) simp_all lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])" by (induct x) auto lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])" by (induct x) auto lemma count_mset_gt_0: "x \ set xs \ count (mset xs) x > 0" by (induction xs) auto lemma count_mset_0_iff [simp]: "count (mset xs) x = 0 \ x \ set xs" by (induction xs) auto lemma mset_single_iff[iff]: "mset xs = {#x#} \ xs = [x]" by (cases xs) auto lemma mset_single_iff_right[iff]: "{#x#} = mset xs \ xs = [x]" by (cases xs) auto lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs" by (induct xs) auto lemma set_mset_comp_mset [simp]: "set_mset \ mset = set" by (simp add: fun_eq_iff) lemma size_mset [simp]: "size (mset xs) = length xs" by (induct xs) simp_all lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" by (induct xs arbitrary: ys) auto lemma mset_filter[simp]: "mset (filter P xs) = {#x \# mset xs. P x #}" by (induct xs) simp_all lemma mset_rev [simp]: "mset (rev xs) = mset xs" by (induct xs) simp_all lemma surj_mset: "surj mset" unfolding surj_def proof (rule allI) fix M show "\xs. M = mset xs" by (induction M) (auto intro: exI[of _ "_ # _"]) qed lemma distinct_count_atmost_1: "distinct x = (\a. count (mset x) a = (if a \ set x then 1 else 0))" proof (induct x) case Nil then show ?case by simp next case (Cons x xs) show ?case (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs using Cons by simp next assume ?rhs then have "x \ set xs" by (simp split: if_splits) moreover from \?rhs\ have "(\a. count (mset xs) a = (if a \ set xs then 1 else 0))" by (auto split: if_splits simp add: count_eq_zero_iff) ultimately show ?lhs using Cons by simp qed qed lemma mset_eq_setD: assumes "mset xs = mset ys" shows "set xs = set ys" proof - from assms have "set_mset (mset xs) = set_mset (mset ys)" by simp then show ?thesis by simp qed lemma set_eq_iff_mset_eq_distinct: \distinct x \ distinct y \ set x = set y \ mset x = mset y\ by (auto simp: multiset_eq_iff distinct_count_atmost_1) lemma set_eq_iff_mset_remdups_eq: \set x = set y \ mset (remdups x) = mset (remdups y)\ apply (rule iffI) apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1]) apply (drule distinct_remdups [THEN distinct_remdups [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]]) apply simp done lemma mset_eq_imp_distinct_iff: \distinct xs \ distinct ys\ if \mset xs = mset ys\ using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD) lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}" by (induct xs) (auto simp add: multiset_eq_iff) lemma mset_eq_length: assumes "mset xs = mset ys" shows "length xs = length ys" using assms by (metis size_mset) lemma mset_eq_length_filter: assumes "mset xs = mset ys" shows "length (filter (\x. z = x) xs) = length (filter (\y. z = y) ys)" using assms by (metis count_mset) lemma fold_multiset_equiv: \List.fold f xs = List.fold f ys\ if f: \\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x\ and \mset xs = mset ys\ using f \mset xs = mset ys\ [symmetric] proof (induction xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then have *: \set ys = set (x # xs)\ by (blast dest: mset_eq_setD) have \\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x\ by (rule Cons.prems(1)) (simp_all add: *) moreover from * have \x \ set ys\ by simp ultimately have \List.fold f ys = List.fold f (remove1 x ys) \ f x\ by (fact fold_remove1_split) moreover from Cons.prems have \List.fold f xs = List.fold f (remove1 x ys)\ by (auto intro: Cons.IH) ultimately show ?case by simp qed lemma fold_permuted_eq: \List.fold (\) xs z = List.fold (\) ys z\ if \mset xs = mset ys\ and \P z\ and P: \\x z. x \ set xs \ P z \ P (x \ z)\ and f: \\x y z. x \ set xs \ y \ set xs \ P z \ x \ (y \ z) = y \ (x \ z)\ for f (infixl \\\ 70) using \P z\ P f \mset xs = mset ys\ [symmetric] proof (induction xs arbitrary: ys z) case Nil then show ?case by simp next case (Cons x xs) then have *: \set ys = set (x # xs)\ by (blast dest: mset_eq_setD) have \P z\ by (fact Cons.prems(1)) moreover have \\x z. x \ set ys \ P z \ P (x \ z)\ by (rule Cons.prems(2)) (simp_all add: *) moreover have \\x y z. x \ set ys \ y \ set ys \ P z \ x \ (y \ z) = y \ (x \ z)\ by (rule Cons.prems(3)) (simp_all add: *) moreover from * have \x \ set ys\ by simp ultimately have \fold (\) ys z = fold (\) (remove1 x ys) (x \ z)\ by (induction ys arbitrary: z) auto moreover from Cons.prems have \fold (\) xs (x \ z) = fold (\) (remove1 x ys) (x \ z)\ by (auto intro: Cons.IH) ultimately show ?case by simp qed lemma mset_shuffles: "zs \ shuffles xs ys \ mset zs = mset xs + mset ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma mset_insort [simp]: "mset (insort x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_map[simp]: "mset (map f xs) = image_mset f (mset xs)" by (induct xs) simp_all global_interpretation mset_set: folding add_mset "{#}" defines mset_set = "folding_on.F add_mset {#}" by standard (simp add: fun_eq_iff) lemma sum_multiset_singleton [simp]: "sum (\n. {#n#}) A = mset_set A" by (induction A rule: infinite_finite_induct) auto lemma count_mset_set [simp]: "finite A \ x \ A \ count (mset_set A) x = 1" (is "PROP ?P") "\ finite A \ count (mset_set A) x = 0" (is "PROP ?Q") "x \ A \ count (mset_set A) x = 0" (is "PROP ?R") proof - have *: "count (mset_set A) x = 0" if "x \ A" for A proof (cases "finite A") case False then show ?thesis by simp next case True from True \x \ A\ show ?thesis by (induct A) auto qed then show "PROP ?P" "PROP ?Q" "PROP ?R" by (auto elim!: Set.set_insert) qed \ \TODO: maybe define \<^const>\mset_set\ also in terms of \<^const>\Abs_multiset\\ lemma elem_mset_set[simp, intro]: "finite A \ x \# mset_set A \ x \ A" by (induct A rule: finite_induct) simp_all lemma mset_set_Union: "finite A \ finite B \ A \ B = {} \ mset_set (A \ B) = mset_set A + mset_set B" by (induction A rule: finite_induct) auto lemma filter_mset_mset_set [simp]: "finite A \ filter_mset P (mset_set A) = mset_set {x\A. P x}" proof (induction A rule: finite_induct) case (insert x A) from insert.hyps have "filter_mset P (mset_set (insert x A)) = filter_mset P (mset_set A) + mset_set (if P x then {x} else {})" by simp also have "filter_mset P (mset_set A) = mset_set {x\A. P x}" by (rule insert.IH) also from insert.hyps have "\ + mset_set (if P x then {x} else {}) = mset_set ({x \ A. P x} \ (if P x then {x} else {}))" (is "_ = mset_set ?A") by (intro mset_set_Union [symmetric]) simp_all also from insert.hyps have "?A = {y\insert x A. P y}" by auto finally show ?case . qed simp_all lemma mset_set_Diff: assumes "finite A" "B \ A" shows "mset_set (A - B) = mset_set A - mset_set B" proof - from assms have "mset_set ((A - B) \ B) = mset_set (A - B) + mset_set B" by (intro mset_set_Union) (auto dest: finite_subset) also from assms have "A - B \ B = A" by blast finally show ?thesis by simp qed lemma mset_set_set: "distinct xs \ mset_set (set xs) = mset xs" by (induction xs) simp_all lemma count_mset_set': "count (mset_set A) x = (if finite A \ x \ A then 1 else 0)" by auto lemma subset_imp_msubset_mset_set: assumes "A \ B" "finite B" shows "mset_set A \# mset_set B" proof (rule mset_subset_eqI) fix x :: 'a from assms have "finite A" by (rule finite_subset) with assms show "count (mset_set A) x \ count (mset_set B) x" by (cases "x \ A"; cases "x \ B") auto qed lemma mset_set_set_mset_msubset: "mset_set (set_mset A) \# A" proof (rule mset_subset_eqI) fix x show "count (mset_set (set_mset A)) x \ count A x" by (cases "x \# A") simp_all qed lemma mset_set_upto_eq_mset_upto: \mset_set {.. by (induction n) (auto simp: ac_simps lessThan_Suc) context linorder begin definition sorted_list_of_multiset :: "'a multiset \ 'a list" where "sorted_list_of_multiset M = fold_mset insort [] M" lemma sorted_list_of_multiset_empty [simp]: "sorted_list_of_multiset {#} = []" by (simp add: sorted_list_of_multiset_def) lemma sorted_list_of_multiset_singleton [simp]: "sorted_list_of_multiset {#x#} = [x]" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed lemma sorted_list_of_multiset_insert [simp]: "sorted_list_of_multiset (add_mset x M) = List.insort x (sorted_list_of_multiset M)" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed end lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M" by (induct M) simp_all lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs" by (induct xs) simp_all lemma finite_set_mset_mset_set[simp]: "finite A \ set_mset (mset_set A) = A" by auto lemma mset_set_empty_iff: "mset_set A = {#} \ A = {} \ infinite A" using finite_set_mset_mset_set by fastforce lemma infinite_set_mset_mset_set: "\ finite A \ set_mset (mset_set A) = {}" by simp lemma set_sorted_list_of_multiset [simp]: "set (sorted_list_of_multiset M) = set_mset M" by (induct M) (simp_all add: set_insort_key) lemma sorted_list_of_mset_set [simp]: "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma mset_upt [simp]: "mset [m.. {#the (map_of xs i). i \# mset (map fst xs)#} = mset (map snd xs)" proof (induction xs) case (Cons x xs) have "{#the (map_of (x # xs) i). i \# mset (map fst (x # xs))#} = add_mset (snd x) {#the (if i = fst x then Some (snd x) else map_of xs i). i \# mset (map fst xs)#}" (is "_ = add_mset _ ?A") by simp also from Cons.prems have "?A = {#the (map_of xs i). i :# mset (map fst xs)#}" by (cases x, intro image_mset_cong) (auto simp: in_multiset_in_set) also from Cons.prems have "\ = mset (map snd xs)" by (intro Cons.IH) simp_all finally show ?case by simp qed simp_all lemma msubset_mset_set_iff[simp]: assumes "finite A" "finite B" shows "mset_set A \# mset_set B \ A \ B" using assms set_mset_mono subset_imp_msubset_mset_set by fastforce lemma mset_set_eq_iff[simp]: assumes "finite A" "finite B" shows "mset_set A = mset_set B \ A = B" using assms by (fastforce dest: finite_set_mset_mset_set) lemma image_mset_mset_set: \<^marker>\contributor \Lukas Bulwahn\\ assumes "inj_on f A" shows "image_mset f (mset_set A) = mset_set (f ` A)" proof cases assume "finite A" from this \inj_on f A\ show ?thesis by (induct A) auto next assume "infinite A" from this \inj_on f A\ have "infinite (f ` A)" using finite_imageD by blast from \infinite A\ \infinite (f ` A)\ show ?thesis by simp qed subsection \More properties of the replicate and repeat operations\ lemma in_replicate_mset[simp]: "x \# replicate_mset n y \ n > 0 \ x = y" unfolding replicate_mset_def by (induct n) auto lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})" by (auto split: if_splits) lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" by (induct n, simp_all) lemma count_le_replicate_mset_subset_eq: "n \ count M x \ replicate_mset n x \# M" by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def) lemma filter_eq_replicate_mset: "{#y \# D. y = x#} = replicate_mset (count D x) x" by (induct D) simp_all lemma replicate_count_mset_eq_filter_eq: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" by (induct xs) auto lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \ n = 0" by (induct n) simp_all lemma replicate_mset_eq_iff: "replicate_mset m a = replicate_mset n b \ m = 0 \ n = 0 \ m = n \ a = b" by (auto simp add: multiset_eq_iff) lemma repeat_mset_cancel1: "repeat_mset a A = repeat_mset a B \ A = B \ a = 0" by (auto simp: multiset_eq_iff) lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \ a = b \ A = {#}" by (auto simp: multiset_eq_iff) lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \ n = 0 \ A = {#}" by (cases n) auto lemma image_replicate_mset [simp]: "image_mset f (replicate_mset n a) = replicate_mset n (f a)" by (induct n) simp_all lemma replicate_mset_msubseteq_iff: "replicate_mset m a \# replicate_mset n b \ m = 0 \ a = b \ m \ n" by (cases m) (auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq) lemma msubseteq_replicate_msetE: assumes "A \# replicate_mset n a" obtains m where "m \ n" and "A = replicate_mset m a" proof (cases "n = 0") case True with assms that show thesis by simp next case False from assms have "set_mset A \ set_mset (replicate_mset n a)" by (rule set_mset_mono) with False have "set_mset A \ {a}" by simp then have "\m. A = replicate_mset m a" proof (induction A) case empty then show ?case by simp next case (add b A) then obtain m where "A = replicate_mset m a" by auto with add.prems show ?case by (auto intro: exI [of _ "Suc m"]) qed then obtain m where A: "A = replicate_mset m a" .. with assms have "m \ n" by (auto simp add: replicate_mset_msubseteq_iff) then show thesis using A .. qed subsection \Big operators\ locale comm_monoid_mset = comm_monoid begin interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) context begin definition F :: "'a multiset \ 'a" where eq_fold: "F M = fold_mset f \<^bold>1 M" lemma empty [simp]: "F {#} = \<^bold>1" by (simp add: eq_fold) lemma singleton [simp]: "F {#x#} = x" proof - interpret comp_fun_commute by standard (simp add: fun_eq_iff left_commute) show ?thesis by (simp add: eq_fold) qed lemma union [simp]: "F (M + N) = F M \<^bold>* F N" proof - interpret comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) show ?thesis by (induct N) (simp_all add: left_commute eq_fold) qed lemma add_mset [simp]: "F (add_mset x N) = x \<^bold>* F N" unfolding add_mset_add_single[of x N] union by (simp add: ac_simps) lemma insert [simp]: shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)" by (simp add: eq_fold) lemma remove: assumes "x \# A" shows "F A = x \<^bold>* F (A - {#x#})" using multi_member_split[OF assms] by auto lemma neutral: "\x\#A. x = \<^bold>1 \ F A = \<^bold>1" by (induct A) simp_all lemma neutral_const [simp]: "F (image_mset (\_. \<^bold>1) A) = \<^bold>1" by (simp add: neutral) private lemma F_image_mset_product: "F {#g x j \<^bold>* F {#g i j. i \# A#}. j \# B#} = F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \# A#}. j \# B#}" by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms) lemma swap: "F (image_mset (\i. F (image_mset (g i) B)) A) = F (image_mset (\j. F (image_mset (\i. g i j) A)) B)" apply (induction A, simp) apply (induction B, auto simp add: F_image_mset_product ac_simps) done lemma distrib: "F (image_mset (\x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)" by (induction A) (auto simp: ac_simps) lemma union_disjoint: "A \# B = {#} \ F (image_mset g (A \# B)) = F (image_mset g A) \<^bold>* F (image_mset g B)" by (induction A) (auto simp: ac_simps) end end lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \ _ \ _)" by standard (simp add: add_ac comp_def) declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp] lemma in_mset_fold_plus_iff[iff]: "x \# fold_mset (+) M NN \ x \# M \ (\N. N \# NN \ x \# N)" by (induct NN) auto context comm_monoid_add begin sublocale sum_mset: comm_monoid_mset plus 0 defines sum_mset = sum_mset.F .. lemma sum_unfold_sum_mset: "sum f A = sum_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) end notation sum_mset ("\\<^sub>#") syntax (ASCII) "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) syntax "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST sum_mset (CONST image_mset (\i. b) A)" context comm_monoid_add begin lemma sum_mset_sum_list: "sum_mset (mset xs) = sum_list xs" by (induction xs) auto end context canonically_ordered_monoid_add begin lemma sum_mset_0_iff [simp]: "sum_mset M = 0 \ (\x \ set_mset M. x = 0)" by (induction M) auto end context ordered_comm_monoid_add begin lemma sum_mset_mono: "sum_mset (image_mset f K) \ sum_mset (image_mset g K)" if "\i. i \# K \ f i \ g i" using that by (induction K) (simp_all add: add_mono) end context cancel_comm_monoid_add begin lemma sum_mset_diff: "sum_mset (M - N) = sum_mset M - sum_mset N" if "N \# M" for M N :: "'a multiset" using that by (auto simp add: subset_mset.le_iff_add) end context semiring_0 begin lemma sum_mset_distrib_left: "c * (\x \# M. f x) = (\x \# M. c * f(x))" by (induction M) (simp_all add: algebra_simps) lemma sum_mset_distrib_right: "(\x \# M. f x) * c = (\x \# M. f x * c)" by (induction M) (simp_all add: algebra_simps) end lemma sum_mset_product: fixes f :: "'a::{comm_monoid_add,times} \ 'b::semiring_0" shows "(\i \# A. f i) * (\i \# B. g i) = (\i\#A. \j\#B. f i * g j)" by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right) context semiring_1 begin lemma sum_mset_replicate_mset [simp]: "sum_mset (replicate_mset n a) = of_nat n * a" by (induction n) (simp_all add: algebra_simps) lemma sum_mset_delta: "sum_mset (image_mset (\x. if x = y then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) lemma sum_mset_delta': "sum_mset (image_mset (\x. if y = x then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) end lemma of_nat_sum_mset [simp]: "of_nat (sum_mset A) = sum_mset (image_mset of_nat A)" by (induction A) auto lemma size_eq_sum_mset: "size M = (\a\#M. 1)" using image_mset_const_eq [of "1::nat" M] by simp lemma size_mset_set [simp]: "size (mset_set A) = card A" by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset) lemma sum_mset_constant [simp]: fixes y :: "'b::semiring_1" shows \(\x\#A. y) = of_nat (size A) * y\ by (induction A) (auto simp: algebra_simps) lemma set_mset_Union_mset[simp]: "set_mset (\\<^sub># MM) = (\M \ set_mset MM. set_mset M)" by (induct MM) auto lemma in_Union_mset_iff[iff]: "x \# \\<^sub># MM \ (\M. M \# MM \ x \# M)" by (induct MM) auto lemma count_sum: "count (sum f A) x = sum (\a. count (f a) x) A" by (induct A rule: infinite_finite_induct) simp_all lemma sum_eq_empty_iff: assumes "finite A" shows "sum f A = {#} \ (\a\A. f a = {#})" using assms by induct simp_all lemma Union_mset_empty_conv[simp]: "\\<^sub># M = {#} \ (\i\#M. i = {#})" by (induction M) auto lemma Union_image_single_mset[simp]: "\\<^sub># (image_mset (\x. {#x#}) m) = m" by(induction m) auto context comm_monoid_mult begin sublocale prod_mset: comm_monoid_mset times 1 defines prod_mset = prod_mset.F .. lemma prod_mset_empty: "prod_mset {#} = 1" by (fact prod_mset.empty) lemma prod_mset_singleton: "prod_mset {#x#} = x" by (fact prod_mset.singleton) lemma prod_mset_Un: "prod_mset (A + B) = prod_mset A * prod_mset B" by (fact prod_mset.union) lemma prod_mset_prod_list: "prod_mset (mset xs) = prod_list xs" by (induct xs) auto lemma prod_mset_replicate_mset [simp]: "prod_mset (replicate_mset n a) = a ^ n" by (induct n) simp_all lemma prod_unfold_prod_mset: "prod f A = prod_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma prod_mset_multiplicity: "prod_mset M = prod (\x. x ^ count M x) (set_mset M)" by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def) lemma prod_mset_delta: "prod_mset (image_mset (\x. if x = y then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_delta': "prod_mset (image_mset (\x. if y = x then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_subset_imp_dvd: assumes "A \# B" shows "prod_mset A dvd prod_mset B" proof - from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) also have "prod_mset \ = prod_mset (B - A) * prod_mset A" by simp also have "prod_mset A dvd \" by simp finally show ?thesis . qed lemma dvd_prod_mset: assumes "x \# A" shows "x dvd prod_mset A" using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp end notation prod_mset ("\\<^sub>#") syntax (ASCII) "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10) syntax "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST prod_mset (CONST image_mset (\i. b) A)" lemma prod_mset_constant [simp]: "(\_\#A. c) = c ^ size A" by (simp add: image_mset_const_eq) lemma (in semidom) prod_mset_zero_iff [iff]: "prod_mset A = 0 \ 0 \# A" by (induct A) auto lemma (in semidom_divide) prod_mset_diff: assumes "B \# A" and "0 \# B" shows "prod_mset (A - B) = prod_mset A div prod_mset B" proof - from assms obtain C where "A = B + C" by (metis subset_mset.add_diff_inverse) with assms show ?thesis by simp qed lemma (in semidom_divide) prod_mset_minus: assumes "a \# A" and "a \ 0" shows "prod_mset (A - {#a#}) = prod_mset A div a" using assms prod_mset_diff [of "{#a#}" A] by auto lemma (in normalization_semidom) normalize_prod_mset_normalize: "normalize (prod_mset (image_mset normalize A)) = normalize (prod_mset A)" proof (induction A) case (add x A) have "normalize (prod_mset (image_mset normalize (add_mset x A))) = normalize (x * normalize (prod_mset (image_mset normalize A)))" by simp also note add.IH finally show ?case by simp qed auto lemma (in algebraic_semidom) is_unit_prod_mset_iff: "is_unit (prod_mset A) \ (\x \# A. is_unit x)" by (induct A) (auto simp: is_unit_mult_iff) lemma (in normalization_semidom_multiplicative) normalize_prod_mset: "normalize (prod_mset A) = prod_mset (image_mset normalize A)" by (induct A) (simp_all add: normalize_mult) lemma (in normalization_semidom_multiplicative) normalized_prod_msetI: assumes "\a. a \# A \ normalize a = a" shows "normalize (prod_mset A) = prod_mset A" proof - from assms have "image_mset normalize A = A" by (induct A) simp_all then show ?thesis by (simp add: normalize_prod_mset) qed subsection \Multiset as order-ignorant lists\ context linorder begin lemma mset_insort [simp]: "mset (insort_key k x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_sort [simp]: "mset (sort_key k xs) = mset xs" by (induct xs) simp_all text \ This lemma shows which properties suffice to show that a function \f\ with \f xs = ys\ behaves like sort. \ lemma properties_for_sort_key: assumes "mset ys = mset xs" and "\k. k \ set ys \ filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" and "sorted (map f ys)" shows "sort_key f xs = ys" using assms proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons.prems(2) have "\k \ set ys. filter (\x. f k = f x) (remove1 x ys) = filter (\x. f k = f x) xs" by (simp add: filter_remove1) with Cons.prems have "sort_key f xs = remove1 x ys" by (auto intro!: Cons.hyps simp add: sorted_map_remove1) moreover from Cons.prems have "x \# mset ys" by auto then have "x \ set ys" by simp ultimately show ?case using Cons.prems by (simp add: insort_key_remove1) qed lemma properties_for_sort: assumes multiset: "mset ys = mset xs" and "sorted ys" shows "sort xs = ys" proof (rule properties_for_sort_key) from multiset show "mset ys = mset xs" . from \sorted ys\ show "sorted (map (\x. x) ys)" by simp from multiset have "length (filter (\y. k = y) ys) = length (filter (\x. k = x) xs)" for k by (rule mset_eq_length_filter) then have "replicate (length (filter (\y. k = y) ys)) k = replicate (length (filter (\x. k = x) xs)) k" for k by simp then show "k \ set ys \ filter (\y. k = y) ys = filter (\x. k = x) xs" for k by (simp add: replicate_length_filter) qed lemma sort_key_inj_key_eq: assumes mset_equal: "mset xs = mset ys" and "inj_on f (set xs)" and "sorted (map f ys)" shows "sort_key f xs = ys" proof (rule properties_for_sort_key) from mset_equal show "mset ys = mset xs" by simp from \sorted (map f ys)\ show "sorted (map f ys)" . show "[x\ys . f k = f x] = [x\xs . f k = f x]" if "k \ set ys" for k proof - from mset_equal have set_equal: "set xs = set ys" by (rule mset_eq_setD) with that have "insert k (set ys) = set ys" by auto with \inj_on f (set xs)\ have inj: "inj_on f (insert k (set ys))" by (simp add: set_equal) from inj have "[x\ys . f k = f x] = filter (HOL.eq k) ys" by (auto intro!: inj_on_filter_key_eq) also have "\ = replicate (count (mset ys) k) k" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = replicate (count (mset xs) k) k" using mset_equal by simp also have "\ = filter (HOL.eq k) xs" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = [x\xs . f k = f x]" using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal) finally show ?thesis . qed qed lemma sort_key_eq_sort_key: assumes "mset xs = mset ys" and "inj_on f (set xs)" shows "sort_key f xs = sort_key f ys" by (rule sort_key_inj_key_eq) (simp_all add: assms) lemma sort_key_by_quicksort: "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] @ [x\xs. f x = f (xs ! (length xs div 2))] @ sort_key f [x\xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") proof (rule properties_for_sort_key) show "mset ?rhs = mset ?lhs" by (rule multiset_eqI) auto show "sorted (map f ?rhs)" by (auto simp add: sorted_append intro: sorted_map_same) next fix l assume "l \ set ?rhs" let ?pivot = "f (xs ! (length xs div 2))" have *: "\x. f l = f x \ f x = f l" by auto have "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) with * have **: "[x \ sort_key f xs . f l = f x] = [x \ xs. f l = f x]" by simp have "\x P. P (f x) ?pivot \ f l = f x \ P (f l) ?pivot \ f l = f x" by auto then have "\P. [x \ sort_key f xs . P (f x) ?pivot \ f l = f x] = [x \ sort_key f xs. P (f l) ?pivot \ f l = f x]" by simp note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] show "[x \ ?rhs. f l = f x] = [x \ ?lhs. f l = f x]" proof (cases "f l" ?pivot rule: linorder_cases) case less then have "f l \ ?pivot" and "\ f l > ?pivot" by auto with less show ?thesis by (simp add: filter_sort [symmetric] ** ***) next case equal then show ?thesis by (simp add: * less_le) next case greater then have "f l \ ?pivot" and "\ f l < ?pivot" by auto with greater show ?thesis by (simp add: filter_sort [symmetric] ** ***) qed qed lemma sort_by_quicksort: "sort xs = sort [x\xs. x < xs ! (length xs div 2)] @ [x\xs. x = xs ! (length xs div 2)] @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") using sort_key_by_quicksort [of "\x. x", symmetric] by simp text \A stable parameterized quicksort\ definition part :: "('b \ 'a) \ 'a \ 'b list \ 'b list \ 'b list \ 'b list" where "part f pivot xs = ([x \ xs. f x < pivot], [x \ xs. f x = pivot], [x \ xs. pivot < f x])" lemma part_code [code]: "part f pivot [] = ([], [], [])" "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in if x' < pivot then (x # lts, eqs, gts) else if x' > pivot then (lts, eqs, x # gts) else (lts, x # eqs, gts))" by (auto simp add: part_def Let_def split_def) lemma sort_key_by_quicksort_code [code]: "sort_key f xs = (case xs of [] \ [] | [x] \ xs | [x, y] \ (if f x \ f y then xs else [y, x]) | _ \ let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" proof (cases xs) case Nil then show ?thesis by simp next case (Cons _ ys) note hyps = Cons show ?thesis proof (cases ys) case Nil with hyps show ?thesis by simp next case (Cons _ zs) note hyps = hyps Cons show ?thesis proof (cases zs) case Nil with hyps show ?thesis by auto next case Cons from sort_key_by_quicksort [of f xs] have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" by (simp only: split_def Let_def part_def fst_conv snd_conv) with hyps Cons show ?thesis by (simp only: list.cases) qed qed qed end hide_const (open) part lemma mset_remdups_subset_eq: "mset (remdups xs) \# mset xs" by (induct xs) (auto intro: subset_mset.order_trans) lemma mset_update: "i < length ls \ mset (ls[i := v]) = add_mset v (mset ls - {#ls ! i#})" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases i) case 0 then show ?thesis by simp next case (Suc i') with Cons show ?thesis by (cases \x = xs ! i'\) auto qed qed lemma mset_swap: "i < length ls \ j < length ls \ mset (ls[j := ls ! i, i := ls ! j]) = mset ls" by (cases "i = j") (simp_all add: mset_update nth_mem_mset) lemma mset_eq_finite: \finite {ys. mset ys = mset xs}\ proof - have \{ys. mset ys = mset xs} \ {ys. set ys \ set xs \ length ys \ length xs}\ by (auto simp add: dest: mset_eq_setD mset_eq_length) moreover have \finite {ys. set ys \ set xs \ length ys \ length xs}\ using finite_lists_length_le by blast ultimately show ?thesis by (rule finite_subset) qed subsection \The multiset order\ definition mult1 :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult1 r = {(N, M). \a M0 K. M = add_mset a M0 \ N = M0 + K \ (\b. b \# K \ (b, a) \ r)}" definition mult :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult r = (mult1 r)\<^sup>+" definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp r M N \ (M, N) \ mult {(x, y). r x y}" declare multp_def[pred_set_conv] lemma mult1I: assumes "M = add_mset a M0" and "N = M0 + K" and "\b. b \# K \ (b, a) \ r" shows "(N, M) \ mult1 r" using assms unfolding mult1_def by blast lemma mult1E: assumes "(N, M) \ mult1 r" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "\b. b \# K \ (b, a) \ r" using assms unfolding mult1_def by blast lemma mono_mult1: assumes "r \ r'" shows "mult1 r \ mult1 r'" unfolding mult1_def using assms by blast lemma mono_mult: assumes "r \ r'" shows "mult r \ mult r'" unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast lemma mono_multp[mono]: "r \ r' \ multp r \ multp r'" unfolding le_fun_def le_bool_def proof (intro allI impI) fix M N :: "'a multiset" assume "\x xa. r x xa \ r' x xa" hence "{(x, y). r x y} \ {(x, y). r' x y}" by blast thus "multp r M N \ multp r' M N" unfolding multp_def by (fact mono_mult[THEN subsetD, rotated]) qed lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) subsubsection \Well-foundedness\ lemma less_add: assumes mult1: "(N, add_mset a M0) \ mult1 r" shows "(\M. (M, M0) \ mult1 r \ N = add_mset a M) \ (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K)" proof - let ?r = "\K a. \b. b \# K \ (b, a) \ r" let ?R = "\N M. \a M0 K. M = add_mset a M0 \ N = M0 + K \ ?r K a" obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'" and N: "N = M0' + K" and r: "?r K a'" using mult1 unfolding mult1_def by auto show ?thesis (is "?case1 \ ?case2") proof - from M0 consider "M0 = M0'" "a = a'" | K' where "M0 = add_mset a' K'" "M0' = add_mset a K'" by atomize_elim (simp only: add_eq_conv_ex) then show ?thesis proof cases case 1 with N r have "?r K a \ N = M0 + K" by simp then have ?case2 .. then show ?thesis .. next case 2 from N 2(2) have n: "N = add_mset a (K' + K)" by simp with r 2(1) have "?R (K' + K) M0" by blast with n have ?case1 by (simp add: mult1_def) then show ?thesis .. qed qed qed lemma all_accessible: assumes "wf r" shows "\M. M \ Wellfounded.acc (mult1 r)" proof let ?R = "mult1 r" let ?W = "Wellfounded.acc ?R" { fix M M0 a assume M0: "M0 \ ?W" and wf_hyp: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" and acc_hyp: "\M. (M, M0) \ ?R \ add_mset a M \ ?W" have "add_mset a M0 \ ?W" proof (rule accI [of "add_mset a M0"]) fix N assume "(N, add_mset a M0) \ ?R" then consider M where "(M, M0) \ ?R" "N = add_mset a M" | K where "\b. b \# K \ (b, a) \ r" "N = M0 + K" by atomize_elim (rule less_add) then show "N \ ?W" proof cases case 1 from acc_hyp have "(M, M0) \ ?R \ add_mset a M \ ?W" .. from this and \(M, M0) \ ?R\ have "add_mset a M \ ?W" .. then show "N \ ?W" by (simp only: \N = add_mset a M\) next case 2 from this(1) have "M0 + K \ ?W" proof (induct K) case empty from M0 show "M0 + {#} \ ?W" by simp next case (add x K) from add.prems have "(x, a) \ r" by simp with wf_hyp have "\M \ ?W. add_mset x M \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "add_mset x (M0 + K) \ ?W" .. then show "M0 + (add_mset x K) \ ?W" by simp qed then show "N \ ?W" by (simp only: 2(2)) qed qed } note tedious_reasoning = this show "M \ ?W" for M proof (induct M) show "{#} \ ?W" proof (rule accI) fix b assume "(b, {#}) \ ?R" with not_less_empty show "b \ ?W" by contradiction qed fix M a assume "M \ ?W" from \wf r\ have "\M \ ?W. add_mset a M \ ?W" proof induct fix a assume r: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" show "\M \ ?W. add_mset a M \ ?W" proof fix M assume "M \ ?W" then show "add_mset a M \ ?W" by (rule acc_induct) (rule tedious_reasoning [OF _ r]) qed qed from this and \M \ ?W\ show "add_mset a M \ ?W" .. qed qed lemma wf_mult1: "wf r \ wf (mult1 r)" by (rule acc_wfI) (rule all_accessible) lemma wf_mult: "wf r \ wf (mult r)" unfolding mult_def by (rule wf_trancl) (rule wf_mult1) lemma wfP_multp: "wfP r \ wfP (multp r)" unfolding multp_def wfP_def by (simp add: wf_mult) subsubsection \Closure-free presentation\ text \One direction.\ lemma mult_implies_one_step: assumes trans: "trans r" and MN: "(M, N) \ mult r" shows "\I J K. N = I + J \ M = I + K \ J \ {#} \ (\k \ set_mset K. \j \ set_mset J. (k, j) \ r)" using MN unfolding mult_def mult1_def proof (induction rule: converse_trancl_induct) case (base y) then show ?case by force next case (step y z) note yz = this(1) and zN = this(2) and N_decomp = this(3) obtain I J K where N: "N = I + J" "z = I + K" "J \ {#}" "\k\#K. \j\#J. (k, j) \ r" using N_decomp by blast obtain a M0 K' where z: "z = add_mset a M0" and y: "y = M0 + K'" and K: "\b. b \# K' \ (b, a) \ r" using yz by blast show ?case proof (cases "a \# K") case True moreover have "\j\#J. (k, j) \ r" if "k \# K'" for k using K N trans True by (meson that transE) ultimately show ?thesis by (rule_tac x = I in exI, rule_tac x = J in exI, rule_tac x = "(K - {#a#}) + K'" in exI) (use z y N in \auto simp del: subset_mset.add_diff_assoc2 dest: in_diffD\) next case False then have "a \# I" by (metis N(2) union_iff union_single_eq_member z) moreover have "M0 = I + K - {#a#}" using N(2) z by force ultimately show ?thesis by (rule_tac x = "I - {#a#}" in exI, rule_tac x = "add_mset a J" in exI, rule_tac x = "K + K'" in exI) (use z y N False K in \auto simp: add.assoc\) qed qed lemma multp_implies_one_step: "transp R \ multp R M N \ \I J K. N = I + J \ M = I + K \ J \ {#} \ (\k\#K. \x\#J. R k x)" by (rule mult_implies_one_step[to_pred]) lemma one_step_implies_mult: assumes "J \ {#}" and "\k \ set_mset K. \j \ set_mset J. (k, j) \ r" shows "(I + K, I + J) \ mult r" using assms proof (induction "size J" arbitrary: I J K) case 0 then show ?case by auto next case (Suc n) note IH = this(1) and size_J = this(2)[THEN sym] obtain J' a where J: "J = add_mset a J'" using size_J by (blast dest: size_eq_Suc_imp_eq_union) show ?case proof (cases "J' = {#}") case True then show ?thesis using J Suc by (fastforce simp add: mult_def mult1_def) next case [simp]: False have K: "K = {#x \# K. (x, a) \ r#} + {#x \# K. (x, a) \ r#}" by simp have "(I + K, (I + {# x \# K. (x, a) \ r #}) + J') \ mult r" using IH[of J' "{# x \# K. (x, a) \ r#}" "I + {# x \# K. (x, a) \ r#}"] J Suc.prems K size_J by (auto simp: ac_simps) moreover have "(I + {#x \# K. (x, a) \ r#} + J', I + J) \ mult r" by (fastforce simp: J mult1_def mult_def) ultimately show ?thesis unfolding mult_def by simp qed qed lemma one_step_implies_multp: "J \ {#} \ \k\#K. \j\#J. R k j \ multp R (I + K) (I + J)" by (rule one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified]) lemma subset_implies_mult: assumes sub: "A \# B" shows "(A, B) \ mult r" proof - have ApBmA: "A + (B - A) = B" using sub by simp have BmA: "B - A \ {#}" using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le) thus ?thesis by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified]) qed lemma subset_implies_multp: "A \# B \ multp r A B" by (rule subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def]) lemma multp_repeat_mset_repeat_msetI: assumes "transp R" and "multp R A B" and "n \ 0" shows "multp R (repeat_mset n A) (repeat_mset n B)" proof - from \transp R\ \multp R A B\ obtain I J K where "B = I + J" and "A = I + K" and "J \ {#}" and "\k \# K. \x \# J. R k x" by (auto dest: multp_implies_one_step) have repeat_n_A_eq: "repeat_mset n A = repeat_mset n I + repeat_mset n K" using \A = I + K\ by simp have repeat_n_B_eq: "repeat_mset n B = repeat_mset n I + repeat_mset n J" using \B = I + J\ by simp show ?thesis unfolding repeat_n_A_eq repeat_n_B_eq proof (rule one_step_implies_multp) from \n \ 0\ show "repeat_mset n J \ {#}" using \J \ {#}\ by (simp add: repeat_mset_eq_empty_iff) next show "\k \# repeat_mset n K. \j \# repeat_mset n J. R k j" using \\k \# K. \x \# J. R k x\ by (metis count_greater_zero_iff nat_0_less_mult_iff repeat_mset.rep_eq) qed qed subsubsection \Monotonicity\ lemma multp_mono_strong: assumes "multp R M1 M2" and "transp R" and S_if_R: "\x y. x \ set_mset M1 \ y \ set_mset M2 \ R x y \ S x y" shows "multp S M1 M2" proof - obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \ {#}" and "\k\#K. \x\#J. R k x" using multp_implies_one_step[OF \transp R\ \multp R M1 M2\] by auto show ?thesis unfolding \M2 = I + J\ \M1 = I + K\ proof (rule one_step_implies_multp[OF \J \ {#}\]) show "\k\#K. \j\#J. S k j" using S_if_R by (metis \M1 = I + K\ \M2 = I + J\ \\k\#K. \x\#J. R k x\ union_iff) qed qed lemma mult_mono_strong: assumes "(M1, M2) \ mult r" and "trans r" and S_if_R: "\x y. x \ set_mset M1 \ y \ set_mset M2 \ (x, y) \ r \ (x, y) \ s" shows "(M1, M2) \ mult s" using assms multp_mono_strong[of "\x y. (x, y) \ r" M1 M2 "\x y. (x, y) \ s", unfolded multp_def transp_trans_eq, simplified] by blast lemma monotone_on_multp_multp_image_mset: assumes "monotone_on A orda ordb f" and "transp orda" shows "monotone_on {M. set_mset M \ A} (multp orda) (multp ordb) (image_mset f)" proof (rule monotone_onI) fix M1 M2 assume M1_in: "M1 \ {M. set_mset M \ A}" and M2_in: "M2 \ {M. set_mset M \ A}" and M1_lt_M2: "multp orda M1 M2" from multp_implies_one_step[OF \transp orda\ M1_lt_M2] obtain I J K where M2_eq: "M2 = I + J" and M1_eq: "M1 = I + K" and J_neq_mempty: "J \ {#}" and ball_K_less: "\k\#K. \x\#J. orda k x" by metis have "multp ordb (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" proof (intro one_step_implies_multp ballI) show "image_mset f J \ {#}" using J_neq_mempty by simp next fix k' assume "k'\#image_mset f K" then obtain k where "k' = f k" and k_in: "k \# K" by auto then obtain j where j_in: "j\#J" and "orda k j" using ball_K_less by auto have "ordb (f k) (f j)" proof (rule \monotone_on A orda ordb f\[THEN monotone_onD, OF _ _ \orda k j\]) show "k \ A" using M1_eq M1_in k_in by auto next show "j \ A" using M2_eq M2_in j_in by auto qed thus "\j\#image_mset f J. ordb k' j" using \j \# J\ \k' = f k\ by auto qed thus "multp ordb (image_mset f M1) (image_mset f M2)" by (simp add: M1_eq M2_eq) qed lemma monotone_multp_multp_image_mset: assumes "monotone orda ordb f" and "transp orda" shows "monotone (multp orda) (multp ordb) (image_mset f)" by (rule monotone_on_multp_multp_image_mset[OF assms, simplified]) lemma multp_image_mset_image_msetI: assumes "multp (\x y. R (f x) (f y)) M1 M2" and "transp R" shows "multp R (image_mset f M1) (image_mset f M2)" proof - from \transp R\ have "transp (\x y. R (f x) (f y))" by (auto intro: transpI dest: transpD) with \multp (\x y. R (f x) (f y)) M1 M2\ obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \ {#}" and "\k\#K. \x\#J. R (f k) (f x)" using multp_implies_one_step by blast have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" proof (rule one_step_implies_multp) show "image_mset f J \ {#}" by (simp add: \J \ {#}\) next show "\k\#image_mset f K. \j\#image_mset f J. R k j" by (simp add: \\k\#K. \x\#J. R (f k) (f x)\) qed thus ?thesis by (simp add: \M1 = I + K\ \M2 = I + J\) qed lemma multp_image_mset_image_msetD: assumes "multp R (image_mset f A) (image_mset f B)" and "transp R" and inj_on_f: "inj_on f (set_mset A \ set_mset B)" shows "multp (\x y. R (f x) (f y)) A B" proof - from assms(1,2) obtain I J K where f_B_eq: "image_mset f B = I + J" and f_A_eq: "image_mset f A = I + K" and J_neq_mempty: "J \ {#}" and ball_K_less: "\k\#K. \x\#J. R k x" by (auto dest: multp_implies_one_step) from f_B_eq obtain I' J' where B_def: "B = I' + J'" and I_def: "I = image_mset f I'" and J_def: "J = image_mset f J'" using image_mset_eq_plusD by blast from inj_on_f have inj_on_f': "inj_on f (set_mset A \ set_mset I')" by (rule inj_on_subset) (auto simp add: B_def) from f_A_eq obtain K' where A_def: "A = I' + K'" and K_def: "K = image_mset f K'" by (auto simp: I_def dest: image_mset_eq_image_mset_plusD[OF _ inj_on_f']) show ?thesis unfolding A_def B_def proof (intro one_step_implies_multp ballI) from J_neq_mempty show "J' \ {#}" by (simp add: J_def) next fix k assume "k \# K'" with ball_K_less obtain j' where "j' \# J" and "R (f k) j'" using K_def by auto moreover then obtain j where "j \# J'" and "f j = j'" using J_def by auto ultimately show "\j\#J'. R (f k) (f j)" by blast qed qed subsubsection \The multiset extension is cancellative for multiset union\ lemma mult_cancel: assumes "trans s" and "irrefl_on (set_mset Z) s" shows "(X + Z, Y + Z) \ mult s \ (X, Y) \ mult s" (is "?L \ ?R") proof assume ?L thus ?R using \irrefl_on (set_mset Z) s\ proof (induct Z) case (add z Z) obtain X' Y' Z' where *: "add_mset z X + Z = Z' + X'" "add_mset z Y + Z = Z' + Y'" "Y' \ {#}" "\x \ set_mset X'. \y \ set_mset Y'. (x, y) \ s" using mult_implies_one_step[OF \trans s\ add(2)] by auto consider Z2 where "Z' = add_mset z Z2" | X2 Y2 where "X' = add_mset z X2" "Y' = add_mset z Y2" using *(1,2) by (metis add_mset_remove_trivial_If insert_iff set_mset_add_mset_insert union_iff) thus ?case proof (cases) case 1 thus ?thesis using * one_step_implies_mult[of Y' X' s Z2] add(3) by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1) elim: irrefl_on_subset) next case 2 then obtain y where "y \ set_mset Y2" "(z, y) \ s" using *(4) \irrefl_on (set_mset (add_mset z Z)) s\ by (auto simp: irrefl_on_def) moreover from this transD[OF \trans s\ _ this(2)] have "x' \ set_mset X2 \ \y \ set_mset Y2. (x', y) \ s" for x' using 2 *(4)[rule_format, of x'] by auto ultimately show ?thesis using * one_step_implies_mult[of Y2 X2 s Z'] 2 add(3) by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1) elim: irrefl_on_subset) qed qed auto next assume ?R then obtain I J K where "Y = I + J" "X = I + K" "J \ {#}" "\k \ set_mset K. \j \ set_mset J. (k, j) \ s" using mult_implies_one_step[OF \trans s\] by blast thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) qed lemma multp_cancel: "transp R \ irreflp_on (set_mset Z) R \ multp R (X + Z) (Y + Z) \ multp R X Y" by (rule mult_cancel[to_pred]) lemma mult_cancel_add_mset: "trans r \ irrefl_on {z} r \ ((add_mset z X, add_mset z Y) \ mult r) = ((X, Y) \ mult r)" by (rule mult_cancel[of _ "{#_#}", simplified]) lemma multp_cancel_add_mset: "transp R \ irreflp_on {z} R \ multp R (add_mset z X) (add_mset z Y) = multp R X Y" by (rule mult_cancel_add_mset[to_pred, folded bot_set_def]) lemma mult_cancel_max0: assumes "trans s" and "irrefl_on (set_mset X \ set_mset Y) s" shows "(X, Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" (is "?L \ ?R") proof - have "(X - X \# Y + X \# Y, Y - X \# Y + X \# Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" proof (rule mult_cancel) from assms show "trans s" by simp next from assms show "irrefl_on (set_mset (X \# Y)) s" by simp qed moreover have "X - X \# Y + X \# Y = X" "Y - X \# Y + X \# Y = Y" by (auto simp flip: count_inject) ultimately show ?thesis by simp qed lemma mult_cancel_max: "trans r \ irrefl_on (set_mset X \ set_mset Y) r \ (X, Y) \ mult r \ (X - Y, Y - X) \ mult r" by (rule mult_cancel_max0[simplified]) lemma multp_cancel_max: "transp R \ irreflp_on (set_mset X \ set_mset Y) R \ multp R X Y \ multp R (X - Y) (Y - X)" by (rule mult_cancel_max[to_pred]) subsubsection \Strict partial-order properties\ lemma mult1_lessE: assumes "(N, M) \ mult1 {(a, b). r a b}" and "asymp r" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ r b a" proof - from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and *: "b \# K \ r b a" for b by (blast elim: mult1E) moreover from * [of a] have "a \# K" using \asymp r\ by (meson asympD) ultimately show thesis by (auto intro: that) qed lemma trans_mult: "trans r \ trans (mult r)" by (simp add: mult_def) lemma transp_multp: "transp r \ transp (multp r)" unfolding multp_def transp_trans_eq by (fact trans_mult[of "{(x, y). r x y}" for r, folded transp_trans]) lemma irrefl_mult: assumes "trans r" "irrefl r" shows "irrefl (mult r)" proof (intro irreflI notI) fix M assume "(M, M) \ mult r" then obtain I J K where "M = I + J" and "M = I + K" and "J \ {#}" and "(\k\set_mset K. \j\set_mset J. (k, j) \ r)" using mult_implies_one_step[OF \trans r\] by blast then have *: "K \ {#}" and **: "\k\set_mset K. \j\set_mset K. (k, j) \ r" by auto have "finite (set_mset K)" by simp hence "set_mset K = {}" using ** proof (induction rule: finite_induct) case empty thus ?case by simp next case (insert x F) have False using \irrefl r\[unfolded irrefl_def, rule_format] using \trans r\[THEN transD] by (metis equals0D insert.IH insert.prems insertE insertI1 insertI2) thus ?case .. qed with * show False by simp qed lemma irreflp_multp: "transp R \ irreflp R \ irreflp (multp R)" by (rule irrefl_mult[of "{(x, y). r x y}" for r, folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def]) instantiation multiset :: (preorder) order begin definition less_multiset :: "'a multiset \ 'a multiset \ bool" where "M < N \ multp (<) M N" definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" where "less_eq_multiset M N \ M < N \ M = N" instance proof intro_classes fix M N :: "'a multiset" show "(M < N) = (M \ N \ \ N \ M)" unfolding less_eq_multiset_def less_multiset_def by (metis irreflp_def irreflp_on_less irreflp_multp transpE transp_on_less transp_multp) next fix M :: "'a multiset" show "M \ M" unfolding less_eq_multiset_def by simp next fix M1 M2 M3 :: "'a multiset" show "M1 \ M2 \ M2 \ M3 \ M1 \ M3" unfolding less_eq_multiset_def less_multiset_def using transp_multp[OF transp_on_less, THEN transpD] by blast next fix M N :: "'a multiset" show "M \ N \ N \ M \ M = N" unfolding less_eq_multiset_def less_multiset_def using transp_multp[OF transp_on_less, THEN transpD] using irreflp_multp[OF transp_on_less irreflp_on_less, unfolded irreflp_def, rule_format] by blast qed end lemma mset_le_irrefl [elim!]: fixes M :: "'a::preorder multiset" shows "M < M \ R" by simp lemma wfP_less_multiset[simp]: assumes wfP_less: "wfP ((<) :: ('a :: preorder) \ 'a \ bool)" shows "wfP ((<) :: 'a multiset \ 'a multiset \ bool)" using wfP_multp[OF wfP_less] less_multiset_def by (metis wfPUNIVI wfP_induct) subsubsection \Strict total-order properties\ lemma total_on_mult: assumes "total_on A r" and "trans r" and "\M. M \ B \ set_mset M \ A" shows "total_on B (mult r)" proof (rule total_onI) fix M1 M2 assume "M1 \ B" and "M2 \ B" and "M1 \ M2" let ?I = "M1 \# M2" show "(M1, M2) \ mult r \ (M2, M1) \ mult r" proof (cases "M1 - ?I = {#} \ M2 - ?I = {#}") case True with \M1 \ M2\ show ?thesis by (metis Diff_eq_empty_iff_mset diff_intersect_left_idem diff_intersect_right_idem subset_implies_mult subset_mset.less_le) next case False from assms(1) have "total_on (set_mset (M1 - ?I)) r" by (meson \M1 \ B\ assms(3) diff_subset_eq_self set_mset_mono total_on_subset) with False obtain greatest1 where greatest1_in: "greatest1 \# M1 - ?I" and greatest1_greatest: "\x \# M1 - ?I. greatest1 \ x \ (x, greatest1) \ r" using Multiset.bex_greatest_element[to_set, of "M1 - ?I" r] by (metis assms(2) subset_UNIV trans_on_subset) from assms(1) have "total_on (set_mset (M2 - ?I)) r" by (meson \M2 \ B\ assms(3) diff_subset_eq_self set_mset_mono total_on_subset) with False obtain greatest2 where greatest2_in: "greatest2 \# M2 - ?I" and greatest2_greatest: "\x \# M2 - ?I. greatest2 \ x \ (x, greatest2) \ r" using Multiset.bex_greatest_element[to_set, of "M2 - ?I" r] by (metis assms(2) subset_UNIV trans_on_subset) have "greatest1 \ greatest2" using greatest1_in \greatest2 \# M2 - ?I\ by (metis diff_intersect_left_idem diff_intersect_right_idem dual_order.eq_iff in_diff_count in_diff_countE le_add_same_cancel2 less_irrefl zero_le) hence "(greatest1, greatest2) \ r \ (greatest2, greatest1) \ r" using \total_on A r\[unfolded total_on_def, rule_format, of greatest1 greatest2] \M1 \ B\ \M2 \ B\ greatest1_in greatest2_in assms(3) by (meson in_diffD in_mono) thus ?thesis proof (elim disjE) assume "(greatest1, greatest2) \ r" have "(?I + (M1 - ?I), ?I + (M2 - ?I)) \ mult r" proof (rule one_step_implies_mult[of "M2 - ?I" "M1 - ?I" r ?I]) show "M2 - ?I \ {#}" using False by force next show "\k\#M1 - ?I. \j\#M2 - ?I. (k, j) \ r" using \(greatest1, greatest2) \ r\ greatest2_in greatest1_greatest by (metis assms(2) transD) qed hence "(M1, M2) \ mult r" by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 subset_mset.inf.cobounded2) thus "(M1, M2) \ mult r \ (M2, M1) \ mult r" .. next assume "(greatest2, greatest1) \ r" have "(?I + (M2 - ?I), ?I + (M1 - ?I)) \ mult r" proof (rule one_step_implies_mult[of "M1 - ?I" "M2 - ?I" r ?I]) show "M1 - M1 \# M2 \ {#}" using False by force next show "\k\#M2 - ?I. \j\#M1 - ?I. (k, j) \ r" using \(greatest2, greatest1) \ r\ greatest1_in greatest2_greatest by (metis assms(2) transD) qed hence "(M2, M1) \ mult r" by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 subset_mset.inf.cobounded2) thus "(M1, M2) \ mult r \ (M2, M1) \ mult r" .. qed qed qed lemma total_mult: "total r \ trans r \ total (mult r)" by (rule total_on_mult[of UNIV r UNIV, simplified]) lemma totalp_on_multp: "totalp_on A R \ transp R \ (\M. M \ B \ set_mset M \ A) \ totalp_on B (multp R)" using total_on_mult[of A "{(x,y). R x y}" B, to_pred] by (simp add: multp_def total_on_def totalp_on_def) lemma totalp_multp: "totalp R \ transp R \ totalp (multp R)" by (rule totalp_on_multp[of UNIV R UNIV, simplified]) subsection \Quasi-executable version of the multiset extension\ text \ Predicate variants of \mult\ and the reflexive closure of \mult\, which are executable whenever the given predicate \P\ is. Together with the standard code equations for \(\#\) and \(-\) this should yield quadratic (with respect to calls to \P\) implementations of \multp_code\ and \multeqp_code\. \ definition multp_code :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp_code P N M = (let Z = M \# N; X = M - Z in X \ {#} \ (let Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x)))" definition multeqp_code :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multeqp_code P N M = (let Z = M \# N; X = M - Z; Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x))" lemma multp_code_iff_mult: assumes "irrefl_on (set_mset N \ set_mset M) R" and "trans R" and [simp]: "\x y. P x y \ (x, y) \ R" shows "multp_code P N M \ (N, M) \ mult R" (is "?L \ ?R") proof - have *: "M \# N + (N - M \# N) = N" "M \# N + (M - M \# N) = M" "(M - M \# N) \# (N - M \# N) = {#}" by (auto simp flip: count_inject) show ?thesis proof assume ?L thus ?R using one_step_implies_mult[of "M - M \# N" "N - M \# N" R "M \# N"] * by (auto simp: multp_code_def Let_def) next { fix I J K :: "'a multiset" assume "(I + J) \# (I + K) = {#}" then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty) } note [dest!] = this assume ?R thus ?L using mult_cancel_max using mult_implies_one_step[OF assms(2), of "N - M \# N" "M - M \# N"] mult_cancel_max[OF assms(2,1)] * by (auto simp: multp_code_def) qed qed lemma multp_code_iff_multp: "irreflp_on (set_mset M \ set_mset N) R \ transp R \ multp_code R M N \ multp R M N" using multp_code_iff_mult[simplified, to_pred, of M N R R] by simp lemma multp_code_eq_multp: assumes "irreflp R" and "transp R" shows "multp_code R = multp R" proof (intro ext) fix M N show "multp_code R M N = multp R M N" proof (rule multp_code_iff_multp) from assms show "irreflp_on (set_mset M \ set_mset N) R" by (auto intro: irreflp_on_subset) next from assms show "transp R" by simp qed qed lemma multeqp_code_iff_reflcl_mult: assumes "irrefl_on (set_mset N \ set_mset M) R" and "trans R" and "\x y. P x y \ (x, y) \ R" shows "multeqp_code P N M \ (N, M) \ (mult R)\<^sup>=" proof - { assume "N \ M" "M - M \# N = {#}" then obtain y where "count N y \ count M y" by (auto simp flip: count_inject) then have "\y. count M y < count N y" using \M - M \# N = {#}\ by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y]) } then have "multeqp_code P N M \ multp_code P N M \ N = M" by (auto simp: multeqp_code_def multp_code_def Let_def in_diff_count) thus ?thesis using multp_code_iff_mult[OF assms] by simp qed lemma multeqp_code_iff_reflclp_multp: "irreflp_on (set_mset M \ set_mset N) R \ transp R \ multeqp_code R M N \ (multp R)\<^sup>=\<^sup>= M N" using multeqp_code_iff_reflcl_mult[simplified, to_pred, of M N R R] by simp lemma multeqp_code_eq_reflclp_multp: assumes "irreflp R" and "transp R" shows "multeqp_code R = (multp R)\<^sup>=\<^sup>=" proof (intro ext) fix M N show "multeqp_code R M N \ (multp R)\<^sup>=\<^sup>= M N" proof (rule multeqp_code_iff_reflclp_multp) from assms show "irreflp_on (set_mset M \ set_mset N) R" by (auto intro: irreflp_on_subset) next from assms show "transp R" by simp qed qed subsubsection \Monotonicity of multiset union\ lemma mult1_union: "(B, D) \ mult1 r \ (C + B, C + D) \ mult1 r" by (force simp: mult1_def) lemma union_le_mono2: "B < D \ C + B < C + (D::'a::preorder multiset)" apply (unfold less_multiset_def multp_def mult_def) apply (erule trancl_induct) apply (blast intro: mult1_union) apply (blast intro: mult1_union trancl_trans) done lemma union_le_mono1: "B < D \ B + C < D + (C::'a::preorder multiset)" apply (subst add.commute [of B C]) apply (subst add.commute [of D C]) apply (erule union_le_mono2) done lemma union_less_mono: fixes A B C D :: "'a::preorder multiset" shows "A < C \ B < D \ A + B < C + D" by (blast intro!: union_le_mono1 union_le_mono2 less_trans) instantiation multiset :: (preorder) ordered_ab_semigroup_add begin instance by standard (auto simp add: less_eq_multiset_def intro: union_le_mono2) end subsubsection \Termination proofs with multiset orders\ lemma multi_member_skip: "x \# XS \ x \# {# y #} + XS" and multi_member_this: "x \# {# x #} + XS" and multi_member_last: "x \# {# x #}" by auto definition "ms_strict = mult pair_less" definition "ms_weak = ms_strict \ Id" lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def by (auto intro: wf_mult1 wf_trancl simp: mult_def) lemma smsI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z + B) \ ms_strict" unfolding ms_strict_def by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) lemma wmsI: "(set_mset A, set_mset B) \ max_strict \ A = {#} \ B = {#} \ (Z + A, Z + B) \ ms_weak" unfolding ms_weak_def ms_strict_def by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) inductive pw_leq where pw_leq_empty: "pw_leq {#} {#}" | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" lemma pw_leq_lstep: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" by (drule pw_leq_step) (rule pw_leq_empty, simp) lemma pw_leq_split: assumes "pw_leq X Y" shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" using assms proof induct case pw_leq_empty thus ?case by auto next case (pw_leq_step x y X Y) then obtain A B Z where [simp]: "X = A + Z" "Y = B + Z" and 1[simp]: "(set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#})" by auto from pw_leq_step consider "x = y" | "(x, y) \ pair_less" unfolding pair_leq_def by auto thus ?case proof cases case [simp]: 1 have "{#x#} + X = A + ({#y#}+Z) \ {#y#} + Y = B + ({#y#}+Z) \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" by auto thus ?thesis by blast next case 2 let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" have "{#x#} + X = ?A' + Z" "{#y#} + Y = ?B' + Z" by auto moreover have "(set_mset ?A', set_mset ?B') \ max_strict" using 1 2 unfolding max_strict_def by (auto elim!: max_ext.cases) ultimately show ?thesis by blast qed qed lemma assumes pwleq: "pw_leq Z Z'" shows ms_strictI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" and ms_weakI1: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" proof - from pw_leq_split[OF pwleq] obtain A' B' Z'' where [simp]: "Z = A' + Z''" "Z' = B' + Z''" and mx_or_empty: "(set_mset A', set_mset B') \ max_strict \ (A' = {#} \ B' = {#})" by blast { assume max: "(set_mset A, set_mset B) \ max_strict" from mx_or_empty have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" proof assume max': "(set_mset A', set_mset B') \ max_strict" with max have "(set_mset (A + A'), set_mset (B + B')) \ max_strict" by (auto simp: max_strict_def intro: max_ext_additive) thus ?thesis by (rule smsI) next assume [simp]: "A' = {#} \ B' = {#}" show ?thesis by (rule smsI) (auto intro: max) qed thus "(Z + A, Z' + B) \ ms_strict" by (simp add: ac_simps) thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) } from mx_or_empty have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add: ac_simps) qed lemma empty_neutral: "{#} + x = x" "x + {#} = x" and nonempty_plus: "{# x #} + rs \ {#}" and nonempty_single: "{# x #} \ {#}" by auto setup \ let fun msetT T = \<^Type>\multiset T\; fun mk_mset T [] = \<^instantiate>\'a = T in term \{#}\\ | mk_mset T [x] = \<^instantiate>\'a = T and x in term \{#x#}\\ | mk_mset T (x :: xs) = \<^Const>\plus \msetT T\ for \mk_mset T [x]\ \mk_mset T xs\\ fun mset_member_tac ctxt m i = if m <= 0 then resolve_tac ctxt @{thms multi_member_this} i ORELSE resolve_tac ctxt @{thms multi_member_last} i else resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i fun mset_nonempty_tac ctxt = resolve_tac ctxt @{thms nonempty_plus} ORELSE' resolve_tac ctxt @{thms nonempty_single} fun regroup_munion_conv ctxt = Function_Lib.regroup_conv ctxt \<^const_abbrev>\empty_mset\ \<^const_name>\plus\ (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral})) fun unfold_pwleq_tac ctxt i = (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st)) ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i) ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i) val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union}, @{thm Un_insert_left}, @{thm Un_empty_left}] in ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset { msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, reduction_pair = @{thm ms_reduction_pair} }) end \ subsection \Legacy theorem bindings\ lemmas multi_count_eq = multiset_eq_iff [symmetric] lemma union_commute: "M + N = N + (M::'a multiset)" by (fact add.commute) lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" by (fact add.assoc) lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" by (fact add.left_commute) lemmas union_ac = union_assoc union_commute union_lcomm add_mset_commute lemma union_right_cancel: "M + K = N + K \ M = (N::'a multiset)" by (fact add_right_cancel) lemma union_left_cancel: "K + M = K + N \ M = (N::'a multiset)" by (fact add_left_cancel) lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" by (fact add_left_imp_eq) lemma mset_subset_trans: "(M::'a multiset) \# K \ K \# N \ M \# N" by (fact subset_mset.less_trans) lemma multiset_inter_commute: "A \# B = B \# A" by (fact subset_mset.inf.commute) lemma multiset_inter_assoc: "A \# (B \# C) = A \# B \# C" by (fact subset_mset.inf.assoc [symmetric]) lemma multiset_inter_left_commute: "A \# (B \# C) = B \# (A \# C)" by (fact subset_mset.inf.left_commute) lemmas multiset_inter_ac = multiset_inter_commute multiset_inter_assoc multiset_inter_left_commute lemma mset_le_not_refl: "\ M < (M::'a::preorder multiset)" by (fact less_irrefl) lemma mset_le_trans: "K < M \ M < N \ K < (N::'a::preorder multiset)" by (fact less_trans) lemma mset_le_not_sym: "M < N \ \ N < (M::'a::preorder multiset)" by (fact less_not_sym) lemma mset_le_asym: "M < N \ (\ P \ N < (M::'a::preorder multiset)) \ P" by (fact less_asym) declaration \ let fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = let val (maybe_opt, ps) = Nitpick_Model.dest_plain_fun t' ||> (~~) ||> map (apsnd (snd o HOLogic.dest_number)) fun elems_for t = (case AList.lookup (=) ps t of SOME n => replicate n t | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) in (case maps elems_for (all_values elem_T) @ (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of [] => \<^Const>\Groups.zero T\ | ts => foldl1 (fn (s, t) => \<^Const>\add_mset elem_T for s t\) ts) end | multiset_postproc _ _ _ _ t = t in Nitpick_Model.register_term_postprocessor \<^typ>\'a multiset\ multiset_postproc end \ subsection \Naive implementation using lists\ code_datatype mset lemma [code]: "{#} = mset []" by simp lemma [code]: "add_mset x (mset xs) = mset (x # xs)" by simp lemma [code]: "Multiset.is_empty (mset xs) \ List.null xs" by (simp add: Multiset.is_empty_def List.null_def) lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)" by simp lemma [code]: "image_mset f (mset xs) = mset (map f xs)" by simp lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" by simp lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add) lemma [code]: "mset xs \# mset ys = mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))" proof - have "\zs. mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (auto simp add: inter_add_right1 inter_add_right2 ac_simps) then show ?thesis by simp qed lemma [code]: "mset xs \# mset ys = mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" proof - have "\zs. mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff) then show ?thesis by simp qed declare in_multiset_in_set [code_unfold] lemma [code]: "count (mset xs) x = fold (\y. if x = y then Suc else id) xs 0" proof - have "\n. fold (\y. if x = y then Suc else id) xs n = count (mset xs) x + n" by (induct xs) simp_all then show ?thesis by simp qed declare set_mset_mset [code] declare sorted_list_of_multiset_mset [code] lemma [code]: \ \not very efficient, but representation-ignorant!\ "mset_set A = mset (sorted_list_of_set A)" apply (cases "finite A") apply simp_all apply (induct A rule: finite_induct) apply simp_all done declare size_mset [code] fun subset_eq_mset_impl :: "'a list \ 'a list \ bool option" where "subset_eq_mset_impl [] ys = Some (ys \ [])" | "subset_eq_mset_impl (Cons x xs) ys = (case List.extract ((=) x) ys of None \ None | Some (ys1,_,ys2) \ subset_eq_mset_impl xs (ys1 @ ys2))" lemma subset_eq_mset_impl: "(subset_eq_mset_impl xs ys = None \ \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some True \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some False \ mset xs = mset ys)" proof (induct xs arbitrary: ys) case (Nil ys) show ?case by (auto simp: subset_mset.zero_less_iff_neq_zero) next case (Cons x xs ys) show ?case proof (cases "List.extract ((=) x) ys") case None hence x: "x \ set ys" by (simp add: extract_None_iff) { assume "mset (x # xs) \# mset ys" from set_mset_mono[OF this] x have False by simp } note nle = this moreover { assume "mset (x # xs) \# mset ys" hence "mset (x # xs) \# mset ys" by auto from nle[OF this] have False . } ultimately show ?thesis using None by auto next case (Some res) obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto) note Some = Some[unfolded res] from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp hence id: "mset ys = add_mset x (mset (ys1 @ ys2))" by auto show ?thesis unfolding subset_eq_mset_impl.simps unfolding Some option.simps split unfolding id using Cons[of "ys1 @ ys2"] unfolding subset_mset_def subseteq_mset_def by auto qed qed lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys \ None" using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys = Some True" using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) instantiation multiset :: (equal) equal begin definition [code del]: "HOL.equal A (B :: 'a multiset) \ A = B" lemma [code]: "HOL.equal (mset xs) (mset ys) \ subset_eq_mset_impl xs ys = Some False" unfolding equal_multiset_def using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) instance by standard (simp add: equal_multiset_def) end declare sum_mset_sum_list [code] lemma [code]: "prod_mset (mset xs) = fold times xs 1" proof - have "\x. fold times xs x = prod_mset (mset xs) * x" by (induct xs) (simp_all add: ac_simps) then show ?thesis by simp qed text \ Exercise for the casual reader: add implementations for \<^term>\(\)\ and \<^term>\(<)\ (multiset order). \ text \Quickcheck generators\ context includes term_syntax begin definition msetify :: "'a::typerep list \ (unit \ Code_Evaluation.term) \ 'a multiset \ (unit \ Code_Evaluation.term)" where [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\} xs" end instantiation multiset :: (random) random begin context includes state_combinator_syntax begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\xs. Pair (msetify xs))" instance .. end end instantiation multiset :: (full_exhaustive) full_exhaustive begin definition full_exhaustive_multiset :: "('a multiset \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\xs. f (msetify xs)) i" instance .. end hide_const (open) msetify subsection \BNF setup\ definition rel_mset where "rel_mset R X Y \ (\xs ys. mset xs = X \ mset ys = Y \ list_all2 R xs ys)" lemma mset_zip_take_Cons_drop_twice: assumes "length xs = length ys" "j \ length xs" shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = add_mset (x,y) (mset (zip xs ys))" using assms proof (induct xs ys arbitrary: x y j rule: list_induct2) case Nil thus ?case by simp next case (Cons x xs y ys) thus ?case proof (cases "j = 0") case True thus ?thesis by simp next case False then obtain k where k: "j = Suc k" by (cases j) simp hence "k \ length xs" using Cons.prems by auto hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = add_mset (x,y) (mset (zip xs ys))" by (rule Cons.hyps(2)) thus ?thesis unfolding k by auto qed qed lemma ex_mset_zip_left: assumes "length xs = length ys" "mset xs' = mset xs" shows "\ys'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" using assms proof (induct xs ys arbitrary: xs' rule: list_induct2) case Nil thus ?case by auto next case (Cons x xs y ys xs') obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) define xsa where "xsa = take j xs' @ drop (Suc j) xs'" have "mset xs' = {#x#} + mset xsa" unfolding xsa_def using j_len nth_j by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' append_take_drop_id mset.simps(2) mset_append) hence ms_x: "mset xsa = mset xs" by (simp add: Cons.prems) then obtain ysa where len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast define ys' where "ys' = take j ysa @ y # drop j ysa" have xs': "xs' = take j xsa @ x # drop j xsa" using ms_x j_len nth_j Cons.prems xsa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons length_drop size_mset) have j_len': "j \ length xsa" using j_len xs' xsa_def by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) have "length ys' = length xs'" unfolding ys'_def using Cons.prems len_a ms_x by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" unfolding xs' ys'_def by (rule trans[OF mset_zip_take_Cons_drop_twice]) (auto simp: len_a ms_a j_len') ultimately show ?case by blast qed lemma list_all2_reorder_left_invariance: assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" shows "\ys'. list_all2 R xs' ys' \ mset ys' = mset ys" proof - have len: "length xs = length ys" using rel list_all2_conv_all_nth by auto obtain ys' where len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" using len ms_x by (metis ex_mset_zip_left) have "list_all2 R xs' ys'" using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) moreover have "mset ys' = mset ys" using len len' ms_xy map_snd_zip mset_map by metis ultimately show ?thesis by blast qed lemma ex_mset: "\xs. mset xs = X" by (induct X) (simp, metis mset.simps(2)) inductive pred_mset :: "('a \ bool) \ 'a multiset \ bool" where "pred_mset P {#}" | "\P a; pred_mset P M\ \ pred_mset P (add_mset a M)" lemma pred_mset_iff: \ \TODO: alias for \<^const>\Multiset.Ball\\ \pred_mset P M \ Multiset.Ball M P\ (is \?P \ ?Q\) proof assume ?P then show ?Q by induction simp_all next assume ?Q then show ?P by (induction M) (auto intro: pred_mset.intros) qed bnf "'a multiset" map: image_mset sets: set_mset bd: natLeq wits: "{#}" rel: rel_mset pred: pred_mset proof - show "image_mset id = id" by (rule image_mset.id) show "image_mset (g \ f) = image_mset g \ image_mset f" for f g unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) show "(\z. z \ set_mset X \ f z = g z) \ image_mset f X = image_mset g X" for f g X by (induct X) simp_all show "set_mset \ image_mset f = (`) f \ set_mset" for f by auto show "card_order natLeq" by (rule natLeq_card_order) show "BNF_Cardinal_Arithmetic.cinfinite natLeq" by (rule natLeq_cinfinite) show "regularCard natLeq" by (rule regularCard_natLeq) show "ordLess2 (card_of (set_mset X)) natLeq" for X by transfer (auto simp: finite_iff_ordLess_natLeq[symmetric]) show "rel_mset R OO rel_mset S \ rel_mset (R OO S)" for R S unfolding rel_mset_def[abs_def] OO_def apply clarify subgoal for X Z Y xs ys' ys zs apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) apply (auto intro: list_all2_trans) done done show "rel_mset R = (\x y. \z. set_mset z \ {(x, y). R x y} \ image_mset fst z = x \ image_mset snd z = y)" for R unfolding rel_mset_def[abs_def] apply (rule ext)+ apply safe apply (rule_tac x = "mset (zip xs ys)" in exI; auto simp: in_set_zip list_all2_iff simp flip: mset_map) apply (rename_tac XY) apply (cut_tac X = XY in ex_mset) apply (erule exE) apply (rename_tac xys) apply (rule_tac x = "map fst xys" in exI) apply (auto simp: mset_map) apply (rule_tac x = "map snd xys" in exI) apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) done show "z \ set_mset {#} \ False" for z by auto show "pred_mset P = (\x. Ball (set_mset x) P)" for P by (simp add: fun_eq_iff pred_mset_iff) qed inductive rel_mset' :: \('a \ 'b \ bool) \ 'a multiset \ 'b multiset \ bool\ where Zero[intro]: "rel_mset' R {#} {#}" | Plus[intro]: "\R a b; rel_mset' R M N\ \ rel_mset' R (add_mset a M) (add_mset b N)" lemma rel_mset_Zero: "rel_mset R {#} {#}" unfolding rel_mset_def Grp_def by auto declare multiset.count[simp] declare count_Abs_multiset[simp] declare multiset.count_inverse[simp] lemma rel_mset_Plus: assumes ab: "R a b" and MN: "rel_mset R M N" shows "rel_mset R (add_mset a M) (add_mset b N)" proof - have "\ya. add_mset a (image_mset fst y) = image_mset fst ya \ add_mset b (image_mset snd y) = image_mset snd ya \ set_mset ya \ {(x, y). R x y}" if "R a b" and "set_mset y \ {(x, y). R x y}" for y using that by (intro exI[of _ "add_mset (a,b) y"]) auto thus ?thesis using assms unfolding multiset.rel_compp_Grp Grp_def by blast qed lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \ rel_mset R M N" by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) lemma rel_mset_size: "rel_mset R M N \ size M = size N" unfolding multiset.rel_compp_Grp Grp_def by auto lemma rel_mset_Zero_iff [simp]: shows "rel_mset rel {#} Y \ Y = {#}" and "rel_mset rel X {#} \ X = {#}" by (auto simp add: rel_mset_Zero dest: rel_mset_size) lemma multiset_induct2[case_names empty addL addR]: assumes empty: "P {#} {#}" and addL: "\a M N. P M N \ P (add_mset a M) N" and addR: "\a M N. P M N \ P M (add_mset a N)" shows "P M N" apply(induct N rule: multiset_induct) apply(induct M rule: multiset_induct, rule empty, erule addL) apply(induct M rule: multiset_induct, erule addR, erule addR) done lemma multiset_induct2_size[consumes 1, case_names empty add]: assumes c: "size M = size N" and empty: "P {#} {#}" and add: "\a b M N a b. P M N \ P (add_mset a M) (add_mset b N)" shows "P M N" using c proof (induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) show ?case proof(cases "M = {#}") case True hence "N = {#}" using less.prems by auto thus ?thesis using True empty by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) have "N \ {#}" using False less.prems by auto then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split) have "size M1 = size N1" using less.prems unfolding M N by auto thus ?thesis using M N less.hyps add by auto qed qed lemma msed_map_invL: assumes "image_mset f (add_mset a M) = N" shows "\N1. N = add_mset (f a) N1 \ image_mset f M = N1" proof - have "f a \# N" using assms multiset.set_map[of f "add_mset a M"] by auto then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis have "image_mset f M = N1" using assms unfolding N by simp thus ?thesis using N by blast qed lemma msed_map_invR: assumes "image_mset f M = add_mset b N" shows "\M1 a. M = add_mset a M1 \ f a = b \ image_mset f M1 = N" proof - obtain a where a: "a \# M" and fa: "f a = b" using multiset.set_map[of f M] unfolding assms by (metis image_iff union_single_eq_member) then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp thus ?thesis using M fa by blast qed lemma msed_rel_invL: assumes "rel_mset R (add_mset a M) N" shows "\N1 b. N = add_mset b N1 \ R a b \ rel_mset R M N1" proof - obtain K where KM: "image_mset fst K = add_mset a M" and KN: "image_mset snd K = N" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a" and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1" using msed_map_invL[OF KN[unfolded K]] by auto have Rab: "R a (snd ab)" using sK a unfolding K by auto have "rel_mset R M N1" using sK K1M K1N1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using N Rab by auto qed lemma msed_rel_invR: assumes "rel_mset R M (add_mset b N)" shows "\M1 a. M = add_mset a M1 \ R a b \ rel_mset R M1 N" proof - obtain K where KN: "image_mset snd K = add_mset b N" and KM: "image_mset fst K = M" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b" and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1" using msed_map_invL[OF KM[unfolded K]] by auto have Rab: "R (fst ab) b" using sK b unfolding K by auto have "rel_mset R M1 N" using sK K1N K1M1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using M Rab by auto qed lemma rel_mset_imp_rel_mset': assumes "rel_mset R M N" shows "rel_mset' R M N" using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) have c: "size M = size N" using rel_mset_size[OF less.prems] . show ?case proof(cases "M = {#}") case True hence "N = {#}" using c by simp thus ?thesis using True rel_mset'.Zero by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1" using msed_rel_invL[OF less.prems[unfolded M]] by auto have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp qed qed lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto text \The main end product for \<^const>\rel_mset\: inductive characterization:\ lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]] subsection \Size setup\ lemma size_multiset_o_map: "size_multiset g \ image_mset f = size_multiset (g \ f)" apply (rule ext) subgoal for x by (induct x) auto done setup \ BNF_LFP_Size.register_size_global \<^type_name>\multiset\ \<^const_name>\size_multiset\ @{thm size_multiset_overloaded_def} @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single size_union} @{thms size_multiset_o_map} \ subsection \Lemmas about Size\ lemma size_mset_SucE: "size A = Suc n \ (\a B. A = {#a#} + B \ size B = n \ P) \ P" by (cases A) (auto simp add: ac_simps) lemma size_Suc_Diff1: "x \# M \ Suc (size (M - {#x#})) = size M" using arg_cong[OF insert_DiffM, of _ _ size] by simp lemma size_Diff_singleton: "x \# M \ size (M - {#x#}) = size M - 1" by (simp flip: size_Suc_Diff1) lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \# A then size A - 1 else size A)" by (simp add: diff_single_trivial size_Diff_singleton) lemma size_Un_Int: "size A + size B = size (A \# B) + size (A \# B)" by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup) lemma size_Un_disjoint: "A \# B = {#} \ size (A \# B) = size A + size B" using size_Un_Int[of A B] by simp lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \# M')" by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1) lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \ size (M - M')" by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono) lemma size_Diff1_less: "x\# M \ size (M - {#x#}) < size M" by (rule Suc_less_SucD) (simp add: size_Suc_Diff1) lemma size_Diff2_less: "x\# M \ y\# M \ size (M - {#x#} - {#y#}) < size M" by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int) lemma size_Diff1_le: "size (M - {#x#}) \ size M" by (cases "x \# M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial) lemma size_psubset: "M \# M' \ size M < size M' \ M \# M'" using less_irrefl subset_mset_def by blast lifting_update multiset.lifting lifting_forget multiset.lifting hide_const (open) wcount end diff --git a/src/HOL/Library/Multiset_Order.thy b/src/HOL/Library/Multiset_Order.thy --- a/src/HOL/Library/Multiset_Order.thy +++ b/src/HOL/Library/Multiset_Order.thy @@ -1,856 +1,856 @@ (* Title: HOL/Library/Multiset_Order.thy Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Martin Desharnais, MPI-INF Saarbruecken *) section \More Theorems about the Multiset Order\ theory Multiset_Order imports Multiset begin subsection \Alternative Characterizations\ subsubsection \The Dershowitz--Manna Ordering\ definition multp\<^sub>D\<^sub>M where "multp\<^sub>D\<^sub>M r M N \ (\X Y. X \ {#} \ X \# N \ M = (N - X) + Y \ (\k. k \# Y \ (\a. a \# X \ r k a)))" lemma multp\<^sub>D\<^sub>M_imp_multp: "multp\<^sub>D\<^sub>M r M N \ multp r M N" proof - assume "multp\<^sub>D\<^sub>M r M N" then obtain X Y where "X \ {#}" and "X \# N" and "M = N - X + Y" and "\k. k \# Y \ (\a. a \# X \ r k a)" unfolding multp\<^sub>D\<^sub>M_def by blast then have "multp r (N - X + Y) (N - X + X)" by (intro one_step_implies_multp) (auto simp: Bex_def trans_def) with \M = N - X + Y\ \X \# N\ show "multp r M N" by (metis subset_mset.diff_add) qed subsubsection \The Huet--Oppen Ordering\ definition multp\<^sub>H\<^sub>O where "multp\<^sub>H\<^sub>O r M N \ M \ N \ (\y. count N y < count M y \ (\x. r y x \ count M x < count N x))" lemma multp_imp_multp\<^sub>H\<^sub>O: assumes "asymp r" and "transp r" shows "multp r M N \ multp\<^sub>H\<^sub>O r M N" unfolding multp_def mult_def proof (induction rule: trancl_induct) case (base P) then show ?case using \asymp r\ by (auto elim!: mult1_lessE simp: count_eq_zero_iff multp\<^sub>H\<^sub>O_def split: if_splits dest!: Suc_lessD) next case (step N P) from step(3) have "M \ N" and **: "\y. count N y < count M y \ (\x. r y x \ count M x < count N x)" by (simp_all add: multp\<^sub>H\<^sub>O_def) from step(2) obtain M0 a K where *: "P = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ r b a" using \asymp r\ by (auto elim: mult1_lessE) from \M \ N\ ** *(1,2,3) have "M \ P" using *(4) \asymp r\ by (metis asympD add_cancel_right_right add_diff_cancel_left' add_mset_add_single count_inI count_union diff_diff_add_mset diff_single_trivial in_diff_count multi_member_last) moreover { assume "count P a \ count M a" with \a \# K\ have "count N a < count M a" unfolding *(1,2) by (auto simp add: not_in_iff) with ** obtain z where z: "r a z" "count M z < count N z" by blast with * have "count N z \ count P z" using \asymp r\ by (metis add_diff_cancel_left' add_mset_add_single asympD diff_diff_add_mset diff_single_trivial in_diff_count not_le_imp_less) with z have "\z. r a z \ count M z < count P z" by auto } note count_a = this { fix y assume count_y: "count P y < count M y" have "\x. r y x \ count M x < count P x" proof (cases "y = a") case True with count_y count_a show ?thesis by auto next case False show ?thesis proof (cases "y \# K") case True with *(4) have "r y a" by simp then show ?thesis by (cases "count P a \ count M a") (auto dest: count_a intro: \transp r\[THEN transpD]) next case False with \y \ a\ have "count P y = count N y" unfolding *(1,2) by (simp add: not_in_iff) with count_y ** obtain z where z: "r y z" "count M z < count N z" by auto show ?thesis proof (cases "z \# K") case True with *(4) have "r z a" by simp with z(1) show ?thesis by (cases "count P a \ count M a") (auto dest!: count_a intro: \transp r\[THEN transpD]) next case False with \a \# K\ have "count N z \ count P z" unfolding * by (auto simp add: not_in_iff) with z show ?thesis by auto qed qed qed } ultimately show ?case unfolding multp\<^sub>H\<^sub>O_def by blast qed lemma multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M: "multp\<^sub>H\<^sub>O r M N \ multp\<^sub>D\<^sub>M r M N" unfolding multp\<^sub>D\<^sub>M_def proof (intro iffI exI conjI) assume "multp\<^sub>H\<^sub>O r M N" then obtain z where z: "count M z < count N z" unfolding multp\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff) define X where "X = N - M" define Y where "Y = M - N" from z show "X \ {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq) from z show "X \# N" unfolding X_def by auto show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force show "\k. k \# Y \ (\a. a \# X \ r k a)" proof (intro allI impI) fix k assume "k \# Y" then have "count N k < count M k" unfolding Y_def by (auto simp add: in_diff_count) with \multp\<^sub>H\<^sub>O r M N\ obtain a where "r k a" and "count M a < count N a" unfolding multp\<^sub>H\<^sub>O_def by blast then show "\a. a \# X \ r k a" unfolding X_def by (auto simp add: in_diff_count) qed qed lemma multp_eq_multp\<^sub>D\<^sub>M: "asymp r \ transp r \ multp r = multp\<^sub>D\<^sub>M r" using multp\<^sub>D\<^sub>M_imp_multp multp_imp_multp\<^sub>H\<^sub>O[THEN multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M] by blast lemma multp_eq_multp\<^sub>H\<^sub>O: "asymp r \ transp r \ multp r = multp\<^sub>H\<^sub>O r" using multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M[THEN multp\<^sub>D\<^sub>M_imp_multp] multp_imp_multp\<^sub>H\<^sub>O by blast lemma multp\<^sub>D\<^sub>M_plus_plusI[simp]: assumes "multp\<^sub>D\<^sub>M R M1 M2" shows "multp\<^sub>D\<^sub>M R (M + M1) (M + M2)" proof - from assms obtain X Y where "X \ {#}" and "X \# M2" and "M1 = M2 - X + Y" and "\k. k \# Y \ (\a. a \# X \ R k a)" unfolding multp\<^sub>D\<^sub>M_def by auto show "multp\<^sub>D\<^sub>M R (M + M1) (M + M2)" unfolding multp\<^sub>D\<^sub>M_def proof (intro exI conjI) show "X \ {#}" using \X \ {#}\ by simp next show "X \# M + M2" using \X \# M2\ by (simp add: subset_mset.add_increasing) next show "M + M1 = M + M2 - X + Y" using \X \# M2\ \M1 = M2 - X + Y\ by (metis multiset_diff_union_assoc union_assoc) next show "\k. k \# Y \ (\a. a \# X \ R k a)" using \\k. k \# Y \ (\a. a \# X \ R k a)\ by simp qed qed lemma multp\<^sub>H\<^sub>O_plus_plus[simp]: "multp\<^sub>H\<^sub>O R (M + M1) (M + M2) \ multp\<^sub>H\<^sub>O R M1 M2" unfolding multp\<^sub>H\<^sub>O_def by simp lemma strict_subset_implies_multp\<^sub>D\<^sub>M: "A \# B \ multp\<^sub>D\<^sub>M r A B" unfolding multp\<^sub>D\<^sub>M_def by (metis add.right_neutral add_diff_cancel_right' empty_iff mset_subset_eq_add_right set_mset_empty subset_mset.lessE) lemma strict_subset_implies_multp\<^sub>H\<^sub>O: "A \# B \ multp\<^sub>H\<^sub>O r A B" unfolding multp\<^sub>H\<^sub>O_def by (simp add: leD mset_subset_eq_count) lemma multp\<^sub>H\<^sub>O_implies_one_step_strong: assumes "multp\<^sub>H\<^sub>O R A B" defines "J \ B - A" and "K \ A - B" shows "J \ {#}" and "\k \# K. \x \# J. R k x" proof - show "J \ {#}" using \multp\<^sub>H\<^sub>O R A B\ by (metis Diff_eq_empty_iff_mset J_def add.right_neutral multp\<^sub>D\<^sub>M_def multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M multp\<^sub>H\<^sub>O_plus_plus subset_mset.add_diff_inverse subset_mset.le_zero_eq) show "\k\#K. \x\#J. R k x" using \multp\<^sub>H\<^sub>O R A B\ by (metis J_def K_def in_diff_count multp\<^sub>H\<^sub>O_def) qed lemma multp\<^sub>H\<^sub>O_minus_inter_minus_inter_iff: fixes M1 M2 :: "_ multiset" shows "multp\<^sub>H\<^sub>O R (M1 - M2) (M2 - M1) \ multp\<^sub>H\<^sub>O R M1 M2" by (metis diff_intersect_left_idem multiset_inter_commute multp\<^sub>H\<^sub>O_plus_plus subset_mset.add_diff_inverse subset_mset.inf.cobounded1) lemma multp\<^sub>H\<^sub>O_iff_set_mset_less\<^sub>H\<^sub>O_set_mset: "multp\<^sub>H\<^sub>O R M1 M2 \ (set_mset (M1 - M2) \ set_mset (M2 - M1) \ (\y \# M1 - M2. (\x \# M2 - M1. R y x)))" unfolding multp\<^sub>H\<^sub>O_minus_inter_minus_inter_iff[of R M1 M2, symmetric] unfolding multp\<^sub>H\<^sub>O_def unfolding count_minus_inter_lt_count_minus_inter_iff unfolding minus_inter_eq_minus_inter_iff by auto subsubsection \Monotonicity\ lemma multp\<^sub>D\<^sub>M_mono_strong: "multp\<^sub>D\<^sub>M R M1 M2 \ (\x y. x \# M1 \ y \# M2 \ R x y \ S x y) \ multp\<^sub>D\<^sub>M S M1 M2" unfolding multp\<^sub>D\<^sub>M_def by (metis add_diff_cancel_left' in_diffD subset_mset.diff_add) lemma multp\<^sub>H\<^sub>O_mono_strong: "multp\<^sub>H\<^sub>O R M1 M2 \ (\x y. x \# M1 \ y \# M2 \ R x y \ S x y) \ multp\<^sub>H\<^sub>O S M1 M2" unfolding multp\<^sub>H\<^sub>O_def by (metis count_inI less_zeroE) subsubsection \Properties of Orders\ paragraph \Asymmetry\ text \The following lemma is a negative result stating that asymmetry of an arbitrary binary relation cannot be simply lifted to @{const multp\<^sub>H\<^sub>O}. It suffices to have four distinct values to build a counterexample.\ lemma asymp_not_liftable_to_multp\<^sub>H\<^sub>O: fixes a b c d :: 'a assumes "distinct [a, b, c, d]" shows "\ (\(R :: 'a \ 'a \ bool). asymp R \ asymp (multp\<^sub>H\<^sub>O R))" proof - define R :: "'a \ 'a \ bool" where "R = (\x y. x = a \ y = c \ x = b \ y = d \ x = c \ y = b \ x = d \ y = a)" from assms(1) have "{#a, b#} \ {#c, d#}" by (metis add_mset_add_single distinct.simps(2) list.set(1) list.simps(15) multi_member_this set_mset_add_mset_insert set_mset_single) from assms(1) have "asymp R" by (auto simp: R_def intro: asymp_onI) moreover have "\ asymp (multp\<^sub>H\<^sub>O R)" unfolding asymp_on_def Set.ball_simps not_all not_imp not_not proof (intro exI conjI) show "multp\<^sub>H\<^sub>O R {#a, b#} {#c, d#}" unfolding multp\<^sub>H\<^sub>O_def using \{#a, b#} \ {#c, d#}\ R_def assms by auto next show "multp\<^sub>H\<^sub>O R {#c, d#} {#a, b#}" unfolding multp\<^sub>H\<^sub>O_def using \{#a, b#} \ {#c, d#}\ R_def assms by auto qed ultimately show ?thesis unfolding not_all not_imp by auto qed text \However, if the binary relation is both asymmetric and transitive, then @{const multp\<^sub>H\<^sub>O} is also asymmetric.\ lemma asymp_on_multp\<^sub>H\<^sub>O: assumes "asymp_on A R" and "transp_on A R" and B_sub_A: "\M. M \ B \ set_mset M \ A" shows "asymp_on B (multp\<^sub>H\<^sub>O R)" proof (rule asymp_onI) fix M1 M2 :: "'a multiset" assume "M1 \ B" "M2 \ B" "multp\<^sub>H\<^sub>O R M1 M2" from \transp_on A R\ B_sub_A have tran: "transp_on (set_mset (M1 - M2)) R" using \M1 \ B\ by (meson in_diffD subset_eq transp_on_subset) from \asymp_on A R\ B_sub_A have asym: "asymp_on (set_mset (M1 - M2)) R" using \M1 \ B\ by (meson in_diffD subset_eq asymp_on_subset) show "\ multp\<^sub>H\<^sub>O R M2 M1" proof (cases "M1 - M2 = {#}") case True then show ?thesis using multp\<^sub>H\<^sub>O_implies_one_step_strong(1) by metis next case False hence "\m\#M1 - M2. \x\#M1 - M2. x \ m \ \ R m x" using Finite_Set.bex_max_element[of "set_mset (M1 - M2)" R, OF finite_set_mset asym tran] by simp with \transp_on A R\ B_sub_A have "\y\#M2 - M1. \x\#M1 - M2. \ R y x" using \multp\<^sub>H\<^sub>O R M1 M2\[THEN multp\<^sub>H\<^sub>O_implies_one_step_strong(2)] using asym[THEN irreflp_on_if_asymp_on, THEN irreflp_onD] by (metis \M1 \ B\ \M2 \ B\ in_diffD subsetD transp_onD) thus ?thesis unfolding multp\<^sub>H\<^sub>O_iff_set_mset_less\<^sub>H\<^sub>O_set_mset by simp qed qed lemma asymp_multp\<^sub>H\<^sub>O: assumes "asymp R" and "transp R" shows "asymp (multp\<^sub>H\<^sub>O R)" using assms asymp_on_multp\<^sub>H\<^sub>O[of UNIV, simplified] by metis paragraph \Irreflexivity\ lemma irreflp_on_multp\<^sub>H\<^sub>O[simp]: "irreflp_on B (multp\<^sub>H\<^sub>O R)" by (simp add: irreflp_onI multp\<^sub>H\<^sub>O_def) paragraph \Transitivity\ lemma transp_on_multp\<^sub>H\<^sub>O: assumes "asymp_on A R" and "transp_on A R" and B_sub_A: "\M. M \ B \ set_mset M \ A" shows "transp_on B (multp\<^sub>H\<^sub>O R)" proof (rule transp_onI) from assms have "asymp_on B (multp\<^sub>H\<^sub>O R)" using asymp_on_multp\<^sub>H\<^sub>O by metis fix M1 M2 M3 assume hyps: "M1 \ B" "M2 \ B" "M3 \ B" "multp\<^sub>H\<^sub>O R M1 M2" "multp\<^sub>H\<^sub>O R M2 M3" from assms have [intro]: "asymp_on (set_mset M1 \ set_mset M2) R" "transp_on (set_mset M1 \ set_mset M2) R" using \M1 \ B\ \M2 \ B\ by (simp_all add: asymp_on_subset transp_on_subset) from assms have "transp_on (set_mset M1) R" by (meson transp_on_subset hyps(1)) from \multp\<^sub>H\<^sub>O R M1 M2\ have "M1 \ M2" and "\y. count M2 y < count M1 y \ (\x. R y x \ count M1 x < count M2 x)" unfolding multp\<^sub>H\<^sub>O_def by simp_all from \multp\<^sub>H\<^sub>O R M2 M3\ have "M2 \ M3" and "\y. count M3 y < count M2 y \ (\x. R y x \ count M2 x < count M3 x)" unfolding multp\<^sub>H\<^sub>O_def by simp_all show "multp\<^sub>H\<^sub>O R M1 M3" proof (rule ccontr) let ?P = "\x. count M3 x < count M1 x \ (\y. R x y \ count M1 y \ count M3 y)" assume "\ multp\<^sub>H\<^sub>O R M1 M3" hence "M1 = M3 \ (\x. ?P x)" unfolding multp\<^sub>H\<^sub>O_def by force thus False proof (elim disjE) assume "M1 = M3" thus False using \asymp_on B (multp\<^sub>H\<^sub>O R)\[THEN asymp_onD] using \M2 \ B\ \M3 \ B\ \multp\<^sub>H\<^sub>O R M1 M2\ \multp\<^sub>H\<^sub>O R M2 M3\ by metis next assume "\x. ?P x" hence "\x \# M1 + M2. ?P x" by (auto simp: count_inI) have "\y \# M1 + M2. ?P y \ (\z \# M1 + M2. R y z \ \ ?P z)" proof (rule Finite_Set.bex_max_element_with_property) show "\x \# M1 + M2. ?P x" using \\x. ?P x\ by (auto simp: count_inI) qed auto then obtain x where "x \# M1 + M2" and "count M3 x < count M1 x" and "\y. R x y \ count M1 y \ count M3 y" and "\y \# M1 + M2. R x y \ count M3 y < count M1 y \ (\z. R y z \ count M1 z < count M3 z)" by force let ?Q = "\x'. R\<^sup>=\<^sup>= x x' \ count M3 x' < count M2 x'" show False proof (cases "\x'. ?Q x'") case True have "\y \# M1 + M2. ?Q y \ (\z \# M1 + M2. R y z \ \ ?Q z)" proof (rule Finite_Set.bex_max_element_with_property) show "\x \# M1 + M2. ?Q x" using \\x. ?Q x\ by (auto simp: count_inI) qed auto then obtain x' where "x' \# M1 + M2" and "R\<^sup>=\<^sup>= x x'" and "count M3 x' < count M2 x'" and maximality_x': "\z \# M1 + M2. R x' z \ \ (R\<^sup>=\<^sup>= x z) \ count M3 z \ count M2 z" by (auto simp: linorder_not_less) with \multp\<^sub>H\<^sub>O R M2 M3\ obtain y' where "R x' y'" and "count M2 y' < count M3 y'" unfolding multp\<^sub>H\<^sub>O_def by auto hence "count M2 y' < count M1 y'" by (smt (verit) \R\<^sup>=\<^sup>= x x'\ \\y. R x y \ count M3 y \ count M1 y\ \count M3 x < count M1 x\ \count M3 x' < count M2 x'\ assms(2) count_inI dual_order.strict_trans1 hyps(1) hyps(2) hyps(3) less_nat_zero_code B_sub_A subsetD sup2E transp_onD) with \multp\<^sub>H\<^sub>O R M1 M2\ obtain y'' where "R y' y''" and "count M1 y'' < count M2 y''" unfolding multp\<^sub>H\<^sub>O_def by auto hence "count M3 y'' < count M2 y''" by (smt (verit, del_insts) \R x' y'\ \R\<^sup>=\<^sup>= x x'\ \\y. R x y \ count M3 y \ count M1 y\ \count M2 y' < count M3 y'\ \count M3 x < count M1 x\ \count M3 x' < count M2 x'\ assms(2) count_greater_zero_iff dual_order.strict_trans1 hyps(1) hyps(2) hyps(3) less_nat_zero_code linorder_not_less B_sub_A subset_iff sup2E transp_onD) moreover have "count M2 y'' \ count M3 y''" proof - have "y'' \# M1 + M2" by (metis \count M1 y'' < count M2 y''\ count_inI not_less_iff_gr_or_eq union_iff) moreover have "R x' y''" by (metis \R x' y'\ \R y' y''\ \count M2 y' < count M1 y'\ \transp_on (set_mset M1 \ set_mset M2) R\ \x' \# M1 + M2\ calculation count_inI nat_neq_iff set_mset_union transp_onD union_iff) moreover have "R\<^sup>=\<^sup>= x y''" using \R\<^sup>=\<^sup>= x x'\ by (metis (mono_tags, opaque_lifting) \transp_on (set_mset M1 \ set_mset M2) R\ \x \# M1 + M2\ \x' \# M1 + M2\ calculation(1) calculation(2) set_mset_union sup2I1 transp_onD transp_on_reflclp) ultimately show ?thesis using maximality_x'[rule_format, of y''] by metis qed ultimately show ?thesis by linarith next case False hence "\x'. R\<^sup>=\<^sup>= x x' \ count M2 x' \ count M3 x'" by auto hence "count M2 x \ count M3 x" by simp hence "count M2 x < count M1 x" using \count M3 x < count M1 x\ by linarith with \multp\<^sub>H\<^sub>O R M1 M2\ obtain y where "R x y" and "count M1 y < count M2 y" unfolding multp\<^sub>H\<^sub>O_def by auto hence "count M3 y < count M2 y" using \\y. R x y \ count M3 y \ count M1 y\ dual_order.strict_trans2 by metis then show ?thesis using False \R x y\ by auto qed qed qed qed lemma transp_multp\<^sub>H\<^sub>O: assumes "asymp R" and "transp R" shows "transp (multp\<^sub>H\<^sub>O R)" using assms transp_on_multp\<^sub>H\<^sub>O[of UNIV, simplified] by metis paragraph \Totality\ lemma totalp_on_multp\<^sub>D\<^sub>M: "totalp_on A R \ (\M. M \ B \ set_mset M \ A) \ totalp_on B (multp\<^sub>D\<^sub>M R)" by (smt (verit, ccfv_SIG) count_inI in_mono multp\<^sub>H\<^sub>O_def multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M not_less_iff_gr_or_eq totalp_onD totalp_onI) lemma totalp_multp\<^sub>D\<^sub>M: "totalp R \ totalp (multp\<^sub>D\<^sub>M R)" by (rule totalp_on_multp\<^sub>D\<^sub>M[of UNIV R UNIV, simplified]) lemma totalp_on_multp\<^sub>H\<^sub>O: "totalp_on A R \ (\M. M \ B \ set_mset M \ A) \ totalp_on B (multp\<^sub>H\<^sub>O R)" by (smt (verit, ccfv_SIG) count_inI in_mono multp\<^sub>H\<^sub>O_def not_less_iff_gr_or_eq totalp_onD totalp_onI) lemma totalp_multp\<^sub>H\<^sub>O: "totalp R \ totalp (multp\<^sub>H\<^sub>O R)" by (rule totalp_on_multp\<^sub>H\<^sub>O[of UNIV R UNIV, simplified]) paragraph \Type Classes\ context preorder begin lemma order_mult: "class.order (\M N. (M, N) \ mult {(x, y). x < y} \ M = N) (\M N. (M, N) \ mult {(x, y). x < y})" (is "class.order ?le ?less") proof - have irrefl: "\M :: 'a multiset. \ ?less M M" proof fix M :: "'a multiset" have "trans {(x'::'a, x). x' < x}" by (rule transI) (blast intro: less_trans) moreover assume "(M, M) \ mult {(x, y). x < y}" ultimately have "\I J K. M = I + J \ M = I + K \ J \ {#} \ (\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by (rule mult_implies_one_step) then obtain I J K where "M = I + J" and "M = I + K" and "J \ {#}" and "(\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by blast then have aux1: "K \ {#}" and aux2: "\k\set_mset K. \j\set_mset K. k < j" by auto have "finite (set_mset K)" by simp moreover note aux2 ultimately have "set_mset K = {}" by (induct rule: finite_induct) (simp, metis (mono_tags) insert_absorb insert_iff insert_not_empty less_irrefl less_trans) with aux1 show False by simp qed have trans: "\K M N :: 'a multiset. ?less K M \ ?less M N \ ?less K N" unfolding mult_def by (blast intro: trancl_trans) show "class.order ?le ?less" by standard (auto simp add: less_eq_multiset_def irrefl dest: trans) qed text \The Dershowitz--Manna ordering:\ definition less_multiset\<^sub>D\<^sub>M where "less_multiset\<^sub>D\<^sub>M M N \ (\X Y. X \ {#} \ X \# N \ M = (N - X) + Y \ (\k. k \# Y \ (\a. a \# X \ k < a)))" text \The Huet--Oppen ordering:\ definition less_multiset\<^sub>H\<^sub>O where "less_multiset\<^sub>H\<^sub>O M N \ M \ N \ (\y. count N y < count M y \ (\x. y < x \ count M x < count N x))" lemma mult_imp_less_multiset\<^sub>H\<^sub>O: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>H\<^sub>O M N" unfolding multp_def[of "(<)", symmetric] using multp_imp_multp\<^sub>H\<^sub>O[of "(<)"] by (simp add: less_multiset\<^sub>H\<^sub>O_def multp\<^sub>H\<^sub>O_def) lemma less_multiset\<^sub>D\<^sub>M_imp_mult: "less_multiset\<^sub>D\<^sub>M M N \ (M, N) \ mult {(x, y). x < y}" unfolding multp_def[of "(<)", symmetric] by (rule multp\<^sub>D\<^sub>M_imp_multp[of "(<)" M N]) (simp add: less_multiset\<^sub>D\<^sub>M_def multp\<^sub>D\<^sub>M_def) lemma less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M: "less_multiset\<^sub>H\<^sub>O M N \ less_multiset\<^sub>D\<^sub>M M N" unfolding less_multiset\<^sub>D\<^sub>M_def less_multiset\<^sub>H\<^sub>O_def unfolding multp\<^sub>D\<^sub>M_def[symmetric] multp\<^sub>H\<^sub>O_def[symmetric] by (rule multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M) lemma mult_less_multiset\<^sub>D\<^sub>M: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>D\<^sub>M M N" unfolding multp_def[of "(<)", symmetric] using multp_eq_multp\<^sub>D\<^sub>M[of "(<)", simplified] by (simp add: multp\<^sub>D\<^sub>M_def less_multiset\<^sub>D\<^sub>M_def) lemma mult_less_multiset\<^sub>H\<^sub>O: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>H\<^sub>O M N" unfolding multp_def[of "(<)", symmetric] using multp_eq_multp\<^sub>H\<^sub>O[of "(<)", simplified] by (simp add: multp\<^sub>H\<^sub>O_def less_multiset\<^sub>H\<^sub>O_def) lemmas mult\<^sub>D\<^sub>M = mult_less_multiset\<^sub>D\<^sub>M[unfolded less_multiset\<^sub>D\<^sub>M_def] lemmas mult\<^sub>H\<^sub>O = mult_less_multiset\<^sub>H\<^sub>O[unfolded less_multiset\<^sub>H\<^sub>O_def] end lemma less_multiset_less_multiset\<^sub>H\<^sub>O: "M < N \ less_multiset\<^sub>H\<^sub>O M N" unfolding less_multiset_def multp_def mult\<^sub>H\<^sub>O less_multiset\<^sub>H\<^sub>O_def .. lemma less_multiset\<^sub>D\<^sub>M: "M < N \ (\X Y. X \ {#} \ X \# N \ M = N - X + Y \ (\k. k \# Y \ (\a. a \# X \ k < a)))" by (rule mult\<^sub>D\<^sub>M[folded multp_def less_multiset_def]) lemma less_multiset\<^sub>H\<^sub>O: "M < N \ M \ N \ (\y. count N y < count M y \ (\x>y. count M x < count N x))" by (rule mult\<^sub>H\<^sub>O[folded multp_def less_multiset_def]) lemma subset_eq_imp_le_multiset: shows "M \# N \ M \ N" unfolding less_eq_multiset_def less_multiset\<^sub>H\<^sub>O by (simp add: less_le_not_le subseteq_mset_def) (* FIXME: "le" should be "less" in this and other names *) lemma le_multiset_right_total: "M < add_mset x M" unfolding less_eq_multiset_def less_multiset\<^sub>H\<^sub>O by simp lemma less_eq_multiset_empty_left[simp]: shows "{#} \ M" by (simp add: subset_eq_imp_le_multiset) lemma ex_gt_imp_less_multiset: "(\y. y \# N \ (\x. x \# M \ x < y)) \ M < N" unfolding less_multiset\<^sub>H\<^sub>O by (metis count_eq_zero_iff count_greater_zero_iff less_le_not_le) lemma less_eq_multiset_empty_right[simp]: "M \ {#} \ \ M \ {#}" by (metis less_eq_multiset_empty_left antisym) (* FIXME: "le" should be "less" in this and other names *) lemma le_multiset_empty_left[simp]: "M \ {#} \ {#} < M" by (simp add: less_multiset\<^sub>H\<^sub>O) (* FIXME: "le" should be "less" in this and other names *) lemma le_multiset_empty_right[simp]: "\ M < {#}" using subset_mset.le_zero_eq less_multiset_def multp_def less_multiset\<^sub>D\<^sub>M by blast (* FIXME: "le" should be "less" in this and other names *) lemma union_le_diff_plus: "P \# M \ N < P \ M - P + N < M" by (drule subset_mset.diff_add[symmetric]) (metis union_le_mono2) instantiation multiset :: (preorder) ordered_ab_semigroup_monoid_add_imp_le begin lemma less_eq_multiset\<^sub>H\<^sub>O: "M \ N \ (\y. count N y < count M y \ (\x. y < x \ count M x < count N x))" by (auto simp: less_eq_multiset_def less_multiset\<^sub>H\<^sub>O) instance by standard (auto simp: less_eq_multiset\<^sub>H\<^sub>O) lemma fixes M N :: "'a multiset" shows less_eq_multiset_plus_left: "N \ (M + N)" and less_eq_multiset_plus_right: "M \ (M + N)" by simp_all lemma fixes M N :: "'a multiset" shows le_multiset_plus_left_nonempty: "M \ {#} \ N < M + N" and le_multiset_plus_right_nonempty: "N \ {#} \ M < M + N" by simp_all end lemma all_lt_Max_imp_lt_mset: "N \ {#} \ (\a \# M. a < Max (set_mset N)) \ M < N" by (meson Max_in[OF finite_set_mset] ex_gt_imp_less_multiset set_mset_eq_empty_iff) lemma lt_imp_ex_count_lt: "M < N \ \y. count M y < count N y" by (meson less_eq_multiset\<^sub>H\<^sub>O less_le_not_le) lemma subset_imp_less_mset: "A \# B \ A < B" by (simp add: order.not_eq_order_implies_strict subset_eq_imp_le_multiset) lemma image_mset_strict_mono: assumes mono_f: "\x \ set_mset M. \y \ set_mset N. x < y \ f x < f y" and less: "M < N" shows "image_mset f M < image_mset f N" proof - obtain Y X where y_nemp: "Y \ {#}" and y_sub_N: "Y \# N" and M_eq: "M = N - Y + X" and ex_y: "\x. x \# X \ (\y. y \# Y \ x < y)" using less[unfolded less_multiset\<^sub>D\<^sub>M] by blast have x_sub_M: "X \# M" using M_eq by simp let ?fY = "image_mset f Y" let ?fX = "image_mset f X" show ?thesis unfolding less_multiset\<^sub>D\<^sub>M proof (intro exI conjI) show "image_mset f M = image_mset f N - ?fY + ?fX" using M_eq[THEN arg_cong, of "image_mset f"] y_sub_N by (metis image_mset_Diff image_mset_union) next obtain y where y: "\x. x \# X \ y x \# Y \ x < y x" using ex_y by moura show "\fx. fx \# ?fX \ (\fy. fy \# ?fY \ fx < fy)" proof (intro allI impI) fix fx assume "fx \# ?fX" then obtain x where fx: "fx = f x" and x_in: "x \# X" by auto hence y_in: "y x \# Y" and y_gt: "x < y x" using y[rule_format, OF x_in] by blast+ hence "f (y x) \# ?fY \ f x < f (y x)" using mono_f y_sub_N x_sub_M x_in by (metis image_eqI in_image_mset mset_subset_eqD) thus "\fy. fy \# ?fY \ fx < fy" unfolding fx by auto qed qed (auto simp: y_nemp y_sub_N image_mset_subseteq_mono) qed lemma image_mset_mono: assumes mono_f: "\x \ set_mset M. \y \ set_mset N. x < y \ f x < f y" and less: "M \ N" shows "image_mset f M \ image_mset f N" by (metis eq_iff image_mset_strict_mono less less_imp_le mono_f order.not_eq_order_implies_strict) lemma mset_lt_single_right_iff[simp]: "M < {#y#} \ (\x \# M. x < y)" for y :: "'a::linorder" proof (rule iffI) assume M_lt_y: "M < {#y#}" show "\x \# M. x < y" proof fix x assume x_in: "x \# M" hence M: "M - {#x#} + {#x#} = M" by (meson insert_DiffM2) hence "\ {#x#} < {#y#} \ x < y" using x_in M_lt_y by (metis diff_single_eq_union le_multiset_empty_left less_add_same_cancel2 mset_le_trans) also have "\ {#y#} < M" using M_lt_y mset_le_not_sym by blast ultimately show "x < y" by (metis (no_types) Max_ge all_lt_Max_imp_lt_mset empty_iff finite_set_mset insertE less_le_trans linorder_less_linear mset_le_not_sym set_mset_add_mset_insert set_mset_eq_empty_iff x_in) qed next assume y_max: "\x \# M. x < y" show "M < {#y#}" by (rule all_lt_Max_imp_lt_mset) (auto intro!: y_max) qed lemma mset_le_single_right_iff[simp]: "M \ {#y#} \ M = {#y#} \ (\x \# M. x < y)" for y :: "'a::linorder" by (meson less_eq_multiset_def mset_lt_single_right_iff) subsubsection \Simplifications\ lemma multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[simp]: assumes "n \ 0" shows "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B) \ multp\<^sub>H\<^sub>O R A B" proof (rule iffI) assume hyp: "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)" hence 1: "repeat_mset n A \ repeat_mset n B" and 2: "\y. n * count B y < n * count A y \ (\x. R y x \ n * count A x < n * count B x)" by (simp_all add: multp\<^sub>H\<^sub>O_def) from 1 \n \ 0\ have "A \ B" by auto moreover from 2 \n \ 0\ have "\y. count B y < count A y \ (\x. R y x \ count A x < count B x)" by auto ultimately show "multp\<^sub>H\<^sub>O R A B" by (simp add: multp\<^sub>H\<^sub>O_def) next assume "multp\<^sub>H\<^sub>O R A B" hence 1: "A \ B" and 2: "\y. count B y < count A y \ (\x. R y x \ count A x < count B x)" by (simp_all add: multp\<^sub>H\<^sub>O_def) from 1 have "repeat_mset n A \ repeat_mset n B" by (simp add: assms repeat_mset_cancel1) moreover from 2 have "\y. n * count B y < n * count A y \ (\x. R y x \ n * count A x < n * count B x)" by auto ultimately show "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)" by (simp add: multp\<^sub>H\<^sub>O_def) qed lemma multp\<^sub>H\<^sub>O_double_double[simp]: "multp\<^sub>H\<^sub>O R (A + A) (B + B) \ multp\<^sub>H\<^sub>O R A B" using multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[of 2] by (simp add: numeral_Bit0) subsection \Simprocs\ lemma mset_le_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \ repeat_mset j u + n) = (repeat_mset (i-j) u + m \ n)" proof - assume "j \ i" then have "j + (i - j) = i" using le_add_diff_inverse by blast then show ?thesis by (metis (no_types) add_le_cancel_left left_add_mult_distrib_mset) qed lemma mset_le_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \ repeat_mset j u + n) = (m \ repeat_mset (j-i) u + n)" proof - assume "i \ j" then have "i + (j - i) = j" using le_add_diff_inverse by blast then show ?thesis by (metis (no_types) add_le_cancel_left left_add_mult_distrib_mset) qed simproc_setup msetless_cancel ("(l::'a::preorder multiset) + m < n" | "(l::'a multiset) < m + n" | "add_mset a m < n" | "m < add_mset a n" | "replicate_mset p a < n" | "m < replicate_mset p a" | "repeat_mset p m < n" | "m < repeat_mset p n") = - \fn phi => Cancel_Simprocs.less_cancel\ + \K Cancel_Simprocs.less_cancel\ simproc_setup msetle_cancel ("(l::'a::preorder multiset) + m \ n" | "(l::'a multiset) \ m + n" | "add_mset a m \ n" | "m \ add_mset a n" | "replicate_mset p a \ n" | "m \ replicate_mset p a" | "repeat_mset p m \ n" | "m \ repeat_mset p n") = - \fn phi => Cancel_Simprocs.less_eq_cancel\ + \K Cancel_Simprocs.less_eq_cancel\ subsection \Additional facts and instantiations\ lemma ex_gt_count_imp_le_multiset: "(\y :: 'a :: order. y \# M + N \ y \ x) \ count M x < count N x \ M < N" unfolding less_multiset\<^sub>H\<^sub>O by (metis count_greater_zero_iff le_imp_less_or_eq less_imp_not_less not_gr_zero union_iff) lemma mset_lt_single_iff[iff]: "{#x#} < {#y#} \ x < y" unfolding less_multiset\<^sub>H\<^sub>O by simp lemma mset_le_single_iff[iff]: "{#x#} \ {#y#} \ x \ y" for x y :: "'a::order" unfolding less_eq_multiset\<^sub>H\<^sub>O by force instance multiset :: (linorder) linordered_cancel_ab_semigroup_add by standard (metis less_eq_multiset\<^sub>H\<^sub>O not_less_iff_gr_or_eq) lemma less_eq_multiset_total: fixes M N :: "'a :: linorder multiset" shows "\ M \ N \ N \ M" by simp instantiation multiset :: (wellorder) wellorder begin lemma wf_less_multiset: "wf {(M :: 'a multiset, N). M < N}" unfolding less_multiset_def multp_def by (auto intro: wf_mult wf) instance by standard (metis less_multiset_def multp_def wf wf_def wf_mult) end instantiation multiset :: (preorder) order_bot begin definition bot_multiset :: "'a multiset" where "bot_multiset = {#}" instance by standard (simp add: bot_multiset_def) end instance multiset :: (preorder) no_top proof standard fix x :: "'a multiset" obtain a :: 'a where True by simp have "x < x + (x + {#a#})" by simp then show "\y. x < y" by blast qed instance multiset :: (preorder) ordered_cancel_comm_monoid_add by standard instantiation multiset :: (linorder) distrib_lattice begin definition inf_multiset :: "'a multiset \ 'a multiset \ 'a multiset" where "inf_multiset A B = (if A < B then A else B)" definition sup_multiset :: "'a multiset \ 'a multiset \ 'a multiset" where "sup_multiset A B = (if B > A then B else A)" instance by intro_classes (auto simp: inf_multiset_def sup_multiset_def) end end diff --git a/src/HOL/Library/adhoc_overloading.ML b/src/HOL/Library/adhoc_overloading.ML --- a/src/HOL/Library/adhoc_overloading.ML +++ b/src/HOL/Library/adhoc_overloading.ML @@ -1,245 +1,245 @@ (* Author: Alexander Krauss, TU Muenchen Author: Christian Sternagel, University of Innsbruck Adhoc overloading of constants based on their types. *) signature ADHOC_OVERLOADING = sig val is_overloaded: Proof.context -> string -> bool val generic_add_overloaded: string -> Context.generic -> Context.generic val generic_remove_overloaded: string -> Context.generic -> Context.generic val generic_add_variant: string -> term -> Context.generic -> Context.generic (*If the list of variants is empty at the end of "generic_remove_variant", then "generic_remove_overloaded" is called implicitly.*) val generic_remove_variant: string -> term -> Context.generic -> Context.generic val show_variants: bool Config.T end structure Adhoc_Overloading: ADHOC_OVERLOADING = struct val show_variants = Attrib.setup_config_bool \<^binding>\show_variants\ (K false); (* errors *) fun err_duplicate_variant oconst = error ("Duplicate variant of " ^ quote oconst); fun err_not_a_variant oconst = error ("Not a variant of " ^ quote oconst); fun err_not_overloaded oconst = error ("Constant " ^ quote oconst ^ " is not declared as overloaded"); fun err_unresolved_overloading ctxt0 (c, T) t instances = let val ctxt = Config.put show_variants true ctxt0 val const_space = Proof_Context.const_space ctxt val prt_const = Pretty.block [Name_Space.pretty ctxt const_space c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)] in error (Pretty.string_of (Pretty.chunks [ Pretty.block [ Pretty.str "Unresolved adhoc overloading of constant", Pretty.brk 1, prt_const, Pretty.brk 1, Pretty.str "in term", Pretty.brk 1, Pretty.block [Pretty.quote (Syntax.pretty_term ctxt t)]], Pretty.block ( (if null instances then [Pretty.str "no instances"] else Pretty.fbreaks ( Pretty.str "multiple instances:" :: map (Pretty.mark Markup.item o Syntax.pretty_term ctxt) instances)))])) end; (* generic data *) fun variants_eq ((v1, T1), (v2, T2)) = Term.aconv_untyped (v1, v2) andalso T1 = T2; structure Overload_Data = Generic_Data ( type T = {variants : (term * typ) list Symtab.table, oconsts : string Termtab.table}; val empty = {variants = Symtab.empty, oconsts = Termtab.empty}; fun merge ({variants = vtab1, oconsts = otab1}, {variants = vtab2, oconsts = otab2}) : T = let fun merge_oconsts _ (oconst1, oconst2) = if oconst1 = oconst2 then oconst1 else err_duplicate_variant oconst1; in {variants = Symtab.merge_list variants_eq (vtab1, vtab2), oconsts = Termtab.join merge_oconsts (otab1, otab2)} end; ); fun map_tables f g = Overload_Data.map (fn {variants = vtab, oconsts = otab} => {variants = f vtab, oconsts = g otab}); val is_overloaded = Symtab.defined o #variants o Overload_Data.get o Context.Proof; val get_variants = Symtab.lookup o #variants o Overload_Data.get o Context.Proof; val get_overloaded = Termtab.lookup o #oconsts o Overload_Data.get o Context.Proof; fun generic_add_overloaded oconst context = if is_overloaded (Context.proof_of context) oconst then context else map_tables (Symtab.update (oconst, [])) I context; fun generic_remove_overloaded oconst context = let fun remove_oconst_and_variants context oconst = let val remove_variants = (case get_variants (Context.proof_of context) oconst of NONE => I | SOME vs => fold (Termtab.remove (op =) o rpair oconst o fst) vs); in map_tables (Symtab.delete_safe oconst) remove_variants context end; in if is_overloaded (Context.proof_of context) oconst then remove_oconst_and_variants context oconst else err_not_overloaded oconst end; local fun generic_variant add oconst t context = let val ctxt = Context.proof_of context; val _ = if is_overloaded ctxt oconst then () else err_not_overloaded oconst; val T = t |> fastype_of; val t' = Term.map_types (K dummyT) t; in if add then let val _ = (case get_overloaded ctxt t' of NONE => () | SOME oconst' => err_duplicate_variant oconst'); in map_tables (Symtab.cons_list (oconst, (t', T))) (Termtab.update (t', oconst)) context end else let val _ = if member variants_eq (the (get_variants ctxt oconst)) (t', T) then () else err_not_a_variant oconst; in map_tables (Symtab.map_entry oconst (remove1 variants_eq (t', T))) (Termtab.delete_safe t') context |> (fn context => (case get_variants (Context.proof_of context) oconst of SOME [] => generic_remove_overloaded oconst context | _ => context)) end end; in val generic_add_variant = generic_variant true; val generic_remove_variant = generic_variant false; end; (* check / uncheck *) fun unifiable_with thy T1 T2 = let val maxidx1 = Term.maxidx_of_typ T1; val T2' = Logic.incr_tvar (maxidx1 + 1) T2; val maxidx2 = Term.maxidx_typ T2' maxidx1; in can (Sign.typ_unify thy (T1, T2')) (Vartab.empty, maxidx2) end; fun get_candidates ctxt (c, T) = get_variants ctxt c |> Option.map (map_filter (fn (t, T') => if unifiable_with (Proof_Context.theory_of ctxt) T T' then SOME t else NONE)); fun insert_variants ctxt t (oconst as Const (c, T)) = (case get_candidates ctxt (c, T) of SOME [] => err_unresolved_overloading ctxt (c, T) t [] | SOME [variant] => variant | _ => oconst) | insert_variants _ _ oconst = oconst; fun insert_overloaded ctxt = let fun proc t = Term.map_types (K dummyT) t |> get_overloaded ctxt |> Option.map (Const o rpair (Term.type_of t)); in Pattern.rewrite_term_top (Proof_Context.theory_of ctxt) [] [proc] end; fun check ctxt = map (fn t => Term.map_aterms (insert_variants ctxt t) t); fun uncheck ctxt ts = if Config.get ctxt show_variants orelse exists (is_none o try Term.type_of) ts then ts else map (insert_overloaded ctxt) ts; fun reject_unresolved ctxt = let val the_candidates = the o get_candidates ctxt; fun check_unresolved t = (case filter (is_overloaded ctxt o fst) (Term.add_consts t []) of [] => t | (cT :: _) => err_unresolved_overloading ctxt cT t (the_candidates cT)); in map check_unresolved end; (* setup *) val _ = Context.>> (Syntax_Phases.term_check 0 "adhoc_overloading" check #> Syntax_Phases.term_check 1 "adhoc_overloading_unresolved_check" reject_unresolved #> Syntax_Phases.term_uncheck 0 "adhoc_overloading" uncheck); (* commands *) fun generic_adhoc_overloading_cmd add = if add then fold (fn (oconst, ts) => generic_add_overloaded oconst #> fold (generic_add_variant oconst) ts) else fold (fn (oconst, ts) => fold (generic_remove_variant oconst) ts); fun adhoc_overloading_cmd' add args phi = let val args' = args |> map (apsnd (map_filter (fn t => let val t' = Morphism.term phi t; in if Term.aconv_untyped (t, t') then SOME t' else NONE end))); in generic_adhoc_overloading_cmd add args' end; fun adhoc_overloading_cmd add raw_args lthy = let fun const_name ctxt = fst o dest_Const o Proof_Context.read_const {proper = false, strict = false} ctxt; (* FIXME {proper = true, strict = true} (!?) *) fun read_term ctxt = singleton (Variable.polymorphic ctxt) o Syntax.read_term ctxt; val args = raw_args |> map (apfst (const_name lthy)) |> map (apsnd (map (read_term lthy))); in - Local_Theory.declaration {syntax = true, pervasive = false} + Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (adhoc_overloading_cmd' add args) lthy end; val _ = Outer_Syntax.local_theory \<^command_keyword>\adhoc_overloading\ "add adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd true); val _ = Outer_Syntax.local_theory \<^command_keyword>\no_adhoc_overloading\ "delete adhoc overloading for constants / fixed variables" (Parse.and_list1 (Parse.const -- Scan.repeat Parse.term) >> adhoc_overloading_cmd false); end; diff --git a/src/HOL/List.thy b/src/HOL/List.thy --- a/src/HOL/List.thy +++ b/src/HOL/List.thy @@ -1,8354 +1,8354 @@ (* Title: HOL/List.thy Author: Tobias Nipkow; proofs tidied by LCP *) section \The datatype of finite lists\ theory List imports Sledgehammer 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" fun successively :: "('a \ 'a \ bool) \ 'a list \ bool" where "successively P [] = True" | "successively P [x] = True" | "successively P (x # y # xs) = (P x y \ successively P (y#xs))" definition distinct_adj where "distinct_adj = successively (\)" 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 "successively (\) [True,False,True,False]" 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(ed) 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)" text \A class-based sorted predicate:\ context linorder begin abbreviation sorted :: "'a list \ bool" where "sorted \ sorted_wrt (\)" lemma sorted_simps: "sorted [] = True" "sorted (x # ys) = ((\y \ set ys. x\y) \ sorted ys)" by auto lemma strict_sorted_simps: "sorted_wrt (<) [] = True" "sorted_wrt (<) (x # ys) = ((\y \ set ys. x sorted_wrt (<) ys)" by auto 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)" lemma strict_sorted_iff: "sorted_wrt (<) l \ sorted l \ distinct l" by (induction l) (auto iff: antisym_conv1) lemma strict_sorted_imp_sorted: "sorted_wrt (<) xs \ sorted xs" by (auto simp: strict_sorted_iff) 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 lemmas Nil_tl = tl_Nil[THEN eq_iff_swap] 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 lemmas Suc_length_conv = length_Suc_conv[THEN eq_iff_swap] 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 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 lemmas Nil_is_append_conv [iff] = append_is_Nil_conv[THEN eq_iff_swap] lemma append_self_conv [iff]: "(xs @ ys = xs) = (ys = [])" by (induct xs) auto lemmas self_append_conv [iff] = append_self_conv[THEN eq_iff_swap] 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 lemmas self_append_conv2 [iff] = append_self_conv2[THEN eq_iff_swap] 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 tl_append_if: "tl (xs @ ys) = (if xs = [] then tl ys else tl xs @ ys)" by (simp) 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, opaque_lifting) 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 auto lemma append_eq_appendI: "\xs @ xs1 = zs; ys = xs1 @ us\ \ xs @ ys = zs @ us" by auto 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 + in K (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 (rule list.map_disc_iff) lemmas Nil_is_map_conv [iff] = map_is_Nil_conv[THEN eq_iff_swap] 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, opaque_lifting) 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 lemmas Nil_is_rev_conv [iff] = rev_is_Nil_conv[THEN eq_iff_swap] lemma rev_singleton_conv [simp]: "(rev xs = [x]) = (xs = [x])" by (cases xs) auto lemma singleton_rev_conv [simp]: "([x] = rev xs) = ([x] = xs)" 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]: assumes "P []" and "\x xs. P xs \ P (xs @ [x])" shows "P xs" proof - have "P (rev (rev xs))" by (rule_tac list = "rev xs" in list.induct, simp_all add: assms) then show ?thesis by simp qed 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 lemma length_Suc_conv_rev: "(length xs = Suc n) = (\y ys. xs = ys @ [y] \ length ys = n)" by (induct xs rule: rev_induct) 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 lemmas set_empty2[iff] = set_empty[THEN eq_iff_swap] 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>\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 lemmas Nil_eq_concat_conv [simp] = concat_eq_Nil_conv[THEN eq_iff_swap] 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 rev_concat: "rev (concat xs) = concat (map rev (rev xs))" by (induct xs) auto lemma length_concat_rev[simp]: "length (concat (rev xs)) = length (concat xs)" by (induction 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) lemma hd_concat: "\xs \ []; hd xs \ []\ \ hd (concat xs) = hd (hd xs)" by (metis concat.simps(2) hd_Cons_tl hd_append2) 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 (Const(\<^const_name>\concat\,T) $ (Const(\<^const_name>\rev\,_) $ xss)) acc = len (Const(\<^const_name>\concat\,T) $ xss) 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 \\<^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 filter_concat: "filter p (concat xs) = concat (map (filter p) 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 lemmas empty_filter_conv = filter_empty_conv[THEN eq_iff_swap] 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) 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) lemmas Cons_eq_filter_iff = filter_eq_Cons_iff[THEN eq_iff_swap] 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>\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 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 hd_Nil_eq_last: "hd Nil = last Nil" unfolding hd_def last_def by simp 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: "hd(rev xs) = last xs" by (metis hd_Cons_tl hd_Nil_eq_last last_snoc rev_eq_Cons_iff rev_is_Nil_conv) lemma last_rev: "last(rev xs) = hd xs" by (metis hd_rev rev_swap) 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_all_iff [simp]: "take n xs = xs \ length xs \ n" by (metis length_take min.order_iff take_all) (* Looks like a good simp rule but can cause looping; too much interaction between take and length lemmas take_all_iff2[simp] = take_all_iff[THEN eq_iff_swap] *) lemma take_eq_Nil[simp]: "(take n xs = []) = (n = 0 \ xs = [])" by(induct xs arbitrary: n)(auto simp: take_Cons split:nat.split) lemmas take_eq_Nil2[simp] = take_eq_Nil[THEN eq_iff_swap] lemma drop_eq_Nil [simp]: "drop n xs = [] \ length xs \ n" by (metis diff_is_0_eq drop_all length_drop list.size(3)) lemmas drop_eq_Nil2 [simp] = drop_eq_Nil[THEN eq_iff_swap] 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_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) 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) lemma drop_butlast: "drop n (butlast xs) = butlast (drop n xs)" by (simp add: butlast_conv_take drop_take ac_simps) lemma butlast_power: "(butlast ^^ n) xs = take (length xs - n) xs" by (induct n) (auto simp: butlast_take) 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 lemmas append_eq_map_conv = map_eq_append_conv[THEN eq_iff_swap] 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: 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_append: "takeWhile P (xs @ ys) = (if \x\set xs. P x then xs @ takeWhile P ys else takeWhile P xs)" using takeWhile_append1[of _ xs P ys] takeWhile_append2[of xs P ys] by auto lemma takeWhile_tail: "\ P x \ takeWhile P (xs @ (x#l)) = takeWhile P xs" by (induct xs) auto lemma takeWhile_eq_Nil_iff: "takeWhile P xs = [] \ xs = [] \ \P (hd xs)" by (cases xs) auto lemma takeWhile_nth: "j < length (takeWhile P xs) \ takeWhile P xs ! j = xs ! j" by (metis nth_append takeWhile_dropWhile_id) lemma takeWhile_takeWhile: "takeWhile Q (takeWhile P xs) = takeWhile (\x. P x \ Q x) xs" by(induct xs) simp_all 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_append: "dropWhile P (xs @ ys) = (if \x\set xs. P x then dropWhile P ys else dropWhile P xs @ ys)" using dropWhile_append1[of _ xs P ys] dropWhile_append2[of xs P ys] by 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 dropWhile_eq_self_iff: "dropWhile P xs = xs \ xs = [] \ \P (hd xs)" by (cases xs) (auto simp: dropWhile_eq_Cons_conv) lemma dropWhile_dropWhile1: "(\x. Q x \ P x) \ dropWhile Q (dropWhile P xs) = dropWhile P xs" by(induct xs) simp_all lemma dropWhile_dropWhile2: "(\x. P x \ Q x) \ takeWhile P (takeWhile Q xs) = takeWhile P xs" by(induct xs) simp_all lemma dropWhile_takeWhile: "(\x. P x \ Q x) \ dropWhile P (takeWhile Q xs) = takeWhile Q (dropWhile P xs)" by (induction 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[simp]: "zip xs ys = [] \ xs = [] \ ys = []" by (cases xs; cases ys) simp_all lemmas Nil_eq_zip_iff[simp] = zip_eq_Nil_iff[THEN eq_iff_swap] 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 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 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 lemma hd_zip: \hd (zip xs ys) = (hd xs, hd ys)\ if \xs \ []\ and \ys \ []\ using that by (cases xs; cases ys) simp_all lemma last_zip: \last (zip xs ys) = (last xs, last ys)\ if \xs \ []\ and \ys \ []\ and \length xs = length ys\ using that by (cases xs rule: rev_cases; cases ys rule: rev_cases) simp_all 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_on) fold_set_fold_remdups: assumes "set xs \ S" shows "Finite_Set.fold f y (set xs) = fold f (remdups xs) y" by (rule sym, use assms in \induct xs arbitrary: y\) (simp_all add: insert_absorb fold_fun_left_comm) lemma (in comp_fun_idem_on) fold_set_fold: assumes "set xs \ S" shows "Finite_Set.fold f y (set xs) = fold f xs y" by (rule sym, use assms in \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') lemma map_upt_eqI: \map f [m.. if \length xs = n - m\ \\i. i < length xs \ xs ! i = f (m + i)\ proof (rule nth_equalityI) from \length xs = n - m\ show \length (map f [m.. by simp next fix i assume \i < length (map f [m.. then have \i < n - m\ by simp with that have \xs ! i = f (m + i)\ by simp with \i < n - m\ show \map f [m.. by simp qed 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) lemmas upto_Nil2[simp] = upto_Nil[THEN eq_iff_swap] 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" proof(induction i j arbitrary: k rule: upto.induct) case (1 i j) then show ?case by (auto simp add: upto_rec1 [of i j] nth_Cons') qed 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: "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>\successively\\ lemma successively_Cons: "successively P (x # xs) \ xs = [] \ P x (hd xs) \ successively P xs" by (cases xs) auto lemma successively_cong [cong]: assumes "\x y. x \ set xs \ y \ set xs \ P x y \ Q x y" "xs = ys" shows "successively P xs \ successively Q ys" unfolding assms(2) [symmetric] using assms(1) by (induction xs) (auto simp: successively_Cons) lemma successively_append_iff: "successively P (xs @ ys) \ successively P xs \ successively P ys \ (xs = [] \ ys = [] \ P (last xs) (hd ys))" by (induction xs) (auto simp: successively_Cons) lemma successively_if_sorted_wrt: "sorted_wrt P xs \ successively P xs" by (induction xs rule: induct_list012) auto lemma successively_iff_sorted_wrt_strong: assumes "\x y z. x \ set xs \ y \ set xs \ z \ set xs \ P x y \ P y z \ P x z" shows "successively P xs \ sorted_wrt P xs" proof assume "successively P xs" from this and assms show "sorted_wrt P xs" proof (induction xs rule: induct_list012) case (3 x y xs) from "3.prems" have "P x y" by auto have IH: "sorted_wrt P (y # xs)" using "3.prems" by(intro "3.IH"(2) list.set_intros(2))(simp, blast intro: list.set_intros(2)) have "P x z" if asm: "z \ set xs" for z proof - from IH and asm have "P y z" by auto with \P x y\ show "P x z" using "3.prems" asm by auto qed with IH and \P x y\ show ?case by auto qed auto qed (use successively_if_sorted_wrt in blast) lemma successively_conv_sorted_wrt: assumes "transp P" shows "successively P xs \ sorted_wrt P xs" using assms unfolding transp_def by (intro successively_iff_sorted_wrt_strong) blast lemma successively_rev [simp]: "successively P (rev xs) \ successively (\x y. P y x) xs" by (induction xs rule: remdups_adj.induct) (auto simp: successively_append_iff successively_Cons) lemma successively_map: "successively P (map f xs) \ successively (\x y. P (f x) (f y)) xs" by (induction xs rule: induct_list012) auto lemma successively_mono: assumes "successively P xs" assumes "\x y. x \ set xs \ y \ set xs \ P x y \ Q x y" shows "successively Q xs" using assms by (induction Q xs rule: successively.induct) auto lemma successively_altdef: "successively = (\P. rec_list True (\x xs b. case xs of [] \ True | y # _ \ P x y \ b))" proof (intro ext) fix P and xs :: "'a list" show "successively P xs = rec_list True (\x xs b. case xs of [] \ True | y # _ \ P x y \ b) xs" by (induction xs) (auto simp: successively_Cons split: list.splits) qed subsubsection \\<^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) lemmas remdups_eq_nil_right_iff [simp] = remdups_eq_nil_iff[THEN eq_iff_swap] 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 have "set (take i xs) \ set (drop (Suc i) xs) = {}" by (metis True d disjoint_insert(1) distinct_append id_take_nth_drop list.set(2)) then show ?thesis using d False anot \i < length xs\ by (simp add: upd_conv_take_nth_drop) 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 \An iff-version of @{thm distinct_concat} is available further down as \distinct_concat_iff\.\ text \It is best to avoid the following 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 less_Suc_eq_le split: nat.split_asm) apply (metis Suc_leI in_set_conv_nth length_pos_if_in_set lessI less_imp_le_nat less_nat_zero_code) apply (metis Suc_le_eq) 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; 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)}" using less_Suc_eq_0_disj by auto 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 using less_Suc_eq_0_disj by auto 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 two: "Suc (Suc 0) \ 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 two by auto qed obtain ys' where "ys = x1 # x2 # ys'" using two f_nth[of 0] f_nth[of 1] by (auto simp: Suc_le_length_iff \f 0 = 0\ \f (Suc 0) = Suc 0\) have Suc0_le_f_Suc: "Suc 0 \ f (Suc i)" for i by (metis Suc_le_mono \f (Suc 0) = Suc 0\ f_mono le0 mono_def) define f' where "f' x = f (Suc x) - 1" for x have f_Suc: "f (Suc i) = Suc (f' i)" for i using Suc0_le_f_Suc[of i] by (auto simp: f'_def) 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 last_remdups_adj [simp]: "last (remdups_adj xs) = last xs" by (induction xs rule: remdups_adj.induct) auto 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.. successively P (remdups_adj xs)" by (induction xs rule: remdups_adj.induct) (auto simp: successively_Cons) lemma successively_remdups_adj_iff: "(\x. x \ set xs \ P x x) \ successively P (remdups_adj xs) \ successively P xs" by (induction xs rule: remdups_adj.induct)(auto simp: successively_Cons) lemma remdups_adj_Cons': "remdups_adj (x # xs) = x # remdups_adj (dropWhile (\y. y = x) xs)" by (induction xs) auto lemma remdups_adj_singleton_iff: "length (remdups_adj xs) = Suc 0 \ xs \ [] \ xs = replicate (length xs) (hd xs)" proof safe assume *: "xs = replicate (length xs) (hd xs)" and [simp]: "xs \ []" show "length (remdups_adj xs) = Suc 0" by (subst *) (auto simp: remdups_adj_replicate) next assume "length (remdups_adj xs) = Suc 0" thus "xs = replicate (length xs) (hd xs)" by (induction xs rule: remdups_adj.induct) (auto split: if_splits) qed auto lemma tl_remdups_adj: "ys \ [] \ tl (remdups_adj ys) = remdups_adj (dropWhile (\x. x = hd ys) (tl ys))" by (cases ys) (simp_all add: remdups_adj_Cons') lemma remdups_adj_append_dropWhile: "remdups_adj (xs @ y # ys) = remdups_adj (xs @ [y]) @ remdups_adj (dropWhile (\x. x = y) ys)" by (subst remdups_adj_append) (simp add: tl_remdups_adj) lemma remdups_adj_append': assumes "xs = [] \ ys = [] \ last xs \ hd ys" shows "remdups_adj (xs @ ys) = remdups_adj xs @ remdups_adj ys" proof - have ?thesis if [simp]: "xs \ []" "ys \ []" and "last xs \ hd ys" proof - obtain x xs' where xs: "xs = xs' @ [x]" by (cases xs rule: rev_cases) auto have "remdups_adj (xs' @ x # ys) = remdups_adj (xs' @ [x]) @ remdups_adj ys" using \last xs \ hd ys\ unfolding xs by (metis (full_types) dropWhile_eq_self_iff last_snoc remdups_adj_append_dropWhile) thus ?thesis by (simp add: xs) qed thus ?thesis using assms by (cases "xs = []"; cases "ys = []") auto qed lemma remdups_adj_append'': "xs \ [] \ remdups_adj (xs @ ys) = remdups_adj xs @ remdups_adj (dropWhile (\y. y = last xs) ys)" by (induction xs rule: remdups_adj.induct) (auto simp: remdups_adj_Cons') subsection \@{const distinct_adj}\ lemma distinct_adj_Nil [simp]: "distinct_adj []" and distinct_adj_singleton [simp]: "distinct_adj [x]" and distinct_adj_Cons_Cons [simp]: "distinct_adj (x # y # xs) \ x \ y \ distinct_adj (y # xs)" by (auto simp: distinct_adj_def) lemma distinct_adj_Cons: "distinct_adj (x # xs) \ xs = [] \ x \ hd xs \ distinct_adj xs" by (cases xs) auto lemma distinct_adj_ConsD: "distinct_adj (x # xs) \ distinct_adj xs" by (cases xs) auto lemma distinct_adj_remdups_adj[simp]: "distinct_adj (remdups_adj xs)" by (induction xs rule: remdups_adj.induct) (auto simp: distinct_adj_Cons) lemma distinct_adj_altdef: "distinct_adj xs \ remdups_adj xs = xs" proof assume "remdups_adj xs = xs" with distinct_adj_remdups_adj[of xs] show "distinct_adj xs" by simp next assume "distinct_adj xs" thus "remdups_adj xs = xs" by (induction xs rule: induct_list012) auto qed lemma distinct_adj_rev [simp]: "distinct_adj (rev xs) \ distinct_adj xs" by (simp add: distinct_adj_def eq_commute) lemma distinct_adj_append_iff: "distinct_adj (xs @ ys) \ distinct_adj xs \ distinct_adj ys \ (xs = [] \ ys = [] \ last xs \ hd ys)" by (auto simp: distinct_adj_def successively_append_iff) lemma distinct_adj_appendD1 [dest]: "distinct_adj (xs @ ys) \ distinct_adj xs" and distinct_adj_appendD2 [dest]: "distinct_adj (xs @ ys) \ distinct_adj ys" by (auto simp: distinct_adj_append_iff) lemma distinct_adj_mapI: "distinct_adj xs \ inj_on f (set xs) \ distinct_adj (map f xs)" unfolding distinct_adj_def successively_map by (erule successively_mono) (auto simp: inj_on_def) lemma distinct_adj_mapD: "distinct_adj (map f xs) \ distinct_adj xs" unfolding distinct_adj_def successively_map by (erule successively_mono) auto lemma distinct_adj_map_iff: "inj_on f (set xs) \ distinct_adj (map f xs) \ distinct_adj xs" using distinct_adj_mapD distinct_adj_mapI by blast subsubsection \\<^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 lemmas find_None_iff2 = find_None_iff[THEN eq_iff_swap] 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 lemmas find_Some_iff2 = find_Some_iff[THEN eq_iff_swap] 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\\ text \This library is intentionally minimal. See the remark about multisets at the point above where @{const count_list} is defined.\ lemma count_list_append[simp]: "count_list (xs @ ys) x = count_list xs x + count_list ys x" by (induction xs) simp_all lemma count_list_0_iff: "count_list xs x = 0 \ x \ set xs" by (induction xs) simp_all lemma count_notin[simp]: "x \ set xs \ count_list xs x = 0" by(simp add: count_list_0_iff) lemma count_le_length: "count_list xs x \ length xs" by (induction xs) auto lemma count_list_map_ge: "count_list xs x \ count_list (map f xs) (f x)" by (induction xs) auto lemma count_list_inj_map: "\ inj_on f (set xs); x \ set xs \ \ count_list (map f xs) (f x) = count_list xs x" by (induction xs) (simp, fastforce) lemma count_list_rev[simp]: "count_list (rev xs) x = count_list xs x" 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 using sum.remove [of X x "count_list xs"] by (auto simp: sum.If_cases simp flip: diff_eq) 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 lemma remove1_split: "a \ set xs \ remove1 a xs = ys \ (\ls rs. xs = ls @ a # rs \ a \ set ls \ ys = ls @ rs)" by (metis remove1.simps(2) remove1_append split_list_first) 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) lemma distinct_concat_iff: "distinct (concat xs) \ distinct (removeAll [] xs) \ (\ys. ys \ set xs \ distinct ys) \ (\ys zs. ys \ set xs \ zs \ set xs \ ys \ zs \ set ys \ set zs = {})" apply (induct xs) apply(simp_all, safe, auto) by (metis Int_iff UN_I empty_iff equals0I set_empty) 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 card_set_1_iff_replicate: "card(set xs) = Suc 0 \ xs \ [] \ (\x. xs = replicate (length xs) x)" by (metis card_1_singleton_iff empty_iff insert_iff replicate_eqI set_empty2 set_replicate) 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 lemmas empty_replicate[simp] = replicate_empty[THEN eq_iff_swap] 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 takeWhile_replicate[simp]: "takeWhile P (replicate n x) = (if P x then replicate n x else [])" using takeWhile_eq_Nil_iff by fastforce lemma dropWhile_replicate[simp]: "dropWhile P (replicate n x) = (if P x then [] else replicate n x)" using dropWhile_eq_self_iff by fastforce 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 = ys @ xs \ \m n zs. concat (replicate m zs) = xs \ concat (replicate n zs) = ys" proof (induction "length (xs @ ys) + length xs" arbitrary: xs ys rule: less_induct) case less consider (1) "length ys < length xs" | (2) "xs = []" | (3) "length xs \ length ys \ xs \ []" by linarith then show ?case proof (cases) case 1 then show ?thesis using less.hyps[OF _ less.prems[symmetric]] nat_add_left_cancel_less by auto next case 2 then have "concat (replicate 0 ys) = xs \ concat (replicate 1 ys) = ys" by simp then show ?thesis by blast next case 3 then have "length xs \ length ys" and "xs \ []" by blast+ from \length xs \ length ys\ and \xs @ ys = ys @ xs\ obtain ws where "ys = xs @ ws" by (auto simp: append_eq_append_conv2) from this and \xs \ []\ have "length ws < length ys" by simp from \xs @ ys = ys @ xs\[unfolded \ys = xs @ ws\] have "xs @ ws = ws @ xs" by simp from less.hyps[OF _ this] \length ws < length ys\ obtain m n' zs where "concat (replicate m zs) = xs" and "concat (replicate n' zs) = ws" by auto then have "concat (replicate (m+n') zs) = ys" using \ys = xs @ ws\ by (simp add: replicate_add) then show ?thesis using \concat (replicate m zs) = xs\ by blast qed 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 assms(3)] 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 by (auto simp add: mod_Suc False Suc.hyps drop_Suc rotate1_hd_tl take_Suc Suc_length_conv) 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_replicate[simp]: "rotate1 (replicate n a) = replicate n a" by (cases n) (simp_all add: replicate_append_same) 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) lemma nth_rotate: \rotate m xs ! n = xs ! ((m + n) mod length xs)\ if \n < length xs\ using that apply (auto simp add: rotate_drop_take nth_append not_less less_diff_conv ac_simps dest!: le_Suc_ex) apply (metis add.commute mod_add_right_eq mod_less) apply (metis (no_types, lifting) Nat.diff_diff_right add.commute add_diff_cancel_right' diff_le_self dual_order.strict_trans2 length_greater_0_conv less_nat_zero_code list.size(3) mod_add_right_eq mod_add_self2 mod_le_divisor mod_less) done lemma nth_rotate1: \rotate1 xs ! n = xs ! (Suc n mod length xs)\ if \n < length xs\ using that nth_rotate [of n xs 1] by simp lemma inj_rotate1: "inj rotate1" proof fix xs ys :: "'a list" show "rotate1 xs = rotate1 ys \ xs = ys" by (cases xs; cases ys; simp) qed lemma surj_rotate1: "surj rotate1" proof (safe, simp_all) fix xs :: "'a list" show "xs \ range rotate1" proof (cases xs rule: rev_exhaust) case Nil hence "xs = rotate1 []" by auto thus ?thesis by fast next case (snoc as a) hence "xs = rotate1 (a#as)" by force thus ?thesis by fast qed qed lemma bij_rotate1: "bij (rotate1 :: 'a list \ 'a list)" using bijI inj_rotate1 surj_rotate1 by blast lemma rotate1_fixpoint_card: "rotate1 xs = xs \ xs = [] \ card(set xs) = 1" by(induction xs) (auto simp: card_insert_if[OF finite_set] append_eq_Cons_conv) 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 (auto simp: in_set_zip subset_iff) 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}" by (induction xs arbitrary: A B) (auto simp add: nths_Cons card_less_Suc card_less_Suc2) lemma drop_eq_nths: "drop n xs = nths xs {i. i \ n}" by (induction xs arbitrary: n) (auto simp add: nths_Cons nths_all drop_Cons' intro: arg_cong2[where f=nths, OF refl]) 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))" by (simp add: assms card_Pow card_distinct distinct_card length_subseqs subseqs_powset) 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" proof (induction "replicate m x" "replicate n x" arbitrary: m n rule: splice.induct) case (2 x xs) then show ?case by (auto simp add: Cons_replicate_eq dest: gr0_implies_Suc) qed auto 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 xs :: \'a list\ have "Suc j < length xs \ xs \ [] \ j < length xs - Suc 0" by (cases xs) 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) lemma same_length_different: assumes "xs \ ys" and "length xs = length ys" shows "\pre x xs' y ys'. x\y \ xs = pre @ [x] @ xs' \ ys = pre @ [y] @ ys'" using assms proof (induction xs arbitrary: ys) case Nil then show ?case by auto next case (Cons x xs) then obtain z zs where ys: "ys = Cons z zs" by (metis length_Suc_conv) show ?case proof (cases "x=z") case True then have "xs \ zs" "length xs = length zs" using Cons.prems ys by auto then obtain pre u xs' v ys' where "u\v" and xs: "xs = pre @ [u] @ xs'" and zs: "zs = pre @ [v] @ys'" using Cons.IH by meson then have "x # xs = (z#pre) @ [u] @ xs' \ ys = (z#pre) @ [v] @ ys'" by (simp add: True ys) with \u\v\ show ?thesis by blast next case False then have "x # xs = [] @ [x] @ xs \ ys = [] @ [z] @ zs" by (simp add: ys) then show ?thesis using False by blast qed qed 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 assumes "sorted_wrt f xs" shows sorted_wrt_take: "sorted_wrt f (take n xs)" and sorted_wrt_drop: "sorted_wrt f (drop n xs)" proof - from assms have "sorted_wrt f (take n xs @ drop n xs)" by simp thus "sorted_wrt f (take n xs)" and "sorted_wrt f (drop n xs)" unfolding sorted_wrt_append by simp_all qed lemma sorted_wrt_filter: "sorted_wrt f xs \ sorted_wrt f (filter P xs)" by (induction xs) auto 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_iff_nth_Suc_transp: assumes "transp P" shows "sorted_wrt P xs \ (\i. Suc i < length xs \ P (xs!i) (xs!(Suc i)))" (is "?L = ?R") proof assume ?L thus ?R by (simp add: sorted_wrt_iff_nth_less) next assume ?R have "i < j \ j < length xs \ P (xs ! i) (xs ! j)" for i j by(induct i j rule: less_Suc_induct)(simp add: \?R\, meson assms transpE transp_on_less) thus ?L by (simp add: sorted_wrt_iff_nth_less) qed 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 (simp add: 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 sorted0: "sorted [] = True" by simp lemma sorted1: "sorted [x] = True" by simp lemma sorted2: "sorted (x # y # zs) = (x \ y \ sorted (y # zs))" by auto lemmas sorted2_simps = sorted1 sorted2 lemma sorted_append: "sorted (xs@ys) = (sorted xs \ sorted ys \ (\x \ set xs. \y \ set ys. x\y))" by (simp add: sorted_wrt_append) lemma sorted_map: "sorted (map f xs) = sorted_wrt (\x y. f x \ f y) xs" by (simp add: sorted_wrt_map) lemma sorted01: "length xs \ 1 \ sorted xs" by (simp add: 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_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_iff_nth_Suc: "sorted xs \ (\i. Suc i < length xs \ xs!i \ xs!(Suc i))" by(simp add: sorted_wrt_iff_nth_Suc_transp) lemma sorted_rev_nth_mono: "sorted (rev xs) \ i \ j \ j < length xs \ xs!j \ xs!i" by (metis local.nle_le order_class.antisym_conv1 sorted_wrt_iff_nth_less sorted_wrt_rev) lemma sorted_rev_iff_nth_mono: "sorted (rev xs) \ (\ i j. i \ j \ j < length xs \ xs!j \ xs!i)" (is "?L = ?R") proof assume ?L thus ?R by (blast intro: sorted_rev_nth_mono) next assume ?R have "rev xs ! k \ rev xs ! l" if asms: "k \ l" "l < length(rev xs)" for k l proof - have "k < length xs" "l < length xs" "length xs - Suc l \ length xs - Suc k" "length xs - Suc k < length xs" using asms by auto thus "rev xs ! k \ rev xs ! l" by (simp add: \?R\ rev_nth) qed thus ?L by (simp add: sorted_iff_nth_mono) qed lemma sorted_rev_iff_nth_Suc: "sorted (rev xs) \ (\i. Suc i < length xs \ xs!(Suc i) \ xs!i)" proof- interpret dual: linorder "(\x y. y \ x)" "(\x y. y < x)" using dual_linorder . show ?thesis using dual_linorder dual.sorted_iff_nth_Suc dual.sorted_iff_nth_mono unfolding sorted_rev_iff_nth_mono by simp qed 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 "sorted xs" shows "sorted (butlast xs)" by (simp add: assms butlast_conv_take sorted_wrt_take) 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 x xs y ys) then show ?case by (cases \x = y\) (auto simp add: insert_eq_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" using assms map_inj_on sorted_distinct_set_unique by fastforce lemma sorted_dropWhile: "sorted xs \ sorted (dropWhile P xs)" by (auto dest: sorted_wrt_drop simp add: dropWhile_eq_drop) lemma sorted_takeWhile: "sorted xs \ sorted (takeWhile P xs)" by (subst takeWhile_eq_take) (auto dest: sorted_wrt_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: "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" by (metis hd_conv_nth hd_dropWhile length_greater_0_conv length_pos_if_in_set local.leI x) 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: order.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_insort_key: "distinct (map f (insort_key f x xs)) = (f x \ f ` set xs \ (distinct (map f 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 sort_key_id_if_sorted: "sorted (map f xs) \ sort_key f xs = xs" by (induction xs) (auto simp add: insort_is_Cons) text \Subsumed by @{thm sort_key_id_if_sorted} but easier to find:\ lemma sorted_sort_id: "sorted xs \ sort xs = xs" by (simp add: sort_key_id_if_sorted) 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_key [simp]: "remove1 x (insort_key f 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) 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) 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_key_list_of_set\\ text\ This function maps (finite) linearly ordered sets to sorted lists. The linear order is obtained by a key function that maps the elements of the set to a type that is linearly ordered. 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\). Note: this is a generalisation of the older \sorted_list_of_set\ that is obtained by setting the key function to the identity. Consequently, new theorems should be added to the locale below. They should also be aliased to more convenient names for use with \sorted_list_of_set\ as seen further below. \ definition (in linorder) sorted_key_list_of_set :: "('b \ 'a) \ 'b set \ 'b list" where "sorted_key_list_of_set f \ folding_on.F (insort_key f) []" locale folding_insort_key = lo?: linorder "less_eq :: 'a \ 'a \ bool" less for less_eq (infix "\" 50) and less (infix "\" 50) + fixes S fixes f :: "'b \ 'a" assumes inj_on: "inj_on f S" begin lemma insort_key_commute: "x \ S \ y \ S \ insort_key f y o insort_key f x = insort_key f x o insort_key f y" proof(rule ext, goal_cases) case (1 xs) with inj_on show ?case by (induction xs) (auto simp: inj_onD) qed sublocale fold_insort_key: folding_on S "insort_key f" "[]" rewrites "folding_on.F (insort_key f) [] = sorted_key_list_of_set f" proof - show "folding_on S (insort_key f)" by standard (simp add: insort_key_commute) qed (simp add: sorted_key_list_of_set_def) lemma idem_if_sorted_distinct: assumes "set xs \ S" and "sorted (map f xs)" "distinct xs" shows "sorted_key_list_of_set f (set xs) = xs" proof(cases "S = {}") case True then show ?thesis using \set xs \ S\ by auto next case False with assms show ?thesis proof(induction xs) case (Cons a xs) with Cons show ?case by (cases xs) auto qed simp qed lemma sorted_key_list_of_set_empty: "sorted_key_list_of_set f {} = []" by (fact fold_insort_key.empty) lemma sorted_key_list_of_set_insert: assumes "insert x A \ S" and "finite A" "x \ A" shows "sorted_key_list_of_set f (insert x A) = insort_key f x (sorted_key_list_of_set f A)" using assms by (fact fold_insort_key.insert) lemma sorted_key_list_of_set_insert_remove [simp]: assumes "insert x A \ S" and "finite A" shows "sorted_key_list_of_set f (insert x A) = insort_key f x (sorted_key_list_of_set f (A - {x}))" using assms by (fact fold_insort_key.insert_remove) lemma sorted_key_list_of_set_eq_Nil_iff [simp]: assumes "A \ S" and "finite A" shows "sorted_key_list_of_set f A = [] \ A = {}" using assms by (auto simp: fold_insort_key.remove) lemma set_sorted_key_list_of_set [simp]: assumes "A \ S" and "finite A" shows "set (sorted_key_list_of_set f A) = A" using assms(2,1) by (induct A rule: finite_induct) (simp_all add: set_insort_key) lemma sorted_sorted_key_list_of_set [simp]: assumes "A \ S" shows "sorted (map f (sorted_key_list_of_set f A))" proof (cases "finite A") case True thus ?thesis using \A \ S\ by (induction A) (simp_all add: sorted_insort_key) next case False thus ?thesis by simp qed lemma distinct_if_distinct_map: "distinct (map f xs) \ distinct xs" using inj_on by (simp add: distinct_map) lemma distinct_sorted_key_list_of_set [simp]: assumes "A \ S" shows "distinct (map f (sorted_key_list_of_set f A))" proof (cases "finite A") case True thus ?thesis using \A \ S\ inj_on by (induction A) (force simp: distinct_insort_key dest: inj_onD)+ next case False thus ?thesis by simp qed lemma length_sorted_key_list_of_set [simp]: assumes "A \ S" shows "length (sorted_key_list_of_set f A) = card A" proof (cases "finite A") case True with assms inj_on show ?thesis using distinct_card[symmetric, OF distinct_sorted_key_list_of_set] by (auto simp: subset_inj_on intro!: card_image) qed auto lemmas sorted_key_list_of_set = set_sorted_key_list_of_set sorted_sorted_key_list_of_set distinct_sorted_key_list_of_set lemma sorted_key_list_of_set_remove: assumes "insert x A \ S" and "finite A" shows "sorted_key_list_of_set f (A - {x}) = remove1 x (sorted_key_list_of_set f A)" proof (cases "x \ A") case False with assms have "x \ set (sorted_key_list_of_set f 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 lemma strict_sorted_key_list_of_set [simp]: "A \ S \ sorted_wrt (\) (map f (sorted_key_list_of_set f A))" by (cases "finite A") (auto simp: strict_sorted_iff subset_inj_on[OF inj_on]) lemma finite_set_strict_sorted: assumes "A \ S" and "finite A" obtains l where "sorted_wrt (\) (map f l)" "set l = A" "length l = card A" using assms by (meson length_sorted_key_list_of_set set_sorted_key_list_of_set strict_sorted_key_list_of_set) lemma (in linorder) strict_sorted_equal: assumes "sorted_wrt (<) xs" and "sorted_wrt (<) ys" and "set ys = set xs" shows "ys = xs" using assms proof (induction xs arbitrary: ys) case (Cons x xs) show ?case proof (cases ys) case Nil then show ?thesis using Cons.prems by auto next case (Cons y ys') then have "xs = ys'" by (metis Cons.prems list.inject sorted_distinct_set_unique strict_sorted_iff) moreover have "x = y" using Cons.prems \xs = ys'\ local.Cons by fastforce ultimately show ?thesis using local.Cons by blast qed qed auto lemma (in linorder) strict_sorted_equal_Uniq: "\\<^sub>\\<^sub>1xs. sorted_wrt (<) xs \ set xs = A" by (simp add: Uniq_def strict_sorted_equal) lemma sorted_key_list_of_set_inject: assumes "A \ S" "B \ S" assumes "sorted_key_list_of_set f A = sorted_key_list_of_set f B" "finite A" "finite B" shows "A = B" using assms set_sorted_key_list_of_set by metis lemma sorted_key_list_of_set_unique: assumes "A \ S" and "finite A" shows "sorted_wrt (\) (map f l) \ set l = A \ length l = card A \ sorted_key_list_of_set f A = l" using assms by (auto simp: strict_sorted_iff card_distinct idem_if_sorted_distinct) end context linorder begin definition "sorted_list_of_set \ sorted_key_list_of_set (\x::'a. x)" text \ We abuse the \rewrites\ functionality of locales to remove trivial assumptions that result from instantiating the key function to the identity. \ sublocale sorted_list_of_set: folding_insort_key "(\)" "(<)" UNIV "(\x. x)" rewrites "sorted_key_list_of_set (\x. x) = sorted_list_of_set" and "\xs. map (\x. x) xs \ xs" and "\X. (X \ UNIV) \ True" and "\x. x \ UNIV \ True" and "\P. (True \ P) \ Trueprop P" and "\P Q. (True \ PROP P \ PROP Q) \ (PROP P \ True \ PROP Q)" proof - show "folding_insort_key (\) (<) UNIV (\x. x)" by standard simp qed (simp_all add: sorted_list_of_set_def) text \Alias theorems for backwards compatibility and ease of use.\ lemmas sorted_list_of_set = sorted_list_of_set.sorted_key_list_of_set and sorted_list_of_set_empty = sorted_list_of_set.sorted_key_list_of_set_empty and sorted_list_of_set_insert = sorted_list_of_set.sorted_key_list_of_set_insert and sorted_list_of_set_insert_remove = sorted_list_of_set.sorted_key_list_of_set_insert_remove and sorted_list_of_set_eq_Nil_iff = sorted_list_of_set.sorted_key_list_of_set_eq_Nil_iff and set_sorted_list_of_set = sorted_list_of_set.set_sorted_key_list_of_set and sorted_sorted_list_of_set = sorted_list_of_set.sorted_sorted_key_list_of_set and distinct_sorted_list_of_set = sorted_list_of_set.distinct_sorted_key_list_of_set and length_sorted_list_of_set = sorted_list_of_set.length_sorted_key_list_of_set and sorted_list_of_set_remove = sorted_list_of_set.sorted_key_list_of_set_remove and strict_sorted_list_of_set = sorted_list_of_set.strict_sorted_key_list_of_set and sorted_list_of_set_inject = sorted_list_of_set.sorted_key_list_of_set_inject and sorted_list_of_set_unique = sorted_list_of_set.sorted_key_list_of_set_unique and finite_set_strict_sorted = sorted_list_of_set.finite_set_strict_sorted 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.fold_insort_key.eq_fold sort_conv_fold fold_set_fold_remdups) qed end lemma sorted_list_of_set_range [simp]: "sorted_list_of_set {m.. {}" shows "sorted_list_of_set A = Min A # sorted_list_of_set (A - {Min A})" using assms by (auto simp: less_le simp flip: sorted_list_of_set.sorted_key_list_of_set_unique intro: Min_in) lemma sorted_list_of_set_greaterThanLessThan: assumes "Suc i < j" shows "sorted_list_of_set {i<.. j" shows "sorted_list_of_set {i<..j} = Suc i # sorted_list_of_set {Suc i<..j}" using sorted_list_of_set_greaterThanLessThan [of i "Suc j"] by (metis assms greaterThanAtMost_def greaterThanLessThan_eq le_imp_less_Suc lessThan_Suc_atMost) lemma nth_sorted_list_of_set_greaterThanLessThan: "n < j - Suc i \ sorted_list_of_set {i<.. sorted_list_of_set {i<..j} ! n = Suc (i+n)" using nth_sorted_list_of_set_greaterThanLessThan [of n "Suc j" i] by (simp add: greaterThanAtMost_def greaterThanLessThan_eq lessThan_Suc_atMost) subsubsection \\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!]: assumes "wf r" shows "wf (lex r)" unfolding lex_def proof (rule wf_UN) show "wf (lexn r i)" for i by (simp add: assms wf_lexn) show "\i j. lexn r i \ lexn r j \ Domain (lexn r i) \ Range (lexn r j) = {}" by (metis DomainE Int_emptyI RangeE lexn_length) qed 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)}" proof (induction n) case (Suc n) then show ?case apply (simp add: image_Collect lex_prod_def, safe, blast) apply (rule_tac x = "ab # xys" in exI, simp) apply (case_tac xys; force) done qed auto text\By Mathias Fleury:\ proposition 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 corollary lex_transI: assumes "trans r" shows "trans (lex r)" using lexn_transI [OF assms] by (clarsimp simp add: lex_def trans_def) (metis lexn_length) 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 (auto simp add: lenlex_def Id_on_def lex_prod_def inv_image_def) lemma total_lenlex: assumes "total r" shows "total (lenlex r)" proof - have "(xs,ys) \ lexn r (length xs) \ (ys,xs) \ lexn r (length xs)" if "xs \ ys" and len: "length xs = length ys" for xs ys proof - obtain pre x xs' y ys' where "x\y" and xs: "xs = pre @ [x] @ xs'" and ys: "ys = pre @ [y] @ys'" by (meson len \xs \ ys\ same_length_different) then consider "(x,y) \ r" | "(y,x) \ r" by (meson UNIV_I assms total_on_def) then show ?thesis by cases (use len in \(force simp add: lexn_conv xs ys)+\) qed then show ?thesis by (fastforce simp: lenlex_def total_on_def lex_def) qed lemma lenlex_transI [intro]: "trans r \ trans (lenlex r)" unfolding lenlex_def by (meson lex_transI trans_inv_image trans_less_than trans_lex_prod) 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" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: lex_conv) (metis hd_append list.sel(1) list.sel(3) tl_append2) next assume ?rhs then show ?lhs by (simp add: lex_conv) (blast intro: Cons_eq_appendI) qed lemma Nil_lenlex_iff1 [simp]: "([], ns) \ lenlex r \ ns \ []" and Nil_lenlex_iff2 [simp]: "(ns,[]) \ lenlex r" by (auto simp: lenlex_def) lemma Cons_lenlex_iff: "((m # ms, n # ns) \ lenlex r) \ length ms < length ns \ length ms = length ns \ (m,n) \ r \ (m = n \ (ms,ns) \ lenlex r)" by (auto simp: lenlex_def) lemma lenlex_irreflexive: "(\x. (x,x) \ r) \ (xs,xs) \ lenlex r" by (induction xs) (auto simp add: Cons_lenlex_iff) lemma lenlex_trans: "\(x,y) \ lenlex r; (y,z) \ lenlex r; trans r\ \ (x,z) \ lenlex r" by (meson lenlex_transI transD) lemma lenlex_length: "(ms, ns) \ lenlex r \ length ms \ length ns" by (auto simp: lenlex_def) 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 lemma irrefl_lex: "irrefl r \ irrefl (lex r)" by (meson irrefl_def lex_take_index) lemma lexl_not_refl [simp]: "irrefl r \ (x,x) \ lex r" by (meson irrefl_def lex_take_index) 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)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: lexord_def) apply (metis hd_append list.sel(1) list.sel(3) tl_append2) done qed (auto simp add: lexord_def; (blast | meson Cons_eq_appendI)) lemmas lexord_simps = lexord_Nil_left lexord_Nil_right lexord_cons_cons lemma lexord_same_pref_iff: "(xs @ ys, xs @ zs) \ lexord r \ (\x \ set xs. (x,x) \ r) \ (ys, zs) \ lexord r" by(induction xs) auto lemma lexord_same_pref_if_irrefl[simp]: "irrefl r \ (xs @ ys, xs @ zs) \ lexord r \ (ys, zs) \ lexord r" by (simp add: irrefl_def lexord_same_pref_iff) lemma lexord_append_rightI: "\ b z. y = b # z \ (x, x @ y) \ lexord r" by (metis append_Nil2 lexord_Nil_left lexord_same_pref_iff) lemma lexord_append_left_rightI: "(a,b) \ r \ (u @ a # x, u @ b # y) \ lexord r" by (simp add: lexord_same_pref_iff) lemma lexord_append_leftI: "(u,v) \ lexord r \ (x @ u, x @ v) \ lexord r" by (simp add: lexord_same_pref_iff) lemma lexord_append_leftD: "\(x @ u, x @ v) \ lexord r; (\a. (a,a) \ r) \ \ (u,v) \ lexord r" by (simp add: lexord_same_pref_iff) 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))" proof - have "(\a v. y = x @ a # v) = (length x < length y \ take (length x) y = x)" by (metis Cons_nth_drop_Suc append_eq_conv_conj drop_all list.simps(3) not_le) moreover have "(\u a b. (a, b) \ r \ (\v. x = u @ a # v) \ (\w. y = u @ b # w)) = (\i take i x = take i y \ (x ! i, y ! i) \ r)" apply safe using less_iff_Suc_add apply auto[1] by (metis id_take_nth_drop) ultimately show ?thesis by (auto simp: lexord_def Let_def) qed \ \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_sufI: assumes "(u,w) \ lexord r" "length w \ length u" shows "(u@v,w@z) \ lexord r" proof- from leD[OF assms(2)] assms(1)[unfolded lexord_take_index_conv[of u w r] min_absorb2[OF assms(2)]] obtain i where "take i u = take i w" and "(u!i,w!i) \ r" and "i < length w" by blast hence "((u@v)!i, (w@z)!i) \ r" unfolding nth_append using less_le_trans[OF \i < length w\ assms(2)] \(u!i,w!i) \ r\ by presburger moreover have "i < min (length (u@v)) (length (w@z))" using assms(2) \i < length w\ by simp moreover have "take i (u@v) = take i (w@z)" using assms(2) \i < length w\ \take i u = take i w\ by simp ultimately show ?thesis using lexord_take_index_conv by blast qed lemma lexord_sufE: assumes "(xs@zs,ys@qs) \ lexord r" "xs \ ys" "length xs = length ys" "length zs = length qs" shows "(xs,ys) \ lexord r" proof- obtain i where "i < length (xs@zs)" and "i < length (ys@qs)" and "take i (xs@zs) = take i (ys@qs)" and "((xs@zs) ! i, (ys@qs) ! i) \ r" using assms(1) lex_take_index[unfolded lexord_lex,of "xs @ zs" "ys @ qs" r] length_append[of xs zs, unfolded assms(3,4), folded length_append[of ys qs]] by blast have "length (take i xs) = length (take i ys)" by (simp add: assms(3)) have "i < length xs" using assms(2,3) le_less_linear take_all[of xs i] take_all[of ys i] \take i (xs @ zs) = take i (ys @ qs)\ append_eq_append_conv take_append by metis hence "(xs ! i, ys ! i) \ r" using \((xs @ zs) ! i, (ys @ qs) ! i) \ r\ assms(3) by (simp add: nth_append) moreover have "take i xs = take i ys" using assms(3) \take i (xs @ zs) = take i (ys @ qs)\ by auto ultimately show ?thesis unfolding lexord_take_index_conv using \i < length xs\ assms(3) by fastforce qed 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 (meson lexord_trans transI) lemma total_lexord: "total r \ total (lexord r)" unfolding total_on_def proof clarsimp fix x y assume "\x y. x \ y \ (x, y) \ r \ (y, x) \ r" and "(x::'a list) \ y" and "(y, x) \ lexord r" then show "(x, y) \ 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 corollary lexord_linear: "(\a b. (a,b) \ r \ a = b \ (b,a) \ r) \ (x,y) \ lexord r \ x = y \ (y,x) \ lexord r" using total_lexord by (metis UNIV_I total_on_def) 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 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 dest: asymD) 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 (auto simp: hyp dest: asymD) qed lemma asym_lex: "asym R \ asym (lex R)" by (meson asymI asymD irrefl_lex lexord_asym lexord_lex) lemma asym_lenlex: "asym R \ asym (lenlex R)" by (simp add: lenlex_def asym_inv_image asym_less_than asym_lex asym_lex_prod) lemma lenlex_append1: assumes len: "(us,xs) \ lenlex R" and eq: "length vs = length ys" shows "(us @ vs, xs @ ys) \ lenlex R" using len proof (induction us) case Nil then show ?case by (simp add: lenlex_def eq) next case (Cons u us) with lex_append_rightI show ?case by (fastforce simp add: lenlex_def eq) qed lemma lenlex_append2 [simp]: assumes "irrefl R" shows "(us @ xs, us @ ys) \ lenlex R \ (xs, ys) \ lenlex R" proof (induction us) case Nil then show ?case by (simp add: lenlex_def) next case (Cons u us) with assms show ?case by (auto simp: lenlex_def irrefl_def) 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, code]: "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, code]: "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: "lexordp xs ys \ lexordp_eq xs ys" by (induction rule: lexordp.induct) simp_all lemma lexordp_eq_pref: "lexordp_eq u (u @ v)" by (metis append_Nil2 lexordp_append_rightI lexordp_eq_refl lexordp_into_lexordp_eq) end declare ord.lexordp_simps [simp, code] declare ord.lexordp_eq_simps [simp, code] 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 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)" proof (induction arbitrary: xs set: Wellfounded.acc) case outer: (1 u) show ?case proof (induct xs rule: acc_induct) case 1 show "xs \ Wellfounded.acc (listrel1 r)" by (simp add: outer.prems) qed (metis (no_types, lifting) Cons_listrel1E2 acc.simps outer.IH) qed 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)" proof (induction set: Wellfounded.acc) case (1 x) then have "\u v. \u \ set x; (v, u) \ r\ \ v \ Wellfounded.acc r" by (metis in_lists_conv_set in_set_conv_decomp listrel1I) then show ?case by (meson acc.intros in_listsI) qed 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: assumes "r \ A \ A" shows "listrel r \ lists A \ lists A" proof clarify show "a \ lists A \ b \ lists A" if "(a, b) \ listrel r" for a b using that assms by (induction rule: listrel.induct, auto) qed lemma listrel_refl_on: assumes "refl_on A r" shows "refl_on (lists A) (listrel r)" proof - have "l \ lists A \ (l, l) \ listrel r" for l using assms unfolding refl_on_def by (induction l, auto intro: listrel.intros) then show ?thesis by (meson assms listrel_subset refl_on_def) qed lemma listrel_sym: "sym r \ sym (listrel r)" by (simp add: listrel_iff_nth sym_def) lemma listrel_trans: assumes "trans r" shows "trans (listrel r)" proof - have "(x, z) \ listrel r" if "(x, y) \ listrel r" "(y, z) \ listrel r" for x y z using that proof induction case (Cons x y xs ys) then show ?case by clarsimp (metis assms listrel.Cons listrel_iff_nth transD) qed auto then show ?thesis using transI by blast qed 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>*" proof (induction rule: rtrancl.induct) case (rtrancl_into_rtrancl a b c) then show ?case by (metis listrel1I2 rtrancl.rtrancl_into_rtrancl) qed auto lemma rtrancl_listrel1_ConsI2: "(x,y) \ r\<^sup>* \ (xs, ys) \ (listrel1 r)\<^sup>* \ (x # xs, y # ys) \ (listrel1 r)\<^sup>*" by (meson in_mono listrel1I1 listrel1_rtrancl_subset_rtrancl_listrel1 rtrancl_listrel1_ConsI1 rtrancl_trans) 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_wrt.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" using le_less_Suc_eq by fastforce 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" by (smt (verit, best) atLeastAtMost_iff) 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)" by (simp add: length_remdups_card_conv) 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 successively_transfer [transfer_rule]: "((A ===> A ===> (=)) ===> list_all2 A ===> (=)) successively successively" unfolding successively_altdef 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 distinct_adj_transfer [transfer_rule]: assumes "bi_unique A" shows "(list_all2 A ===> (=)) distinct_adj distinct_adj" unfolding rel_fun_def proof (intro allI impI) fix xs ys assume "list_all2 A xs ys" thus "distinct_adj xs \ distinct_adj ys" proof (induction rule: list_all2_induct) case (Cons x xs y ys) show ?case by (metis Cons assms bi_unique_def distinct_adj_Cons list.rel_sel) qed auto qed 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" proof (rule rel_funI, rule rel_setI) show "\l \ lists X; rel_set A X Y\ \ \y\lists Y. list_all2 A l y" for X Y l proof (induction l rule: lists.induct) case (Cons a l) then show ?case by (simp only: rel_set_def list_all2_Cons1, metis lists.Cons) qed auto show "\l \ lists Y; rel_set A X Y\ \ \x\lists X. list_all2 A x l" for X Y l proof (induction l rule: lists.induct) case (Cons a l) then show ?case by (simp only: rel_set_def list_all2_Cons2, metis lists.Cons) qed auto qed 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/Nat.thy b/src/HOL/Nat.thy --- a/src/HOL/Nat.thy +++ b/src/HOL/Nat.thy @@ -1,2583 +1,2583 @@ (* Title: HOL/Nat.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Natural numbers\ theory Nat imports Inductive Typedef Fun Rings begin subsection \Type \ind\\ typedecl ind axiomatization Zero_Rep :: ind and Suc_Rep :: "ind \ ind" \ \The axiom of infinity in 2 parts:\ where Suc_Rep_inject: "Suc_Rep x = Suc_Rep y \ x = y" and Suc_Rep_not_Zero_Rep: "Suc_Rep x \ Zero_Rep" subsection \Type nat\ text \Type definition\ inductive Nat :: "ind \ bool" where Zero_RepI: "Nat Zero_Rep" | Suc_RepI: "Nat i \ Nat (Suc_Rep i)" typedef nat = "{n. Nat n}" morphisms Rep_Nat Abs_Nat using Nat.Zero_RepI by auto lemma Nat_Rep_Nat: "Nat (Rep_Nat n)" using Rep_Nat by simp lemma Nat_Abs_Nat_inverse: "Nat n \ Rep_Nat (Abs_Nat n) = n" using Abs_Nat_inverse by simp lemma Nat_Abs_Nat_inject: "Nat n \ Nat m \ Abs_Nat n = Abs_Nat m \ n = m" using Abs_Nat_inject by simp instantiation nat :: zero begin definition Zero_nat_def: "0 = Abs_Nat Zero_Rep" instance .. end definition Suc :: "nat \ nat" where "Suc n = Abs_Nat (Suc_Rep (Rep_Nat n))" lemma Suc_not_Zero: "Suc m \ 0" by (simp add: Zero_nat_def Suc_def Suc_RepI Zero_RepI Nat_Abs_Nat_inject Suc_Rep_not_Zero_Rep Nat_Rep_Nat) lemma Zero_not_Suc: "0 \ Suc m" by (rule not_sym) (rule Suc_not_Zero) lemma Suc_Rep_inject': "Suc_Rep x = Suc_Rep y \ x = y" by (rule iffI, rule Suc_Rep_inject) simp_all lemma nat_induct0: assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" proof - have "P (Abs_Nat (Rep_Nat n))" using assms unfolding Zero_nat_def Suc_def by (iprover intro: Nat_Rep_Nat [THEN Nat.induct] elim: Nat_Abs_Nat_inverse [THEN subst]) then show ?thesis by (simp add: Rep_Nat_inverse) qed free_constructors case_nat for "0 :: nat" | Suc pred where "pred (0 :: nat) = (0 :: nat)" proof atomize_elim fix n show "n = 0 \ (\m. n = Suc m)" by (induction n rule: nat_induct0) auto next fix n m show "(Suc n = Suc m) = (n = m)" by (simp add: Suc_def Nat_Abs_Nat_inject Nat_Rep_Nat Suc_RepI Suc_Rep_inject' Rep_Nat_inject) next fix n show "0 \ Suc n" by (simp add: Suc_not_Zero) qed \ \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "0 :: nat" Suc by (erule nat_induct0) auto setup \Sign.parent_path\ \ \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "nat"\ declare old.nat.inject[iff del] and old.nat.distinct(1)[simp del, induct_simp del] lemmas induct = old.nat.induct lemmas inducts = old.nat.inducts lemmas rec = old.nat.rec lemmas simps = nat.inject nat.distinct nat.case nat.rec setup \Sign.parent_path\ abbreviation rec_nat :: "'a \ (nat \ 'a \ 'a) \ nat \ 'a" where "rec_nat \ old.rec_nat" declare nat.sel[code del] hide_const (open) Nat.pred \ \hide everything related to the selector\ hide_fact nat.case_eq_if nat.collapse nat.expand nat.sel nat.exhaust_sel nat.split_sel nat.split_sel_asm lemma nat_exhaust [case_names 0 Suc, cases type: nat]: "(y = 0 \ P) \ (\nat. y = Suc nat \ P) \ P" \ \for backward compatibility -- names of variables differ\ by (rule old.nat.exhaust) lemma nat_induct [case_names 0 Suc, induct type: nat]: fixes n assumes "P 0" and "\n. P n \ P (Suc n)" shows "P n" \ \for backward compatibility -- names of variables differ\ using assms by (rule nat.induct) hide_fact nat_exhaust nat_induct0 ML \ val nat_basic_lfp_sugar = let val ctr_sugar = the (Ctr_Sugar.ctr_sugar_of_global \<^theory> \<^type_name>\nat\); val recx = Logic.varify_types_global \<^term>\rec_nat\; val C = body_type (fastype_of recx); in {T = HOLogic.natT, fp_res_index = 0, C = C, fun_arg_Tsss = [[], [[HOLogic.natT, C]]], ctr_sugar = ctr_sugar, recx = recx, rec_thms = @{thms nat.rec}} end; \ setup \ let fun basic_lfp_sugars_of _ [\<^typ>\nat\] _ _ ctxt = ([], [0], [nat_basic_lfp_sugar], [], [], [], TrueI (*dummy*), [], false, ctxt) | basic_lfp_sugars_of bs arg_Ts callers callssss ctxt = BNF_LFP_Rec_Sugar.default_basic_lfp_sugars_of bs arg_Ts callers callssss ctxt; in BNF_LFP_Rec_Sugar.register_lfp_rec_extension {nested_simps = [], special_endgame_tac = K (K (K (K no_tac))), is_new_datatype = K (K true), basic_lfp_sugars_of = basic_lfp_sugars_of, rewrite_nested_rec_call = NONE} end \ text \Injectiveness and distinctness lemmas\ lemma inj_Suc [simp]: "inj_on Suc N" by (simp add: inj_on_def) lemma bij_betw_Suc [simp]: "bij_betw Suc M N \ Suc ` M = N" by (simp add: bij_betw_def) lemma Suc_neq_Zero: "Suc m = 0 \ R" by (rule notE) (rule Suc_not_Zero) lemma Zero_neq_Suc: "0 = Suc m \ R" by (rule Suc_neq_Zero) (erule sym) lemma Suc_inject: "Suc x = Suc y \ x = y" by (rule inj_Suc [THEN injD]) lemma n_not_Suc_n: "n \ Suc n" by (induct n) simp_all lemma Suc_n_not_n: "Suc n \ n" by (rule not_sym) (rule n_not_Suc_n) text \A special form of induction for reasoning about \<^term>\m < n\ and \<^term>\m - n\.\ lemma diff_induct: assumes "\x. P x 0" and "\y. P 0 (Suc y)" and "\x y. P x y \ P (Suc x) (Suc y)" shows "P m n" proof (induct n arbitrary: m) case 0 show ?case by (rule assms(1)) next case (Suc n) show ?case proof (induct m) case 0 show ?case by (rule assms(2)) next case (Suc m) from \P m n\ show ?case by (rule assms(3)) qed qed subsection \Arithmetic operators\ instantiation nat :: comm_monoid_diff begin primrec plus_nat where add_0: "0 + n = (n::nat)" | add_Suc: "Suc m + n = Suc (m + n)" lemma add_0_right [simp]: "m + 0 = m" for m :: nat by (induct m) simp_all lemma add_Suc_right [simp]: "m + Suc n = Suc (m + n)" by (induct m) simp_all declare add_0 [code] lemma add_Suc_shift [code]: "Suc m + n = m + Suc n" by simp primrec minus_nat where diff_0 [code]: "m - 0 = (m::nat)" | diff_Suc: "m - Suc n = (case m - n of 0 \ 0 | Suc k \ k)" declare diff_Suc [simp del] lemma diff_0_eq_0 [simp, code]: "0 - n = 0" for n :: nat by (induct n) (simp_all add: diff_Suc) lemma diff_Suc_Suc [simp, code]: "Suc m - Suc n = m - n" by (induct n) (simp_all add: diff_Suc) instance proof fix n m q :: nat show "(n + m) + q = n + (m + q)" by (induct n) simp_all show "n + m = m + n" by (induct n) simp_all show "m + n - m = n" by (induct m) simp_all show "n - m - q = n - (m + q)" by (induct q) (simp_all add: diff_Suc) show "0 + n = n" by simp show "0 - n = 0" by simp qed end hide_fact (open) add_0 add_0_right diff_0 instantiation nat :: comm_semiring_1_cancel begin definition One_nat_def [simp]: "1 = Suc 0" primrec times_nat where mult_0: "0 * n = (0::nat)" | mult_Suc: "Suc m * n = n + (m * n)" lemma mult_0_right [simp]: "m * 0 = 0" for m :: nat by (induct m) simp_all lemma mult_Suc_right [simp]: "m * Suc n = m + (m * n)" by (induct m) (simp_all add: add.left_commute) lemma add_mult_distrib: "(m + n) * k = (m * k) + (n * k)" for m n k :: nat by (induct m) (simp_all add: add.assoc) instance proof fix k n m q :: nat show "0 \ (1::nat)" by simp show "1 * n = n" by simp show "n * m = m * n" by (induct n) simp_all show "(n * m) * q = n * (m * q)" by (induct n) (simp_all add: add_mult_distrib) show "(n + m) * q = n * q + m * q" by (rule add_mult_distrib) show "k * (m - n) = (k * m) - (k * n)" by (induct m n rule: diff_induct) simp_all qed end subsubsection \Addition\ text \Reasoning about \m + 0 = 0\, etc.\ lemma add_is_0 [iff]: "m + n = 0 \ m = 0 \ n = 0" for m n :: nat by (cases m) simp_all lemma add_is_1: "m + n = Suc 0 \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (cases m) simp_all lemma one_is_add: "Suc 0 = m + n \ m = Suc 0 \ n = 0 \ m = 0 \ n = Suc 0" by (rule trans, rule eq_commute, rule add_is_1) lemma add_eq_self_zero: "m + n = m \ n = 0" for m n :: nat by (induct m) simp_all lemma plus_1_eq_Suc: "plus 1 = Suc" by (simp add: fun_eq_iff) lemma Suc_eq_plus1: "Suc n = n + 1" by simp lemma Suc_eq_plus1_left: "Suc n = 1 + n" by simp subsubsection \Difference\ lemma Suc_diff_diff [simp]: "(Suc m - n) - Suc k = m - n - k" by (simp add: diff_diff_add) lemma diff_Suc_1 [simp]: "Suc n - 1 = n" by simp subsubsection \Multiplication\ lemma mult_is_0 [simp]: "m * n = 0 \ m = 0 \ n = 0" for m n :: nat by (induct m) auto lemma mult_eq_1_iff [simp]: "m * n = Suc 0 \ m = Suc 0 \ n = Suc 0" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (induct n) auto qed lemma one_eq_mult_iff [simp]: "Suc 0 = m * n \ m = Suc 0 \ n = Suc 0" by (simp add: eq_commute flip: mult_eq_1_iff) lemma nat_mult_eq_1_iff [simp]: "m * n = 1 \ m = 1 \ n = 1" and nat_1_eq_mult_iff [simp]: "1 = m * n \ m = 1 \ n = 1" for m n :: nat by auto lemma mult_cancel1 [simp]: "k * m = k * n \ m = n \ k = 0" for k m n :: nat proof - have "k \ 0 \ k * m = k * n \ m = n" proof (induct n arbitrary: m) case 0 then show "m = 0" by simp next case (Suc n) then show "m = Suc n" by (cases m) (simp_all add: eq_commute [of 0]) qed then show ?thesis by auto qed lemma mult_cancel2 [simp]: "m * k = n * k \ m = n \ k = 0" for k m n :: nat by (simp add: mult.commute) lemma Suc_mult_cancel1: "Suc k * m = Suc k * n \ m = n" by (subst mult_cancel1) simp subsection \Orders on \<^typ>\nat\\ subsubsection \Operation definition\ instantiation nat :: linorder begin primrec less_eq_nat where "(0::nat) \ n \ True" | "Suc m \ n \ (case n of 0 \ False | Suc n \ m \ n)" declare less_eq_nat.simps [simp del] lemma le0 [iff]: "0 \ n" for n :: nat by (simp add: less_eq_nat.simps) lemma [code]: "0 \ n \ True" for n :: nat by simp definition less_nat where less_eq_Suc_le: "n < m \ Suc n \ m" lemma Suc_le_mono [iff]: "Suc n \ Suc m \ n \ m" by (simp add: less_eq_nat.simps(2)) lemma Suc_le_eq [code]: "Suc m \ n \ m < n" unfolding less_eq_Suc_le .. lemma le_0_eq [iff]: "n \ 0 \ n = 0" for n :: nat by (induct n) (simp_all add: less_eq_nat.simps(2)) lemma not_less0 [iff]: "\ n < 0" for n :: nat by (simp add: less_eq_Suc_le) lemma less_nat_zero_code [code]: "n < 0 \ False" for n :: nat by simp lemma Suc_less_eq [iff]: "Suc m < Suc n \ m < n" by (simp add: less_eq_Suc_le) lemma less_Suc_eq_le [code]: "m < Suc n \ m \ n" by (simp add: less_eq_Suc_le) lemma Suc_less_eq2: "Suc n < m \ (\m'. m = Suc m' \ n < m')" by (cases m) auto lemma le_SucI: "m \ n \ m \ Suc n" by (induct m arbitrary: n) (simp_all add: less_eq_nat.simps(2) split: nat.splits) lemma Suc_leD: "Suc m \ n \ m \ n" by (cases n) (auto intro: le_SucI) lemma less_SucI: "m < n \ m < Suc n" by (simp add: less_eq_Suc_le) (erule Suc_leD) lemma Suc_lessD: "Suc m < n \ m < n" by (simp add: less_eq_Suc_le) (erule Suc_leD) instance proof fix n m q :: nat show "n < m \ n \ m \ \ m \ n" proof (induct n arbitrary: m) case 0 then show ?case by (cases m) (simp_all add: less_eq_Suc_le) next case (Suc n) then show ?case by (cases m) (simp_all add: less_eq_Suc_le) qed show "n \ n" by (induct n) simp_all then show "n = m" if "n \ m" and "m \ n" using that by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) show "n \ q" if "n \ m" and "m \ q" using that proof (induct n arbitrary: m q) case 0 show ?case by simp next case (Suc n) then show ?case by (simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits, clarify, simp_all (no_asm_use) add: less_eq_nat.simps(2) split: nat.splits) qed show "n \ m \ m \ n" by (induct n arbitrary: m) (simp_all add: less_eq_nat.simps(2) split: nat.splits) qed end instantiation nat :: order_bot begin definition bot_nat :: nat where "bot_nat = 0" instance by standard (simp add: bot_nat_def) end instance nat :: no_top by standard (auto intro: less_Suc_eq_le [THEN iffD2]) subsubsection \Introduction properties\ lemma lessI [iff]: "n < Suc n" by (simp add: less_Suc_eq_le) lemma zero_less_Suc [iff]: "0 < Suc n" by (simp add: less_Suc_eq_le) subsubsection \Elimination properties\ lemma less_not_refl: "\ n < n" for n :: nat by (rule order_less_irrefl) lemma less_not_refl2: "n < m \ m \ n" for m n :: nat by (rule not_sym) (rule less_imp_neq) lemma less_not_refl3: "s < t \ s \ t" for s t :: nat by (rule less_imp_neq) lemma less_irrefl_nat: "n < n \ R" for n :: nat by (rule notE, rule less_not_refl) lemma less_zeroE: "n < 0 \ R" for n :: nat by (rule notE) (rule not_less0) lemma less_Suc_eq: "m < Suc n \ m < n \ m = n" unfolding less_Suc_eq_le le_less .. lemma less_Suc0 [iff]: "(n < Suc 0) = (n = 0)" by (simp add: less_Suc_eq) lemma less_one [iff]: "n < 1 \ n = 0" for n :: nat unfolding One_nat_def by (rule less_Suc0) lemma Suc_mono: "m < n \ Suc m < Suc n" by simp text \"Less than" is antisymmetric, sort of.\ lemma less_antisym: "\ n < m \ n < Suc m \ m = n" unfolding not_less less_Suc_eq_le by (rule antisym) lemma nat_neq_iff: "m \ n \ m < n \ n < m" for m n :: nat by (rule linorder_neq_iff) subsubsection \Inductive (?) properties\ lemma Suc_lessI: "m < n \ Suc m \ n \ Suc m < n" unfolding less_eq_Suc_le [of m] le_less by simp lemma lessE: assumes major: "i < k" and 1: "k = Suc i \ P" and 2: "\j. i < j \ k = Suc j \ P" shows P proof - from major have "\j. i \ j \ k = Suc j" unfolding less_eq_Suc_le by (induct k) simp_all then have "(\j. i < j \ k = Suc j) \ k = Suc i" by (auto simp add: less_le) with 1 2 show P by auto qed lemma less_SucE: assumes major: "m < Suc n" and less: "m < n \ P" and eq: "m = n \ P" shows P proof (rule major [THEN lessE]) show "Suc n = Suc m \ P" using eq by blast show "\j. \m < j; Suc n = Suc j\ \ P" by (blast intro: less) qed lemma Suc_lessE: assumes major: "Suc i < k" and minor: "\j. i < j \ k = Suc j \ P" shows P proof (rule major [THEN lessE]) show "k = Suc (Suc i) \ P" using lessI minor by iprover show "\j. \Suc i < j; k = Suc j\ \ P" using Suc_lessD minor by iprover qed lemma Suc_less_SucD: "Suc m < Suc n \ m < n" by simp lemma less_trans_Suc: assumes le: "i < j" shows "j < k \ Suc i < k" proof (induct k) case 0 then show ?case by simp next case (Suc k) with le show ?case by simp (auto simp add: less_Suc_eq dest: Suc_lessD) qed text \Can be used with \less_Suc_eq\ to get \<^prop>\n = m \ n < m\.\ lemma not_less_eq: "\ m < n \ n < Suc m" by (simp only: not_less less_Suc_eq_le) lemma not_less_eq_eq: "\ m \ n \ Suc n \ m" by (simp only: not_le Suc_le_eq) text \Properties of "less than or equal".\ lemma le_imp_less_Suc: "m \ n \ m < Suc n" by (simp only: less_Suc_eq_le) lemma Suc_n_not_le_n: "\ Suc n \ n" by (simp add: not_le less_Suc_eq_le) lemma le_Suc_eq: "m \ Suc n \ m \ n \ m = Suc n" by (simp add: less_Suc_eq_le [symmetric] less_Suc_eq) lemma le_SucE: "m \ Suc n \ (m \ n \ R) \ (m = Suc n \ R) \ R" by (drule le_Suc_eq [THEN iffD1], iprover+) lemma Suc_leI: "m < n \ Suc m \ n" by (simp only: Suc_le_eq) text \Stronger version of \Suc_leD\.\ lemma Suc_le_lessD: "Suc m \ n \ m < n" by (simp only: Suc_le_eq) lemma less_imp_le_nat: "m < n \ m \ n" for m n :: nat unfolding less_eq_Suc_le by (rule Suc_leD) text \For instance, \(Suc m < Suc n) = (Suc m \ n) = (m < n)\\ lemmas le_simps = less_imp_le_nat less_Suc_eq_le Suc_le_eq text \Equivalence of \m \ n\ and \m < n \ m = n\\ lemma less_or_eq_imp_le: "m < n \ m = n \ m \ n" for m n :: nat unfolding le_less . lemma le_eq_less_or_eq: "m \ n \ m < n \ m = n" for m n :: nat by (rule le_less) text \Useful with \blast\.\ lemma eq_imp_le: "m = n \ m \ n" for m n :: nat by auto lemma le_refl: "n \ n" for n :: nat by simp lemma le_trans: "i \ j \ j \ k \ i \ k" for i j k :: nat by (rule order_trans) lemma le_antisym: "m \ n \ n \ m \ m = n" for m n :: nat by (rule antisym) lemma nat_less_le: "m < n \ m \ n \ m \ n" for m n :: nat by (rule less_le) lemma le_neq_implies_less: "m \ n \ m \ n \ m < n" for m n :: nat unfolding less_le .. lemma nat_le_linear: "m \ n \ n \ m" for m n :: nat by (rule linear) lemmas linorder_neqE_nat = linorder_neqE [where 'a = nat] lemma le_less_Suc_eq: "m \ n \ n < Suc m \ n = m" unfolding less_Suc_eq_le by auto lemma not_less_less_Suc_eq: "\ n < m \ n < Suc m \ n = m" unfolding not_less by (rule le_less_Suc_eq) lemmas not_less_simps = not_less_less_Suc_eq le_less_Suc_eq lemma not0_implies_Suc: "n \ 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr0_implies_Suc: "n > 0 \ \m. n = Suc m" by (cases n) simp_all lemma gr_implies_not0: "m < n \ n \ 0" for m n :: nat by (cases n) simp_all lemma neq0_conv[iff]: "n \ 0 \ 0 < n" for n :: nat by (cases n) simp_all text \This theorem is useful with \blast\\ lemma gr0I: "(n = 0 \ False) \ 0 < n" for n :: nat by (rule neq0_conv[THEN iffD1]) iprover lemma gr0_conv_Suc: "0 < n \ (\m. n = Suc m)" by (fast intro: not0_implies_Suc) lemma not_gr0 [iff]: "\ 0 < n \ n = 0" for n :: nat using neq0_conv by blast lemma Suc_le_D: "Suc n \ m' \ \m. m' = Suc m" by (induct m') simp_all text \Useful in certain inductive arguments\ lemma less_Suc_eq_0_disj: "m < Suc n \ m = 0 \ (\j. m = Suc j \ j < n)" by (cases m) simp_all lemma All_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma All_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) lemma Ex_less_Suc: "(\i < Suc n. P i) = (P n \ (\i < n. P i))" by (auto simp: less_Suc_eq) lemma Ex_less_Suc2: "(\i < Suc n. P i) = (P 0 \ (\i < n. P(Suc i)))" by (auto simp: less_Suc_eq_0_disj) text \@{term mono} (non-strict) doesn't imply increasing, as the function could be constant\ lemma strict_mono_imp_increasing: fixes n::nat assumes "strict_mono f" shows "f n \ n" proof (induction n) case 0 then show ?case by auto next case (Suc n) then show ?case unfolding not_less_eq_eq [symmetric] using Suc_n_not_le_n assms order_trans strict_mono_less_eq by blast qed subsubsection \Monotonicity of Addition\ lemma Suc_pred [simp]: "n > 0 \ Suc (n - Suc 0) = n" by (simp add: diff_Suc split: nat.split) lemma Suc_diff_1 [simp]: "0 < n \ Suc (n - 1) = n" unfolding One_nat_def by (rule Suc_pred) lemma nat_add_left_cancel_le [simp]: "k + m \ k + n \ m \ n" for k m n :: nat by (induct k) simp_all lemma nat_add_left_cancel_less [simp]: "k + m < k + n \ m < n" for k m n :: nat by (induct k) simp_all lemma add_gr_0 [iff]: "m + n > 0 \ m > 0 \ n > 0" for m n :: nat by (auto dest: gr0_implies_Suc) text \strict, in 1st argument\ lemma add_less_mono1: "i < j \ i + k < j + k" for i j k :: nat by (induct k) simp_all text \strict, in both arguments\ lemma add_less_mono: fixes i j k l :: nat assumes "i < j" "k < l" shows "i + k < j + l" proof - have "i + k < j + k" by (simp add: add_less_mono1 assms) also have "... < j + l" using \i < j\ by (induction j) (auto simp: assms) finally show ?thesis . qed lemma less_imp_Suc_add: "m < n \ \k. n = Suc (m + k)" proof (induct n) case 0 then show ?case by simp next case Suc then show ?case by (simp add: order_le_less) (blast elim!: less_SucE intro!: Nat.add_0_right [symmetric] add_Suc_right [symmetric]) qed lemma le_Suc_ex: "k \ l \ (\n. l = k + n)" for k l :: nat by (auto simp: less_Suc_eq_le[symmetric] dest: less_imp_Suc_add) lemma less_natE: assumes \m < n\ obtains q where \n = Suc (m + q)\ using assms by (auto dest: less_imp_Suc_add intro: that) text \strict, in 1st argument; proof is by induction on \k > 0\\ lemma mult_less_mono2: fixes i j :: nat assumes "i < j" and "0 < k" shows "k * i < k * j" using \0 < k\ proof (induct k) case 0 then show ?case by simp next case (Suc k) with \i < j\ show ?case by (cases k) (simp_all add: add_less_mono) qed text \Addition is the inverse of subtraction: if \<^term>\n \ m\ then \<^term>\n + (m - n) = m\.\ lemma add_diff_inverse_nat: "\ m < n \ n + (m - n) = m" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma nat_le_iff_add: "m \ n \ (\k. n = m + k)" for m n :: nat using nat_add_left_cancel_le[of m 0] by (auto dest: le_Suc_ex) text \The naturals form an ordered \semidom\ and a \dioid\.\ instance nat :: linordered_semidom proof fix m n q :: nat show "0 < (1::nat)" by simp show "m \ n \ q + m \ q + n" by simp show "m < n \ 0 < q \ q * m < q * n" by (simp add: mult_less_mono2) show "m \ 0 \ n \ 0 \ m * n \ 0" by simp show "n \ m \ (m - n) + n = m" by (simp add: add_diff_inverse_nat add.commute linorder_not_less) qed instance nat :: dioid by standard (rule nat_le_iff_add) declare le0[simp del] \ \This is now @{thm zero_le}\ declare le_0_eq[simp del] \ \This is now @{thm le_zero_eq}\ declare not_less0[simp del] \ \This is now @{thm not_less_zero}\ declare not_gr0[simp del] \ \This is now @{thm not_gr_zero}\ instance nat :: ordered_cancel_comm_monoid_add .. instance nat :: ordered_cancel_comm_monoid_diff .. subsubsection \\<^term>\min\ and \<^term>\max\\ global_interpretation bot_nat_0: ordering_top \(\)\ \(>)\ \0::nat\ by standard simp global_interpretation max_nat: semilattice_neutr_order max \0::nat\ \(\)\ \(>)\ by standard (simp add: max_def) lemma mono_Suc: "mono Suc" by (rule monoI) simp lemma min_0L [simp]: "min 0 n = 0" for n :: nat by (rule min_absorb1) simp lemma min_0R [simp]: "min n 0 = 0" for n :: nat by (rule min_absorb2) simp lemma min_Suc_Suc [simp]: "min (Suc m) (Suc n) = Suc (min m n)" by (simp add: mono_Suc min_of_mono) lemma min_Suc1: "min (Suc n) m = (case m of 0 \ 0 | Suc m' \ Suc(min n m'))" by (simp split: nat.split) lemma min_Suc2: "min m (Suc n) = (case m of 0 \ 0 | Suc m' \ Suc(min m' n))" by (simp split: nat.split) lemma max_0L [simp]: "max 0 n = n" for n :: nat by (fact max_nat.left_neutral) lemma max_0R [simp]: "max n 0 = n" for n :: nat by (fact max_nat.right_neutral) lemma max_Suc_Suc [simp]: "max (Suc m) (Suc n) = Suc (max m n)" by (simp add: mono_Suc max_of_mono) lemma max_Suc1: "max (Suc n) m = (case m of 0 \ Suc n | Suc m' \ Suc (max n m'))" by (simp split: nat.split) lemma max_Suc2: "max m (Suc n) = (case m of 0 \ Suc n | Suc m' \ Suc (max m' n))" by (simp split: nat.split) lemma nat_mult_min_left: "min m n * q = min (m * q) (n * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_min_right: "m * min n q = min (m * n) (m * q)" for m n q :: nat by (simp add: min_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) lemma nat_add_max_left: "max m n + q = max (m + q) (n + q)" for m n q :: nat by (simp add: max_def) lemma nat_add_max_right: "m + max n q = max (m + n) (m + q)" for m n q :: nat by (simp add: max_def) lemma nat_mult_max_left: "max m n * q = max (m * q) (n * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_right_le_imp_le mult_right_less_imp_less le_less_trans) lemma nat_mult_max_right: "m * max n q = max (m * n) (m * q)" for m n q :: nat by (simp add: max_def not_le) (auto dest: mult_left_le_imp_le mult_left_less_imp_less le_less_trans) subsubsection \Additional theorems about \<^term>\(\)\\ text \Complete induction, aka course-of-values induction\ instance nat :: wellorder proof fix P and n :: nat assume step: "(\m. m < n \ P m) \ P n" for n :: nat have "\q. q \ n \ P q" proof (induct n) case (0 n) have "P 0" by (rule step) auto with 0 show ?case by auto next case (Suc m n) then have "n \ m \ n = Suc m" by (simp add: le_Suc_eq) then show ?case proof assume "n \ m" then show "P n" by (rule Suc(1)) next assume n: "n = Suc m" show "P n" by (rule step) (rule Suc(1), simp add: n le_simps) qed qed then show "P n" by auto qed lemma Least_eq_0[simp]: "P 0 \ Least P = 0" for P :: "nat \ bool" by (rule Least_equality[OF _ le0]) lemma Least_Suc: assumes "P n" "\ P 0" shows "(LEAST n. P n) = Suc (LEAST m. P (Suc m))" proof (cases n) case (Suc m) show ?thesis proof (rule antisym) show "(LEAST x. P x) \ Suc (LEAST x. P (Suc x))" using assms Suc by (force intro: LeastI Least_le) have \
: "P (LEAST x. P x)" by (blast intro: LeastI assms) show "Suc (LEAST m. P (Suc m)) \ (LEAST n. P n)" proof (cases "(LEAST n. P n)") case 0 then show ?thesis using \
by (simp add: assms) next case Suc with \
show ?thesis by (auto simp: Least_le) qed qed qed (use assms in auto) lemma Least_Suc2: "P n \ Q m \ \ P 0 \ \k. P (Suc k) = Q k \ Least P = Suc (Least Q)" by (erule (1) Least_Suc [THEN ssubst]) simp lemma ex_least_nat_le: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\k\n. (\i P i) \ P k" proof (cases n) case (Suc m) with assms show ?thesis by (blast intro: Least_le LeastI_ex dest: not_less_Least) qed (use assms in auto) lemma ex_least_nat_less: fixes P :: "nat \ bool" assumes "P n" "\ P 0" shows "\ki\k. \ P i) \ P (Suc k)" proof (cases n) case (Suc m) then obtain k where k: "k \ n" "\i P i" "P k" using ex_least_nat_le [OF assms] by blast show ?thesis by (cases k) (use assms k less_eq_Suc_le in auto) qed (use assms in auto) lemma nat_less_induct: fixes P :: "nat \ bool" assumes "\n. \m. m < n \ P m \ P n" shows "P n" using assms less_induct by blast lemma measure_induct_rule [case_names less]: fixes f :: "'a \ 'b::wellorder" assumes step: "\x. (\y. f y < f x \ P y) \ P x" shows "P a" by (induct m \ "f a" arbitrary: a rule: less_induct) (auto intro: step) text \old style induction rules:\ lemma measure_induct: fixes f :: "'a \ 'b::wellorder" shows "(\x. \y. f y < f x \ P y \ P x) \ P a" by (rule measure_induct_rule [of f P a]) iprover lemma full_nat_induct: assumes step: "\n. (\m. Suc m \ n \ P m) \ P n" shows "P n" by (rule less_induct) (auto intro: step simp:le_simps) text\An induction rule for establishing binary relations\ lemma less_Suc_induct [consumes 1]: assumes less: "i < j" and step: "\i. P i (Suc i)" and trans: "\i j k. i < j \ j < k \ P i j \ P j k \ P i k" shows "P i j" proof - from less obtain k where j: "j = Suc (i + k)" by (auto dest: less_imp_Suc_add) have "P i (Suc (i + k))" proof (induct k) case 0 show ?case by (simp add: step) next case (Suc k) have "0 + i < Suc k + i" by (rule add_less_mono1) simp then have "i < Suc (i + k)" by (simp add: add.commute) from trans[OF this lessI Suc step] show ?case by simp qed then show "P i j" by (simp add: j) qed text \ The method of infinite descent, frequently used in number theory. Provided by Roelof Oosterhuis. \P n\ is true for all natural numbers if \<^item> case ``0'': given \n = 0\ prove \P n\ \<^item> case ``smaller'': given \n > 0\ and \\ P n\ prove there exists a smaller natural number \m\ such that \\ P m\. \ lemma infinite_descent: "(\n. \ P n \ \m P m) \ P n" for P :: "nat \ bool" \ \compact version without explicit base case\ by (induct n rule: less_induct) auto lemma infinite_descent0 [case_names 0 smaller]: fixes P :: "nat \ bool" assumes "P 0" and "\n. n > 0 \ \ P n \ \m. m < n \ \ P m" shows "P n" proof (rule infinite_descent) fix n show "\ P n \ \m P m" using assms by (cases "n > 0") auto qed text \ Infinite descent using a mapping to \nat\: \P x\ is true for all \x \ D\ if there exists a \V \ D \ nat\ and \<^item> case ``0'': given \V x = 0\ prove \P x\ \<^item> ``smaller'': given \V x > 0\ and \\ P x\ prove there exists a \y \ D\ such that \V y < V x\ and \\ P y\. \ corollary infinite_descent0_measure [case_names 0 smaller]: fixes V :: "'a \ nat" assumes 1: "\x. V x = 0 \ P x" and 2: "\x. V x > 0 \ \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof (induct n rule: infinite_descent0) case 0 with 1 show "P x" by auto next case (smaller n) then obtain x where *: "V x = n " and "V x > 0 \ \ P x" by auto with 2 obtain y where "V y < V x \ \ P y" by auto with * obtain m where "m = V y \ m < n \ \ P y" by auto then show ?case by auto qed ultimately show "P x" by auto qed text \Again, without explicit base case:\ lemma infinite_descent_measure: fixes V :: "'a \ nat" assumes "\x. \ P x \ \y. V y < V x \ \ P y" shows "P x" proof - from assms obtain n where "n = V x" by auto moreover have "\x. V x = n \ P x" proof - have "\m < V x. \y. V y = m \ \ P y" if "\ P x" for x using assms and that by auto then show "\x. V x = n \ P x" by (induct n rule: infinite_descent, auto) qed ultimately show "P x" by auto qed text \A (clumsy) way of lifting \<\ monotonicity to \\\ monotonicity\ lemma less_mono_imp_le_mono: fixes f :: "nat \ nat" and i j :: nat assumes "\i j::nat. i < j \ f i < f j" and "i \ j" shows "f i \ f j" using assms by (auto simp add: order_le_less) text \non-strict, in 1st argument\ lemma add_le_mono1: "i \ j \ i + k \ j + k" for i j k :: nat by (rule add_right_mono) text \non-strict, in both arguments\ lemma add_le_mono: "i \ j \ k \ l \ i + k \ j + l" for i j k l :: nat by (rule add_mono) lemma le_add2: "n \ m + n" for m n :: nat by simp lemma le_add1: "n \ n + m" for m n :: nat by simp lemma less_add_Suc1: "i < Suc (i + m)" by (rule le_less_trans, rule le_add1, rule lessI) lemma less_add_Suc2: "i < Suc (m + i)" by (rule le_less_trans, rule le_add2, rule lessI) lemma less_iff_Suc_add: "m < n \ (\k. n = Suc (m + k))" by (iprover intro!: less_add_Suc1 less_imp_Suc_add) lemma trans_le_add1: "i \ j \ i \ j + m" for i j m :: nat by (rule le_trans, assumption, rule le_add1) lemma trans_le_add2: "i \ j \ i \ m + j" for i j m :: nat by (rule le_trans, assumption, rule le_add2) lemma trans_less_add1: "i < j \ i < j + m" for i j m :: nat by (rule less_le_trans, assumption, rule le_add1) lemma trans_less_add2: "i < j \ i < m + j" for i j m :: nat by (rule less_le_trans, assumption, rule le_add2) lemma add_lessD1: "i + j < k \ i < k" for i j k :: nat by (rule le_less_trans [of _ "i+j"]) (simp_all add: le_add1) lemma not_add_less1 [iff]: "\ i + j < i" for i j :: nat by simp lemma not_add_less2 [iff]: "\ j + i < i" for i j :: nat by simp lemma add_leD1: "m + k \ n \ m \ n" for k m n :: nat by (rule order_trans [of _ "m + k"]) (simp_all add: le_add1) lemma add_leD2: "m + k \ n \ k \ n" for k m n :: nat by (force simp add: add.commute dest: add_leD1) lemma add_leE: "m + k \ n \ (m \ n \ k \ n \ R) \ R" for k m n :: nat by (blast dest: add_leD1 add_leD2) text \needs \\k\ for \ac_simps\ to work\ lemma less_add_eq_less: "\k. k < l \ m + l = k + n \ m < n" for l m n :: nat by (force simp del: add_Suc_right simp add: less_iff_Suc_add add_Suc_right [symmetric] ac_simps) subsubsection \More results about difference\ lemma Suc_diff_le: "n \ m \ Suc m - n = Suc (m - n)" by (induct m n rule: diff_induct) simp_all lemma diff_less_Suc: "m - n < Suc m" by (induct m n rule: diff_induct) (auto simp: less_Suc_eq) lemma diff_le_self [simp]: "m - n \ m" for m n :: nat by (induct m n rule: diff_induct) (simp_all add: le_SucI) lemma less_imp_diff_less: "j < k \ j - n < k" for j k n :: nat by (rule le_less_trans, rule diff_le_self) lemma diff_Suc_less [simp]: "0 < n \ n - Suc i < n" by (cases n) (auto simp add: le_simps) lemma diff_add_assoc: "k \ j \ (i + j) - k = i + (j - k)" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc) lemma add_diff_assoc [simp]: "k \ j \ i + (j - k) = i + j - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc) lemma diff_add_assoc2: "k \ j \ (j + i) - k = (j - k) + i" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.diff_add_assoc2) lemma add_diff_assoc2 [simp]: "k \ j \ j - k + i = j + i - k" for i j k :: nat by (fact ordered_cancel_comm_monoid_diff_class.add_diff_assoc2) lemma le_imp_diff_is_add: "i \ j \ (j - i = k) = (j = k + i)" for i j k :: nat by auto lemma diff_is_0_eq [simp]: "m - n = 0 \ m \ n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma diff_is_0_eq' [simp]: "m \ n \ m - n = 0" for m n :: nat by (rule iffD2, rule diff_is_0_eq) lemma zero_less_diff [simp]: "0 < n - m \ m < n" for m n :: nat by (induct m n rule: diff_induct) simp_all lemma less_imp_add_positive: assumes "i < j" shows "\k::nat. 0 < k \ i + k = j" proof from assms show "0 < j - i \ i + (j - i) = j" by (simp add: order_less_imp_le) qed text \a nice rewrite for bounded subtraction\ lemma nat_minus_add_max: "n - m + m = max n m" for m n :: nat by (simp add: max_def not_le order_less_imp_le) lemma nat_diff_split: "P (a - b) \ (a < b \ P 0) \ (\d. a = b + d \ P d)" for a b :: nat \ \elimination of \-\ on \nat\\ by (cases "a < b") (auto simp add: not_less le_less dest!: add_eq_self_zero [OF sym]) lemma nat_diff_split_asm: "P (a - b) \ \ (a < b \ \ P 0 \ (\d. a = b + d \ \ P d))" for a b :: nat \ \elimination of \-\ on \nat\ in assumptions\ by (auto split: nat_diff_split) lemma Suc_pred': "0 < n \ n = Suc(n - 1)" by simp lemma add_eq_if: "m + n = (if m = 0 then n else Suc ((m - 1) + n))" unfolding One_nat_def by (cases m) simp_all lemma mult_eq_if: "m * n = (if m = 0 then 0 else n + ((m - 1) * n))" for m n :: nat by (cases m) simp_all lemma Suc_diff_eq_diff_pred: "0 < n \ Suc m - n = m - (n - 1)" by (cases n) simp_all lemma diff_Suc_eq_diff_pred: "m - Suc n = (m - 1) - n" by (cases m) simp_all lemma Let_Suc [simp]: "Let (Suc n) f \ f (Suc n)" by (fact Let_def) subsubsection \Monotonicity of multiplication\ lemma mult_le_mono1: "i \ j \ i * k \ j * k" for i j k :: nat by (simp add: mult_right_mono) lemma mult_le_mono2: "i \ j \ k * i \ k * j" for i j k :: nat by (simp add: mult_left_mono) text \\\\ monotonicity, BOTH arguments\ lemma mult_le_mono: "i \ j \ k \ l \ i * k \ j * l" for i j k l :: nat by (simp add: mult_mono) lemma mult_less_mono1: "i < j \ 0 < k \ i * k < j * k" for i j k :: nat by (simp add: mult_strict_right_mono) text \Differs from the standard \zero_less_mult_iff\ in that there are no negative numbers.\ lemma nat_0_less_mult_iff [simp]: "0 < m * n \ 0 < m \ 0 < n" for m n :: nat proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma one_le_mult_iff [simp]: "Suc 0 \ m * n \ Suc 0 \ m \ Suc 0 \ n" proof (induct m) case 0 then show ?case by simp next case (Suc m) then show ?case by (cases n) simp_all qed lemma mult_less_cancel2 [simp]: "m * k < n * k \ 0 < k \ m < n" for k m n :: nat proof (intro iffI conjI) assume m: "m * k < n * k" then show "0 < k" by (cases k) auto show "m < n" proof (cases k) case 0 then show ?thesis using m by auto next case (Suc k') then show ?thesis using m by (simp flip: linorder_not_le) (blast intro: add_mono mult_le_mono1) qed next assume "0 < k \ m < n" then show "m * k < n * k" by (blast intro: mult_less_mono1) qed lemma mult_less_cancel1 [simp]: "k * m < k * n \ 0 < k \ m < n" for k m n :: nat by (simp add: mult.commute [of k]) lemma mult_le_cancel1 [simp]: "k * m \ k * n \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma mult_le_cancel2 [simp]: "m * k \ n * k \ (0 < k \ m \ n)" for k m n :: nat by (simp add: linorder_not_less [symmetric], auto) lemma Suc_mult_less_cancel1: "Suc k * m < Suc k * n \ m < n" by (subst mult_less_cancel1) simp lemma Suc_mult_le_cancel1: "Suc k * m \ Suc k * n \ m \ n" by (subst mult_le_cancel1) simp lemma le_square: "m \ m * m" for m :: nat by (cases m) (auto intro: le_add1) lemma le_cube: "m \ m * (m * m)" for m :: nat by (cases m) (auto intro: le_add1) text \Lemma for \gcd\\ lemma mult_eq_self_implies_10: fixes m n :: nat assumes "m = m * n" shows "n = 1 \ m = 0" proof (rule disjCI) assume "m \ 0" show "n = 1" proof (cases n "1::nat" rule: linorder_cases) case greater show ?thesis using assms mult_less_mono2 [OF greater, of m] \m \ 0\ by auto qed (use assms \m \ 0\ in auto) qed lemma mono_times_nat: fixes n :: nat assumes "n > 0" shows "mono (times n)" proof fix m q :: nat assume "m \ q" with assms show "n * m \ n * q" by simp qed text \The lattice order on \<^typ>\nat\.\ instantiation nat :: distrib_lattice begin definition "(inf :: nat \ nat \ nat) = min" definition "(sup :: nat \ nat \ nat) = max" instance by intro_classes (auto simp add: inf_nat_def sup_nat_def max_def not_le min_def intro: order_less_imp_le antisym elim!: order_trans order_less_trans) end subsection \Natural operation of natural numbers on functions\ text \ We use the same logical constant for the power operations on functions and relations, in order to share the same syntax. \ consts compow :: "nat \ 'a \ 'a" abbreviation compower :: "'a \ nat \ 'a" (infixr "^^" 80) where "f ^^ n \ compow n f" notation (latex output) compower ("(_\<^bsup>_\<^esup>)" [1000] 1000) text \\f ^^ n = f \ \ \ f\, the \n\-fold composition of \f\\ overloading funpow \ "compow :: nat \ ('a \ 'a) \ ('a \ 'a)" begin primrec funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where "funpow 0 f = id" | "funpow (Suc n) f = f \ funpow n f" end lemma funpow_0 [simp]: "(f ^^ 0) x = x" by simp lemma funpow_Suc_right: "f ^^ Suc n = f ^^ n \ f" proof (induct n) case 0 then show ?case by simp next fix n assume "f ^^ Suc n = f ^^ n \ f" then show "f ^^ Suc (Suc n) = f ^^ Suc n \ f" by (simp add: o_assoc) qed lemmas funpow_simps_right = funpow.simps(1) funpow_Suc_right text \For code generation.\ context begin qualified definition funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" where funpow_code_def [code_abbrev]: "funpow = compow" lemma [code]: "funpow (Suc n) f = f \ funpow n f" "funpow 0 f = id" by (simp_all add: funpow_code_def) end lemma funpow_add: "f ^^ (m + n) = f ^^ m \ f ^^ n" by (induct m) simp_all lemma funpow_mult: "(f ^^ m) ^^ n = f ^^ (m * n)" for f :: "'a \ 'a" by (induct n) (simp_all add: funpow_add) lemma funpow_swap1: "f ((f ^^ n) x) = (f ^^ n) (f x)" proof - have "f ((f ^^ n) x) = (f ^^ (n + 1)) x" by simp also have "\ = (f ^^ n \ f ^^ 1) x" by (simp only: funpow_add) also have "\ = (f ^^ n) (f x)" by simp finally show ?thesis . qed lemma comp_funpow: "comp f ^^ n = comp (f ^^ n)" for f :: "'a \ 'a" by (induct n) simp_all lemma Suc_funpow[simp]: "Suc ^^ n = ((+) n)" by (induct n) simp_all lemma id_funpow[simp]: "id ^^ n = id" by (induct n) simp_all lemma funpow_mono: "mono f \ A \ B \ (f ^^ n) A \ (f ^^ n) B" for f :: "'a \ ('a::order)" by (induct n arbitrary: A B) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right mono_def) lemma funpow_mono2: assumes "mono f" and "i \ j" and "x \ y" and "x \ f x" shows "(f ^^ i) x \ (f ^^ j) y" using assms(2,3) proof (induct j arbitrary: y) case 0 then show ?case by simp next case (Suc j) show ?case proof(cases "i = Suc j") case True with assms(1) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right monoD funpow_mono) next case False with assms(1,4) Suc show ?thesis by (simp del: funpow.simps add: funpow_simps_right le_eq_less_or_eq less_Suc_eq_le) (simp add: Suc.hyps monoD order_subst1) qed qed lemma inj_fn[simp]: fixes f::"'a \ 'a" assumes "inj f" shows "inj (f^^n)" proof (induction n) case Suc thus ?case using inj_compose[OF assms Suc.IH] by (simp del: comp_apply) qed simp lemma surj_fn[simp]: fixes f::"'a \ 'a" assumes "surj f" shows "surj (f^^n)" proof (induction n) case Suc thus ?case by (simp add: comp_surj[OF Suc.IH assms] del: comp_apply) qed simp lemma bij_fn[simp]: fixes f::"'a \ 'a" assumes "bij f" shows "bij (f^^n)" by (rule bijI[OF inj_fn[OF bij_is_inj[OF assms]] surj_fn[OF bij_is_surj[OF assms]]]) lemma bij_betw_funpow: \<^marker>\contributor \Lars Noschinski\\ assumes "bij_betw f S S" shows "bij_betw (f ^^ n) S S" proof (induct n) case 0 then show ?case by (auto simp: id_def[symmetric]) next case (Suc n) then show ?case unfolding funpow.simps using assms by (rule bij_betw_trans) qed subsection \Kleene iteration\ lemma Kleene_iter_lpfp: fixes f :: "'a::order_bot \ 'a" assumes "mono f" and "f p \ p" shows "(f ^^ k) bot \ p" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma lfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) bot = (f ^^ k) bot" shows "lfp f = (f ^^ k) bot" proof (rule antisym) show "lfp f \ (f ^^ k) bot" proof (rule lfp_lowerbound) show "f ((f ^^ k) bot) \ (f ^^ k) bot" using assms(2) by simp qed show "(f ^^ k) bot \ lfp f" using Kleene_iter_lpfp[OF assms(1)] lfp_unfold[OF assms(1)] by simp qed lemma mono_pow: "mono f \ mono (f ^^ n)" for f :: "'a \ 'a::complete_lattice" by (induct n) (auto simp: mono_def) lemma lfp_funpow: assumes f: "mono f" shows "lfp (f ^^ Suc n) = lfp f" proof (rule antisym) show "lfp f \ lfp (f ^^ Suc n)" proof (rule lfp_lowerbound) have "f (lfp (f ^^ Suc n)) = lfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: lfp_rolling f mono_pow comp_def) then show "f (lfp (f ^^ Suc n)) \ lfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (lfp f) = lfp f" for n by (induct n) (auto intro: f lfp_fixpoint) then show "lfp (f ^^ Suc n) \ lfp f" by (intro lfp_lowerbound) (simp del: funpow.simps) qed lemma gfp_funpow: assumes f: "mono f" shows "gfp (f ^^ Suc n) = gfp f" proof (rule antisym) show "gfp f \ gfp (f ^^ Suc n)" proof (rule gfp_upperbound) have "f (gfp (f ^^ Suc n)) = gfp (\x. f ((f ^^ n) x))" unfolding funpow_Suc_right by (simp add: gfp_rolling f mono_pow comp_def) then show "f (gfp (f ^^ Suc n)) \ gfp (f ^^ Suc n)" by (simp add: comp_def) qed have "(f ^^ n) (gfp f) = gfp f" for n by (induct n) (auto intro: f gfp_fixpoint) then show "gfp (f ^^ Suc n) \ gfp f" by (intro gfp_upperbound) (simp del: funpow.simps) qed lemma Kleene_iter_gpfp: fixes f :: "'a::order_top \ 'a" assumes "mono f" and "p \ f p" shows "p \ (f ^^ k) top" proof (induct k) case 0 show ?case by simp next case Suc show ?case using monoD[OF assms(1) Suc] assms(2) by simp qed lemma gfp_Kleene_iter: assumes "mono f" and "(f ^^ Suc k) top = (f ^^ k) top" shows "gfp f = (f ^^ k) top" (is "?lhs = ?rhs") proof (rule antisym) have "?rhs \ f ?rhs" using assms(2) by simp then show "?rhs \ ?lhs" by (rule gfp_upperbound) show "?lhs \ ?rhs" using Kleene_iter_gpfp[OF assms(1)] gfp_unfold[OF assms(1)] by simp qed subsection \Embedding of the naturals into any \semiring_1\: \<^term>\of_nat\\ context semiring_1 begin definition of_nat :: "nat \ 'a" where "of_nat n = (plus 1 ^^ n) 0" lemma of_nat_simps [simp]: shows of_nat_0: "of_nat 0 = 0" and of_nat_Suc: "of_nat (Suc m) = 1 + of_nat m" by (simp_all add: of_nat_def) lemma of_nat_1 [simp]: "of_nat 1 = 1" by (simp add: of_nat_def) lemma of_nat_add [simp]: "of_nat (m + n) = of_nat m + of_nat n" by (induct m) (simp_all add: ac_simps) lemma of_nat_mult [simp]: "of_nat (m * n) = of_nat m * of_nat n" by (induct m) (simp_all add: ac_simps distrib_right) lemma mult_of_nat_commute: "of_nat x * y = y * of_nat x" by (induct x) (simp_all add: algebra_simps) primrec of_nat_aux :: "('a \ 'a) \ nat \ 'a \ 'a" where "of_nat_aux inc 0 i = i" | "of_nat_aux inc (Suc n) i = of_nat_aux inc n (inc i)" \ \tail recursive\ lemma of_nat_code: "of_nat n = of_nat_aux (\i. i + 1) n 0" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "\i. of_nat_aux (\i. i + 1) n (i + 1) = of_nat_aux (\i. i + 1) n i + 1" by (induct n) simp_all from this [of 0] have "of_nat_aux (\i. i + 1) n 1 = of_nat_aux (\i. i + 1) n 0 + 1" by simp with Suc show ?case by (simp add: add.commute) qed lemma of_nat_of_bool [simp]: "of_nat (of_bool P) = of_bool P" by auto end declare of_nat_code [code] context semiring_1_cancel begin lemma of_nat_diff: \of_nat (m - n) = of_nat m - of_nat n\ if \n \ m\ proof - from that obtain q where \m = n + q\ by (blast dest: le_Suc_ex) then show ?thesis by simp qed end text \Class for unital semirings with characteristic zero. Includes non-ordered rings like the complex numbers.\ class semiring_char_0 = semiring_1 + assumes inj_of_nat: "inj of_nat" begin lemma of_nat_eq_iff [simp]: "of_nat m = of_nat n \ m = n" by (auto intro: inj_of_nat injD) text \Special cases where either operand is zero\ lemma of_nat_0_eq_iff [simp]: "0 = of_nat n \ 0 = n" by (fact of_nat_eq_iff [of 0 n, unfolded of_nat_0]) lemma of_nat_eq_0_iff [simp]: "of_nat m = 0 \ m = 0" by (fact of_nat_eq_iff [of m 0, unfolded of_nat_0]) lemma of_nat_1_eq_iff [simp]: "1 = of_nat n \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_eq_1_iff [simp]: "of_nat n = 1 \ n=1" using of_nat_eq_iff by fastforce lemma of_nat_neq_0 [simp]: "of_nat (Suc n) \ 0" unfolding of_nat_eq_0_iff by simp lemma of_nat_0_neq [simp]: "0 \ of_nat (Suc n)" unfolding of_nat_0_eq_iff by simp end class ring_char_0 = ring_1 + semiring_char_0 context linordered_nonzero_semiring begin lemma of_nat_0_le_iff [simp]: "0 \ of_nat n" by (induct n) simp_all lemma of_nat_less_0_iff [simp]: "\ of_nat m < 0" by (simp add: not_less) lemma of_nat_mono[simp]: "i \ j \ of_nat i \ of_nat j" by (auto simp: le_iff_add intro!: add_increasing2) lemma of_nat_less_iff [simp]: "of_nat m < of_nat n \ m < n" proof(induct m n rule: diff_induct) case (1 m) then show ?case by auto next case (2 n) then show ?case by (simp add: add_pos_nonneg) next case (3 m n) then show ?case by (auto simp: add_commute [of 1] add_mono1 not_less add_right_mono leD) qed lemma of_nat_le_iff [simp]: "of_nat m \ of_nat n \ m \ n" by (simp add: not_less [symmetric] linorder_not_less [symmetric]) lemma less_imp_of_nat_less: "m < n \ of_nat m < of_nat n" by simp lemma of_nat_less_imp_less: "of_nat m < of_nat n \ m < n" by simp text \Every \linordered_nonzero_semiring\ has characteristic zero.\ subclass semiring_char_0 by standard (auto intro!: injI simp add: order.eq_iff) text \Special cases where either operand is zero\ lemma of_nat_le_0_iff [simp]: "of_nat m \ 0 \ m = 0" by (rule of_nat_le_iff [of _ 0, simplified]) lemma of_nat_0_less_iff [simp]: "0 < of_nat n \ 0 < n" by (rule of_nat_less_iff [of 0, simplified]) end context linordered_nonzero_semiring begin lemma of_nat_max: "of_nat (max x y) = max (of_nat x) (of_nat y)" by (auto simp: max_def ord_class.max_def) lemma of_nat_min: "of_nat (min x y) = min (of_nat x) (of_nat y)" by (auto simp: min_def ord_class.min_def) end context linordered_semidom begin subclass linordered_nonzero_semiring .. subclass semiring_char_0 .. end context linordered_idom begin lemma abs_of_nat [simp]: "\of_nat n\ = of_nat n" by (simp add: abs_if) lemma sgn_of_nat [simp]: "sgn (of_nat n) = of_bool (n > 0)" by simp end lemma of_nat_id [simp]: "of_nat n = n" by (induct n) simp_all lemma of_nat_eq_id [simp]: "of_nat = id" by (auto simp add: fun_eq_iff) subsection \The set of natural numbers\ context semiring_1 begin definition Nats :: "'a set" ("\") where "\ = range of_nat" lemma of_nat_in_Nats [simp]: "of_nat n \ \" by (simp add: Nats_def) lemma Nats_0 [simp]: "0 \ \" using of_nat_0 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_1 [simp]: "1 \ \" using of_nat_1 [symmetric] unfolding Nats_def by (rule range_eqI) lemma Nats_add [simp]: "a \ \ \ b \ \ \ a + b \ \" unfolding Nats_def using of_nat_add [symmetric] by (blast intro: range_eqI) lemma Nats_mult [simp]: "a \ \ \ b \ \ \ a * b \ \" unfolding Nats_def using of_nat_mult [symmetric] by (blast intro: range_eqI) lemma Nats_cases [cases set: Nats]: assumes "x \ \" obtains (of_nat) n where "x = of_nat n" unfolding Nats_def proof - from \x \ \\ have "x \ range of_nat" unfolding Nats_def . then obtain n where "x = of_nat n" .. then show thesis .. qed lemma Nats_induct [case_names of_nat, induct set: Nats]: "x \ \ \ (\n. P (of_nat n)) \ P x" by (rule Nats_cases) auto end lemma Nats_diff [simp]: fixes a:: "'a::linordered_idom" assumes "a \ \" "b \ \" "b \ a" shows "a - b \ \" proof - obtain i where i: "a = of_nat i" using Nats_cases assms by blast obtain j where j: "b = of_nat j" using Nats_cases assms by blast have "j \ i" using \b \ a\ i j of_nat_le_iff by blast then have *: "of_nat i - of_nat j = (of_nat (i-j) :: 'a)" by (simp add: of_nat_diff) then show ?thesis by (simp add: * i j) qed subsection \Further arithmetic facts concerning the natural numbers\ lemma subst_equals: assumes "t = s" and "u = t" shows "u = s" using assms(2,1) by (rule trans) locale nat_arith begin lemma add1: "(A::'a::comm_monoid_add) \ k + a \ A + b \ k + (a + b)" by (simp only: ac_simps) lemma add2: "(B::'a::comm_monoid_add) \ k + b \ a + B \ k + (a + b)" by (simp only: ac_simps) lemma suc1: "A == k + a \ Suc A \ k + Suc a" by (simp only: add_Suc_right) lemma rule0: "(a::'a::comm_monoid_add) \ a + 0" by (simp only: add_0_right) end ML_file \Tools/nat_arith.ML\ simproc_setup nateq_cancel_sums ("(l::nat) + m = n" | "(l::nat) = m + n" | "Suc m = n" | "m = Suc n") = - \fn phi => try o Nat_Arith.cancel_eq_conv\ + \K (try o Nat_Arith.cancel_eq_conv)\ simproc_setup natless_cancel_sums ("(l::nat) + m < n" | "(l::nat) < m + n" | "Suc m < n" | "m < Suc n") = - \fn phi => try o Nat_Arith.cancel_less_conv\ + \K (try o Nat_Arith.cancel_less_conv)\ simproc_setup natle_cancel_sums ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "Suc m \ n" | "m \ Suc n") = - \fn phi => try o Nat_Arith.cancel_le_conv\ + \K (try o Nat_Arith.cancel_le_conv)\ simproc_setup natdiff_cancel_sums ("(l::nat) + m - n" | "(l::nat) - (m + n)" | "Suc m - n" | "m - Suc n") = - \fn phi => try o Nat_Arith.cancel_diff_conv\ + \K (try o Nat_Arith.cancel_diff_conv)\ context order begin lemma lift_Suc_mono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_antimono_le: assumes mono: "\n. f n \ f (Suc n)" and "n \ n'" shows "f n \ f n'" proof (cases "n < n'") case True then show ?thesis by (induct n n' rule: less_Suc_induct) (auto intro: mono) next case False with \n \ n'\ show ?thesis by auto qed lemma lift_Suc_mono_less: assumes mono: "\n. f n < f (Suc n)" and "n < n'" shows "f n < f n'" using \n < n'\ by (induct n n' rule: less_Suc_induct) (auto intro: mono) lemma lift_Suc_mono_less_iff: "(\n. f n < f (Suc n)) \ f n < f m \ n < m" by (blast intro: less_asym' lift_Suc_mono_less [of f] dest: linorder_not_less[THEN iffD1] le_eq_less_or_eq [THEN iffD1]) end lemma mono_iff_le_Suc: "mono f \ (\n. f n \ f (Suc n))" unfolding mono_def by (auto intro: lift_Suc_mono_le [of f]) lemma antimono_iff_le_Suc: "antimono f \ (\n. f (Suc n) \ f n)" unfolding antimono_def by (auto intro: lift_Suc_antimono_le [of f]) lemma strict_mono_Suc_iff: "strict_mono f \ (\n. f n < f (Suc n))" proof (intro iffI strict_monoI) assume *: "\n. f n < f (Suc n)" fix m n :: nat assume "m < n" thus "f m < f n" by (induction rule: less_Suc_induct) (use * in auto) qed (auto simp: strict_mono_def) lemma strict_mono_add: "strict_mono (\n::'a::linordered_semidom. n + k)" by (auto simp: strict_mono_def) lemma mono_nat_linear_lb: fixes f :: "nat \ nat" assumes "\m n. m < n \ f m < f n" shows "f m + k \ f (m + k)" proof (induct k) case 0 then show ?case by simp next case (Suc k) then have "Suc (f m + k) \ Suc (f (m + k))" by simp also from assms [of "m + k" "Suc (m + k)"] have "Suc (f (m + k)) \ f (Suc (m + k))" by (simp add: Suc_le_eq) finally show ?case by simp qed text \Subtraction laws, mostly by Clemens Ballarin\ lemma diff_less_mono: fixes a b c :: nat assumes "a < b" and "c \ a" shows "a - c < b - c" proof - from assms obtain d e where "b = c + (d + e)" and "a = c + e" and "d > 0" by (auto dest!: le_Suc_ex less_imp_Suc_add simp add: ac_simps) then show ?thesis by simp qed lemma less_diff_conv: "i < j - k \ i + k < j" for i j k :: nat by (cases "k \ j") (auto simp add: not_le dest: less_imp_Suc_add le_Suc_ex) lemma less_diff_conv2: "k \ j \ j - k < i \ j < i + k" for j k i :: nat by (auto dest: le_Suc_ex) lemma le_diff_conv: "j - k \ i \ j \ i + k" for j k i :: nat by (cases "k \ j") (auto simp add: not_le dest!: less_imp_Suc_add le_Suc_ex) lemma diff_diff_cancel [simp]: "i \ n \ n - (n - i) = i" for i n :: nat by (auto dest: le_Suc_ex) lemma diff_less [simp]: "0 < n \ 0 < m \ m - n < m" for i n :: nat by (auto dest: less_imp_Suc_add) text \Simplification of relational expressions involving subtraction\ lemma diff_diff_eq: "k \ m \ k \ n \ m - k - (n - k) = m - n" for m n k :: nat by (auto dest!: le_Suc_ex) hide_fact (open) diff_diff_eq lemma eq_diff_iff: "k \ m \ k \ n \ m - k = n - k \ m = n" for m n k :: nat by (auto dest: le_Suc_ex) lemma less_diff_iff: "k \ m \ k \ n \ m - k < n - k \ m < n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff: "k \ m \ k \ n \ m - k \ n - k \ m \ n" for m n k :: nat by (auto dest!: le_Suc_ex) lemma le_diff_iff': "a \ c \ b \ c \ c - a \ c - b \ b \ a" for a b c :: nat by (force dest: le_Suc_ex) text \(Anti)Monotonicity of subtraction -- by Stephan Merz\ lemma diff_le_mono: "m \ n \ m - l \ n - l" for m n l :: nat by (auto dest: less_imp_le less_imp_Suc_add split: nat_diff_split) lemma diff_le_mono2: "m \ n \ l - n \ l - m" for m n l :: nat by (auto dest: less_imp_le le_Suc_ex less_imp_Suc_add less_le_trans split: nat_diff_split) lemma diff_less_mono2: "m < n \ m < l \ l - n < l - m" for m n l :: nat by (auto dest: less_imp_Suc_add split: nat_diff_split) lemma diffs0_imp_equal: "m - n = 0 \ n - m = 0 \ m = n" for m n :: nat by (simp split: nat_diff_split) lemma min_diff: "min (m - i) (n - i) = min m n - i" for m n i :: nat by (cases m n rule: le_cases) (auto simp add: not_le min.absorb1 min.absorb2 min.absorb_iff1 [symmetric] diff_le_mono) lemma inj_on_diff_nat: fixes k :: nat assumes "\n. n \ N \ k \ n" shows "inj_on (\n. n - k) N" proof (rule inj_onI) fix x y assume a: "x \ N" "y \ N" "x - k = y - k" with assms have "x - k + k = y - k + k" by auto with a assms show "x = y" by (auto simp add: eq_diff_iff) qed text \Rewriting to pull differences out\ lemma diff_diff_right [simp]: "k \ j \ i - (j - k) = i + k - j" for i j k :: nat by (fact diff_diff_right) lemma diff_Suc_diff_eq1 [simp]: assumes "k \ j" shows "i - Suc (j - k) = i + k - Suc j" proof - from assms have *: "Suc (j - k) = Suc j - k" by (simp add: Suc_diff_le) from assms have "k \ Suc j" by (rule order_trans) simp with diff_diff_right [of k "Suc j" i] * show ?thesis by simp qed lemma diff_Suc_diff_eq2 [simp]: assumes "k \ j" shows "Suc (j - k) - i = Suc j - (k + i)" proof - from assms obtain n where "j = k + n" by (auto dest: le_Suc_ex) moreover have "Suc n - i = (k + Suc n) - (k + i)" using add_diff_cancel_left [of k "Suc n" i] by simp ultimately show ?thesis by simp qed lemma Suc_diff_Suc: assumes "n < m" shows "Suc (m - Suc n) = m - n" proof - from assms obtain q where "m = n + Suc q" by (auto dest: less_imp_Suc_add) moreover define r where "r = Suc q" ultimately have "Suc (m - Suc n) = r" and "m = n + r" by simp_all then show ?thesis by simp qed lemma one_less_mult: "Suc 0 < n \ Suc 0 < m \ Suc 0 < m * n" using less_1_mult [of n m] by (simp add: ac_simps) lemma n_less_m_mult_n: "0 < n \ Suc 0 < m \ n < m * n" using mult_strict_right_mono [of 1 m n] by simp lemma n_less_n_mult_m: "0 < n \ Suc 0 < m \ n < n * m" using mult_strict_left_mono [of 1 m n] by simp text \Induction starting beyond zero\ lemma nat_induct_at_least [consumes 1, case_names base Suc]: "P n" if "n \ m" "P m" "\n. n \ m \ P n \ P (Suc n)" proof - define q where "q = n - m" with \n \ m\ have "n = m + q" by simp moreover have "P (m + q)" by (induction q) (use that in simp_all) ultimately show "P n" by simp qed lemma nat_induct_non_zero [consumes 1, case_names 1 Suc]: "P n" if "n > 0" "P 1" "\n. n > 0 \ P n \ P (Suc n)" proof - from \n > 0\ have "n \ 1" by (cases n) simp_all moreover note \P 1\ moreover have "\n. n \ 1 \ P n \ P (Suc n)" using \\n. n > 0 \ P n \ P (Suc n)\ by (simp add: Suc_le_eq) ultimately show "P n" by (rule nat_induct_at_least) qed text \Specialized induction principles that work "backwards":\ lemma inc_induct [consumes 1, case_names base step]: assumes less: "i \ j" and base: "P j" and step: "\n. i \ n \ n < j \ P (Suc n) \ P n" shows "P i" using less step proof (induct "j - i" arbitrary: i) case (0 i) then have "i = j" by simp with base show ?case by simp next case (Suc d n) from Suc.hyps have "n \ j" by auto with Suc have "n < j" by (simp add: less_le) from \Suc d = j - n\ have "d + 1 = j - n" by simp then have "d + 1 - 1 = j - n - 1" by simp then have "d = j - n - 1" by simp then have "d = j - (n + 1)" by (simp add: diff_diff_eq) then have "d = j - Suc n" by simp moreover from \n < j\ have "Suc n \ j" by (simp add: Suc_le_eq) ultimately have "P (Suc n)" proof (rule Suc.hyps) fix q assume "Suc n \ q" then have "n \ q" by (simp add: Suc_le_eq less_imp_le) moreover assume "q < j" moreover assume "P (Suc q)" ultimately show "P q" by (rule Suc.prems) qed with order_refl \n < j\ show "P n" by (rule Suc.prems) qed lemma strict_inc_induct [consumes 1, case_names base step]: assumes less: "i < j" and base: "\i. j = Suc i \ P i" and step: "\i. i < j \ P (Suc i) \ P i" shows "P i" using less proof (induct "j - i - 1" arbitrary: i) case (0 i) from \i < j\ obtain n where "j = i + n" and "n > 0" by (auto dest!: less_imp_Suc_add) with 0 have "j = Suc i" by (auto intro: order_antisym simp add: Suc_le_eq) with base show ?case by simp next case (Suc d i) from \Suc d = j - i - 1\ have *: "Suc d = j - Suc i" by (simp add: diff_diff_add) then have "Suc d - 1 = j - Suc i - 1" by simp then have "d = j - Suc i - 1" by simp moreover from * have "j - Suc i \ 0" by auto then have "Suc i < j" by (simp add: not_le) ultimately have "P (Suc i)" by (rule Suc.hyps) with \i < j\ show "P i" by (rule step) qed lemma zero_induct_lemma: "P k \ (\n. P (Suc n) \ P n) \ P (k - i)" using inc_induct[of "k - i" k P, simplified] by blast lemma zero_induct: "P k \ (\n. P (Suc n) \ P n) \ P 0" using inc_induct[of 0 k P] by blast text \Further induction rule similar to @{thm inc_induct}.\ lemma dec_induct [consumes 1, case_names base step]: "i \ j \ P i \ (\n. i \ n \ n < j \ P n \ P (Suc n)) \ P j" proof (induct j arbitrary: i) case 0 then show ?case by simp next case (Suc j) from Suc.prems consider "i \ j" | "i = Suc j" by (auto simp add: le_Suc_eq) then show ?case proof cases case 1 moreover have "j < Suc j" by simp moreover have "P j" using \i \ j\ \P i\ proof (rule Suc.hyps) fix q assume "i \ q" moreover assume "q < j" then have "q < Suc j" by (simp add: less_Suc_eq) moreover assume "P q" ultimately show "P (Suc q)" by (rule Suc.prems) qed ultimately show "P (Suc j)" by (rule Suc.prems) next case 2 with \P i\ show "P (Suc j)" by simp qed qed lemma transitive_stepwise_le: assumes "m \ n" "\x. R x x" "\x y z. R x y \ R y z \ R x z" and "\n. R n (Suc n)" shows "R m n" using \m \ n\ by (induction rule: dec_induct) (use assms in blast)+ subsubsection \Greatest operator\ lemma ex_has_greatest_nat: "P (k::nat) \ \y. P y \ y \ b \ \x. P x \ (\y. P y \ y \ x)" proof (induction "b-k" arbitrary: b k rule: less_induct) case less show ?case proof cases assume "\n>k. P n" then obtain n where "n>k" "P n" by blast have "n \ b" using \P n\ less.prems(2) by auto hence "b-n < b-k" by(rule diff_less_mono2[OF \k less_le_trans[OF \k]]) from less.hyps[OF this \P n\ less.prems(2)] show ?thesis . next assume "\ (\n>k. P n)" hence "\y. P y \ y \ k" by (auto simp: not_less) thus ?thesis using less.prems(1) by auto qed qed lemma fixes k::nat assumes "P k" and minor: "\y. P y \ y \ b" shows GreatestI_nat: "P (Greatest P)" and Greatest_le_nat: "k \ Greatest P" proof - obtain x where "P x" "\y. P y \ y \ x" using assms ex_has_greatest_nat by blast with \P k\ show "P (Greatest P)" "k \ Greatest P" using GreatestI2_order by blast+ qed lemma GreatestI_ex_nat: "\ \k::nat. P k; \y. P y \ y \ b \ \ P (Greatest P)" by (blast intro: GreatestI_nat) subsection \Monotonicity of \funpow\\ lemma funpow_increasing: "m \ n \ mono f \ (f ^^ n) \ \ (f ^^ m) \" for f :: "'a::{lattice,order_top} \ 'a" by (induct rule: inc_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma funpow_decreasing: "m \ n \ mono f \ (f ^^ m) \ \ (f ^^ n) \" for f :: "'a::{lattice,order_bot} \ 'a" by (induct rule: dec_induct) (auto simp del: funpow.simps(2) simp add: funpow_Suc_right intro: order_trans[OF _ funpow_mono]) lemma mono_funpow: "mono Q \ mono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_bot} \ 'a" by (auto intro!: funpow_decreasing simp: mono_def) lemma antimono_funpow: "mono Q \ antimono (\i. (Q ^^ i) \)" for Q :: "'a::{lattice,order_top} \ 'a" by (auto intro!: funpow_increasing simp: antimono_def) subsection \The divides relation on \<^typ>\nat\\ lemma dvd_1_left [iff]: "Suc 0 dvd k" by (simp add: dvd_def) lemma dvd_1_iff_1 [simp]: "m dvd Suc 0 \ m = Suc 0" by (simp add: dvd_def) lemma nat_dvd_1_iff_1 [simp]: "m dvd 1 \ m = 1" for m :: nat by (simp add: dvd_def) lemma dvd_antisym: "m dvd n \ n dvd m \ m = n" for m n :: nat unfolding dvd_def by (force dest: mult_eq_self_implies_10 simp add: mult.assoc) lemma dvd_diff_nat [simp]: "k dvd m \ k dvd n \ k dvd (m - n)" for k m n :: nat unfolding dvd_def by (blast intro: right_diff_distrib' [symmetric]) lemma dvd_diffD: fixes k m n :: nat assumes "k dvd m - n" "k dvd n" "n \ m" shows "k dvd m" proof - have "k dvd n + (m - n)" using assms by (blast intro: dvd_add) with assms show ?thesis by simp qed lemma dvd_diffD1: "k dvd m - n \ k dvd m \ n \ m \ k dvd n" for k m n :: nat by (drule_tac m = m in dvd_diff_nat) auto lemma dvd_mult_cancel: fixes m n k :: nat assumes "k * m dvd k * n" and "0 < k" shows "m dvd n" proof - from assms(1) obtain q where "k * n = (k * m) * q" .. then have "k * n = k * (m * q)" by (simp add: ac_simps) with \0 < k\ have "n = m * q" by (auto simp add: mult_left_cancel) then show ?thesis .. qed lemma dvd_mult_cancel1: fixes m n :: nat assumes "0 < m" shows "m * n dvd m \ n = 1" proof assume "m * n dvd m" then have "m * n dvd m * 1" by simp then have "n dvd 1" by (iprover intro: assms dvd_mult_cancel) then show "n = 1" by auto qed auto lemma dvd_mult_cancel2: "0 < m \ n * m dvd m \ n = 1" for m n :: nat using dvd_mult_cancel1 [of m n] by (simp add: ac_simps) lemma dvd_imp_le: "k dvd n \ 0 < n \ k \ n" for k n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma nat_dvd_not_less: "0 < m \ m < n \ \ n dvd m" for m n :: nat by (auto elim!: dvdE) (auto simp add: gr0_conv_Suc) lemma less_eq_dvd_minus: fixes m n :: nat assumes "m \ n" shows "m dvd n \ m dvd n - m" proof - from assms have "n = m + (n - m)" by simp then obtain q where "n = m + q" .. then show ?thesis by (simp add: add.commute [of m]) qed lemma dvd_minus_self: "m dvd n - m \ n < m \ m dvd n" for m n :: nat by (cases "n < m") (auto elim!: dvdE simp add: not_less le_imp_diff_is_add dest: less_imp_le) lemma dvd_minus_add: fixes m n q r :: nat assumes "q \ n" "q \ r * m" shows "m dvd n - q \ m dvd n + (r * m - q)" proof - have "m dvd n - q \ m dvd r * m + (n - q)" using dvd_add_times_triv_left_iff [of m r] by simp also from assms have "\ \ m dvd r * m + n - q" by simp also from assms have "\ \ m dvd (r * m - q) + n" by simp also have "\ \ m dvd n + (r * m - q)" by (simp add: add.commute) finally show ?thesis . qed subsection \Aliasses\ lemma nat_mult_1: "1 * n = n" for n :: nat by (fact mult_1_left) lemma nat_mult_1_right: "n * 1 = n" for n :: nat by (fact mult_1_right) lemma diff_mult_distrib: "(m - n) * k = (m * k) - (n * k)" for k m n :: nat by (fact left_diff_distrib') lemma diff_mult_distrib2: "k * (m - n) = (k * m) - (k * n)" for k m n :: nat by (fact right_diff_distrib') (*Used in AUTO2 and Groups.le_diff_conv2 (with variables renamed) doesn't work for some reason*) lemma le_diff_conv2: "k \ j \ (i \ j - k) = (i + k \ j)" for i j k :: nat by (fact le_diff_conv2) lemma diff_self_eq_0 [simp]: "m - m = 0" for m :: nat by (fact diff_cancel) lemma diff_diff_left [simp]: "i - j - k = i - (j + k)" for i j k :: nat by (fact diff_diff_add) lemma diff_commute: "i - j - k = i - k - j" for i j k :: nat by (fact diff_right_commute) lemma diff_add_inverse: "(n + m) - n = m" for m n :: nat by (fact add_diff_cancel_left') lemma diff_add_inverse2: "(m + n) - n = m" for m n :: nat by (fact add_diff_cancel_right') lemma diff_cancel: "(k + m) - (k + n) = m - n" for k m n :: nat by (fact add_diff_cancel_left) lemma diff_cancel2: "(m + k) - (n + k) = m - n" for k m n :: nat by (fact add_diff_cancel_right) lemma diff_add_0: "n - (n + m) = 0" for m n :: nat by (fact diff_add_zero) lemma add_mult_distrib2: "k * (m + n) = (k * m) + (k * n)" for k m n :: nat by (fact distrib_left) lemmas nat_distrib = add_mult_distrib distrib_left diff_mult_distrib diff_mult_distrib2 subsection \Size of a datatype value\ class size = fixes size :: "'a \ nat" \ \see further theory \Wellfounded\\ instantiation nat :: size begin definition size_nat where [simp, code]: "size (n::nat) = n" instance .. end lemmas size_nat = size_nat_def lemma size_neq_size_imp_neq: "size x \ size y \ x \ y" by (erule contrapos_nn) (rule arg_cong) subsection \Code module namespace\ code_identifier code_module Nat \ (SML) Arith and (OCaml) Arith and (Haskell) Arith hide_const (open) of_nat_aux end diff --git a/src/HOL/Nominal/nominal_inductive.ML b/src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML +++ b/src/HOL/Nominal/nominal_inductive.ML @@ -1,698 +1,698 @@ (* Title: HOL/Nominal/nominal_inductive.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. *) signature NOMINAL_INDUCTIVE = sig val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state val prove_eqvt: string -> string list -> local_theory -> local_theory end structure NominalInductive : NOMINAL_INDUCTIVE = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val fresh_prod = @{thm fresh_prod}; val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else insert (op aconv o apply2 fst) (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} => EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)), REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion; fun first_order_matchs pats objs = Thm.first_order_match (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); fun first_order_mrs ths th = ths MRS Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th; fun prove_strong_ind s avoids lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val monos = Inductive.get_monos lthy; val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP lthy (hd names))))); val (raw_induct', ctxt') = lthy |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse error ("Duplicate variable names for case " ^ quote a)); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); val avoids' = if null induct_cases then replicate (length intrs) ("", []) else map (fn name => (name, the_default [] (AList.lookup op = avoids name))) induct_cases; fun mk_avoids params (name, ps) = let val k = length params - 1 in map (fn x => case find_index (equal x o fst) params of ~1 => error ("No such variable in case " ^ quote name ^ " of inductive definition: " ^ quote x) | i => (Bound (k - i), snd (nth params i))) ps end; val prems = map (fn (prem, avoid) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, fold (add_binders thy 0) (prems @ [concl]) [] @ map (apfst (incr_boundvars 1)) (mk_avoids params avoid), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); val atomTs = distinct op = (maps (map snd o #2) prems); val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy ("fs_" ^ Long_Name.base_name (fst (dest_Type T)))) atomTs)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_distinct [] = [] | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) => if T = U then SOME (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) else NONE) xs @ mk_distinct xs; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh bvars @ mk_distinct bvars @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') in (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem))) end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val vc_compat = map (fn (params, bvars, prems, (p, ts)) => map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (mk_distinct bvars @ maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_const U T $ u $ t)) bvars) (ts ~~ binder_types (fastype_of p)))) prems; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn aT => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs; fun eqvt_ss ctxt = put_simpset HOL_basic_ss ctxt addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; val perm_bij = Global_Theory.get_thms thy "perm_bij"; val fs_atoms = map (fn aT => Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "1")) atomTs; val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'"; val fresh_atm = Global_Theory.get_thms thy "fresh_atm"; val swap_simps = Global_Theory.get_thms thy "swap_simps"; val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh"; fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = let (** protect terms to avoid that fresh_prod interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, NominalDatatype.fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn {context = goal_ctxt, ...} => EVERY [resolve_tac goal_ctxt exists_fresh' 1, resolve_tac goal_ctxt fs_atoms 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn goal_ctxt => EVERY [eresolve_tac goal_ctxt [exE] 1, full_simp_tac (put_simpset HOL_ss goal_ctxt addsimps (fresh_prod :: fresh_atm)) 1, full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps [@{thm id_apply}]) 1, REPEAT (eresolve_tac goal_ctxt [conjE] 1)]) [ex] ctxt in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; fun mk_ind_proof ctxt thss = Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} => let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} => resolve_tac goal_ctxt1 [raw_induct] 1 THEN EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => [REPEAT (resolve_tac goal_ctxt1 [allI] 1), simp_tac (eqvt_ss goal_ctxt1) 1, SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} => let val (params', (pis, z)) = chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||> split_last; val bvars' = map (fn (Bound i, T) => (nth params' (length params' - i), T) | (t, T) => (t, T)) bvars; val pi_bvars = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val (freshs1, freshs2, ctxt') = fold (obtain_fresh_name (ts @ pi_bvars)) (map snd bvars') ([], [], goal_ctxt2); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold (concat_perm #> map) pis' pis; val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) (Vartab.empty, Vartab.empty); val ihyp' = Thm.instantiate (TVars.empty, Vars.make (map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t)) (map (Envir.subst_term env) vs ~~ map (fold_rev (NominalDatatype.mk_perm []) (rev pis' @ pis)) params' @ [z]))) ihyp; fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (eqvt_ss ctxt') (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt') (rev pis' @ pis) th)); val (gprems1, gprems2) = split_list (map (fn (th, t) => if null (preds_of ps t) then (SOME th, mk_pi th) else (map_thm ctxt' (split_conj (K o I) names) (eresolve_tac ctxt' [conjunct1] 1) monos NONE th, mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th)))) (gprems ~~ oprems)) |>> map_filter I; val vc_compat_ths' = map (fn th => let val th' = first_order_mrs gprems1 th; val (bop, lhs, rhs) = (case Thm.concl_of th' of _ $ (fresh $ lhs $ rhs) => (fn t => fn u => fresh $ t $ u, lhs, rhs) | _ $ (_ $ (_ $ lhs $ rhs)) => (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); val th'' = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) (fold_rev (NominalDatatype.mk_perm []) pis rhs))) (fn {context = goal_ctxt3, ...} => simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps (fresh_bij @ perm_bij)) 1 THEN resolve_tac goal_ctxt3 [th'] 1) in Simplifier.simplify (eqvt_ss ctxt' addsimps fresh_atm) th'' end) vc_compat_ths; val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) (** we have to pre-simplify the rewrite rules **) val swap_simps_simpset = put_simpset HOL_ss ctxt' addsimps swap_simps @ map (Simplifier.simplify (put_simpset HOL_ss ctxt' addsimps swap_simps)) (vc_compat_ths'' @ freshs2'); val th = Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn {context = goal_ctxt4, ...} => EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1, resolve_tac goal_ctxt4 [ihyp'] 1, REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems) (simp_tac swap_simps_simpset 1), REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss goal_ctxt4 addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac goal_ctxt4 gprems2 1)])); val final = Goal.prove ctxt' [] [] (Thm.term_of concl) (fn {context = goal_ctxt5, ...} => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5 addsimps vc_compat_ths'' @ freshs2' @ perm_fresh_fresh @ fresh_atm) 1); val final' = Proof_Context.export ctxt' goal_ctxt2 [final]; in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1]) (prems ~~ thss ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN asm_full_simp_tac goal_ctxt 1) end) |> singleton (Proof_Context.export ctxt lthy); (** strong case analysis rule **) val cases_prems = map (fn ((name, avoids), rule) => let val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] lthy; val prem :: prems = Logic.strip_imp_prems rule'; val concl = Logic.strip_imp_concl rule' in (prem, List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), concl, fold_map (fn (prem, (_, avoid)) => fn ctxt => let val prems = Logic.strip_assums_hyp prem; val params = Logic.strip_params prem; val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) = if member (op = o apsnd fst) bnds (Bound i) then let val ([s'], ctxt') = Variable.variant_fixes [s] ctxt; val t = Free (s', T) in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts); val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params (0, 0, ctxt, [], [], [], []) in ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt') end) (prems ~~ avoids) ctxt') end) (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~ elims); val cases_prems' = map (fn (prem, args, concl, (prems, _)) => let fun mk_prem (ps, [], _, prems) = Logic.list_all (ps, Logic.list_implies (prems, concl)) | mk_prem (ps, qs, _, prems) = Logic.list_all (ps, Logic.mk_implies (Logic.list_implies (mk_distinct qs @ maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) args) qs, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop prems))), concl)) in map mk_prem prems end) cases_prems; fun cases_eqvt_simpset ctxt = put_simpset HOL_ss ctxt addsimps eqvt_thms @ swap_simps @ perm_pi_simp addsimprocs [NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; fun simp_fresh_atm ctxt = Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps fresh_atm); fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt))), prems') = (name, Goal.prove ctxt [] (prem :: prems') concl (fn {prems = hyp :: hyps, context = ctxt1} => EVERY (resolve_tac ctxt1 [hyp RS elim] 1 :: map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} => if null qs then resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1 else let val params' = map (Thm.term_of o #2 o nth (rev params)) is; val tab = params' ~~ map fst qs; val (hyps1, hyps2) = chop (length args) case_hyps; (* turns a = t and [x1 # t, ..., xn # t] *) (* into [x1 # a, ..., xn # a] *) fun inst_fresh th' ths = let val (ths1, ths2) = chop (length qs) ths in (map (fn th => let val (cf, ct) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)); val arg_cong' = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ct)] [NONE, SOME ct, SOME cf] (arg_cong RS iffD2); val inst = Thm.first_order_match (ct, Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th'))) in [th', th] MRS Thm.instantiate inst arg_cong' end) ths1, ths2) end; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length args * length qs) (map (first_order_mrs hyps2) vc_compat_ths); val vc_compat_ths' = NominalDatatype.mk_not_sym vc_compat_ths1 @ flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); val (freshs1, freshs2, ctxt3) = fold (obtain_fresh_name (args @ map fst qs @ params')) (map snd qs) ([], [], ctxt2); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis = map (NominalDatatype.perm_of_pair) ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); val mk_pis = fold_rev (mk_perm_bool ctxt3) (map (Thm.cterm_of ctxt3) pis); val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms (fn x as Free _ => if member (op =) args x then x else (case AList.lookup op = tab x of SOME y => y | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps)); val inst = Thm.first_order_match (Thm.dest_arg (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); val th = Goal.prove ctxt3 [] [] (Thm.term_of concl) (fn {context = ctxt4, ...} => resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} => let val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; val case_simpset = cases_eqvt_simpset ctxt5 addsimps freshs2' @ map (simp_fresh_atm ctxt5) (vc_compat_ths' @ fresh_hyps'); val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; val hyps1' = map (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; val hyps2' = map (mk_pis #> Simplifier.simplify case_simpset) hyps2; val case_hyps' = hyps1' @ hyps2' in simp_tac case_simpset 1 THEN REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN resolve_tac ctxt5 case_hyps' 1) end) ctxt4 1) val final = Proof_Context.export ctxt3 ctxt2 [th] in resolve_tac ctxt2 final 1 end) ctxt1 1) (thss ~~ hyps ~~ prems))) |> singleton (Proof_Context.export ctxt lthy)) in ctxt'' |> Proof.theorem NONE (fn thss => fn lthy1 => let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr lthy1)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1; val strong_cases = map (mk_cases_proof ##> Inductive.rulify lthy1) (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); val strong_induct_atts = - map (Attrib.internal o K) + map (Attrib.internal \<^here> o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects lthy1 (1 upto length names) strong_induct'; in lthy2 |> Local_Theory.notes [((rec_qualified (Binding.name "strong_inducts"), []), strong_inducts |> map (fn th => ([th], - [Attrib.internal (K ind_case_names), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> + [Attrib.internal \<^here> (K ind_case_names), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> Local_Theory.notes (map (fn ((name, elim), (_, cases)) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), - [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) + [Attrib.internal \<^here> (K (Rule_Cases.case_names (map snd cases))), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) (strong_cases ~~ induct_cases')) |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; fun prove_eqvt s xatoms lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val intrs = map (atomize_intr lthy) intrs; val monos = Inductive.get_monos lthy; val intrs' = Inductive.unpartition_rules intrs (map (fn (((s, ths), (_, k)), th) => (s, ths ~~ Inductive.infer_intro_vars thy th k ths)) (Inductive.partition_rules raw_induct intrs ~~ Inductive.arities_of raw_induct ~~ elims)); val k = length (Inductive.params_of raw_induct); val atoms' = NominalAtoms.atoms_of thy; val atoms = if null xatoms then atoms' else let val atoms = map (Sign.intern_type thy) xatoms in (case duplicates op = atoms of [] => () | xs => error ("Duplicate atoms: " ^ commas xs); case subtract (op =) atoms' atoms of [] => () | xs => error ("No such atoms: " ^ commas xs); atoms) end; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val (([t], [pi]), ctxt1) = lthy |> Variable.import_terms false [Thm.concl_of raw_induct] ||>> Variable.variant_fixes ["pi"]; fun eqvt_simpset ctxt = put_simpset HOL_basic_ss ctxt addsimps (NominalThmDecls.get_eqvt_thms ctxt @ perm_pi_simp) addsimprocs [mk_perm_bool_simproc names, NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val ps = map (fst o HOLogic.dest_imp) (HOLogic.dest_conj (HOLogic.dest_Trueprop t)); fun eqvt_tac ctxt pi (intr, vs) st = let fun eqvt_err s = let val ([t], ctxt') = Variable.import_terms true [Thm.prop_of intr] ctxt in error ("Could not prove equivariance for introduction rule\n" ^ Syntax.string_of_term ctxt' t ^ "\n" ^ s) end; val res = SUBPROOF (fn {context = goal_ctxt, prems, params, ...} => let val prems' = map (fn th => the_default th (map_thm goal_ctxt (split_conj (K I) names) (eresolve_tac goal_ctxt [conjunct2] 1) monos NONE th)) prems; val prems'' = map (fn th => Simplifier.simplify (eqvt_simpset goal_ctxt) (mk_perm_bool goal_ctxt (Thm.cterm_of goal_ctxt pi) th)) prems'; val intr' = infer_instantiate goal_ctxt (map (#1 o dest_Var) vs ~~ map (Thm.cterm_of goal_ctxt o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params) intr in (resolve_tac goal_ctxt [intr'] THEN_ALL_NEW (TRY o resolve_tac goal_ctxt prems'')) 1 end) ctxt 1 st in case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of NONE => eqvt_err ("Rule does not match goal\n" ^ Syntax.string_of_term ctxt (hd (Thm.prems_of st))) | SOME (th, _) => Seq.single th end; val thss = map (fn atom => let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) in map (fn th => zero_var_indexes (th RS mp)) (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt1 [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => let val (h, ts) = strip_comb p; val (ts1, ts2) = chop k ts in HOLogic.mk_imp (p, list_comb (h, ts1 @ map (NominalDatatype.mk_perm [] pi') ts2)) end) ps))) (fn {context = goal_ctxt, ...} => EVERY (resolve_tac goal_ctxt [raw_induct] 1 :: map (fn intr_vs => full_simp_tac (eqvt_simpset goal_ctxt) 1 THEN eqvt_tac goal_ctxt pi' intr_vs) intrs')) |> singleton (Proof_Context.export ctxt1 lthy))) end) atoms in lthy |> Local_Theory.notes (map (fn (name, ths) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), - [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) + @{attributes [eqvt]}), [(ths, [])])) (names ~~ transp thss)) |> snd end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive\ "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\avoids\ |-- Parse.and_list1 (Parse.name -- (\<^keyword>\:\ |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) => prove_strong_ind name avoids)); val _ = Outer_Syntax.local_theory \<^command_keyword>\equivariance\ "prove equivariance for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\[\ |-- Parse.list1 Parse.name --| \<^keyword>\]\) [] >> (fn (name, atoms) => prove_eqvt name atoms)); end diff --git a/src/HOL/Nominal/nominal_inductive2.ML b/src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML +++ b/src/HOL/Nominal/nominal_inductive2.ML @@ -1,505 +1,505 @@ (* Title: HOL/Nominal/nominal_inductive2.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. Experimental version that allows to avoid lists of atoms. *) signature NOMINAL_INDUCTIVE2 = sig val prove_strong_ind: string -> string option -> (string * string list) list -> local_theory -> Proof.state end structure NominalInductive2 : NOMINAL_INDUCTIVE2 = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (Conv.prems_conv ~1 o atomize_conv) ctxt)); fun fresh_postprocess ctxt = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim}, @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else AList.map_default op = (T, []) (insert op aconv (incr_boundvars (~i) u))) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn {context = goal_ctxt, ...} => EVERY [cut_facts_tac [th] 1, eresolve_tac goal_ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac goal_ctxt monos)), REPEAT_DETERM (resolve_tac goal_ctxt [impI] 1 THEN (assume_tac goal_ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; fun abs_params params t = let val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params) in (Logic.list_all (params, t), (rev vs, subst_bounds (vs, t))) end; fun inst_params thy (vs, p) th cts = let val env = Pattern.first_order_match thy (p, Thm.prop_of th) (Vartab.empty, Vartab.empty) in Thm.instantiate (TVars.empty, Vars.make (map (dest_Var o Envir.subst_term env) vs ~~ cts)) th end; fun prove_strong_ind s alt_name avoids lthy = let val thy = Proof_Context.theory_of lthy; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global lthy (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct lthy raw_induct; val elims = map (atomize_induct lthy) elims; val monos = Inductive.get_monos lthy; val eqvt_thms = NominalThmDecls.get_eqvt_thms lthy; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP lthy (hd names))))); val induct_cases' = if null induct_cases then replicate (length intrs) "" else induct_cases; val (raw_induct', ctxt') = lthy |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); fun mk_avoids params name sets = let val (_, ctxt') = Proof_Context.add_fixes (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) lthy; fun mk s = let val t = Syntax.read_term ctxt' s; val t' = fold_rev absfree params t |> funpow (length params) (fn Abs (_, _, t) => t) in (t', HOLogic.dest_setT (fastype_of t)) end handle TERM _ => error ("Expression " ^ quote s ^ " to be avoided in case " ^ quote name ^ " is not a set type"); fun add_set p [] = [p] | add_set (t, T) ((u, U) :: ps) = if T = U then let val S = HOLogic.mk_setT T in (Const (\<^const_name>\sup\, S --> S --> S) $ u $ t, T) :: ps end else (u, U) :: add_set (t, T) ps in fold (mk #> add_set) sets [] end; val prems = map (fn (prem, name) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, if null avoids then map (fn (T, ts) => (HOLogic.mk_set T ts, T)) (fold (add_binders thy 0) (prems @ [concl]) []) else case AList.lookup op = avoids name of NONE => [] | SOME sets => map (apfst (incr_boundvars 1)) (mk_avoids params name sets), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ induct_cases'); val atomTs = distinct op = (maps (map snd o #2) prems); val atoms = map (fst o dest_Type) atomTs; val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn a => Sign.intern_class thy ("fs_" ^ Long_Name.base_name a)) atoms)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh sets @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); in abs_params params' prem end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) => map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const U T $ u $ t)) sets) (ts ~~ binder_types (fastype_of p)) @ map (fn (u, U) => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |> split_list) prems |> split_list; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn a => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name a ^ "2")) atoms; fun eqvt_ss ctxt = put_simpset HOL_basic_ss ctxt addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij"; val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms; val at_insts = map (NominalAtoms.at_inst_of thy) atoms; val dj_thms = maps (fn a => map (NominalAtoms.dj_thm_of thy a) (remove (op =) a atoms)) atoms; val finite_ineq = map2 (fn th => fn th' => th' RS (th RS @{thm pt_set_finite_ineq})) pt_insts at_insts; val perm_set_forget = map (fn th => th RS @{thm dj_perm_set_forget}) dj_thms; val perm_freshs_freshs = atomTs ~~ map2 (fn th => fn th' => th' RS (th RS @{thm pt_freshs_freshs})) pt_insts at_insts; fun obtain_fresh_name ts sets (T, fin) (freshs, ths1, ths2, ths3, ctxt) = let val thy = Proof_Context.theory_of ctxt; (** protect terms to avoid that fresh_star_prod_set interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts); val atom = fst (dest_Type T); val {at_inst, ...} = NominalAtoms.the_atom_info thy atom; val fs_atom = Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name atom ^ "1"); val avoid_th = Thm.instantiate' [SOME (Thm.global_ctyp_of thy (fastype_of p))] [SOME (Thm.global_cterm_of thy p)] ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding}); val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result (fn ctxt' => EVERY [resolve_tac ctxt' [avoid_th] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1, full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, rotate_tac 1 1, REPEAT (eresolve_tac ctxt' [conjE] 1)]) [] ctxt; val (Ts1, _ :: Ts2) = chop_prefix (not o equal T) (map snd sets); val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2); val (pis1, pis2) = chop (length Ts1) (map Bound (length pTs - 1 downto 0)); val _ $ (f $ (_ $ pi $ l) $ r) = Thm.prop_of th2 val th2' = Goal.prove ctxt' [] [] (Logic.list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop (f $ fold_rev (NominalDatatype.mk_perm (rev pTs)) (pis1 @ pi :: pis2) l $ r))) (fn {context = goal_ctxt, ...} => cut_facts_tac [th2] 1 THEN full_simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps perm_set_forget) 1) |> Simplifier.simplify (eqvt_ss ctxt') in (freshs @ [Thm.term_of cx], ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt') end; fun mk_ind_proof ctxt thss = Goal.prove ctxt [] prems' concl' (fn {prems = ihyps, context = goal_ctxt} => let val th = Goal.prove goal_ctxt [] [] concl (fn {context = goal_ctxt1, ...} => resolve_tac goal_ctxt1 [raw_induct] 1 THEN EVERY (maps (fn (((((_, sets, oprems, _), vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) => [REPEAT (resolve_tac goal_ctxt1 [allI] 1), simp_tac (eqvt_ss goal_ctxt1) 1, SUBPROOF (fn {prems = gprems, params, concl, context = goal_ctxt2, ...} => let val (cparams', (pis, z)) = chop (length params - length atomTs - 1) (map #2 params) ||> (map Thm.term_of #> split_last); val params' = map Thm.term_of cparams' val sets' = map (apfst (curry subst_bounds (rev params'))) sets; val pi_sets = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) sets'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val gprems1 = map_filter (fn (th, t) => if null (preds_of ps t) then SOME th else map_thm goal_ctxt2 (split_conj (K o I) names) (eresolve_tac goal_ctxt2 [conjunct1] 1) monos NONE th) (gprems ~~ oprems); val vc_compat_ths' = map2 (fn th => fn p => let val th' = gprems1 MRS inst_params thy p th cparams'; val (h, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of th')) in Goal.prove goal_ctxt2 [] [] (HOLogic.mk_Trueprop (list_comb (h, map (fold_rev (NominalDatatype.mk_perm []) pis) ts))) (fn {context = goal_ctxt3, ...} => simp_tac (put_simpset HOL_basic_ss goal_ctxt3 addsimps (fresh_star_bij @ finite_ineq)) 1 THEN resolve_tac goal_ctxt3 [th'] 1) end) vc_compat_ths vc_compat_vs; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length sets) vc_compat_ths'; val vc_compat_ths1' = map (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.rewrite (eqvt_ss goal_ctxt2))))) vc_compat_ths1; val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold (obtain_fresh_name ts sets) (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], goal_ctxt2); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold_rev (concat_perm #> map) pis' pis; val ihyp' = inst_params thy vs_ihypt ihyp (map (fold_rev (NominalDatatype.mk_perm []) (pis' @ pis) #> Thm.global_cterm_of thy) params' @ [Thm.global_cterm_of thy z]); fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt'' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (eqvt_ss ctxt'') (fold_rev (mk_perm_bool ctxt'' o Thm.cterm_of ctxt'') (pis' @ pis) th)); val gprems2 = map (fn (th, t) => if null (preds_of ps t) then mk_pi th else mk_pi (the (map_thm ctxt'' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt'' (length pis'')) monos (SOME t) th))) (gprems ~~ oprems); val perm_freshs_freshs' = map (fn (th, (_, T)) => th RS the (AList.lookup op = perm_freshs_freshs T)) (fresh_ths2 ~~ sets); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn {context = goal_ctxt4, ...} => EVERY ([simp_tac (eqvt_ss goal_ctxt4) 1, resolve_tac goal_ctxt4 [ihyp'] 1] @ map (fn th => resolve_tac goal_ctxt4 [th] 1) fresh_ths3 @ [REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss goal_ctxt4 addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac goal_ctxt4 gprems2 1)])); val final = Goal.prove ctxt'' [] [] (Thm.term_of concl) (fn {context = goal_ctxt5, ...} => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss goal_ctxt5 addsimps vc_compat_ths1' @ fresh_ths1 @ perm_freshs_freshs') 1); val final' = Proof_Context.export ctxt'' goal_ctxt2 [final]; in resolve_tac goal_ctxt2 final' 1 end) goal_ctxt1 1]) (prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac goal_ctxt [conjE] 1) THEN REPEAT (REPEAT (resolve_tac goal_ctxt [conjI, impI] 1) THEN eresolve_tac goal_ctxt [impE] 1 THEN assume_tac goal_ctxt 1 THEN REPEAT (eresolve_tac goal_ctxt @{thms allE_Nil} 1) THEN asm_full_simp_tac goal_ctxt 1) end) |> fresh_postprocess ctxt |> singleton (Proof_Context.export ctxt lthy); in ctxt'' |> Proof.theorem NONE (fn thss => fn lthy1 => let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr lthy1)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof lthy1 thss' |> Inductive.rulify lthy1; val strong_induct_atts = - map (Attrib.internal o K) + map (Attrib.internal \<^here> o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val (induct_name, inducts_name) = case alt_name of NONE => (rec_qualified (Binding.name "strong_induct"), rec_qualified (Binding.name "strong_inducts")) | SOME s => (Binding.name s, Binding.name (s ^ "s")); val ((_, [strong_induct']), lthy2) = lthy1 |> Local_Theory.note ((induct_name, strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects lthy2 (1 upto length names) strong_induct' in lthy2 |> Local_Theory.notes [((inducts_name, []), strong_inducts |> map (fn th => ([th], - [Attrib.internal (K ind_case_names), - Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd + [Attrib.internal \<^here> (K ind_case_names), + Attrib.internal \<^here> (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive2\ "prove strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.option (\<^keyword>\(\ |-- Parse.!!! (Parse.name --| \<^keyword>\)\)) -- (Scan.optional (\<^keyword>\avoids\ |-- Parse.enum1 "|" (Parse.name -- (\<^keyword>\:\ |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) => prove_strong_ind name rule_name avoids)); end diff --git a/src/HOL/Nonstandard_Analysis/NSA.thy b/src/HOL/Nonstandard_Analysis/NSA.thy --- a/src/HOL/Nonstandard_Analysis/NSA.thy +++ b/src/HOL/Nonstandard_Analysis/NSA.thy @@ -1,1672 +1,1672 @@ (* Title: HOL/Nonstandard_Analysis/NSA.thy Author: Jacques D. Fleuriot, University of Cambridge Author: Lawrence C Paulson, University of Cambridge *) section \Infinite Numbers, Infinitesimals, Infinitely Close Relation\ theory NSA imports HyperDef "HOL-Library.Lub_Glb" begin definition hnorm :: "'a::real_normed_vector star \ real star" where [transfer_unfold]: "hnorm = *f* norm" definition Infinitesimal :: "('a::real_normed_vector) star set" where "Infinitesimal = {x. \r \ Reals. 0 < r \ hnorm x < r}" definition HFinite :: "('a::real_normed_vector) star set" where "HFinite = {x. \r \ Reals. hnorm x < r}" definition HInfinite :: "('a::real_normed_vector) star set" where "HInfinite = {x. \r \ Reals. r < hnorm x}" definition approx :: "'a::real_normed_vector star \ 'a star \ bool" (infixl "\" 50) where "x \ y \ x - y \ Infinitesimal" \ \the ``infinitely close'' relation\ definition st :: "hypreal \ hypreal" where "st = (\x. SOME r. x \ HFinite \ r \ \ \ r \ x)" \ \the standard part of a hyperreal\ definition monad :: "'a::real_normed_vector star \ 'a star set" where "monad x = {y. x \ y}" definition galaxy :: "'a::real_normed_vector star \ 'a star set" where "galaxy x = {y. (x + -y) \ HFinite}" lemma SReal_def: "\ \ {x. \r. x = hypreal_of_real r}" by (simp add: Reals_def image_def) subsection \Nonstandard Extension of the Norm Function\ definition scaleHR :: "real star \ 'a star \ 'a::real_normed_vector star" where [transfer_unfold]: "scaleHR = starfun2 scaleR" lemma Standard_hnorm [simp]: "x \ Standard \ hnorm x \ Standard" by (simp add: hnorm_def) lemma star_of_norm [simp]: "star_of (norm x) = hnorm (star_of x)" by transfer (rule refl) lemma hnorm_ge_zero [simp]: "\x::'a::real_normed_vector star. 0 \ hnorm x" by transfer (rule norm_ge_zero) lemma hnorm_eq_zero [simp]: "\x::'a::real_normed_vector star. hnorm x = 0 \ x = 0" by transfer (rule norm_eq_zero) lemma hnorm_triangle_ineq: "\x y::'a::real_normed_vector star. hnorm (x + y) \ hnorm x + hnorm y" by transfer (rule norm_triangle_ineq) lemma hnorm_triangle_ineq3: "\x y::'a::real_normed_vector star. \hnorm x - hnorm y\ \ hnorm (x - y)" by transfer (rule norm_triangle_ineq3) lemma hnorm_scaleR: "\x::'a::real_normed_vector star. hnorm (a *\<^sub>R x) = \star_of a\ * hnorm x" by transfer (rule norm_scaleR) lemma hnorm_scaleHR: "\a (x::'a::real_normed_vector star). hnorm (scaleHR a x) = \a\ * hnorm x" by transfer (rule norm_scaleR) lemma hnorm_mult_ineq: "\x y::'a::real_normed_algebra star. hnorm (x * y) \ hnorm x * hnorm y" by transfer (rule norm_mult_ineq) lemma hnorm_mult: "\x y::'a::real_normed_div_algebra star. hnorm (x * y) = hnorm x * hnorm y" by transfer (rule norm_mult) lemma hnorm_hyperpow: "\(x::'a::{real_normed_div_algebra} star) n. hnorm (x pow n) = hnorm x pow n" by transfer (rule norm_power) lemma hnorm_one [simp]: "hnorm (1::'a::real_normed_div_algebra star) = 1" by transfer (rule norm_one) lemma hnorm_zero [simp]: "hnorm (0::'a::real_normed_vector star) = 0" by transfer (rule norm_zero) lemma zero_less_hnorm_iff [simp]: "\x::'a::real_normed_vector star. 0 < hnorm x \ x \ 0" by transfer (rule zero_less_norm_iff) lemma hnorm_minus_cancel [simp]: "\x::'a::real_normed_vector star. hnorm (- x) = hnorm x" by transfer (rule norm_minus_cancel) lemma hnorm_minus_commute: "\a b::'a::real_normed_vector star. hnorm (a - b) = hnorm (b - a)" by transfer (rule norm_minus_commute) lemma hnorm_triangle_ineq2: "\a b::'a::real_normed_vector star. hnorm a - hnorm b \ hnorm (a - b)" by transfer (rule norm_triangle_ineq2) lemma hnorm_triangle_ineq4: "\a b::'a::real_normed_vector star. hnorm (a - b) \ hnorm a + hnorm b" by transfer (rule norm_triangle_ineq4) lemma abs_hnorm_cancel [simp]: "\a::'a::real_normed_vector star. \hnorm a\ = hnorm a" by transfer (rule abs_norm_cancel) lemma hnorm_of_hypreal [simp]: "\r. hnorm (of_hypreal r::'a::real_normed_algebra_1 star) = \r\" by transfer (rule norm_of_real) lemma nonzero_hnorm_inverse: "\a::'a::real_normed_div_algebra star. a \ 0 \ hnorm (inverse a) = inverse (hnorm a)" by transfer (rule nonzero_norm_inverse) lemma hnorm_inverse: "\a::'a::{real_normed_div_algebra, division_ring} star. hnorm (inverse a) = inverse (hnorm a)" by transfer (rule norm_inverse) lemma hnorm_divide: "\a b::'a::{real_normed_field, field} star. hnorm (a / b) = hnorm a / hnorm b" by transfer (rule norm_divide) lemma hypreal_hnorm_def [simp]: "\r::hypreal. hnorm r = \r\" by transfer (rule real_norm_def) lemma hnorm_add_less: "\(x::'a::real_normed_vector star) y r s. hnorm x < r \ hnorm y < s \ hnorm (x + y) < r + s" by transfer (rule norm_add_less) lemma hnorm_mult_less: "\(x::'a::real_normed_algebra star) y r s. hnorm x < r \ hnorm y < s \ hnorm (x * y) < r * s" by transfer (rule norm_mult_less) lemma hnorm_scaleHR_less: "\x\ < r \ hnorm y < s \ hnorm (scaleHR x y) < r * s" by (simp only: hnorm_scaleHR) (simp add: mult_strict_mono') subsection \Closure Laws for the Standard Reals\ lemma Reals_add_cancel: "x + y \ \ \ y \ \ \ x \ \" by (drule (1) Reals_diff) simp lemma SReal_hrabs: "x \ \ \ \x\ \ \" for x :: hypreal by (simp add: Reals_eq_Standard) lemma SReal_hypreal_of_real [simp]: "hypreal_of_real x \ \" by (simp add: Reals_eq_Standard) lemma SReal_divide_numeral: "r \ \ \ r / (numeral w::hypreal) \ \" by simp text \\\\ is not in Reals because it is an infinitesimal\ lemma SReal_epsilon_not_mem: "\ \ \" by (auto simp: SReal_def hypreal_of_real_not_eq_epsilon [symmetric]) lemma SReal_omega_not_mem: "\ \ \" by (auto simp: SReal_def hypreal_of_real_not_eq_omega [symmetric]) lemma SReal_UNIV_real: "{x. hypreal_of_real x \ \} = (UNIV::real set)" by simp lemma SReal_iff: "x \ \ \ (\y. x = hypreal_of_real y)" by (simp add: SReal_def) lemma hypreal_of_real_image: "hypreal_of_real `(UNIV::real set) = \" by (simp add: Reals_eq_Standard Standard_def) lemma inv_hypreal_of_real_image: "inv hypreal_of_real ` \ = UNIV" by (simp add: Reals_eq_Standard Standard_def inj_star_of) lemma SReal_dense: "x \ \ \ y \ \ \ x < y \ \r \ Reals. x < r \ r < y" for x y :: hypreal using dense by (fastforce simp add: SReal_def) subsection \Set of Finite Elements is a Subring of the Extended Reals\ lemma HFinite_add: "x \ HFinite \ y \ HFinite \ x + y \ HFinite" unfolding HFinite_def by (blast intro!: Reals_add hnorm_add_less) lemma HFinite_mult: "x \ HFinite \ y \ HFinite \ x * y \ HFinite" for x y :: "'a::real_normed_algebra star" unfolding HFinite_def by (blast intro!: Reals_mult hnorm_mult_less) lemma HFinite_scaleHR: "x \ HFinite \ y \ HFinite \ scaleHR x y \ HFinite" by (auto simp: HFinite_def intro!: Reals_mult hnorm_scaleHR_less) lemma HFinite_minus_iff: "- x \ HFinite \ x \ HFinite" by (simp add: HFinite_def) lemma HFinite_star_of [simp]: "star_of x \ HFinite" by (simp add: HFinite_def) (metis SReal_hypreal_of_real gt_ex star_of_less star_of_norm) lemma SReal_subset_HFinite: "(\::hypreal set) \ HFinite" by (auto simp add: SReal_def) lemma HFiniteD: "x \ HFinite \ \t \ Reals. hnorm x < t" by (simp add: HFinite_def) lemma HFinite_hrabs_iff [iff]: "\x\ \ HFinite \ x \ HFinite" for x :: hypreal by (simp add: HFinite_def) lemma HFinite_hnorm_iff [iff]: "hnorm x \ HFinite \ x \ HFinite" for x :: hypreal by (simp add: HFinite_def) lemma HFinite_numeral [simp]: "numeral w \ HFinite" unfolding star_numeral_def by (rule HFinite_star_of) text \As always with numerals, \0\ and \1\ are special cases.\ lemma HFinite_0 [simp]: "0 \ HFinite" unfolding star_zero_def by (rule HFinite_star_of) lemma HFinite_1 [simp]: "1 \ HFinite" unfolding star_one_def by (rule HFinite_star_of) lemma hrealpow_HFinite: "x \ HFinite \ x ^ n \ HFinite" for x :: "'a::{real_normed_algebra,monoid_mult} star" by (induct n) (auto intro: HFinite_mult) lemma HFinite_bounded: fixes x y :: hypreal assumes "x \ HFinite" and y: "y \ x" "0 \ y" shows "y \ HFinite" proof (cases "x \ 0") case True then have "y = 0" using y by auto then show ?thesis by simp next case False then show ?thesis using assms le_less_trans by (auto simp: HFinite_def) qed subsection \Set of Infinitesimals is a Subring of the Hyperreals\ lemma InfinitesimalI: "(\r. r \ \ \ 0 < r \ hnorm x < r) \ x \ Infinitesimal" by (simp add: Infinitesimal_def) lemma InfinitesimalD: "x \ Infinitesimal \ \r \ Reals. 0 < r \ hnorm x < r" by (simp add: Infinitesimal_def) lemma InfinitesimalI2: "(\r. 0 < r \ hnorm x < star_of r) \ x \ Infinitesimal" by (auto simp add: Infinitesimal_def SReal_def) lemma InfinitesimalD2: "x \ Infinitesimal \ 0 < r \ hnorm x < star_of r" by (auto simp add: Infinitesimal_def SReal_def) lemma Infinitesimal_zero [iff]: "0 \ Infinitesimal" by (simp add: Infinitesimal_def) lemma Infinitesimal_add: assumes "x \ Infinitesimal" "y \ Infinitesimal" shows "x + y \ Infinitesimal" proof (rule InfinitesimalI) show "hnorm (x + y) < r" if "r \ \" and "0 < r" for r :: "real star" proof - have "hnorm x < r/2" "hnorm y < r/2" using InfinitesimalD SReal_divide_numeral assms half_gt_zero that by blast+ then show ?thesis using hnorm_add_less by fastforce qed qed lemma Infinitesimal_minus_iff [simp]: "- x \ Infinitesimal \ x \ Infinitesimal" by (simp add: Infinitesimal_def) lemma Infinitesimal_hnorm_iff: "hnorm x \ Infinitesimal \ x \ Infinitesimal" by (simp add: Infinitesimal_def) lemma Infinitesimal_hrabs_iff [iff]: "\x\ \ Infinitesimal \ x \ Infinitesimal" for x :: hypreal by (simp add: abs_if) lemma Infinitesimal_of_hypreal_iff [simp]: "(of_hypreal x::'a::real_normed_algebra_1 star) \ Infinitesimal \ x \ Infinitesimal" by (subst Infinitesimal_hnorm_iff [symmetric]) simp lemma Infinitesimal_diff: "x \ Infinitesimal \ y \ Infinitesimal \ x - y \ Infinitesimal" using Infinitesimal_add [of x "- y"] by simp lemma Infinitesimal_mult: fixes x y :: "'a::real_normed_algebra star" assumes "x \ Infinitesimal" "y \ Infinitesimal" shows "x * y \ Infinitesimal" proof (rule InfinitesimalI) show "hnorm (x * y) < r" if "r \ \" and "0 < r" for r :: "real star" proof - have "hnorm x < 1" "hnorm y < r" using assms that by (auto simp add: InfinitesimalD) then show ?thesis using hnorm_mult_less by fastforce qed qed lemma Infinitesimal_HFinite_mult: fixes x y :: "'a::real_normed_algebra star" assumes "x \ Infinitesimal" "y \ HFinite" shows "x * y \ Infinitesimal" proof (rule InfinitesimalI) obtain t where "hnorm y < t" "t \ Reals" using HFiniteD \y \ HFinite\ by blast then have "t > 0" using hnorm_ge_zero le_less_trans by blast show "hnorm (x * y) < r" if "r \ \" and "0 < r" for r :: "real star" proof - have "hnorm x < r/t" by (meson InfinitesimalD Reals_divide \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero le_less_trans that) then have "hnorm (x * y) < (r / t) * t" using \hnorm y < t\ hnorm_mult_less by blast then show ?thesis using \0 < t\ by auto qed qed lemma Infinitesimal_HFinite_scaleHR: assumes "x \ Infinitesimal" "y \ HFinite" shows "scaleHR x y \ Infinitesimal" proof (rule InfinitesimalI) obtain t where "hnorm y < t" "t \ Reals" using HFiniteD \y \ HFinite\ by blast then have "t > 0" using hnorm_ge_zero le_less_trans by blast show "hnorm (scaleHR x y) < r" if "r \ \" and "0 < r" for r :: "real star" proof - have "\x\ * hnorm y < (r / t) * t" by (metis InfinitesimalD Reals_divide \0 < t\ \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero hypreal_hnorm_def mult_strict_mono' that) then show ?thesis by (simp add: \0 < t\ hnorm_scaleHR less_imp_not_eq2) qed qed lemma Infinitesimal_HFinite_mult2: fixes x y :: "'a::real_normed_algebra star" assumes "x \ Infinitesimal" "y \ HFinite" shows "y * x \ Infinitesimal" proof (rule InfinitesimalI) obtain t where "hnorm y < t" "t \ Reals" using HFiniteD \y \ HFinite\ by blast then have "t > 0" using hnorm_ge_zero le_less_trans by blast show "hnorm (y * x) < r" if "r \ \" and "0 < r" for r :: "real star" proof - have "hnorm x < r/t" by (meson InfinitesimalD Reals_divide \hnorm y < t\ \t \ \\ assms(1) divide_pos_pos hnorm_ge_zero le_less_trans that) then have "hnorm (y * x) < t * (r / t)" using \hnorm y < t\ hnorm_mult_less by blast then show ?thesis using \0 < t\ by auto qed qed lemma Infinitesimal_scaleR2: assumes "x \ Infinitesimal" shows "a *\<^sub>R x \ Infinitesimal" by (metis HFinite_star_of Infinitesimal_HFinite_mult2 Infinitesimal_hnorm_iff assms hnorm_scaleR hypreal_hnorm_def star_of_norm) lemma Compl_HFinite: "- HFinite = HInfinite" proof - have "r < hnorm x" if *: "\s. s \ \ \ s \ hnorm x" and "r \ \" for x :: "'a star" and r :: hypreal using * [of "r+1"] \r \ \\ by auto then show ?thesis by (auto simp add: HInfinite_def HFinite_def linorder_not_less) qed lemma HInfinite_inverse_Infinitesimal: "x \ HInfinite \ inverse x \ Infinitesimal" for x :: "'a::real_normed_div_algebra star" by (simp add: HInfinite_def InfinitesimalI hnorm_inverse inverse_less_imp_less) lemma inverse_Infinitesimal_iff_HInfinite: "x \ 0 \ inverse x \ Infinitesimal \ x \ HInfinite" for x :: "'a::real_normed_div_algebra star" by (metis Compl_HFinite Compl_iff HInfinite_inverse_Infinitesimal InfinitesimalD Infinitesimal_HFinite_mult Reals_1 hnorm_one left_inverse less_irrefl zero_less_one) lemma HInfiniteI: "(\r. r \ \ \ r < hnorm x) \ x \ HInfinite" by (simp add: HInfinite_def) lemma HInfiniteD: "x \ HInfinite \ r \ \ \ r < hnorm x" by (simp add: HInfinite_def) lemma HInfinite_mult: fixes x y :: "'a::real_normed_div_algebra star" assumes "x \ HInfinite" "y \ HInfinite" shows "x * y \ HInfinite" proof (rule HInfiniteI, simp only: hnorm_mult) have "x \ 0" using Compl_HFinite HFinite_0 assms by blast show "r < hnorm x * hnorm y" if "r \ \" for r :: "real star" proof - have "r = r * 1" by simp also have "\ < hnorm x * hnorm y" by (meson HInfiniteD Reals_1 \x \ 0\ assms le_numeral_extra(1) mult_strict_mono that zero_less_hnorm_iff) finally show ?thesis . qed qed lemma hypreal_add_zero_less_le_mono: "r < x \ 0 \ y \ r < x + y" for r x y :: hypreal by simp lemma HInfinite_add_ge_zero: "x \ HInfinite \ 0 \ y \ 0 \ x \ x + y \ HInfinite" for x y :: hypreal by (auto simp: abs_if add.commute HInfinite_def) lemma HInfinite_add_ge_zero2: "x \ HInfinite \ 0 \ y \ 0 \ x \ y + x \ HInfinite" for x y :: hypreal by (auto intro!: HInfinite_add_ge_zero simp add: add.commute) lemma HInfinite_add_gt_zero: "x \ HInfinite \ 0 < y \ 0 < x \ x + y \ HInfinite" for x y :: hypreal by (blast intro: HInfinite_add_ge_zero order_less_imp_le) lemma HInfinite_minus_iff: "- x \ HInfinite \ x \ HInfinite" by (simp add: HInfinite_def) lemma HInfinite_add_le_zero: "x \ HInfinite \ y \ 0 \ x \ 0 \ x + y \ HInfinite" for x y :: hypreal by (metis (no_types, lifting) HInfinite_add_ge_zero2 HInfinite_minus_iff add.inverse_distrib_swap neg_0_le_iff_le) lemma HInfinite_add_lt_zero: "x \ HInfinite \ y < 0 \ x < 0 \ x + y \ HInfinite" for x y :: hypreal by (blast intro: HInfinite_add_le_zero order_less_imp_le) lemma not_Infinitesimal_not_zero: "x \ Infinitesimal \ x \ 0" by auto lemma HFinite_diff_Infinitesimal_hrabs: "x \ HFinite - Infinitesimal \ \x\ \ HFinite - Infinitesimal" for x :: hypreal by blast lemma hnorm_le_Infinitesimal: "e \ Infinitesimal \ hnorm x \ e \ x \ Infinitesimal" by (auto simp: Infinitesimal_def abs_less_iff) lemma hnorm_less_Infinitesimal: "e \ Infinitesimal \ hnorm x < e \ x \ Infinitesimal" by (erule hnorm_le_Infinitesimal, erule order_less_imp_le) lemma hrabs_le_Infinitesimal: "e \ Infinitesimal \ \x\ \ e \ x \ Infinitesimal" for x :: hypreal by (erule hnorm_le_Infinitesimal) simp lemma hrabs_less_Infinitesimal: "e \ Infinitesimal \ \x\ < e \ x \ Infinitesimal" for x :: hypreal by (erule hnorm_less_Infinitesimal) simp lemma Infinitesimal_interval: "e \ Infinitesimal \ e' \ Infinitesimal \ e' < x \ x < e \ x \ Infinitesimal" for x :: hypreal by (auto simp add: Infinitesimal_def abs_less_iff) lemma Infinitesimal_interval2: "e \ Infinitesimal \ e' \ Infinitesimal \ e' \ x \ x \ e \ x \ Infinitesimal" for x :: hypreal by (auto intro: Infinitesimal_interval simp add: order_le_less) lemma lemma_Infinitesimal_hyperpow: "x \ Infinitesimal \ 0 < N \ \x pow N\ \ \x\" for x :: hypreal apply (clarsimp simp: Infinitesimal_def) by (metis Reals_1 abs_ge_zero hyperpow_Suc_le_self2 hyperpow_hrabs hypnat_gt_zero_iff2 zero_less_one) lemma Infinitesimal_hyperpow: "x \ Infinitesimal \ 0 < N \ x pow N \ Infinitesimal" for x :: hypreal using hrabs_le_Infinitesimal lemma_Infinitesimal_hyperpow by blast lemma hrealpow_hyperpow_Infinitesimal_iff: "(x ^ n \ Infinitesimal) \ x pow (hypnat_of_nat n) \ Infinitesimal" by (simp only: hyperpow_hypnat_of_nat) lemma Infinitesimal_hrealpow: "x \ Infinitesimal \ 0 < n \ x ^ n \ Infinitesimal" for x :: hypreal by (simp add: hrealpow_hyperpow_Infinitesimal_iff Infinitesimal_hyperpow) lemma not_Infinitesimal_mult: "x \ Infinitesimal \ y \ Infinitesimal \ x * y \ Infinitesimal" for x y :: "'a::real_normed_div_algebra star" by (metis (no_types, lifting) inverse_Infinitesimal_iff_HInfinite ComplI Compl_HFinite Infinitesimal_HFinite_mult divide_inverse eq_divide_imp inverse_inverse_eq mult_zero_right) lemma Infinitesimal_mult_disj: "x * y \ Infinitesimal \ x \ Infinitesimal \ y \ Infinitesimal" for x y :: "'a::real_normed_div_algebra star" using not_Infinitesimal_mult by blast lemma HFinite_Infinitesimal_not_zero: "x \ HFinite-Infinitesimal \ x \ 0" by blast lemma HFinite_Infinitesimal_diff_mult: "x \ HFinite - Infinitesimal \ y \ HFinite - Infinitesimal \ x * y \ HFinite - Infinitesimal" for x y :: "'a::real_normed_div_algebra star" by (simp add: HFinite_mult not_Infinitesimal_mult) lemma Infinitesimal_subset_HFinite: "Infinitesimal \ HFinite" using HFinite_def InfinitesimalD Reals_1 zero_less_one by blast lemma Infinitesimal_star_of_mult: "x \ Infinitesimal \ x * star_of r \ Infinitesimal" for x :: "'a::real_normed_algebra star" by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult]) lemma Infinitesimal_star_of_mult2: "x \ Infinitesimal \ star_of r * x \ Infinitesimal" for x :: "'a::real_normed_algebra star" by (erule HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult2]) subsection \The Infinitely Close Relation\ lemma mem_infmal_iff: "x \ Infinitesimal \ x \ 0" by (simp add: Infinitesimal_def approx_def) lemma approx_minus_iff: "x \ y \ x - y \ 0" by (simp add: approx_def) lemma approx_minus_iff2: "x \ y \ - y + x \ 0" by (simp add: approx_def add.commute) lemma approx_refl [iff]: "x \ x" by (simp add: approx_def Infinitesimal_def) lemma approx_sym: "x \ y \ y \ x" by (metis Infinitesimal_minus_iff approx_def minus_diff_eq) lemma approx_trans: assumes "x \ y" "y \ z" shows "x \ z" proof - have "x - y \ Infinitesimal" "z - y \ Infinitesimal" using assms approx_def approx_sym by auto then have "x - z \ Infinitesimal" using Infinitesimal_diff by force then show ?thesis by (simp add: approx_def) qed lemma approx_trans2: "r \ x \ s \ x \ r \ s" by (blast intro: approx_sym approx_trans) lemma approx_trans3: "x \ r \ x \ s \ r \ s" by (blast intro: approx_sym approx_trans) lemma approx_reorient: "x \ y \ y \ x" by (blast intro: approx_sym) text \Reorientation simplification procedure: reorients (polymorphic) \0 = x\, \1 = x\, \nnn = x\ provided \x\ isn't \0\, \1\ or a numeral.\ simproc_setup approx_reorient_simproc ("0 \ x" | "1 \ y" | "numeral w \ z" | "- 1 \ y" | "- numeral w \ r") = \ let val rule = @{thm approx_reorient} RS eq_reflection - fun proc phi ss ct = + fun proc ct = case Thm.term_of ct of _ $ t $ u => if can HOLogic.dest_number u then NONE else if can HOLogic.dest_number t then SOME rule else NONE | _ => NONE - in proc end + in K (K proc) end \ lemma Infinitesimal_approx_minus: "x - y \ Infinitesimal \ x \ y" by (simp add: approx_minus_iff [symmetric] mem_infmal_iff) lemma approx_monad_iff: "x \ y \ monad x = monad y" apply (simp add: monad_def set_eq_iff) using approx_reorient approx_trans by blast lemma Infinitesimal_approx: "x \ Infinitesimal \ y \ Infinitesimal \ x \ y" by (simp add: Infinitesimal_diff approx_def) lemma approx_add: "a \ b \ c \ d \ a + c \ b + d" proof (unfold approx_def) assume inf: "a - b \ Infinitesimal" "c - d \ Infinitesimal" have "a + c - (b + d) = (a - b) + (c - d)" by simp also have "... \ Infinitesimal" using inf by (rule Infinitesimal_add) finally show "a + c - (b + d) \ Infinitesimal" . qed lemma approx_minus: "a \ b \ - a \ - b" by (metis approx_def approx_sym minus_diff_eq minus_diff_minus) lemma approx_minus2: "- a \ - b \ a \ b" by (auto dest: approx_minus) lemma approx_minus_cancel [simp]: "- a \ - b \ a \ b" by (blast intro: approx_minus approx_minus2) lemma approx_add_minus: "a \ b \ c \ d \ a + - c \ b + - d" by (blast intro!: approx_add approx_minus) lemma approx_diff: "a \ b \ c \ d \ a - c \ b - d" using approx_add [of a b "- c" "- d"] by simp lemma approx_mult1: "a \ b \ c \ HFinite \ a * c \ b * c" for a b c :: "'a::real_normed_algebra star" by (simp add: approx_def Infinitesimal_HFinite_mult left_diff_distrib [symmetric]) lemma approx_mult2: "a \ b \ c \ HFinite \ c * a \ c * b" for a b c :: "'a::real_normed_algebra star" by (simp add: approx_def Infinitesimal_HFinite_mult2 right_diff_distrib [symmetric]) lemma approx_mult_subst: "u \ v * x \ x \ y \ v \ HFinite \ u \ v * y" for u v x y :: "'a::real_normed_algebra star" by (blast intro: approx_mult2 approx_trans) lemma approx_mult_subst2: "u \ x * v \ x \ y \ v \ HFinite \ u \ y * v" for u v x y :: "'a::real_normed_algebra star" by (blast intro: approx_mult1 approx_trans) lemma approx_mult_subst_star_of: "u \ x * star_of v \ x \ y \ u \ y * star_of v" for u x y :: "'a::real_normed_algebra star" by (auto intro: approx_mult_subst2) lemma approx_eq_imp: "a = b \ a \ b" by (simp add: approx_def) lemma Infinitesimal_minus_approx: "x \ Infinitesimal \ - x \ x" by (blast intro: Infinitesimal_minus_iff [THEN iffD2] mem_infmal_iff [THEN iffD1] approx_trans2) lemma bex_Infinitesimal_iff: "(\y \ Infinitesimal. x - z = y) \ x \ z" by (simp add: approx_def) lemma bex_Infinitesimal_iff2: "(\y \ Infinitesimal. x = z + y) \ x \ z" by (force simp add: bex_Infinitesimal_iff [symmetric]) lemma Infinitesimal_add_approx: "y \ Infinitesimal \ x + y = z \ x \ z" using approx_sym bex_Infinitesimal_iff2 by blast lemma Infinitesimal_add_approx_self: "y \ Infinitesimal \ x \ x + y" by (simp add: Infinitesimal_add_approx) lemma Infinitesimal_add_approx_self2: "y \ Infinitesimal \ x \ y + x" by (auto dest: Infinitesimal_add_approx_self simp add: add.commute) lemma Infinitesimal_add_minus_approx_self: "y \ Infinitesimal \ x \ x + - y" by (blast intro!: Infinitesimal_add_approx_self Infinitesimal_minus_iff [THEN iffD2]) lemma Infinitesimal_add_cancel: "y \ Infinitesimal \ x + y \ z \ x \ z" using Infinitesimal_add_approx approx_trans by blast lemma Infinitesimal_add_right_cancel: "y \ Infinitesimal \ x \ z + y \ x \ z" by (metis Infinitesimal_add_approx_self approx_monad_iff) lemma approx_add_left_cancel: "d + b \ d + c \ b \ c" by (metis add_diff_cancel_left bex_Infinitesimal_iff) lemma approx_add_right_cancel: "b + d \ c + d \ b \ c" by (simp add: approx_def) lemma approx_add_mono1: "b \ c \ d + b \ d + c" by (simp add: approx_add) lemma approx_add_mono2: "b \ c \ b + a \ c + a" by (simp add: add.commute approx_add_mono1) lemma approx_add_left_iff [simp]: "a + b \ a + c \ b \ c" by (fast elim: approx_add_left_cancel approx_add_mono1) lemma approx_add_right_iff [simp]: "b + a \ c + a \ b \ c" by (simp add: add.commute) lemma approx_HFinite: "x \ HFinite \ x \ y \ y \ HFinite" by (metis HFinite_add Infinitesimal_subset_HFinite approx_sym subsetD bex_Infinitesimal_iff2) lemma approx_star_of_HFinite: "x \ star_of D \ x \ HFinite" by (rule approx_sym [THEN [2] approx_HFinite], auto) lemma approx_mult_HFinite: "a \ b \ c \ d \ b \ HFinite \ d \ HFinite \ a * c \ b * d" for a b c d :: "'a::real_normed_algebra star" by (meson approx_HFinite approx_mult2 approx_mult_subst2 approx_sym) lemma scaleHR_left_diff_distrib: "\a b x. scaleHR (a - b) x = scaleHR a x - scaleHR b x" by transfer (rule scaleR_left_diff_distrib) lemma approx_scaleR1: "a \ star_of b \ c \ HFinite \ scaleHR a c \ b *\<^sub>R c" unfolding approx_def by (metis Infinitesimal_HFinite_scaleHR scaleHR_def scaleHR_left_diff_distrib star_scaleR_def starfun2_star_of) lemma approx_scaleR2: "a \ b \ c *\<^sub>R a \ c *\<^sub>R b" by (simp add: approx_def Infinitesimal_scaleR2 scaleR_right_diff_distrib [symmetric]) lemma approx_scaleR_HFinite: "a \ star_of b \ c \ d \ d \ HFinite \ scaleHR a c \ b *\<^sub>R d" by (meson approx_HFinite approx_scaleR1 approx_scaleR2 approx_sym approx_trans) lemma approx_mult_star_of: "a \ star_of b \ c \ star_of d \ a * c \ star_of b * star_of d" for a c :: "'a::real_normed_algebra star" by (blast intro!: approx_mult_HFinite approx_star_of_HFinite HFinite_star_of) lemma approx_SReal_mult_cancel_zero: fixes a x :: hypreal assumes "a \ \" "a \ 0" "a * x \ 0" shows "x \ 0" proof - have "inverse a \ HFinite" using Reals_inverse SReal_subset_HFinite assms(1) by blast then show ?thesis using assms by (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) qed lemma approx_mult_SReal1: "a \ \ \ x \ 0 \ x * a \ 0" for a x :: hypreal by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult1) lemma approx_mult_SReal2: "a \ \ \ x \ 0 \ a * x \ 0" for a x :: hypreal by (auto dest: SReal_subset_HFinite [THEN subsetD] approx_mult2) lemma approx_mult_SReal_zero_cancel_iff [simp]: "a \ \ \ a \ 0 \ a * x \ 0 \ x \ 0" for a x :: hypreal by (blast intro: approx_SReal_mult_cancel_zero approx_mult_SReal2) lemma approx_SReal_mult_cancel: fixes a w z :: hypreal assumes "a \ \" "a \ 0" "a * w \ a * z" shows "w \ z" proof - have "inverse a \ HFinite" using Reals_inverse SReal_subset_HFinite assms(1) by blast then show ?thesis using assms by (auto dest: approx_mult2 simp add: mult.assoc [symmetric]) qed lemma approx_SReal_mult_cancel_iff1 [simp]: "a \ \ \ a \ 0 \ a * w \ a * z \ w \ z" for a w z :: hypreal by (meson SReal_subset_HFinite approx_SReal_mult_cancel approx_mult2 subsetD) lemma approx_le_bound: fixes z :: hypreal assumes "z \ f" " f \ g" "g \ z" shows "f \ z" proof - obtain y where "z \ g + y" and "y \ Infinitesimal" "f = g + y" using assms bex_Infinitesimal_iff2 by auto then have "z - g \ Infinitesimal" using assms(3) hrabs_le_Infinitesimal by auto then show ?thesis by (metis approx_def approx_trans2 assms(2)) qed lemma approx_hnorm: "x \ y \ hnorm x \ hnorm y" for x y :: "'a::real_normed_vector star" proof (unfold approx_def) assume "x - y \ Infinitesimal" then have "hnorm (x - y) \ Infinitesimal" by (simp only: Infinitesimal_hnorm_iff) moreover have "(0::real star) \ Infinitesimal" by (rule Infinitesimal_zero) moreover have "0 \ \hnorm x - hnorm y\" by (rule abs_ge_zero) moreover have "\hnorm x - hnorm y\ \ hnorm (x - y)" by (rule hnorm_triangle_ineq3) ultimately have "\hnorm x - hnorm y\ \ Infinitesimal" by (rule Infinitesimal_interval2) then show "hnorm x - hnorm y \ Infinitesimal" by (simp only: Infinitesimal_hrabs_iff) qed subsection \Zero is the Only Infinitesimal that is also a Real\ lemma Infinitesimal_less_SReal: "x \ \ \ y \ Infinitesimal \ 0 < x \ y < x" for x y :: hypreal using InfinitesimalD by fastforce lemma Infinitesimal_less_SReal2: "y \ Infinitesimal \ \r \ Reals. 0 < r \ y < r" for y :: hypreal by (blast intro: Infinitesimal_less_SReal) lemma SReal_not_Infinitesimal: "0 < y \ y \ \ ==> y \ Infinitesimal" for y :: hypreal by (auto simp add: Infinitesimal_def abs_if) lemma SReal_minus_not_Infinitesimal: "y < 0 \ y \ \ \ y \ Infinitesimal" for y :: hypreal using Infinitesimal_minus_iff Reals_minus SReal_not_Infinitesimal neg_0_less_iff_less by blast lemma SReal_Int_Infinitesimal_zero: "\ Int Infinitesimal = {0::hypreal}" proof - have "x = 0" if "x \ \" "x \ Infinitesimal" for x :: "real star" using that SReal_minus_not_Infinitesimal SReal_not_Infinitesimal not_less_iff_gr_or_eq by blast then show ?thesis by auto qed lemma SReal_Infinitesimal_zero: "x \ \ \ x \ Infinitesimal \ x = 0" for x :: hypreal using SReal_Int_Infinitesimal_zero by blast lemma SReal_HFinite_diff_Infinitesimal: "x \ \ \ x \ 0 \ x \ HFinite - Infinitesimal" for x :: hypreal by (auto dest: SReal_Infinitesimal_zero SReal_subset_HFinite [THEN subsetD]) lemma hypreal_of_real_HFinite_diff_Infinitesimal: "hypreal_of_real x \ 0 \ hypreal_of_real x \ HFinite - Infinitesimal" by (rule SReal_HFinite_diff_Infinitesimal) auto lemma star_of_Infinitesimal_iff_0 [iff]: "star_of x \ Infinitesimal \ x = 0" proof show "x = 0" if "star_of x \ Infinitesimal" proof - have "hnorm (star_n (\n. x)) \ Standard" by (metis Reals_eq_Standard SReal_iff star_of_def star_of_norm) then show ?thesis by (metis InfinitesimalD2 less_irrefl star_of_norm that zero_less_norm_iff) qed qed auto lemma star_of_HFinite_diff_Infinitesimal: "x \ 0 \ star_of x \ HFinite - Infinitesimal" by simp lemma numeral_not_Infinitesimal [simp]: "numeral w \ (0::hypreal) \ (numeral w :: hypreal) \ Infinitesimal" by (fast dest: Reals_numeral [THEN SReal_Infinitesimal_zero]) text \Again: \1\ is a special case, but not \0\ this time.\ lemma one_not_Infinitesimal [simp]: "(1::'a::{real_normed_vector,zero_neq_one} star) \ Infinitesimal" by (metis star_of_Infinitesimal_iff_0 star_one_def zero_neq_one) lemma approx_SReal_not_zero: "y \ \ \ x \ y \ y \ 0 \ x \ 0" for x y :: hypreal using SReal_Infinitesimal_zero approx_sym mem_infmal_iff by auto lemma HFinite_diff_Infinitesimal_approx: "x \ y \ y \ HFinite - Infinitesimal \ x \ HFinite - Infinitesimal" by (meson Diff_iff approx_HFinite approx_sym approx_trans3 mem_infmal_iff) text \The premise \y \ 0\ is essential; otherwise \x / y = 0\ and we lose the \HFinite\ premise.\ lemma Infinitesimal_ratio: "y \ 0 \ y \ Infinitesimal \ x / y \ HFinite \ x \ Infinitesimal" for x y :: "'a::{real_normed_div_algebra,field} star" using Infinitesimal_HFinite_mult by fastforce lemma Infinitesimal_SReal_divide: "x \ Infinitesimal \ y \ \ \ x / y \ Infinitesimal" for x y :: hypreal by (metis HFinite_star_of Infinitesimal_HFinite_mult Reals_inverse SReal_iff divide_inverse) section \Standard Part Theorem\ text \ Every finite \x \ R*\ is infinitely close to a unique real number (i.e. a member of \Reals\). \ subsection \Uniqueness: Two Infinitely Close Reals are Equal\ lemma star_of_approx_iff [simp]: "star_of x \ star_of y \ x = y" by (metis approx_def right_minus_eq star_of_Infinitesimal_iff_0 star_of_simps(2)) lemma SReal_approx_iff: "x \ \ \ y \ \ \ x \ y \ x = y" for x y :: hypreal by (meson Reals_diff SReal_Infinitesimal_zero approx_def approx_refl right_minus_eq) lemma numeral_approx_iff [simp]: "(numeral v \ (numeral w :: 'a::{numeral,real_normed_vector} star)) = (numeral v = (numeral w :: 'a))" by (metis star_of_approx_iff star_of_numeral) text \And also for \0 \ #nn\ and \1 \ #nn\, \#nn \ 0\ and \#nn \ 1\.\ lemma [simp]: "(numeral w \ (0::'a::{numeral,real_normed_vector} star)) = (numeral w = (0::'a))" "((0::'a::{numeral,real_normed_vector} star) \ numeral w) = (numeral w = (0::'a))" "(numeral w \ (1::'b::{numeral,one,real_normed_vector} star)) = (numeral w = (1::'b))" "((1::'b::{numeral,one,real_normed_vector} star) \ numeral w) = (numeral w = (1::'b))" "\ (0 \ (1::'c::{zero_neq_one,real_normed_vector} star))" "\ (1 \ (0::'c::{zero_neq_one,real_normed_vector} star))" unfolding star_numeral_def star_zero_def star_one_def star_of_approx_iff by (auto intro: sym) lemma star_of_approx_numeral_iff [simp]: "star_of k \ numeral w \ k = numeral w" by (subst star_of_approx_iff [symmetric]) auto lemma star_of_approx_zero_iff [simp]: "star_of k \ 0 \ k = 0" by (simp_all add: star_of_approx_iff [symmetric]) lemma star_of_approx_one_iff [simp]: "star_of k \ 1 \ k = 1" by (simp_all add: star_of_approx_iff [symmetric]) lemma approx_unique_real: "r \ \ \ s \ \ \ r \ x \ s \ x \ r = s" for r s :: hypreal by (blast intro: SReal_approx_iff [THEN iffD1] approx_trans2) subsection \Existence of Unique Real Infinitely Close\ subsubsection \Lifting of the Ub and Lub Properties\ lemma hypreal_of_real_isUb_iff: "isUb \ (hypreal_of_real ` Q) (hypreal_of_real Y) = isUb UNIV Q Y" for Q :: "real set" and Y :: real by (simp add: isUb_def setle_def) lemma hypreal_of_real_isLub_iff: "isLub \ (hypreal_of_real ` Q) (hypreal_of_real Y) = isLub (UNIV :: real set) Q Y" (is "?lhs = ?rhs") for Q :: "real set" and Y :: real proof assume ?lhs then show ?rhs by (simp add: isLub_def leastP_def) (metis hypreal_of_real_isUb_iff mem_Collect_eq setge_def star_of_le) next assume ?rhs then show ?lhs apply (simp add: isLub_def leastP_def hypreal_of_real_isUb_iff setge_def) by (metis SReal_iff hypreal_of_real_isUb_iff isUb_def star_of_le) qed lemma lemma_isUb_hypreal_of_real: "isUb \ P Y \ \Yo. isUb \ P (hypreal_of_real Yo)" by (auto simp add: SReal_iff isUb_def) lemma lemma_isLub_hypreal_of_real: "isLub \ P Y \ \Yo. isLub \ P (hypreal_of_real Yo)" by (auto simp add: isLub_def leastP_def isUb_def SReal_iff) lemma SReal_complete: fixes P :: "hypreal set" assumes "isUb \ P Y" "P \ \" "P \ {}" shows "\t. isLub \ P t" proof - obtain Q where "P = hypreal_of_real ` Q" by (metis \P \ \\ hypreal_of_real_image subset_imageE) then show ?thesis by (metis assms(1) \P \ {}\ equals0I hypreal_of_real_isLub_iff hypreal_of_real_isUb_iff image_empty lemma_isUb_hypreal_of_real reals_complete) qed text \Lemmas about lubs.\ lemma lemma_st_part_lub: fixes x :: hypreal assumes "x \ HFinite" shows "\t. isLub \ {s. s \ \ \ s < x} t" proof - obtain t where t: "t \ \" "hnorm x < t" using HFiniteD assms by blast then have "isUb \ {s. s \ \ \ s < x} t" by (simp add: abs_less_iff isUbI le_less_linear less_imp_not_less setleI) moreover have "\y. y \ \ \ y < x" using t by (rule_tac x = "-t" in exI) (auto simp add: abs_less_iff) ultimately show ?thesis using SReal_complete by fastforce qed lemma hypreal_setle_less_trans: "S *<= x \ x < y \ S *<= y" for x y :: hypreal by (meson le_less_trans less_imp_le setle_def) lemma hypreal_gt_isUb: "isUb R S x \ x < y \ y \ R \ isUb R S y" for x y :: hypreal using hypreal_setle_less_trans isUb_def by blast lemma lemma_SReal_ub: "x \ \ \ isUb \ {s. s \ \ \ s < x} x" for x :: hypreal by (auto intro: isUbI setleI order_less_imp_le) lemma lemma_SReal_lub: fixes x :: hypreal assumes "x \ \" shows "isLub \ {s. s \ \ \ s < x} x" proof - have "x \ y" if "isUb \ {s \ \. s < x} y" for y proof - have "y \ \" using isUbD2a that by blast show ?thesis proof (cases x y rule: linorder_cases) case greater then obtain r where "y < r" "r < x" using dense by blast then show ?thesis using isUbD [OF that] by simp (meson SReal_dense \y \ \\ assms greater not_le) qed auto qed with assms show ?thesis by (simp add: isLubI2 isUbI setgeI setleI) qed lemma lemma_st_part_major: fixes x r t :: hypreal assumes x: "x \ HFinite" and r: "r \ \" "0 < r" and t: "isLub \ {s. s \ \ \ s < x} t" shows "\x - t\ < r" proof - have "t \ \" using isLubD1a t by blast have lemma_st_part_gt_ub: "x < r \ r \ \ \ isUb \ {s. s \ \ \ s < x} r" for r :: hypreal by (auto dest: order_less_trans intro: order_less_imp_le intro!: isUbI setleI) have "isUb \ {s \ \. s < x} t" by (simp add: t isLub_isUb) then have "\ r + t < x" by (metis (mono_tags, lifting) Reals_add \t \ \\ add_le_same_cancel2 isUbD leD mem_Collect_eq r) then have "x - t \ r" by simp moreover have "\ x < t - r" using lemma_st_part_gt_ub isLub_le_isUb \t \ \\ r t x by fastforce then have "- (x - t) \ r" by linarith moreover have False if "x - t = r \ - (x - t) = r" proof - have "x \ \" by (metis \t \ \\ \r \ \\ that Reals_add_cancel Reals_minus_iff add_uminus_conv_diff) then have "isLub \ {s \ \. s < x} x" by (rule lemma_SReal_lub) then show False using r t that x isLub_unique by force qed ultimately show ?thesis using abs_less_iff dual_order.order_iff_strict by blast qed lemma lemma_st_part_major2: "x \ HFinite \ isLub \ {s. s \ \ \ s < x} t \ \r \ Reals. 0 < r \ \x - t\ < r" for x t :: hypreal by (blast dest!: lemma_st_part_major) text\Existence of real and Standard Part Theorem.\ lemma lemma_st_part_Ex: "x \ HFinite \ \t\Reals. \r \ Reals. 0 < r \ \x - t\ < r" for x :: hypreal by (meson isLubD1a lemma_st_part_lub lemma_st_part_major2) lemma st_part_Ex: "x \ HFinite \ \t\Reals. x \ t" for x :: hypreal by (metis InfinitesimalI approx_def hypreal_hnorm_def lemma_st_part_Ex) text \There is a unique real infinitely close.\ lemma st_part_Ex1: "x \ HFinite \ \!t::hypreal. t \ \ \ x \ t" by (meson SReal_approx_iff approx_trans2 st_part_Ex) subsection \Finite, Infinite and Infinitesimal\ lemma HFinite_Int_HInfinite_empty [simp]: "HFinite Int HInfinite = {}" using Compl_HFinite by blast lemma HFinite_not_HInfinite: assumes x: "x \ HFinite" shows "x \ HInfinite" using Compl_HFinite x by blast lemma not_HFinite_HInfinite: "x \ HFinite \ x \ HInfinite" using Compl_HFinite by blast lemma HInfinite_HFinite_disj: "x \ HInfinite \ x \ HFinite" by (blast intro: not_HFinite_HInfinite) lemma HInfinite_HFinite_iff: "x \ HInfinite \ x \ HFinite" by (blast dest: HFinite_not_HInfinite not_HFinite_HInfinite) lemma HFinite_HInfinite_iff: "x \ HFinite \ x \ HInfinite" by (simp add: HInfinite_HFinite_iff) lemma HInfinite_diff_HFinite_Infinitesimal_disj: "x \ Infinitesimal \ x \ HInfinite \ x \ HFinite - Infinitesimal" by (fast intro: not_HFinite_HInfinite) lemma HFinite_inverse: "x \ HFinite \ x \ Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" using HInfinite_inverse_Infinitesimal not_HFinite_HInfinite by force lemma HFinite_inverse2: "x \ HFinite - Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" by (blast intro: HFinite_inverse) text \Stronger statement possible in fact.\ lemma Infinitesimal_inverse_HFinite: "x \ Infinitesimal \ inverse x \ HFinite" for x :: "'a::real_normed_div_algebra star" using HFinite_HInfinite_iff HInfinite_inverse_Infinitesimal by fastforce lemma HFinite_not_Infinitesimal_inverse: "x \ HFinite - Infinitesimal \ inverse x \ HFinite - Infinitesimal" for x :: "'a::real_normed_div_algebra star" using HFinite_Infinitesimal_not_zero HFinite_inverse2 Infinitesimal_HFinite_mult2 by fastforce lemma approx_inverse: fixes x y :: "'a::real_normed_div_algebra star" assumes "x \ y" and y: "y \ HFinite - Infinitesimal" shows "inverse x \ inverse y" proof - have x: "x \ HFinite - Infinitesimal" using HFinite_diff_Infinitesimal_approx assms(1) y by blast with y HFinite_inverse2 have "inverse x \ HFinite" "inverse y \ HFinite" by blast+ then have "inverse y * x \ 1" by (metis Diff_iff approx_mult2 assms(1) left_inverse not_Infinitesimal_not_zero y) then show ?thesis by (metis (no_types, lifting) DiffD2 HFinite_Infinitesimal_not_zero Infinitesimal_mult_disj x approx_def approx_sym left_diff_distrib left_inverse) qed (*Used for NSLIM_inverse, NSLIMSEQ_inverse*) lemmas star_of_approx_inverse = star_of_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] lemmas hypreal_of_real_approx_inverse = hypreal_of_real_HFinite_diff_Infinitesimal [THEN [2] approx_inverse] lemma inverse_add_Infinitesimal_approx: "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (x + h) \ inverse x" for x h :: "'a::real_normed_div_algebra star" by (auto intro: approx_inverse approx_sym Infinitesimal_add_approx_self) lemma inverse_add_Infinitesimal_approx2: "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (h + x) \ inverse x" for x h :: "'a::real_normed_div_algebra star" by (metis add.commute inverse_add_Infinitesimal_approx) lemma inverse_add_Infinitesimal_approx_Infinitesimal: "x \ HFinite - Infinitesimal \ h \ Infinitesimal \ inverse (x + h) - inverse x \ h" for x h :: "'a::real_normed_div_algebra star" by (meson Infinitesimal_approx bex_Infinitesimal_iff inverse_add_Infinitesimal_approx) lemma Infinitesimal_square_iff: "x \ Infinitesimal \ x * x \ Infinitesimal" for x :: "'a::real_normed_div_algebra star" using Infinitesimal_mult Infinitesimal_mult_disj by auto declare Infinitesimal_square_iff [symmetric, simp] lemma HFinite_square_iff [simp]: "x * x \ HFinite \ x \ HFinite" for x :: "'a::real_normed_div_algebra star" using HFinite_HInfinite_iff HFinite_mult HInfinite_mult by blast lemma HInfinite_square_iff [simp]: "x * x \ HInfinite \ x \ HInfinite" for x :: "'a::real_normed_div_algebra star" by (auto simp add: HInfinite_HFinite_iff) lemma approx_HFinite_mult_cancel: "a \ HFinite - Infinitesimal \ a * w \ a * z \ w \ z" for a w z :: "'a::real_normed_div_algebra star" by (metis DiffD2 Infinitesimal_mult_disj bex_Infinitesimal_iff right_diff_distrib) lemma approx_HFinite_mult_cancel_iff1: "a \ HFinite - Infinitesimal \ a * w \ a * z \ w \ z" for a w z :: "'a::real_normed_div_algebra star" by (auto intro: approx_mult2 approx_HFinite_mult_cancel) lemma HInfinite_HFinite_add_cancel: "x + y \ HInfinite \ y \ HFinite \ x \ HInfinite" using HFinite_add HInfinite_HFinite_iff by blast lemma HInfinite_HFinite_add: "x \ HInfinite \ y \ HFinite \ x + y \ HInfinite" by (metis (no_types, opaque_lifting) HFinite_HInfinite_iff HFinite_add HFinite_minus_iff add.commute add_minus_cancel) lemma HInfinite_ge_HInfinite: "x \ HInfinite \ x \ y \ 0 \ x \ y \ HInfinite" for x y :: hypreal by (auto intro: HFinite_bounded simp add: HInfinite_HFinite_iff) lemma Infinitesimal_inverse_HInfinite: "x \ Infinitesimal \ x \ 0 \ inverse x \ HInfinite" for x :: "'a::real_normed_div_algebra star" by (metis Infinitesimal_HFinite_mult not_HFinite_HInfinite one_not_Infinitesimal right_inverse) lemma HInfinite_HFinite_not_Infinitesimal_mult: "x \ HInfinite \ y \ HFinite - Infinitesimal \ x * y \ HInfinite" for x y :: "'a::real_normed_div_algebra star" by (metis (no_types, opaque_lifting) HFinite_HInfinite_iff HFinite_Infinitesimal_not_zero HFinite_inverse2 HFinite_mult mult.assoc mult.right_neutral right_inverse) lemma HInfinite_HFinite_not_Infinitesimal_mult2: "x \ HInfinite \ y \ HFinite - Infinitesimal \ y * x \ HInfinite" for x y :: "'a::real_normed_div_algebra star" by (metis Diff_iff HInfinite_HFinite_iff HInfinite_inverse_Infinitesimal Infinitesimal_HFinite_mult2 divide_inverse mult_zero_right nonzero_eq_divide_eq) lemma HInfinite_gt_SReal: "x \ HInfinite \ 0 < x \ y \ \ \ y < x" for x y :: hypreal by (auto dest!: bspec simp add: HInfinite_def abs_if order_less_imp_le) lemma HInfinite_gt_zero_gt_one: "x \ HInfinite \ 0 < x \ 1 < x" for x :: hypreal by (auto intro: HInfinite_gt_SReal) lemma not_HInfinite_one [simp]: "1 \ HInfinite" by (simp add: HInfinite_HFinite_iff) lemma approx_hrabs_disj: "\x\ \ x \ \x\ \ -x" for x :: hypreal by (simp add: abs_if) subsection \Theorems about Monads\ lemma monad_hrabs_Un_subset: "monad \x\ \ monad x \ monad (- x)" for x :: hypreal by (simp add: abs_if) lemma Infinitesimal_monad_eq: "e \ Infinitesimal \ monad (x + e) = monad x" by (fast intro!: Infinitesimal_add_approx_self [THEN approx_sym] approx_monad_iff [THEN iffD1]) lemma mem_monad_iff: "u \ monad x \ - u \ monad (- x)" by (simp add: monad_def) lemma Infinitesimal_monad_zero_iff: "x \ Infinitesimal \ x \ monad 0" by (auto intro: approx_sym simp add: monad_def mem_infmal_iff) lemma monad_zero_minus_iff: "x \ monad 0 \ - x \ monad 0" by (simp add: Infinitesimal_monad_zero_iff [symmetric]) lemma monad_zero_hrabs_iff: "x \ monad 0 \ \x\ \ monad 0" for x :: hypreal using Infinitesimal_monad_zero_iff by blast lemma mem_monad_self [simp]: "x \ monad x" by (simp add: monad_def) subsection \Proof that \<^term>\x \ y\ implies \<^term>\\x\ \ \y\\\ lemma approx_subset_monad: "x \ y \ {x, y} \ monad x" by (simp (no_asm)) (simp add: approx_monad_iff) lemma approx_subset_monad2: "x \ y \ {x, y} \ monad y" using approx_subset_monad approx_sym by auto lemma mem_monad_approx: "u \ monad x \ x \ u" by (simp add: monad_def) lemma approx_mem_monad: "x \ u \ u \ monad x" by (simp add: monad_def) lemma approx_mem_monad2: "x \ u \ x \ monad u" using approx_mem_monad approx_sym by blast lemma approx_mem_monad_zero: "x \ y \ x \ monad 0 \ y \ monad 0" using approx_trans monad_def by blast lemma Infinitesimal_approx_hrabs: "x \ y \ x \ Infinitesimal \ \x\ \ \y\" for x y :: hypreal using approx_hnorm by fastforce lemma less_Infinitesimal_less: "0 < x \ x \ Infinitesimal \ e \ Infinitesimal \ e < x" for x :: hypreal using Infinitesimal_interval less_linear by blast lemma Ball_mem_monad_gt_zero: "0 < x \ x \ Infinitesimal \ u \ monad x \ 0 < u" for u x :: hypreal by (metis bex_Infinitesimal_iff2 less_Infinitesimal_less less_add_same_cancel2 mem_monad_approx) lemma Ball_mem_monad_less_zero: "x < 0 \ x \ Infinitesimal \ u \ monad x \ u < 0" for u x :: hypreal by (metis Ball_mem_monad_gt_zero approx_monad_iff less_asym linorder_neqE_linordered_idom mem_infmal_iff mem_monad_approx mem_monad_self) lemma lemma_approx_gt_zero: "0 < x \ x \ Infinitesimal \ x \ y \ 0 < y" for x y :: hypreal by (blast dest: Ball_mem_monad_gt_zero approx_subset_monad) lemma lemma_approx_less_zero: "x < 0 \ x \ Infinitesimal \ x \ y \ y < 0" for x y :: hypreal by (blast dest: Ball_mem_monad_less_zero approx_subset_monad) lemma approx_hrabs: "x \ y \ \x\ \ \y\" for x y :: hypreal by (drule approx_hnorm) simp lemma approx_hrabs_zero_cancel: "\x\ \ 0 \ x \ 0" for x :: hypreal using mem_infmal_iff by blast lemma approx_hrabs_add_Infinitesimal: "e \ Infinitesimal \ \x\ \ \x + e\" for e x :: hypreal by (fast intro: approx_hrabs Infinitesimal_add_approx_self) lemma approx_hrabs_add_minus_Infinitesimal: "e \ Infinitesimal ==> \x\ \ \x + -e\" for e x :: hypreal by (fast intro: approx_hrabs Infinitesimal_add_minus_approx_self) lemma hrabs_add_Infinitesimal_cancel: "e \ Infinitesimal \ e' \ Infinitesimal \ \x + e\ = \y + e'\ \ \x\ \ \y\" for e e' x y :: hypreal by (metis approx_hrabs_add_Infinitesimal approx_trans2) lemma hrabs_add_minus_Infinitesimal_cancel: "e \ Infinitesimal \ e' \ Infinitesimal \ \x + -e\ = \y + -e'\ \ \x\ \ \y\" for e e' x y :: hypreal by (meson Infinitesimal_minus_iff hrabs_add_Infinitesimal_cancel) subsection \More \<^term>\HFinite\ and \<^term>\Infinitesimal\ Theorems\ text \ Interesting slightly counterintuitive theorem: necessary for proving that an open interval is an NS open set. \ lemma Infinitesimal_add_hypreal_of_real_less: assumes "x < y" and u: "u \ Infinitesimal" shows "hypreal_of_real x + u < hypreal_of_real y" proof - have "\u\ < hypreal_of_real y - hypreal_of_real x" using InfinitesimalD \x < y\ u by fastforce then show ?thesis by (simp add: abs_less_iff) qed lemma Infinitesimal_add_hrabs_hypreal_of_real_less: "x \ Infinitesimal \ \hypreal_of_real r\ < hypreal_of_real y \ \hypreal_of_real r + x\ < hypreal_of_real y" by (metis Infinitesimal_add_hypreal_of_real_less approx_hrabs_add_Infinitesimal approx_sym bex_Infinitesimal_iff2 star_of_abs star_of_less) lemma Infinitesimal_add_hrabs_hypreal_of_real_less2: "x \ Infinitesimal \ \hypreal_of_real r\ < hypreal_of_real y \ \x + hypreal_of_real r\ < hypreal_of_real y" using Infinitesimal_add_hrabs_hypreal_of_real_less by fastforce lemma hypreal_of_real_le_add_Infininitesimal_cancel: assumes le: "hypreal_of_real x + u \ hypreal_of_real y + v" and u: "u \ Infinitesimal" and v: "v \ Infinitesimal" shows "hypreal_of_real x \ hypreal_of_real y" proof (rule ccontr) assume "\ hypreal_of_real x \ hypreal_of_real y" then have "hypreal_of_real y + (v - u) < hypreal_of_real x" by (simp add: Infinitesimal_add_hypreal_of_real_less Infinitesimal_diff u v) then show False by (simp add: add_diff_eq add_le_imp_le_diff le leD) qed lemma hypreal_of_real_le_add_Infininitesimal_cancel2: "u \ Infinitesimal \ v \ Infinitesimal \ hypreal_of_real x + u \ hypreal_of_real y + v \ x \ y" by (blast intro: star_of_le [THEN iffD1] intro!: hypreal_of_real_le_add_Infininitesimal_cancel) lemma hypreal_of_real_less_Infinitesimal_le_zero: "hypreal_of_real x < e \ e \ Infinitesimal \ hypreal_of_real x \ 0" by (metis Infinitesimal_interval eq_iff le_less_linear star_of_Infinitesimal_iff_0 star_of_eq_0) lemma Infinitesimal_add_not_zero: "h \ Infinitesimal \ x \ 0 \ star_of x + h \ 0" by (metis Infinitesimal_add_approx_self star_of_approx_zero_iff) lemma monad_hrabs_less: "y \ monad x \ 0 < hypreal_of_real e \ \y - x\ < hypreal_of_real e" by (simp add: Infinitesimal_approx_minus approx_sym less_Infinitesimal_less mem_monad_approx) lemma mem_monad_SReal_HFinite: "x \ monad (hypreal_of_real a) \ x \ HFinite" using HFinite_star_of approx_HFinite mem_monad_approx by blast subsection \Theorems about Standard Part\ lemma st_approx_self: "x \ HFinite \ st x \ x" by (metis (no_types, lifting) approx_refl approx_trans3 someI_ex st_def st_part_Ex st_part_Ex1) lemma st_SReal: "x \ HFinite \ st x \ \" by (metis (mono_tags, lifting) approx_sym someI_ex st_def st_part_Ex) lemma st_HFinite: "x \ HFinite \ st x \ HFinite" by (erule st_SReal [THEN SReal_subset_HFinite [THEN subsetD]]) lemma st_unique: "r \ \ \ r \ x \ st x = r" by (meson SReal_subset_HFinite approx_HFinite approx_unique_real st_SReal st_approx_self subsetD) lemma st_SReal_eq: "x \ \ \ st x = x" by (metis approx_refl st_unique) lemma st_hypreal_of_real [simp]: "st (hypreal_of_real x) = hypreal_of_real x" by (rule SReal_hypreal_of_real [THEN st_SReal_eq]) lemma st_eq_approx: "x \ HFinite \ y \ HFinite \ st x = st y \ x \ y" by (auto dest!: st_approx_self elim!: approx_trans3) lemma approx_st_eq: assumes x: "x \ HFinite" and y: "y \ HFinite" and xy: "x \ y" shows "st x = st y" proof - have "st x \ x" "st y \ y" "st x \ \" "st y \ \" by (simp_all add: st_approx_self st_SReal x y) with xy show ?thesis by (fast elim: approx_trans approx_trans2 SReal_approx_iff [THEN iffD1]) qed lemma st_eq_approx_iff: "x \ HFinite \ y \ HFinite \ x \ y \ st x = st y" by (blast intro: approx_st_eq st_eq_approx) lemma st_Infinitesimal_add_SReal: "x \ \ \ e \ Infinitesimal \ st (x + e) = x" by (simp add: Infinitesimal_add_approx_self st_unique) lemma st_Infinitesimal_add_SReal2: "x \ \ \ e \ Infinitesimal \ st (e + x) = x" by (metis add.commute st_Infinitesimal_add_SReal) lemma HFinite_st_Infinitesimal_add: "x \ HFinite \ \e \ Infinitesimal. x = st(x) + e" by (blast dest!: st_approx_self [THEN approx_sym] bex_Infinitesimal_iff2 [THEN iffD2]) lemma st_add: "x \ HFinite \ y \ HFinite \ st (x + y) = st x + st y" by (simp add: st_unique st_SReal st_approx_self approx_add) lemma st_numeral [simp]: "st (numeral w) = numeral w" by (rule Reals_numeral [THEN st_SReal_eq]) lemma st_neg_numeral [simp]: "st (- numeral w) = - numeral w" using st_unique by auto lemma st_0 [simp]: "st 0 = 0" by (simp add: st_SReal_eq) lemma st_1 [simp]: "st 1 = 1" by (simp add: st_SReal_eq) lemma st_neg_1 [simp]: "st (- 1) = - 1" by (simp add: st_SReal_eq) lemma st_minus: "x \ HFinite \ st (- x) = - st x" by (simp add: st_unique st_SReal st_approx_self approx_minus) lemma st_diff: "\x \ HFinite; y \ HFinite\ \ st (x - y) = st x - st y" by (simp add: st_unique st_SReal st_approx_self approx_diff) lemma st_mult: "\x \ HFinite; y \ HFinite\ \ st (x * y) = st x * st y" by (simp add: st_unique st_SReal st_approx_self approx_mult_HFinite) lemma st_Infinitesimal: "x \ Infinitesimal \ st x = 0" by (simp add: st_unique mem_infmal_iff) lemma st_not_Infinitesimal: "st(x) \ 0 \ x \ Infinitesimal" by (fast intro: st_Infinitesimal) lemma st_inverse: "x \ HFinite \ st x \ 0 \ st (inverse x) = inverse (st x)" by (simp add: approx_inverse st_SReal st_approx_self st_not_Infinitesimal st_unique) lemma st_divide [simp]: "x \ HFinite \ y \ HFinite \ st y \ 0 \ st (x / y) = st x / st y" by (simp add: divide_inverse st_mult st_not_Infinitesimal HFinite_inverse st_inverse) lemma st_idempotent [simp]: "x \ HFinite \ st (st x) = st x" by (blast intro: st_HFinite st_approx_self approx_st_eq) lemma Infinitesimal_add_st_less: "x \ HFinite \ y \ HFinite \ u \ Infinitesimal \ st x < st y \ st x + u < st y" by (metis Infinitesimal_add_hypreal_of_real_less SReal_iff st_SReal star_of_less) lemma Infinitesimal_add_st_le_cancel: "x \ HFinite \ y \ HFinite \ u \ Infinitesimal \ st x \ st y + u \ st x \ st y" by (meson Infinitesimal_add_st_less leD le_less_linear) lemma st_le: "x \ HFinite \ y \ HFinite \ x \ y \ st x \ st y" by (metis approx_le_bound approx_sym linear st_SReal st_approx_self st_part_Ex1) lemma st_zero_le: "0 \ x \ x \ HFinite \ 0 \ st x" by (metis HFinite_0 st_0 st_le) lemma st_zero_ge: "x \ 0 \ x \ HFinite \ st x \ 0" by (metis HFinite_0 st_0 st_le) lemma st_hrabs: "x \ HFinite \ \st x\ = st \x\" by (simp add: order_class.order.antisym st_zero_ge linorder_not_le st_zero_le abs_if st_minus linorder_not_less) subsection \Alternative Definitions using Free Ultrafilter\ subsubsection \\<^term>\HFinite\\ lemma HFinite_FreeUltrafilterNat: assumes "star_n X \ HFinite" shows "\u. eventually (\n. norm (X n) < u) \" proof - obtain r where "hnorm (star_n X) < hypreal_of_real r" using HFiniteD SReal_iff assms by fastforce then have "\\<^sub>F n in \. norm (X n) < r" by (simp add: hnorm_def star_n_less star_of_def starfun_star_n) then show ?thesis .. qed lemma FreeUltrafilterNat_HFinite: assumes "eventually (\n. norm (X n) < u) \" shows "star_n X \ HFinite" proof - have "hnorm (star_n X) < hypreal_of_real u" by (simp add: assms hnorm_def star_n_less star_of_def starfun_star_n) then show ?thesis by (meson HInfiniteD SReal_hypreal_of_real less_asym not_HFinite_HInfinite) qed lemma HFinite_FreeUltrafilterNat_iff: "star_n X \ HFinite \ (\u. eventually (\n. norm (X n) < u) \)" using FreeUltrafilterNat_HFinite HFinite_FreeUltrafilterNat by blast subsubsection \\<^term>\HInfinite\\ text \Exclude this type of sets from free ultrafilter for Infinite numbers!\ lemma FreeUltrafilterNat_const_Finite: "eventually (\n. norm (X n) = u) \ \ star_n X \ HFinite" by (simp add: FreeUltrafilterNat_HFinite [where u = "u+1"] eventually_mono) lemma HInfinite_FreeUltrafilterNat: assumes "star_n X \ HInfinite" shows "\\<^sub>F n in \. u < norm (X n)" proof - have "\ (\\<^sub>F n in \. norm (X n) < u + 1)" using FreeUltrafilterNat_HFinite HFinite_HInfinite_iff assms by auto then show ?thesis by (auto simp flip: FreeUltrafilterNat.eventually_not_iff elim: eventually_mono) qed lemma FreeUltrafilterNat_HInfinite: assumes "\u. eventually (\n. u < norm (X n)) \" shows "star_n X \ HInfinite" proof - { fix u assume "\\<^sub>Fn in \. norm (X n) < u" "\\<^sub>Fn in \. u < norm (X n)" then have "\\<^sub>F x in \. False" by eventually_elim auto then have False by (simp add: eventually_False FreeUltrafilterNat.proper) } then show ?thesis using HFinite_FreeUltrafilterNat HInfinite_HFinite_iff assms by blast qed lemma HInfinite_FreeUltrafilterNat_iff: "star_n X \ HInfinite \ (\u. eventually (\n. u < norm (X n)) \)" using HInfinite_FreeUltrafilterNat FreeUltrafilterNat_HInfinite by blast subsubsection \\<^term>\Infinitesimal\\ lemma ball_SReal_eq: "(\x::hypreal \ Reals. P x) \ (\x::real. P (star_of x))" by (auto simp: SReal_def) lemma Infinitesimal_FreeUltrafilterNat_iff: "(star_n X \ Infinitesimal) = (\u>0. eventually (\n. norm (X n) < u) \)" (is "?lhs = ?rhs") proof - have "?lhs \ (\r>0. hnorm (star_n X) < hypreal_of_real r)" by (simp add: Infinitesimal_def ball_SReal_eq) also have "... \ ?rhs" by (simp add: hnorm_def starfun_star_n star_of_def star_less_def starP2_star_n) finally show ?thesis . qed text \Infinitesimals as smaller than \1/n\ for all \n::nat (> 0)\.\ lemma lemma_Infinitesimal: "(\r. 0 < r \ x < r) \ (\n. x < inverse (real (Suc n)))" by (meson inverse_positive_iff_positive less_trans of_nat_0_less_iff reals_Archimedean zero_less_Suc) lemma lemma_Infinitesimal2: "(\r \ Reals. 0 < r \ x < r) \ (\n. x < inverse(hypreal_of_nat (Suc n)))" (is "_ = ?rhs") proof (intro iffI strip) assume R: ?rhs fix r::hypreal assume "r \ \" "0 < r" then obtain n y where "inverse (real (Suc n)) < y" and r: "r = hypreal_of_real y" by (metis SReal_iff reals_Archimedean star_of_0_less) then have "inverse (1 + hypreal_of_nat n) < hypreal_of_real y" by (metis of_nat_Suc star_of_inverse star_of_less star_of_nat_def) then show "x < r" by (metis R r le_less_trans less_imp_le of_nat_Suc) qed (meson Reals_inverse Reals_of_nat of_nat_0_less_iff positive_imp_inverse_positive zero_less_Suc) lemma Infinitesimal_hypreal_of_nat_iff: "Infinitesimal = {x. \n. hnorm x < inverse (hypreal_of_nat (Suc n))}" using Infinitesimal_def lemma_Infinitesimal2 by auto subsection \Proof that \\\ is an infinite number\ text \It will follow that \\\ is an infinitesimal number.\ lemma Suc_Un_eq: "{n. n < Suc m} = {n. n < m} Un {n. n = m}" by (auto simp add: less_Suc_eq) text \Prove that any segment is finite and hence cannot belong to \\\.\ lemma finite_real_of_nat_segment: "finite {n::nat. real n < real (m::nat)}" by auto lemma finite_real_of_nat_less_real: "finite {n::nat. real n < u}" proof - obtain m where "u < real m" using reals_Archimedean2 by blast then have "{n. real n < u} \ {.. u}" by (metis infinite_nat_iff_unbounded leD le_nat_floor mem_Collect_eq) lemma finite_rabs_real_of_nat_le_real: "finite {n::nat. \real n\ \ u}" by (simp add: finite_real_of_nat_le_real) lemma rabs_real_of_nat_le_real_FreeUltrafilterNat: "\ eventually (\n. \real n\ \ u) \" by (blast intro!: FreeUltrafilterNat.finite finite_rabs_real_of_nat_le_real) lemma FreeUltrafilterNat_nat_gt_real: "eventually (\n. u < real n) \" proof - have "{n::nat. \ u < real n} = {n. real n \ u}" by auto then show ?thesis by (auto simp add: FreeUltrafilterNat.finite' finite_real_of_nat_le_real) qed text \The complement of \{n. \real n\ \ u} = {n. u < \real n\}\ is in \\\ by property of (free) ultrafilters.\ text \\<^term>\\\ is a member of \<^term>\HInfinite\.\ theorem HInfinite_omega [simp]: "\ \ HInfinite" proof - have "\\<^sub>F n in \. u < norm (1 + real n)" for u using FreeUltrafilterNat_nat_gt_real [of "u-1"] eventually_mono by fastforce then show ?thesis by (simp add: omega_def FreeUltrafilterNat_HInfinite) qed text \Epsilon is a member of Infinitesimal.\ lemma Infinitesimal_epsilon [simp]: "\ \ Infinitesimal" by (auto intro!: HInfinite_inverse_Infinitesimal HInfinite_omega simp add: epsilon_inverse_omega) lemma HFinite_epsilon [simp]: "\ \ HFinite" by (auto intro: Infinitesimal_subset_HFinite [THEN subsetD]) lemma epsilon_approx_zero [simp]: "\ \ 0" by (simp add: mem_infmal_iff [symmetric]) text \Needed for proof that we define a hyperreal \[ hypreal_of_real a\ given that \\n. |X n - a| < 1/n\. Used in proof of \NSLIM \ LIM\.\ lemma real_of_nat_less_inverse_iff: "0 < u \ u < inverse (real(Suc n)) \ real(Suc n) < inverse u" using less_imp_inverse_less by force lemma finite_inverse_real_of_posnat_gt_real: "0 < u \ finite {n. u < inverse (real (Suc n))}" proof (simp only: real_of_nat_less_inverse_iff) have "{n. 1 + real n < inverse u} = {n. real n < inverse u - 1}" by fastforce then show "finite {n. real (Suc n) < inverse u}" using finite_real_of_nat_less_real [of "inverse u - 1"] by auto qed lemma finite_inverse_real_of_posnat_ge_real: assumes "0 < u" shows "finite {n. u \ inverse (real (Suc n))}" proof - have "\na. u \ inverse (1 + real na) \ na \ ceiling (inverse u)" by (smt (verit, best) assms ceiling_less_cancel ceiling_of_nat inverse_inverse_eq inverse_le_iff_le) then show ?thesis apply (auto simp add: finite_nat_set_iff_bounded_le) by (meson assms inverse_positive_iff_positive le_nat_iff less_imp_le zero_less_ceiling) qed lemma inverse_real_of_posnat_ge_real_FreeUltrafilterNat: "0 < u \ \ eventually (\n. u \ inverse(real(Suc n))) \" by (blast intro!: FreeUltrafilterNat.finite finite_inverse_real_of_posnat_ge_real) lemma FreeUltrafilterNat_inverse_real_of_posnat: "0 < u \ eventually (\n. inverse(real(Suc n)) < u) \" by (drule inverse_real_of_posnat_ge_real_FreeUltrafilterNat) (simp add: FreeUltrafilterNat.eventually_not_iff not_le[symmetric]) text \Example of an hypersequence (i.e. an extended standard sequence) whose term with an hypernatural suffix is an infinitesimal i.e. the whn'nth term of the hypersequence is a member of Infinitesimal\ lemma SEQ_Infinitesimal: "( *f* (\n::nat. inverse(real(Suc n)))) whn \ Infinitesimal" by (simp add: hypnat_omega_def starfun_star_n star_n_inverse Infinitesimal_FreeUltrafilterNat_iff FreeUltrafilterNat_inverse_real_of_posnat del: of_nat_Suc) text \Example where we get a hyperreal from a real sequence for which a particular property holds. The theorem is used in proofs about equivalence of nonstandard and standard neighbourhoods. Also used for equivalence of nonstandard ans standard definitions of pointwise limit.\ text \\|X(n) - x| < 1/n \ [] - hypreal_of_real x| \ Infinitesimal\\ lemma real_seq_to_hypreal_Infinitesimal: "\n. norm (X n - x) < inverse (real (Suc n)) \ star_n X - star_of x \ Infinitesimal" unfolding star_n_diff star_of_def Infinitesimal_FreeUltrafilterNat_iff star_n_inverse by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat intro: order_less_trans elim!: eventually_mono) lemma real_seq_to_hypreal_approx: "\n. norm (X n - x) < inverse (real (Suc n)) \ star_n X \ star_of x" by (metis bex_Infinitesimal_iff real_seq_to_hypreal_Infinitesimal) lemma real_seq_to_hypreal_approx2: "\n. norm (x - X n) < inverse(real(Suc n)) \ star_n X \ star_of x" by (metis norm_minus_commute real_seq_to_hypreal_approx) lemma real_seq_to_hypreal_Infinitesimal2: "\n. norm(X n - Y n) < inverse(real(Suc n)) \ star_n X - star_n Y \ Infinitesimal" unfolding Infinitesimal_FreeUltrafilterNat_iff star_n_diff by (auto dest!: FreeUltrafilterNat_inverse_real_of_posnat intro: order_less_trans elim!: eventually_mono) end diff --git a/src/HOL/Num.thy b/src/HOL/Num.thy --- a/src/HOL/Num.thy +++ b/src/HOL/Num.thy @@ -1,1553 +1,1553 @@ (* Title: HOL/Num.thy Author: Florian Haftmann Author: Brian Huffman *) section \Binary Numerals\ theory Num imports BNF_Least_Fixpoint Transfer begin subsection \The \num\ type\ datatype num = One | Bit0 num | Bit1 num text \Increment function for type \<^typ>\num\\ primrec inc :: "num \ num" where "inc One = Bit0 One" | "inc (Bit0 x) = Bit1 x" | "inc (Bit1 x) = Bit0 (inc x)" text \Converting between type \<^typ>\num\ and type \<^typ>\nat\\ primrec nat_of_num :: "num \ nat" where "nat_of_num One = Suc 0" | "nat_of_num (Bit0 x) = nat_of_num x + nat_of_num x" | "nat_of_num (Bit1 x) = Suc (nat_of_num x + nat_of_num x)" primrec num_of_nat :: "nat \ num" where "num_of_nat 0 = One" | "num_of_nat (Suc n) = (if 0 < n then inc (num_of_nat n) else One)" lemma nat_of_num_pos: "0 < nat_of_num x" by (induct x) simp_all lemma nat_of_num_neq_0: " nat_of_num x \ 0" by (induct x) simp_all lemma nat_of_num_inc: "nat_of_num (inc x) = Suc (nat_of_num x)" by (induct x) simp_all lemma num_of_nat_double: "0 < n \ num_of_nat (n + n) = Bit0 (num_of_nat n)" by (induct n) simp_all text \Type \<^typ>\num\ is isomorphic to the strictly positive natural numbers.\ lemma nat_of_num_inverse: "num_of_nat (nat_of_num x) = x" by (induct x) (simp_all add: num_of_nat_double nat_of_num_pos) lemma num_of_nat_inverse: "0 < n \ nat_of_num (num_of_nat n) = n" by (induct n) (simp_all add: nat_of_num_inc) lemma num_eq_iff: "x = y \ nat_of_num x = nat_of_num y" apply safe apply (drule arg_cong [where f=num_of_nat]) apply (simp add: nat_of_num_inverse) done lemma num_induct [case_names One inc]: fixes P :: "num \ bool" assumes One: "P One" and inc: "\x. P x \ P (inc x)" shows "P x" proof - obtain n where n: "Suc n = nat_of_num x" by (cases "nat_of_num x") (simp_all add: nat_of_num_neq_0) have "P (num_of_nat (Suc n))" proof (induct n) case 0 from One show ?case by simp next case (Suc n) then have "P (inc (num_of_nat (Suc n)))" by (rule inc) then show "P (num_of_nat (Suc (Suc n)))" by simp qed with n show "P x" by (simp add: nat_of_num_inverse) qed text \ From now on, there are two possible models for \<^typ>\num\: as positive naturals (rule \num_induct\) and as digit representation (rules \num.induct\, \num.cases\). \ subsection \Numeral operations\ instantiation num :: "{plus,times,linorder}" begin definition [code del]: "m + n = num_of_nat (nat_of_num m + nat_of_num n)" definition [code del]: "m * n = num_of_nat (nat_of_num m * nat_of_num n)" definition [code del]: "m \ n \ nat_of_num m \ nat_of_num n" definition [code del]: "m < n \ nat_of_num m < nat_of_num n" instance by standard (auto simp add: less_num_def less_eq_num_def num_eq_iff) end lemma nat_of_num_add: "nat_of_num (x + y) = nat_of_num x + nat_of_num y" unfolding plus_num_def by (intro num_of_nat_inverse add_pos_pos nat_of_num_pos) lemma nat_of_num_mult: "nat_of_num (x * y) = nat_of_num x * nat_of_num y" unfolding times_num_def by (intro num_of_nat_inverse mult_pos_pos nat_of_num_pos) lemma add_num_simps [simp, code]: "One + One = Bit0 One" "One + Bit0 n = Bit1 n" "One + Bit1 n = Bit0 (n + One)" "Bit0 m + One = Bit1 m" "Bit0 m + Bit0 n = Bit0 (m + n)" "Bit0 m + Bit1 n = Bit1 (m + n)" "Bit1 m + One = Bit0 (m + One)" "Bit1 m + Bit0 n = Bit1 (m + n)" "Bit1 m + Bit1 n = Bit0 (m + n + One)" by (simp_all add: num_eq_iff nat_of_num_add) lemma mult_num_simps [simp, code]: "m * One = m" "One * n = n" "Bit0 m * Bit0 n = Bit0 (Bit0 (m * n))" "Bit0 m * Bit1 n = Bit0 (m * Bit1 n)" "Bit1 m * Bit0 n = Bit0 (Bit1 m * n)" "Bit1 m * Bit1 n = Bit1 (m + n + Bit0 (m * n))" by (simp_all add: num_eq_iff nat_of_num_add nat_of_num_mult distrib_right distrib_left) lemma eq_num_simps: "One = One \ True" "One = Bit0 n \ False" "One = Bit1 n \ False" "Bit0 m = One \ False" "Bit1 m = One \ False" "Bit0 m = Bit0 n \ m = n" "Bit0 m = Bit1 n \ False" "Bit1 m = Bit0 n \ False" "Bit1 m = Bit1 n \ m = n" by simp_all lemma le_num_simps [simp, code]: "One \ n \ True" "Bit0 m \ One \ False" "Bit1 m \ One \ False" "Bit0 m \ Bit0 n \ m \ n" "Bit0 m \ Bit1 n \ m \ n" "Bit1 m \ Bit1 n \ m \ n" "Bit1 m \ Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma less_num_simps [simp, code]: "m < One \ False" "One < Bit0 n \ True" "One < Bit1 n \ True" "Bit0 m < Bit0 n \ m < n" "Bit0 m < Bit1 n \ m \ n" "Bit1 m < Bit1 n \ m < n" "Bit1 m < Bit0 n \ m < n" using nat_of_num_pos [of n] nat_of_num_pos [of m] by (auto simp add: less_eq_num_def less_num_def) lemma le_num_One_iff: "x \ num.One \ x = num.One" by (simp add: antisym_conv) text \Rules using \One\ and \inc\ as constructors.\ lemma add_One: "x + One = inc x" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma add_One_commute: "One + n = n + One" by (induct n) simp_all lemma add_inc: "x + inc y = inc (x + y)" by (simp add: num_eq_iff nat_of_num_add nat_of_num_inc) lemma mult_inc: "x * inc y = x * y + x" by (simp add: num_eq_iff nat_of_num_mult nat_of_num_add nat_of_num_inc) text \The \<^const>\num_of_nat\ conversion.\ lemma num_of_nat_One: "n \ 1 \ num_of_nat n = One" by (cases n) simp_all lemma num_of_nat_plus_distrib: "0 < m \ 0 < n \ num_of_nat (m + n) = num_of_nat m + num_of_nat n" by (induct n) (auto simp add: add_One add_One_commute add_inc) text \A double-and-decrement function.\ primrec BitM :: "num \ num" where "BitM One = One" | "BitM (Bit0 n) = Bit1 (BitM n)" | "BitM (Bit1 n) = Bit1 (Bit0 n)" lemma BitM_plus_one: "BitM n + One = Bit0 n" by (induct n) simp_all lemma one_plus_BitM: "One + BitM n = Bit0 n" unfolding add_One_commute BitM_plus_one .. lemma BitM_inc_eq: \Num.BitM (Num.inc n) = Num.Bit1 n\ by (induction n) simp_all lemma inc_BitM_eq: \Num.inc (Num.BitM n) = num.Bit0 n\ by (simp add: BitM_plus_one[symmetric] add_One) text \Squaring and exponentiation.\ primrec sqr :: "num \ num" where "sqr One = One" | "sqr (Bit0 n) = Bit0 (Bit0 (sqr n))" | "sqr (Bit1 n) = Bit1 (Bit0 (sqr n + n))" primrec pow :: "num \ num \ num" where "pow x One = x" | "pow x (Bit0 y) = sqr (pow x y)" | "pow x (Bit1 y) = sqr (pow x y) * x" lemma nat_of_num_sqr: "nat_of_num (sqr x) = nat_of_num x * nat_of_num x" by (induct x) (simp_all add: algebra_simps nat_of_num_add) lemma sqr_conv_mult: "sqr x = x * x" by (simp add: num_eq_iff nat_of_num_sqr nat_of_num_mult) lemma num_double [simp]: "num.Bit0 num.One * n = num.Bit0 n" by (simp add: num_eq_iff nat_of_num_mult) subsection \Binary numerals\ text \ We embed binary representations into a generic algebraic structure using \numeral\. \ class numeral = one + semigroup_add begin primrec numeral :: "num \ 'a" where numeral_One: "numeral One = 1" | numeral_Bit0: "numeral (Bit0 n) = numeral n + numeral n" | numeral_Bit1: "numeral (Bit1 n) = numeral n + numeral n + 1" lemma numeral_code [code]: "numeral One = 1" "numeral (Bit0 n) = (let m = numeral n in m + m)" "numeral (Bit1 n) = (let m = numeral n in m + m + 1)" by (simp_all add: Let_def) lemma one_plus_numeral_commute: "1 + numeral x = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) next case Bit1 then show ?case by (simp add: add.assoc [symmetric]) (simp add: add.assoc) qed lemma numeral_inc: "numeral (inc x) = numeral x + 1" proof (induct x) case One then show ?case by simp next case Bit0 then show ?case by simp next case (Bit1 x) have "numeral x + (1 + numeral x) + 1 = numeral x + (numeral x + 1) + 1" by (simp only: one_plus_numeral_commute) with Bit1 show ?case by (simp add: add.assoc) qed declare numeral.simps [simp del] abbreviation "Numeral1 \ numeral One" declare numeral_One [code_post] end text \Numeral syntax.\ syntax "_Numeral" :: "num_const \ 'a" ("_") ML_file \Tools/numeral.ML\ parse_translation \ let fun numeral_tr [(c as Const (\<^syntax_const>\_constrain\, _)) $ t $ u] = c $ numeral_tr [t] $ u | numeral_tr [Const (num, _)] = (Numeral.mk_number_syntax o #value o Lexicon.read_num) num | numeral_tr ts = raise TERM ("numeral_tr", ts); in [(\<^syntax_const>\_Numeral\, K numeral_tr)] end \ typed_print_translation \ let fun num_tr' ctxt T [n] = let val k = Numeral.dest_num_syntax n; val t' = Syntax.const \<^syntax_const>\_Numeral\ $ Syntax.free (string_of_int k); in (case T of Type (\<^type_name>\fun\, [_, T']) => if Printer.type_emphasis ctxt T' then Syntax.const \<^syntax_const>\_constrain\ $ t' $ Syntax_Phases.term_of_typ ctxt T' else t' | _ => if T = dummyT then t' else raise Match) end; in [(\<^const_syntax>\numeral\, num_tr')] end \ subsection \Class-specific numeral rules\ text \\<^const>\numeral\ is a morphism.\ subsubsection \Structures with addition: class \numeral\\ context numeral begin lemma numeral_add: "numeral (m + n) = numeral m + numeral n" by (induct n rule: num_induct) (simp_all only: numeral_One add_One add_inc numeral_inc add.assoc) lemma numeral_plus_numeral: "numeral m + numeral n = numeral (m + n)" by (rule numeral_add [symmetric]) lemma numeral_plus_one: "numeral n + 1 = numeral (n + One)" using numeral_add [of n One] by (simp add: numeral_One) lemma one_plus_numeral: "1 + numeral n = numeral (One + n)" using numeral_add [of One n] by (simp add: numeral_One) lemma one_add_one: "1 + 1 = 2" using numeral_add [of One One] by (simp add: numeral_One) lemmas add_numeral_special = numeral_plus_one one_plus_numeral one_add_one end subsubsection \Structures with negation: class \neg_numeral\\ class neg_numeral = numeral + group_add begin lemma uminus_numeral_One: "- Numeral1 = - 1" by (simp add: numeral_One) text \Numerals form an abelian subgroup.\ inductive is_num :: "'a \ bool" where "is_num 1" | "is_num x \ is_num (- x)" | "is_num x \ is_num y \ is_num (x + y)" lemma is_num_numeral: "is_num (numeral k)" by (induct k) (simp_all add: numeral.simps is_num.intros) lemma is_num_add_commute: "is_num x \ is_num y \ x + y = y + x" proof(induction x rule: is_num.induct) case 1 then show ?case proof (induction y rule: is_num.induct) case 1 then show ?case by simp next case (2 y) then have "y + (1 + - y) + y = y + (- y + 1) + y" by (simp add: add.assoc) then have "y + (1 + - y) = y + (- y + 1)" by simp then show ?case by (rule add_left_imp_eq[of y]) next case (3 x y) then have "1 + (x + y) = x + 1 + y" by (simp add: add.assoc [symmetric]) then show ?case using 3 by (simp add: add.assoc) qed next case (2 x) then have "x + (- x + y) + x = x + (y + - x) + x" by (simp add: add.assoc) then have "x + (- x + y) = x + (y + - x)" by simp then show ?case by (rule add_left_imp_eq[of x]) next case (3 x z) moreover have "x + (y + z) = (x + y) + z" by (simp add: add.assoc[symmetric]) ultimately show ?case by (simp add: add.assoc) qed lemma is_num_add_left_commute: "is_num x \ is_num y \ x + (y + z) = y + (x + z)" by (simp only: add.assoc [symmetric] is_num_add_commute) lemmas is_num_normalize = add.assoc is_num_add_commute is_num_add_left_commute is_num.intros is_num_numeral minus_add definition dbl :: "'a \ 'a" where "dbl x = x + x" definition dbl_inc :: "'a \ 'a" where "dbl_inc x = x + x + 1" definition dbl_dec :: "'a \ 'a" where "dbl_dec x = x + x - 1" definition sub :: "num \ num \ 'a" where "sub k l = numeral k - numeral l" lemma numeral_BitM: "numeral (BitM n) = numeral (Bit0 n) - 1" by (simp only: BitM_plus_one [symmetric] numeral_add numeral_One eq_diff_eq) lemma sub_inc_One_eq: \Num.sub (Num.inc n) num.One = numeral n\ by (simp_all add: sub_def diff_eq_eq numeral_inc numeral.numeral_One) lemma dbl_simps [simp]: "dbl (- numeral k) = - dbl (numeral k)" "dbl 0 = 0" "dbl 1 = 2" "dbl (- 1) = - 2" "dbl (numeral k) = numeral (Bit0 k)" by (simp_all add: dbl_def numeral.simps minus_add) lemma dbl_inc_simps [simp]: "dbl_inc (- numeral k) = - dbl_dec (numeral k)" "dbl_inc 0 = 1" "dbl_inc 1 = 3" "dbl_inc (- 1) = - 1" "dbl_inc (numeral k) = numeral (Bit1 k)" by (simp_all add: dbl_inc_def dbl_dec_def numeral.simps numeral_BitM is_num_normalize algebra_simps del: add_uminus_conv_diff) lemma dbl_dec_simps [simp]: "dbl_dec (- numeral k) = - dbl_inc (numeral k)" "dbl_dec 0 = - 1" "dbl_dec 1 = 1" "dbl_dec (- 1) = - 3" "dbl_dec (numeral k) = numeral (BitM k)" by (simp_all add: dbl_dec_def dbl_inc_def numeral.simps numeral_BitM is_num_normalize) lemma sub_num_simps [simp]: "sub One One = 0" "sub One (Bit0 l) = - numeral (BitM l)" "sub One (Bit1 l) = - numeral (Bit0 l)" "sub (Bit0 k) One = numeral (BitM k)" "sub (Bit1 k) One = numeral (Bit0 k)" "sub (Bit0 k) (Bit0 l) = dbl (sub k l)" "sub (Bit0 k) (Bit1 l) = dbl_dec (sub k l)" "sub (Bit1 k) (Bit0 l) = dbl_inc (sub k l)" "sub (Bit1 k) (Bit1 l) = dbl (sub k l)" by (simp_all add: dbl_def dbl_dec_def dbl_inc_def sub_def numeral.simps numeral_BitM is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_simps: "numeral m + - numeral n = sub m n" "- numeral m + numeral n = sub n m" "- numeral m + - numeral n = - (numeral m + numeral n)" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma add_neg_numeral_special: "1 + - numeral m = sub One m" "- numeral m + 1 = sub One m" "numeral m + - 1 = sub m One" "- 1 + numeral n = sub n One" "- 1 + - numeral n = - numeral (inc n)" "- numeral m + - 1 = - numeral (inc m)" "1 + - 1 = 0" "- 1 + 1 = 0" "- 1 + - 1 = - 2" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize right_minus numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_simps: "numeral m - numeral n = sub m n" "numeral m - - numeral n = numeral (m + n)" "- numeral m - numeral n = - numeral (m + n)" "- numeral m - - numeral n = sub n m" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize del: add_uminus_conv_diff add: diff_conv_add_uminus) lemma diff_numeral_special: "1 - numeral n = sub One n" "numeral m - 1 = sub m One" "1 - - numeral n = numeral (One + n)" "- numeral m - 1 = - numeral (m + One)" "- 1 - numeral n = - numeral (inc n)" "numeral m - - 1 = numeral (inc m)" "- 1 - - numeral n = sub n One" "- numeral m - - 1 = sub One m" "1 - 1 = 0" "- 1 - 1 = - 2" "1 - - 1 = 2" "- 1 - - 1 = 0" by (simp_all add: sub_def numeral_add numeral.simps is_num_normalize numeral_inc del: add_uminus_conv_diff add: diff_conv_add_uminus) end subsubsection \Structures with multiplication: class \semiring_numeral\\ class semiring_numeral = semiring + monoid_mult begin subclass numeral .. lemma numeral_mult: "numeral (m * n) = numeral m * numeral n" by (induct n rule: num_induct) (simp_all add: numeral_One mult_inc numeral_inc numeral_add distrib_left) lemma numeral_times_numeral: "numeral m * numeral n = numeral (m * n)" by (rule numeral_mult [symmetric]) lemma mult_2: "2 * z = z + z" by (simp add: one_add_one [symmetric] distrib_right) lemma mult_2_right: "z * 2 = z + z" by (simp add: one_add_one [symmetric] distrib_left) lemma left_add_twice: "a + (a + b) = 2 * a + b" by (simp add: mult_2 ac_simps) end subsubsection \Structures with a zero: class \semiring_1\\ context semiring_1 begin subclass semiring_numeral .. lemma of_nat_numeral [simp]: "of_nat (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_class.numeral.simps of_nat_add of_nat_1) end lemma nat_of_num_numeral [code_abbrev]: "nat_of_num = numeral" proof fix n have "numeral n = nat_of_num n" by (induct n) (simp_all add: numeral.simps) then show "nat_of_num n = numeral n" by simp qed lemma nat_of_num_code [code]: "nat_of_num One = 1" "nat_of_num (Bit0 n) = (let m = nat_of_num n in m + m)" "nat_of_num (Bit1 n) = (let m = nat_of_num n in Suc (m + m))" by (simp_all add: Let_def) subsubsection \Equality: class \semiring_char_0\\ context semiring_char_0 begin lemma numeral_eq_iff: "numeral m = numeral n \ m = n" by (simp only: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] of_nat_eq_iff num_eq_iff) lemma numeral_eq_one_iff: "numeral n = 1 \ n = One" by (rule numeral_eq_iff [of n One, unfolded numeral_One]) lemma one_eq_numeral_iff: "1 = numeral n \ One = n" by (rule numeral_eq_iff [of One n, unfolded numeral_One]) lemma numeral_neq_zero: "numeral n \ 0" by (simp add: of_nat_numeral [symmetric] nat_of_num_numeral [symmetric] nat_of_num_pos) lemma zero_neq_numeral: "0 \ numeral n" unfolding eq_commute [of 0] by (rule numeral_neq_zero) lemmas eq_numeral_simps [simp] = numeral_eq_iff numeral_eq_one_iff one_eq_numeral_iff numeral_neq_zero zero_neq_numeral end subsubsection \Comparisons: class \linordered_nonzero_semiring\\ context linordered_nonzero_semiring begin lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" proof - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) then show ?thesis by simp qed lemma one_le_numeral: "1 \ numeral n" using numeral_le_iff [of num.One n] by (simp add: numeral_One) lemma numeral_le_one_iff: "numeral n \ 1 \ n \ num.One" using numeral_le_iff [of n num.One] by (simp add: numeral_One) lemma numeral_less_iff: "numeral m < numeral n \ m < n" proof - have "of_nat (numeral m) < of_nat (numeral n) \ m < n" unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. then show ?thesis by simp qed lemma not_numeral_less_one: "\ numeral n < 1" using numeral_less_iff [of n num.One] by (simp add: numeral_One) lemma one_less_numeral_iff: "1 < numeral n \ num.One < n" using numeral_less_iff [of num.One n] by (simp add: numeral_One) lemma zero_le_numeral: "0 \ numeral n" using dual_order.trans one_le_numeral zero_le_one by blast lemma zero_less_numeral: "0 < numeral n" using less_linear not_numeral_less_one order.strict_trans zero_less_one by blast lemma not_numeral_le_zero: "\ numeral n \ 0" by (simp add: not_le zero_less_numeral) lemma not_numeral_less_zero: "\ numeral n < 0" by (simp add: not_less zero_le_numeral) lemmas le_numeral_extra = zero_le_one not_one_le_zero order_refl [of 0] order_refl [of 1] lemmas less_numeral_extra = zero_less_one not_one_less_zero less_irrefl [of 0] less_irrefl [of 1] lemmas le_numeral_simps [simp] = numeral_le_iff one_le_numeral numeral_le_one_iff zero_le_numeral not_numeral_le_zero lemmas less_numeral_simps [simp] = numeral_less_iff one_less_numeral_iff not_numeral_less_one zero_less_numeral not_numeral_less_zero lemma min_0_1 [simp]: fixes min' :: "'a \ 'a \ 'a" defines "min' \ min" shows "min' 0 1 = 0" "min' 1 0 = 0" "min' 0 (numeral x) = 0" "min' (numeral x) 0 = 0" "min' 1 (numeral x) = 1" "min' (numeral x) 1 = 1" by (simp_all add: min'_def min_def le_num_One_iff) lemma max_0_1 [simp]: fixes max' :: "'a \ 'a \ 'a" defines "max' \ max" shows "max' 0 1 = 1" "max' 1 0 = 1" "max' 0 (numeral x) = numeral x" "max' (numeral x) 0 = numeral x" "max' 1 (numeral x) = numeral x" "max' (numeral x) 1 = numeral x" by (simp_all add: max'_def max_def le_num_One_iff) end text \Unfold \min\ and \max\ on numerals.\ lemmas max_number_of [simp] = max_def [of "numeral u" "numeral v"] max_def [of "numeral u" "- numeral v"] max_def [of "- numeral u" "numeral v"] max_def [of "- numeral u" "- numeral v"] for u v lemmas min_number_of [simp] = min_def [of "numeral u" "numeral v"] min_def [of "numeral u" "- numeral v"] min_def [of "- numeral u" "numeral v"] min_def [of "- numeral u" "- numeral v"] for u v subsubsection \Multiplication and negation: class \ring_1\\ context ring_1 begin subclass neg_numeral .. lemma mult_neg_numeral_simps: "- numeral m * - numeral n = numeral (m * n)" "- numeral m * numeral n = - numeral (m * n)" "numeral m * - numeral n = - numeral (m * n)" by (simp_all only: mult_minus_left mult_minus_right minus_minus numeral_mult) lemma mult_minus1 [simp]: "- 1 * z = - z" by (simp add: numeral.simps) lemma mult_minus1_right [simp]: "z * - 1 = - z" by (simp add: numeral.simps) lemma minus_sub_one_diff_one [simp]: \- sub m One - 1 = - numeral m\ proof - have \sub m One + 1 = numeral m\ by (simp flip: eq_diff_eq add: diff_numeral_special) then have \- (sub m One + 1) = - numeral m\ by simp then show ?thesis by simp qed end subsubsection \Equality using \iszero\ for rings with non-zero characteristic\ context ring_1 begin definition iszero :: "'a \ bool" where "iszero z \ z = 0" lemma iszero_0 [simp]: "iszero 0" by (simp add: iszero_def) lemma not_iszero_1 [simp]: "\ iszero 1" by (simp add: iszero_def) lemma not_iszero_Numeral1: "\ iszero Numeral1" by (simp add: numeral_One) lemma not_iszero_neg_1 [simp]: "\ iszero (- 1)" by (simp add: iszero_def) lemma not_iszero_neg_Numeral1: "\ iszero (- Numeral1)" by (simp add: numeral_One) lemma iszero_neg_numeral [simp]: "iszero (- numeral w) \ iszero (numeral w)" unfolding iszero_def by (rule neg_equal_0_iff_equal) lemma eq_iff_iszero_diff: "x = y \ iszero (x - y)" unfolding iszero_def by (rule eq_iff_diff_eq_0) text \ The \eq_numeral_iff_iszero\ lemmas are not declared \[simp]\ by default, because for rings of characteristic zero, better simp rules are possible. For a type like integers mod \n\, type-instantiated versions of these rules should be added to the simplifier, along with a type-specific rule for deciding propositions of the form \iszero (numeral w)\. bh: Maybe it would not be so bad to just declare these as simp rules anyway? I should test whether these rules take precedence over the \ring_char_0\ rules in the simplifier. \ lemma eq_numeral_iff_iszero: "numeral x = numeral y \ iszero (sub x y)" "numeral x = - numeral y \ iszero (numeral (x + y))" "- numeral x = numeral y \ iszero (numeral (x + y))" "- numeral x = - numeral y \ iszero (sub y x)" "numeral x = 1 \ iszero (sub x One)" "1 = numeral y \ iszero (sub One y)" "- numeral x = 1 \ iszero (numeral (x + One))" "1 = - numeral y \ iszero (numeral (One + y))" "numeral x = 0 \ iszero (numeral x)" "0 = numeral y \ iszero (numeral y)" "- numeral x = 0 \ iszero (numeral x)" "0 = - numeral y \ iszero (numeral y)" unfolding eq_iff_iszero_diff diff_numeral_simps diff_numeral_special by simp_all end subsubsection \Equality and negation: class \ring_char_0\\ context ring_char_0 begin lemma not_iszero_numeral [simp]: "\ iszero (numeral w)" by (simp add: iszero_def) lemma neg_numeral_eq_iff: "- numeral m = - numeral n \ m = n" by simp lemma numeral_neq_neg_numeral: "numeral m \ - numeral n" by (simp add: eq_neg_iff_add_eq_0 numeral_plus_numeral) lemma neg_numeral_neq_numeral: "- numeral m \ numeral n" by (rule numeral_neq_neg_numeral [symmetric]) lemma zero_neq_neg_numeral: "0 \ - numeral n" by simp lemma neg_numeral_neq_zero: "- numeral n \ 0" by simp lemma one_neq_neg_numeral: "1 \ - numeral n" using numeral_neq_neg_numeral [of One n] by (simp add: numeral_One) lemma neg_numeral_neq_one: "- numeral n \ 1" using neg_numeral_neq_numeral [of n One] by (simp add: numeral_One) lemma neg_one_neq_numeral: "- 1 \ numeral n" using neg_numeral_neq_numeral [of One n] by (simp add: numeral_One) lemma numeral_neq_neg_one: "numeral n \ - 1" using numeral_neq_neg_numeral [of n One] by (simp add: numeral_One) lemma neg_one_eq_numeral_iff: "- 1 = - numeral n \ n = One" using neg_numeral_eq_iff [of One n] by (auto simp add: numeral_One) lemma numeral_eq_neg_one_iff: "- numeral n = - 1 \ n = One" using neg_numeral_eq_iff [of n One] by (auto simp add: numeral_One) lemma neg_one_neq_zero: "- 1 \ 0" by simp lemma zero_neq_neg_one: "0 \ - 1" by simp lemma neg_one_neq_one: "- 1 \ 1" using neg_numeral_neq_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemma one_neq_neg_one: "1 \ - 1" using numeral_neq_neg_numeral [of One One] by (simp only: numeral_One not_False_eq_True) lemmas eq_neg_numeral_simps [simp] = neg_numeral_eq_iff numeral_neq_neg_numeral neg_numeral_neq_numeral one_neq_neg_numeral neg_numeral_neq_one zero_neq_neg_numeral neg_numeral_neq_zero neg_one_neq_numeral numeral_neq_neg_one neg_one_eq_numeral_iff numeral_eq_neg_one_iff neg_one_neq_zero zero_neq_neg_one neg_one_neq_one one_neq_neg_one end subsubsection \Structures with negation and order: class \linordered_idom\\ context linordered_idom begin subclass ring_char_0 .. lemma neg_numeral_le_iff: "- numeral m \ - numeral n \ n \ m" by (simp only: neg_le_iff_le numeral_le_iff) lemma neg_numeral_less_iff: "- numeral m < - numeral n \ n < m" by (simp only: neg_less_iff_less numeral_less_iff) lemma neg_numeral_less_zero: "- numeral n < 0" by (simp only: neg_less_0_iff_less zero_less_numeral) lemma neg_numeral_le_zero: "- numeral n \ 0" by (simp only: neg_le_0_iff_le zero_le_numeral) lemma not_zero_less_neg_numeral: "\ 0 < - numeral n" by (simp only: not_less neg_numeral_le_zero) lemma not_zero_le_neg_numeral: "\ 0 \ - numeral n" by (simp only: not_le neg_numeral_less_zero) lemma neg_numeral_less_numeral: "- numeral m < numeral n" using neg_numeral_less_zero zero_less_numeral by (rule less_trans) lemma neg_numeral_le_numeral: "- numeral m \ numeral n" by (simp only: less_imp_le neg_numeral_less_numeral) lemma not_numeral_less_neg_numeral: "\ numeral m < - numeral n" by (simp only: not_less neg_numeral_le_numeral) lemma not_numeral_le_neg_numeral: "\ numeral m \ - numeral n" by (simp only: not_le neg_numeral_less_numeral) lemma neg_numeral_less_one: "- numeral m < 1" by (rule neg_numeral_less_numeral [of m One, unfolded numeral_One]) lemma neg_numeral_le_one: "- numeral m \ 1" by (rule neg_numeral_le_numeral [of m One, unfolded numeral_One]) lemma not_one_less_neg_numeral: "\ 1 < - numeral m" by (simp only: not_less neg_numeral_le_one) lemma not_one_le_neg_numeral: "\ 1 \ - numeral m" by (simp only: not_le neg_numeral_less_one) lemma not_numeral_less_neg_one: "\ numeral m < - 1" using not_numeral_less_neg_numeral [of m One] by (simp add: numeral_One) lemma not_numeral_le_neg_one: "\ numeral m \ - 1" using not_numeral_le_neg_numeral [of m One] by (simp add: numeral_One) lemma neg_one_less_numeral: "- 1 < numeral m" using neg_numeral_less_numeral [of One m] by (simp add: numeral_One) lemma neg_one_le_numeral: "- 1 \ numeral m" using neg_numeral_le_numeral [of One m] by (simp add: numeral_One) lemma neg_numeral_less_neg_one_iff: "- numeral m < - 1 \ m \ One" by (cases m) simp_all lemma neg_numeral_le_neg_one: "- numeral m \ - 1" by simp lemma not_neg_one_less_neg_numeral: "\ - 1 < - numeral m" by simp lemma not_neg_one_le_neg_numeral_iff: "\ - 1 \ - numeral m \ m \ One" by (cases m) simp_all lemma sub_non_negative: "sub n m \ 0 \ n \ m" by (simp only: sub_def le_diff_eq) simp lemma sub_positive: "sub n m > 0 \ n > m" by (simp only: sub_def less_diff_eq) simp lemma sub_non_positive: "sub n m \ 0 \ n \ m" by (simp only: sub_def diff_le_eq) simp lemma sub_negative: "sub n m < 0 \ n < m" by (simp only: sub_def diff_less_eq) simp lemmas le_neg_numeral_simps [simp] = neg_numeral_le_iff neg_numeral_le_numeral not_numeral_le_neg_numeral neg_numeral_le_zero not_zero_le_neg_numeral neg_numeral_le_one not_one_le_neg_numeral neg_one_le_numeral not_numeral_le_neg_one neg_numeral_le_neg_one not_neg_one_le_neg_numeral_iff lemma le_minus_one_simps [simp]: "- 1 \ 0" "- 1 \ 1" "\ 0 \ - 1" "\ 1 \ - 1" by simp_all lemmas less_neg_numeral_simps [simp] = neg_numeral_less_iff neg_numeral_less_numeral not_numeral_less_neg_numeral neg_numeral_less_zero not_zero_less_neg_numeral neg_numeral_less_one not_one_less_neg_numeral neg_one_less_numeral not_numeral_less_neg_one neg_numeral_less_neg_one_iff not_neg_one_less_neg_numeral lemma less_minus_one_simps [simp]: "- 1 < 0" "- 1 < 1" "\ 0 < - 1" "\ 1 < - 1" by (simp_all add: less_le) lemma abs_numeral [simp]: "\numeral n\ = numeral n" by simp lemma abs_neg_numeral [simp]: "\- numeral n\ = numeral n" by (simp only: abs_minus_cancel abs_numeral) lemma abs_neg_one [simp]: "\- 1\ = 1" by simp end subsubsection \Natural numbers\ lemma numeral_num_of_nat: "numeral (num_of_nat n) = n" if "n > 0" using that nat_of_num_numeral num_of_nat_inverse by simp lemma Suc_1 [simp]: "Suc 1 = 2" unfolding Suc_eq_plus1 by (rule one_add_one) lemma Suc_numeral [simp]: "Suc (numeral n) = numeral (n + One)" unfolding Suc_eq_plus1 by (rule numeral_plus_one) definition pred_numeral :: "num \ nat" where "pred_numeral k = numeral k - 1" declare [[code drop: pred_numeral]] lemma numeral_eq_Suc: "numeral k = Suc (pred_numeral k)" by (simp add: pred_numeral_def) lemma eval_nat_numeral: "numeral One = Suc 0" "numeral (Bit0 n) = Suc (numeral (BitM n))" "numeral (Bit1 n) = Suc (numeral (Bit0 n))" by (simp_all add: numeral.simps BitM_plus_one) lemma pred_numeral_simps [simp]: "pred_numeral One = 0" "pred_numeral (Bit0 k) = numeral (BitM k)" "pred_numeral (Bit1 k) = numeral (Bit0 k)" by (simp_all only: pred_numeral_def eval_nat_numeral diff_Suc_Suc diff_0) lemma pred_numeral_inc [simp]: "pred_numeral (Num.inc k) = numeral k" by (simp only: pred_numeral_def numeral_inc diff_add_inverse2) lemma numeral_2_eq_2: "2 = Suc (Suc 0)" by (simp add: eval_nat_numeral) lemma numeral_3_eq_3: "3 = Suc (Suc (Suc 0))" by (simp add: eval_nat_numeral) lemma numeral_1_eq_Suc_0: "Numeral1 = Suc 0" by (simp only: numeral_One One_nat_def) lemma Suc_nat_number_of_add: "Suc (numeral v + n) = numeral (v + One) + n" by simp lemma numerals: "Numeral1 = (1::nat)" "2 = Suc (Suc 0)" by (rule numeral_One) (rule numeral_2_eq_2) lemmas numeral_nat = eval_nat_numeral BitM.simps One_nat_def text \Comparisons involving \<^term>\Suc\.\ lemma eq_numeral_Suc [simp]: "numeral k = Suc n \ pred_numeral k = n" by (simp add: numeral_eq_Suc) lemma Suc_eq_numeral [simp]: "Suc n = numeral k \ n = pred_numeral k" by (simp add: numeral_eq_Suc) lemma less_numeral_Suc [simp]: "numeral k < Suc n \ pred_numeral k < n" by (simp add: numeral_eq_Suc) lemma less_Suc_numeral [simp]: "Suc n < numeral k \ n < pred_numeral k" by (simp add: numeral_eq_Suc) lemma le_numeral_Suc [simp]: "numeral k \ Suc n \ pred_numeral k \ n" by (simp add: numeral_eq_Suc) lemma le_Suc_numeral [simp]: "Suc n \ numeral k \ n \ pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_Suc_numeral [simp]: "Suc n - numeral k = n - pred_numeral k" by (simp add: numeral_eq_Suc) lemma diff_numeral_Suc [simp]: "numeral k - Suc n = pred_numeral k - n" by (simp add: numeral_eq_Suc) lemma max_Suc_numeral [simp]: "max (Suc n) (numeral k) = Suc (max n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma max_numeral_Suc [simp]: "max (numeral k) (Suc n) = Suc (max (pred_numeral k) n)" by (simp add: numeral_eq_Suc) lemma min_Suc_numeral [simp]: "min (Suc n) (numeral k) = Suc (min n (pred_numeral k))" by (simp add: numeral_eq_Suc) lemma min_numeral_Suc [simp]: "min (numeral k) (Suc n) = Suc (min (pred_numeral k) n)" by (simp add: numeral_eq_Suc) text \For \<^term>\case_nat\ and \<^term>\rec_nat\.\ lemma case_nat_numeral [simp]: "case_nat a f (numeral v) = (let pv = pred_numeral v in f pv)" by (simp add: numeral_eq_Suc) lemma case_nat_add_eq_if [simp]: "case_nat a f ((numeral v) + n) = (let pv = pred_numeral v in f (pv + n))" by (simp add: numeral_eq_Suc) lemma rec_nat_numeral [simp]: "rec_nat a f (numeral v) = (let pv = pred_numeral v in f pv (rec_nat a f pv))" by (simp add: numeral_eq_Suc Let_def) lemma rec_nat_add_eq_if [simp]: "rec_nat a f (numeral v + n) = (let pv = pred_numeral v in f (pv + n) (rec_nat a f (pv + n)))" by (simp add: numeral_eq_Suc Let_def) text \Case analysis on \<^term>\n < 2\.\ lemma less_2_cases: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) lemma less_2_cases_iff: "n < 2 \ n = 0 \ n = Suc 0" by (auto simp add: numeral_2_eq_2) text \Removal of Small Numerals: 0, 1 and (in additive positions) 2.\ text \bh: Are these rules really a good idea? LCP: well, it already happens for 0 and 1!\ lemma add_2_eq_Suc [simp]: "2 + n = Suc (Suc n)" by simp lemma add_2_eq_Suc' [simp]: "n + 2 = Suc (Suc n)" by simp text \Can be used to eliminate long strings of Sucs, but not by default.\ lemma Suc3_eq_add_3: "Suc (Suc (Suc n)) = 3 + n" by simp lemmas nat_1_add_1 = one_add_one [where 'a=nat] (* legacy *) context semiring_numeral begin lemma numeral_add_unfold_funpow: \numeral k + a = ((+) 1 ^^ numeral k) a\ proof (rule sym, induction k arbitrary: a) case One then show ?case by (simp add: numeral_One Num.numeral_One) next case (Bit0 k) then show ?case by (simp add: numeral_Bit0 Num.numeral_Bit0 ac_simps funpow_add) next case (Bit1 k) then show ?case by (simp add: numeral_Bit1 Num.numeral_Bit1 ac_simps funpow_add) qed end context semiring_1 begin lemma numeral_unfold_funpow: \numeral k = ((+) 1 ^^ numeral k) 0\ using numeral_add_unfold_funpow [of k 0] by simp end context includes lifting_syntax begin lemma transfer_rule_numeral: \((=) ===> R) numeral numeral\ if [transfer_rule]: \R 0 0\ \R 1 1\ \(R ===> R ===> R) (+) (+)\ for R :: \'a::{semiring_numeral,monoid_add} \ 'b::{semiring_numeral,monoid_add} \ bool\ proof - have "((=) ===> R) (\k. ((+) 1 ^^ numeral k) 0) (\k. ((+) 1 ^^ numeral k) 0)" by transfer_prover moreover have \numeral = (\k. ((+) (1::'a) ^^ numeral k) 0)\ using numeral_add_unfold_funpow [where ?'a = 'a, of _ 0] by (simp add: fun_eq_iff) moreover have \numeral = (\k. ((+) (1::'b) ^^ numeral k) 0)\ using numeral_add_unfold_funpow [where ?'a = 'b, of _ 0] by (simp add: fun_eq_iff) ultimately show ?thesis by simp qed end subsection \Particular lemmas concerning \<^term>\2\\ context linordered_field begin subclass field_char_0 .. lemma half_gt_zero_iff: "0 < a / 2 \ 0 < a" by (auto simp add: field_simps) lemma half_gt_zero [simp]: "0 < a \ 0 < a / 2" by (simp add: half_gt_zero_iff) end subsection \Numeral equations as default simplification rules\ declare (in numeral) numeral_One [simp] declare (in numeral) numeral_plus_numeral [simp] declare (in numeral) add_numeral_special [simp] declare (in neg_numeral) add_neg_numeral_simps [simp] declare (in neg_numeral) add_neg_numeral_special [simp] declare (in neg_numeral) diff_numeral_simps [simp] declare (in neg_numeral) diff_numeral_special [simp] declare (in semiring_numeral) numeral_times_numeral [simp] declare (in ring_1) mult_neg_numeral_simps [simp] subsubsection \Special Simplification for Constants\ text \These distributive laws move literals inside sums and differences.\ lemmas distrib_right_numeral [simp] = distrib_right [of _ _ "numeral v"] for v lemmas distrib_left_numeral [simp] = distrib_left [of "numeral v"] for v lemmas left_diff_distrib_numeral [simp] = left_diff_distrib [of _ _ "numeral v"] for v lemmas right_diff_distrib_numeral [simp] = right_diff_distrib [of "numeral v"] for v text \These are actually for fields, like real\ lemmas zero_less_divide_iff_numeral [simp, no_atp] = zero_less_divide_iff [of "numeral w"] for w lemmas divide_less_0_iff_numeral [simp, no_atp] = divide_less_0_iff [of "numeral w"] for w lemmas zero_le_divide_iff_numeral [simp, no_atp] = zero_le_divide_iff [of "numeral w"] for w lemmas divide_le_0_iff_numeral [simp, no_atp] = divide_le_0_iff [of "numeral w"] for w text \Replaces \inverse #nn\ by \1/#nn\. It looks strange, but then other simprocs simplify the quotient.\ lemmas inverse_eq_divide_numeral [simp] = inverse_eq_divide [of "numeral w"] for w lemmas inverse_eq_divide_neg_numeral [simp] = inverse_eq_divide [of "- numeral w"] for w text \These laws simplify inequalities, moving unary minus from a term into the literal.\ lemmas equation_minus_iff_numeral [no_atp] = equation_minus_iff [of "numeral v"] for v lemmas minus_equation_iff_numeral [no_atp] = minus_equation_iff [of _ "numeral v"] for v lemmas le_minus_iff_numeral [no_atp] = le_minus_iff [of "numeral v"] for v lemmas minus_le_iff_numeral [no_atp] = minus_le_iff [of _ "numeral v"] for v lemmas less_minus_iff_numeral [no_atp] = less_minus_iff [of "numeral v"] for v lemmas minus_less_iff_numeral [no_atp] = minus_less_iff [of _ "numeral v"] for v (* FIXME maybe simproc *) text \Cancellation of constant factors in comparisons (\<\ and \\\)\ lemmas mult_less_cancel_left_numeral [simp, no_atp] = mult_less_cancel_left [of "numeral v"] for v lemmas mult_less_cancel_right_numeral [simp, no_atp] = mult_less_cancel_right [of _ "numeral v"] for v lemmas mult_le_cancel_left_numeral [simp, no_atp] = mult_le_cancel_left [of "numeral v"] for v lemmas mult_le_cancel_right_numeral [simp, no_atp] = mult_le_cancel_right [of _ "numeral v"] for v text \Multiplying out constant divisors in comparisons (\<\, \\\ and \=\)\ named_theorems divide_const_simps "simplification rules to simplify comparisons involving constant divisors" lemmas le_divide_eq_numeral1 [simp,divide_const_simps] = pos_le_divide_eq [of "numeral w", OF zero_less_numeral] neg_le_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_le_eq_numeral1 [simp,divide_const_simps] = pos_divide_le_eq [of "numeral w", OF zero_less_numeral] neg_divide_le_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas less_divide_eq_numeral1 [simp,divide_const_simps] = pos_less_divide_eq [of "numeral w", OF zero_less_numeral] neg_less_divide_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas divide_less_eq_numeral1 [simp,divide_const_simps] = pos_divide_less_eq [of "numeral w", OF zero_less_numeral] neg_divide_less_eq [of "- numeral w", OF neg_numeral_less_zero] for w lemmas eq_divide_eq_numeral1 [simp,divide_const_simps] = eq_divide_eq [of _ _ "numeral w"] eq_divide_eq [of _ _ "- numeral w"] for w lemmas divide_eq_eq_numeral1 [simp,divide_const_simps] = divide_eq_eq [of _ "numeral w"] divide_eq_eq [of _ "- numeral w"] for w subsubsection \Optional Simplification Rules Involving Constants\ text \Simplify quotients that are compared with a literal constant.\ lemmas le_divide_eq_numeral [divide_const_simps] = le_divide_eq [of "numeral w"] le_divide_eq [of "- numeral w"] for w lemmas divide_le_eq_numeral [divide_const_simps] = divide_le_eq [of _ _ "numeral w"] divide_le_eq [of _ _ "- numeral w"] for w lemmas less_divide_eq_numeral [divide_const_simps] = less_divide_eq [of "numeral w"] less_divide_eq [of "- numeral w"] for w lemmas divide_less_eq_numeral [divide_const_simps] = divide_less_eq [of _ _ "numeral w"] divide_less_eq [of _ _ "- numeral w"] for w lemmas eq_divide_eq_numeral [divide_const_simps] = eq_divide_eq [of "numeral w"] eq_divide_eq [of "- numeral w"] for w lemmas divide_eq_eq_numeral [divide_const_simps] = divide_eq_eq [of _ _ "numeral w"] divide_eq_eq [of _ _ "- numeral w"] for w text \Not good as automatic simprules because they cause case splits.\ lemmas [divide_const_simps] = le_divide_eq_1 divide_le_eq_1 less_divide_eq_1 divide_less_eq_1 subsection \Setting up simprocs\ lemma mult_numeral_1: "Numeral1 * a = a" for a :: "'a::semiring_numeral" by simp lemma mult_numeral_1_right: "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp lemma divide_numeral_1: "a / Numeral1 = a" for a :: "'a::field" by simp lemma inverse_numeral_1: "inverse Numeral1 = (Numeral1::'a::division_ring)" by simp text \ Theorem lists for the cancellation simprocs. The use of a binary numeral for 1 reduces the number of special cases. \ lemma mult_1s_semiring_numeral: "Numeral1 * a = a" "a * Numeral1 = a" for a :: "'a::semiring_numeral" by simp_all lemma mult_1s_ring_1: "- Numeral1 * b = - b" "b * - Numeral1 = - b" for b :: "'a::ring_1" by simp_all lemmas mult_1s = mult_1s_semiring_numeral mult_1s_ring_1 setup \ Reorient_Proc.add (fn Const (\<^const_name>\numeral\, _) $ _ => true | Const (\<^const_name>\uminus\, _) $ (Const (\<^const_name>\numeral\, _) $ _) => true | _ => false) \ simproc_setup reorient_numeral ("numeral w = x" | "- numeral w = y") = - Reorient_Proc.proc + \K Reorient_Proc.proc\ subsubsection \Simplification of arithmetic operations on integer constants\ lemmas arith_special = (* already declared simp above *) add_numeral_special add_neg_numeral_special diff_numeral_special lemmas arith_extra_simps = (* rules already in simpset *) numeral_plus_numeral add_neg_numeral_simps add_0_left add_0_right minus_zero diff_numeral_simps diff_0 diff_0_right numeral_times_numeral mult_neg_numeral_simps mult_zero_left mult_zero_right abs_numeral abs_neg_numeral text \ For making a minimal simpset, one must include these default simprules. Also include \simp_thms\. \ lemmas arith_simps = add_num_simps mult_num_simps sub_num_simps BitM.simps dbl_simps dbl_inc_simps dbl_dec_simps abs_zero abs_one arith_extra_simps lemmas more_arith_simps = neg_le_iff_le minus_zero left_minus right_minus mult_1_left mult_1_right mult_minus_left mult_minus_right minus_add_distrib minus_minus mult.assoc lemmas of_nat_simps = of_nat_0 of_nat_1 of_nat_Suc of_nat_add of_nat_mult text \Simplification of relational operations.\ lemmas eq_numeral_extra = zero_neq_one one_neq_zero lemmas rel_simps = le_num_simps less_num_simps eq_num_simps le_numeral_simps le_neg_numeral_simps le_minus_one_simps le_numeral_extra less_numeral_simps less_neg_numeral_simps less_minus_one_simps less_numeral_extra eq_numeral_simps eq_neg_numeral_simps eq_numeral_extra lemma Let_numeral [simp]: "Let (numeral v) f = f (numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. lemma Let_neg_numeral [simp]: "Let (- numeral v) f = f (- numeral v)" \ \Unfold all \let\s involving constants\ unfolding Let_def .. declaration \ let fun number_of ctxt T n = if not (Sign.of_sort (Proof_Context.theory_of ctxt) (T, \<^sort>\numeral\)) then raise CTERM ("number_of", []) else Numeral.mk_cnumber (Thm.ctyp_of ctxt T) n; in K ( Lin_Arith.set_number_of number_of #> Lin_Arith.add_simps @{thms arith_simps more_arith_simps rel_simps pred_numeral_simps arith_special numeral_One of_nat_simps uminus_numeral_One Suc_numeral Let_numeral Let_neg_numeral Let_0 Let_1 le_Suc_numeral le_numeral_Suc less_Suc_numeral less_numeral_Suc Suc_eq_numeral eq_numeral_Suc mult_Suc mult_Suc_right of_nat_numeral}) end \ subsubsection \Simplification of arithmetic when nested to the right\ lemma add_numeral_left [simp]: "numeral v + (numeral w + z) = (numeral(v + w) + z)" by (simp_all add: add.assoc [symmetric]) lemma add_neg_numeral_left [simp]: "numeral v + (- numeral w + y) = (sub v w + y)" "- numeral v + (numeral w + y) = (sub w v + y)" "- numeral v + (- numeral w + y) = (- numeral(v + w) + y)" by (simp_all add: add.assoc [symmetric]) lemma mult_numeral_left_semiring_numeral: "numeral v * (numeral w * z) = (numeral(v * w) * z :: 'a::semiring_numeral)" by (simp add: mult.assoc [symmetric]) lemma mult_numeral_left_ring_1: "- numeral v * (numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "numeral v * (- numeral w * y) = (- numeral(v * w) * y :: 'a::ring_1)" "- numeral v * (- numeral w * y) = (numeral(v * w) * y :: 'a::ring_1)" by (simp_all add: mult.assoc [symmetric]) lemmas mult_numeral_left [simp] = mult_numeral_left_semiring_numeral mult_numeral_left_ring_1 hide_const (open) One Bit0 Bit1 BitM inc pow sqr sub dbl dbl_inc dbl_dec subsection \Code module namespace\ code_identifier code_module Num \ (SML) Arith and (OCaml) Arith and (Haskell) Arith subsection \Printing of evaluated natural numbers as numerals\ lemma [code_post]: "Suc 0 = 1" "Suc 1 = 2" "Suc (numeral n) = numeral (Num.inc n)" by (simp_all add: numeral_inc) lemmas [code_post] = Num.inc.simps subsection \More on auxiliary conversion\ context semiring_1 begin lemma numeral_num_of_nat_unfold: \numeral (num_of_nat n) = (if n = 0 then 1 else of_nat n)\ by (induction n) (simp_all add: numeral_inc ac_simps) lemma num_of_nat_numeral_eq [simp]: \num_of_nat (numeral q) = q\ proof (induction q) case One then show ?case by simp next case (Bit0 q) then have "num_of_nat (numeral (num.Bit0 q)) = num_of_nat (numeral q + numeral q)" by (simp only: Num.numeral_Bit0 Num.numeral_add) also have "\ = num.Bit0 (num_of_nat (numeral q))" by (rule num_of_nat_double) simp finally show ?case using Bit0.IH by simp next case (Bit1 q) then have "num_of_nat (numeral (num.Bit1 q)) = num_of_nat (numeral q + numeral q + 1)" by (simp only: Num.numeral_Bit1 Num.numeral_add) also have "\ = num_of_nat (numeral q + numeral q) + num_of_nat 1" by (rule num_of_nat_plus_distrib) auto also have "\ = num.Bit0 (num_of_nat (numeral q)) + num_of_nat 1" by (subst num_of_nat_double) auto finally show ?case using Bit1.IH by simp qed end end diff --git a/src/HOL/Numeral_Simprocs.thy b/src/HOL/Numeral_Simprocs.thy --- a/src/HOL/Numeral_Simprocs.thy +++ b/src/HOL/Numeral_Simprocs.thy @@ -1,299 +1,299 @@ (* Author: Various *) section \Combination and Cancellation Simprocs for Numeral Expressions\ theory Numeral_Simprocs imports Parity begin ML_file \~~/src/Provers/Arith/assoc_fold.ML\ ML_file \~~/src/Provers/Arith/cancel_numerals.ML\ ML_file \~~/src/Provers/Arith/combine_numerals.ML\ ML_file \~~/src/Provers/Arith/cancel_numeral_factor.ML\ ML_file \~~/src/Provers/Arith/extract_common_term.ML\ lemmas semiring_norm = Let_def arith_simps diff_nat_numeral rel_simps if_False if_True add_Suc add_numeral_left add_neg_numeral_left mult_numeral_left numeral_One [symmetric] uminus_numeral_One [symmetric] Suc_eq_plus1 eq_numeral_iff_iszero not_iszero_Numeral1 text \For \combine_numerals\\ lemma left_add_mult_distrib: "i*u + (j*u + k) = (i+j)*u + (k::nat)" by (simp add: add_mult_distrib) text \For \cancel_numerals\\ lemma nat_diff_add_eq1: "j <= (i::nat) ==> ((i*u + m) - (j*u + n)) = (((i-j)*u + m) - n)" by (simp split: nat_diff_split add: add_mult_distrib) lemma nat_diff_add_eq2: "i <= (j::nat) ==> ((i*u + m) - (j*u + n)) = (m - ((j-i)*u + n))" by (simp split: nat_diff_split add: add_mult_distrib) lemma nat_eq_add_iff1: "j <= (i::nat) ==> (i*u + m = j*u + n) = ((i-j)*u + m = n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_eq_add_iff2: "i <= (j::nat) ==> (i*u + m = j*u + n) = (m = (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_less_add_iff1: "j <= (i::nat) ==> (i*u + m < j*u + n) = ((i-j)*u + m < n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_less_add_iff2: "i <= (j::nat) ==> (i*u + m < j*u + n) = (m < (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_le_add_iff1: "j <= (i::nat) ==> (i*u + m <= j*u + n) = ((i-j)*u + m <= n)" by (auto split: nat_diff_split simp add: add_mult_distrib) lemma nat_le_add_iff2: "i <= (j::nat) ==> (i*u + m <= j*u + n) = (m <= (j-i)*u + n)" by (auto split: nat_diff_split simp add: add_mult_distrib) text \For \cancel_numeral_factors\\ lemma nat_mult_le_cancel1: "(0::nat) < k ==> (k*m <= k*n) = (m<=n)" by auto lemma nat_mult_less_cancel1: "(0::nat) < k ==> (k*m < k*n) = (m (k*m = k*n) = (m=n)" by auto lemma nat_mult_div_cancel1: "(0::nat) < k ==> (k*m) div (k*n) = (m div n)" by auto lemma nat_mult_dvd_cancel_disj[simp]: "(k*m) dvd (k*n) = (k=0 \ m dvd (n::nat))" by (auto simp: dvd_eq_mod_eq_0 mod_mult_mult1) lemma nat_mult_dvd_cancel1: "0 < k \ (k*m) dvd (k*n::nat) = (m dvd n)" by(auto) text \For \cancel_factor\\ lemmas nat_mult_le_cancel_disj = mult_le_cancel1 lemmas nat_mult_less_cancel_disj = mult_less_cancel1 lemma nat_mult_eq_cancel_disj: fixes k m n :: nat shows "k * m = k * n \ k = 0 \ m = n" by (fact mult_cancel_left) lemma nat_mult_div_cancel_disj: fixes k m n :: nat shows "(k * m) div (k * n) = (if k = 0 then 0 else m div n)" by (fact div_mult_mult1_if) lemma numeral_times_minus_swap: fixes x:: "'a::comm_ring_1" shows "numeral w * -x = x * - numeral w" by (simp add: ac_simps) ML_file \Tools/numeral_simprocs.ML\ simproc_setup semiring_assoc_fold ("(a::'a::comm_semiring_1_cancel) * b") = - \fn phi => Numeral_Simprocs.assoc_fold\ + \K Numeral_Simprocs.assoc_fold\ (* TODO: see whether the type class can be generalized further *) simproc_setup int_combine_numerals ("(i::'a::comm_ring_1) + j" | "(i::'a::comm_ring_1) - j") = - \fn phi => Numeral_Simprocs.combine_numerals\ + \K Numeral_Simprocs.combine_numerals\ simproc_setup field_combine_numerals ("(i::'a::{field,ring_char_0}) + j" |"(i::'a::{field,ring_char_0}) - j") = - \fn phi => Numeral_Simprocs.field_combine_numerals\ + \K Numeral_Simprocs.field_combine_numerals\ simproc_setup inteq_cancel_numerals ("(l::'a::comm_ring_1) + m = n" |"(l::'a::comm_ring_1) = m + n" |"(l::'a::comm_ring_1) - m = n" |"(l::'a::comm_ring_1) = m - n" |"(l::'a::comm_ring_1) * m = n" |"(l::'a::comm_ring_1) = m * n" |"- (l::'a::comm_ring_1) = m" |"(l::'a::comm_ring_1) = - m") = - \fn phi => Numeral_Simprocs.eq_cancel_numerals\ + \K Numeral_Simprocs.eq_cancel_numerals\ simproc_setup intless_cancel_numerals ("(l::'a::linordered_idom) + m < n" |"(l::'a::linordered_idom) < m + n" |"(l::'a::linordered_idom) - m < n" |"(l::'a::linordered_idom) < m - n" |"(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n" |"- (l::'a::linordered_idom) < m" |"(l::'a::linordered_idom) < - m") = - \fn phi => Numeral_Simprocs.less_cancel_numerals\ + \K Numeral_Simprocs.less_cancel_numerals\ simproc_setup intle_cancel_numerals ("(l::'a::linordered_idom) + m \ n" |"(l::'a::linordered_idom) \ m + n" |"(l::'a::linordered_idom) - m \ n" |"(l::'a::linordered_idom) \ m - n" |"(l::'a::linordered_idom) * m \ n" |"(l::'a::linordered_idom) \ m * n" |"- (l::'a::linordered_idom) \ m" |"(l::'a::linordered_idom) \ - m") = - \fn phi => Numeral_Simprocs.le_cancel_numerals\ + \K Numeral_Simprocs.le_cancel_numerals\ simproc_setup ring_eq_cancel_numeral_factor ("(l::'a::{idom,ring_char_0}) * m = n" |"(l::'a::{idom,ring_char_0}) = m * n") = - \fn phi => Numeral_Simprocs.eq_cancel_numeral_factor\ + \K Numeral_Simprocs.eq_cancel_numeral_factor\ simproc_setup ring_less_cancel_numeral_factor ("(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n") = - \fn phi => Numeral_Simprocs.less_cancel_numeral_factor\ + \K Numeral_Simprocs.less_cancel_numeral_factor\ simproc_setup ring_le_cancel_numeral_factor ("(l::'a::linordered_idom) * m <= n" |"(l::'a::linordered_idom) <= m * n") = - \fn phi => Numeral_Simprocs.le_cancel_numeral_factor\ + \K Numeral_Simprocs.le_cancel_numeral_factor\ (* TODO: remove comm_ring_1 constraint if possible *) simproc_setup int_div_cancel_numeral_factors ("((l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) * m) div n" |"(l::'a::{euclidean_semiring_cancel,comm_ring_1,ring_char_0}) div (m * n)") = - \fn phi => Numeral_Simprocs.div_cancel_numeral_factor\ + \K Numeral_Simprocs.div_cancel_numeral_factor\ simproc_setup divide_cancel_numeral_factor ("((l::'a::{field,ring_char_0}) * m) / n" |"(l::'a::{field,ring_char_0}) / (m * n)" |"((numeral v)::'a::{field,ring_char_0}) / (numeral w)") = - \fn phi => Numeral_Simprocs.divide_cancel_numeral_factor\ + \K Numeral_Simprocs.divide_cancel_numeral_factor\ simproc_setup ring_eq_cancel_factor ("(l::'a::idom) * m = n" | "(l::'a::idom) = m * n") = - \fn phi => Numeral_Simprocs.eq_cancel_factor\ + \K Numeral_Simprocs.eq_cancel_factor\ simproc_setup linordered_ring_le_cancel_factor ("(l::'a::linordered_idom) * m <= n" |"(l::'a::linordered_idom) <= m * n") = - \fn phi => Numeral_Simprocs.le_cancel_factor\ + \K Numeral_Simprocs.le_cancel_factor\ simproc_setup linordered_ring_less_cancel_factor ("(l::'a::linordered_idom) * m < n" |"(l::'a::linordered_idom) < m * n") = - \fn phi => Numeral_Simprocs.less_cancel_factor\ + \K Numeral_Simprocs.less_cancel_factor\ simproc_setup int_div_cancel_factor ("((l::'a::euclidean_semiring_cancel) * m) div n" |"(l::'a::euclidean_semiring_cancel) div (m * n)") = - \fn phi => Numeral_Simprocs.div_cancel_factor\ + \K Numeral_Simprocs.div_cancel_factor\ simproc_setup int_mod_cancel_factor ("((l::'a::euclidean_semiring_cancel) * m) mod n" |"(l::'a::euclidean_semiring_cancel) mod (m * n)") = - \fn phi => Numeral_Simprocs.mod_cancel_factor\ + \K Numeral_Simprocs.mod_cancel_factor\ simproc_setup dvd_cancel_factor ("((l::'a::idom) * m) dvd n" |"(l::'a::idom) dvd (m * n)") = - \fn phi => Numeral_Simprocs.dvd_cancel_factor\ + \K Numeral_Simprocs.dvd_cancel_factor\ simproc_setup divide_cancel_factor ("((l::'a::field) * m) / n" |"(l::'a::field) / (m * n)") = - \fn phi => Numeral_Simprocs.divide_cancel_factor\ + \K Numeral_Simprocs.divide_cancel_factor\ ML_file \Tools/nat_numeral_simprocs.ML\ simproc_setup nat_combine_numerals ("(i::nat) + j" | "Suc (i + j)") = - \fn phi => Nat_Numeral_Simprocs.combine_numerals\ + \K Nat_Numeral_Simprocs.combine_numerals\ simproc_setup nateq_cancel_numerals ("(l::nat) + m = n" | "(l::nat) = m + n" | "(l::nat) * m = n" | "(l::nat) = m * n" | "Suc m = n" | "m = Suc n") = - \fn phi => Nat_Numeral_Simprocs.eq_cancel_numerals\ + \K Nat_Numeral_Simprocs.eq_cancel_numerals\ simproc_setup natless_cancel_numerals ("(l::nat) + m < n" | "(l::nat) < m + n" | "(l::nat) * m < n" | "(l::nat) < m * n" | "Suc m < n" | "m < Suc n") = - \fn phi => Nat_Numeral_Simprocs.less_cancel_numerals\ + \K Nat_Numeral_Simprocs.less_cancel_numerals\ simproc_setup natle_cancel_numerals ("(l::nat) + m \ n" | "(l::nat) \ m + n" | "(l::nat) * m \ n" | "(l::nat) \ m * n" | "Suc m \ n" | "m \ Suc n") = - \fn phi => Nat_Numeral_Simprocs.le_cancel_numerals\ + \K Nat_Numeral_Simprocs.le_cancel_numerals\ simproc_setup natdiff_cancel_numerals ("((l::nat) + m) - n" | "(l::nat) - (m + n)" | "(l::nat) * m - n" | "(l::nat) - m * n" | "Suc m - n" | "m - Suc n") = - \fn phi => Nat_Numeral_Simprocs.diff_cancel_numerals\ + \K Nat_Numeral_Simprocs.diff_cancel_numerals\ simproc_setup nat_eq_cancel_numeral_factor ("(l::nat) * m = n" | "(l::nat) = m * n") = - \fn phi => Nat_Numeral_Simprocs.eq_cancel_numeral_factor\ + \K Nat_Numeral_Simprocs.eq_cancel_numeral_factor\ simproc_setup nat_less_cancel_numeral_factor ("(l::nat) * m < n" | "(l::nat) < m * n") = - \fn phi => Nat_Numeral_Simprocs.less_cancel_numeral_factor\ + \K Nat_Numeral_Simprocs.less_cancel_numeral_factor\ simproc_setup nat_le_cancel_numeral_factor ("(l::nat) * m <= n" | "(l::nat) <= m * n") = - \fn phi => Nat_Numeral_Simprocs.le_cancel_numeral_factor\ + \K Nat_Numeral_Simprocs.le_cancel_numeral_factor\ simproc_setup nat_div_cancel_numeral_factor ("((l::nat) * m) div n" | "(l::nat) div (m * n)") = - \fn phi => Nat_Numeral_Simprocs.div_cancel_numeral_factor\ + \K Nat_Numeral_Simprocs.div_cancel_numeral_factor\ simproc_setup nat_dvd_cancel_numeral_factor ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") = - \fn phi => Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\ + \K Nat_Numeral_Simprocs.dvd_cancel_numeral_factor\ simproc_setup nat_eq_cancel_factor ("(l::nat) * m = n" | "(l::nat) = m * n") = - \fn phi => Nat_Numeral_Simprocs.eq_cancel_factor\ + \K Nat_Numeral_Simprocs.eq_cancel_factor\ simproc_setup nat_less_cancel_factor ("(l::nat) * m < n" | "(l::nat) < m * n") = - \fn phi => Nat_Numeral_Simprocs.less_cancel_factor\ + \K Nat_Numeral_Simprocs.less_cancel_factor\ simproc_setup nat_le_cancel_factor ("(l::nat) * m <= n" | "(l::nat) <= m * n") = - \fn phi => Nat_Numeral_Simprocs.le_cancel_factor\ + \K Nat_Numeral_Simprocs.le_cancel_factor\ simproc_setup nat_div_cancel_factor ("((l::nat) * m) div n" | "(l::nat) div (m * n)") = - \fn phi => Nat_Numeral_Simprocs.div_cancel_factor\ + \K Nat_Numeral_Simprocs.div_cancel_factor\ simproc_setup nat_dvd_cancel_factor ("((l::nat) * m) dvd n" | "(l::nat) dvd (m * n)") = - \fn phi => Nat_Numeral_Simprocs.dvd_cancel_factor\ + \K Nat_Numeral_Simprocs.dvd_cancel_factor\ declaration \ K (Lin_Arith.add_simprocs [\<^simproc>\semiring_assoc_fold\, \<^simproc>\int_combine_numerals\, \<^simproc>\inteq_cancel_numerals\, \<^simproc>\intless_cancel_numerals\, \<^simproc>\intle_cancel_numerals\, \<^simproc>\field_combine_numerals\, \<^simproc>\nat_combine_numerals\, \<^simproc>\nateq_cancel_numerals\, \<^simproc>\natless_cancel_numerals\, \<^simproc>\natle_cancel_numerals\, \<^simproc>\natdiff_cancel_numerals\, Numeral_Simprocs.field_divide_cancel_numeral_factor]) \ end diff --git a/src/HOL/Product_Type.thy b/src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy +++ b/src/HOL/Product_Type.thy @@ -1,1376 +1,1376 @@ (* Title: HOL/Product_Type.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section \Cartesian products\ theory Product_Type imports Typedef Inductive Fun keywords "inductive_set" "coinductive_set" :: thy_defn begin subsection \\<^typ>\bool\ is a datatype\ free_constructors (discs_sels) case_bool for True | False by auto text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype True False by (auto intro: bool_induct) setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "bool"\ lemmas induct = old.bool.induct lemmas inducts = old.bool.inducts lemmas rec = old.bool.rec lemmas simps = bool.distinct bool.case bool.rec setup \Sign.parent_path\ declare case_split [cases type: bool] \ \prefer plain propositional version\ lemma [code]: "HOL.equal False P \ \ P" and [code]: "HOL.equal True P \ P" and [code]: "HOL.equal P False \ \ P" and [code]: "HOL.equal P True \ P" and [code nbe]: "HOL.equal P P \ True" by (simp_all add: equal) lemma If_case_cert: assumes "CASE \ (\b. If b f g)" shows "(CASE True \ f) &&& (CASE False \ g)" using assms by simp_all setup \Code.declare_case_global @{thm If_case_cert}\ code_printing constant "HOL.equal :: bool \ bool \ bool" \ (Haskell) infix 4 "==" | class_instance "bool" :: "equal" \ (Haskell) - subsection \The \unit\ type\ typedef unit = "{True}" by auto definition Unity :: unit ("'(')") where "() = Abs_unit True" lemma unit_eq [no_atp]: "u = ()" by (induct u) (simp add: Unity_def) text \ Simplification procedure for @{thm [source] unit_eq}. Cannot use this rule directly --- it loops! \ simproc_setup unit_eq ("x::unit") = \ - fn _ => fn _ => fn ct => + K (K (fn ct => if HOLogic.is_unit (Thm.term_of ct) then NONE - else SOME (mk_meta_eq @{thm unit_eq}) + else SOME (mk_meta_eq @{thm unit_eq}))) \ free_constructors case_unit for "()" by auto text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype "()" by simp setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "unit"\ lemmas induct = old.unit.induct lemmas inducts = old.unit.inducts lemmas rec = old.unit.rec lemmas simps = unit.case unit.rec setup \Sign.parent_path\ lemma unit_all_eq1: "(\x::unit. PROP P x) \ PROP P ()" by simp lemma unit_all_eq2: "(\x::unit. PROP P) \ PROP P" by (rule triv_forall_equality) text \ This rewrite counters the effect of simproc \unit_eq\ on @{term [source] "\u::unit. f u"}, replacing it by @{term [source] f} rather than by @{term [source] "\u. f ()"}. \ lemma unit_abs_eta_conv [simp]: "(\u::unit. f ()) = f" by (rule ext) simp lemma UNIV_unit: "UNIV = {()}" by auto instantiation unit :: default begin definition "default = ()" instance .. end instantiation unit :: "{complete_boolean_algebra,complete_linorder,wellorder}" begin definition less_eq_unit :: "unit \ unit \ bool" where "(_::unit) \ _ \ True" lemma less_eq_unit [iff]: "u \ v" for u v :: unit by (simp add: less_eq_unit_def) definition less_unit :: "unit \ unit \ bool" where "(_::unit) < _ \ False" lemma less_unit [iff]: "\ u < v" for u v :: unit by (simp_all add: less_eq_unit_def less_unit_def) definition bot_unit :: unit where [code_unfold]: "\ = ()" definition top_unit :: unit where [code_unfold]: "\ = ()" definition inf_unit :: "unit \ unit \ unit" where [simp]: "_ \ _ = ()" definition sup_unit :: "unit \ unit \ unit" where [simp]: "_ \ _ = ()" definition Inf_unit :: "unit set \ unit" where [simp]: "\_ = ()" definition Sup_unit :: "unit set \ unit" where [simp]: "\_ = ()" definition uminus_unit :: "unit \ unit" where [simp]: "- _ = ()" declare less_eq_unit_def [abs_def, code_unfold] less_unit_def [abs_def, code_unfold] inf_unit_def [abs_def, code_unfold] sup_unit_def [abs_def, code_unfold] Inf_unit_def [abs_def, code_unfold] Sup_unit_def [abs_def, code_unfold] uminus_unit_def [abs_def, code_unfold] instance by intro_classes auto end lemma [code]: "HOL.equal u v \ True" for u v :: unit unfolding equal unit_eq [of u] unit_eq [of v] by (rule iffI TrueI refl)+ code_printing type_constructor unit \ (SML) "unit" and (OCaml) "unit" and (Haskell) "()" and (Scala) "Unit" | constant Unity \ (SML) "()" and (OCaml) "()" and (Haskell) "()" and (Scala) "()" | class_instance unit :: equal \ (Haskell) - | constant "HOL.equal :: unit \ unit \ bool" \ (Haskell) infix 4 "==" code_reserved SML unit code_reserved OCaml unit code_reserved Scala Unit subsection \The product type\ subsubsection \Type definition\ definition Pair_Rep :: "'a \ 'b \ 'a \ 'b \ bool" where "Pair_Rep a b = (\x y. x = a \ y = b)" definition "prod = {f. \a b. f = Pair_Rep (a::'a) (b::'b)}" typedef ('a, 'b) prod ("(_ \/ _)" [21, 20] 20) = "prod :: ('a \ 'b \ bool) set" unfolding prod_def by auto type_notation (ASCII) prod (infixr "*" 20) definition Pair :: "'a \ 'b \ 'a \ 'b" where "Pair a b = Abs_prod (Pair_Rep a b)" lemma prod_cases: "(\a b. P (Pair a b)) \ P p" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) free_constructors case_prod for Pair fst snd proof - fix P :: bool and p :: "'a \ 'b" show "(\x1 x2. p = Pair x1 x2 \ P) \ P" by (cases p) (auto simp add: prod_def Pair_def Pair_Rep_def) next fix a c :: 'a and b d :: 'b have "Pair_Rep a b = Pair_Rep c d \ a = c \ b = d" by (auto simp add: Pair_Rep_def fun_eq_iff) moreover have "Pair_Rep a b \ prod" and "Pair_Rep c d \ prod" by (auto simp add: prod_def) ultimately show "Pair a b = Pair c d \ a = c \ b = d" by (simp add: Pair_def Abs_prod_inject) qed text \Avoid name clashes by prefixing the output of \old_rep_datatype\ with \old\.\ setup \Sign.mandatory_path "old"\ old_rep_datatype Pair by (erule prod_cases) (rule prod.inject) setup \Sign.parent_path\ text \But erase the prefix for properties that are not generated by \free_constructors\.\ setup \Sign.mandatory_path "prod"\ declare old.prod.inject [iff del] lemmas induct = old.prod.induct lemmas inducts = old.prod.inducts lemmas rec = old.prod.rec lemmas simps = prod.inject prod.case prod.rec setup \Sign.parent_path\ declare prod.case [nitpick_simp del] declare old.prod.case_cong_weak [cong del] declare prod.case_eq_if [mono] declare prod.split [no_atp] declare prod.split_asm [no_atp] text \ @{thm [source] prod.split} could be declared as \[split]\ done after the Splitter has been speeded up significantly; precompute the constants involved and don't do anything unless the current goal contains one of those constants. \ subsubsection \Tuple syntax\ text \ Patterns -- extends pre-defined type \<^typ>\pttrn\ used in abstractions. \ nonterminal tuple_args and patterns syntax "_tuple" :: "'a \ tuple_args \ 'a \ 'b" ("(1'(_,/ _'))") "_tuple_arg" :: "'a \ tuple_args" ("_") "_tuple_args" :: "'a \ tuple_args \ tuple_args" ("_,/ _") "_pattern" :: "pttrn \ patterns \ pttrn" ("'(_,/ _')") "" :: "pttrn \ patterns" ("_") "_patterns" :: "pttrn \ patterns \ patterns" ("_,/ _") "_unit" :: pttrn ("'(')") translations "(x, y)" \ "CONST Pair x y" "_pattern x y" \ "CONST Pair x y" "_patterns x y" \ "CONST Pair x y" "_tuple x (_tuple_args y z)" \ "_tuple x (_tuple_arg (_tuple y z))" "\(x, y, zs). b" \ "CONST case_prod (\x (y, zs). b)" "\(x, y). b" \ "CONST case_prod (\x y. b)" "_abs (CONST Pair x y) t" \ "\(x, y). t" \ \This rule accommodates tuples in \case C \ (x, y) \ \ \\: The \(x, y)\ is parsed as \Pair x y\ because it is \logic\, not \pttrn\.\ "\(). b" \ "CONST case_unit b" "_abs (CONST Unity) t" \ "\(). t" text \print \<^term>\case_prod f\ as \<^term>\\(x, y). f x y\ and \<^term>\case_prod (\x. f x)\ as \<^term>\\(x, y). f x y\\ typed_print_translation \ let fun case_prod_guess_names_tr' T [Abs (x, _, Abs _)] = raise Match | case_prod_guess_names_tr' T [Abs (x, xT, t)] = (case (head_of t) of Const (\<^const_syntax>\case_prod\, _) => raise Match | _ => let val (_ :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 1 t $ Bound 0); val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, xT, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end) | case_prod_guess_names_tr' T [t] = (case head_of t of Const (\<^const_syntax>\case_prod\, _) => raise Match | _ => let val (xT :: yT :: _) = binder_types (domain_type T) handle Bind => raise Match; val (y, t') = Syntax_Trans.atomic_abs_tr' ("y", yT, incr_boundvars 2 t $ Bound 1 $ Bound 0); val (x', t'') = Syntax_Trans.atomic_abs_tr' ("x", xT, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end) | case_prod_guess_names_tr' _ _ = raise Match; in [(\<^const_syntax>\case_prod\, K case_prod_guess_names_tr')] end \ text \Reconstruct pattern from (nested) \<^const>\case_prod\s, avoiding eta-contraction of body; required for enclosing "let", if "let" does not avoid eta-contraction, which has been observed to occur.\ print_translation \ let fun case_prod_tr' [Abs (x, T, t as (Abs abs))] = (* case_prod (\x y. t) \ \(x, y) t *) let val (y, t') = Syntax_Trans.atomic_abs_tr' abs; val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ y) $ t'' end | case_prod_tr' [Abs (x, T, (s as Const (\<^const_syntax>\case_prod\, _) $ t))] = (* case_prod (\x. (case_prod (\y z. t))) \ \(x, y, z). t *) let val Const (\<^syntax_const>\_abs\, _) $ (Const (\<^syntax_const>\_pattern\, _) $ y $ z) $ t' = case_prod_tr' [t]; val (x', t'') = Syntax_Trans.atomic_abs_tr' (x, T, t'); in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x' $ (Syntax.const \<^syntax_const>\_patterns\ $ y $ z)) $ t'' end | case_prod_tr' [Const (\<^const_syntax>\case_prod\, _) $ t] = (* case_prod (case_prod (\x y z. t)) \ \((x, y), z). t *) case_prod_tr' [(case_prod_tr' [t])] (* inner case_prod_tr' creates next pattern *) | case_prod_tr' [Const (\<^syntax_const>\_abs\, _) $ x_y $ Abs abs] = (* case_prod (\pttrn z. t) \ \(pttrn, z). t *) let val (z, t) = Syntax_Trans.atomic_abs_tr' abs in Syntax.const \<^syntax_const>\_abs\ $ (Syntax.const \<^syntax_const>\_pattern\ $ x_y $ z) $ t end | case_prod_tr' _ = raise Match; in [(\<^const_syntax>\case_prod\, K case_prod_tr')] end \ subsubsection \Code generator setup\ code_printing type_constructor prod \ (SML) infix 2 "*" and (OCaml) infix 2 "*" and (Haskell) "!((_),/ (_))" and (Scala) "((_),/ (_))" | constant Pair \ (SML) "!((_),/ (_))" and (OCaml) "!((_),/ (_))" and (Haskell) "!((_),/ (_))" and (Scala) "!((_),/ (_))" | class_instance prod :: equal \ (Haskell) - | constant "HOL.equal :: 'a \ 'b \ 'a \ 'b \ bool" \ (Haskell) infix 4 "==" | constant fst \ (Haskell) "fst" | constant snd \ (Haskell) "snd" subsubsection \Fundamental operations and properties\ lemma Pair_inject: "(a, b) = (a', b') \ (a = a' \ b = b' \ R) \ R" by simp lemma surj_pair [simp]: "\x y. p = (x, y)" by (cases p) simp lemma fst_eqD: "fst (x, y) = a \ x = a" by simp lemma snd_eqD: "snd (x, y) = a \ y = a" by simp lemma case_prod_unfold [nitpick_unfold]: "case_prod = (\c p. c (fst p) (snd p))" by (simp add: fun_eq_iff split: prod.split) lemma case_prod_conv [simp, code]: "(case (a, b) of (c, d) \ f c d) = f a b" by (fact prod.case) lemmas surjective_pairing = prod.collapse [symmetric] lemma prod_eq_iff: "s = t \ fst s = fst t \ snd s = snd t" by (cases s, cases t) simp lemma prod_eqI [intro?]: "fst p = fst q \ snd p = snd q \ p = q" by (simp add: prod_eq_iff) lemma case_prodI: "f a b \ case (a, b) of (c, d) \ f c d" by (rule prod.case [THEN iffD2]) lemma case_prodD: "(case (a, b) of (c, d) \ f c d) \ f a b" by (rule prod.case [THEN iffD1]) lemma case_prod_Pair [simp]: "case_prod Pair = id" by (simp add: fun_eq_iff split: prod.split) lemma case_prod_eta: "(\(x, y). f (x, y)) = f" \ \Subsumes the old \split_Pair\ when \<^term>\f\ is the identity function.\ by (simp add: fun_eq_iff split: prod.split) (* This looks like a sensible simp-rule but appears to do more harm than good: lemma case_prod_const [simp]: "(\(_,_). c) = (\_. c)" by(rule case_prod_eta) *) lemma case_prod_comp: "(case x of (a, b) \ (f \ g) a b) = f (g (fst x)) (snd x)" by (cases x) simp lemma The_case_prod: "The (case_prod P) = (THE xy. P (fst xy) (snd xy))" by (simp add: case_prod_unfold) lemma cond_case_prod_eta: "(\x y. f x y = g (x, y)) \ (\(x, y). f x y) = g" by (simp add: case_prod_eta) lemma split_paired_all [no_atp]: "(\x. PROP P x) \ (\a b. PROP P (a, b))" proof fix a b assume "\x. PROP P x" then show "PROP P (a, b)" . next fix x assume "\a b. PROP P (a, b)" from \PROP P (fst x, snd x)\ show "PROP P x" by simp qed text \ The rule @{thm [source] split_paired_all} does not work with the Simplifier because it also affects premises in congrence rules, where this can lead to premises of the form \\a b. \ = ?P(a, b)\ which cannot be solved by reflexivity. \ lemmas split_tupled_all = split_paired_all unit_all_eq2 ML \ (* replace parameters of product type by individual component parameters *) local (* filtering with exists_paired_all is an essential optimization *) fun exists_paired_all (Const (\<^const_name>\Pure.all\, _) $ Abs (_, T, t)) = can HOLogic.dest_prodT T orelse exists_paired_all t | exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u | exists_paired_all (Abs (_, _, t)) = exists_paired_all t | exists_paired_all _ = false; val ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm split_paired_all}, @{thm unit_all_eq2}, @{thm unit_abs_eta_conv}] addsimprocs [\<^simproc>\unit_eq\]); in fun split_all_tac ctxt = SUBGOAL (fn (t, i) => if exists_paired_all t then safe_full_simp_tac (put_simpset ss ctxt) i else no_tac); fun unsafe_split_all_tac ctxt = SUBGOAL (fn (t, i) => if exists_paired_all t then full_simp_tac (put_simpset ss ctxt) i else no_tac); fun split_all ctxt th = if exists_paired_all (Thm.prop_of th) then full_simplify (put_simpset ss ctxt) th else th; end; \ setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_all_tac", split_all_tac))\ lemma split_paired_All [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" \ \\[iff]\ is not a good idea because it makes \blast\ loop\ by fast lemma split_paired_Ex [simp, no_atp]: "(\x. P x) \ (\a b. P (a, b))" by fast lemma split_paired_The [no_atp]: "(THE x. P x) = (THE (a, b). P (a, b))" \ \Can't be added to simpset: loops!\ by (simp add: case_prod_eta) text \ Simplification procedure for @{thm [source] cond_case_prod_eta}. Using @{thm [source] case_prod_eta} as a rewrite rule is not general enough, and using @{thm [source] cond_case_prod_eta} directly would render some existing proofs very inefficient; similarly for \prod.case_eq_if\. \ ML \ local val cond_case_prod_eta_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms cond_case_prod_eta}); fun Pair_pat k 0 (Bound m) = (m = k) | Pair_pat k i (Const (\<^const_name>\Pair\, _) $ Bound m $ t) = i > 0 andalso m = k + i andalso Pair_pat k (i - 1) t | Pair_pat _ _ _ = false; fun no_args k i (Abs (_, _, t)) = no_args (k + 1) i t | no_args k i (t $ u) = no_args k i t andalso no_args k i u | no_args k i (Bound m) = m < k orelse m > k + i | no_args _ _ _ = true; fun split_pat tp i (Abs (_, _, t)) = if tp 0 i t then SOME (i, t) else NONE | split_pat tp i (Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t)) = split_pat tp (i + 1) t | split_pat tp i _ = NONE; fun metaeq ctxt lhs rhs = mk_meta_eq (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs))) (K (simp_tac (put_simpset cond_case_prod_eta_ss ctxt) 1))); fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k + 1) i t | beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse (beta_term_pat k i t andalso beta_term_pat k i u) | beta_term_pat k i t = no_args k i t; fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg | eta_term_pat _ _ _ = false; fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) | subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg else (subst arg k i t $ subst arg k i u) | subst arg k i t = t; in fun beta_proc ctxt (s as Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t) $ arg) = (case split_pat beta_term_pat 1 t of SOME (i, f) => SOME (metaeq ctxt s (subst arg 0 i f)) | NONE => NONE) | beta_proc _ _ = NONE; fun eta_proc ctxt (s as Const (\<^const_name>\case_prod\, _) $ Abs (_, _, t)) = (case split_pat eta_term_pat 1 t of SOME (_, ft) => SOME (metaeq ctxt s (let val f $ _ = ft in f end)) | NONE => NONE) | eta_proc _ _ = NONE; end; \ simproc_setup case_prod_beta ("case_prod f z") = - \fn _ => fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => beta_proc ctxt (Thm.term_of ct))\ simproc_setup case_prod_eta ("case_prod f") = - \fn _ => fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct)\ + \K (fn ctxt => fn ct => eta_proc ctxt (Thm.term_of ct))\ lemma case_prod_beta': "(\(x,y). f x y) = (\x. f (fst x) (snd x))" by (auto simp: fun_eq_iff) text \ \<^medskip> \<^const>\case_prod\ used as a logical connective or set former. \<^medskip> These rules are for use with \blast\; could instead call \simp\ using @{thm [source] prod.split} as rewrite.\ lemma case_prodI2: "\p. (\a b. p = (a, b) \ c a b) \ case p of (a, b) \ c a b" by (simp add: split_tupled_all) lemma case_prodI2': "\p. (\a b. (a, b) = p \ c a b x) \ (case p of (a, b) \ c a b) x" by (simp add: split_tupled_all) lemma case_prodE [elim!]: "(case p of (a, b) \ c a b) \ (\x y. p = (x, y) \ c x y \ Q) \ Q" by (induct p) simp lemma case_prodE' [elim!]: "(case p of (a, b) \ c a b) z \ (\x y. p = (x, y) \ c x y z \ Q) \ Q" by (induct p) simp lemma case_prodE2: assumes q: "Q (case z of (a, b) \ P a b)" and r: "\x y. z = (x, y) \ Q (P x y) \ R" shows R proof (rule r) show "z = (fst z, snd z)" by simp then show "Q (P (fst z) (snd z))" using q by (simp add: case_prod_unfold) qed lemma case_prodD': "(case (a, b) of (c, d) \ R c d) c \ R a b c" by simp lemma mem_case_prodI: "z \ c a b \ z \ (case (a, b) of (d, e) \ c d e)" by simp lemma mem_case_prodI2 [intro!]: "\p. (\a b. p = (a, b) \ z \ c a b) \ z \ (case p of (a, b) \ c a b)" by (simp only: split_tupled_all) simp declare mem_case_prodI [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI2' [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI2 [intro!] \ \postponed to maintain traditional declaration order!\ declare case_prodI [intro!] \ \postponed to maintain traditional declaration order!\ lemma mem_case_prodE [elim!]: assumes "z \ case_prod c p" obtains x y where "p = (x, y)" and "z \ c x y" using assms by (rule case_prodE2) ML \ local (* filtering with exists_p_split is an essential optimization *) fun exists_p_split (Const (\<^const_name>\case_prod\,_) $ _ $ (Const (\<^const_name>\Pair\,_)$_$_)) = true | exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u | exists_p_split (Abs (_, _, t)) = exists_p_split t | exists_p_split _ = false; in fun split_conv_tac ctxt = SUBGOAL (fn (t, i) => if exists_p_split t then safe_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms case_prod_conv}) i else no_tac); end; \ (* This prevents applications of splitE for already splitted arguments leading to quite time-consuming computations (in particular for nested tuples) *) setup \map_theory_claset (fn ctxt => ctxt addSbefore ("split_conv_tac", split_conv_tac))\ lemma split_eta_SetCompr [simp, no_atp]: "(\u. \x y. u = (x, y) \ P (x, y)) = P" by (rule ext) fast lemma split_eta_SetCompr2 [simp, no_atp]: "(\u. \x y. u = (x, y) \ P x y) = case_prod P" by (rule ext) fast lemma split_part [simp]: "(\(a,b). P \ Q a b) = (\ab. P \ case_prod Q ab)" \ \Allows simplifications of nested splits in case of independent predicates.\ by (rule ext) blast (* Do NOT make this a simp rule as it a) only helps in special situations b) can lead to nontermination in the presence of split_def *) lemma split_comp_eq: fixes f :: "'a \ 'b \ 'c" and g :: "'d \ 'a" shows "(\u. f (g (fst u)) (snd u)) = case_prod (\x. f (g x))" by (rule ext) auto lemma pair_imageI [intro]: "(a, b) \ A \ f a b \ (\(a, b). f a b) ` A" by (rule image_eqI [where x = "(a, b)"]) auto lemma Collect_const_case_prod[simp]: "{(a,b). P} = (if P then UNIV else {})" by auto lemma The_split_eq [simp]: "(THE (x',y'). x = x' \ y = y') = (x, y)" by blast (* the following would be slightly more general, but cannot be used as rewrite rule: ### Cannot add premise as rewrite rule because it contains (type) unknowns: ### ?y = .x Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" by (rtac some_equality 1) by ( Simp_tac 1) by (split_all_tac 1) by (Asm_full_simp_tac 1) qed "The_split_eq"; *) lemma case_prod_beta: "case_prod f p = f (fst p) (snd p)" by (fact prod.case_eq_if) lemma prod_cases3 [cases type]: obtains (fields) a b c where "y = (a, b, c)" proof (cases y) case (Pair a b) with that show ?thesis by (cases b) blast qed lemma prod_induct3 [case_names fields, induct type]: "(\a b c. P (a, b, c)) \ P x" by (cases x) blast lemma prod_cases4 [cases type]: obtains (fields) a b c d where "y = (a, b, c, d)" proof (cases y) case (fields a b c) with that show ?thesis by (cases c) blast qed lemma prod_induct4 [case_names fields, induct type]: "(\a b c d. P (a, b, c, d)) \ P x" by (cases x) blast lemma prod_cases5 [cases type]: obtains (fields) a b c d e where "y = (a, b, c, d, e)" proof (cases y) case (fields a b c d) with that show ?thesis by (cases d) blast qed lemma prod_induct5 [case_names fields, induct type]: "(\a b c d e. P (a, b, c, d, e)) \ P x" by (cases x) blast lemma prod_cases6 [cases type]: obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" proof (cases y) case (fields a b c d e) with that show ?thesis by (cases e) blast qed lemma prod_induct6 [case_names fields, induct type]: "(\a b c d e f. P (a, b, c, d, e, f)) \ P x" by (cases x) blast lemma prod_cases7 [cases type]: obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" proof (cases y) case (fields a b c d e f) with that show ?thesis by (cases f) blast qed lemma prod_induct7 [case_names fields, induct type]: "(\a b c d e f g. P (a, b, c, d, e, f, g)) \ P x" by (cases x) blast definition internal_case_prod :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "internal_case_prod \ case_prod" lemma internal_case_prod_conv: "internal_case_prod c (a, b) = c a b" by (simp only: internal_case_prod_def case_prod_conv) ML_file \Tools/split_rule.ML\ hide_const internal_case_prod subsubsection \Derived operations\ definition curry :: "('a \ 'b \ 'c) \ 'a \ 'b \ 'c" where "curry = (\c x y. c (x, y))" lemma curry_conv [simp, code]: "curry f a b = f (a, b)" by (simp add: curry_def) lemma curryI [intro!]: "f (a, b) \ curry f a b" by (simp add: curry_def) lemma curryD [dest!]: "curry f a b \ f (a, b)" by (simp add: curry_def) lemma curryE: "curry f a b \ (f (a, b) \ Q) \ Q" by (simp add: curry_def) lemma curry_case_prod [simp]: "curry (case_prod f) = f" by (simp add: curry_def case_prod_unfold) lemma case_prod_curry [simp]: "case_prod (curry f) = f" by (simp add: curry_def case_prod_unfold) lemma curry_K: "curry (\x. c) = (\x y. c)" by (simp add: fun_eq_iff) text \The composition-uncurry combinator.\ definition scomp :: "('a \ 'b \ 'c) \ ('b \ 'c \ 'd) \ 'a \ 'd" (infixl "\\" 60) where "f \\ g = (\x. case_prod g (f x))" no_notation scomp (infixl "\\" 60) bundle state_combinator_syntax begin notation fcomp (infixl "\>" 60) notation scomp (infixl "\\" 60) end context includes state_combinator_syntax begin lemma scomp_unfold: "(\\) = (\f g x. g (fst (f x)) (snd (f x)))" by (simp add: fun_eq_iff scomp_def case_prod_unfold) lemma scomp_apply [simp]: "(f \\ g) x = case_prod g (f x)" by (simp add: scomp_unfold case_prod_unfold) lemma Pair_scomp: "Pair x \\ f = f x" by (simp add: fun_eq_iff) lemma scomp_Pair: "x \\ Pair = x" by (simp add: fun_eq_iff) lemma scomp_scomp: "(f \\ g) \\ h = f \\ (\x. g x \\ h)" by (simp add: fun_eq_iff scomp_unfold) lemma scomp_fcomp: "(f \\ g) \> h = f \\ (\x. g x \> h)" by (simp add: fun_eq_iff scomp_unfold fcomp_def) lemma fcomp_scomp: "(f \> g) \\ h = f \> (g \\ h)" by (simp add: fun_eq_iff scomp_unfold) end code_printing constant scomp \ (Eval) infixl 3 "#->" text \ \<^term>\map_prod\ --- action of the product functor upon functions. \ definition map_prod :: "('a \ 'c) \ ('b \ 'd) \ 'a \ 'b \ 'c \ 'd" where "map_prod f g = (\(x, y). (f x, g y))" lemma map_prod_simp [simp, code]: "map_prod f g (a, b) = (f a, g b)" by (simp add: map_prod_def) functor map_prod: map_prod by (auto simp add: split_paired_all) lemma fst_map_prod [simp]: "fst (map_prod f g x) = f (fst x)" by (cases x) simp_all lemma snd_map_prod [simp]: "snd (map_prod f g x) = g (snd x)" by (cases x) simp_all lemma fst_comp_map_prod [simp]: "fst \ map_prod f g = f \ fst" by (rule ext) simp_all lemma snd_comp_map_prod [simp]: "snd \ map_prod f g = g \ snd" by (rule ext) simp_all lemma map_prod_compose: "map_prod (f1 \ f2) (g1 \ g2) = (map_prod f1 g1 \ map_prod f2 g2)" by (rule ext) (simp add: map_prod.compositionality comp_def) lemma map_prod_ident [simp]: "map_prod (\x. x) (\y. y) = (\z. z)" by (rule ext) (simp add: map_prod.identity) lemma map_prod_imageI [intro]: "(a, b) \ R \ (f a, g b) \ map_prod f g ` R" by (rule image_eqI) simp_all lemma prod_fun_imageE [elim!]: assumes major: "c \ map_prod f g ` R" and cases: "\x y. c = (f x, g y) \ (x, y) \ R \ P" shows P proof (rule major [THEN imageE]) fix x assume "c = map_prod f g x" "x \ R" then show P using cases by (cases x) simp qed definition apfst :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" where "apfst f = map_prod f id" definition apsnd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" where "apsnd f = map_prod id f" lemma apfst_conv [simp, code]: "apfst f (x, y) = (f x, y)" by (simp add: apfst_def) lemma apsnd_conv [simp, code]: "apsnd f (x, y) = (x, f y)" by (simp add: apsnd_def) lemma fst_apfst [simp]: "fst (apfst f x) = f (fst x)" by (cases x) simp lemma fst_comp_apfst [simp]: "fst \ apfst f = f \ fst" by (simp add: fun_eq_iff) lemma fst_apsnd [simp]: "fst (apsnd f x) = fst x" by (cases x) simp lemma fst_comp_apsnd [simp]: "fst \ apsnd f = fst" by (simp add: fun_eq_iff) lemma snd_apfst [simp]: "snd (apfst f x) = snd x" by (cases x) simp lemma snd_comp_apfst [simp]: "snd \ apfst f = snd" by (simp add: fun_eq_iff) lemma snd_apsnd [simp]: "snd (apsnd f x) = f (snd x)" by (cases x) simp lemma snd_comp_apsnd [simp]: "snd \ apsnd f = f \ snd" by (simp add: fun_eq_iff) lemma apfst_compose: "apfst f (apfst g x) = apfst (f \ g) x" by (cases x) simp lemma apsnd_compose: "apsnd f (apsnd g x) = apsnd (f \ g) x" by (cases x) simp lemma apfst_apsnd [simp]: "apfst f (apsnd g x) = (f (fst x), g (snd x))" by (cases x) simp lemma apsnd_apfst [simp]: "apsnd f (apfst g x) = (g (fst x), f (snd x))" by (cases x) simp lemma apfst_id [simp]: "apfst id = id" by (simp add: fun_eq_iff) lemma apsnd_id [simp]: "apsnd id = id" by (simp add: fun_eq_iff) lemma apfst_eq_conv [simp]: "apfst f x = apfst g x \ f (fst x) = g (fst x)" by (cases x) simp lemma apsnd_eq_conv [simp]: "apsnd f x = apsnd g x \ f (snd x) = g (snd x)" by (cases x) simp lemma apsnd_apfst_commute: "apsnd f (apfst g p) = apfst g (apsnd f p)" by simp context begin local_setup \Local_Theory.map_background_naming (Name_Space.mandatory_path "prod")\ definition swap :: "'a \ 'b \ 'b \ 'a" where "swap p = (snd p, fst p)" end lemma swap_simp [simp]: "prod.swap (x, y) = (y, x)" by (simp add: prod.swap_def) lemma swap_swap [simp]: "prod.swap (prod.swap p) = p" by (cases p) simp lemma swap_comp_swap [simp]: "prod.swap \ prod.swap = id" by (simp add: fun_eq_iff) lemma pair_in_swap_image [simp]: "(y, x) \ prod.swap ` A \ (x, y) \ A" by (auto intro!: image_eqI) lemma inj_swap [simp]: "inj_on prod.swap A" by (rule inj_onI) auto lemma swap_inj_on: "inj_on (\(i, j). (j, i)) A" by (rule inj_onI) auto lemma surj_swap [simp]: "surj prod.swap" by (rule surjI [of _ prod.swap]) simp lemma bij_swap [simp]: "bij prod.swap" by (simp add: bij_def) lemma case_swap [simp]: "(case prod.swap p of (y, x) \ f x y) = (case p of (x, y) \ f x y)" by (cases p) simp lemma fst_swap [simp]: "fst (prod.swap x) = snd x" by (cases x) simp lemma snd_swap [simp]: "snd (prod.swap x) = fst x" by (cases x) simp text \Disjoint union of a family of sets -- Sigma.\ definition Sigma :: "'a set \ ('a \ 'b set) \ ('a \ 'b) set" where "Sigma A B \ \x\A. \y\B x. {Pair x y}" abbreviation Times :: "'a set \ 'b set \ ('a \ 'b) set" (infixr "\" 80) where "A \ B \ Sigma A (\_. B)" hide_const (open) Times bundle no_Set_Product_syntax begin no_notation Product_Type.Times (infixr "\" 80) end bundle Set_Product_syntax begin notation Product_Type.Times (infixr "\" 80) end syntax "_Sigma" :: "pttrn \ 'a set \ 'b set \ ('a \ 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) translations "SIGMA x:A. B" \ "CONST Sigma A (\x. B)" lemma SigmaI [intro!]: "a \ A \ b \ B a \ (a, b) \ Sigma A B" unfolding Sigma_def by blast lemma SigmaE [elim!]: "c \ Sigma A B \ (\x y. x \ A \ y \ B x \ c = (x, y) \ P) \ P" \ \The general elimination rule.\ unfolding Sigma_def by blast text \ Elimination of \<^term>\(a, b) \ A \ B\ -- introduces no eigenvariables. \ lemma SigmaD1: "(a, b) \ Sigma A B \ a \ A" by blast lemma SigmaD2: "(a, b) \ Sigma A B \ b \ B a" by blast lemma SigmaE2: "(a, b) \ Sigma A B \ (a \ A \ b \ B a \ P) \ P" by blast lemma Sigma_cong: "A = B \ (\x. x \ B \ C x = D x) \ (SIGMA x:A. C x) = (SIGMA x:B. D x)" by auto lemma Sigma_mono: "A \ C \ (\x. x \ A \ B x \ D x) \ Sigma A B \ Sigma C D" by blast lemma Sigma_empty1 [simp]: "Sigma {} B = {}" by blast lemma Sigma_empty2 [simp]: "A \ {} = {}" by blast lemma UNIV_Times_UNIV [simp]: "UNIV \ UNIV = UNIV" by auto lemma Compl_Times_UNIV1 [simp]: "- (UNIV \ A) = UNIV \ (-A)" by auto lemma Compl_Times_UNIV2 [simp]: "- (A \ UNIV) = (-A) \ UNIV" by auto lemma mem_Sigma_iff [iff]: "(a, b) \ Sigma A B \ a \ A \ b \ B a" by blast lemma mem_Times_iff: "x \ A \ B \ fst x \ A \ snd x \ B" by (induct x) simp lemma Sigma_empty_iff: "(SIGMA i:I. X i) = {} \ (\i\I. X i = {})" by auto lemma Times_subset_cancel2: "x \ C \ A \ C \ B \ C \ A \ B" by blast lemma Times_eq_cancel2: "x \ C \ A \ C = B \ C \ A = B" by (blast elim: equalityE) lemma Collect_case_prod_Sigma: "{(x, y). P x \ Q x y} = (SIGMA x:Collect P. Collect (Q x))" by blast lemma Collect_case_prod [simp]: "{(a, b). P a \ Q b} = Collect P \ Collect Q " by (fact Collect_case_prod_Sigma) lemma Collect_case_prodD: "x \ Collect (case_prod A) \ A (fst x) (snd x)" by auto lemma Collect_case_prod_mono: "A \ B \ Collect (case_prod A) \ Collect (case_prod B)" by auto (auto elim!: le_funE) lemma Collect_split_mono_strong: "X = fst ` A \ Y = snd ` A \ \a\X. \b \ Y. P a b \ Q a b \ A \ Collect (case_prod P) \ A \ Collect (case_prod Q)" by fastforce lemma UN_Times_distrib: "(\(a, b)\A \ B. E a \ F b) = \(E ` A) \ \(F ` B)" \ \Suggested by Pierre Chartier\ by blast lemma split_paired_Ball_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast lemma split_paired_Bex_Sigma [simp, no_atp]: "(\z\Sigma A B. P z) \ (\x\A. \y\B x. P (x, y))" by blast lemma Sigma_Un_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast lemma Sigma_Un_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast lemma Sigma_Int_distrib1: "Sigma (I \ J) C = Sigma I C \ Sigma J C" by blast lemma Sigma_Int_distrib2: "(SIGMA i:I. A i \ B i) = Sigma I A \ Sigma I B" by blast lemma Sigma_Diff_distrib1: "Sigma (I - J) C = Sigma I C - Sigma J C" by blast lemma Sigma_Diff_distrib2: "(SIGMA i:I. A i - B i) = Sigma I A - Sigma I B" by blast lemma Sigma_Union: "Sigma (\X) B = (\A\X. Sigma A B)" by blast lemma Pair_vimage_Sigma: "Pair x -` Sigma A f = (if x \ A then f x else {})" by auto text \ Non-dependent versions are needed to avoid the need for higher-order matching, especially when the rules are re-oriented. \ lemma Times_Un_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Un_distrib1) lemma Times_Int_distrib1: "(A \ B) \ C = A \ C \ B \ C " by (fact Sigma_Int_distrib1) lemma Times_Diff_distrib1: "(A - B) \ C = A \ C - B \ C " by (fact Sigma_Diff_distrib1) lemma Times_empty [simp]: "A \ B = {} \ A = {} \ B = {}" by auto lemma times_subset_iff: "A \ C \ B \ D \ A={} \ C={} \ A \ B \ C \ D" by blast lemma times_eq_iff: "A \ B = C \ D \ A = C \ B = D \ (A = {} \ B = {}) \ (C = {} \ D = {})" by auto lemma fst_image_times [simp]: "fst ` (A \ B) = (if B = {} then {} else A)" by force lemma snd_image_times [simp]: "snd ` (A \ B) = (if A = {} then {} else B)" by force lemma fst_image_Sigma: "fst ` (Sigma A B) = {x \ A. B(x) \ {}}" by force lemma snd_image_Sigma: "snd ` (Sigma A B) = (\ x \ A. B x)" by force lemma vimage_fst: "fst -` A = A \ UNIV" by auto lemma vimage_snd: "snd -` A = UNIV \ A" by auto lemma insert_Times_insert [simp]: "insert a A \ insert b B = insert (a,b) (A \ insert b B \ insert a A \ B)" by blast lemma vimage_Times: "f -` (A \ B) = (fst \ f) -` A \ (snd \ f) -` B" proof (rule set_eqI) show "x \ f -` (A \ B) \ x \ (fst \ f) -` A \ (snd \ f) -` B" for x by (cases "f x") (auto split: prod.split) qed lemma Times_Int_Times: "A \ B \ C \ D = (A \ C) \ (B \ D)" by auto lemma image_paired_Times: "(\(x,y). (f x, g y)) ` (A \ B) = (f ` A) \ (g ` B)" by auto lemma product_swap: "prod.swap ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) lemma swap_product: "(\(i, j). (j, i)) ` (A \ B) = B \ A" by (auto simp add: set_eq_iff) lemma image_split_eq_Sigma: "(\x. (f x, g x)) ` A = Sigma (f ` A) (\x. g ` (f -` {x} \ A))" proof (safe intro!: imageI) fix a b assume *: "a \ A" "b \ A" and eq: "f a = f b" show "(f b, g a) \ (\x. (f x, g x)) ` A" using * eq[symmetric] by auto qed simp_all lemma subset_fst_snd: "A \ (fst ` A \ snd ` A)" by force lemma inj_on_apfst [simp]: "inj_on (apfst f) (A \ UNIV) \ inj_on f A" by (auto simp add: inj_on_def) lemma inj_apfst [simp]: "inj (apfst f) \ inj f" using inj_on_apfst[of f UNIV] by simp lemma inj_on_apsnd [simp]: "inj_on (apsnd f) (UNIV \ A) \ inj_on f A" by (auto simp add: inj_on_def) lemma inj_apsnd [simp]: "inj (apsnd f) \ inj f" using inj_on_apsnd[of f UNIV] by simp context begin qualified definition product :: "'a set \ 'b set \ ('a \ 'b) set" where [code_abbrev]: "product A B = A \ B" lemma member_product: "x \ Product_Type.product A B \ x \ A \ B" by (simp add: product_def) end text \The following \<^const>\map_prod\ lemmas are due to Joachim Breitner:\ lemma map_prod_inj_on: assumes "inj_on f A" and "inj_on g B" shows "inj_on (map_prod f g) (A \ B)" proof (rule inj_onI) fix x :: "'a \ 'c" fix y :: "'a \ 'c" assume "x \ A \ B" then have "fst x \ A" and "snd x \ B" by auto assume "y \ A \ B" then have "fst y \ A" and "snd y \ B" by auto assume "map_prod f g x = map_prod f g y" then have "fst (map_prod f g x) = fst (map_prod f g y)" by auto then have "f (fst x) = f (fst y)" by (cases x, cases y) auto with \inj_on f A\ and \fst x \ A\ and \fst y \ A\ have "fst x = fst y" by (auto dest: inj_onD) moreover from \map_prod f g x = map_prod f g y\ have "snd (map_prod f g x) = snd (map_prod f g y)" by auto then have "g (snd x) = g (snd y)" by (cases x, cases y) auto with \inj_on g B\ and \snd x \ B\ and \snd y \ B\ have "snd x = snd y" by (auto dest: inj_onD) ultimately show "x = y" by (rule prod_eqI) qed lemma map_prod_surj: fixes f :: "'a \ 'b" and g :: "'c \ 'd" assumes "surj f" and "surj g" shows "surj (map_prod f g)" unfolding surj_def proof fix y :: "'b \ 'd" from \surj f\ obtain a where "fst y = f a" by (auto elim: surjE) moreover from \surj g\ obtain b where "snd y = g b" by (auto elim: surjE) ultimately have "(fst y, snd y) = map_prod f g (a,b)" by auto then show "\x. y = map_prod f g x" by auto qed lemma map_prod_surj_on: assumes "f ` A = A'" and "g ` B = B'" shows "map_prod f g ` (A \ B) = A' \ B'" unfolding image_def proof (rule set_eqI, rule iffI) fix x :: "'a \ 'c" assume "x \ {y::'a \ 'c. \x::'b \ 'd\A \ B. y = map_prod f g x}" then obtain y where "y \ A \ B" and "x = map_prod f g y" by blast from \image f A = A'\ and \y \ A \ B\ have "f (fst y) \ A'" by auto moreover from \image g B = B'\ and \y \ A \ B\ have "g (snd y) \ B'" by auto ultimately have "(f (fst y), g (snd y)) \ (A' \ B')" by auto with \x = map_prod f g y\ show "x \ A' \ B'" by (cases y) auto next fix x :: "'a \ 'c" assume "x \ A' \ B'" then have "fst x \ A'" and "snd x \ B'" by auto from \image f A = A'\ and \fst x \ A'\ have "fst x \ image f A" by auto then obtain a where "a \ A" and "fst x = f a" by (rule imageE) moreover from \image g B = B'\ and \snd x \ B'\ obtain b where "b \ B" and "snd x = g b" by auto ultimately have "(fst x, snd x) = map_prod f g (a, b)" by auto moreover from \a \ A\ and \b \ B\ have "(a , b) \ A \ B" by auto ultimately have "\y \ A \ B. x = map_prod f g y" by auto then show "x \ {x. \y \ A \ B. x = map_prod f g y}" by auto qed subsection \Simproc for rewriting a set comprehension into a pointfree expression\ ML_file \Tools/set_comprehension_pointfree.ML\ setup \ Code_Preproc.map_pre (fn ctxt => ctxt addsimprocs [Simplifier.make_simproc \<^context> "set comprehension" {lhss = [\<^term>\Collect P\], proc = K Set_Comprehension_Pointfree.code_simproc}]) \ subsection \Lemmas about disjointness\ lemma disjnt_Times1_iff [simp]: "disjnt (C \ A) (C \ B) \ C = {} \ disjnt A B" by (auto simp: disjnt_def) lemma disjnt_Times2_iff [simp]: "disjnt (A \ C) (B \ C) \ C = {} \ disjnt A B" by (auto simp: disjnt_def) lemma disjnt_Sigma_iff: "disjnt (Sigma A C) (Sigma B C) \ (\i \ A\B. C i = {}) \ disjnt A B" by (auto simp: disjnt_def) subsection \Inductively defined sets\ (* simplify {(x1, ..., xn). (x1, ..., xn) : S} to S *) simproc_setup Collect_mem ("Collect t") = \ - fn _ => fn ctxt => fn ct => + K (fn ctxt => fn ct => (case Thm.term_of ct of S as Const (\<^const_name>\Collect\, Type (\<^type_name>\fun\, [_, T])) $ t => let val (u, _, ps) = HOLogic.strip_ptupleabs t in (case u of (c as Const (\<^const_name>\Set.member\, _)) $ q $ S' => (case try (HOLogic.strip_ptuple ps) q of NONE => NONE | SOME ts => if not (Term.is_open S') andalso ts = map Bound (length ps downto 0) then let val simp = full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_paired_all}, @{thm case_prod_conv}]) 1 in SOME (Goal.prove ctxt [] [] (Const (\<^const_name>\Pure.eq\, T --> T --> propT) $ S $ S') (K (EVERY [resolve_tac ctxt [eq_reflection] 1, resolve_tac ctxt @{thms subset_antisym} 1, resolve_tac ctxt @{thms subsetI} 1, dresolve_tac ctxt @{thms CollectD} 1, simp, resolve_tac ctxt @{thms subsetI} 1, resolve_tac ctxt @{thms CollectI} 1, simp]))) end else NONE) | _ => NONE) end - | _ => NONE) + | _ => NONE)) \ ML_file \Tools/inductive_set.ML\ subsection \Legacy theorem bindings and duplicates\ lemmas fst_conv = prod.sel(1) lemmas snd_conv = prod.sel(2) lemmas split_def = case_prod_unfold lemmas split_beta' = case_prod_beta' lemmas split_beta = prod.case_eq_if lemmas split_conv = case_prod_conv lemmas split = case_prod_conv hide_const (open) prod end diff --git a/src/HOL/Real_Asymp/exp_log_expression.ML b/src/HOL/Real_Asymp/exp_log_expression.ML --- a/src/HOL/Real_Asymp/exp_log_expression.ML +++ b/src/HOL/Real_Asymp/exp_log_expression.ML @@ -1,689 +1,689 @@ signature EXP_LOG_EXPRESSION = sig exception DUP datatype expr = ConstExpr of term | X | Uminus of expr | Add of expr * expr | Minus of expr * expr | Inverse of expr | Mult of expr * expr | Div of expr * expr | Ln of expr | Exp of expr | Power of expr * term | LnPowr of expr * expr | ExpLn of expr | Powr of expr * expr | Powr_Nat of expr * expr | Powr' of expr * term | Root of expr * term | Absolute of expr | Sgn of expr | Min of expr * expr | Max of expr * expr | Floor of expr | Ceiling of expr | Frac of expr | NatMod of expr * expr | Sin of expr | Cos of expr | ArcTan of expr | Custom of string * term * expr list val preproc_term_conv : Proof.context -> conv val expr_to_term : expr -> term val reify : Proof.context -> term -> expr * thm val reify_simple : Proof.context -> term -> expr * thm type custom_handler = Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis val register_custom : binding -> term -> custom_handler -> local_theory -> local_theory val register_custom_from_thm : binding -> thm -> custom_handler -> local_theory -> local_theory val expand_custom : Proof.context -> string -> custom_handler option val to_mathematica : expr -> string val to_maple : expr -> string val to_maxima : expr -> string val to_sympy : expr -> string val to_sage : expr -> string val reify_mathematica : Proof.context -> term -> string val reify_maple : Proof.context -> term -> string val reify_maxima : Proof.context -> term -> string val reify_sympy : Proof.context -> term -> string val reify_sage : Proof.context -> term -> string val limit_mathematica : string -> string val limit_maple : string -> string val limit_maxima : string -> string val limit_sympy : string -> string val limit_sage : string -> string end structure Exp_Log_Expression : EXP_LOG_EXPRESSION = struct datatype expr = ConstExpr of term | X | Uminus of expr | Add of expr * expr | Minus of expr * expr | Inverse of expr | Mult of expr * expr | Div of expr * expr | Ln of expr | Exp of expr | Power of expr * term | LnPowr of expr * expr | ExpLn of expr | Powr of expr * expr | Powr_Nat of expr * expr | Powr' of expr * term | Root of expr * term | Absolute of expr | Sgn of expr | Min of expr * expr | Max of expr * expr | Floor of expr | Ceiling of expr | Frac of expr | NatMod of expr * expr | Sin of expr | Cos of expr | ArcTan of expr | Custom of string * term * expr list type custom_handler = Lazy_Eval.eval_ctxt -> term -> thm list * Asymptotic_Basis.basis -> thm * Asymptotic_Basis.basis type entry = { name : string, pat : term, net_pat : term, expand : custom_handler } type entry' = { pat : term, net_pat : term, expand : custom_handler } exception DUP structure Custom_Funs = Generic_Data ( type T = { name_table : entry' Name_Space.table, net : entry Item_Net.T } fun eq_entry ({name = name1, ...}, {name = name2, ...}) = (name1 = name2) val empty = { name_table = Name_Space.empty_table "exp_log_custom_function", net = Item_Net.init eq_entry (fn {net_pat, ...} => [net_pat]) } fun merge ({name_table = tbl1, net = net1}, {name_table = tbl2, net = net2}) = {name_table = Name_Space.join_tables (fn _ => raise DUP) (tbl1, tbl2), net = Item_Net.merge (net1, net2)} ) fun rewrite' ctxt old_prems bounds thms ct = let val thy = Proof_Context.theory_of ctxt fun apply_rule t thm = let val lhs = thm |> Thm.concl_of |> Logic.dest_equals |> fst val _ = Pattern.first_order_match thy (lhs, t) (Vartab.empty, Vartab.empty) val insts = (lhs, t) |> apply2 (Thm.cterm_of ctxt) |> Thm.first_order_match val thm = Thm.instantiate insts thm val prems = Thm.prems_of thm val frees = fold Term.add_frees prems [] in if exists (member op = bounds o fst) frees then NONE else let val thm' = thm OF (map (Thm.assume o Thm.cterm_of ctxt) prems) val prems' = fold (insert op aconv) prems old_prems val crhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd |> Thm.cterm_of ctxt in SOME (thm', crhs, prems') end end handle Pattern.MATCH => NONE fun rewrite_subterm prems ct (Abs (x, _, _)) = let val ((v, ct'), ctxt') = Variable.dest_abs_cterm ct ctxt; val (thm, prems) = rewrite' ctxt' prems (x :: bounds) thms ct' in if Thm.is_reflexive thm then (Thm.reflexive ct, prems) else (Thm.abstract_rule x v thm, prems) end | rewrite_subterm prems ct (_ $ _) = let val (cs, ct) = Thm.dest_comb ct val (thm, prems') = rewrite' ctxt prems bounds thms cs val (thm', prems'') = rewrite' ctxt prems' bounds thms ct in (Thm.combination thm thm', prems'') end | rewrite_subterm prems ct _ = (Thm.reflexive ct, prems) val t = Thm.term_of ct in case get_first (apply_rule t) thms of NONE => rewrite_subterm old_prems ct t | SOME (thm, rhs, prems) => case rewrite' ctxt prems bounds thms rhs of (thm', prems) => (Thm.transitive thm thm', prems) end fun rewrite ctxt thms ct = let val thm1 = Thm.eta_long_conversion ct val rhs = thm1 |> Thm.cprop_of |> Thm.dest_comb |> snd val (thm2, prems) = rewrite' ctxt [] [] thms rhs val rhs = thm2 |> Thm.cprop_of |> Thm.dest_comb |> snd val thm3 = Thm.eta_conversion rhs val thm = Thm.transitive thm1 (Thm.transitive thm2 thm3) in fold (fn prem => fn thm => Thm.implies_intr (Thm.cterm_of ctxt prem) thm) prems thm end fun preproc_term_conv ctxt = let val thms = Named_Theorems.get ctxt \<^named_theorems>\real_asymp_reify_simps\ val thms = map (fn thm => thm RS @{thm HOL.eq_reflection}) thms in rewrite ctxt thms end fun register_custom' binding pat expand context = let val n = pat |> fastype_of |> strip_type |> fst |> length val maxidx = Term.maxidx_of_term pat val vars = map (fn i => Var ((Name.uu_, maxidx + i), \<^typ>\real\)) (1 upto n) val net_pat = Library.foldl betapply (pat, vars) val {name_table = tbl, net = net} = Custom_Funs.get context val entry' = {pat = pat, net_pat = net_pat, expand = expand} val (name, tbl) = Name_Space.define context true (binding, entry') tbl val entry = {name = name, pat = pat, net_pat = net_pat, expand = expand} val net = Item_Net.update entry net in Custom_Funs.put {name_table = tbl, net = net} context end fun register_custom binding pat expand = let fun decl phi = register_custom' binding (Morphism.term phi pat) expand in - Local_Theory.declaration {syntax = false, pervasive = false} decl + Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} decl end fun register_custom_from_thm binding thm expand = let val pat = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_comb |> snd in register_custom binding pat expand end fun expand_custom ctxt name = let val {name_table, ...} = Custom_Funs.get (Context.Proof ctxt) in case Name_Space.lookup name_table name of NONE => NONE | SOME {expand, ...} => SOME expand end fun expr_to_term e = let fun expr_to_term' (ConstExpr c) = c | expr_to_term' X = Bound 0 | expr_to_term' (Add (a, b)) = \<^term>\(+) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Mult (a, b)) = \<^term>\(*) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Minus (a, b)) = \<^term>\(-) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Div (a, b)) = \<^term>\(/) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Uminus a) = \<^term>\uminus :: real => _\ $ expr_to_term' a | expr_to_term' (Inverse a) = \<^term>\inverse :: real => _\ $ expr_to_term' a | expr_to_term' (Ln a) = \<^term>\ln :: real => _\ $ expr_to_term' a | expr_to_term' (Exp a) = \<^term>\exp :: real => _\ $ expr_to_term' a | expr_to_term' (Powr (a,b)) = \<^term>\(powr) :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Powr_Nat (a,b)) = \<^term>\powr_nat :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (LnPowr (a,b)) = \<^term>\ln :: real => _\ $ (\<^term>\(powr) :: real => _\ $ expr_to_term' a $ expr_to_term' b) | expr_to_term' (ExpLn a) = \<^term>\exp :: real => _\ $ (\<^term>\ln :: real => _\ $ expr_to_term' a) | expr_to_term' (Powr' (a,b)) = \<^term>\(powr) :: real => _\ $ expr_to_term' a $ b | expr_to_term' (Power (a,b)) = \<^term>\(^) :: real => _\ $ expr_to_term' a $ b | expr_to_term' (Floor a) = \<^term>\Multiseries_Expansion.rfloor\ $ expr_to_term' a | expr_to_term' (Ceiling a) = \<^term>\Multiseries_Expansion.rceil\ $ expr_to_term' a | expr_to_term' (Frac a) = \<^term>\Archimedean_Field.frac :: real \ real\ $ expr_to_term' a | expr_to_term' (NatMod (a,b)) = \<^term>\Multiseries_Expansion.rnatmod\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Root (a,b)) = \<^term>\root :: nat \ real \ _\ $ b $ expr_to_term' a | expr_to_term' (Sin a) = \<^term>\sin :: real => _\ $ expr_to_term' a | expr_to_term' (ArcTan a) = \<^term>\arctan :: real => _\ $ expr_to_term' a | expr_to_term' (Cos a) = \<^term>\cos :: real => _\ $ expr_to_term' a | expr_to_term' (Absolute a) = \<^term>\abs :: real => _\ $ expr_to_term' a | expr_to_term' (Sgn a) = \<^term>\sgn :: real => _\ $ expr_to_term' a | expr_to_term' (Min (a,b)) = \<^term>\min :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Max (a,b)) = \<^term>\max :: real => _\ $ expr_to_term' a $ expr_to_term' b | expr_to_term' (Custom (_, t, args)) = Envir.beta_eta_contract ( fold (fn e => fn t => betapply (t, expr_to_term' e )) args t) in Abs ("x", \<^typ>\real\, expr_to_term' e) end fun reify_custom ctxt t = let val thy = Proof_Context.theory_of ctxt val t = Envir.beta_eta_contract t val t' = Envir.beta_eta_contract (Term.abs ("x", \<^typ>\real\) t) val {net, ...} = Custom_Funs.get (Context.Proof ctxt) val entries = Item_Net.retrieve_matching net (Term.subst_bound (Free ("x", \<^typ>\real\), t)) fun go {pat, name, ...} = let val n = pat |> fastype_of |> strip_type |> fst |> length val maxidx = Term.maxidx_of_term t' val vs = map (fn i => (Name.uu_, maxidx + i)) (1 upto n) val args = map (fn v => Var (v, \<^typ>\real => real\) $ Bound 0) vs val pat' = Envir.beta_eta_contract (Term.abs ("x", \<^typ>\real\) (Library.foldl betapply (pat, args))) val (T_insts, insts) = Pattern.match thy (pat', t') (Vartab.empty, Vartab.empty) fun map_option _ [] acc = SOME (rev acc) | map_option f (x :: xs) acc = case f x of NONE => NONE | SOME y => map_option f xs (y :: acc) val t = Envir.subst_term (T_insts, insts) pat in Option.map (pair (name, t)) (map_option (Option.map snd o Vartab.lookup insts) vs []) end handle Pattern.MATCH => NONE in get_first go entries end fun reify_aux ctxt t' t = let fun is_const t = fastype_of (Abs ("x", \<^typ>\real\, t)) = \<^typ>\real \ real\ andalso not (exists_subterm (fn t => t = Bound 0) t) fun is_const' t = not (exists_subterm (fn t => t = Bound 0) t) fun reify'' (\<^term>\(+) :: real => _\ $ s $ t) = Add (reify' s, reify' t) | reify'' (\<^term>\(-) :: real => _\ $ s $ t) = Minus (reify' s, reify' t) | reify'' (\<^term>\(*) :: real => _\ $ s $ t) = Mult (reify' s, reify' t) | reify'' (\<^term>\(/) :: real => _\ $ s $ t) = Div (reify' s, reify' t) | reify'' (\<^term>\uminus :: real => _\ $ s) = Uminus (reify' s) | reify'' (\<^term>\inverse :: real => _\ $ s) = Inverse (reify' s) | reify'' (\<^term>\ln :: real => _\ $ (\<^term>\(powr) :: real => _\ $ s $ t)) = LnPowr (reify' s, reify' t) | reify'' (\<^term>\exp :: real => _\ $ (\<^term>\ln :: real => _\ $ s)) = ExpLn (reify' s) | reify'' (\<^term>\ln :: real => _\ $ s) = Ln (reify' s) | reify'' (\<^term>\exp :: real => _\ $ s) = Exp (reify' s) | reify'' (\<^term>\(powr) :: real => _\ $ s $ t) = (if is_const t then Powr' (reify' s, t) else Powr (reify' s, reify' t)) | reify'' (\<^term>\powr_nat :: real => _\ $ s $ t) = Powr_Nat (reify' s, reify' t) | reify'' (\<^term>\(^) :: real => _\ $ s $ t) = (if is_const' t then Power (reify' s, t) else raise TERM ("reify", [t'])) | reify'' (\<^term>\root\ $ s $ t) = (if is_const' s then Root (reify' t, s) else raise TERM ("reify", [t'])) | reify'' (\<^term>\abs :: real => _\ $ s) = Absolute (reify' s) | reify'' (\<^term>\sgn :: real => _\ $ s) = Sgn (reify' s) | reify'' (\<^term>\min :: real => _\ $ s $ t) = Min (reify' s, reify' t) | reify'' (\<^term>\max :: real => _\ $ s $ t) = Max (reify' s, reify' t) | reify'' (\<^term>\Multiseries_Expansion.rfloor\ $ s) = Floor (reify' s) | reify'' (\<^term>\Multiseries_Expansion.rceil\ $ s) = Ceiling (reify' s) | reify'' (\<^term>\Archimedean_Field.frac :: real \ real\ $ s) = Frac (reify' s) | reify'' (\<^term>\Multiseries_Expansion.rnatmod\ $ s $ t) = NatMod (reify' s, reify' t) | reify'' (\<^term>\sin :: real => _\ $ s) = Sin (reify' s) | reify'' (\<^term>\arctan :: real => _\ $ s) = ArcTan (reify' s) | reify'' (\<^term>\cos :: real => _\ $ s) = Cos (reify' s) | reify'' (Bound 0) = X | reify'' t = case reify_custom ctxt t of SOME ((name, t), ts) => let val args = map (reify_aux ctxt t') ts in Custom (name, t, args) end | NONE => raise TERM ("reify", [t']) and reify' t = if is_const t then ConstExpr t else reify'' t in case Envir.eta_long [] t of Abs (_, \<^typ>\real\, t'') => reify' t'' | _ => raise TERM ("reify", [t]) end fun reify ctxt t = let val thm = preproc_term_conv ctxt (Thm.cterm_of ctxt t) val rhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd in (reify_aux ctxt t rhs, thm) end fun reify_simple_aux ctxt t' t = let fun is_const t = fastype_of (Abs ("x", \<^typ>\real\, t)) = \<^typ>\real \ real\ andalso not (exists_subterm (fn t => t = Bound 0) t) fun is_const' t = not (exists_subterm (fn t => t = Bound 0) t) fun reify'' (\<^term>\(+) :: real => _\ $ s $ t) = Add (reify'' s, reify'' t) | reify'' (\<^term>\(-) :: real => _\ $ s $ t) = Minus (reify'' s, reify'' t) | reify'' (\<^term>\(*) :: real => _\ $ s $ t) = Mult (reify'' s, reify'' t) | reify'' (\<^term>\(/) :: real => _\ $ s $ t) = Div (reify'' s, reify'' t) | reify'' (\<^term>\uminus :: real => _\ $ s) = Uminus (reify'' s) | reify'' (\<^term>\inverse :: real => _\ $ s) = Inverse (reify'' s) | reify'' (\<^term>\ln :: real => _\ $ s) = Ln (reify'' s) | reify'' (\<^term>\exp :: real => _\ $ s) = Exp (reify'' s) | reify'' (\<^term>\(powr) :: real => _\ $ s $ t) = Powr (reify'' s, reify'' t) | reify'' (\<^term>\powr_nat :: real => _\ $ s $ t) = Powr_Nat (reify'' s, reify'' t) | reify'' (\<^term>\(^) :: real => _\ $ s $ t) = (if is_const' t then Power (reify'' s, t) else raise TERM ("reify", [t'])) | reify'' (\<^term>\root\ $ s $ t) = (if is_const' s then Root (reify'' t, s) else raise TERM ("reify", [t'])) | reify'' (\<^term>\abs :: real => _\ $ s) = Absolute (reify'' s) | reify'' (\<^term>\sgn :: real => _\ $ s) = Sgn (reify'' s) | reify'' (\<^term>\min :: real => _\ $ s $ t) = Min (reify'' s, reify'' t) | reify'' (\<^term>\max :: real => _\ $ s $ t) = Max (reify'' s, reify'' t) | reify'' (\<^term>\Multiseries_Expansion.rfloor\ $ s) = Floor (reify'' s) | reify'' (\<^term>\Multiseries_Expansion.rceil\ $ s) = Ceiling (reify'' s) | reify'' (\<^term>\Archimedean_Field.frac :: real \ real\ $ s) = Frac (reify'' s) | reify'' (\<^term>\Multiseries_Expansion.rnatmod\ $ s $ t) = NatMod (reify'' s, reify'' t) | reify'' (\<^term>\sin :: real => _\ $ s) = Sin (reify'' s) | reify'' (\<^term>\cos :: real => _\ $ s) = Cos (reify'' s) | reify'' (Bound 0) = X | reify'' t = if is_const t then ConstExpr t else case reify_custom ctxt t of SOME ((name, t), ts) => let val args = map (reify_aux ctxt t') ts in Custom (name, t, args) end | NONE => raise TERM ("reify", [t']) in case Envir.eta_long [] t of Abs (_, \<^typ>\real\, t'') => reify'' t'' | _ => raise TERM ("reify", [t]) end fun reify_simple ctxt t = let val thm = preproc_term_conv ctxt (Thm.cterm_of ctxt t) val rhs = thm |> Thm.concl_of |> Logic.dest_equals |> snd in (reify_simple_aux ctxt t rhs, thm) end fun simple_print_const (Free (x, _)) = x | simple_print_const (\<^term>\uminus :: real => real\ $ a) = "(-" ^ simple_print_const a ^ ")" | simple_print_const (\<^term>\(+) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "+" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\(-) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "-" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\(*) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "*" ^ simple_print_const b ^ ")" | simple_print_const (\<^term>\inverse :: real => _\ $ a) = "(1 / " ^ simple_print_const a ^ ")" | simple_print_const (\<^term>\(/) :: real => _\ $ a $ b) = "(" ^ simple_print_const a ^ "/" ^ simple_print_const b ^ ")" | simple_print_const t = Int.toString (snd (HOLogic.dest_number t)) fun to_mathematica (Add (a, b)) = "(" ^ to_mathematica a ^ " + " ^ to_mathematica b ^ ")" | to_mathematica (Minus (a, b)) = "(" ^ to_mathematica a ^ " - " ^ to_mathematica b ^ ")" | to_mathematica (Mult (a, b)) = "(" ^ to_mathematica a ^ " * " ^ to_mathematica b ^ ")" | to_mathematica (Div (a, b)) = "(" ^ to_mathematica a ^ " / " ^ to_mathematica b ^ ")" | to_mathematica (Powr (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ ")" | to_mathematica (Powr_Nat (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ ")" | to_mathematica (Powr' (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica (ConstExpr b) ^ ")" | to_mathematica (LnPowr (a, b)) = "Log[" ^ to_mathematica a ^ " ^ " ^ to_mathematica b ^ "]" | to_mathematica (ExpLn a) = "Exp[Ln[" ^ to_mathematica a ^ "]]" | to_mathematica (Power (a, b)) = "(" ^ to_mathematica a ^ " ^ " ^ to_mathematica (ConstExpr b) ^ ")" | to_mathematica (Root (a, \<^term>\2::real\)) = "Sqrt[" ^ to_mathematica a ^ "]" | to_mathematica (Root (a, b)) = "Surd[" ^ to_mathematica a ^ ", " ^ to_mathematica (ConstExpr b) ^ "]" | to_mathematica (Uminus a) = "(-" ^ to_mathematica a ^ ")" | to_mathematica (Inverse a) = "(1/(" ^ to_mathematica a ^ "))" | to_mathematica (Exp a) = "Exp[" ^ to_mathematica a ^ "]" | to_mathematica (Ln a) = "Log[" ^ to_mathematica a ^ "]" | to_mathematica (Sin a) = "Sin[" ^ to_mathematica a ^ "]" | to_mathematica (Cos a) = "Cos[" ^ to_mathematica a ^ "]" | to_mathematica (ArcTan a) = "ArcTan[" ^ to_mathematica a ^ "]" | to_mathematica (Absolute a) = "Abs[" ^ to_mathematica a ^ "]" | to_mathematica (Sgn a) = "Sign[" ^ to_mathematica a ^ "]" | to_mathematica (Min (a, b)) = "Min[" ^ to_mathematica a ^ ", " ^ to_mathematica b ^ "]" | to_mathematica (Max (a, b)) = "Max[" ^ to_mathematica a ^ ", " ^ to_mathematica b ^ "]" | to_mathematica (Floor a) = "Floor[" ^ to_mathematica a ^ "]" | to_mathematica (Ceiling a) = "Ceiling[" ^ to_mathematica a ^ "]" | to_mathematica (Frac a) = "Mod[" ^ to_mathematica a ^ ", 1]" | to_mathematica (ConstExpr t) = simple_print_const t | to_mathematica X = "X" (* TODO: correct translation of frac() for Maple and Sage *) fun to_maple (Add (a, b)) = "(" ^ to_maple a ^ " + " ^ to_maple b ^ ")" | to_maple (Minus (a, b)) = "(" ^ to_maple a ^ " - " ^ to_maple b ^ ")" | to_maple (Mult (a, b)) = "(" ^ to_maple a ^ " * " ^ to_maple b ^ ")" | to_maple (Div (a, b)) = "(" ^ to_maple a ^ " / " ^ to_maple b ^ ")" | to_maple (Powr (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (Powr_Nat (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (Powr' (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple (ConstExpr b) ^ ")" | to_maple (LnPowr (a, b)) = "ln(" ^ to_maple a ^ " ^ " ^ to_maple b ^ ")" | to_maple (ExpLn a) = "ln(exp(" ^ to_maple a ^ "))" | to_maple (Power (a, b)) = "(" ^ to_maple a ^ " ^ " ^ to_maple (ConstExpr b) ^ ")" | to_maple (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_maple a ^ ")" | to_maple (Root (a, b)) = "root(" ^ to_maple a ^ ", " ^ to_maple (ConstExpr b) ^ ")" | to_maple (Uminus a) = "(-" ^ to_maple a ^ ")" | to_maple (Inverse a) = "(1/(" ^ to_maple a ^ "))" | to_maple (Exp a) = "exp(" ^ to_maple a ^ ")" | to_maple (Ln a) = "ln(" ^ to_maple a ^ ")" | to_maple (Sin a) = "sin(" ^ to_maple a ^ ")" | to_maple (Cos a) = "cos(" ^ to_maple a ^ ")" | to_maple (ArcTan a) = "arctan(" ^ to_maple a ^ ")" | to_maple (Absolute a) = "abs(" ^ to_maple a ^ ")" | to_maple (Sgn a) = "signum(" ^ to_maple a ^ ")" | to_maple (Min (a, b)) = "min(" ^ to_maple a ^ ", " ^ to_maple b ^ ")" | to_maple (Max (a, b)) = "max(" ^ to_maple a ^ ", " ^ to_maple b ^ ")" | to_maple (Floor a) = "floor(" ^ to_maple a ^ ")" | to_maple (Ceiling a) = "ceil(" ^ to_maple a ^ ")" | to_maple (Frac a) = "frac(" ^ to_maple a ^ ")" | to_maple (ConstExpr t) = simple_print_const t | to_maple X = "x" fun to_maxima (Add (a, b)) = "(" ^ to_maxima a ^ " + " ^ to_maxima b ^ ")" | to_maxima (Minus (a, b)) = "(" ^ to_maxima a ^ " - " ^ to_maxima b ^ ")" | to_maxima (Mult (a, b)) = "(" ^ to_maxima a ^ " * " ^ to_maxima b ^ ")" | to_maxima (Div (a, b)) = "(" ^ to_maxima a ^ " / " ^ to_maxima b ^ ")" | to_maxima (Powr (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Powr_Nat (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Powr' (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (ExpLn a) = "exp (log (" ^ to_maxima a ^ "))" | to_maxima (LnPowr (a, b)) = "log(" ^ to_maxima a ^ " ^ " ^ to_maxima b ^ ")" | to_maxima (Power (a, b)) = "(" ^ to_maxima a ^ " ^ " ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_maxima a ^ ")" | to_maxima (Root (a, b)) = to_maxima a ^ "^(1/" ^ to_maxima (ConstExpr b) ^ ")" | to_maxima (Uminus a) = "(-" ^ to_maxima a ^ ")" | to_maxima (Inverse a) = "(1/(" ^ to_maxima a ^ "))" | to_maxima (Exp a) = "exp(" ^ to_maxima a ^ ")" | to_maxima (Ln a) = "log(" ^ to_maxima a ^ ")" | to_maxima (Sin a) = "sin(" ^ to_maxima a ^ ")" | to_maxima (Cos a) = "cos(" ^ to_maxima a ^ ")" | to_maxima (ArcTan a) = "atan(" ^ to_maxima a ^ ")" | to_maxima (Absolute a) = "abs(" ^ to_maxima a ^ ")" | to_maxima (Sgn a) = "signum(" ^ to_maxima a ^ ")" | to_maxima (Min (a, b)) = "min(" ^ to_maxima a ^ ", " ^ to_maxima b ^ ")" | to_maxima (Max (a, b)) = "max(" ^ to_maxima a ^ ", " ^ to_maxima b ^ ")" | to_maxima (Floor a) = "floor(" ^ to_maxima a ^ ")" | to_maxima (Ceiling a) = "ceil(" ^ to_maxima a ^ ")" | to_maxima (Frac a) = let val x = to_maxima a in "(" ^ x ^ " - floor(" ^ x ^ "))" end | to_maxima (ConstExpr t) = simple_print_const t | to_maxima X = "x" fun to_sympy (Add (a, b)) = "(" ^ to_sympy a ^ " + " ^ to_sympy b ^ ")" | to_sympy (Minus (a, b)) = "(" ^ to_sympy a ^ " - " ^ to_sympy b ^ ")" | to_sympy (Mult (a, b)) = "(" ^ to_sympy a ^ " * " ^ to_sympy b ^ ")" | to_sympy (Div (a, b)) = "(" ^ to_sympy a ^ " / " ^ to_sympy b ^ ")" | to_sympy (Powr (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Powr_Nat (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Powr' (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (ExpLn a) = "exp (log (" ^ to_sympy a ^ "))" | to_sympy (LnPowr (a, b)) = "log(" ^ to_sympy a ^ " ** " ^ to_sympy b ^ ")" | to_sympy (Power (a, b)) = "(" ^ to_sympy a ^ " ** " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_sympy a ^ ")" | to_sympy (Root (a, b)) = "root(" ^ to_sympy a ^ ", " ^ to_sympy (ConstExpr b) ^ ")" | to_sympy (Uminus a) = "(-" ^ to_sympy a ^ ")" | to_sympy (Inverse a) = "(1/(" ^ to_sympy a ^ "))" | to_sympy (Exp a) = "exp(" ^ to_sympy a ^ ")" | to_sympy (Ln a) = "log(" ^ to_sympy a ^ ")" | to_sympy (Sin a) = "sin(" ^ to_sympy a ^ ")" | to_sympy (Cos a) = "cos(" ^ to_sympy a ^ ")" | to_sympy (ArcTan a) = "atan(" ^ to_sympy a ^ ")" | to_sympy (Absolute a) = "abs(" ^ to_sympy a ^ ")" | to_sympy (Sgn a) = "sign(" ^ to_sympy a ^ ")" | to_sympy (Min (a, b)) = "min(" ^ to_sympy a ^ ", " ^ to_sympy b ^ ")" | to_sympy (Max (a, b)) = "max(" ^ to_sympy a ^ ", " ^ to_sympy b ^ ")" | to_sympy (Floor a) = "floor(" ^ to_sympy a ^ ")" | to_sympy (Ceiling a) = "ceiling(" ^ to_sympy a ^ ")" | to_sympy (Frac a) = "frac(" ^ to_sympy a ^ ")" | to_sympy (ConstExpr t) = simple_print_const t | to_sympy X = "x" fun to_sage (Add (a, b)) = "(" ^ to_sage a ^ " + " ^ to_sage b ^ ")" | to_sage (Minus (a, b)) = "(" ^ to_sage a ^ " - " ^ to_sage b ^ ")" | to_sage (Mult (a, b)) = "(" ^ to_sage a ^ " * " ^ to_sage b ^ ")" | to_sage (Div (a, b)) = "(" ^ to_sage a ^ " / " ^ to_sage b ^ ")" | to_sage (Powr (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Powr_Nat (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Powr' (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage (ConstExpr b) ^ ")" | to_sage (ExpLn a) = "exp (log (" ^ to_sage a ^ "))" | to_sage (LnPowr (a, b)) = "log(" ^ to_sage a ^ " ^ " ^ to_sage b ^ ")" | to_sage (Power (a, b)) = "(" ^ to_sage a ^ " ^ " ^ to_sage (ConstExpr b) ^ ")" | to_sage (Root (a, \<^term>\2::real\)) = "sqrt(" ^ to_sage a ^ ")" | to_sage (Root (a, b)) = to_sage a ^ "^(1/" ^ to_sage (ConstExpr b) ^ ")" | to_sage (Uminus a) = "(-" ^ to_sage a ^ ")" | to_sage (Inverse a) = "(1/(" ^ to_sage a ^ "))" | to_sage (Exp a) = "exp(" ^ to_sage a ^ ")" | to_sage (Ln a) = "log(" ^ to_sage a ^ ")" | to_sage (Sin a) = "sin(" ^ to_sage a ^ ")" | to_sage (Cos a) = "cos(" ^ to_sage a ^ ")" | to_sage (ArcTan a) = "atan(" ^ to_sage a ^ ")" | to_sage (Absolute a) = "abs(" ^ to_sage a ^ ")" | to_sage (Sgn a) = "sign(" ^ to_sage a ^ ")" | to_sage (Min (a, b)) = "min(" ^ to_sage a ^ ", " ^ to_sage b ^ ")" | to_sage (Max (a, b)) = "max(" ^ to_sage a ^ ", " ^ to_sage b ^ ")" | to_sage (Floor a) = "floor(" ^ to_sage a ^ ")" | to_sage (Ceiling a) = "ceil(" ^ to_sage a ^ ")" | to_sage (Frac a) = "frac(" ^ to_sage a ^ ")" | to_sage (ConstExpr t) = simple_print_const t | to_sage X = "x" fun reify_mathematica ctxt = to_mathematica o fst o reify_simple ctxt fun reify_maple ctxt = to_maple o fst o reify_simple ctxt fun reify_maxima ctxt = to_maxima o fst o reify_simple ctxt fun reify_sympy ctxt = to_sympy o fst o reify_simple ctxt fun reify_sage ctxt = to_sage o fst o reify_simple ctxt fun limit_mathematica s = "Limit[" ^ s ^ ", X -> Infinity]" fun limit_maple s = "limit(" ^ s ^ ", x = infinity);" fun limit_maxima s = "limit(" ^ s ^ ", x, inf);" fun limit_sympy s = "limit(" ^ s ^ ", x, oo)" fun limit_sage s = "limit(" ^ s ^ ", x = Infinity)" end diff --git a/src/HOL/Set.thy b/src/HOL/Set.thy --- a/src/HOL/Set.thy +++ b/src/HOL/Set.thy @@ -1,2044 +1,2042 @@ (* Title: HOL/Set.thy Author: Tobias Nipkow Author: Lawrence C Paulson Author: Markus Wenzel *) section \Set theory for higher-order logic\ theory Set imports Lattices Boolean_Algebras begin subsection \Sets as predicates\ typedecl 'a set axiomatization Collect :: "('a \ bool) \ 'a set" \ \comprehension\ and member :: "'a \ 'a set \ bool" \ \membership\ where mem_Collect_eq [iff, code_unfold]: "member a (Collect P) = P a" and Collect_mem_eq [simp]: "Collect (\x. member x A) = A" notation member ("'(\')") and member ("(_/ \ _)" [51, 51] 50) abbreviation not_member where "not_member x A \ \ (x \ A)" \ \non-membership\ notation not_member ("'(\')") and not_member ("(_/ \ _)" [51, 51] 50) notation (ASCII) member ("'(:')") and member ("(_/ : _)" [51, 51] 50) and not_member ("'(~:')") and not_member ("(_/ ~: _)" [51, 51] 50) text \Set comprehensions\ syntax "_Coll" :: "pttrn \ bool \ 'a set" ("(1{_./ _})") translations "{x. P}" \ "CONST Collect (\x. P)" syntax (ASCII) "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/: _)./ _})") syntax "_Collect" :: "pttrn \ 'a set \ bool \ 'a set" ("(1{(_/ \ _)./ _})") translations "{p:A. P}" \ "CONST Collect (\p. p \ A \ P)" lemma CollectI: "P a \ a \ {x. P x}" by simp lemma CollectD: "a \ {x. P x} \ P a" by simp lemma Collect_cong: "(\x. P x = Q x) \ {x. P x} = {x. Q x}" by simp text \ Simproc for pulling \x = t\ in \{x. \ \ x = t \ \}\ to the front (and similarly for \t = x\): \ simproc_setup defined_Collect ("{x. P x \ Q x}") = \ - fn _ => Quantifier1.rearrange_Collect + K (Quantifier1.rearrange_Collect (fn ctxt => resolve_tac ctxt @{thms Collect_cong} 1 THEN resolve_tac ctxt @{thms iffI} 1 THEN ALLGOALS (EVERY' [REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}, - DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})])) + DEPTH_SOLVE_1 o (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms conjI})]))) \ lemmas CollectE = CollectD [elim_format] lemma set_eqI: assumes "\x. x \ A \ x \ B" shows "A = B" proof - from assms have "{x. x \ A} = {x. x \ B}" by simp then show ?thesis by simp qed lemma set_eq_iff: "A = B \ (\x. x \ A \ x \ B)" by (auto intro:set_eqI) lemma Collect_eqI: assumes "\x. P x = Q x" shows "Collect P = Collect Q" using assms by (auto intro: set_eqI) text \Lifting of predicate class instances\ instantiation set :: (type) boolean_algebra begin definition less_eq_set where "A \ B \ (\x. member x A) \ (\x. member x B)" definition less_set where "A < B \ (\x. member x A) < (\x. member x B)" definition inf_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition sup_set where "A \ B = Collect ((\x. member x A) \ (\x. member x B))" definition bot_set where "\ = Collect \" definition top_set where "\ = Collect \" definition uminus_set where "- A = Collect (- (\x. member x A))" definition minus_set where "A - B = Collect ((\x. member x A) - (\x. member x B))" instance by standard (simp_all add: less_eq_set_def less_set_def inf_set_def sup_set_def bot_set_def top_set_def uminus_set_def minus_set_def less_le_not_le sup_inf_distrib1 diff_eq set_eqI fun_eq_iff del: inf_apply sup_apply bot_apply top_apply minus_apply uminus_apply) end text \Set enumerations\ abbreviation empty :: "'a set" ("{}") where "{} \ bot" definition insert :: "'a \ 'a set \ 'a set" where insert_compr: "insert a B = {x. x = a \ x \ B}" syntax "_Finset" :: "args \ 'a set" ("{(_)}") translations "{x, xs}" \ "CONST insert x {xs}" "{x}" \ "CONST insert x {}" subsection \Subsets and bounded quantifiers\ abbreviation subset :: "'a set \ 'a set \ bool" where "subset \ less" abbreviation subset_eq :: "'a set \ 'a set \ bool" where "subset_eq \ less_eq" notation subset ("'(\')") and subset ("(_/ \ _)" [51, 51] 50) and subset_eq ("'(\')") and subset_eq ("(_/ \ _)" [51, 51] 50) abbreviation (input) supset :: "'a set \ 'a set \ bool" where "supset \ greater" abbreviation (input) supset_eq :: "'a set \ 'a set \ bool" where "supset_eq \ greater_eq" notation supset ("'(\')") and supset ("(_/ \ _)" [51, 51] 50) and supset_eq ("'(\')") and supset_eq ("(_/ \ _)" [51, 51] 50) notation (ASCII output) subset ("'(<')") and subset ("(_/ < _)" [51, 51] 50) and subset_eq ("'(<=')") and subset_eq ("(_/ <= _)" [51, 51] 50) definition Ball :: "'a set \ ('a \ bool) \ bool" where "Ball A P \ (\x. x \ A \ P x)" \ \bounded universal quantifiers\ definition Bex :: "'a set \ ('a \ bool) \ bool" where "Bex A P \ (\x. x \ A \ P x)" \ \bounded existential quantifiers\ syntax (ASCII) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3ALL (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3EX (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3EX! (_/:_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST (_/:_)./ _)" [0, 0, 10] 10) syntax (input) "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3! (_/:_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3? (_/:_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3?! (_/:_)./ _)" [0, 0, 10] 10) syntax "_Ball" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn \ 'a set \ bool \ bool" ("(3\(_/\_)./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn \ 'a set \ bool \ bool" ("(3\!(_/\_)./ _)" [0, 0, 10] 10) "_Bleast" :: "id \ 'a set \ bool \ 'a" ("(3LEAST(_/\_)./ _)" [0, 0, 10] 10) translations "\x\A. P" \ "CONST Ball A (\x. P)" "\x\A. P" \ "CONST Bex A (\x. P)" "\!x\A. P" \ "\!x. x \ A \ P" "LEAST x:A. P" \ "LEAST x. x \ A \ P" syntax (ASCII output) "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3ALL _<=_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3EX _<=_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3EX! _<=_./ _)" [0, 0, 10] 10) syntax "_setlessAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setlessEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleAll" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx" :: "[idt, 'a, bool] \ bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_setleEx1" :: "[idt, 'a, bool] \ bool" ("(3\!_\_./ _)" [0, 0, 10] 10) translations "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\A\B. P" \ "\A. A \ B \ P" "\!A\B. P" \ "\!A. A \ B \ P" print_translation \ let val All_binder = Mixfix.binder_name \<^const_syntax>\All\; val Ex_binder = Mixfix.binder_name \<^const_syntax>\Ex\; val impl = \<^const_syntax>\HOL.implies\; val conj = \<^const_syntax>\HOL.conj\; val sbset = \<^const_syntax>\subset\; val sbset_eq = \<^const_syntax>\subset_eq\; val trans = [((All_binder, impl, sbset), \<^syntax_const>\_setlessAll\), ((All_binder, impl, sbset_eq), \<^syntax_const>\_setleAll\), ((Ex_binder, conj, sbset), \<^syntax_const>\_setlessEx\), ((Ex_binder, conj, sbset_eq), \<^syntax_const>\_setleEx\)]; fun mk v (v', T) c n P = if v = v' andalso not (Term.exists_subterm (fn Free (x, _) => x = v | _ => false) n) then Syntax.const c $ Syntax_Trans.mark_bound_body (v', T) $ n $ P else raise Match; fun tr' q = (q, fn _ => (fn [Const (\<^syntax_const>\_bound\, _) $ Free (v, Type (\<^type_name>\set\, _)), Const (c, _) $ (Const (d, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (v', T)) $ n) $ P] => (case AList.lookup (=) trans (q, c, d) of NONE => raise Match | SOME l => mk v (v', T) l n P) | _ => raise Match)); in [tr' All_binder, tr' Ex_binder] end \ text \ \<^medskip> Translate between \{e | x1\xn. P}\ and \{u. \x1\xn. u = e \ P}\; \{y. \x1\xn. y = e \ P}\ is only translated if \[0..n] \ bvs e\. \ syntax "_Setcompr" :: "'a \ idts \ bool \ 'a set" ("(1{_ |/_./ _})") parse_translation \ let val ex_tr = snd (Syntax_Trans.mk_binder_tr ("EX ", \<^const_syntax>\Ex\)); fun nvars (Const (\<^syntax_const>\_idts\, _) $ _ $ idts) = nvars idts + 1 | nvars _ = 1; fun setcompr_tr ctxt [e, idts, b] = let val eq = Syntax.const \<^const_syntax>\HOL.eq\ $ Bound (nvars idts) $ e; val P = Syntax.const \<^const_syntax>\HOL.conj\ $ eq $ b; val exP = ex_tr ctxt [idts, P]; in Syntax.const \<^const_syntax>\Collect\ $ absdummy dummyT exP end; in [(\<^syntax_const>\_Setcompr\, setcompr_tr)] end \ print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Ball\ \<^syntax_const>\_Ball\, Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Bex\ \<^syntax_const>\_Bex\] \ \ \to avoid eta-contraction of body\ print_translation \ let val ex_tr' = snd (Syntax_Trans.mk_binder_tr' (\<^const_syntax>\Ex\, "DUMMY")); fun setcompr_tr' ctxt [Abs (abs as (_, _, P))] = let fun check (Const (\<^const_syntax>\Ex\, _) $ Abs (_, _, P), n) = check (P, n + 1) | check (Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\HOL.eq\, _) $ Bound m $ e) $ P, n) = n > 0 andalso m = n andalso not (loose_bvar1 (P, n)) andalso subset (=) (0 upto (n - 1), add_loose_bnos (e, 0, [])) | check _ = false; fun tr' (_ $ abs) = let val _ $ idts $ (_ $ (_ $ _ $ e) $ Q) = ex_tr' ctxt [abs] in Syntax.const \<^syntax_const>\_Setcompr\ $ e $ idts $ Q end; in if check (P, 0) then tr' P else let val (x as _ $ Free(xN, _), t) = Syntax_Trans.atomic_abs_tr' abs; val M = Syntax.const \<^syntax_const>\_Coll\ $ x $ t; in case t of Const (\<^const_syntax>\HOL.conj\, _) $ (Const (\<^const_syntax>\Set.member\, _) $ (Const (\<^syntax_const>\_bound\, _) $ Free (yN, _)) $ A) $ P => if xN = yN then Syntax.const \<^syntax_const>\_Collect\ $ x $ A $ P else M | _ => M end end; in [(\<^const_syntax>\Collect\, setcompr_tr')] end \ simproc_setup defined_Bex ("\x\A. P x \ Q x") = \ - fn _ => Quantifier1.rearrange_Bex - (fn ctxt => unfold_tac ctxt @{thms Bex_def}) + K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def})) \ simproc_setup defined_All ("\x\A. P x \ Q x") = \ - fn _ => Quantifier1.rearrange_Ball - (fn ctxt => unfold_tac ctxt @{thms Ball_def}) + K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def})) \ lemma ballI [intro!]: "(\x. x \ A \ P x) \ \x\A. P x" by (simp add: Ball_def) lemmas strip = impI allI ballI lemma bspec [dest?]: "\x\A. P x \ x \ A \ P x" by (simp add: Ball_def) text \Gives better instantiation for bound:\ setup \ map_theory_claset (fn ctxt => ctxt addbefore ("bspec", fn ctxt' => dresolve_tac ctxt' @{thms bspec} THEN' assume_tac ctxt')) \ ML \ structure Simpdata = struct open Simpdata; val mksimps_pairs = [(\<^const_name>\Ball\, @{thms bspec})] @ mksimps_pairs; end; open Simpdata; \ declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (mksimps mksimps_pairs))\ lemma ballE [elim]: "\x\A. P x \ (P x \ Q) \ (x \ A \ Q) \ Q" unfolding Ball_def by blast lemma bexI [intro]: "P x \ x \ A \ \x\A. P x" \ \Normally the best argument order: \P x\ constrains the choice of \x \ A\.\ unfolding Bex_def by blast lemma rev_bexI [intro?]: "x \ A \ P x \ \x\A. P x" \ \The best argument order when there is only one \x \ A\.\ unfolding Bex_def by blast lemma bexCI: "(\x\A. \ P x \ P a) \ a \ A \ \x\A. P x" unfolding Bex_def by blast lemma bexE [elim!]: "\x\A. P x \ (\x. x \ A \ P x \ Q) \ Q" unfolding Bex_def by blast lemma ball_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \trivial rewrite rule.\ by (simp add: Ball_def) lemma bex_triv [simp]: "(\x\A. P) \ ((\x. x \ A) \ P)" \ \Dual form for existentials.\ by (simp add: Bex_def) lemma bex_triv_one_point1 [simp]: "(\x\A. x = a) \ a \ A" by blast lemma bex_triv_one_point2 [simp]: "(\x\A. a = x) \ a \ A" by blast lemma bex_one_point1 [simp]: "(\x\A. x = a \ P x) \ a \ A \ P a" by blast lemma bex_one_point2 [simp]: "(\x\A. a = x \ P x) \ a \ A \ P a" by blast lemma ball_one_point1 [simp]: "(\x\A. x = a \ P x) \ (a \ A \ P a)" by blast lemma ball_one_point2 [simp]: "(\x\A. a = x \ P x) \ (a \ A \ P a)" by blast lemma ball_conj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast lemma bex_disj_distrib: "(\x\A. P x \ Q x) \ (\x\A. P x) \ (\x\A. Q x)" by blast text \Congruence rules\ lemma ball_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Ball_def) lemma ball_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Ball_def) lemma bex_cong: "\ A = B; \x. x \ B \ P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: Bex_def cong: conj_cong) lemma bex_cong_simp [cong]: "\ A = B; \x. x \ B =simp=> P x \ Q x \ \ (\x\A. P x) \ (\x\B. Q x)" by (simp add: simp_implies_def Bex_def cong: conj_cong) lemma bex1_def: "(\!x\X. P x) \ (\x\X. P x) \ (\x\X. \y\X. P x \ P y \ x = y)" by auto subsection \Basic operations\ subsubsection \Subsets\ lemma subsetI [intro!]: "(\x. x \ A \ x \ B) \ A \ B" by (simp add: less_eq_set_def le_fun_def) text \ \<^medskip> Map the type \'a set \ anything\ to just \'a\; for overloading constants whose first argument has type \'a set\. \ lemma subsetD [elim, intro?]: "A \ B \ c \ A \ c \ B" by (simp add: less_eq_set_def le_fun_def) \ \Rule in Modus Ponens style.\ lemma rev_subsetD [intro?,no_atp]: "c \ A \ A \ B \ c \ B" \ \The same, with reversed premises for use with @{method erule} -- cf. @{thm rev_mp}.\ by (rule subsetD) lemma subsetCE [elim,no_atp]: "A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" \ \Classical elimination rule.\ by (auto simp add: less_eq_set_def le_fun_def) lemma subset_eq: "A \ B \ (\x\A. x \ B)" by blast lemma contra_subsetD [no_atp]: "A \ B \ c \ B \ c \ A" by blast lemma subset_refl: "A \ A" by (fact order_refl) (* already [iff] *) lemma subset_trans: "A \ B \ B \ C \ A \ C" by (fact order_trans) lemma subset_not_subset_eq [code]: "A \ B \ A \ B \ \ B \ A" by (fact less_le_not_le) lemma eq_mem_trans: "a = b \ b \ A \ a \ A" by simp lemmas basic_trans_rules [trans] = order_trans_rules rev_subsetD subsetD eq_mem_trans subsubsection \Equality\ lemma subset_antisym [intro!]: "A \ B \ B \ A \ A = B" \ \Anti-symmetry of the subset relation.\ by (iprover intro: set_eqI subsetD) text \\<^medskip> Equality rules from ZF set theory -- are they appropriate here?\ lemma equalityD1: "A = B \ A \ B" by simp lemma equalityD2: "A = B \ B \ A" by simp text \ \<^medskip> Be careful when adding this to the claset as \subset_empty\ is in the simpset: \<^prop>\A = {}\ goes to \<^prop>\{} \ A\ and \<^prop>\A \ {}\ and then back to \<^prop>\A = {}\! \ lemma equalityE: "A = B \ (A \ B \ B \ A \ P) \ P" by simp lemma equalityCE [elim]: "A = B \ (c \ A \ c \ B \ P) \ (c \ A \ c \ B \ P) \ P" by blast lemma eqset_imp_iff: "A = B \ x \ A \ x \ B" by simp lemma eqelem_imp_iff: "x = y \ x \ A \ y \ A" by simp subsubsection \The empty set\ lemma empty_def: "{} = {x. False}" by (simp add: bot_set_def bot_fun_def) lemma empty_iff [simp]: "c \ {} \ False" by (simp add: empty_def) lemma emptyE [elim!]: "a \ {} \ P" by simp lemma empty_subsetI [iff]: "{} \ A" \ \One effect is to delete the ASSUMPTION \<^prop>\{} \ A\\ by blast lemma equals0I: "(\y. y \ A \ False) \ A = {}" by blast lemma equals0D: "A = {} \ a \ A" \ \Use for reasoning about disjointness: \A \ B = {}\\ by blast lemma ball_empty [simp]: "Ball {} P \ True" by (simp add: Ball_def) lemma bex_empty [simp]: "Bex {} P \ False" by (simp add: Bex_def) subsubsection \The universal set -- UNIV\ abbreviation UNIV :: "'a set" where "UNIV \ top" lemma UNIV_def: "UNIV = {x. True}" by (simp add: top_set_def top_fun_def) lemma UNIV_I [simp]: "x \ UNIV" by (simp add: UNIV_def) declare UNIV_I [intro] \ \unsafe makes it less likely to cause problems\ lemma UNIV_witness [intro?]: "\x. x \ UNIV" by simp lemma subset_UNIV: "A \ UNIV" by (fact top_greatest) (* already simp *) text \ \<^medskip> Eta-contracting these two rules (to remove \P\) causes them to be ignored because of their interaction with congruence rules. \ lemma ball_UNIV [simp]: "Ball UNIV P \ All P" by (simp add: Ball_def) lemma bex_UNIV [simp]: "Bex UNIV P \ Ex P" by (simp add: Bex_def) lemma UNIV_eq_I: "(\x. x \ A) \ UNIV = A" by auto lemma UNIV_not_empty [iff]: "UNIV \ {}" by (blast elim: equalityE) lemma empty_not_UNIV[simp]: "{} \ UNIV" by blast subsubsection \The Powerset operator -- Pow\ definition Pow :: "'a set \ 'a set set" where Pow_def: "Pow A = {B. B \ A}" lemma Pow_iff [iff]: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma PowI: "A \ B \ A \ Pow B" by (simp add: Pow_def) lemma PowD: "A \ Pow B \ A \ B" by (simp add: Pow_def) lemma Pow_bottom: "{} \ Pow B" by simp lemma Pow_top: "A \ Pow A" by simp lemma Pow_not_empty: "Pow A \ {}" using Pow_top by blast subsubsection \Set complement\ lemma Compl_iff [simp]: "c \ - A \ c \ A" by (simp add: fun_Compl_def uminus_set_def) lemma ComplI [intro!]: "(c \ A \ False) \ c \ - A" by (simp add: fun_Compl_def uminus_set_def) blast text \ \<^medskip> This form, with negated conclusion, works well with the Classical prover. Negated assumptions behave like formulae on the right side of the notional turnstile \dots \ lemma ComplD [dest!]: "c \ - A \ c \ A" by simp lemmas ComplE = ComplD [elim_format] lemma Compl_eq: "- A = {x. \ x \ A}" by blast subsubsection \Binary intersection\ abbreviation inter :: "'a set \ 'a set \ 'a set" (infixl "\" 70) where "(\) \ inf" notation (ASCII) inter (infixl "Int" 70) lemma Int_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: inf_set_def inf_fun_def) lemma Int_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Int_def by blast lemma IntI [intro!]: "c \ A \ c \ B \ c \ A \ B" by simp lemma IntD1: "c \ A \ B \ c \ A" by simp lemma IntD2: "c \ A \ B \ c \ B" by simp lemma IntE [elim!]: "c \ A \ B \ (c \ A \ c \ B \ P) \ P" by simp subsubsection \Binary union\ abbreviation union :: "'a set \ 'a set \ 'a set" (infixl "\" 65) where "union \ sup" notation (ASCII) union (infixl "Un" 65) lemma Un_def: "A \ B = {x. x \ A \ x \ B}" by (simp add: sup_set_def sup_fun_def) lemma Un_iff [simp]: "c \ A \ B \ c \ A \ c \ B" unfolding Un_def by blast lemma UnI1 [elim?]: "c \ A \ c \ A \ B" by simp lemma UnI2 [elim?]: "c \ B \ c \ A \ B" by simp text \\<^medskip> Classical introduction rule: no commitment to \A\ vs. \B\.\ lemma UnCI [intro!]: "(c \ B \ c \ A) \ c \ A \ B" by auto lemma UnE [elim!]: "c \ A \ B \ (c \ A \ P) \ (c \ B \ P) \ P" unfolding Un_def by blast lemma insert_def: "insert a B = {x. x = a} \ B" by (simp add: insert_compr Un_def) subsubsection \Set difference\ lemma Diff_iff [simp]: "c \ A - B \ c \ A \ c \ B" by (simp add: minus_set_def fun_diff_def) lemma DiffI [intro!]: "c \ A \ c \ B \ c \ A - B" by simp lemma DiffD1: "c \ A - B \ c \ A" by simp lemma DiffD2: "c \ A - B \ c \ B \ P" by simp lemma DiffE [elim!]: "c \ A - B \ (c \ A \ c \ B \ P) \ P" by simp lemma set_diff_eq: "A - B = {x. x \ A \ x \ B}" by blast lemma Compl_eq_Diff_UNIV: "- A = (UNIV - A)" by blast subsubsection \Augmenting a set -- \<^const>\insert\\ lemma insert_iff [simp]: "a \ insert b A \ a = b \ a \ A" unfolding insert_def by blast lemma insertI1: "a \ insert a B" by simp lemma insertI2: "a \ B \ a \ insert b B" by simp lemma insertE [elim!]: "a \ insert b A \ (a = b \ P) \ (a \ A \ P) \ P" unfolding insert_def by blast lemma insertCI [intro!]: "(a \ B \ a = b) \ a \ insert b B" \ \Classical introduction rule.\ by auto lemma subset_insert_iff: "A \ insert x B \ (if x \ A then A - {x} \ B else A \ B)" by auto lemma set_insert: assumes "x \ A" obtains B where "A = insert x B" and "x \ B" proof show "A = insert x (A - {x})" using assms by blast show "x \ A - {x}" by blast qed lemma insert_ident: "x \ A \ x \ B \ insert x A = insert x B \ A = B" by auto lemma insert_eq_iff: assumes "a \ A" "b \ B" shows "insert a A = insert b B \ (if a = b then A = B else \C. A = insert b C \ b \ C \ B = insert a C \ a \ C)" (is "?L \ ?R") proof show ?R if ?L proof (cases "a = b") case True with assms \?L\ show ?R by (simp add: insert_ident) next case False let ?C = "A - {b}" have "A = insert b ?C \ b \ ?C \ B = insert a ?C \ a \ ?C" using assms \?L\ \a \ b\ by auto then show ?R using \a \ b\ by auto qed show ?L if ?R using that by (auto split: if_splits) qed lemma insert_UNIV: "insert x UNIV = UNIV" by auto subsubsection \Singletons, using insert\ lemma singletonI [intro!]: "a \ {a}" \ \Redundant? But unlike \insertCI\, it proves the subgoal immediately!\ by (rule insertI1) lemma singletonD [dest!]: "b \ {a} \ b = a" by blast lemmas singletonE = singletonD [elim_format] lemma singleton_iff: "b \ {a} \ b = a" by blast lemma singleton_inject [dest!]: "{a} = {b} \ a = b" by blast lemma singleton_insert_inj_eq [iff]: "{b} = insert a A \ a = b \ A \ {b}" by blast lemma singleton_insert_inj_eq' [iff]: "insert a A = {b} \ a = b \ A \ {b}" by blast lemma subset_singletonD: "A \ {x} \ A = {} \ A = {x}" by fast lemma subset_singleton_iff: "X \ {a} \ X = {} \ X = {a}" by blast lemma subset_singleton_iff_Uniq: "(\a. A \ {a}) \ (\\<^sub>\\<^sub>1x. x \ A)" unfolding Uniq_def by blast lemma singleton_conv [simp]: "{x. x = a} = {a}" by blast lemma singleton_conv2 [simp]: "{x. a = x} = {a}" by blast lemma Diff_single_insert: "A - {x} \ B \ A \ insert x B" by blast lemma subset_Diff_insert: "A \ B - insert x C \ A \ B - C \ x \ A" by blast lemma doubleton_eq_iff: "{a, b} = {c, d} \ a = c \ b = d \ a = d \ b = c" by (blast elim: equalityE) lemma Un_singleton_iff: "A \ B = {x} \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto lemma singleton_Un_iff: "{x} = A \ B \ A = {} \ B = {x} \ A = {x} \ B = {} \ A = {x} \ B = {x}" by auto subsubsection \Image of a set under a function\ text \Frequently \b\ does not have the syntactic form of \f x\.\ definition image :: "('a \ 'b) \ 'a set \ 'b set" (infixr "`" 90) where "f ` A = {y. \x\A. y = f x}" lemma image_eqI [simp, intro]: "b = f x \ x \ A \ b \ f ` A" unfolding image_def by blast lemma imageI: "x \ A \ f x \ f ` A" by (rule image_eqI) (rule refl) lemma rev_image_eqI: "x \ A \ b = f x \ b \ f ` A" \ \This version's more effective when we already have the required \x\.\ by (rule image_eqI) lemma imageE [elim!]: assumes "b \ (\x. f x) ` A" \ \The eta-expansion gives variable-name preservation.\ obtains x where "b = f x" and "x \ A" using assms unfolding image_def by blast lemma Compr_image_eq: "{x \ f ` A. P x} = f ` {x \ A. P (f x)}" by auto lemma image_Un: "f ` (A \ B) = f ` A \ f ` B" by blast lemma image_iff: "z \ f ` A \ (\x\A. z = f x)" by blast lemma image_subsetI: "(\x. x \ A \ f x \ B) \ f ` A \ B" \ \Replaces the three steps \subsetI\, \imageE\, \hypsubst\, but breaks too many existing proofs.\ by blast lemma image_subset_iff: "f ` A \ B \ (\x\A. f x \ B)" \ \This rewrite rule would confuse users if made default.\ by blast lemma subset_imageE: assumes "B \ f ` A" obtains C where "C \ A" and "B = f ` C" proof - from assms have "B = f ` {a \ A. f a \ B}" by fast moreover have "{a \ A. f a \ B} \ A" by blast ultimately show thesis by (blast intro: that) qed lemma subset_image_iff: "B \ f ` A \ (\AA\A. B = f ` AA)" by (blast elim: subset_imageE) lemma image_ident [simp]: "(\x. x) ` Y = Y" by blast lemma image_empty [simp]: "f ` {} = {}" by blast lemma image_insert [simp]: "f ` insert a B = insert (f a) (f ` B)" by blast lemma image_constant: "x \ A \ (\x. c) ` A = {c}" by auto lemma image_constant_conv: "(\x. c) ` A = (if A = {} then {} else {c})" by auto lemma image_image: "f ` (g ` A) = (\x. f (g x)) ` A" by blast lemma insert_image [simp]: "x \ A \ insert (f x) (f ` A) = f ` A" by blast lemma image_is_empty [iff]: "f ` A = {} \ A = {}" by blast lemma empty_is_image [iff]: "{} = f ` A \ A = {}" by blast lemma image_Collect: "f ` {x. P x} = {f x | x. P x}" \ \NOT suitable as a default simp rule: the RHS isn't simpler than the LHS, with its implicit quantifier and conjunction. Also image enjoys better equational properties than does the RHS.\ by blast lemma if_image_distrib [simp]: "(\x. if P x then f x else g x) ` S = f ` (S \ {x. P x}) \ g ` (S \ {x. \ P x})" by auto lemma image_cong: "f ` M = g ` N" if "M = N" "\x. x \ N \ f x = g x" using that by (simp add: image_def) lemma image_cong_simp [cong]: "f ` M = g ` N" if "M = N" "\x. x \ N =simp=> f x = g x" using that image_cong [of M N f g] by (simp add: simp_implies_def) lemma image_Int_subset: "f ` (A \ B) \ f ` A \ f ` B" by blast lemma image_diff_subset: "f ` A - f ` B \ f ` (A - B)" by blast lemma Setcompr_eq_image: "{f x |x. x \ A} = f ` A" by blast lemma setcompr_eq_image: "{f x |x. P x} = f ` {x. P x}" by auto lemma ball_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by simp lemma bex_imageD: "\x\f ` A. P x \ \x\A. P (f x)" by auto lemma image_add_0 [simp]: "(+) (0::'a::comm_monoid_add) ` S = S" by auto theorem Cantors_theorem: "\f. f ` A = Pow A" proof assume "\f. f ` A = Pow A" then obtain f where f: "f ` A = Pow A" .. let ?X = "{a \ A. a \ f a}" have "?X \ Pow A" by blast then have "?X \ f ` A" by (simp only: f) then obtain x where "x \ A" and "f x = ?X" by blast then show False by blast qed text \\<^medskip> Range of a function -- just an abbreviation for image!\ abbreviation range :: "('a \ 'b) \ 'b set" \ \of function\ where "range f \ f ` UNIV" lemma range_eqI: "b = f x \ b \ range f" by simp lemma rangeI: "f x \ range f" by simp lemma rangeE [elim?]: "b \ range (\x. f x) \ (\x. b = f x \ P) \ P" by (rule imageE) lemma range_subsetD: "range f \ B \ f i \ B" by blast lemma full_SetCompr_eq: "{u. \x. u = f x} = range f" by auto lemma range_composition: "range (\x. f (g x)) = f ` range g" by auto lemma range_constant [simp]: "range (\_. x) = {x}" by (simp add: image_constant) lemma range_eq_singletonD: "range f = {a} \ f x = a" by auto subsubsection \Some rules with \if\\ text \Elimination of \{x. \ \ x = t \ \}\.\ lemma Collect_conv_if: "{x. x = a \ P x} = (if P a then {a} else {})" by auto lemma Collect_conv_if2: "{x. a = x \ P x} = (if P a then {a} else {})" by auto text \ Rewrite rules for boolean case-splitting: faster than \if_split [split]\. \ lemma if_split_eq1: "(if Q then x else y) = b \ (Q \ x = b) \ (\ Q \ y = b)" by (rule if_split) lemma if_split_eq2: "a = (if Q then x else y) \ (Q \ a = x) \ (\ Q \ a = y)" by (rule if_split) text \ Split ifs on either side of the membership relation. Not for \[simp]\ -- can cause goals to blow up! \ lemma if_split_mem1: "(if Q then x else y) \ b \ (Q \ x \ b) \ (\ Q \ y \ b)" by (rule if_split) lemma if_split_mem2: "(a \ (if Q then x else y)) \ (Q \ a \ x) \ (\ Q \ a \ y)" by (rule if_split [where P = "\S. a \ S"]) lemmas split_ifs = if_bool_eq_conj if_split_eq1 if_split_eq2 if_split_mem1 if_split_mem2 (*Would like to add these, but the existing code only searches for the outer-level constant, which in this case is just Set.member; we instead need to use term-nets to associate patterns with rules. Also, if a rule fails to apply, then the formula should be kept. [("uminus", Compl_iff RS iffD1), ("minus", [Diff_iff RS iffD1]), ("Int", [IntD1,IntD2]), ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] *) subsection \Further operations and lemmas\ subsubsection \The ``proper subset'' relation\ lemma psubsetI [intro!]: "A \ B \ A \ B \ A \ B" unfolding less_le by blast lemma psubsetE [elim!]: "A \ B \ (A \ B \ \ B \ A \ R) \ R" unfolding less_le by blast lemma psubset_insert_iff: "A \ insert x B \ (if x \ B then A \ B else if x \ A then A - {x} \ B else A \ B)" by (auto simp add: less_le subset_insert_iff) lemma psubset_eq: "A \ B \ A \ B \ A \ B" by (simp only: less_le) lemma psubset_imp_subset: "A \ B \ A \ B" by (simp add: psubset_eq) lemma psubset_trans: "A \ B \ B \ C \ A \ C" unfolding less_le by (auto dest: subset_antisym) lemma psubsetD: "A \ B \ c \ A \ c \ B" unfolding less_le by (auto dest: subsetD) lemma psubset_subset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma subset_psubset_trans: "A \ B \ B \ C \ A \ C" by (auto simp add: psubset_eq) lemma psubset_imp_ex_mem: "A \ B \ \b. b \ B - A" unfolding less_le by blast lemma atomize_ball: "(\x. x \ A \ P x) \ Trueprop (\x\A. P x)" by (simp only: Ball_def atomize_all atomize_imp) lemmas [symmetric, rulify] = atomize_ball and [symmetric, defn] = atomize_ball lemma image_Pow_mono: "f ` A \ B \ image f ` Pow A \ Pow B" by blast lemma image_Pow_surj: "f ` A = B \ image f ` Pow A = Pow B" by (blast elim: subset_imageE) subsubsection \Derived rules involving subsets.\ text \\insert\.\ lemma subset_insertI: "B \ insert a B" by (rule subsetI) (erule insertI2) lemma subset_insertI2: "A \ B \ A \ insert b B" by blast lemma subset_insert: "x \ A \ A \ insert x B \ A \ B" by blast text \\<^medskip> Finite Union -- the least upper bound of two sets.\ lemma Un_upper1: "A \ A \ B" by (fact sup_ge1) lemma Un_upper2: "B \ A \ B" by (fact sup_ge2) lemma Un_least: "A \ C \ B \ C \ A \ B \ C" by (fact sup_least) text \\<^medskip> Finite Intersection -- the greatest lower bound of two sets.\ lemma Int_lower1: "A \ B \ A" by (fact inf_le1) lemma Int_lower2: "A \ B \ B" by (fact inf_le2) lemma Int_greatest: "C \ A \ C \ B \ C \ A \ B" by (fact inf_greatest) text \\<^medskip> Set difference.\ lemma Diff_subset[simp]: "A - B \ A" by blast lemma Diff_subset_conv: "A - B \ C \ A \ B \ C" by blast subsubsection \Equalities involving union, intersection, inclusion, etc.\ text \\{}\.\ lemma Collect_const [simp]: "{s. P} = (if P then UNIV else {})" \ \supersedes \Collect_False_empty\\ by auto lemma subset_empty [simp]: "A \ {} \ A = {}" by (fact bot_unique) lemma not_psubset_empty [iff]: "\ (A < {})" by (fact not_less_bot) (* FIXME: already simp *) lemma Collect_subset [simp]: "{x\A. P x} \ A" by auto lemma Collect_empty_eq [simp]: "Collect P = {} \ (\x. \ P x)" by blast lemma empty_Collect_eq [simp]: "{} = Collect P \ (\x. \ P x)" by blast lemma Collect_neg_eq: "{x. \ P x} = - {x. P x}" by blast lemma Collect_disj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_imp_eq: "{x. P x \ Q x} = - {x. P x} \ {x. Q x}" by blast lemma Collect_conj_eq: "{x. P x \ Q x} = {x. P x} \ {x. Q x}" by blast lemma Collect_mono_iff: "Collect P \ Collect Q \ (\x. P x \ Q x)" by blast text \\<^medskip> \insert\.\ lemma insert_is_Un: "insert a A = {a} \ A" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a {}\\ by blast lemma insert_not_empty [simp]: "insert a A \ {}" and empty_not_insert [simp]: "{} \ insert a A" by blast+ lemma insert_absorb: "a \ A \ insert a A = A" \ \\[simp]\ causes recursive calls when there are nested inserts\ \ \with \<^emph>\quadratic\ running time\ by blast lemma insert_absorb2 [simp]: "insert x (insert x A) = insert x A" by blast lemma insert_commute: "insert x (insert y A) = insert y (insert x A)" by blast lemma insert_subset [simp]: "insert x A \ B \ x \ B \ A \ B" by blast lemma mk_disjoint_insert: "a \ A \ \B. A = insert a B \ a \ B" \ \use new \B\ rather than \A - {a}\ to avoid infinite unfolding\ by (rule exI [where x = "A - {a}"]) blast lemma insert_Collect: "insert a (Collect P) = {u. u \ a \ P u}" by auto lemma insert_inter_insert [simp]: "insert a A \ insert a B = insert a (A \ B)" by blast lemma insert_disjoint [simp]: "insert a A \ B = {} \ a \ B \ A \ B = {}" "{} = insert a A \ B \ a \ B \ {} = A \ B" by auto lemma disjoint_insert [simp]: "B \ insert a A = {} \ a \ B \ B \ A = {}" "{} = A \ insert b B \ b \ A \ {} = A \ B" by auto text \\<^medskip> \Int\\ lemma Int_absorb: "A \ A = A" by (fact inf_idem) (* already simp *) lemma Int_left_absorb: "A \ (A \ B) = A \ B" by (fact inf_left_idem) lemma Int_commute: "A \ B = B \ A" by (fact inf_commute) lemma Int_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact inf_left_commute) lemma Int_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact inf_assoc) lemmas Int_ac = Int_assoc Int_left_absorb Int_commute Int_left_commute \ \Intersection is an AC-operator\ lemma Int_absorb1: "B \ A \ A \ B = B" by (fact inf_absorb2) lemma Int_absorb2: "A \ B \ A \ B = A" by (fact inf_absorb1) lemma Int_empty_left: "{} \ B = {}" by (fact inf_bot_left) (* already simp *) lemma Int_empty_right: "A \ {} = {}" by (fact inf_bot_right) (* already simp *) lemma disjoint_eq_subset_Compl: "A \ B = {} \ A \ - B" by blast lemma disjoint_iff: "A \ B = {} \ (\x. x\A \ x \ B)" by blast lemma disjoint_iff_not_equal: "A \ B = {} \ (\x\A. \y\B. x \ y)" by blast lemma Int_UNIV_left: "UNIV \ B = B" by (fact inf_top_left) (* already simp *) lemma Int_UNIV_right: "A \ UNIV = A" by (fact inf_top_right) (* already simp *) lemma Int_Un_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact inf_sup_distrib1) lemma Int_Un_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact inf_sup_distrib2) lemma Int_UNIV [simp]: "A \ B = UNIV \ A = UNIV \ B = UNIV" by (fact inf_eq_top_iff) (* already simp *) lemma Int_subset_iff [simp]: "C \ A \ B \ C \ A \ C \ B" by (fact le_inf_iff) lemma Int_Collect: "x \ A \ {x. P x} \ x \ A \ P x" by blast text \\<^medskip> \Un\.\ lemma Un_absorb: "A \ A = A" by (fact sup_idem) (* already simp *) lemma Un_left_absorb: "A \ (A \ B) = A \ B" by (fact sup_left_idem) lemma Un_commute: "A \ B = B \ A" by (fact sup_commute) lemma Un_left_commute: "A \ (B \ C) = B \ (A \ C)" by (fact sup_left_commute) lemma Un_assoc: "(A \ B) \ C = A \ (B \ C)" by (fact sup_assoc) lemmas Un_ac = Un_assoc Un_left_absorb Un_commute Un_left_commute \ \Union is an AC-operator\ lemma Un_absorb1: "A \ B \ A \ B = B" by (fact sup_absorb2) lemma Un_absorb2: "B \ A \ A \ B = A" by (fact sup_absorb1) lemma Un_empty_left: "{} \ B = B" by (fact sup_bot_left) (* already simp *) lemma Un_empty_right: "A \ {} = A" by (fact sup_bot_right) (* already simp *) lemma Un_UNIV_left: "UNIV \ B = UNIV" by (fact sup_top_left) (* already simp *) lemma Un_UNIV_right: "A \ UNIV = UNIV" by (fact sup_top_right) (* already simp *) lemma Un_insert_left [simp]: "(insert a B) \ C = insert a (B \ C)" by blast lemma Un_insert_right [simp]: "A \ (insert a B) = insert a (A \ B)" by blast lemma Int_insert_left: "(insert a B) \ C = (if a \ C then insert a (B \ C) else B \ C)" by auto lemma Int_insert_left_if0 [simp]: "a \ C \ (insert a B) \ C = B \ C" by auto lemma Int_insert_left_if1 [simp]: "a \ C \ (insert a B) \ C = insert a (B \ C)" by auto lemma Int_insert_right: "A \ (insert a B) = (if a \ A then insert a (A \ B) else A \ B)" by auto lemma Int_insert_right_if0 [simp]: "a \ A \ A \ (insert a B) = A \ B" by auto lemma Int_insert_right_if1 [simp]: "a \ A \ A \ (insert a B) = insert a (A \ B)" by auto lemma Un_Int_distrib: "A \ (B \ C) = (A \ B) \ (A \ C)" by (fact sup_inf_distrib1) lemma Un_Int_distrib2: "(B \ C) \ A = (B \ A) \ (C \ A)" by (fact sup_inf_distrib2) lemma Un_Int_crazy: "(A \ B) \ (B \ C) \ (C \ A) = (A \ B) \ (B \ C) \ (C \ A)" by blast lemma subset_Un_eq: "A \ B \ A \ B = B" by (fact le_iff_sup) lemma Un_empty [iff]: "A \ B = {} \ A = {} \ B = {}" by (fact sup_eq_bot_iff) (* FIXME: already simp *) lemma Un_subset_iff [simp]: "A \ B \ C \ A \ C \ B \ C" by (fact le_sup_iff) lemma Un_Diff_Int: "(A - B) \ (A \ B) = A" by blast lemma Diff_Int2: "A \ C - B \ C = A \ C - B" by blast lemma subset_UnE: assumes "C \ A \ B" obtains A' B' where "A' \ A" "B' \ B" "C = A' \ B'" proof show "C \ A \ A" "C \ B \ B" "C = (C \ A) \ (C \ B)" using assms by blast+ qed lemma Un_Int_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto lemma Int_Un_eq [simp]: "(S \ T) \ S = S" "(S \ T) \ T = T" "S \ (S \ T) = S" "T \ (S \ T) = T" by auto text \\<^medskip> Set complement\ lemma Compl_disjoint [simp]: "A \ - A = {}" by (fact inf_compl_bot) lemma Compl_disjoint2 [simp]: "- A \ A = {}" by (fact compl_inf_bot) lemma Compl_partition: "A \ - A = UNIV" by (fact sup_compl_top) lemma Compl_partition2: "- A \ A = UNIV" by (fact compl_sup_top) lemma double_complement: "- (-A) = A" for A :: "'a set" by (fact double_compl) (* already simp *) lemma Compl_Un: "- (A \ B) = (- A) \ (- B)" by (fact compl_sup) (* already simp *) lemma Compl_Int: "- (A \ B) = (- A) \ (- B)" by (fact compl_inf) (* already simp *) lemma subset_Compl_self_eq: "A \ - A \ A = {}" by blast lemma Un_Int_assoc_eq: "(A \ B) \ C = A \ (B \ C) \ C \ A" \ \Halmos, Naive Set Theory, page 16.\ by blast lemma Compl_UNIV_eq: "- UNIV = {}" by (fact compl_top_eq) (* already simp *) lemma Compl_empty_eq: "- {} = UNIV" by (fact compl_bot_eq) (* already simp *) lemma Compl_subset_Compl_iff [iff]: "- A \ - B \ B \ A" by (fact compl_le_compl_iff) (* FIXME: already simp *) lemma Compl_eq_Compl_iff [iff]: "- A = - B \ A = B" for A B :: "'a set" by (fact compl_eq_compl_iff) (* FIXME: already simp *) lemma Compl_insert: "- insert x A = (- A) - {x}" by blast text \\<^medskip> Bounded quantifiers. The following are not added to the default simpset because (a) they duplicate the body and (b) there are no similar rules for \Int\. \ lemma ball_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast lemma bex_Un: "(\x \ A \ B. P x) \ (\x\A. P x) \ (\x\B. P x)" by blast text \\<^medskip> Set difference.\ lemma Diff_eq: "A - B = A \ (- B)" by blast lemma Diff_eq_empty_iff [simp]: "A - B = {} \ A \ B" by blast lemma Diff_cancel [simp]: "A - A = {}" by blast lemma Diff_idemp [simp]: "(A - B) - B = A - B" for A B :: "'a set" by blast lemma Diff_triv: "A \ B = {} \ A - B = A" by (blast elim: equalityE) lemma empty_Diff [simp]: "{} - A = {}" by blast lemma Diff_empty [simp]: "A - {} = A" by blast lemma Diff_UNIV [simp]: "A - UNIV = {}" by blast lemma Diff_insert0 [simp]: "x \ A \ A - insert x B = A - B" by blast lemma Diff_insert: "A - insert a B = A - B - {a}" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma Diff_insert2: "A - insert a B = A - {a} - B" \ \NOT SUITABLE FOR REWRITING since \{a} \ insert a 0\\ by blast lemma insert_Diff_if: "insert x A - B = (if x \ B then A - B else insert x (A - B))" by auto lemma insert_Diff1 [simp]: "x \ B \ insert x A - B = A - B" by blast lemma insert_Diff_single[simp]: "insert a (A - {a}) = insert a A" by blast lemma insert_Diff: "a \ A \ insert a (A - {a}) = A" by blast lemma Diff_insert_absorb: "x \ A \ (insert x A) - {x} = A" by auto lemma Diff_disjoint [simp]: "A \ (B - A) = {}" by blast lemma Diff_partition: "A \ B \ A \ (B - A) = B" by blast lemma double_diff: "A \ B \ B \ C \ B - (C - A) = A" by blast lemma Un_Diff_cancel [simp]: "A \ (B - A) = A \ B" by blast lemma Un_Diff_cancel2 [simp]: "(B - A) \ A = B \ A" by blast lemma Diff_Un: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Int: "A - (B \ C) = (A - B) \ (A - C)" by blast lemma Diff_Diff_Int: "A - (A - B) = A \ B" by blast lemma Un_Diff: "(A \ B) - C = (A - C) \ (B - C)" by blast lemma Int_Diff: "(A \ B) - C = A \ (B - C)" by blast lemma Diff_Int_distrib: "C \ (A - B) = (C \ A) - (C \ B)" by blast lemma Diff_Int_distrib2: "(A - B) \ C = (A \ C) - (B \ C)" by blast lemma Diff_Compl [simp]: "A - (- B) = A \ B" by auto lemma Compl_Diff_eq [simp]: "- (A - B) = - A \ B" by blast lemma subset_Compl_singleton [simp]: "A \ - {b} \ b \ A" by blast text \\<^medskip> Quantification over type \<^typ>\bool\.\ lemma bool_induct: "P True \ P False \ P x" by (cases x) auto lemma all_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_induct) lemma bool_contrapos: "P x \ \ P False \ P True" by (cases x) auto lemma ex_bool_eq: "(\b. P b) \ P True \ P False" by (auto intro: bool_contrapos) lemma UNIV_bool: "UNIV = {False, True}" by (auto intro: bool_induct) text \\<^medskip> \Pow\\ lemma Pow_empty [simp]: "Pow {} = {{}}" by (auto simp add: Pow_def) lemma Pow_singleton_iff [simp]: "Pow X = {Y} \ X = {} \ Y = {}" by blast (* somewhat slow *) lemma Pow_insert: "Pow (insert a A) = Pow A \ (insert a ` Pow A)" by (blast intro: image_eqI [where ?x = "u - {a}" for u]) lemma Pow_Compl: "Pow (- A) = {- B | B. A \ Pow B}" by (blast intro: exI [where ?x = "- u" for u]) lemma Pow_UNIV [simp]: "Pow UNIV = UNIV" by blast lemma Un_Pow_subset: "Pow A \ Pow B \ Pow (A \ B)" by blast lemma Pow_Int_eq [simp]: "Pow (A \ B) = Pow A \ Pow B" by blast text \\<^medskip> Miscellany.\ lemma Int_Diff_disjoint: "A \ B \ (A - B) = {}" by blast lemma Int_Diff_Un: "A \ B \ (A - B) = A" by blast lemma set_eq_subset: "A = B \ A \ B \ B \ A" by blast lemma subset_iff: "A \ B \ (\t. t \ A \ t \ B)" by blast lemma subset_iff_psubset_eq: "A \ B \ A \ B \ A = B" unfolding less_le by blast lemma all_not_in_conv [simp]: "(\x. x \ A) \ A = {}" by blast lemma ex_in_conv: "(\x. x \ A) \ A \ {}" by blast lemma ball_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\P. (\x\{}. P x) \ True" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\ (\x\A. P x)) \ (\x\A. \ P x)" by auto lemma bex_simps [simp, no_atp]: "\A P Q. (\x\A. P x \ Q) \ ((\x\A. P x) \ Q)" "\A P Q. (\x\A. P \ Q x) \ (P \ (\x\A. Q x))" "\P. (\x\{}. P x) \ False" "\P. (\x\UNIV. P x) \ (\x. P x)" "\a B P. (\x\insert a B. P x) \ (P a \ (\x\B. P x))" "\P Q. (\x\Collect Q. P x) \ (\x. Q x \ P x)" "\A P f. (\x\f`A. P x) \ (\x\A. P (f x))" "\A P. (\(\x\A. P x)) \ (\x\A. \ P x)" by auto lemma ex_image_cong_iff [simp, no_atp]: "(\x. x\f`A) \ A \ {}" "(\x. x\f`A \ P x) \ (\x\A. P (f x))" by auto subsubsection \Monotonicity of various operations\ lemma image_mono: "A \ B \ f ` A \ f ` B" by blast lemma Pow_mono: "A \ B \ Pow A \ Pow B" by blast lemma insert_mono: "C \ D \ insert a C \ insert a D" by blast lemma Un_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact sup_mono) lemma Int_mono: "A \ C \ B \ D \ A \ B \ C \ D" by (fact inf_mono) lemma Diff_mono: "A \ C \ D \ B \ A - B \ C - D" by blast lemma Compl_anti_mono: "A \ B \ - B \ - A" by (fact compl_mono) text \\<^medskip> Monotonicity of implications.\ lemma in_mono: "A \ B \ x \ A \ x \ B" by (rule impI) (erule subsetD) lemma conj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma disj_mono: "P1 \ Q1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_mono: "Q1 \ P1 \ P2 \ Q2 \ (P1 \ P2) \ (Q1 \ Q2)" by iprover lemma imp_refl: "P \ P" .. lemma not_mono: "Q \ P \ \ P \ \ Q" by iprover lemma ex_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma all_mono: "(\x. P x \ Q x) \ (\x. P x) \ (\x. Q x)" by iprover lemma Collect_mono: "(\x. P x \ Q x) \ Collect P \ Collect Q" by blast lemma Int_Collect_mono: "A \ B \ (\x. x \ A \ P x \ Q x) \ A \ Collect P \ B \ Collect Q" by blast lemmas basic_monos = subset_refl imp_refl disj_mono conj_mono ex_mono Collect_mono in_mono lemma eq_to_mono: "a = b \ c = d \ b \ d \ a \ c" by iprover subsubsection \Inverse image of a function\ definition vimage :: "('a \ 'b) \ 'b set \ 'a set" (infixr "-`" 90) where "f -` B \ {x. f x \ B}" lemma vimage_eq [simp]: "a \ f -` B \ f a \ B" unfolding vimage_def by blast lemma vimage_singleton_eq: "a \ f -` {b} \ f a = b" by simp lemma vimageI [intro]: "f a = b \ b \ B \ a \ f -` B" unfolding vimage_def by blast lemma vimageI2: "f a \ A \ a \ f -` A" unfolding vimage_def by fast lemma vimageE [elim!]: "a \ f -` B \ (\x. f a = x \ x \ B \ P) \ P" unfolding vimage_def by blast lemma vimageD: "a \ f -` A \ f a \ A" unfolding vimage_def by fast lemma vimage_empty [simp]: "f -` {} = {}" by blast lemma vimage_Compl: "f -` (- A) = - (f -` A)" by blast lemma vimage_Un [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by blast lemma vimage_Int [simp]: "f -` (A \ B) = (f -` A) \ (f -` B)" by fast lemma vimage_Collect_eq [simp]: "f -` Collect P = {y. P (f y)}" by blast lemma vimage_Collect: "(\x. P (f x) = Q x) \ f -` (Collect P) = Collect Q" by blast lemma vimage_insert: "f -` (insert a B) = (f -` {a}) \ (f -` B)" \ \NOT suitable for rewriting because of the recurrence of \{a}\.\ by blast lemma vimage_Diff: "f -` (A - B) = (f -` A) - (f -` B)" by blast lemma vimage_UNIV [simp]: "f -` UNIV = UNIV" by blast lemma vimage_mono: "A \ B \ f -` A \ f -` B" \ \monotonicity\ by blast lemma vimage_image_eq: "f -` (f ` A) = {y. \x\A. f x = f y}" by (blast intro: sym) lemma image_vimage_subset: "f ` (f -` A) \ A" by blast lemma image_vimage_eq [simp]: "f ` (f -` A) = A \ range f" by blast lemma image_subset_iff_subset_vimage: "f ` A \ B \ A \ f -` B" by blast lemma subset_vimage_iff: "A \ f -` B \ (\x\A. f x \ B)" by auto lemma vimage_const [simp]: "((\x. c) -` A) = (if c \ A then UNIV else {})" by auto lemma vimage_if [simp]: "((\x. if x \ B then c else d) -` A) = (if c \ A then (if d \ A then UNIV else B) else if d \ A then - B else {})" by (auto simp add: vimage_def) lemma vimage_inter_cong: "(\ w. w \ S \ f w = g w) \ f -` y \ S = g -` y \ S" by auto lemma vimage_ident [simp]: "(\x. x) -` Y = Y" by blast subsubsection \Singleton sets\ definition is_singleton :: "'a set \ bool" where "is_singleton A \ (\x. A = {x})" lemma is_singletonI [simp, intro!]: "is_singleton {x}" unfolding is_singleton_def by simp lemma is_singletonI': "A \ {} \ (\x y. x \ A \ y \ A \ x = y) \ is_singleton A" unfolding is_singleton_def by blast lemma is_singletonE: "is_singleton A \ (\x. A = {x} \ P) \ P" unfolding is_singleton_def by blast subsubsection \Getting the contents of a singleton set\ definition the_elem :: "'a set \ 'a" where "the_elem X = (THE x. X = {x})" lemma the_elem_eq [simp]: "the_elem {x} = x" by (simp add: the_elem_def) lemma is_singleton_the_elem: "is_singleton A \ A = {the_elem A}" by (auto simp: is_singleton_def) lemma the_elem_image_unique: assumes "A \ {}" and *: "\y. y \ A \ f y = f x" shows "the_elem (f ` A) = f x" unfolding the_elem_def proof (rule the1_equality) from \A \ {}\ obtain y where "y \ A" by auto with * have "f x = f y" by simp with \y \ A\ have "f x \ f ` A" by blast with * show "f ` A = {f x}" by auto then show "\!x. f ` A = {x}" by auto qed subsubsection \Monad operation\ definition bind :: "'a set \ ('a \ 'b set) \ 'b set" where "bind A f = {x. \B \ f`A. x \ B}" hide_const (open) bind lemma bind_bind: "Set.bind (Set.bind A B) C = Set.bind A (\x. Set.bind (B x) C)" for A :: "'a set" by (auto simp: bind_def) lemma empty_bind [simp]: "Set.bind {} f = {}" by (simp add: bind_def) lemma nonempty_bind_const: "A \ {} \ Set.bind A (\_. B) = B" by (auto simp: bind_def) lemma bind_const: "Set.bind A (\_. B) = (if A = {} then {} else B)" by (auto simp: bind_def) lemma bind_singleton_conv_image: "Set.bind A (\x. {f x}) = f ` A" by (auto simp: bind_def) subsubsection \Operations for execution\ definition is_empty :: "'a set \ bool" where [code_abbrev]: "is_empty A \ A = {}" hide_const (open) is_empty definition remove :: "'a \ 'a set \ 'a set" where [code_abbrev]: "remove x A = A - {x}" hide_const (open) remove lemma member_remove [simp]: "x \ Set.remove y A \ x \ A \ x \ y" by (simp add: remove_def) definition filter :: "('a \ bool) \ 'a set \ 'a set" where [code_abbrev]: "filter P A = {a \ A. P a}" hide_const (open) filter lemma member_filter [simp]: "x \ Set.filter P A \ x \ A \ P x" by (simp add: filter_def) instantiation set :: (equal) equal begin definition "HOL.equal A B \ A \ B \ B \ A" instance by standard (auto simp add: equal_set_def) end text \Misc\ definition pairwise :: "('a \ 'a \ bool) \ 'a set \ bool" where "pairwise R S \ (\x \ S. \y \ S. x \ y \ R x y)" lemma pairwise_alt: "pairwise R S \ (\x\S. \y\S-{x}. R x y)" by (auto simp add: pairwise_def) lemma pairwise_trivial [simp]: "pairwise (\i j. j \ i) I" by (auto simp: pairwise_def) lemma pairwiseI [intro?]: "pairwise R S" if "\x y. x \ S \ y \ S \ x \ y \ R x y" using that by (simp add: pairwise_def) lemma pairwiseD: "R x y" and "R y x" if "pairwise R S" "x \ S" and "y \ S" and "x \ y" using that by (simp_all add: pairwise_def) lemma pairwise_empty [simp]: "pairwise P {}" by (simp add: pairwise_def) lemma pairwise_singleton [simp]: "pairwise P {A}" by (simp add: pairwise_def) lemma pairwise_insert: "pairwise r (insert x s) \ (\y. y \ s \ y \ x \ r x y \ r y x) \ pairwise r s" by (force simp: pairwise_def) lemma pairwise_subset: "pairwise P S \ T \ S \ pairwise P T" by (force simp: pairwise_def) lemma pairwise_mono: "\pairwise P A; \x y. P x y \ Q x y; B \ A\ \ pairwise Q B" by (fastforce simp: pairwise_def) lemma pairwise_imageI: "pairwise P (f ` A)" if "\x y. x \ A \ y \ A \ x \ y \ f x \ f y \ P (f x) (f y)" using that by (auto intro: pairwiseI) lemma pairwise_image: "pairwise r (f ` s) \ pairwise (\x y. (f x \ f y) \ r (f x) (f y)) s" by (force simp: pairwise_def) definition disjnt :: "'a set \ 'a set \ bool" where "disjnt A B \ A \ B = {}" lemma disjnt_self_iff_empty [simp]: "disjnt S S \ S = {}" by (auto simp: disjnt_def) lemma disjnt_iff: "disjnt A B \ (\x. \ (x \ A \ x \ B))" by (force simp: disjnt_def) lemma disjnt_sym: "disjnt A B \ disjnt B A" using disjnt_iff by blast lemma disjnt_empty1 [simp]: "disjnt {} A" and disjnt_empty2 [simp]: "disjnt A {}" by (auto simp: disjnt_def) lemma disjnt_insert1 [simp]: "disjnt (insert a X) Y \ a \ Y \ disjnt X Y" by (simp add: disjnt_def) lemma disjnt_insert2 [simp]: "disjnt Y (insert a X) \ a \ Y \ disjnt Y X" by (simp add: disjnt_def) lemma disjnt_subset1 : "\disjnt X Y; Z \ X\ \ disjnt Z Y" by (auto simp: disjnt_def) lemma disjnt_subset2 : "\disjnt X Y; Z \ Y\ \ disjnt X Z" by (auto simp: disjnt_def) lemma disjnt_Un1 [simp]: "disjnt (A \ B) C \ disjnt A C \ disjnt B C" by (auto simp: disjnt_def) lemma disjnt_Un2 [simp]: "disjnt C (A \ B) \ disjnt C A \ disjnt C B" by (auto simp: disjnt_def) lemma disjnt_Diff1: "disjnt (X-Y) (U-V)" and disjnt_Diff2: "disjnt (U-V) (X-Y)" if "X \ V" using that by (auto simp: disjnt_def) lemma disjoint_image_subset: "\pairwise disjnt \; \X. X \ \ \ f X \ X\ \ pairwise disjnt (f `\)" unfolding disjnt_def pairwise_def by fast lemma pairwise_disjnt_iff: "pairwise disjnt \ \ (\x. \\<^sub>\\<^sub>1 X. X \ \ \ x \ X)" by (auto simp: Uniq_def disjnt_iff pairwise_def) lemma disjnt_insert: \<^marker>\contributor \Lars Hupel\\ \disjnt (insert x M) N\ if \x \ N\ \disjnt M N\ using that by (simp add: disjnt_def) lemma Int_emptyI: "(\x. x \ A \ x \ B \ False) \ A \ B = {}" by blast lemma in_image_insert_iff: assumes "\C. C \ B \ x \ C" shows "A \ insert x ` B \ x \ A \ A - {x} \ B" (is "?P \ ?Q") proof assume ?P then show ?Q using assms by auto next assume ?Q then have "x \ A" and "A - {x} \ B" by simp_all from \A - {x} \ B\ have "insert x (A - {x}) \ insert x ` B" by (rule imageI) also from \x \ A\ have "insert x (A - {x}) = A" by auto finally show ?P . qed hide_const (open) member not_member lemmas equalityI = subset_antisym lemmas set_mp = subsetD lemmas set_rev_mp = rev_subsetD ML \ val Ball_def = @{thm Ball_def} val Bex_def = @{thm Bex_def} val CollectD = @{thm CollectD} val CollectE = @{thm CollectE} val CollectI = @{thm CollectI} val Collect_conj_eq = @{thm Collect_conj_eq} val Collect_mem_eq = @{thm Collect_mem_eq} val IntD1 = @{thm IntD1} val IntD2 = @{thm IntD2} val IntE = @{thm IntE} val IntI = @{thm IntI} val Int_Collect = @{thm Int_Collect} val UNIV_I = @{thm UNIV_I} val UNIV_witness = @{thm UNIV_witness} val UnE = @{thm UnE} val UnI1 = @{thm UnI1} val UnI2 = @{thm UnI2} val ballE = @{thm ballE} val ballI = @{thm ballI} val bexCI = @{thm bexCI} val bexE = @{thm bexE} val bexI = @{thm bexI} val bex_triv = @{thm bex_triv} val bspec = @{thm bspec} val contra_subsetD = @{thm contra_subsetD} val equalityCE = @{thm equalityCE} val equalityD1 = @{thm equalityD1} val equalityD2 = @{thm equalityD2} val equalityE = @{thm equalityE} val equalityI = @{thm equalityI} val imageE = @{thm imageE} val imageI = @{thm imageI} val image_Un = @{thm image_Un} val image_insert = @{thm image_insert} val insert_commute = @{thm insert_commute} val insert_iff = @{thm insert_iff} val mem_Collect_eq = @{thm mem_Collect_eq} val rangeE = @{thm rangeE} val rangeI = @{thm rangeI} val range_eqI = @{thm range_eqI} val subsetCE = @{thm subsetCE} val subsetD = @{thm subsetD} val subsetI = @{thm subsetI} val subset_refl = @{thm subset_refl} val subset_trans = @{thm subset_trans} val vimageD = @{thm vimageD} val vimageE = @{thm vimageE} val vimageI = @{thm vimageI} val vimageI2 = @{thm vimageI2} val vimage_Collect = @{thm vimage_Collect} val vimage_Int = @{thm vimage_Int} val vimage_Un = @{thm vimage_Un} \ end diff --git a/src/HOL/Statespace/state_space.ML b/src/HOL/Statespace/state_space.ML --- a/src/HOL/Statespace/state_space.ML +++ b/src/HOL/Statespace/state_space.ML @@ -1,734 +1,736 @@ (* Title: HOL/Statespace/state_space.ML Author: Norbert Schirmer, TU Muenchen, 2007 Author: Norbert Schirmer, Apple, 2021 *) signature STATE_SPACE = sig val distinct_compsN : string val getN : string val putN : string val injectN : string val namespaceN : string val projectN : string val valuetypesN : string val namespace_definition : bstring -> typ -> (xstring, string) Expression.expr * (binding * string option * mixfix) list -> string list -> string list -> theory -> theory val define_statespace : string list -> string -> ((string * bool) * (string list * bstring * (string * string) list)) list -> (string * string) list -> theory -> theory val define_statespace_i : string option -> string list -> string -> ((string * bool) * (typ list * bstring * (string * string) list)) list -> (string * typ) list -> theory -> theory val statespace_decl : ((string list * bstring) * (((string * bool) * (string list * xstring * (bstring * bstring) list)) list * (bstring * string) list)) parser val neq_x_y : Proof.context -> term -> term -> thm option val distinctNameSolver : Simplifier.solver val distinctTree_tac : Proof.context -> int -> tactic val distinct_simproc : Simplifier.simproc val is_statespace : Context.generic -> xstring -> bool val get_comp' : Context.generic -> string -> (typ * string list) option val get_comp : Context.generic -> string -> (typ * string) option (* legacy wrapper, eventually replace by primed version *) val get_comps : Context.generic -> (typ * string list) Termtab.table val silent : bool Config.T val gen_lookup_tr : Proof.context -> term -> string -> term val lookup_swap_tr : Proof.context -> term list -> term val lookup_tr : Proof.context -> term list -> term val lookup_tr' : Proof.context -> term list -> term val gen'_update_tr : bool -> bool -> Proof.context -> string -> term -> term -> term val gen_update_tr : (* legacy wrapper, eventually replace by primed version *) bool -> Proof.context -> string -> term -> term -> term val update_tr : Proof.context -> term list -> term val update_tr' : Proof.context -> term list -> term val trace_data: Context.generic -> unit end; structure StateSpace : STATE_SPACE = struct (* Names *) val distinct_compsN = "distinct_names" val namespaceN = "_namespace" val valuetypesN = "_valuetypes" val projectN = "project" val injectN = "inject" val getN = "get" val putN = "put" val project_injectL = "StateSpaceLocale.project_inject"; (* Library *) fun fold1 f xs = fold f (tl xs) (hd xs) fun fold1' f [] x = x | fold1' f xs _ = fold1 f xs fun sorted_subset eq [] ys = true | sorted_subset eq (x::xs) [] = false | sorted_subset eq (x::xs) (y::ys) = if eq (x,y) then sorted_subset eq xs ys else sorted_subset eq (x::xs) ys; fun comps_of_distinct_thm thm = Thm.prop_of thm |> (fn (_$(_$t)) => DistinctTreeProver.dest_tree t) |> map (fst o dest_Free) |> sort_strings; fun insert_tagged_distinct_thms tagged_thm tagged_thms = let fun ins t1 [] = [t1] | ins (t1 as (names1, _)) ((t2 as (names2, _))::thms) = if Ord_List.subset string_ord (names1, names2) then t2::thms else if Ord_List.subset string_ord (names2, names1) then ins t1 thms else t2 :: ins t1 thms in ins tagged_thm tagged_thms end fun join_tagged_distinct_thms tagged_thms1 tagged_thms2 = tagged_thms1 |> fold insert_tagged_distinct_thms tagged_thms2 fun tag_distinct_thm thm = (comps_of_distinct_thm thm, thm) val tag_distinct_thms = map tag_distinct_thm fun join_distinct_thms (thms1, thms2) = if pointer_eq (thms1, thms2) then thms1 else join_tagged_distinct_thms (tag_distinct_thms thms1) (tag_distinct_thms thms2) |> map snd fun insert_distinct_thm thm thms = join_distinct_thms (thms, [thm]) fun join_declinfo_entry name (T1:typ, names1:string list) (T2, names2) = let fun typ_info T names = @{make_string} T ^ " " ^ Pretty.string_of (Pretty.str_list "(" ")" names); in if T1 = T2 then (T1, distinct (op =) (names1 @ names2)) else error ("statespace component '" ^ name ^ "' disagrees on type:\n " ^ typ_info T1 names1 ^ " vs. " ^ typ_info T2 names2 ) end fun guess_name (Free (x,_)) = x | guess_name _ = "unknown" val join_declinfo = Termtab.join (fn t => uncurry (join_declinfo_entry (guess_name t))) (* A component might appear in *different* statespaces within the same context. However, it must be mapped to the same type. Note that this information is currently only properly maintained within contexts where all components are actually "fixes" and not arbitrary terms. Moreover, on the theory level the info stays empty. This means that on the theory level we do not make use of the syntax and the tree-based distinct simprocs. N.B: It might still make sense (from a performance perspective) to overcome this limitation and even use the simprocs when a statespace is interpreted for concrete names like HOL-strings. Once the distinct-theorem is proven by the interpretation the simproc can use the positions in the tree to derive distinctness of two strings vs. HOL-string comparison. *) type statespace_info = {args: (string * sort) list, (* type arguments *) parents: (typ list * string * string option list) list, (* type instantiation, state-space name, component renamings *) components: (string * typ) list, types: typ list (* range types of state space *) }; structure Data = Generic_Data ( type T = (typ * string list) Termtab.table * (*declinfo: type, names of statespaces of component*) thm list Symtab.table * (*distinctthm: minimal list of maximal distinctness-assumptions for a component name*) statespace_info Symtab.table; val empty = (Termtab.empty, Symtab.empty, Symtab.empty); fun merge ((declinfo1, distinctthm1, statespaces1), (declinfo2, distinctthm2, statespaces2)) = (join_declinfo (declinfo1, declinfo2), Symtab.join (K join_distinct_thms) (distinctthm1, distinctthm2), Symtab.merge (K true) (statespaces1, statespaces2)); ); val get_declinfo = #1 o Data.get val get_distinctthm = #2 o Data.get val get_statespaces = #3 o Data.get val map_declinfo = Data.map o @{apply 3(1)} val map_distinctthm = Data.map o @{apply 3(2)} val map_statespaces = Data.map o @{apply 3(3)} fun trace_data context = tracing ("StateSpace.Data: " ^ @{make_string} {declinfo = get_declinfo context, distinctthm = get_distinctthm context, statespaces = get_statespaces context}) fun update_declinfo (n,v) = map_declinfo (fn declinfo => let val vs = apsnd single v in Termtab.map_default (n, vs) (join_declinfo_entry (guess_name n) vs) declinfo end); fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun prove_interpretation_in ctxt_tac (name, expr) thy = thy |> Interpretation.global_sublocale_cmd (name, Position.none) (expression_no_pos expr) [] |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => SIMPLE_METHOD (ctxt_tac ctxt)), Position.no_range), NONE) |> Proof_Context.theory_of fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) Binding.empty [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun is_statespace context name = Symtab.defined (get_statespaces context) (Locale.intern (Context.theory_of context) name) fun add_statespace name args parents components types = map_statespaces (Symtab.update_new (name, {args=args,parents=parents, components=components,types=types})) val get_statespace = Symtab.lookup o get_statespaces val the_statespace = the oo get_statespace fun mk_free ctxt name = if Variable.is_fixed ctxt name orelse Variable.is_declared ctxt name then let val n' = Variable.intern_fixed ctxt name |> perhaps Long_Name.dest_hidden; val free = Free (n', Proof_Context.infer_type ctxt (n', dummyT)) in SOME (free) end else (tracing ("variable not fixed or declared: " ^ name); NONE) fun get_dist_thm context name = Symtab.lookup_list (get_distinctthm context) name |> map (Thm.transfer'' context) fun get_dist_thm2 ctxt x y = (let val dist_thms = [x, y] |> map (#1 o dest_Free) |> maps (get_dist_thm (Context.Proof ctxt)); fun get_paths dist_thm = let val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val tree = Thm.term_of ctree; val x_path = the (DistinctTreeProver.find_tree x tree); val y_path = the (DistinctTreeProver.find_tree y tree); in SOME (dist_thm, x_path, y_path) end handle Option.Option => NONE val result = get_first get_paths dist_thms in result end) fun get_comp' context name = mk_free (Context.proof_of context) name |> Option.mapPartial (fn t => let val declinfo = get_declinfo context in case Termtab.lookup declinfo t of NONE => (* during syntax phase, types of fixes might not yet be constrained completely *) AList.lookup (fn (x, Free (n,_)) => n = x | _ => false) (Termtab.dest declinfo) name | some => some end) (* legacy wrapper *) fun get_comp ctxt name = get_comp' ctxt name |> Option.map (apsnd (fn ns => if null ns then "" else hd ns)) val get_comps = get_declinfo (*** Tactics ***) fun neq_x_y ctxt x y = (let val (dist_thm, x_path, y_path) = the (get_dist_thm2 ctxt x y); val thm = DistinctTreeProver.distinctTreeProver ctxt dist_thm x_path y_path; in SOME thm end handle Option.Option => NONE) fun distinctTree_tac ctxt = SUBGOAL (fn (goal, i) => (case goal of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ (x as Free _) $ (y as Free _))) => (case neq_x_y ctxt x y of SOME neq => resolve_tac ctxt [neq] i | NONE => no_tac) | _ => no_tac)); val distinctNameSolver = mk_solver "distinctNameSolver" distinctTree_tac; val distinct_simproc = Simplifier.make_simproc \<^context> "StateSpace.distinct_simproc" {lhss = [\<^term>\x = y\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\,_) $ (x as Free _) $ (y as Free _) => Option.map (fn neq => DistinctTreeProver.neq_to_eq_False OF [neq]) (neq_x_y ctxt x y) | _ => NONE)}; fun interprete_parent name dist_thm_name parent_expr thy = let fun solve_tac ctxt = CSUBGOAL (fn (goal, i) => let val distinct_thm = Proof_Context.get_thm ctxt dist_thm_name; val rule = DistinctTreeProver.distinct_implProver ctxt distinct_thm goal; in resolve_tac ctxt [rule] i end); fun tac ctxt = Locale.intro_locales_tac {strict = true, eager = true} ctxt [] THEN ALLGOALS (solve_tac ctxt); in thy |> prove_interpretation_in tac (name, parent_expr) end; fun namespace_definition name nameT parent_expr parent_comps new_comps thy = let val all_comps = parent_comps @ new_comps; val vars = (map (fn n => (Binding.name n, NONE, NoSyn)) all_comps); val dist_thm_name = distinct_compsN; val dist_thm_full_name = dist_thm_name; fun type_attr phi = Thm.declaration_attribute (fn thm => fn context => (case context of Context.Theory _ => context | Context.Proof ctxt => let val declinfo = get_declinfo context val tt = get_distinctthm context; val all_names = comps_of_distinct_thm thm; val thm0 = Thm.trim_context thm; fun upd name = Symtab.map_default (name, [thm0]) (insert_distinct_thm thm0) val tt' = tt |> fold upd all_names; val context' = Context_Position.set_visible false ctxt addsimprocs [distinct_simproc] |> Context_Position.restore_visible ctxt |> Context.Proof |> map_declinfo (K declinfo) |> map_distinctthm (K tt'); in context' end)); - val attr = Attrib.internal type_attr; + val attr = Attrib.internal \<^here> type_attr; val assume = ((Binding.name dist_thm_name, [attr]), [(HOLogic.Trueprop $ (Const (\<^const_name>\all_distinct\, Type (\<^type_name>\tree\, [nameT]) --> HOLogic.boolT) $ DistinctTreeProver.mk_tree (fn n => Free (n, nameT)) nameT (sort fast_string_ord all_comps)), [])]); in thy |> add_locale name ([], vars) [Element.Assumes [assume]] |> Proof_Context.theory_of |> interprete_parent name dist_thm_full_name parent_expr end; fun encode_dot x = if x = #"." then #"_" else x; fun encode_type (TFree (s, _)) = s | encode_type (TVar ((s,i),_)) = "?" ^ s ^ string_of_int i | encode_type (Type (n,Ts)) = let val Ts' = fold1' (fn x => fn y => x ^ "_" ^ y) (map encode_type Ts) ""; val n' = String.map encode_dot n; in if Ts'="" then n' else Ts' ^ "_" ^ n' end; fun project_name T = projectN ^"_"^encode_type T; fun inject_name T = injectN ^"_"^encode_type T; fun add_declaration name decl thy = thy |> Named_Target.init [] name - |> (fn lthy => Local_Theory.declaration {syntax = true, pervasive = false} (decl lthy) lthy) + |> (fn lthy => + Local_Theory.declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} + (decl lthy) lthy) |> Local_Theory.exit_global; fun parent_components thy (Ts, pname, renaming) = let fun rename [] xs = xs | rename (NONE::rs) (x::xs) = x::rename rs xs | rename (SOME r::rs) ((x,T)::xs) = (r,T)::rename rs xs; val {args, parents, components, ...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val parent_comps = maps (fn (Ts',n,rs) => parent_components thy (map subst Ts', n, rs)) parents; val all_comps = rename renaming (parent_comps @ map (apsnd subst) components); in all_comps end; fun statespace_definition state_type args name parents parent_comps components thy = let val full_name = Sign.full_bname thy name; val all_comps = parent_comps @ components; val components' = map (fn (n,T) => (n,(T,full_name))) components; fun parent_expr (prefix, (_, n, rs)) = (suffix namespaceN n, (prefix, (Expression.Positional rs,[]))); val parents_expr = map parent_expr parents; fun distinct_types Ts = let val tab = fold (fn T => fn tab => Typtab.update (T,()) tab) Ts Typtab.empty; in map fst (Typtab.dest tab) end; val Ts = distinct_types (map snd all_comps); val arg_names = map fst args; val valueN = singleton (Name.variant_list arg_names) "'value"; val nameN = singleton (Name.variant_list (valueN :: arg_names)) "'name"; val valueT = TFree (valueN, Sign.defaultS thy); val nameT = TFree (nameN, Sign.defaultS thy); val stateT = nameT --> valueT; fun projectT T = valueT --> T; fun injectT T = T --> valueT; val locinsts = map (fn T => (project_injectL, ((encode_type T,false),(Expression.Positional [SOME (Free (project_name T,projectT T)), SOME (Free ((inject_name T,injectT T)))],[])))) Ts; val locs = maps (fn T => [(Binding.name (project_name T),NONE,NoSyn), (Binding.name (inject_name T),NONE,NoSyn)]) Ts; val constrains = maps (fn T => [(project_name T,projectT T),(inject_name T,injectT T)]) Ts; fun interprete_parent_valuetypes (prefix, (Ts, pname, _)) thy = let val {args,types,...} = the_statespace (Context.Theory thy) pname; val inst = map fst args ~~ Ts; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst o fst); val pars = maps ((fn T => [project_name T,inject_name T]) o subst) types; val expr = ([(suffix valuetypesN name, (prefix, (Expression.Positional (map SOME pars),[])))],[]); in prove_interpretation_in (fn ctxt => ALLGOALS (solve_tac ctxt (Assumption.all_prems_of ctxt))) (suffix valuetypesN name, expr) thy end; fun interprete_parent (prefix, (_, pname, rs)) = let val expr = ([(pname, (prefix, (Expression.Positional rs,[])))],[]) in prove_interpretation_in (fn ctxt => Locale.intro_locales_tac {strict = true, eager = false} ctxt []) (full_name, expr) end; fun declare_declinfo updates lthy phi ctxt = let fun upd_prf ctxt = let fun upd (n,v) = let val nT = Proof_Context.infer_type (Local_Theory.target_of lthy) (n, dummyT) in Context.proof_map (update_declinfo (Morphism.term phi (Free (n,nT)),v)) end; val ctxt' = ctxt |> fold upd updates in ctxt' end; in Context.mapping I upd_prf ctxt end; fun string_of_typ T = Print_Mode.setmp [] (Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy))) T; val fixestate = (case state_type of NONE => [] | SOME s => let val fx = Element.Fixes [(Binding.name s,SOME (string_of_typ stateT),NoSyn)]; val cs = Element.Constrains (map (fn (n,T) => (n,string_of_typ T)) ((map (fn (n,_) => (n,nameT)) all_comps) @ constrains)) in [fx,cs] end ) in thy |> namespace_definition (suffix namespaceN name) nameT (parents_expr,[]) (map fst parent_comps) (map fst components) |> Context.theory_map (add_statespace full_name args (map snd parents) components []) |> add_locale (suffix valuetypesN name) (locinsts,locs) [] |> Proof_Context.theory_of |> fold interprete_parent_valuetypes parents |> add_locale_cmd name ([(suffix namespaceN full_name ,(("",false),(Expression.Named [],[]))), (suffix valuetypesN full_name,(("",false),(Expression.Named [],[])))],[]) fixestate |> Proof_Context.theory_of |> fold interprete_parent parents |> add_declaration full_name (declare_declinfo components') end; (* prepare arguments *) fun read_typ ctxt raw_T env = let val ctxt' = fold (Variable.declare_typ o TFree) env ctxt; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun cert_typ ctxt raw_T env = let val thy = Proof_Context.theory_of ctxt; val T = Type.no_tvars (Sign.certify_typ thy raw_T) handle TYPE (msg, _, _) => error msg; val env' = Term.add_tfreesT T env; in (T, env') end; fun gen_define_statespace prep_typ state_space args name parents comps thy = let (* - args distinct - only args may occur in comps and parent-instantiations - number of insts must match parent args - no duplicate renamings - renaming should occur in namespace *) val _ = writeln ("Defining statespace " ^ quote name ^ " ..."); val ctxt = Proof_Context.init_global thy; fun add_parent (prefix, (Ts, pname, rs)) env = let val prefix' = (case prefix of ("", mandatory) => (pname, mandatory) | _ => prefix); val full_pname = Sign.full_bname thy pname; val {args,components,...} = (case get_statespace (Context.Theory thy) full_pname of SOME r => r | NONE => error ("Undefined statespace " ^ quote pname)); val (Ts',env') = fold_map (prep_typ ctxt) Ts env handle ERROR msg => cat_error msg ("The error(s) above occurred in parent statespace specification " ^ quote pname); val err_insts = if length args <> length Ts' then ["number of type instantiation(s) does not match arguments of parent statespace " ^ quote pname] else []; val rnames = map fst rs val err_dup_renamings = (case duplicates (op =) rnames of [] => [] | dups => ["Duplicate renaming(s) for " ^ commas dups]) val cnames = map fst components; val err_rename_unknowns = (case subtract (op =) cnames rnames of [] => [] | rs => ["Unknown components " ^ commas rs]); val rs' = map (AList.lookup (op =) rs o fst) components; val errs =err_insts @ err_dup_renamings @ err_rename_unknowns in if null errs then ((prefix', (Ts', full_pname, rs')), env') else error (cat_lines (errs @ ["in parent statespace " ^ quote pname])) end; val (parents',env) = fold_map add_parent parents []; val err_dup_args = (case duplicates (op =) args of [] => [] | dups => ["Duplicate type argument(s) " ^ commas dups]); val err_dup_components = (case duplicates (op =) (map fst comps) of [] => [] | dups => ["Duplicate state-space components " ^ commas dups]); fun prep_comp (n,T) env = let val (T', env') = prep_typ ctxt T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n,T'), env') end; val (comps',env') = fold_map prep_comp comps env; val err_extra_frees = (case subtract (op =) args (map fst env') of [] => [] | extras => ["Extra free type variable(s) " ^ commas extras]); val defaultS = Sign.defaultS thy; val args' = map (fn x => (x, AList.lookup (op =) env x |> the_default defaultS)) args; fun fst_eq ((x:string,_),(y,_)) = x = y; fun snd_eq ((_,t:typ),(_,u)) = t = u; val raw_parent_comps = maps (parent_components thy o snd) parents'; fun check_type (n,T) = (case distinct (snd_eq) (filter (curry fst_eq (n,T)) raw_parent_comps) of [] => [] | [_] => [] | rs => ["Different types for component " ^ quote n ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o snd) rs)]) val err_dup_types = maps check_type (duplicates fst_eq raw_parent_comps) val parent_comps = distinct (fst_eq) raw_parent_comps; val all_comps = parent_comps @ comps'; val err_comp_in_parent = (case duplicates (op =) (map fst all_comps) of [] => [] | xs => ["Components already defined in parents: " ^ commas_quote xs]); val errs = err_dup_args @ err_dup_components @ err_extra_frees @ err_dup_types @ err_comp_in_parent; in if null errs then thy |> statespace_definition state_space args' name parents' parent_comps comps' else error (cat_lines errs) end handle ERROR msg => cat_error msg ("Failed to define statespace " ^ quote name); val define_statespace = gen_define_statespace read_typ NONE; val define_statespace_i = gen_define_statespace cert_typ; (*** parse/print - translations ***) val silent = Attrib.setup_config_bool \<^binding>\statespace_silent\ (K false); fun gen_lookup_tr ctxt s n = (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.free (project_name T) $ Syntax.free n $ s | NONE => if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.lookup\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ s else raise TERM ("StateSpace.gen_lookup_tr: component " ^ quote n ^ " not defined", [])); fun lookup_tr ctxt [s, x] = (case Term_Position.strip_positions x of Free (n,_) => gen_lookup_tr ctxt s n | _ => raise Match); fun lookup_swap_tr ctxt [Free (n,_),s] = gen_lookup_tr ctxt s n; fun lookup_tr' ctxt [_ $ Free (prj, _), n as (_ $ Free (name, _)), s] = (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if prj = project_name T then Syntax.const "_statespace_lookup" $ s $ n else raise Match | NONE => raise Match) | lookup_tr' _ _ = raise Match; fun gen'_update_tr const_val id ctxt n v s = let fun pname T = if id then \<^const_name>\Fun.id\ else project_name T; fun iname T = if id then \<^const_name>\Fun.id\ else inject_name T; val v = if const_val then (Syntax.const \<^const_name>\K_statefun\ $ v) else v in (case get_comp' (Context.Proof ctxt) n of SOME (T, _) => Syntax.const \<^const_name>\StateFun.update\ $ Syntax.free (pname T) $ Syntax.free (iname T) $ Syntax.free n $ v $ s | NONE => if Config.get ctxt silent then Syntax.const \<^const_name>\StateFun.update\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.const \<^const_syntax>\undefined\ $ Syntax.free n $ v $ s else raise TERM ("StateSpace.gen_update_tr: component " ^ n ^ " not defined", [])) end; val gen_update_tr = gen'_update_tr true fun update_tr ctxt [s, x, v] = (case Term_Position.strip_positions x of Free (n, _) => gen'_update_tr true false ctxt n v s | _ => raise Match); fun update_tr' ctxt [_ $ Free (prj, _), _ $ Free (inj, _), n as (_ $ Free (name, _)), (Const (k, _) $ v), s] = if Long_Name.base_name k = Long_Name.base_name \<^const_name>\K_statefun\ then (case get_comp' (Context.Proof ctxt) name of SOME (T, _) => if inj = inject_name T andalso prj = project_name T then Syntax.const "_statespace_update" $ s $ n $ v else raise Match | NONE => raise Match) else raise Match | update_tr' _ _ = raise Match; (*** outer syntax *) local val type_insts = Parse.typ >> single || \<^keyword>\(\ |-- Parse.!!! (Parse.list1 Parse.typ --| \<^keyword>\)\) val comp = Parse.name -- (\<^keyword>\::\ |-- Parse.!!! Parse.typ); fun plus1_unless test scan = scan ::: Scan.repeat (\<^keyword>\+\ |-- Scan.unless test (Parse.!!! scan)); val mapsto = \<^keyword>\=\; val rename = Parse.name -- (mapsto |-- Parse.name); val renames = Scan.optional (\<^keyword>\[\ |-- Parse.!!! (Parse.list1 rename --| \<^keyword>\]\)) []; val parent = Parse_Spec.locale_prefix -- ((type_insts -- Parse.name) || (Parse.name >> pair [])) -- renames >> (fn ((prefix, (insts, name)), renames) => (prefix, (insts, name, renames))); in val statespace_decl = Parse.type_args -- Parse.name -- (\<^keyword>\=\ |-- ((Scan.repeat1 comp >> pair []) || (plus1_unless comp parent -- Scan.optional (\<^keyword>\+\ |-- Parse.!!! (Scan.repeat1 comp)) []))); val _ = Outer_Syntax.command \<^command_keyword>\statespace\ "define state-space as locale context" (statespace_decl >> (fn ((args, name), (parents, comps)) => Toplevel.theory (define_statespace args name parents comps))); end; end; diff --git a/src/HOL/Tools/BNF/bnf_def.ML b/src/HOL/Tools/BNF/bnf_def.ML --- a/src/HOL/Tools/BNF/bnf_def.ML +++ b/src/HOL/Tools/BNF/bnf_def.ML @@ -1,2187 +1,2187 @@ (* Title: HOL/Tools/BNF/bnf_def.ML Author: Dmitriy Traytel, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Author: Jan van Brügge, TU Muenchen Copyright 2012, 2013, 2014, 2022 Definition of bounded natural functors. *) signature BNF_DEF = sig type bnf type nonemptiness_witness = {I: int list, wit: term, prop: thm list} val morph_bnf: morphism -> bnf -> bnf val morph_bnf_defs: morphism -> bnf -> bnf val permute_deads: (typ list -> typ list) -> bnf -> bnf val transfer_bnf: theory -> bnf -> bnf val bnf_of: Proof.context -> string -> bnf option val bnf_of_global: theory -> string -> bnf option val bnf_interpretation: string -> (bnf -> local_theory -> local_theory) -> theory -> theory val interpret_bnf: (string -> bool) -> bnf -> local_theory -> local_theory val register_bnf_raw: string -> bnf -> local_theory -> local_theory val register_bnf: (string -> bool) -> string -> bnf -> local_theory -> local_theory val name_of_bnf: bnf -> binding val T_of_bnf: bnf -> typ val live_of_bnf: bnf -> int val lives_of_bnf: bnf -> typ list val dead_of_bnf: bnf -> int val deads_of_bnf: bnf -> typ list val bd_of_bnf: bnf -> term val nwits_of_bnf: bnf -> int val mapN: string val predN: string val relN: string val setN: string val mk_setN: int -> string val mk_witN: int -> string val map_of_bnf: bnf -> term val pred_of_bnf: bnf -> term val rel_of_bnf: bnf -> term val sets_of_bnf: bnf -> term list val mk_T_of_bnf: typ list -> typ list -> bnf -> typ val mk_bd_of_bnf: typ list -> typ list -> bnf -> term val mk_map_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_pred_of_bnf: typ list -> typ list -> bnf -> term val mk_rel_of_bnf: typ list -> typ list -> typ list -> bnf -> term val mk_sets_of_bnf: typ list list -> typ list list -> bnf -> term list val mk_wits_of_bnf: typ list list -> typ list list -> bnf -> (int list * term) list val bd_Card_order_of_bnf: bnf -> thm val bd_Cinfinite_of_bnf: bnf -> thm val bd_Cnotzero_of_bnf: bnf -> thm val bd_card_order_of_bnf: bnf -> thm val bd_cinfinite_of_bnf: bnf -> thm val bd_regularCard_of_bnf: bnf -> thm val collect_set_map_of_bnf: bnf -> thm val in_bd_of_bnf: bnf -> thm val in_cong_of_bnf: bnf -> thm val in_mono_of_bnf: bnf -> thm val in_rel_of_bnf: bnf -> thm val inj_map_of_bnf: bnf -> thm val inj_map_strong_of_bnf: bnf -> thm val le_rel_OO_of_bnf: bnf -> thm val map_comp0_of_bnf: bnf -> thm val map_comp_of_bnf: bnf -> thm val map_cong0_of_bnf: bnf -> thm val map_cong_of_bnf: bnf -> thm val map_cong_pred_of_bnf: bnf -> thm val map_cong_simp_of_bnf: bnf -> thm val map_def_of_bnf: bnf -> thm val map_id0_of_bnf: bnf -> thm val map_id_of_bnf: bnf -> thm val map_ident0_of_bnf: bnf -> thm val map_ident_of_bnf: bnf -> thm val map_transfer_of_bnf: bnf -> thm val pred_cong0_of_bnf: bnf -> thm val pred_cong_of_bnf: bnf -> thm val pred_cong_simp_of_bnf: bnf -> thm val pred_def_of_bnf: bnf -> thm val pred_map_of_bnf: bnf -> thm val pred_mono_strong0_of_bnf: bnf -> thm val pred_mono_strong_of_bnf: bnf -> thm val pred_mono_of_bnf: bnf -> thm val pred_set_of_bnf: bnf -> thm val pred_rel_of_bnf: bnf -> thm val pred_transfer_of_bnf: bnf -> thm val pred_True_of_bnf: bnf -> thm val rel_Grp_of_bnf: bnf -> thm val rel_OO_Grp_of_bnf: bnf -> thm val rel_OO_of_bnf: bnf -> thm val rel_cong0_of_bnf: bnf -> thm val rel_cong_of_bnf: bnf -> thm val rel_cong_simp_of_bnf: bnf -> thm val rel_conversep_of_bnf: bnf -> thm val rel_def_of_bnf: bnf -> thm val rel_eq_of_bnf: bnf -> thm val rel_flip_of_bnf: bnf -> thm val rel_map_of_bnf: bnf -> thm list val rel_mono_of_bnf: bnf -> thm val rel_mono_strong0_of_bnf: bnf -> thm val rel_mono_strong_of_bnf: bnf -> thm val rel_eq_onp_of_bnf: bnf -> thm val rel_refl_of_bnf: bnf -> thm val rel_refl_strong_of_bnf: bnf -> thm val rel_reflp_of_bnf: bnf -> thm val rel_symp_of_bnf: bnf -> thm val rel_transfer_of_bnf: bnf -> thm val rel_transp_of_bnf: bnf -> thm val set_bd_of_bnf: bnf -> thm list val set_defs_of_bnf: bnf -> thm list val set_map0_of_bnf: bnf -> thm list val set_map_of_bnf: bnf -> thm list val set_transfer_of_bnf: bnf -> thm list val wit_thms_of_bnf: bnf -> thm list val wit_thmss_of_bnf: bnf -> thm list list val mk_map: int -> typ list -> typ list -> term -> term val mk_pred: typ list -> term -> term val mk_rel: int -> typ list -> typ list -> term -> term val mk_set: typ list -> term -> term val build_map: Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_rel: (string * (int * term)) list -> Proof.context -> typ list -> typ list -> (typ * typ -> term) -> typ * typ -> term val build_set: Proof.context -> typ -> typ -> term val flatten_type_args_of_bnf: bnf -> 'a -> 'a list -> 'a list val map_flattened_map_args: Proof.context -> string -> (term list -> 'a list) -> term list -> 'a list val mk_witness: int list * term -> thm list -> nonemptiness_witness val mk_wit_goals: term list -> term list -> term list -> int list * term -> term list val minimize_wits: (''a list * 'b) list -> (''a list * 'b) list val wits_of_bnf: bnf -> nonemptiness_witness list val zip_axioms: 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list -> 'a -> 'a -> 'a -> 'a list datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline datatype fact_policy = Dont_Note | Note_Some | Note_All val bnf_internals: bool Config.T val bnf_timing: bool Config.T val user_policy: fact_policy -> Proof.context -> fact_policy val note_bnf_thms: fact_policy -> (binding -> binding) -> binding -> bnf -> local_theory -> bnf * local_theory val note_bnf_defs: bnf -> local_theory -> bnf * local_theory val print_bnfs: Proof.context -> unit val prepare_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> 'a -> typ) -> (Proof.context -> 'b -> term) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * 'a) * 'b) * 'b list) * 'b) * 'b list) * 'b option) * 'b option -> Proof.context -> string * term list * ((Proof.context -> thm list -> tactic) option * term list list) * ((thm list -> thm list list) -> thm list list -> Proof.context -> bnf * local_theory) * local_theory * thm list val define_bnf_consts: inline_policy -> fact_policy -> bool -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> ((typ list * typ list * typ list * typ) * (term * term list * term * (int list * term) list * term * term) * (thm * thm list * thm * thm list * thm * thm) * ((typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term -> term) * (typ list -> typ list -> typ -> typ) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term) * (typ list -> typ list -> typ list -> term) * (typ list -> typ list -> term))) * local_theory val bnf_def: inline_policy -> (Proof.context -> fact_policy) -> bool -> (binding -> binding) -> (Proof.context -> tactic) list -> (Proof.context -> tactic) -> typ list option -> binding -> binding -> binding -> binding list -> ((((((binding * typ) * term) * term list) * term) * term list) * term option) * term option -> local_theory -> bnf * local_theory val bnf_cmd: (((((((binding * string) * string) * string list) * string) * string list) * string option) * string option) * (Proof.context -> Plugin_Name.filter) -> Proof.context -> Proof.state end; structure BNF_Def : BNF_DEF = struct open BNF_Util open BNF_Tactics open BNF_Def_Tactics val fundefcong_attrs = @{attributes [fundef_cong]}; val mono_attrs = @{attributes [mono]}; type axioms = { map_id0: thm, map_comp0: thm, map_cong0: thm, set_map0: thm list, bd_card_order: thm, bd_cinfinite: thm, bd_regularCard: thm, set_bd: thm list, le_rel_OO: thm, rel_OO_Grp: thm, pred_set: thm }; fun mk_axioms' ((((((((((id, comp), cong), map), c_o), cinf), creg), set_bd), le_rel_OO), rel), pred) = {map_id0 = id, map_comp0 = comp, map_cong0 = cong, set_map0 = map, bd_card_order = c_o, bd_cinfinite = cinf, bd_regularCard = creg, set_bd = set_bd, le_rel_OO = le_rel_OO, rel_OO_Grp = rel, pred_set = pred}; fun dest_cons [] = raise List.Empty | dest_cons (x :: xs) = (x, xs); fun mk_axioms n thms = thms |> map the_single |> dest_cons ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||>> dest_cons ||>> chop n ||>> dest_cons ||>> dest_cons ||> the_single |> mk_axioms'; fun zip_axioms mid mcomp mcong smap bdco bdinf bdreg sbd le_rel_OO rel pred = [mid, mcomp, mcong] @ smap @ [bdco, bdinf, bdreg] @ sbd @ [le_rel_OO, rel, pred]; fun map_axioms f {map_id0, map_comp0, map_cong0, set_map0, bd_card_order, bd_cinfinite, bd_regularCard, set_bd, le_rel_OO, rel_OO_Grp, pred_set} = {map_id0 = f map_id0, map_comp0 = f map_comp0, map_cong0 = f map_cong0, set_map0 = map f set_map0, bd_card_order = f bd_card_order, bd_cinfinite = f bd_cinfinite, bd_regularCard = f bd_regularCard, set_bd = map f set_bd, le_rel_OO = f le_rel_OO, rel_OO_Grp = f rel_OO_Grp, pred_set = f pred_set}; val morph_axioms = map_axioms o Morphism.thm; type defs = { map_def: thm, set_defs: thm list, rel_def: thm, pred_def: thm } fun mk_defs map sets rel pred = {map_def = map, set_defs = sets, rel_def = rel, pred_def = pred}; fun map_defs f {map_def, set_defs, rel_def, pred_def} = {map_def = f map_def, set_defs = map f set_defs, rel_def = f rel_def, pred_def = f pred_def}; val morph_defs = map_defs o Morphism.thm; type facts = { bd_Card_order: thm, bd_Cinfinite: thm, bd_Cnotzero: thm, collect_set_map: thm lazy, in_bd: thm lazy, in_cong: thm lazy, in_mono: thm lazy, in_rel: thm lazy, inj_map: thm lazy, inj_map_strong: thm lazy, map_comp: thm lazy, map_cong: thm lazy, map_cong_simp: thm lazy, map_cong_pred: thm lazy, map_id: thm lazy, map_ident0: thm lazy, map_ident: thm lazy, map_ident_strong: thm lazy, map_transfer: thm lazy, rel_eq: thm lazy, rel_flip: thm lazy, set_map: thm lazy list, rel_cong0: thm lazy, rel_cong: thm lazy, rel_cong_simp: thm lazy, rel_map: thm list lazy, rel_mono: thm lazy, rel_mono_strong0: thm lazy, rel_mono_strong: thm lazy, set_transfer: thm list lazy, rel_Grp: thm lazy, rel_conversep: thm lazy, rel_OO: thm lazy, rel_refl: thm lazy, rel_refl_strong: thm lazy, rel_reflp: thm lazy, rel_symp: thm lazy, rel_transp: thm lazy, rel_transfer: thm lazy, rel_eq_onp: thm lazy, pred_transfer: thm lazy, pred_True: thm lazy, pred_map: thm lazy, pred_rel: thm lazy, pred_mono_strong0: thm lazy, pred_mono_strong: thm lazy, pred_mono: thm lazy, pred_cong0: thm lazy, pred_cong: thm lazy, pred_cong_simp: thm lazy }; fun mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp = { bd_Card_order = bd_Card_order, bd_Cinfinite = bd_Cinfinite, bd_Cnotzero = bd_Cnotzero, collect_set_map = collect_set_map, in_bd = in_bd, in_cong = in_cong, in_mono = in_mono, in_rel = in_rel, inj_map = inj_map, inj_map_strong = inj_map_strong, map_comp = map_comp, map_cong = map_cong, map_cong_simp = map_cong_simp, map_cong_pred = map_cong_pred, map_id = map_id, map_ident0 = map_ident0, map_ident = map_ident, map_ident_strong = map_ident_strong, map_transfer = map_transfer, rel_eq = rel_eq, rel_flip = rel_flip, set_map = set_map, rel_cong0 = rel_cong0, rel_cong = rel_cong, rel_cong_simp = rel_cong_simp, rel_map = rel_map, rel_mono = rel_mono, rel_mono_strong0 = rel_mono_strong0, rel_mono_strong = rel_mono_strong, rel_transfer = rel_transfer, rel_Grp = rel_Grp, rel_conversep = rel_conversep, rel_OO = rel_OO, rel_refl = rel_refl, rel_refl_strong = rel_refl_strong, rel_reflp = rel_reflp, rel_symp = rel_symp, rel_transp = rel_transp, set_transfer = set_transfer, rel_eq_onp = rel_eq_onp, pred_transfer = pred_transfer, pred_True = pred_True, pred_map = pred_map, pred_rel = pred_rel, pred_mono_strong0 = pred_mono_strong0, pred_mono_strong = pred_mono_strong, pred_mono = pred_mono, pred_cong0 = pred_cong0, pred_cong = pred_cong, pred_cong_simp = pred_cong_simp}; fun map_facts f { bd_Card_order, bd_Cinfinite, bd_Cnotzero, collect_set_map, in_bd, in_cong, in_mono, in_rel, inj_map, inj_map_strong, map_comp, map_cong, map_cong_simp, map_cong_pred, map_id, map_ident0, map_ident, map_ident_strong, map_transfer, rel_eq, rel_flip, set_map, rel_cong0, rel_cong, rel_cong_simp, rel_map, rel_mono, rel_mono_strong0, rel_mono_strong, rel_transfer, rel_Grp, rel_conversep, rel_OO, rel_refl, rel_refl_strong, rel_reflp, rel_symp, rel_transp, set_transfer, rel_eq_onp, pred_transfer, pred_True, pred_map, pred_rel, pred_mono_strong0, pred_mono_strong, pred_mono, pred_cong0, pred_cong, pred_cong_simp} = {bd_Card_order = f bd_Card_order, bd_Cinfinite = f bd_Cinfinite, bd_Cnotzero = f bd_Cnotzero, collect_set_map = Lazy.map f collect_set_map, in_bd = Lazy.map f in_bd, in_cong = Lazy.map f in_cong, in_mono = Lazy.map f in_mono, in_rel = Lazy.map f in_rel, inj_map = Lazy.map f inj_map, inj_map_strong = Lazy.map f inj_map_strong, map_comp = Lazy.map f map_comp, map_cong = Lazy.map f map_cong, map_cong_simp = Lazy.map f map_cong_simp, map_cong_pred = Lazy.map f map_cong_pred, map_id = Lazy.map f map_id, map_ident0 = Lazy.map f map_ident0, map_ident = Lazy.map f map_ident, map_ident_strong = Lazy.map f map_ident_strong, map_transfer = Lazy.map f map_transfer, rel_eq = Lazy.map f rel_eq, rel_flip = Lazy.map f rel_flip, set_map = map (Lazy.map f) set_map, rel_cong0 = Lazy.map f rel_cong0, rel_cong = Lazy.map f rel_cong, rel_cong_simp = Lazy.map f rel_cong_simp, rel_map = Lazy.map (map f) rel_map, rel_mono = Lazy.map f rel_mono, rel_mono_strong0 = Lazy.map f rel_mono_strong0, rel_mono_strong = Lazy.map f rel_mono_strong, rel_transfer = Lazy.map f rel_transfer, rel_Grp = Lazy.map f rel_Grp, rel_conversep = Lazy.map f rel_conversep, rel_OO = Lazy.map f rel_OO, rel_refl = Lazy.map f rel_refl, rel_refl_strong = Lazy.map f rel_refl_strong, rel_reflp = Lazy.map f rel_reflp, rel_symp = Lazy.map f rel_symp, rel_transp = Lazy.map f rel_transp, set_transfer = Lazy.map (map f) set_transfer, rel_eq_onp = Lazy.map f rel_eq_onp, pred_transfer = Lazy.map f pred_transfer, pred_True = Lazy.map f pred_True, pred_map = Lazy.map f pred_map, pred_rel = Lazy.map f pred_rel, pred_mono_strong0 = Lazy.map f pred_mono_strong0, pred_mono_strong = Lazy.map f pred_mono_strong, pred_mono = Lazy.map f pred_mono, pred_cong0 = Lazy.map f pred_cong0, pred_cong = Lazy.map f pred_cong, pred_cong_simp = Lazy.map f pred_cong_simp}; val morph_facts = map_facts o Morphism.thm; type nonemptiness_witness = { I: int list, wit: term, prop: thm list }; fun mk_witness (I, wit) prop = {I = I, wit = wit, prop = prop}; fun map_witness f g {I, wit, prop} = {I = I, wit = f wit, prop = map g prop}; fun morph_witness phi = map_witness (Morphism.term phi) (Morphism.thm phi); datatype bnf = BNF of { name: binding, T: typ, live: int, lives: typ list, (*source type variables of map*) lives': typ list, (*target type variables of map*) dead: int, deads: typ list, map: term, sets: term list, bd: term, axioms: axioms, defs: defs, facts: facts, nwits: int, wits: nonemptiness_witness list, rel: term, pred: term }; (* getters *) fun rep_bnf (BNF bnf) = bnf; val name_of_bnf = #name o rep_bnf; val T_of_bnf = #T o rep_bnf; fun mk_T_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf in Term.typ_subst_atomic ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#T bnf_rep) end; val live_of_bnf = #live o rep_bnf; val lives_of_bnf = #lives o rep_bnf; val dead_of_bnf = #dead o rep_bnf; val deads_of_bnf = #deads o rep_bnf; val axioms_of_bnf = #axioms o rep_bnf; val facts_of_bnf = #facts o rep_bnf; val nwits_of_bnf = #nwits o rep_bnf; val wits_of_bnf = #wits o rep_bnf; fun flatten_type_args_of_bnf bnf dead_x xs = let val Type (_, Ts) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val deads = deads_of_bnf bnf; in permute_like_unique (op =) (deads @ lives) Ts (replicate (length deads) dead_x @ xs) end; (*terms*) val map_of_bnf = #map o rep_bnf; val sets_of_bnf = #sets o rep_bnf; fun mk_map_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#map bnf_rep) end; fun mk_sets_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; in map2 (fn (Ds, Ts) => Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts))) (Dss ~~ Tss) (#sets bnf_rep) end; val bd_of_bnf = #bd o rep_bnf; fun mk_bd_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#bd bnf_rep) end; fun mk_wits_of_bnf Dss Tss bnf = let val bnf_rep = rep_bnf bnf; val wits = map (fn x => (#I x, #wit x)) (#wits bnf_rep); in map2 (fn (Ds, Ts) => apsnd (Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)))) (Dss ~~ Tss) wits end; val rel_of_bnf = #rel o rep_bnf; fun mk_rel_of_bnf Ds Ts Us bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts) @ (#lives' bnf_rep ~~ Us)) (#rel bnf_rep) end; val pred_of_bnf = #pred o rep_bnf; fun mk_pred_of_bnf Ds Ts bnf = let val bnf_rep = rep_bnf bnf; in Term.subst_atomic_types ((#deads bnf_rep ~~ Ds) @ (#lives bnf_rep ~~ Ts)) (#pred bnf_rep) end; (*thms*) val bd_Card_order_of_bnf = #bd_Card_order o #facts o rep_bnf; val bd_Cinfinite_of_bnf = #bd_Cinfinite o #facts o rep_bnf; val bd_Cnotzero_of_bnf = #bd_Cnotzero o #facts o rep_bnf; val bd_card_order_of_bnf = #bd_card_order o #axioms o rep_bnf; val bd_cinfinite_of_bnf = #bd_cinfinite o #axioms o rep_bnf; val bd_regularCard_of_bnf = #bd_regularCard o #axioms o rep_bnf; val collect_set_map_of_bnf = Lazy.force o #collect_set_map o #facts o rep_bnf; val in_bd_of_bnf = Lazy.force o #in_bd o #facts o rep_bnf; val in_cong_of_bnf = Lazy.force o #in_cong o #facts o rep_bnf; val in_mono_of_bnf = Lazy.force o #in_mono o #facts o rep_bnf; val in_rel_of_bnf = Lazy.force o #in_rel o #facts o rep_bnf; val inj_map_of_bnf = Lazy.force o #inj_map o #facts o rep_bnf; val inj_map_strong_of_bnf = Lazy.force o #inj_map_strong o #facts o rep_bnf; val le_rel_OO_of_bnf = #le_rel_OO o #axioms o rep_bnf; val map_comp0_of_bnf = #map_comp0 o #axioms o rep_bnf; val map_comp_of_bnf = Lazy.force o #map_comp o #facts o rep_bnf; val map_cong0_of_bnf = #map_cong0 o #axioms o rep_bnf; val map_cong_of_bnf = Lazy.force o #map_cong o #facts o rep_bnf; val map_cong_pred_of_bnf = Lazy.force o #map_cong_pred o #facts o rep_bnf; val map_cong_simp_of_bnf = Lazy.force o #map_cong_simp o #facts o rep_bnf; val map_def_of_bnf = #map_def o #defs o rep_bnf; val map_id0_of_bnf = #map_id0 o #axioms o rep_bnf; val map_id_of_bnf = Lazy.force o #map_id o #facts o rep_bnf; val map_ident0_of_bnf = Lazy.force o #map_ident0 o #facts o rep_bnf; val map_ident_of_bnf = Lazy.force o #map_ident o #facts o rep_bnf; val map_ident_strong_of_bnf = Lazy.force o #map_ident_strong o #facts o rep_bnf; val map_transfer_of_bnf = Lazy.force o #map_transfer o #facts o rep_bnf; val rel_eq_onp_of_bnf = Lazy.force o #rel_eq_onp o #facts o rep_bnf; val pred_def_of_bnf = #pred_def o #defs o rep_bnf; val pred_map_of_bnf = Lazy.force o #pred_map o #facts o rep_bnf; val pred_mono_strong0_of_bnf = Lazy.force o #pred_mono_strong0 o #facts o rep_bnf; val pred_mono_strong_of_bnf = Lazy.force o #pred_mono_strong o #facts o rep_bnf; val pred_mono_of_bnf = Lazy.force o #pred_mono o #facts o rep_bnf; val pred_cong0_of_bnf = Lazy.force o #pred_cong0 o #facts o rep_bnf; val pred_cong_of_bnf = Lazy.force o #pred_cong o #facts o rep_bnf; val pred_cong_simp_of_bnf = Lazy.force o #pred_cong_simp o #facts o rep_bnf; val pred_rel_of_bnf = Lazy.force o #pred_rel o #facts o rep_bnf; val pred_set_of_bnf = #pred_set o #axioms o rep_bnf; val pred_transfer_of_bnf = Lazy.force o #pred_transfer o #facts o rep_bnf; val pred_True_of_bnf = Lazy.force o #pred_True o #facts o rep_bnf; val rel_Grp_of_bnf = Lazy.force o #rel_Grp o #facts o rep_bnf; val rel_OO_Grp_of_bnf = #rel_OO_Grp o #axioms o rep_bnf; val rel_OO_of_bnf = Lazy.force o #rel_OO o #facts o rep_bnf; val rel_cong0_of_bnf = Lazy.force o #rel_cong0 o #facts o rep_bnf; val rel_cong_of_bnf = Lazy.force o #rel_cong o #facts o rep_bnf; val rel_cong_simp_of_bnf = Lazy.force o #rel_cong_simp o #facts o rep_bnf; val rel_conversep_of_bnf = Lazy.force o #rel_conversep o #facts o rep_bnf; val rel_def_of_bnf = #rel_def o #defs o rep_bnf; val rel_eq_of_bnf = Lazy.force o #rel_eq o #facts o rep_bnf; val rel_flip_of_bnf = Lazy.force o #rel_flip o #facts o rep_bnf; val rel_map_of_bnf = Lazy.force o #rel_map o #facts o rep_bnf; val rel_mono_of_bnf = Lazy.force o #rel_mono o #facts o rep_bnf; val rel_mono_strong0_of_bnf = Lazy.force o #rel_mono_strong0 o #facts o rep_bnf; val rel_mono_strong_of_bnf = Lazy.force o #rel_mono_strong o #facts o rep_bnf; val rel_refl_of_bnf = Lazy.force o #rel_refl o #facts o rep_bnf; val rel_refl_strong_of_bnf = Lazy.force o #rel_refl_strong o #facts o rep_bnf; val rel_reflp_of_bnf = Lazy.force o #rel_reflp o #facts o rep_bnf; val rel_symp_of_bnf = Lazy.force o #rel_symp o #facts o rep_bnf; val rel_transfer_of_bnf = Lazy.force o #rel_transfer o #facts o rep_bnf; val rel_transp_of_bnf = Lazy.force o #rel_transp o #facts o rep_bnf; val set_bd_of_bnf = #set_bd o #axioms o rep_bnf; val set_defs_of_bnf = #set_defs o #defs o rep_bnf; val set_map0_of_bnf = #set_map0 o #axioms o rep_bnf; val set_map_of_bnf = map Lazy.force o #set_map o #facts o rep_bnf; val set_transfer_of_bnf = Lazy.force o #set_transfer o #facts o rep_bnf; val wit_thms_of_bnf = maps #prop o wits_of_bnf; val wit_thmss_of_bnf = map #prop o wits_of_bnf; fun mk_bnf name T live lives lives' dead deads map sets bd axioms defs facts wits rel pred = BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = length wits, wits = wits, rel = rel, pred = pred}; fun map_bnf f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 f11 f12 f13 f14 f15 f16 f17 (BNF {name = name, T = T, live = live, lives = lives, lives' = lives', dead = dead, deads = deads, map = map, sets = sets, bd = bd, axioms = axioms, defs = defs, facts = facts, nwits = nwits, wits = wits, rel = rel, pred = pred}) = BNF {name = f1 name, T = f2 T, live = f3 live, lives = f4 lives, lives' = f5 lives', dead = f6 dead, deads = f7 deads, map = f8 map, sets = f9 sets, bd = f10 bd, axioms = f11 axioms, defs = f12 defs, facts = f13 facts, nwits = f14 nwits, wits = f15 wits, rel = f16 rel, pred = f17 pred}; fun morph_bnf phi = let val Tphi = Morphism.typ phi; val tphi = Morphism.term phi; in map_bnf (Morphism.binding phi) Tphi I (map Tphi) (map Tphi) I (map Tphi) tphi (map tphi) tphi (morph_axioms phi) (morph_defs phi) (morph_facts phi) I (map (morph_witness phi)) tphi tphi end; fun morph_bnf_defs phi = map_bnf I I I I I I I I I I I (morph_defs phi) I I I I I; fun permute_deads perm = map_bnf I I I I I I perm I I I I I I I I I I; val transfer_bnf = morph_bnf o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = bnf Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun bnf_of_generic context = Option.map (transfer_bnf (Context.theory_of context)) o Symtab.lookup (Data.get context); val bnf_of = bnf_of_generic o Context.Proof; val bnf_of_global = bnf_of_generic o Context.Theory; (* Utilities *) fun normalize_set insts instA set = let val (T, T') = dest_funT (fastype_of set); val A = fst (Term.dest_TVar (HOLogic.dest_setT T')); val params = Term.add_tvar_namesT T []; in Term.subst_TVars ((A :: params) ~~ (instA :: insts)) set end; fun normalize_rel ctxt instTs instA instB rel = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of rel, Library.foldr (op -->) (instTs, mk_pred2T instA instB)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) rel end handle Type.TYPE_MATCH => error "Bad relator"; fun normalize_pred ctxt instTs instA pred = let val thy = Proof_Context.theory_of ctxt; val tyenv = Sign.typ_match thy (fastype_of pred, Library.foldr (op -->) (instTs, mk_pred1T instA)) Vartab.empty; in Envir.subst_term (tyenv, Vartab.empty) pred end handle Type.TYPE_MATCH => error "Bad predicator"; fun normalize_wit insts CA As wit = let fun strip_param (Ts, T as Type (\<^type_name>\fun\, [T1, T2])) = if Type.raw_instance (CA, T) then (Ts, T) else strip_param (T1 :: Ts, T2) | strip_param x = x; val (Ts, T) = strip_param ([], fastype_of wit); val subst = Term.add_tvar_namesT T [] ~~ insts; fun find y = find_index (fn x => x = y) As; in (map (find o Term.typ_subst_TVars subst) (rev Ts), Term.subst_TVars subst wit) end; fun minimize_wits wits = let fun minimize done [] = done | minimize done ((I, wit) :: todo) = if exists (fn (J, _) => subset (op =) (J, I)) (done @ todo) then minimize done todo else minimize ((I, wit) :: done) todo; in minimize [] wits end; fun mk_map live Ts Us t = let val (Type (_, Ts0), Type (_, Us0)) = strip_typeN (live + 1) (fastype_of t) |>> List.last in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun mk_pred Ts t = let val Type (_, Ts0) = domain_type (body_fun_type (fastype_of t)) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_set = mk_pred; fun mk_rel live Ts Us t = let val [Type (_, Ts0), Type (_, Us0)] = binder_types (snd (strip_typeN live (fastype_of t))) in Term.subst_atomic_types (Ts0 @ Us0 ~~ Ts @ Us) t end; fun build_map_or_rel mk const of_bnf dest pre_cst_table ctxt simple_Ts simple_Us build_simple = let fun build (TU as (T, U)) = if exists (curry (op =) T) simple_Ts orelse exists (curry (op =) U) simple_Us then build_simple TU else if T = U andalso not (exists_subtype_in simple_Ts T) andalso not (exists_subtype_in simple_Us U) then const T else (case TU of (Type (s, Ts), Type (s', Us)) => if s = s' then let fun recurse (live, cst0) = let val cst = mk live Ts Us cst0; val TUs' = map dest (fst (strip_typeN live (fastype_of cst))); in Term.list_comb (cst, map build TUs') end; in (case AList.lookup (op =) pre_cst_table s of NONE => (case bnf_of ctxt s of SOME bnf => recurse (live_of_bnf bnf, of_bnf bnf) | NONE => build_simple TU) | SOME entry => recurse entry) end else build_simple TU | _ => build_simple TU); in build end; val build_map = build_map_or_rel mk_map HOLogic.id_const map_of_bnf dest_funT [(\<^type_name>\set\, (1, \<^term>\image\))]; val build_rel = build_map_or_rel mk_rel HOLogic.eq_const rel_of_bnf dest_pred2T o append [(\<^type_name>\set\, (1, \<^term>\rel_set\)), (\<^type_name>\fun\, (2, \<^term>\rel_fun\))]; fun build_set ctxt A = let fun build T = Abs (Name.uu, T, if T = A then HOLogic.mk_set A [Bound 0] else (case T of Type (s, Ts) => let val sets = map (mk_set Ts) (sets_of_bnf (the (bnf_of ctxt s))) |> filter (exists_subtype_in [A] o range_type o fastype_of); val set_apps = map (fn set => Term.betapply (set, Bound 0)) sets; fun recurse set_app = let val Type (\<^type_name>\set\, [elemT]) = fastype_of set_app in if elemT = A then set_app else mk_UNION set_app (build elemT) end; in if null set_apps then HOLogic.mk_set A [] else Library.foldl1 mk_union (map recurse set_apps) end | _ => HOLogic.mk_set A [])); in build end; fun map_flattened_map_args ctxt s map_args fs = let val flat_fs = flatten_type_args_of_bnf (the (bnf_of ctxt s)) Term.dummy fs; val flat_fs' = map_args flat_fs; in permute_like_unique (op aconv) flat_fs fs flat_fs' end; (* Names *) val mapN = "map"; val setN = "set"; fun mk_setN i = setN ^ nonzero_string_of_int i; val bdN = "bd"; val witN = "wit"; fun mk_witN i = witN ^ nonzero_string_of_int i; val relN = "rel"; val predN = "pred"; val bd_Card_orderN = "bd_Card_order"; val bd_CinfiniteN = "bd_Cinfinite"; val bd_CnotzeroN = "bd_Cnotzero"; val bd_card_orderN = "bd_card_order"; val bd_cinfiniteN = "bd_cinfinite"; val bd_regularCardN = "bd_regularCard"; val collect_set_mapN = "collect_set_map"; val in_bdN = "in_bd"; val in_monoN = "in_mono"; val in_relN = "in_rel"; val inj_mapN = "inj_map"; val inj_map_strongN = "inj_map_strong"; val map_comp0N = "map_comp0"; val map_compN = "map_comp"; val map_cong0N = "map_cong0"; val map_congN = "map_cong"; val map_cong_simpN = "map_cong_simp"; val map_cong_predN = "map_cong_pred"; val map_id0N = "map_id0"; val map_idN = "map_id"; val map_identN = "map_ident"; val map_ident_strongN = "map_ident_strong"; val map_transferN = "map_transfer"; val pred_mono_strong0N = "pred_mono_strong0"; val pred_mono_strongN = "pred_mono_strong"; val pred_monoN = "pred_mono"; val pred_transferN = "pred_transfer"; val pred_TrueN = "pred_True"; val pred_mapN = "pred_map"; val pred_relN = "pred_rel"; val pred_setN = "pred_set"; val pred_congN = "pred_cong"; val pred_cong_simpN = "pred_cong_simp"; val rel_GrpN = "rel_Grp"; val rel_comppN = "rel_compp"; val rel_compp_GrpN = "rel_compp_Grp"; val rel_congN = "rel_cong"; val rel_cong_simpN = "rel_cong_simp"; val rel_conversepN = "rel_conversep"; val rel_eqN = "rel_eq"; val rel_eq_onpN = "rel_eq_onp"; val rel_flipN = "rel_flip"; val rel_mapN = "rel_map"; val rel_monoN = "rel_mono"; val rel_mono_strong0N = "rel_mono_strong0"; val rel_mono_strongN = "rel_mono_strong"; val rel_reflN = "rel_refl"; val rel_refl_strongN = "rel_refl_strong"; val rel_reflpN = "rel_reflp"; val rel_sympN = "rel_symp"; val rel_transferN = "rel_transfer"; val rel_transpN = "rel_transp"; val set_bdN = "set_bd"; val set_map0N = "set_map0"; val set_mapN = "set_map"; val set_transferN = "set_transfer"; datatype inline_policy = Dont_Inline | Hardly_Inline | Smart_Inline | Do_Inline; datatype fact_policy = Dont_Note | Note_Some | Note_All; val bnf_internals = Attrib.setup_config_bool \<^binding>\bnf_internals\ (K false); val bnf_timing = Attrib.setup_config_bool \<^binding>\bnf_timing\ (K false); fun user_policy policy ctxt = if Config.get ctxt bnf_internals then Note_All else policy; val smart_max_inline_term_size = 25; (*FUDGE*) fun note_bnf_thms fact_policy qualify0 bnf_b bnf lthy = let val axioms = axioms_of_bnf bnf; val facts = facts_of_bnf bnf; val wits = wits_of_bnf bnf; val qualify = let val qs = Binding.path_of bnf_b; in fold_rev (fn (s, mand) => Binding.qualify mand s) qs #> qualify0 end; fun note_if_note_all (noted0, lthy0) = let val witNs = if length wits = 1 then [witN] else map mk_witN (1 upto length wits); val notes = [(bd_card_orderN, [#bd_card_order axioms]), (bd_cinfiniteN, [#bd_cinfinite axioms]), (bd_regularCardN, [#bd_regularCard axioms]), (bd_Card_orderN, [#bd_Card_order facts]), (bd_CinfiniteN, [#bd_Cinfinite facts]), (bd_CnotzeroN, [#bd_Cnotzero facts]), (collect_set_mapN, [Lazy.force (#collect_set_map facts)]), (in_bdN, [Lazy.force (#in_bd facts)]), (in_monoN, [Lazy.force (#in_mono facts)]), (map_comp0N, [#map_comp0 axioms]), (rel_mono_strong0N, [Lazy.force (#rel_mono_strong0 facts)]), (pred_mono_strong0N, [Lazy.force (#pred_mono_strong0 facts)]), (set_map0N, #set_map0 axioms), (set_bdN, #set_bd axioms)] @ (witNs ~~ wit_thmss_of_bnf bnf) |> map (fn (thmN, thms) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), []), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; fun note_unless_dont_note (noted0, lthy0) = let val notes = [(in_relN, [Lazy.force (#in_rel facts)], []), (inj_mapN, [Lazy.force (#inj_map facts)], []), (inj_map_strongN, [Lazy.force (#inj_map_strong facts)], []), (map_compN, [Lazy.force (#map_comp facts)], []), (map_cong0N, [#map_cong0 axioms], []), (map_congN, [Lazy.force (#map_cong facts)], fundefcong_attrs), (map_cong_simpN, [Lazy.force (#map_cong_simp facts)], []), (map_cong_predN, [Lazy.force (#map_cong_pred facts)], []), (map_idN, [Lazy.force (#map_id facts)], []), (map_id0N, [#map_id0 axioms], []), (map_transferN, [Lazy.force (#map_transfer facts)], []), (map_identN, [Lazy.force (#map_ident facts)], []), (map_ident_strongN, [Lazy.force (#map_ident_strong facts)], []), (pred_monoN, [Lazy.force (#pred_mono facts)], mono_attrs), (pred_mono_strongN, [Lazy.force (#pred_mono_strong facts)], []), (pred_congN, [Lazy.force (#pred_cong facts)], fundefcong_attrs), (pred_cong_simpN, [Lazy.force (#pred_cong_simp facts)], []), (pred_mapN, [Lazy.force (#pred_map facts)], []), (pred_relN, [Lazy.force (#pred_rel facts)], []), (pred_transferN, [Lazy.force (#pred_transfer facts)], []), (pred_TrueN, [Lazy.force (#pred_True facts)], []), (pred_setN, [#pred_set axioms], []), (rel_comppN, [Lazy.force (#rel_OO facts)], []), (rel_compp_GrpN, no_refl [#rel_OO_Grp axioms], []), (rel_conversepN, [Lazy.force (#rel_conversep facts)], []), (rel_eqN, [Lazy.force (#rel_eq facts)], []), (rel_eq_onpN, [Lazy.force (#rel_eq_onp facts)], []), (rel_flipN, [Lazy.force (#rel_flip facts)], []), (rel_GrpN, [Lazy.force (#rel_Grp facts)], []), (rel_mapN, Lazy.force (#rel_map facts), []), (rel_monoN, [Lazy.force (#rel_mono facts)], mono_attrs), (rel_mono_strongN, [Lazy.force (#rel_mono_strong facts)], []), (rel_congN, [Lazy.force (#rel_cong facts)], fundefcong_attrs), (rel_cong_simpN, [Lazy.force (#rel_cong_simp facts)], []), (rel_reflN, [Lazy.force (#rel_refl facts)], []), (rel_refl_strongN, [Lazy.force (#rel_refl_strong facts)], []), (rel_reflpN, [Lazy.force (#rel_reflp facts)], []), (rel_sympN, [Lazy.force (#rel_symp facts)], []), (rel_transpN, [Lazy.force (#rel_transp facts)], []), (rel_transferN, [Lazy.force (#rel_transfer facts)], []), (set_mapN, map Lazy.force (#set_map facts), []), (set_transferN, Lazy.force (#set_transfer facts), [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify (Binding.qualify true (Binding.name_of bnf_b) (Binding.name thmN)), attrs), [(thms, [])])); in Local_Theory.notes notes lthy0 |>> append noted0 end; in ([], lthy) |> fact_policy = Note_All ? note_if_note_all |> fact_policy <> Dont_Note ? note_unless_dont_note |>> (fn [] => bnf | noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun note_bnf_defs bnf lthy = let fun mk_def_binding cst_of = Thm.def_binding (Binding.qualified_name (dest_Const (cst_of bnf) |> fst)); val notes = [(mk_def_binding map_of_bnf, map_def_of_bnf bnf), (mk_def_binding rel_of_bnf, rel_def_of_bnf bnf), (mk_def_binding pred_of_bnf, pred_def_of_bnf bnf)] @ @{map 2} (pair o mk_def_binding o K) (sets_of_bnf bnf) (set_defs_of_bnf bnf) |> map (fn (b, thm) => ((b, []), [([thm], [])])); in lthy |> Local_Theory.notes notes |>> (fn noted => morph_bnf (substitute_noted_thm noted) bnf) end; fun mk_wit_goals zs bs sets (I, wit) = let val xs = map (nth bs) I; fun wit_goal i = let val z = nth zs i; val set_wit = nth sets i $ Term.list_comb (wit, xs); val concl = HOLogic.mk_Trueprop (if member (op =) I i then HOLogic.mk_eq (z, nth bs i) else \<^term>\False\); in fold_rev Logic.all (z :: xs) (Logic.mk_implies (mk_Trueprop_mem (z, set_wit), concl)) end; in map wit_goal (0 upto length sets - 1) end; (* Define new BNFs *) fun define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy = let val live = length set_rhss; val def_qualify = Binding.qualify false (Binding.name_of bnf_b); fun mk_prefix_binding pre = Binding.prefix_name (pre ^ "_") bnf_b; fun maybe_define user_specified (b, rhs) lthy = let val inline = (user_specified orelse fact_policy = Dont_Note) andalso (case const_policy of Dont_Inline => false | Hardly_Inline => Term.is_Free rhs orelse Term.is_Const rhs | Smart_Inline => Term.size_of_term rhs <= smart_max_inline_term_size | Do_Inline => true); in if inline then ((rhs, Drule.reflexive_thm), lthy) else let val b = b () in apfst (apsnd snd) ((if internal then Local_Theory.define_internal else Local_Theory.define) ((b, NoSyn), ((Binding.concealed (Thm.def_binding b), []), rhs)) lthy) end end; val map_bind_def = (fn () => def_qualify (if Binding.is_empty map_b then mk_prefix_binding mapN else map_b), map_rhs); val set_binds_defs = let fun set_name i get_b = (case try (nth set_bs) (i - 1) of SOME b => if Binding.is_empty b then get_b else K b | NONE => get_b) #> def_qualify; val bs = if live = 1 then [set_name 1 (fn () => mk_prefix_binding setN)] else map (fn i => set_name i (fn () => mk_prefix_binding (mk_setN i))) (1 upto live); in bs ~~ set_rhss end; val bd_bind_def = (fn () => def_qualify (mk_prefix_binding bdN), bd_rhs); val (((bnf_map_term, raw_map_def), (bnf_set_terms, raw_set_defs)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> maybe_define true map_bind_def ||>> apfst split_list o fold_map (maybe_define true) set_binds_defs ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val ((bnf_bd_term, raw_bd_def), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> maybe_define true bd_bind_def ||> `Local_Theory.end_nested; val phi' = Proof_Context.export_morphism lthy_old lthy; val bnf_map_def = Morphism.thm phi raw_map_def; val bnf_set_defs = map (Morphism.thm phi) raw_set_defs; val bnf_bd_def = Morphism.thm phi' raw_bd_def; val bnf_map = Morphism.term phi bnf_map_term; (*TODO: handle errors*) (*simple shape analysis of a map function*) val ((alphas, betas), (Calpha, _)) = fastype_of bnf_map |> strip_typeN live |>> map_split dest_funT ||> dest_funT handle TYPE _ => error "Bad map function"; val Calpha_params = map TVar (Term.add_tvarsT Calpha []); val bnf_T = Morphism.typ phi T_rhs; val bad_args = Term.add_tfreesT bnf_T []; val _ = null bad_args orelse error ("Locally fixed type arguments " ^ commas_quote (map (Syntax.string_of_typ no_defs_lthy o TFree) bad_args)); val bnf_sets = map2 (normalize_set Calpha_params) alphas (map (Morphism.term phi) bnf_set_terms); val bnf_bd = Term.subst_TVars (Term.add_tvar_namesT bnf_T [] ~~ Calpha_params) (Morphism.term phi' bnf_bd_term); (*TODO: assert Ds = (TVars of bnf_map) \ (alphas @ betas) as sets*) val deads = (case Ds_opt of NONE => subtract (op =) (alphas @ betas) (map TVar (Term.add_tvars bnf_map [])) | SOME Ds => map (Morphism.typ phi) Ds); (*TODO: further checks of type of bnf_map*) (*TODO: check types of bnf_sets*) (*TODO: check type of bnf_bd*) (*TODO: check type of bnf_rel*) fun mk_bnf_map Ds As' Bs' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As') @ (betas ~~ Bs')) bnf_map; fun mk_bnf_t Ds As' = Term.subst_atomic_types ((deads ~~ Ds) @ (alphas ~~ As')); fun mk_bnf_T Ds As' = Term.typ_subst_atomic ((deads ~~ Ds) @ (alphas ~~ As')); val (((As, Bs), unsorted_Ds), names_lthy) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees (length deads); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val RTs = map2 (curry HOLogic.mk_prodT) As Bs; val pred2RTs = map2 mk_pred2T As Bs; val (Rs, Rs') = names_lthy |> mk_Frees' "R" pred2RTs |> fst; val CA = mk_bnf_T Ds As Calpha; val CR = mk_bnf_T Ds RTs Calpha; val setRs = @{map 3} (fn R => fn T => fn U => HOLogic.Collect_const (HOLogic.mk_prodT (T, U)) $ HOLogic.mk_case_prod R) Rs As Bs; (*Grp (in (Collect (split R1) .. Collect (split Rn))) (map fst .. fst)^--1 OO Grp (in (Collect (split R1) .. Collect (split Rn))) (map snd .. snd)*) val rel_spec = let val map1 = Term.list_comb (mk_bnf_map Ds RTs As, map fst_const RTs); val map2 = Term.list_comb (mk_bnf_map Ds RTs Bs, map snd_const RTs); val bnf_in = mk_in setRs (map (mk_bnf_t Ds RTs) bnf_sets) CR; in mk_rel_compp (mk_conversep (mk_Grp bnf_in map1), mk_Grp bnf_in map2) |> fold_rev Term.absfree Rs' end; val rel_rhs = the_default rel_spec rel_rhs_opt; val rel_bind_def = (fn () => def_qualify (if Binding.is_empty rel_b then mk_prefix_binding relN else rel_b), rel_rhs); val pred_spec = if live = 0 then Term.absdummy (mk_bnf_T Ds As Calpha) \<^term>\True\ else let val sets = map (mk_bnf_t Ds As) bnf_sets; val argTs = map mk_pred1T As; val T = mk_bnf_T Ds As Calpha; val ((Ps, Ps'), x) = lthy |> mk_Frees' "P" argTs ||>> yield_singleton (mk_Frees "x") T |> fst; val conjs = map2 (fn set => fn P => mk_Ball (set $ x) P) sets Ps; in fold_rev Term.absfree Ps' (Term.absfree (dest_Free x) (Library.foldr1 HOLogic.mk_conj conjs)) end; val pred_rhs = the_default pred_spec pred_rhs_opt; val pred_bind_def = (fn () => def_qualify (if Binding.is_empty pred_b then mk_prefix_binding predN else pred_b), pred_rhs); val wit_rhss = if null wit_rhss then [fold_rev Term.absdummy As (Term.list_comb (mk_bnf_map Ds As As, map2 (fn T => fn i => Term.absdummy T (Bound i)) As (live downto 1)) $ Const (\<^const_name>\undefined\, CA))] else wit_rhss; val nwits = length wit_rhss; val wit_binds_defs = let val bs = if nwits = 1 then [fn () => def_qualify (mk_prefix_binding witN)] else map (fn i => fn () => def_qualify (mk_prefix_binding (mk_witN i))) (1 upto nwits); in bs ~~ wit_rhss end; val ((((bnf_rel_term, raw_rel_def), (bnf_pred_term, raw_pred_def)), (bnf_wit_terms, raw_wit_defs)), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> maybe_define (is_some rel_rhs_opt) rel_bind_def ||>> maybe_define (is_some pred_rhs_opt) pred_bind_def ||>> apfst split_list o fold_map (maybe_define (not (null wit_rhss))) wit_binds_defs ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val bnf_rel_def = Morphism.thm phi raw_rel_def; val bnf_rel = Morphism.term phi bnf_rel_term; fun mk_bnf_rel Ds As' Bs' = normalize_rel lthy (map2 mk_pred2T As' Bs') (mk_bnf_T Ds As' Calpha) (mk_bnf_T Ds Bs' Calpha) bnf_rel; val bnf_pred_def = Morphism.thm phi raw_pred_def; val bnf_pred = Morphism.term phi bnf_pred_term; fun mk_bnf_pred Ds As' = normalize_pred lthy (map mk_pred1T As') (mk_bnf_T Ds As' Calpha) bnf_pred; val bnf_wit_defs = map (Morphism.thm phi) raw_wit_defs; val bnf_wits = map (normalize_wit Calpha_params Calpha alphas o Morphism.term phi) bnf_wit_terms; fun mk_rel_spec Ds' As' Bs' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As') @ (Bs ~~ Bs')) rel_spec; fun mk_pred_spec Ds' As' = Term.subst_atomic_types ((Ds ~~ Ds') @ (As ~~ As')) pred_spec; in (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map, mk_bnf_t, mk_bnf_T, mk_bnf_rel, mk_bnf_pred, mk_rel_spec, mk_pred_spec)), lthy) end; fun prepare_def const_policy mk_fact_policy internal qualify prep_typ prep_term Ds_opt map_b rel_b pred_b set_bs (((((((raw_bnf_b, raw_bnf_T), raw_map), raw_sets), raw_bd), raw_wits), raw_rel_opt), raw_pred_opt) no_defs_lthy = let val fact_policy = mk_fact_policy no_defs_lthy; val bnf_b = qualify raw_bnf_b; val live = length raw_sets; val T_rhs = prep_typ no_defs_lthy raw_bnf_T; val map_rhs = prep_term no_defs_lthy raw_map; val set_rhss = map (prep_term no_defs_lthy) raw_sets; val bd_rhs = prep_term no_defs_lthy raw_bd; val wit_rhss = map (prep_term no_defs_lthy) raw_wits; val rel_rhs_opt = Option.map (prep_term no_defs_lthy) raw_rel_opt; val pred_rhs_opt = Option.map (prep_term no_defs_lthy) raw_pred_opt; fun err T = error ("Trying to register the type " ^ quote (Syntax.string_of_typ no_defs_lthy T) ^ " as unnamed BNF"); val (bnf_b, key) = if Binding.is_empty bnf_b then (case T_rhs of Type (C, Ts) => if forall (can dest_TFree) Ts andalso not (has_duplicates (op =) Ts) then (Binding.qualified_name C, C) else err T_rhs | T => err T) else (bnf_b, Local_Theory.full_name no_defs_lthy bnf_b); val (((alphas, betas, deads, Calpha), (bnf_map, bnf_sets, bnf_bd, bnf_wits, bnf_rel, bnf_pred), (bnf_map_def, bnf_set_defs, bnf_bd_def, bnf_wit_defs, bnf_rel_def, bnf_pred_def), (mk_bnf_map_Ds, mk_bnf_t_Ds, mk_bnf_T_Ds, _, _, mk_rel_spec, mk_pred_spec)), lthy) = define_bnf_consts const_policy fact_policy internal Ds_opt map_b rel_b pred_b set_bs (((((((bnf_b, T_rhs), map_rhs), set_rhss), bd_rhs), wit_rhss), rel_rhs_opt), pred_rhs_opt) no_defs_lthy; val dead = length deads; val (((((((As', Bs'), Cs), unsorted_Ds), Es), B1Ts), B2Ts), (Ts, T)) = lthy |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees dead ||>> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees live ||> fst o mk_TFrees 1 ||> the_single ||> `(replicate live); val Ds = map2 (resort_tfree_or_tvar o Type.sort_of_atyp) deads unsorted_Ds; val mk_bnf_map = mk_bnf_map_Ds Ds; val mk_bnf_t = mk_bnf_t_Ds Ds; val mk_bnf_T = mk_bnf_T_Ds Ds; val pred1PTs = map mk_pred1T As'; val pred1QTs = map mk_pred1T Bs'; val pred2RTs = map2 mk_pred2T As' Bs'; val pred2RTsAsCs = map2 mk_pred2T As' Cs; val pred2RTsBsCs = map2 mk_pred2T Bs' Cs; val pred2RTsBsEs = map2 mk_pred2T Bs' Es; val pred2RTsCsBs = map2 mk_pred2T Cs Bs'; val pred2RTsCsEs = map2 mk_pred2T Cs Es; val pred2RT's = map2 mk_pred2T Bs' As'; val self_pred2RTs = map2 mk_pred2T As' As'; val transfer_domRTs = map2 mk_pred2T As' B1Ts; val transfer_ranRTs = map2 mk_pred2T Bs' B2Ts; val CA' = mk_bnf_T As' Calpha; val CB' = mk_bnf_T Bs' Calpha; val CC' = mk_bnf_T Cs Calpha; val CE' = mk_bnf_T Es Calpha; val CB1 = mk_bnf_T B1Ts Calpha; val CB2 = mk_bnf_T B2Ts Calpha; val bnf_map_AsAs = mk_bnf_map As' As'; val bnf_map_AsBs = mk_bnf_map As' Bs'; val bnf_map_AsCs = mk_bnf_map As' Cs; val bnf_map_BsCs = mk_bnf_map Bs' Cs; val bnf_sets_As = map (mk_bnf_t As') bnf_sets; val bnf_sets_Bs = map (mk_bnf_t Bs') bnf_sets; val bnf_bd_As = mk_bnf_t As' bnf_bd; fun mk_bnf_rel RTs CA CB = normalize_rel lthy RTs CA CB bnf_rel; fun mk_bnf_pred PTs CA = normalize_pred lthy PTs CA bnf_pred; val ((((((((((((((((((((((((((fs, fs'), gs), hs), is), x), x'), y), y'), zs), zs'), ys), As), As_copy), bs), (Ps, Ps')), Ps_copy), Qs), Rs), Rs_copy), Ss), S_AsCs), S_CsBs), S_BsEs), transfer_domRs), transfer_ranRs), _) = lthy |> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "f" (map2 (curry op -->) As' Bs') ||>> mk_Frees "g" (map2 (curry op -->) Bs' Cs) ||>> mk_Frees "h" (map2 (curry op -->) As' Ts) ||>> mk_Frees "i" (map2 (curry op -->) As' Cs) ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "x") CA' ||>> yield_singleton (mk_Frees "y") CB' ||>> yield_singleton (mk_Frees "y") CB' ||>> mk_Frees "z" As' ||>> mk_Frees "z" As' ||>> mk_Frees "y" Bs' ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "A" (map HOLogic.mk_setT As') ||>> mk_Frees "b" As' ||>> mk_Frees' "P" pred1PTs ||>> mk_Frees "P" pred1PTs ||>> mk_Frees "Q" pred1QTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "R" pred2RTs ||>> mk_Frees "S" pred2RTsBsCs ||>> mk_Frees "S" pred2RTsAsCs ||>> mk_Frees "S" pred2RTsCsBs ||>> mk_Frees "S" pred2RTsBsEs ||>> mk_Frees "R" transfer_domRTs ||>> mk_Frees "S" transfer_ranRTs; val fs_copy = map2 (retype_const_or_free o fastype_of) fs gs; val x_copy = retype_const_or_free CA' y'; val y_copy = retype_const_or_free CB' x'; val rel = mk_bnf_rel pred2RTs CA' CB'; val pred = mk_bnf_pred pred1PTs CA'; val pred' = mk_bnf_pred pred1QTs CB'; val relCsEs = mk_bnf_rel pred2RTsCsEs CC' CE'; val relAsAs = mk_bnf_rel self_pred2RTs CA' CA'; val bnf_wit_As = map (apsnd (mk_bnf_t As')) bnf_wits; val map_id0_goal = let val bnf_map_app_id = Term.list_comb (bnf_map_AsAs, map HOLogic.id_const As') in mk_Trueprop_eq (bnf_map_app_id, HOLogic.id_const CA') end; val map_comp0_goal = let val bnf_map_app_comp = Term.list_comb (bnf_map_AsCs, map2 (curry HOLogic.mk_comp) gs fs); val comp_bnf_map_app = HOLogic.mk_comp (Term.list_comb (bnf_map_BsCs, gs), Term.list_comb (bnf_map_AsBs, fs)); in fold_rev Logic.all (fs @ gs) (mk_Trueprop_eq (bnf_map_app_comp, comp_bnf_map_app)) end; fun mk_map_cong_prem mk_implies x z set f f_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (f $ z, f_copy $ z))); val map_cong0_goal = let val prems = @{map 4} (mk_map_cong_prem Logic.mk_implies x) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x); in fold_rev Logic.all (x :: fs @ fs_copy) (Logic.list_implies (prems, eq)) end; val set_map0s_goal = let fun mk_goal setA setB f = let val set_comp_map = HOLogic.mk_comp (setB, Term.list_comb (bnf_map_AsBs, fs)); val image_comp_set = HOLogic.mk_comp (mk_image f, setA); in fold_rev Logic.all fs (mk_Trueprop_eq (set_comp_map, image_comp_set)) end; in @{map 3} mk_goal bnf_sets_As bnf_sets_Bs fs end; val card_order_bd_goal = HOLogic.mk_Trueprop (mk_card_order bnf_bd_As); val cinfinite_bd_goal = HOLogic.mk_Trueprop (mk_cinfinite bnf_bd_As); val regularCard_bd_goal = HOLogic.mk_Trueprop (mk_regularCard bnf_bd_As); val set_bds_goal = let fun mk_goal set = Logic.all x (HOLogic.mk_Trueprop (mk_ordLess (mk_card_of (set $ x)) bnf_bd_As)); in map mk_goal bnf_sets_As end; val relAsCs = mk_bnf_rel pred2RTsAsCs CA' CC'; val relBsCs = mk_bnf_rel pred2RTsBsCs CB' CC'; val relCsBs = mk_bnf_rel pred2RTsCsBs CC' CB'; val rel_OO_lhs = Term.list_comb (relAsCs, map2 (curry mk_rel_compp) Rs Ss); val rel_OO_rhs = mk_rel_compp (Term.list_comb (rel, Rs), Term.list_comb (relBsCs, Ss)); val le_rel_OO_goal = fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_rhs rel_OO_lhs)); val rel_OO_Grp_goal = fold_rev Logic.all Rs (mk_Trueprop_eq (Term.list_comb (rel, Rs), Term.list_comb (mk_rel_spec Ds As' Bs', Rs))); val pred_set_goal = fold_rev Logic.all Ps (mk_Trueprop_eq (Term.list_comb (pred, Ps), Term.list_comb (mk_pred_spec Ds As', Ps))); val goals = zip_axioms map_id0_goal map_comp0_goal map_cong0_goal set_map0s_goal card_order_bd_goal cinfinite_bd_goal regularCard_bd_goal set_bds_goal le_rel_OO_goal rel_OO_Grp_goal pred_set_goal; val mk_wit_goals = mk_wit_goals bs zs bnf_sets_As; fun triv_wit_tac ctxt = mk_trivial_wit_tac ctxt bnf_wit_defs; val wit_goalss = (if null raw_wits then SOME triv_wit_tac else NONE, map mk_wit_goals bnf_wit_As); fun after_qed mk_wit_thms thms lthy = let val (axioms, nontriv_wit_thms) = apfst (mk_axioms live) (chop (length goals) thms); val bd_Card_order = #bd_card_order axioms RS @{thm conjunct2[OF card_order_on_Card_order]}; val bd_Cinfinite = @{thm conjI} OF [#bd_cinfinite axioms, bd_Card_order]; val bd_Cnotzero = bd_Cinfinite RS @{thm Cinfinite_Cnotzero}; fun mk_collect_set_map () = let val defT = mk_bnf_T Ts Calpha --> HOLogic.mk_setT T; val collect_map = HOLogic.mk_comp (mk_collect (map (mk_bnf_t Ts) bnf_sets) defT, Term.list_comb (mk_bnf_map As' Ts, hs)); val image_collect = mk_collect (map2 (fn h => fn set => HOLogic.mk_comp (mk_image h, set)) hs bnf_sets_As) defT; (*collect {set1 ... setm} o map f1 ... fm = collect {f1` o set1 ... fm` o setm}*) val goal = fold_rev Logic.all hs (mk_Trueprop_eq (collect_map, image_collect)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_collect_set_map_tac ctxt (#set_map0 axioms)) |> Thm.close_derivation \<^here> end; val collect_set_map = Lazy.lazy mk_collect_set_map; fun mk_in_mono () = let val prems_mono = map2 (HOLogic.mk_Trueprop oo mk_leq) As As_copy; val in_mono_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_mono, HOLogic.mk_Trueprop (mk_leq (mk_in As bnf_sets_As CA') (mk_in As_copy bnf_sets_As CA')))); in Goal.prove_sorry lthy [] [] in_mono_goal (fn {context = ctxt, prems = _} => mk_in_mono_tac ctxt live) |> Thm.close_derivation \<^here> end; val in_mono = Lazy.lazy mk_in_mono; fun mk_in_cong () = let val prems_cong = map2 (curry mk_Trueprop_eq) As As_copy; val in_cong_goal = fold_rev Logic.all (As @ As_copy) (Logic.list_implies (prems_cong, mk_Trueprop_eq (mk_in As bnf_sets_As CA', mk_in As_copy bnf_sets_As CA'))); in Goal.prove_sorry lthy [] [] in_cong_goal (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val in_cong = Lazy.lazy mk_in_cong; val map_id = Lazy.lazy (fn () => mk_map_id (#map_id0 axioms)); val map_ident0 = Lazy.lazy (fn () => mk_map_ident lthy (#map_id0 axioms)); val map_ident = Lazy.lazy (fn () => mk_map_ident lthy (Lazy.force map_id)); val map_ident_strong = Lazy.lazy (fn () => mk_map_ident_strong lthy (#map_cong0 axioms) (Lazy.force map_id)); val map_comp = Lazy.lazy (fn () => mk_map_comp (#map_comp0 axioms)); fun mk_map_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_map_cong_prem mk_implies x_copy) zs bnf_sets_As fs fs_copy; val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies (prem0 :: prems, eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt @{thms simp_implies_def} THEN mk_map_cong_tac ctxt (#map_cong0 axioms)) |> Thm.close_derivation \<^here> end; val map_cong = Lazy.lazy (mk_map_cong Logic.mk_implies); val map_cong_simp = Lazy.lazy (mk_map_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_inj_map () = let val prems = map (HOLogic.mk_Trueprop o mk_inj) fs; val concl = HOLogic.mk_Trueprop (mk_inj (Term.list_comb (bnf_map_AsBs, fs))); val goal = fold_rev Logic.all fs (Logic.list_implies (prems, concl)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_tac ctxt live (Lazy.force map_id) (Lazy.force map_comp) (#map_cong0 axioms) (Lazy.force map_cong)) |> Thm.close_derivation \<^here> end; val inj_map = Lazy.lazy mk_inj_map; val set_map = map (fn thm => Lazy.lazy (fn () => mk_set_map thm)) (#set_map0 axioms); val wit_thms = if null nontriv_wit_thms then mk_wit_thms (map Lazy.force set_map) else nontriv_wit_thms; fun mk_in_bd () = let val bdT = fst (dest_relT (fastype_of bnf_bd_As)); val bdTs = replicate live bdT; val bd_bnfT = mk_bnf_T bdTs Calpha; val surj_imp_ordLeq_inst = (if live = 0 then TrueI else let val ranTs = map (fn AT => mk_sumT (AT, HOLogic.unitT)) As'; val funTs = map (fn T => bdT --> T) ranTs; val ran_bnfT = mk_bnf_T ranTs Calpha; val (revTs, Ts) = `rev (bd_bnfT :: funTs); val cTs = map (SOME o Thm.ctyp_of lthy) [ran_bnfT, Library.foldr1 HOLogic.mk_prodT Ts]; val tinst = fold (fn T => fn t => HOLogic.mk_case_prod (Term.absdummy T t)) (tl revTs) (Term.absdummy (hd revTs) (Term.list_comb (mk_bnf_map bdTs ranTs, map Bound (live - 1 downto 0)) $ Bound live)); val cts = [NONE, SOME (Thm.cterm_of lthy tinst)]; in Thm.instantiate' cTs cts @{thm surj_imp_ordLeq} end); val bd = mk_cexp (if live = 0 then ctwo else mk_csum (Library.foldr1 (uncurry mk_csum) (map mk_card_of As)) ctwo) (mk_csum bnf_bd_As (mk_card_of (HOLogic.mk_UNIV bd_bnfT))); val in_bd_goal = fold_rev Logic.all As (HOLogic.mk_Trueprop (mk_ordLeq (mk_card_of (mk_in As bnf_sets_As CA')) bd)); val weak_set_bds = map (fn thm => @{thm ordLess_imp_ordLeq} OF [thm]) (#set_bd axioms); in Goal.prove_sorry lthy [] [] in_bd_goal (fn {context = ctxt, prems = _} => mk_in_bd_tac ctxt live surj_imp_ordLeq_inst (Lazy.force map_comp) (Lazy.force map_id) (#map_cong0 axioms) (map Lazy.force set_map) weak_set_bds (#bd_card_order axioms) bd_Card_order bd_Cinfinite bd_Cnotzero) |> Thm.close_derivation \<^here> end; val in_bd = Lazy.lazy mk_in_bd; val rel_OO_Grp = #rel_OO_Grp axioms; val rel_OO_Grps = no_refl [rel_OO_Grp]; fun mk_rel_Grp () = let val lhs = Term.list_comb (rel, map2 mk_Grp As fs); val rhs = mk_Grp (mk_in As bnf_sets_As CA') (Term.list_comb (bnf_map_AsBs, fs)); val goal = fold_rev Logic.all (As @ fs) (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_Grp_tac ctxt rel_OO_Grps (#map_id0 axioms) (#map_cong0 axioms) (Lazy.force map_id) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_Grp = Lazy.lazy mk_rel_Grp; fun mk_rel_prems f = map2 (HOLogic.mk_Trueprop oo f) Rs Rs_copy; fun mk_rel_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (rel, Rs), Term.list_comb (rel, Rs_copy))); fun mk_rel_mono () = let val mono_prems = mk_rel_prems mk_leq; val mono_concl = mk_rel_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_tac ctxt rel_OO_Grps (Lazy.force in_mono)) |> Thm.close_derivation \<^here> end; fun mk_rel_cong0 () = let val cong_prems = mk_rel_prems (curry HOLogic.mk_eq); val cong_concl = mk_rel_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Rs_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val rel_mono = Lazy.lazy mk_rel_mono; val rel_cong0 = Lazy.lazy mk_rel_cong0; fun mk_rel_eq () = Goal.prove_sorry lthy [] [] (mk_Trueprop_eq (Term.list_comb (relAsAs, map HOLogic.eq_const As'), HOLogic.eq_const CA')) (fn {context = ctxt, prems = _} => mk_rel_eq_tac ctxt live (Lazy.force rel_Grp) (Lazy.force rel_cong0) (#map_id0 axioms)) |> Thm.close_derivation \<^here>; val rel_eq = Lazy.lazy mk_rel_eq; fun mk_rel_conversep () = let val relBsAs = mk_bnf_rel pred2RT's CB' CA'; val lhs = Term.list_comb (relBsAs, map mk_conversep Rs); val rhs = mk_conversep (Term.list_comb (rel, Rs)); val le_goal = fold_rev Logic.all Rs (HOLogic.mk_Trueprop (mk_leq lhs rhs)); val le_thm = Goal.prove_sorry lthy [] [] le_goal (fn {context = ctxt, prems = _} => mk_rel_conversep_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> val goal = fold_rev Logic.all Rs (mk_Trueprop_eq (lhs, rhs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_conversep_tac ctxt le_thm (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val rel_conversep = Lazy.lazy mk_rel_conversep; fun mk_rel_OO () = Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Rs @ Ss) (HOLogic.mk_Trueprop (mk_leq rel_OO_lhs rel_OO_rhs))) (fn {context = ctxt, prems = _} => mk_rel_OO_le_tac ctxt rel_OO_Grps (Lazy.force rel_eq) (#map_cong0 axioms) (Lazy.force map_comp) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> (fn thm => @{thm antisym} OF [thm, #le_rel_OO axioms]); val rel_OO = Lazy.lazy mk_rel_OO; fun mk_in_rel () = trans OF [rel_OO_Grp, @{thm OO_Grp_alt}] RS @{thm predicate2_eqD}; val in_rel = Lazy.lazy mk_in_rel; fun mk_rel_flip () = unfold_thms lthy @{thms conversep_iff} (Lazy.force rel_conversep RS @{thm predicate2_eqD}); val rel_flip = Lazy.lazy mk_rel_flip; fun mk_rel_mono_strong0 () = let fun mk_prem setA setB R S a b = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (mk_Ball (setB $ y) (Term.absfree (dest_Free b) (HOLogic.mk_imp (R $ a $ b, S $ a $ b)))))); val prems = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs) $ x $ y) :: @{map 6} mk_prem bnf_sets_As bnf_sets_Bs Rs Rs_copy zs ys; val concl = HOLogic.mk_Trueprop (Term.list_comb (rel, Rs_copy) $ x $ y); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: y :: Rs @ Rs_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_rel_mono_strong0_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> end; val rel_mono_strong0 = Lazy.lazy mk_rel_mono_strong0; val rel_mono_strong = Lazy.map (Object_Logic.rulify lthy) rel_mono_strong0; fun mk_rel_cong_prem mk_implies x x' z z' set set' R R_copy = Logic.all z (Logic.all z' (mk_implies (mk_Trueprop_mem (z, set $ x), mk_implies (mk_Trueprop_mem (z', set' $ x'), mk_Trueprop_eq (R $ z $ z', R_copy $ z $ z'))))); fun mk_rel_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prem1 = mk_Trueprop_eq (y, y_copy); val prems = @{map 6} (mk_rel_cong_prem mk_implies x_copy y_copy) zs ys bnf_sets_As bnf_sets_Bs Rs Rs_copy; val eq = mk_Trueprop_eq (Term.list_comb (rel, Rs) $ x $ y, Term.list_comb (rel, Rs_copy) $ x_copy $ y_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prem1 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prem1 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 2 prems) (Lazy.force rel_mono_strong))) |> Thm.close_derivation \<^here> end; val rel_cong = Lazy.lazy (mk_rel_cong Logic.mk_implies); val rel_cong_simp = Lazy.lazy (mk_rel_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_pred_prems f = map2 (HOLogic.mk_Trueprop oo f) Ps Ps_copy; fun mk_pred_concl f = HOLogic.mk_Trueprop (f (Term.list_comb (pred, Ps), Term.list_comb (pred, Ps_copy))); fun mk_pred_cong0 () = let val cong_prems = mk_pred_prems (curry HOLogic.mk_eq); val cong_concl = mk_pred_concl HOLogic.mk_eq; in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (cong_prems, cong_concl))) (fn {context = ctxt, prems = _} => (TRY o hyp_subst_tac ctxt THEN' rtac ctxt refl) 1) |> Thm.close_derivation \<^here> end; val pred_cong0 = Lazy.lazy mk_pred_cong0; fun mk_rel_eq_onp () = let val lhs = Term.list_comb (relAsAs, map mk_eq_onp Ps); val rhs = mk_eq_onp (Term.list_comb (pred, Ps)); in Goal.prove_sorry lthy (map fst Ps') [] (mk_Trueprop_eq (lhs, rhs)) (fn {context = ctxt, prems = _} => mk_rel_eq_onp_tac ctxt (#pred_set axioms) (#map_id0 axioms) (Lazy.force rel_Grp)) |> Thm.close_derivation \<^here> end; val rel_eq_onp = Lazy.lazy mk_rel_eq_onp; val pred_rel = Lazy.map (fn thm => thm RS sym RS @{thm eq_onp_eqD}) rel_eq_onp; fun mk_pred_mono_strong0 () = let fun mk_prem setA P Q a = HOLogic.mk_Trueprop (mk_Ball (setA $ x) (Term.absfree (dest_Free a) (HOLogic.mk_imp (P $ a, Q $ a)))); val prems = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps) $ x) :: @{map 4} mk_prem bnf_sets_As Ps Ps_copy zs; val concl = HOLogic.mk_Trueprop (Term.list_comb (pred, Ps_copy) $ x); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (x :: Ps @ Ps_copy) (Logic.list_implies (prems, concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_strong0_tac ctxt (Lazy.force pred_rel) (Lazy.force rel_mono_strong0)) |> Thm.close_derivation \<^here> end; val pred_mono_strong0 = Lazy.lazy mk_pred_mono_strong0; val pred_mono_strong = Lazy.map (Object_Logic.rulify lthy) pred_mono_strong0; fun mk_pred_mono () = let val mono_prems = mk_pred_prems mk_leq; val mono_concl = mk_pred_concl (uncurry mk_leq); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (Ps @ Ps_copy) (Logic.list_implies (mono_prems, mono_concl))) (fn {context = ctxt, prems = _} => mk_pred_mono_tac ctxt (Lazy.force rel_eq_onp) (Lazy.force rel_mono)) |> Thm.close_derivation \<^here> end; val pred_mono = Lazy.lazy mk_pred_mono; fun mk_pred_cong_prem mk_implies x z set P P_copy = Logic.all z (mk_implies (mk_Trueprop_mem (z, set $ x), mk_Trueprop_eq (P $ z, P_copy $ z))); fun mk_pred_cong mk_implies () = let val prem0 = mk_Trueprop_eq (x, x_copy); val prems = @{map 4} (mk_pred_cong_prem mk_implies x_copy) zs bnf_sets_As Ps Ps_copy; val eq = mk_Trueprop_eq (Term.list_comb (pred, Ps) $ x, Term.list_comb (pred, Ps_copy) $ x_copy); in fold (Variable.add_free_names lthy) (eq :: prem0 :: prems) [] |> (fn vars => Goal.prove_sorry lthy vars (prem0 :: prems) eq (fn {context = ctxt, prems} => mk_rel_cong_tac ctxt (chop 1 prems) (Lazy.force pred_mono_strong))) |> Thm.close_derivation \<^here> end; val pred_cong = Lazy.lazy (mk_pred_cong Logic.mk_implies); val pred_cong_simp = Lazy.lazy (mk_pred_cong (fn (a, b) => \<^term>\simp_implies\ $ a $ b)); fun mk_map_cong_pred () = let val prem0 = mk_Trueprop_eq (x, x_copy); fun mk_eq f g z = Term.absfree (dest_Free z) (HOLogic.mk_eq (f $ z, g $ z)); val prem = HOLogic.mk_Trueprop (Term.list_comb (pred, @{map 3} mk_eq fs fs_copy zs) $ x_copy); val eq = mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs_copy) $ x_copy); val goal = fold_rev Logic.all (x :: x_copy :: fs @ fs_copy) (Logic.list_implies ([prem0, prem], eq)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [#pred_set axioms] THEN HEADGOAL (EVERY' [REPEAT_DETERM o etac ctxt conjE, etac ctxt (Lazy.force map_cong) THEN_ALL_NEW (etac ctxt @{thm bspec} THEN' assume_tac ctxt)])) |> Thm.close_derivation \<^here> end; val map_cong_pred = Lazy.lazy mk_map_cong_pred; fun mk_rel_map () = let fun mk_goal lhs rhs = fold_rev Logic.all ([x, y] @ S_CsBs @ S_AsCs @ is @ gs) (mk_Trueprop_eq (lhs, rhs)); val lhss = [Term.list_comb (relCsBs, S_CsBs) $ (Term.list_comb (bnf_map_AsCs, is) $ x) $ y, Term.list_comb (relAsCs, S_AsCs) $ x $ (Term.list_comb (bnf_map_BsCs, gs) $ y)]; val rhss = [Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p f (HOLogic.id_const T) $ P) is S_CsBs Bs') $ x $ y, Term.list_comb (rel, @{map 3} (fn f => fn P => fn T => mk_vimage2p (HOLogic.id_const T) f $ P) gs S_AsCs As') $ x $ y]; val goals = map2 mk_goal lhss rhss; in goals |> map (fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_map0_tac ctxt live (Lazy.force rel_OO) (Lazy.force rel_conversep) (Lazy.force rel_Grp) (Lazy.force map_id))) |> map (unfold_thms lthy @{thms vimage2p_def[of id, simplified id_apply] vimage2p_def[of _ id, simplified id_apply]}) |> map (Thm.close_derivation \<^here>) end; val rel_map = Lazy.lazy mk_rel_map; fun mk_rel_refl () = @{thm ge_eq_refl[OF ord_eq_le_trans]} OF [Lazy.force rel_eq RS sym, Lazy.force rel_mono OF (replicate live @{thm refl_ge_eq})]; val rel_refl = Lazy.lazy mk_rel_refl; fun mk_rel_refl_strong () = (rule_by_tactic lthy (ALLGOALS (Object_Logic.full_atomize_tac lthy)) ((Lazy.force rel_eq RS @{thm predicate2_eqD}) RS @{thm iffD2[OF _ refl]} RS Lazy.force rel_mono_strong)) OF (replicate live @{thm diag_imp_eq_le}) val rel_refl_strong = Lazy.lazy mk_rel_refl_strong; fun mk_rel_preserves mk_prop prop_conv_thm thm () = let val Rs = map2 retype_const_or_free self_pred2RTs Rs; val prems = map (HOLogic.mk_Trueprop o mk_prop) Rs; val goal = HOLogic.mk_Trueprop (mk_prop (Term.list_comb (relAsAs, Rs))); val vars = fold (Variable.add_free_names lthy) (goal :: prems) []; in Goal.prove_sorry lthy vars [] (Logic.list_implies (prems, goal)) (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt [prop_conv_thm] THEN HEADGOAL (rtac ctxt (Lazy.force thm RS sym RS @{thm ord_eq_le_trans}) THEN' rtac ctxt (Lazy.force rel_mono) THEN_ALL_NEW assume_tac ctxt)) |> Thm.close_derivation \<^here> end; val rel_reflp = Lazy.lazy (mk_rel_preserves mk_reflp @{thm reflp_eq} rel_eq); val rel_symp = Lazy.lazy (mk_rel_preserves mk_symp @{thm symp_conversep} rel_conversep); val rel_transp = Lazy.lazy (mk_rel_preserves mk_transp @{thm transp_relcompp} rel_OO); fun mk_pred_True () = let val lhs = Term.list_comb (pred, map (fn T => absdummy T \<^term>\True\) As'); val rhs = absdummy CA' \<^term>\True\; val goal = mk_Trueprop_eq (lhs, rhs); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (EVERY' (map (rtac ctxt) [ext, Lazy.force pred_rel RS trans, Lazy.force rel_cong0 RS fun_cong RS fun_cong RS trans OF replicate live @{thm eq_onp_True}, Lazy.force rel_eq RS fun_cong RS fun_cong RS trans, @{thm eqTrueI[OF refl]}]))) |> Thm.close_derivation \<^here> end; val pred_True = Lazy.lazy mk_pred_True; fun mk_pred_map () = let val lhs = Term.list_comb (pred', Qs) $ (Term.list_comb (bnf_map_AsBs, fs) $ x); val rhs = Term.list_comb (pred, @{map 2} (curry HOLogic.mk_comp) Qs fs) $ x; val goal = mk_Trueprop_eq (lhs, rhs); val vars = Variable.add_free_names lthy goal []; val pred_set = #pred_set axioms RS fun_cong RS sym; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt (pred_set RSN (2, pred_set RSN (2, box_equals)))) THEN unfold_thms_tac ctxt (@{thms Ball_image_comp ball_empty} @ map Lazy.force set_map) THEN HEADGOAL (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; val pred_map = Lazy.lazy mk_pred_map; fun mk_map_transfer () = let val rels = map2 mk_rel_fun transfer_domRs transfer_ranRs; val rel = mk_rel_fun (Term.list_comb (mk_bnf_rel transfer_domRTs CA' CB1, transfer_domRs)) (Term.list_comb (mk_bnf_rel transfer_ranRTs CB' CB2, transfer_ranRs)); val concl = HOLogic.mk_Trueprop (fold_rev mk_rel_fun rels rel $ bnf_map_AsBs $ mk_bnf_map B1Ts B2Ts); in Goal.prove_sorry lthy [] [] (fold_rev Logic.all (transfer_domRs @ transfer_ranRs) concl) (fn {context = ctxt, prems = _} => mk_map_transfer_tac ctxt (Lazy.force rel_mono) (Lazy.force in_rel) (map Lazy.force set_map) (#map_cong0 axioms) (Lazy.force map_comp)) |> Thm.close_derivation \<^here> end; val map_transfer = Lazy.lazy mk_map_transfer; fun mk_pred_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map (fn T => mk_rel_fun T iff) Rs; val prem_elems = mk_rel_fun (Term.list_comb (rel, Rs)) iff; val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ pred $ pred'); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_pred_transfer_tac ctxt live (Lazy.force in_rel) (Lazy.force pred_map) (Lazy.force pred_cong)) |> Thm.close_derivation \<^here> end; val pred_transfer = Lazy.lazy mk_pred_transfer; fun mk_rel_transfer () = let val iff = HOLogic.eq_const HOLogic.boolT; val prem_rels = map2 (fn T1 => fn T2 => mk_rel_fun T1 (mk_rel_fun T2 iff)) S_AsCs S_BsEs; val prem_elems = mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsAsCs CA' CC', S_AsCs)) (mk_rel_fun (Term.list_comb (mk_bnf_rel pred2RTsBsEs CB' CE', S_BsEs)) iff); val goal = HOLogic.mk_Trueprop (fold_rev mk_rel_fun prem_rels prem_elems $ rel $ relCsEs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_transfer_tac ctxt (Lazy.force in_rel) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val rel_transfer = Lazy.lazy mk_rel_transfer; fun mk_set_transfer () = let val rel_sets = map2 (fn A => fn B => mk_rel 1 [A] [B] \<^term>\rel_set\) As' Bs'; val rel_Rs = Term.list_comb (rel, Rs); val goals = @{map 4} (fn R => fn rel_set => fn setA => fn setB => HOLogic.mk_Trueprop (mk_rel_fun rel_Rs (rel_set $ R) $ setA $ setB)) Rs rel_sets bnf_sets_As bnf_sets_Bs; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_transfer_tac ctxt (Lazy.force in_rel) (map Lazy.force set_map)) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_transfer = Lazy.lazy mk_set_transfer; fun mk_inj_map_strong () = let val assms = @{map 5} (fn setA => fn z => fn f => fn z' => fn f' => fold_rev Logic.all [z, z'] (Logic.mk_implies (mk_Trueprop_mem (z, setA $ x), Logic.mk_implies (mk_Trueprop_mem (z', setA $ x'), Logic.mk_implies (mk_Trueprop_eq (f $ z, f' $ z'), mk_Trueprop_eq (z, z')))))) bnf_sets_As zs fs zs' fs'; val concl = Logic.mk_implies (mk_Trueprop_eq (Term.list_comb (bnf_map_AsBs, fs) $ x, Term.list_comb (bnf_map_AsBs, fs') $ x'), mk_Trueprop_eq (x, x')); val goal = fold_rev Logic.all (x :: x' :: fs @ fs') (fold_rev (curry Logic.mk_implies) assms concl); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_inj_map_strong_tac ctxt (Lazy.force rel_eq) (Lazy.force rel_map) (Lazy.force rel_mono_strong)) |> Thm.close_derivation \<^here> end; val inj_map_strong = Lazy.lazy mk_inj_map_strong; val defs = mk_defs bnf_map_def bnf_set_defs bnf_rel_def bnf_pred_def; val facts = mk_facts bd_Card_order bd_Cinfinite bd_Cnotzero collect_set_map in_bd in_cong in_mono in_rel inj_map inj_map_strong map_comp map_cong map_cong_simp map_cong_pred map_id map_ident0 map_ident map_ident_strong map_transfer rel_eq rel_flip set_map rel_cong0 rel_cong rel_cong_simp rel_map rel_mono rel_mono_strong0 rel_mono_strong set_transfer rel_Grp rel_conversep rel_OO rel_refl rel_refl_strong rel_reflp rel_symp rel_transp rel_transfer rel_eq_onp pred_transfer pred_True pred_map pred_rel pred_mono_strong0 pred_mono_strong pred_mono pred_cong0 pred_cong pred_cong_simp; val wits = map2 mk_witness bnf_wits wit_thms; val bnf_rel = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas) @ (Bs' ~~ betas)) rel; val bnf_pred = Term.subst_atomic_types ((Ds ~~ deads) @ (As' ~~ alphas)) pred; val bnf = mk_bnf bnf_b Calpha live alphas betas dead deads bnf_map bnf_sets bnf_bd axioms defs facts wits bnf_rel bnf_pred; in note_bnf_thms fact_policy qualify bnf_b bnf lthy end; val one_step_defs = no_reflexive (bnf_map_def :: bnf_bd_def :: bnf_set_defs @ bnf_wit_defs @ [bnf_rel_def, bnf_pred_def]); in (key, goals, wit_goalss, after_qed, lthy, one_step_defs) end; structure BNF_Plugin = Plugin(type T = bnf); fun bnf_interpretation name f = BNF_Plugin.interpretation name (fn bnf => fn lthy => f (transfer_bnf (Proof_Context.theory_of lthy) bnf) lthy); val interpret_bnf = BNF_Plugin.data; fun register_bnf_raw key bnf = - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Symtab.update (key, morph_bnf phi bnf))); fun register_bnf plugins key bnf = register_bnf_raw key bnf #> interpret_bnf plugins bnf; fun bnf_def const_policy fact_policy internal qualify tacs wit_tac Ds map_b rel_b pred_b set_bs raw_csts = (fn (_, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, one_step_defs) => let fun mk_wits_tac ctxt set_maps = TRYALL Goal.conjunction_tac THEN (case triv_tac_opt of SOME tac => tac ctxt set_maps | NONE => unfold_thms_tac ctxt one_step_defs THEN wit_tac ctxt); val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_wit_thms set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => mk_wits_tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); in map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] []) goals (map (fn tac => fn {context = ctxt, prems = _} => unfold_thms_tac ctxt one_step_defs THEN tac ctxt) tacs) |> (fn thms => after_qed mk_wit_thms (map single thms) lthy) end) o prepare_def const_policy fact_policy internal qualify (K I) (K I) Ds map_b rel_b pred_b set_bs raw_csts; fun bnf_cmd (raw_csts, raw_plugins) = (fn (key, goals, (triv_tac_opt, wit_goalss), after_qed, lthy, defs) => let val plugins = raw_plugins lthy; val wit_goals = map Logic.mk_conjunction_balanced wit_goalss; fun mk_triv_wit_thms tac set_maps = Goal.prove_sorry lthy [] [] (Logic.mk_conjunction_balanced wit_goals) (fn {context = ctxt, prems = _} => TRYALL Goal.conjunction_tac THEN tac ctxt set_maps) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length wit_goals) |> map2 (Conjunction.elim_balanced o length) wit_goalss |> (map o map) (Thm.forall_elim_vars 0); val (mk_wit_thms, nontriv_wit_goals) = (case triv_tac_opt of NONE => (fn _ => [], map (map (rpair [])) wit_goalss) | SOME tac => (mk_triv_wit_thms tac, [])); in lthy |> Proof.theorem NONE (uncurry (register_bnf plugins key) oo after_qed mk_wit_thms) (map (single o rpair []) goals @ nontriv_wit_goals) |> Proof.unfolding ([[(@{thm OO_Grp_alt} :: @{thm mem_Collect_eq} :: defs, [])]]) |> Proof.refine_singleton (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (TRYALL (rtac ctxt refl)))) end) o prepare_def Do_Inline (user_policy Note_Some) false I Syntax.read_typ Syntax.read_term NONE Binding.empty Binding.empty Binding.empty [] raw_csts; fun print_bnfs ctxt = let fun pretty_set sets i = Pretty.block [Pretty.str (mk_setN (i + 1) ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt (nth sets i))]; fun pretty_bnf (key, BNF {T, map, sets, bd, live, lives, dead, deads, ...}) = Pretty.big_list (Pretty.string_of (Pretty.block [Pretty.str key, Pretty.str ":", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt T)])) ([Pretty.block [Pretty.str "live:", Pretty.brk 1, Pretty.str (string_of_int live), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) lives)], Pretty.block [Pretty.str "dead:", Pretty.brk 1, Pretty.str (string_of_int dead), Pretty.brk 3, Pretty.list "[" "]" (List.map (Syntax.pretty_typ ctxt) deads)], Pretty.block [Pretty.str (mapN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt map)]] @ List.map (pretty_set sets) (0 upto length sets - 1) @ [Pretty.block [Pretty.str (bdN ^ ":"), Pretty.brk 1, Pretty.quote (Syntax.pretty_term ctxt bd)]]); in Pretty.big_list "Registered bounded natural functors:" (map pretty_bnf (sort_by fst (Symtab.dest (Data.get (Context.Proof ctxt))))) |> Pretty.writeln end; val _ = Outer_Syntax.command \<^command_keyword>\print_bnfs\ "print all bounded natural functors" (Scan.succeed (Toplevel.keep (print_bnfs o Toplevel.context_of))); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\bnf\ "register a type as a bounded natural functor" (parse_opt_binding_colon -- Parse.typ --| (Parse.reserved "map" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "sets" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "bd") Parse.term)) [] --| (Parse.reserved "bd" -- \<^keyword>\:\) -- Parse.term -- Scan.optional ((Parse.reserved "wits" -- \<^keyword>\:\) |-- Scan.repeat1 (Scan.unless (Parse.reserved "rel" || Parse.reserved "plugins") Parse.term)) [] -- Scan.option ((Parse.reserved "rel" -- \<^keyword>\:\) |-- Parse.term) -- Scan.option ((Parse.reserved "pred" -- \<^keyword>\:\) |-- Parse.term) -- Scan.optional Plugin_Name.parse_filter (K Plugin_Name.default_filter) >> bnf_cmd); end; diff --git a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML @@ -1,2927 +1,2927 @@ (* Title: HOL/Tools/BNF/bnf_fp_def_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013, 2014 Sugared datatype and codatatype constructions. *) signature BNF_FP_DEF_SUGAR = sig type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list} type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list} type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list} type fp_sugar = {T: typ, BT: typ, X: typ, fp: BNF_Util.fp_kind, fp_res_index: int, fp_res: BNF_FP_Util.fp_result, pre_bnf: BNF_Def.bnf, fp_bnf: BNF_Def.bnf, absT_info: BNF_Comp.absT_info, fp_nesting_bnfs: BNF_Def.bnf list, live_nesting_bnfs: BNF_Def.bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option} val co_induct_of: 'a list -> 'a val strong_co_induct_of: 'a list -> 'a val morph_fp_bnf_sugar: morphism -> fp_bnf_sugar -> fp_bnf_sugar val morph_fp_co_induct_sugar: morphism -> fp_co_induct_sugar -> fp_co_induct_sugar val morph_fp_ctr_sugar: morphism -> fp_ctr_sugar -> fp_ctr_sugar val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar val fp_sugar_of: Proof.context -> string -> fp_sugar option val fp_sugar_of_global: theory -> string -> fp_sugar option val fp_sugars_of: Proof.context -> fp_sugar list val fp_sugars_of_global: theory -> fp_sugar list val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) -> theory -> theory val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val merge_type_args: BNF_Util.fp_kind -> ''a list * ''a list -> ''a list val type_args_named_constrained_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'a val type_binding_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'b val mixfix_of_spec: ((('a * 'b) * 'c) * 'd) * 'e -> 'b val mixfixed_ctr_specs_of_spec: (('a * 'b) * 'c) * 'd -> 'b val map_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'b val rel_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'c val pred_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'd val sel_default_eqs_of_spec: 'a * 'b -> 'b val mk_parametricity_goal: Proof.context -> term list -> term -> term -> term val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list -> 'a list val mk_ctor: typ list -> term -> term val mk_dtor: typ list -> term -> term val mk_bnf_sets: BNF_Def.bnf -> string * term list val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list val massage_simple_notes: string -> (bstring * 'a list * (int -> 'b)) list -> ((binding * 'c list) * ('a list * 'b) list) list val massage_multi_notes: string list -> typ list -> (string * 'a list list * (string -> 'b)) list -> ((binding * 'b) * ('a list * 'c list) list) list val define_ctrs_dtrs_for_type: string -> typ -> term -> term -> thm -> thm -> int -> int list -> term -> binding list -> mixfix list -> typ list list -> local_theory -> (term list list * term list * thm * thm list) * local_theory val wrap_ctrs: (string -> bool) -> BNF_Util.fp_kind -> bool -> string -> thm -> int -> int list -> thm -> thm -> binding list -> binding list list -> term list -> term list -> thm -> thm list -> local_theory -> Ctr_Sugar.ctr_sugar * local_theory val derive_map_set_rel_pred_thms: (string -> bool) -> BNF_Util.fp_kind -> int -> typ list -> typ list -> typ -> typ -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> string -> BNF_Def.bnf -> BNF_Def.bnf list -> typ -> term -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm -> thm list -> thm list -> thm list -> typ list list -> Ctr_Sugar.ctr_sugar -> local_theory -> (thm list * thm list * thm list list * thm list * thm list * thm list * thm list * thm list * thm list * thm list * thm list list list list * thm list list list list * thm list * thm list * thm list * thm list * thm list) * local_theory type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list) val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list) val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms val mk_co_recs_prelims: Proof.context -> BNF_Util.fp_kind -> typ list list list -> typ list -> typ list -> typ list -> typ list -> int list -> int list list -> term list -> term list * (typ list list * typ list list list list * term list list * term list list list list) option * (string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list)) option val repair_nullary_single_ctr: typ list list -> typ list list val mk_corec_p_pred_types: typ list -> int list -> typ list list val mk_corec_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list -> int list list -> term -> typ list list * (typ list list list list * typ list list list * typ list list list list * typ list) val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory -> (term * thm) * local_theory val define_rec: typ list list * typ list list list list * term list list * term list list list list -> (string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context -> (term * thm) * Proof.context val define_corec: 'a * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list -> term list -> term -> local_theory -> (term * thm) * local_theory val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list -> (string * term list) list -> term -> term -> typ list -> typ list -> term list * ((term * (term * term)) list * (int * term)) list * term val finish_induct_prem: Proof.context -> int -> term list -> term list * ((term * (term * term)) list * (int * term)) list * term -> term val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term -> term -> term -> int -> term list -> term list list -> term list -> term list list -> typ list list -> term val mk_induct_attrs: term list list -> Token.src list val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list -> Token.src list * Token.src list val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list -> ('a * typ list list list list * term list list * 'b) option -> thm -> thm list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list -> term list -> thm list -> Proof.context -> lfp_sugar_thms val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list -> thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> (thm list * thm) list val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list -> string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> int list list -> int list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list -> thm list -> gfp_sugar_thms val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> Ctr_Sugar.ctr_options * ((((((binding option * (typ * sort)) list * binding) * mixfix) * ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * term list) list -> local_theory -> local_theory val co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * Proof.context) -> ((Proof.context -> Plugin_Name.filter) * bool) * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list -> Proof.context -> local_theory val parse_ctr_arg: (binding * string) parser val parse_ctr_specs: ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list parser val parse_spec: ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) parser val parse_co_datatype: (Ctr_Sugar.ctr_options_cmd * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list) parser val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> (local_theory -> local_theory) parser end; structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR = struct open Ctr_Sugar open BNF_FP_Rec_Sugar_Util open BNF_Util open BNF_Comp open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar_Tactics val Eq_prefix = "Eq_"; val case_transferN = "case_transfer"; val ctor_iff_dtorN = "ctor_iff_dtor"; val ctr_transferN = "ctr_transfer"; val disc_transferN = "disc_transfer"; val sel_transferN = "sel_transfer"; val corec_codeN = "corec_code"; val corec_transferN = "corec_transfer"; val map_disc_iffN = "map_disc_iff"; val map_o_corecN = "map_o_corec"; val map_selN = "map_sel"; val pred_injectN = "pred_inject"; val rec_o_mapN = "rec_o_map"; val rec_transferN = "rec_transfer"; val set0N = "set0"; val set_casesN = "set_cases"; val set_introsN = "set_intros"; val set_inductN = "set_induct"; val set_selN = "set_sel"; type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list}; type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list}; type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list}; type fp_sugar = {T: typ, BT: typ, X: typ, fp: fp_kind, fp_res_index: int, fp_res: fp_result, pre_bnf: bnf, fp_bnf: bnf, absT_info: absT_info, fp_nesting_bnfs: bnf list, live_nesting_bnfs: bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option}; fun co_induct_of (i :: _) = i; fun strong_co_induct_of [_, s] = s; fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases} : fp_bnf_sugar) = {map_thms = map (Morphism.thm phi) map_thms, map_disc_iffs = map (Morphism.thm phi) map_disc_iffs, map_selss = map (map (Morphism.thm phi)) map_selss, rel_injects = map (Morphism.thm phi) rel_injects, rel_distincts = map (Morphism.thm phi) rel_distincts, rel_sels = map (Morphism.thm phi) rel_sels, rel_intros = map (Morphism.thm phi) rel_intros, rel_cases = map (Morphism.thm phi) rel_cases, pred_injects = map (Morphism.thm phi) pred_injects, set_thms = map (Morphism.thm phi) set_thms, set_selssss = map (map (map (map (Morphism.thm phi)))) set_selssss, set_introssss = map (map (map (map (Morphism.thm phi)))) set_introssss, set_cases = map (Morphism.thm phi) set_cases}; fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms, co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts} : fp_co_induct_sugar) = {co_rec = Morphism.term phi co_rec, common_co_inducts = map (Morphism.thm phi) common_co_inducts, co_inducts = map (Morphism.thm phi) co_inducts, co_rec_def = Morphism.thm phi co_rec_def, co_rec_thms = map (Morphism.thm phi) co_rec_thms, co_rec_discs = map (Morphism.thm phi) co_rec_discs, co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs, co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss, co_rec_codes = map (Morphism.thm phi) co_rec_codes, co_rec_transfers = map (Morphism.thm phi) co_rec_transfers, co_rec_o_maps = map (Morphism.thm phi) co_rec_o_maps, common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts, rel_co_inducts = map (Morphism.thm phi) rel_co_inducts, common_set_inducts = map (Morphism.thm phi) common_set_inducts, set_inducts = map (Morphism.thm phi) set_inducts}; fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctor_iff_dtor, ctr_defs, ctr_sugar, ctr_transfers, case_transfers, disc_transfers, sel_transfers} : fp_ctr_sugar) = {ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss, ctor_iff_dtor = Morphism.thm phi ctor_iff_dtor, ctr_defs = map (Morphism.thm phi) ctr_defs, ctr_sugar = morph_ctr_sugar phi ctr_sugar, ctr_transfers = map (Morphism.thm phi) ctr_transfers, case_transfers = map (Morphism.thm phi) case_transfers, disc_transfers = map (Morphism.thm phi) disc_transfers, sel_transfers = map (Morphism.thm phi) sel_transfers}; fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar, fp_co_induct_sugar} : fp_sugar) = {T = Morphism.typ phi T, BT = Morphism.typ phi BT, X = Morphism.typ phi X, fp = fp, fp_res = morph_fp_result phi fp_res, fp_res_index = fp_res_index, pre_bnf = morph_bnf phi pre_bnf, fp_bnf = morph_bnf phi fp_bnf, absT_info = morph_absT_info phi absT_info, fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs, live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs, fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar, fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar, fp_co_induct_sugar = Option.map (morph_fp_co_induct_sugar phi) fp_co_induct_sugar}; val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = fp_sugar Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun fp_sugar_of_generic context = Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); fun fp_sugars_of_generic context = Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; val fp_sugar_of = fp_sugar_of_generic o Context.Proof; val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory; val fp_sugars_of = fp_sugars_of_generic o Context.Proof; val fp_sugars_of_global = fp_sugars_of_generic o Context.Theory; structure FP_Sugar_Plugin = Plugin(type T = fp_sugar list); fun fp_sugars_interpretation name f = FP_Sugar_Plugin.interpretation name (fn fp_sugars => fn lthy => f (map (transfer_fp_sugar (Proof_Context.theory_of lthy)) fp_sugars) lthy); val interpret_fp_sugars = FP_Sugar_Plugin.data; val register_fp_sugars_raw = fold (fn fp_sugar as {T = Type (s, _), ...} => - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar)))); fun register_fp_sugars plugins fp_sugars = register_fp_sugars_raw fp_sugars #> interpret_fp_sugars plugins fp_sugars; fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs map_thmss common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss co_rec_disc_iffss co_rec_codess co_rec_transferss common_rel_co_inducts rel_co_inductss common_set_inducts set_inductss co_rec_o_mapss noted = let val fp_sugars = map_index (fn (kk, T) => {T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, fp_bnf = nth (#bnfs fp_res) kk, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = nth ctrXs_Tsss kk, ctor_iff_dtor = nth ctor_iff_dtors kk, ctr_defs = nth ctr_defss kk, ctr_sugar = nth ctr_sugars kk, ctr_transfers = nth ctr_transferss kk, case_transfers = nth case_transferss kk, disc_transfers = nth disc_transferss kk, sel_transfers = nth sel_transferss kk}, fp_bnf_sugar = {map_thms = nth map_thmss kk, map_disc_iffs = nth map_disc_iffss kk, map_selss = nth map_selsss kk, rel_injects = nth rel_injectss kk, rel_distincts = nth rel_distinctss kk, rel_sels = nth rel_selss kk, rel_intros = nth rel_intross kk, rel_cases = nth rel_casess kk, pred_injects = nth pred_injectss kk, set_thms = nth set_thmss kk, set_selssss = nth set_selsssss kk, set_introssss = nth set_introsssss kk, set_cases = nth set_casess kk}, fp_co_induct_sugar = SOME {co_rec = nth co_recs kk, common_co_inducts = common_co_inducts, co_inducts = nth co_inductss kk, co_rec_def = nth co_rec_defs kk, co_rec_thms = nth co_rec_thmss kk, co_rec_discs = nth co_rec_discss kk, co_rec_disc_iffs = nth co_rec_disc_iffss kk, co_rec_selss = nth co_rec_selsss kk, co_rec_codes = nth co_rec_codess kk, co_rec_transfers = nth co_rec_transferss kk, co_rec_o_maps = nth co_rec_o_mapss kk, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = nth rel_co_inductss kk, common_set_inducts = common_set_inducts, set_inducts = nth set_inductss kk}} |> morph_fp_sugar (substitute_noted_thm noted)) Ts; in register_fp_sugars_raw fp_sugars #> fold (interpret_bnf plugins) (#bnfs fp_res) #> interpret_fp_sugars plugins fp_sugars end; fun quasi_unambiguous_case_names names = let val ps = map (`Long_Name.base_name) names; val dups = Library.duplicates (op =) (map fst ps); fun underscore s = let val ss = Long_Name.explode s in space_implode "_" (drop (length ss - 2) ss) end; in map (fn (base, full) => if member (op =) dups base then underscore full else base) ps |> Name.variant_list [] end; fun zipper_map f = let fun zed _ [] = [] | zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys; in zed [] end; fun cannot_merge_types fp = error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters"); fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp; fun merge_type_args fp (As, As') = if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp; fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs; fun type_binding_of_spec (((((_, b), _), _), _), _) = b; fun mixfix_of_spec ((((_, mx), _), _), _) = mx; fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs; fun map_binding_of_spec ((_, (b, _, _)), _) = b; fun rel_binding_of_spec ((_, (_, b, _)), _) = b; fun pred_binding_of_spec ((_, (_, _, b)), _) = b; fun sel_default_eqs_of_spec (_, ts) = ts; fun ctr_sugar_kind_of_fp_kind Least_FP = Datatype | ctr_sugar_kind_of_fp_kind Greatest_FP = Codatatype; fun uncurry_thm 0 thm = thm | uncurry_thm 1 thm = thm | uncurry_thm n thm = rotate_prems ~1 (uncurry_thm (n - 1) (rotate_prems 1 (conjI RS thm))); fun choose_binary_fun fs AB = find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs; fun build_binary_fun_app fs t u = Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u)); fun build_the_rel ctxt Rs Ts A B = build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B); fun build_rel_app ctxt Rs Ts t u = build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u; fun build_set_app ctxt A t = Term.betapply (build_set ctxt A (fastype_of t), t); fun mk_parametricity_goal ctxt Rs t u = let val prem = build_the_rel ctxt Rs [] (fastype_of t) (fastype_of u) in HOLogic.mk_Trueprop (prem $ t $ u) end; val name_of_set = name_of_const "set function" domain_type; val fundefcong_attrs = @{attributes [fundef_cong]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs)); fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss); fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) = p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss; fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) = Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1)); fun flip_rels ctxt n thm = let val Rs = Term.add_vars (Thm.prop_of thm) []; val Rs' = rev (drop (length Rs - n) Rs); in infer_instantiate ctxt (map (fn f => (fst f, Thm.cterm_of ctxt (mk_flip f))) Rs') thm end; fun mk_ctor_or_dtor get_T Ts t = let val Type (_, Ts0) = get_T (fastype_of t) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_ctor = mk_ctor_or_dtor range_type; val mk_dtor = mk_ctor_or_dtor domain_type; fun mk_bnf_sets bnf = let val Type (T_name, Us) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val sets = sets_of_bnf bnf; fun mk_set U = (case find_index (curry (op =) U) lives of ~1 => Term.dummy | i => nth sets i); in (T_name, map mk_set Us) end; fun mk_xtor_co_recs thy fp fpTs Cs ts0 = let val nn = length fpTs; val (fpTs0, Cs0) = map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0 |> split_list; val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs); in map (Term.subst_TVars rho) ts0 end; fun liveness_of_fp_bnf n bnf = (case T_of_bnf bnf of Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts | _ => replicate n false); fun add_nesting_bnf_names Us = let fun add (Type (s, Ts)) ss = let val (needs, ss') = fold_map add Ts ss in if exists I needs then (true, insert (op =) s ss') else (false, ss') end | add T ss = (member (op =) Us T, ss); in snd oo add end; fun nesting_bnfs ctxt ctr_Tsss Us = map_filter (bnf_of ctxt) (fold (fold (fold (add_nesting_bnf_names Us))) ctr_Tsss []); fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p; fun massage_simple_notes base = filter_out (null o #2) #> map (fn (thmN, thms, f_attrs) => ((Binding.qualify true base (Binding.name thmN), []), map_index (fn (i, thm) => ([thm], f_attrs i)) thms)); fun massage_multi_notes b_names Ts = maps (fn (thmN, thmss, attrs) => @{map 3} (fn b_name => fn Type (T_name, _) => fn thms => ((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])])) b_names Ts thmss) #> filter_out (null o fst o hd o snd); fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy = let val ctor_absT = domain_type (fastype_of ctor); val (((w, xss), u'), _) = lthy |> yield_singleton (mk_Frees "w") ctor_absT ||>> mk_Freess "x" ctr_Tss ||>> yield_singleton Variable.variant_fixes fp_b_name; val u = Free (u', fpT); val ctor_iff_dtor_thm = let val goal = fold_rev Logic.all [w, u] (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT]) (Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor) |> Thm.close_derivation \<^here> end; val ctr_rhss = map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs)) ks xss; val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn b => fn mx => fn rhs => Local_Theory.define ((b, mx), ((Thm.make_def_binding (Config.get lthy bnf_internals) b, []), rhs)) #>> apsnd snd) ctr_bindings ctr_mixfixes ctr_rhss ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val ctr_defs = map (Morphism.thm phi) raw_ctr_defs; val ctrs0 = map (Morphism.term phi) raw_ctrs; in ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) end; fun wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs lthy = let val sumEN_thm' = unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms); fun exhaust_tac {context = ctxt, prems = _} = mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'; val inject_tacss = map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} => mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms; val half_distinct_tacss = map (map (fn (def, def') => fn {context = ctxt, ...} => mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def'])) (mk_half_pairss (`I ctr_defs)); val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss; fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs); val ctr_specs = @{map 3} ctr_spec_of disc_bindings ctrs0 sel_bindingss; val (ctr_sugar as {case_cong, ...}, lthy) = free_constructors (ctr_sugar_kind_of_fp_kind fp) tacss ((((plugins, discs_sels), standard_binding), ctr_specs), sel_default_eqs) lthy; val anonymous_notes = [([case_cong], fundefcong_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = if Config.get lthy bnf_internals then [(ctor_iff_dtorN, [ctor_iff_dtor_thm], K [])] |> massage_simple_notes fp_b_name else []; in (ctr_sugar, lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd) end; fun derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps fp_nested_rel_eq_onps fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm extra_unfolds_map extra_unfolds_set extra_unfolds_rel ctr_Tss ({casex, case_thms, discs, selss, sel_defs, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss, injects, distincts, distinct_discsss, ...} : ctr_sugar) lthy = let val n = length ctr_Tss; val ms = map length ctr_Tss; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpBT = B_ify_T fpT; val live_AsBs = filter (op <>) (As ~~ Bs); val live_As = map fst live_AsBs; val fTs = map (op -->) live_AsBs; val ((((((((xss, yss), fs), Ps), Rs), ta), tb), thesis), names_lthy) = lthy |> fold (fold Variable.declare_typ) [As, Bs] |> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss) ||>> mk_Frees "f" fTs ||>> mk_Frees "P" (map mk_pred1T live_As) ||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> yield_singleton (mk_Frees "a") fpT ||>> yield_singleton (mk_Frees "b") fpBT ||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT; val ctrAs = map (mk_ctr As) ctrs; val ctrBs = map (mk_ctr Bs) ctrs; val ctr_defs' = map2 (fn m => fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) ms ctr_defs; val ABfs = live_AsBs ~~ fs; fun derive_rel_case relAsBs rel_inject_thms rel_distinct_thms = let val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb; fun mk_assms ctrA ctrB ctxt = let val argA_Ts = binder_types (fastype_of ctrA); val argB_Ts = binder_types (fastype_of ctrB); val ((argAs, argBs), names_ctxt) = ctxt |> mk_Frees "x" argA_Ts ||>> mk_Frees "y" argB_Ts; val ctrA_args = list_comb (ctrA, argAs); val ctrB_args = list_comb (ctrB, argBs); in (fold_rev Logic.all (argAs @ argBs) (Logic.list_implies (mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) :: map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs, thesis)), names_ctxt) end; val (assms, names_lthy) = @{fold_map 2} mk_assms ctrAs ctrBs names_lthy; val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_case_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; fun derive_case_transfer rel_case_thm = let val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy; val caseA = mk_case As C casex; val caseB = mk_case Bs E casex; val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB; in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_case_transfer_tac ctxt rel_case_thm case_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; in if live = 0 then if plugins transfer_plugin then let val relAsBs = HOLogic.eq_const fpT; val rel_case_thm = derive_rel_case relAsBs [] []; val case_transfer_thm = derive_case_transfer rel_case_thm; val notes = [(case_transferN, [case_transfer_thm], K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> Local_Theory.notes notes; val subst = Morphism.thm (substitute_noted_thm noted); in (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [subst case_transfer_thm], [], []), lthy') end else (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy) else let val mapx = mk_map live As Bs (map_of_bnf fp_bnf); val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf); val setAs = map (mk_set As) (sets_of_bnf fp_bnf); val discAs = map (mk_disc_or_sel As) discs; val discBs = map (mk_disc_or_sel Bs) discs; val selAss = map (map (mk_disc_or_sel As)) selss; val selBss = map (map (mk_disc_or_sel Bs)) selss; val map_ctor_thm = if fp = Least_FP then fp_map_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val (y as Free (y_s, _), _) = lthy |> yield_singleton (mk_Frees "y") y_T; val ctor_cong = infer_instantiate' lthy [NONE, NONE, SOME (Thm.cterm_of lthy ctorB)] arg_cong; val fp_map_thm' = fp_map_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y))]) |> unfold_thms lthy [dtor_ctor]; in (fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans)) |> Drule.generalize (Names.empty, Names.make1_set y_s) end; val map_thms = let fun mk_goal ctrA ctrB xs ys = let val fmap = list_comb (mapx, fs); fun mk_arg (x as Free (_, T)) (Free (_, U)) = if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x; val xs' = map2 mk_arg xs ys; in mk_Trueprop_eq (fmap $ list_comb (ctrA, xs), list_comb (ctrB, xs')) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_tac ctxt abs_inverses pre_map_def map_ctor_thm live_nesting_map_id0s ctr_defs' extra_unfolds_map) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val set0_thms = let fun mk_goal A setA ctrA xs = let val sets = map (build_set_app lthy A) (filter (exists_subtype_in [A] o fastype_of) xs); in mk_Trueprop_eq (setA $ list_comb (ctrA, xs), (if null sets then HOLogic.mk_set A [] else Library.foldl1 mk_union sets)) end; val goals = @{map 2} (fn live_A => fn setA => map2 (mk_goal live_A setA) ctrAs xss) live_As setAs |> flat; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_set_thms fp_nesting_set_maps live_nesting_set_maps ctr_defs' extra_unfolds_set) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_thms = set0_thms |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left}); val rel_ctor_thm = if fp = Least_FP then fp_rel_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val z_T = domain_type (fastype_of ctorB); val ((y as Free (y_s, _), z as Free (z_s, _)), _) = lthy |> yield_singleton (mk_Frees "y") y_T ||>> yield_singleton (mk_Frees "z") z_T; in fp_rel_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))]) |> unfold_thms lthy [dtor_ctor] |> Drule.generalize (Names.empty, Names.make2_set y_s z_s) end; val rel_inject_thms = let fun mk_goal ctrA ctrB xs ys = let val lhs = list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys); val conjuncts = map2 (build_rel_app lthy Rs []) xs ys; in HOLogic.mk_Trueprop (if null conjuncts then lhs else HOLogic.mk_eq (lhs, Library.foldr1 HOLogic.mk_conj conjuncts)) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val half_rel_distinct_thmss = let fun mk_goal ((ctrA, xs), (ctrB, ys)) = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys))); val rel_infos = (ctrAs ~~ xss, ctrBs ~~ yss); val goalss = map (map mk_goal) (mk_half_pairss rel_infos); val goals = flat goalss; in unflat goalss (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_flip = rel_flip_of_bnf fp_bnf; fun mk_other_half_rel_distinct_thm thm = flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2); val other_half_rel_distinct_thmss = map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss; val (rel_distinct_thms, _) = join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss; fun mk_rel_intro_thm m thm = uncurry_thm m (thm RS iffD2) handle THM _ => thm; val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms; val rel_code_thms = map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @ map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms; val ctr_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val (set_cases_thms, set_cases_attrss) = let fun mk_prems assms elem t ctxt = (case fastype_of t of Type (type_name, xs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val new_var = not (T = fastype_of elem); val (x, ctxt') = if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt); in mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt' |>> map (new_var ? Logic.all x) end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt)) | T => rpair ctxt (if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else [])); in split_list (map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)); val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy; val premss = map (fn ctr => let val (args, names_lthy) = mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy; in flat (zipper_map (fn (prev_args, arg, next_args) => let val (args_with_elem, args_without_elem) = if fastype_of arg = A then (prev_args @ [elem] @ next_args, prev_args @ next_args) else `I (prev_args @ [arg] @ next_args); in mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))] elem arg names_lthy |> fst |> map (fold_rev Logic.all args_without_elem) end) args) end) ctrAs; val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars (flat premss) goal (fn {context = ctxt, prems} => mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms) |> Thm.close_derivation \<^here> |> rotate_prems ~1; val cases_set_attr = - Attrib.internal (K (Induct.cases_pred (name_of_set set))); + Attrib.internal \<^here> (K (Induct.cases_pred (name_of_set set))); val ctr_names = quasi_unambiguous_case_names (flat (map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs))); in (* TODO: @{attributes [elim?]} *) (thm, [Attrib.consumes 1, cases_set_attr, Attrib.case_names ctr_names]) end) setAs) end; val (set_intros_thmssss, set_intros_thms) = let fun mk_goals A setA ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (y, ctxt') = yield_singleton (mk_Frees "y") T ctxt; val assm = mk_Trueprop_mem (y, set $ t); in apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args y ctxt') end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt)) | T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt)); val (goalssss, _) = fold_map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)) in @{fold_map 2} (fn ctr => fn xs => fold_map (mk_goals A set (Term.list_comb (ctr, xs))) xs) ctrAs xss end) setAs lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_sel_thms = let val n = length discAs; fun mk_conjunct n k discA selAs discB selBs = (if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @ (if null selAs then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA $ ta, discB $ tb], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs []) (map (rapp ta) selAs) (map (rapp tb) selBs)))]); val goals = if n = 0 then [] else [mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb, (case flat (@{map 5} (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of [] => \<^term>\True\ | conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))]; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs)) |> Thm.close_derivation \<^here>; in map prove goals end; val (rel_case_thm, rel_case_attrs) = let val thm = derive_rel_case relAsBs rel_inject_thms rel_distinct_thms; val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs); in (thm, [Attrib.case_names ctr_names, Attrib.consumes 1] @ @{attributes [cases pred]}) end; val case_transfer_thm = derive_case_transfer rel_case_thm; val sel_transfer_thms = if null selAss then [] else let val shared_sels = foldl1 (uncurry (inter (op =))) (map (op ~~) (selAss ~~ selBss)); val goals = map (uncurry (mk_parametricity_goal names_lthy Rs)) shared_sels; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sel_transfer_tac ctxt n sel_defs case_transfer_thm) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val disc_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt (the_single rel_sel_thms) (the_single exhaust_discs) (flat (flat distinct_discsss))) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val map_disc_iff_thms = let val discsB = map (mk_disc_or_sel Bs) discs; val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs; fun mk_goal (discA_t, discB) = if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then NONE else SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t)); val goals = map_filter mk_goal (discsA_t ~~ discsB); in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val (map_sel_thmss, map_sel_thms) = let fun mk_goal discA selA selB = let val prem = Term.betapply (discA, ta); val lhs = selB $ (Term.list_comb (mapx, fs) $ ta); val lhsT = fastype_of lhs; val map_rhsT = map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT; val map_rhs = build_map lthy [] [] (the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT); val rhs = (case map_rhs of Const (\<^const_name>\id\, _) => selA $ ta | _ => map_rhs $ (selA $ ta)); val concl = mk_Trueprop_eq (lhs, rhs); in if is_refl_bool prem then concl else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl) end; val goalss = @{map 3} (map2 o mk_goal) discAs selAss selBss; val goals = flat goalss; in `(unflat goalss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms (flat sel_thmss) live_nesting_map_id0s) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val (set_sel_thmssss, set_sel_thms) = let fun mk_goal setA discA selA ctxt = let val prem = Term.betapply (discA, ta); val sel_rangeT = range_type (fastype_of selA); val A = HOLogic.dest_setT (range_type (fastype_of setA)); fun travese_nested_types t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; val assm = mk_Trueprop_mem (x, set $ a); in travese_nested_types x ctxt' |>> map (Logic.mk_implies o pair assm) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt)); val (concls, ctxt') = if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt) else travese_nested_types (selA $ ta) ctxt; in if exists_subtype_in [A] sel_rangeT then if is_refl_bool prem then (concls, ctxt') else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt') else ([], ctxt) end; val (goalssss, _) = fold_map (fn set => @{fold_map 2} (fold_map o mk_goal set) discAs selAss) setAs names_lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val pred_injects = let val rel_eq_onp_with_tops_of = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} lthy))); val eq_onps = map rel_eq_onp_with_tops_of (map rel_eq_onp_of_bnf fp_bnfs @ fp_nesting_rel_eq_onps @ live_nesting_rel_eq_onps @ fp_nested_rel_eq_onps); val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) live_As); val cts = map (SOME o Thm.cterm_of lthy) (map mk_eq_onp Ps); val get_rhs = Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> snd; val pred_eq_onp_conj = List.foldr (fn (_, thm) => thm RS @{thm eq_onp_live_step}) @{thm refl[of True]}; fun predify_rel_inject rel_inject = let val conjuncts = try (get_rhs #> HOLogic.dest_conj) rel_inject |> the_default []; fun postproc thm = if null conjuncts then thm RS (@{thm eq_onp_same_args} RS iffD1) else @{thm box_equals} OF [thm, @{thm eq_onp_same_args}, pred_eq_onp_conj conjuncts |> unfold_thms lthy @{thms simp_thms(21)}]; in rel_inject |> Thm.instantiate' cTs cts |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))) |> unfold_thms lthy eq_onps |> postproc |> unfold_thms lthy @{thms top_conj} end; in rel_inject_thms |> map (unfold_thms lthy [@{thm conj_assoc}]) |> map predify_rel_inject |> Proof_Context.export names_lthy lthy end; val anonymous_notes = [(rel_code_thms, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = (if Config.get lthy bnf_internals then [(set0N, set0_thms, K [])] else []) @ [(case_transferN, [case_transfer_thm], K []), (ctr_transferN, ctr_transfer_thms, K []), (disc_transferN, disc_transfer_thms, K []), (sel_transferN, sel_transfer_thms, K []), (mapN, map_thms, K (nitpicksimp_attrs @ simp_attrs)), (map_disc_iffN, map_disc_iff_thms, K simp_attrs), (map_selN, map_sel_thms, K []), (pred_injectN, pred_injects, K simp_attrs), (rel_casesN, [rel_case_thm], K rel_case_attrs), (rel_distinctN, rel_distinct_thms, K simp_attrs), (rel_injectN, rel_inject_thms, K simp_attrs), (rel_introsN, rel_intro_thms, K []), (rel_selN, rel_sel_thms, K []), (setN, set_thms, K (case_fp fp nitpicksimp_attrs [] @ simp_attrs)), (set_casesN, set_cases_thms, nth set_cases_attrss), (set_introsN, set_intros_thms, K []), (set_selN, set_sel_thms, K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) map_thms) |> fp = Least_FP ? uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) rel_code_thms) |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) set0_thms) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (rel_code_thms @ map_thms @ set_thms)) |> Local_Theory.notes (anonymous_notes @ notes); val subst = Morphism.thm (substitute_noted_thm noted); in ((map subst map_thms, map subst map_disc_iff_thms, map (map subst) map_sel_thmss, map subst rel_inject_thms, map subst rel_distinct_thms, map subst rel_sel_thms, map subst rel_intro_thms, [subst rel_case_thm], map subst pred_injects, map subst set_thms, map (map (map (map subst))) set_sel_thmssss, map (map (map (map subst))) set_intros_thmssss, map subst set_cases_thms, map subst ctr_transfer_thms, [subst case_transfer_thm], map subst disc_transfer_thms, map subst sel_transfer_thms), lthy') end end; type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list); fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) = ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs), (map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms; val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism; type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list); fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair), corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs), (corec_selsss, corec_sel_attrs)) = ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs, coinduct_attrs_pair), map (map (Morphism.thm phi)) corecss, map (map (Morphism.thm phi)) corec_discss, (map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs), (map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms; val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism; fun unzip_recT (Type (\<^type_name>\prod\, [_, TFree x])) (T as Type (\<^type_name>\prod\, Ts as [_, TFree y])) = if x = y then [T] else Ts | unzip_recT _ (Type (\<^type_name>\prod\, Ts as [_, TFree _])) = Ts | unzip_recT _ T = [T]; fun mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts = let val Css = map2 replicate ns Cs; val x_Tssss = @{map 6} (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_rec_fun_T => map2 (map2 unzip_recT) ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T))) absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts; val x_Tsss' = map (map flat_rec_arg_args) x_Tssss; val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css; val ((fss, xssss), _) = ctxt |> mk_Freess "f" f_Tss ||>> mk_Freessss "x" x_Tssss; in (f_Tss, x_Tssss, fss, xssss) end; fun unzip_corecT (Type (\<^type_name>\sum\, _)) T = [T] | unzip_corecT _ (Type (\<^type_name>\sum\, Ts)) = Ts | unzip_corecT _ T = [T]; (*avoid "'a itself" arguments in corecursors*) fun repair_nullary_single_ctr [[]] = [[HOLogic.unitT]] | repair_nullary_single_ctr Tss = Tss; fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts = let val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; val g_absTs = map range_type fun_Ts; val g_Tsss = map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) Cs ctr_Tsss' g_Tsss; val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; in (q_Tssss, g_Tsss, g_Tssss, g_absTs) end; fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs; fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec = (mk_corec_p_pred_types Cs ns, mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss (binder_fun_types (fastype_of dtor_corec))); fun mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts = let val p_Tss = mk_corec_p_pred_types Cs ns; val (q_Tssss, g_Tsss, g_Tssss, corec_types) = mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts; val (((((Free (x, _), cs), pss), qssss), gssss), _) = ctxt |> yield_singleton (mk_Frees "x") dummyT ||>> mk_Frees "a" Cs ||>> mk_Freess "p" p_Tss ||>> mk_Freessss "q" q_Tssss ||>> mk_Freessss "g" g_Tssss; val cpss = map2 (map o rapp) cs pss; fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd); fun build_dtor_corec_arg _ [] [cg] = cg | build_dtor_corec_arg T [cq] [cg, cg'] = mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg) (build_sum_inj Inr_const (fastype_of cg', T) $ cg'); val pgss = @{map 3} flat_corec_preds_predsss_gettersss pss qssss gssss; val cqssss = map2 (map o map o map o rapp) cs qssss; val cgssss = map2 (map o map o map o rapp) cs gssss; val cqgsss = @{map 3} (@{map 3} (@{map 3} build_dtor_corec_arg)) g_Tsss cqssss cgssss; in (x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types)) end; fun mk_co_recs_prelims ctxt fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 = let val thy = Proof_Context.theory_of ctxt; val (xtor_co_rec_fun_Ts, xtor_co_recs) = mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd); val (recs_args_types, corecs_args_types) = if fp = Least_FP then mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (rpair NONE o SOME) else mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (pair NONE o SOME); in (xtor_co_recs, recs_args_types, corecs_args_types) end; fun mk_preds_getterss_join c cps absT abs cqgss = let val n = length cqgss; val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqgss; in Term.lambda c (mk_IfN absT cps ts) end; fun define_co_rec_as fp Cs fpT b rhs lthy0 = let val thy = Proof_Context.theory_of lthy0; val ((cst, (_, def)), (lthy', lthy)) = lthy0 |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((b, NoSyn), ((Thm.make_def_binding (Config.get lthy0 bnf_internals) b, []), rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst); val def' = Morphism.thm phi def; in ((cst', def'), lthy') end; fun define_rec (_, _, fss, xssss) mk_binding fpTs Cs reps ctor_rec = let val nn = length fpTs; val (ctor_rec_absTs, fpT) = strip_typeN nn (fastype_of ctor_rec) |>> map domain_type ||> domain_type; in define_co_rec_as Least_FP Cs fpT (mk_binding recN) (fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_rec, @{map 4} (fn ctor_rec_absT => fn rep => fn fs => fn xsss => mk_case_absumprod ctor_rec_absT rep fs (map (map HOLogic.mk_tuple) xsss) (map flat_rec_arg_args xsss)) ctor_rec_absTs reps fss xssss))) end; fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = let val nn = length fpTs; val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); in define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN) (fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec, @{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss))) end; fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) = (case AList.lookup (op =) setss_fp_nesting T_name of NONE => [] | SOME raw_sets0 => let val (Xs_Ts, (Ts, raw_sets)) = filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0)) |> split_list ||> split_list; val sets = map (mk_set Ts0) raw_sets; val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts; val xysets = map (pair x) (ys ~~ sets); val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts; in flat (map2 (map o apfst o cons) xysets ppremss) end) | mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X = [([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))] | mk_induct_raw_prem_prems _ _ _ _ _ = []; fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts = let val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts; val pprems = flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts); val y = Term.list_comb (ctr, map alter_x xs); val p' = enforce_type names_ctxt domain_type (fastype_of y) p; in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end; fun close_induct_prem_prem nn ps xs t = fold_rev Logic.all (map Free (drop (nn + length xs) (rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t; fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) = let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) => mk_Trueprop_mem (y, set $ x')) xysets, HOLogic.mk_Trueprop (p' $ x))) end; fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) = fold_rev Logic.all xs (Logic.list_implies (map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl)); fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts = let fun build_the_rel T Xs_T = build_rel [] ctxt [] [] (fn (T, X) => nth rs' (find_index (fn Xs => member (op =) Xs X) Xss) |> enforce_type ctxt domain_type T) (T, Xs_T) |> Term.subst_atomic_types (flat Xss ~~ flat fpTss); fun build_rel_app usel vsel Xs_T = fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T); in (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @ (if null usels then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))]) end; fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss = @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n) (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss |> flat |> Library.foldr1 HOLogic.mk_conj handle List.Empty => \<^term>\True\; fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss = fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr, HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss))); fun postproc_co_induct ctxt nn prop prop_conj = Drule.zero_var_indexes #> `(conj_dests nn) #>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop)) ##> (fn thm => Thm.permute_prems 0 (~ nn) (if nn = 1 then thm RS prop else funpow nn (fn thm => unfold_thms ctxt @{thms conj_assoc} (thm RS prop_conj)) thm)); fun mk_induct_attrs ctrss = let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss); in [Attrib.case_names induct_cases] end; fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val fpB_Ts = map B_ify_T fpA_Ts; val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss; val ctrBss = map (map B_ify) ctrAss; val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> mk_Freesss "a" ctrAs_Tsss ||>> mk_Freesss "b" ctrBs_Tsss; val prems = let fun mk_prem ctrA ctrB argAs argBs = fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies) (map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs) (HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts (Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs))))); in flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs)); val vars = Variable.add_free_names ctxt goal []; val rel_induct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm, mk_induct_attrs ctrAss) end; fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions abs_inverses ctrss ctr_defss recs rec_defs ctxt = let val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss; val nn = length pre_bnfs; val ns = map length ctr_Tsss; val mss = map (map length) ctr_Tsss; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val (((ps, xsss), us'), names_ctxt) = ctxt |> mk_Frees "P" (map mk_pred1T fpTs) ||>> mk_Freesss "x" ctr_Tsss ||>> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs; val (induct_thms, induct_thm) = let val raw_premss = @{map 4} (@{map 3} o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting) ps ctrss ctr_Tsss ctrXs_Tsss; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)); val goal = Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps))) (raw_premss, concl); val vars = Variable.add_free_names ctxt goal []; val kksss = map (map (map (fst o snd) o #2)) raw_premss; val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss); val thm = Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses abs_inverses fp_nesting_set_maps pre_set_defss) |> Thm.close_derivation \<^here>; in `(conj_dests nn) thm end; val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss; fun mk_rec_thmss (_, x_Tssss, fss, _) recs rec_defs ctor_rec_thms = let val frecs = map (lists_bmoc fss) recs; fun mk_goal frec xctr f xs fxs = fold_rev (fold_rev Logic.all) (xs :: fss) (mk_Trueprop_eq (frec $ xctr, Term.list_comb (f, fxs))); fun maybe_tick (T, U) u f = if try (fst o HOLogic.dest_prodT) U = SOME T then Term.lambda u (HOLogic.mk_prod (u, f $ u)) else f; fun build_rec (x as Free (_, T)) U = if T = U then x else let val build_simple = indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs (fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk)); in build_map ctxt [] [] build_simple (T, U) $ x end; val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss; val goalss = @{map 5} (@{map 4} o mk_goal) frecs xctrss fss xsss fxsss; val tacss = @{map 4} (map ooo mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs) ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss end; val rec_thmss = mk_rec_thmss (the rec_args_typess) recs rec_defs ctor_rec_thms; in ((induct_thms, induct_thm, mk_induct_attrs ctrss), (rec_thmss, nitpicksimp_attrs @ simp_attrs)) end; fun mk_coinduct_attrs fpTs ctrss discss mss = let val fp_b_names = map base_name_of_typ fpTs; fun mk_coinduct_concls ms discs ctrs = let fun mk_disc_concl disc = [name_of_disc disc]; fun mk_ctr_concl 0 _ = [] | mk_ctr_concl _ ctr = [name_of_ctr ctr]; val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]]; val ctr_concls = map2 mk_ctr_concl ms ctrs; in flat (map2 append disc_concls ctr_concls) end; val coinduct_cases = quasi_unambiguous_case_names (map (prefix Eq_prefix) fp_b_names); val coinduct_conclss = @{map 3} (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss; val coinduct_case_names_attr = Attrib.case_names coinduct_cases; val coinduct_case_concl_attrs = map2 (fn casex => fn concls => Attrib.case_conclusion (casex, concls)) coinduct_cases coinduct_conclss; val common_coinduct_attrs = coinduct_case_names_attr :: coinduct_case_concl_attrs; val coinduct_attrs = Attrib.consumes 1 :: coinduct_case_names_attr :: coinduct_case_concl_attrs; in (coinduct_attrs, common_coinduct_attrs) end; fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list) abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpB_Ts = map B_ify_T fpA_Ts; val (Rs, IRs, fpAs, fpBs, _) = let val fp_names = map base_name_of_typ fpA_Ts; val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> Variable.variant_fixes fp_names ||>> Variable.variant_fixes (map (suffix "'") fp_names); in (Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts, names_ctxt) end; val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) = let val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss); fun mk_selsss ts Ts = map2 (map o map o rapp) ts (map (map (map (mk_disc_or_sel Ts))) selsss); in ((mk_discss fpAs As, mk_selsss fpAs As), (mk_discss fpBs Bs, mk_selsss fpBs Bs)) end; val prems = let fun mk_prem_ctr_concls n k discA_t selA_ts discB_t selB_ts = (if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @ (case (selA_ts, selB_ts) of ([], []) => [] | (_ :: _, _ :: _) => [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA_t, discB_t], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]); fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss = Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n) (1 upto n) discA_ts selA_tss discB_ts selB_tss)) handle List.Empty => \<^term>\True\; fun mk_prem IR tA tB n discA_ts selA_tss discB_ts selB_tss = fold_rev Logic.all [tA, tB] (Logic.mk_implies (HOLogic.mk_Trueprop (IR $ tA $ tB), HOLogic.mk_Trueprop (mk_prem_concl n discA_ts selA_tss discB_ts selB_tss))); in @{map 8} mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts))); val vars = Variable.add_free_names ctxt goal []; val rel_coinduct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars) (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm, mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss) end; fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = let fun mk_prems A Ps ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val X = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt; val assm = mk_Trueprop_mem (x, set $ a); in (case build_binary_fun_app Ps x a of NONE => mk_prems A Ps ctr_args x ctxt' |>> map (Logic.all x o Logic.mk_implies o pair assm) | SOME f => ([Logic.all x (Logic.mk_implies (assm, Logic.mk_implies (HOLogic.mk_Trueprop f, HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))], ctxt')) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt) else ([], ctxt)); fun mk_prems_for_ctr A Ps ctr ctxt = let val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt; in fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt' |>> map (fold_rev Logic.all args) o flat |>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr))) end; fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt = let val ((x, fp), ctxt') = ctxt |> yield_singleton (mk_Frees "x") A ||>> yield_singleton (mk_Frees "a") fpT; val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x) (the (build_binary_fun_app Ps x fp))); in fold_map (mk_prems_for_ctr A Ps) ctrs ctxt' |>> split_list |>> map_prod flat flat |>> apfst (rpair concl) end; fun mk_thm ctxt fpTs ctrss sets = let val A = HOLogic.dest_setT (range_type (fastype_of (hd sets))); val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt; val (((prems, concl), case_names), ctxt'') = fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt' |>> apfst split_list o split_list |>> apfst (apfst flat) |>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj)) |>> apsnd flat; val vars = fold (Variable.add_free_names ctxt) (concl :: prems) []; val thm = Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> Thm.close_derivation \<^here>; val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names); - val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets; + val induct_set_attrs = map (Attrib.internal \<^here> o K o Induct.induct_pred o name_of_set) sets; in (thm, case_names_attr :: induct_set_attrs) end val consumes_attr = Attrib.consumes 1; in map (mk_thm ctxt fpTs ctrss #> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr)) (transpose setss) end; fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt = let val n = Thm.nprems_of coind; val m = Thm.nprems_of (hd rel_monos) - n; fun mk_inst phi = (phi, Thm.cterm_of ctxt (mk_union (Var phi, HOLogic.eq_const (fst (dest_pred2T (#2 phi)))))); val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map mk_inst; fun mk_unfold rel_eq rel_mono = let val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl]; val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset}); in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS eqTrueI end; val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)}; in Thm.instantiate (TVars.empty, Vars.make insts) coind |> unfold_thms ctxt unfolds end; fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) = let val nn = length pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val exhausts = map #exhaust ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val (((rs, us'), vs'), _) = ctxt |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs) ||>> Variable.variant_fixes fp_b_names ||>> Variable.variant_fixes (map (suffix "'") fp_b_names); val us = map2 (curry Free) us' fpTs; val udiscss = map2 (map o rapp) us discss; val uselsss = map2 (map o map o rapp) us selsss; val vs = map2 (curry Free) vs' fpTs; val vdiscss = map2 (map o rapp) vs discss; val vselsss = map2 (map o map o rapp) vs selsss; val uvrs = @{map 3} (fn r => fn u => fn v => r $ u $ v) rs us vs; val uv_eqs = map2 (curry HOLogic.mk_eq) us vs; val strong_rs = @{map 4} (fn u => fn v => fn uvr => fn uv_eq => fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v))) uvrs us vs)) fun mk_goal rs0' = Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs) (map alter_r rs0')) uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss, concl); val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else [])); fun prove dtor_coinduct' goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss)) |> Thm.close_derivation \<^here>; val rel_eqs = map rel_eq_of_bnf pre_bnfs; val rel_monos = map rel_mono_of_bnf pre_bnfs; val dtor_coinducts = [dtor_coinduct] @ (if strong then [mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p ctxt] else []); in map2 (postproc_co_induct ctxt nn mp @{thm conj_commute[THEN iffD1]} oo prove) dtor_coinducts goals end; fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs = let fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec = iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]]; val ctor_dtor_corec_thms = @{map 3} mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms; val pre_map_defs = map map_def_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val ctrss = map #ctrs ctr_sugars; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val discIss = map #discIs ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss ctr_sugars; fun mk_maybe_not pos = not pos ? HOLogic.mk_not; val gcorecs = map (lists_bmoc pgss) corecs; val corec_thmss = let val (us', _) = ctxt |> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; fun mk_goal c cps gcorec n k ctr m cfs' = fold_rev (fold_rev Logic.all) ([c] :: pgss) (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps, mk_Trueprop_eq (gcorec $ c, Term.list_comb (ctr, take m cfs')))); val mk_U = typ_subst_nonatomic (map2 (fn C => fn fpT => (mk_sumT (fpT, C), fpT)) Cs fpTs); fun tack (c, u) f = let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in Term.lambda x' (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ x') end; fun build_corec cqg = let val T = fastype_of cqg in if exists_subtype_in Cs T then let val U = mk_U T; val build_simple = indexify fst (map2 (curry mk_sumT) fpTs Cs) (fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk)); in build_map ctxt [] [] build_simple (T, U) $ cqg end else cqg end; val cqgsss' = map (map (map build_corec)) cqgsss; val goalss = @{map 8} (@{map 4} oooo mk_goal) cs cpss gcorecs ns kss ctrss mss cqgsss'; val tacss = @{map 4} (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s) ctor_dtor_corec_thms pre_map_defs abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss |> map (map (unfold_thms ctxt @{thms case_sum_if})) end; val corec_disc_iff_thmss = let fun mk_goal c cps gcorec n k disc = mk_Trueprop_eq (disc $ (gcorec $ c), if n = 1 then \<^Const>\True\ else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps)); val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; fun mk_case_split' cp = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt cp)] @{thm case_split}; val case_splitss' = map (map mk_case_split') cpss; val tacss = @{map 3} (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss; fun prove goal tac = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (tac o #context)) |> Thm.close_derivation \<^here>; fun proves [_] [_] = [] | proves goals tacs = map2 prove goals tacs; in map2 proves goalss tacss end; fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs); val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss; fun mk_corec_sel_thm corec_thm sel sel_thm = let val (domT, ranT) = dest_funT (fastype_of sel); val arg_cong' = Thm.instantiate' (map (SOME o Thm.ctyp_of ctxt) [domT, ranT]) [NONE, NONE, SOME (Thm.cterm_of ctxt sel)] arg_cong |> Thm.varifyT_global; val sel_thm' = sel_thm RSN (2, trans); in corec_thm RS arg_cong' RS sel_thm' end; fun mk_corec_sel_thms corec_thmss = @{map 3} (@{map 3} (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss; val corec_sel_thmsss = mk_corec_sel_thms corec_thmss; in ((coinduct_thms_pairs, mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, simp_attrs), (corec_sel_thmsss, simp_attrs)) end; fun define_co_datatypes prepare_plugins prepare_constraint prepare_typ prepare_term fp construct_fp ((raw_plugins, discs_sels0), specs) lthy = let val plugins = prepare_plugins lthy raw_plugins; val discs_sels = discs_sels0 orelse fp = Greatest_FP; val nn = length specs; val fp_bs = map type_binding_of_spec specs; val fp_b_names = map Binding.name_of fp_bs; val fp_common_name = mk_common_name fp_b_names; val map_bs = map map_binding_of_spec specs; val rel_bs = map rel_binding_of_spec specs; val pred_bs = map pred_binding_of_spec specs; fun prepare_type_arg (_, (ty, c)) = let val TFree (s, _) = prepare_typ lthy ty in TFree (s, prepare_constraint lthy c) end; val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs; val unsorted_Ass0 = map (map (resort_tfree_or_tvar \<^sort>\type\)) Ass0; val unsorted_As = Library.foldr1 (merge_type_args fp) unsorted_Ass0; val num_As = length unsorted_As; val set_boss = map (map fst o type_args_named_constrained_of_spec) specs; val set_bss = map (map (the_default Binding.empty)) set_boss; fun add_fake_type spec = Typedecl.basic_typedecl {final = true} (type_binding_of_spec spec, num_As, Mixfix.reset_pos (mixfix_of_spec spec)); val (fake_T_names, fake_lthy) = fold_map add_fake_type specs lthy; val qsoty = quote o Syntax.string_of_typ fake_lthy; val _ = (case Library.duplicates (op =) unsorted_As of [] => () | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^ "datatype specification")); val bad_args = map (Logic.type_map (singleton (Variable.polymorphic lthy))) unsorted_As |> filter_out Term.is_TVar; val _ = null bad_args orelse error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^ "datatype specification"); val mixfixes = map mixfix_of_spec specs; val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => () | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b))); val mx_ctr_specss = map mixfixed_ctr_specs_of_spec specs; val ctr_specss = map (map fst) mx_ctr_specss; val ctr_mixfixess = map (map snd) mx_ctr_specss; val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss; val ctr_bindingss = map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of_ctr_spec)) fp_b_names ctr_specss; val ctr_argsss = map (map args_of_ctr_spec) ctr_specss; val sel_bindingsss = map (map (map fst)) ctr_argsss; val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss; val raw_sel_default_eqss = map sel_default_eqs_of_spec specs; val (As :: _) :: fake_ctr_Tsss = burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0); val As' = map dest_TFree As; val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss []; val _ = (case subtract (op =) As' rhs_As' of [] => () | extras => error ("Extra type variables on right-hand side: " ^ commas (map (qsoty o TFree) extras))); val fake_Ts = map (fn s => Type (s, As)) fake_T_names; val ((((Bs0, Cs), Es), Xs), _) = lthy |> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As |> mk_TFrees num_As ||>> mk_TFrees nn ||>> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) = s = s' andalso (Ts = Ts' orelse error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^ " (expected " ^ qsoty T' ^ ")")) | eq_fpT_check _ _ = false; fun freeze_fp (T as Type (s, Ts)) = (case find_index (eq_fpT_check T) fake_Ts of ~1 => Type (s, map freeze_fp Ts) | kk => nth Xs kk) | freeze_fp T = T; val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts); val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val _ = let fun add_deps i = fold (fn T => fold_index (fn (j, X) => (i <> j andalso exists_subtype_in [X] T) ? insert (op =) (i, j)) Xs); val add_missing_nodes = fold (AList.default (op =) o rpair []) (0 upto nn - 1); val deps = fold_index (uncurry (fold o add_deps)) ctrXs_Tsss [] |> AList.group (op =) |> add_missing_nodes; val G = Int_Graph.make (map (apfst (rpair ())) deps); val sccs = map (sort int_ord) (Int_Graph.strong_conn G); val str_of_scc = prefix (co_prefix fp ^ "datatype ") o space_implode " and " o map (suffix " = \" o Long_Name.base_name); fun warn [_] = () | warn sccs = warning ("Defined types not fully mutually " ^ co_prefix fp ^ "recursive\n\ \Alternative specification:\n" ^ cat_lines (map (prefix " " o str_of_scc o map (nth fp_b_names)) sccs)); in warn (order_strong_conn (op =) Int_Graph.make Int_Graph.topological_order deps sccs) end; val killed_As = map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE) (As ~~ transpose set_boss); val (((pre_bnfs, absT_infos), _), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels, xtor_co_rec_thms, xtor_rel_co_induct, dtor_set_inducts, xtor_co_rec_transfers, xtor_co_rec_o_maps, ...}, lthy)) = fixpoint_bnf false I (construct_fp mixfixes map_bs rel_bs pred_bs set_bss) fp_bs (map dest_TFree As) (map dest_TFree killed_As) (map dest_TFree Xs) ctrXs_repTs empty_comp_cache lthy handle BAD_DEAD (X, X_backdrop) => (case X_backdrop of Type (bad_tc, _) => let val fake_T = qsoty (unfreeze_fp X); val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " in type expression " ^ fake_T_backdrop) else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) bad_tc) then error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) else error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) end); val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_nesting_rel_eq_onps = map rel_eq_onp_of_bnf fp_nesting_bnfs; val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val live_nesting_rel_eq_onps = map rel_eq_onp_of_bnf live_nesting_bnfs; val liveness = liveness_of_fp_bnf num_As any_fp_bnf; val live = live_of_bnf any_fp_bnf; val _ = if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs @ pred_bs) then warning "Map function, relator, and predicator names ignored" else (); val Bs = @{map 3} (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree_or_tvar S B else A) liveness As Bs0; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val live_AsBs = filter (op <>) (As ~~ Bs); val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_injects = map #abs_inject absT_infos; val abs_inverses = map #abs_inverse absT_infos; val type_definitions = map #type_definition absT_infos; val ctors = map (mk_ctor As) ctors0; val dtors = map (mk_dtor As) dtors0; val fpTs = map (domain_type o fastype_of) dtors; val fpBTs = map B_ify_T fpTs; val real_unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fpTs); val ctr_Tsss = map (map (map real_unfreeze_fp)) ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun define_ctrs_dtrs_for_type_etc fp_bnf fp_b fpT C E ctor dtor xtor_co_rec ctor_dtor dtor_ctor ctor_inject pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm n ks ms abs abs_inject type_definition ctr_bindings ctr_mixfixes ctr_Tss disc_bindings sel_bindingss raw_sel_default_eqs lthy = let val fp_b_name = Binding.name_of fp_b; val ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) = define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy; val ctrs = map (mk_ctr As) ctrs0; val sel_default_eqs = let val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss; val sel_bTs = flat sel_bindingss ~~ flat sel_Tss |> filter_out (Binding.is_empty o fst) |> distinct (Binding.eq_name o apply2 fst); val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy in map (prepare_term sel_default_lthy) raw_sel_default_eqs end; fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b); fun massage_res (ctr_sugar, maps_sets_rels) = (maps_sets_rels, (ctrs, xss, ctor_iff_dtor_thm, ctr_defs, ctr_sugar)); in (wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs #> (fn (ctr_sugar, lthy) => derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps [] fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm [] [] [] ctr_Tss ctr_sugar lthy |>> pair ctr_sugar) ##>> (if fp = Least_FP then define_rec (the recs_args_types) mk_binding fpTs Cs reps else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec #>> apfst massage_res, lthy) end; fun wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types (wrap_one_etcs, lthy) = fold_map I wrap_one_etcs lthy |>> apsnd split_list o apfst (apsnd @{split_list 5} o apfst @{split_list 17} o split_list) o split_list; fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects rel_distincts set_thmss = injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @ set_thmss; fun mk_co_rec_transfer_goals lthy co_recs = let val BE_ify = Term.subst_atomic_types (live_AsBs @ (Cs ~~ Es)); val ((Rs, Ss), names_lthy) = lthy |> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> mk_Frees "S" (map2 mk_pred2T Cs Es); val co_recBs = map BE_ify co_recs; in (Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy) end; fun derive_rec_transfer_thms lthy recs rec_defs (SOME (_, _, _, xsssss)) = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy recs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced nn end; fun derive_rec_o_map_thmss lthy recs rec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_gs = AList.find (op =) (fs ~~ liveness) true; val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs)); val num_rec_args = length rec_arg_Ts; val g_Ts = map B_ify_T rec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names g_Ts; val grecs = map (fn recx => Term.list_comb (Term.map_types B_ify_T recx, gs)) recs; val rec_o_map_lhss = map2 (curry HOLogic.mk_comp) grecs gmaps; val ABfs = (As ~~ Bs) ~~ fs; fun mk_rec_arg_arg (x as Free (_, T)) = let val U = B_ify_T T in if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x end; fun mk_rec_o_map_arg rec_arg_T h = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val xs' = map mk_rec_arg_arg xs; in fold_rev Term.lambda xs (Term.list_comb (h, xs')) end; fun mk_rec_o_map_rhs recx = let val args = map2 mk_rec_o_map_arg rec_arg_Ts gs in Term.list_comb (recx, args) end; val rec_o_map_rhss = map mk_rec_o_map_rhs recs; val rec_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss; val rec_o_map_thms = @{map 3} (fn goal => fn rec_def => fn ctor_rec_o_map => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses ctor_rec_o_map) |> Thm.close_derivation \<^here>) rec_o_map_goals rec_defs xtor_co_rec_o_maps; in map single rec_o_map_thms end; fun derive_note_induct_recs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (ctrss, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (recs, rec_defs)), lthy) = let val ((induct_thms, induct_thm, induct_attrs), (rec_thmss, rec_attrs)) = derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy; val rec_transfer_thmss = map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types); - val induct_type_attr = Attrib.internal o K o Induct.induct_type; - val induct_pred_attr = Attrib.internal o K o Induct.induct_pred; + val induct_type_attr = Attrib.internal \<^here> o K o Induct.induct_type; + val induct_pred_attr = Attrib.internal \<^here> o K o Induct.induct_pred; val ((rel_induct_thmss, common_rel_induct_thms), (rel_induct_attrs, common_rel_induct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) = derive_rel_induct_thms_for_types lthy nn fpTs As Bs ctrss ctr_Tsss (map #exhaust ctr_sugars) xtor_rel_co_induct ctr_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs; in ((map single rel_induct_thms, single common_rel_induct_thm), (rel_induct_attrs, rel_induct_attrs)) end; val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (if nn > 1 then [(inductN, [induct_thm], K induct_attrs), (rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]), (recN, rec_thmss, K rec_attrs), (rec_o_mapN, rec_o_map_thmss, K []), (rec_transferN, rec_transfer_thmss, K []), (rel_inductN, rel_induct_thmss, K (rel_induct_attrs @ [induct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational recs (flat rec_thmss) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (flat rec_thmss)) |> Local_Theory.notes (common_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms (fst (dest_Type (hd fpTs)) ^ implode (map (prefix "_") (tl fp_b_names))) inductN |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars recs rec_defs map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn []) (replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss (replicate nn []) (replicate nn []) rec_transfer_thmss common_rel_induct_thms rel_induct_thmss [] (replicate nn []) rec_o_map_thmss end; fun derive_corec_transfer_thms lthy corecs corec_defs = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy corecs; val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs (flat pgss) pss qssss gssss) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; fun derive_map_o_corec_thmss lthy0 lthy2 corecs corec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_fs = AList.find (op =) (fs ~~ liveness) true; val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0; val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs)); val num_rec_args = length corec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names corec_arg_Ts; val gcorecs = map (fn corecx => Term.list_comb (corecx, gs)) corecs; val map_o_corec_lhss = map2 (curry HOLogic.mk_comp) fmaps gcorecs; val ABfs = (As ~~ Bs) ~~ fs; fun mk_map_o_corec_arg corec_argB_T g = let val T = range_type (fastype_of g); val U = range_type corec_argB_T; in if T = U then g else HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g) end; fun mk_map_o_corec_rhs corecx = let val args = map2 (mk_map_o_corec_arg o B_ify_T) corec_arg_Ts gs in Term.list_comb (B_ify corecx, args) end; val map_o_corec_rhss = map mk_map_o_corec_rhs corecs; val map_o_corec_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) map_o_corec_lhss map_o_corec_rhss; val map_o_corec_thms = @{map 3} (fn goal => fn corec_def => fn dtor_map_o_corec => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt corec_def pre_map_defs live_nesting_map_ident0s abs_inverses dtor_map_o_corec) |> Thm.close_derivation \<^here>) map_o_corec_goals corec_defs xtor_co_rec_o_maps; in map single map_o_corec_thms end; fun derive_note_coinduct_corecs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (_, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (corecs, corec_defs)), lthy) = let val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)], (coinduct_attrs, common_coinduct_attrs)), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) = derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs; fun distinct_prems ctxt thm = Goal.prove (*no sorry*) ctxt [] [] (thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies) (fn _ => HEADGOAL (cut_tac thm THEN' assume_tac ctxt) THEN ALLGOALS eq_assume_tac); fun eq_ifIN _ [thm] = thm | eq_ifIN ctxt (thm :: thms) = distinct_prems ctxt (@{thm eq_ifI} OF (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]}) [thm, eq_ifIN ctxt thms])); val corec_code_thms = map (eq_ifIN lthy) corec_thmss; val corec_sel_thmss = map flat corec_sel_thmsss; - val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type; - val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred; + val coinduct_type_attr = Attrib.internal \<^here> o K o Induct.coinduct_type; + val coinduct_pred_attr = Attrib.internal \<^here> o K o Induct.coinduct_pred; val flat_corec_thms = append oo append; val corec_transfer_thmss = map single (derive_corec_transfer_thms lthy corecs corec_defs); val ((rel_coinduct_thmss, common_rel_coinduct_thms), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_coinduct_thms, common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = derive_rel_coinduct_thms_for_types lthy nn fpTs ns As Bs mss ctr_sugars abs_inverses abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss xtor_rel_co_induct live_nesting_rel_eqs; in ((map single rel_coinduct_thms, single common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) end; val map_o_corec_thmss = derive_map_o_corec_thmss lthy lthy corecs corec_defs; val (set_induct_thms, set_induct_attrss) = derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars) (map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_inducts (map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) dtor_ctors abs_inverses |> split_list; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars (@{map 3} flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss) map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (set_inductN, set_induct_thms, nth set_induct_attrss) :: (if nn > 1 then [(coinductN, [coinduct_thm], K common_coinduct_attrs), (coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs), (rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(coinductN, map single coinduct_thms, fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]), (coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs), (corecN, corec_thmss, K []), (corec_codeN, map single corec_code_thms, K (nitpicksimp_attrs)), (corec_discN, corec_disc_thmss, K []), (corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs), (corec_selN, corec_sel_thmss, K corec_sel_attrs), (corec_transferN, corec_transfer_thmss, K []), (map_o_corecN, map_o_corec_thmss, K []), (rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> fold (Spec_Rules.add Binding.empty Spec_Rules.equational corecs) [flat corec_sel_thmss, flat corec_thmss] |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) corec_code_thms) |> Local_Theory.notes (common_notes @ notes) |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars corecs corec_defs map_thmss [coinduct_thm, coinduct_strong_thm] (transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss corec_disc_iff_thmss (map single corec_code_thms) corec_transfer_thmss common_rel_coinduct_thms rel_coinduct_thmss set_induct_thms (replicate nn (if nn = 1 then set_induct_thms else [])) map_o_corec_thmss end; val lthy = lthy |> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs |> @{fold_map 29} define_ctrs_dtrs_for_type_etc fp_bnfs fp_bs fpTs Cs Es ctors dtors xtor_co_recs ctor_dtors dtor_ctors ctor_injects pre_map_defs pre_set_defss pre_rel_defs xtor_maps xtor_setss xtor_rels ns kss mss abss abs_injects type_definitions ctr_bindingss ctr_mixfixess ctr_Tsss disc_bindingss sel_bindingsss raw_sel_default_eqss |> wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types |> case_fp fp derive_note_induct_recs_thms_for_types derive_note_coinduct_corecs_thms_for_types; val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^ co_prefix fp ^ "datatype")); in lthy end; fun co_datatypes fp = define_co_datatypes (K I) (K I) (K I) (K I) fp; fun co_datatype_cmd fp construct_fp options lthy = define_co_datatypes Plugin_Name.make_filter Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term fp construct_fp options lthy handle EMPTY_DATATYPE s => error ("Cannot define empty datatype " ^ quote s); val parse_ctr_arg = \<^keyword>\(\ |-- parse_binding_colon -- Parse.typ --| \<^keyword>\)\ || Parse.typ >> pair Binding.empty; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.binding parse_ctr_arg -- Parse.opt_mixfix); val parse_spec = parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- parse_ctr_specs) -- parse_map_rel_pred_bindings -- parse_sel_default_eqs; val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec; fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp; end; diff --git a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_n2m_sugar.ML @@ -1,519 +1,519 @@ (* Title: HOL/Tools/BNF/bnf_fp_n2m_sugar.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2013 Suggared flattening of nested to mutual (co)recursion. *) signature BNF_FP_N2M_SUGAR = sig val unfold_lets_splits: term -> term val unfold_splits_lets: term -> term val dest_map: Proof.context -> string -> term -> term * term list val dest_pred: Proof.context -> string -> term -> term * term list val mutualize_fp_sugars: (string -> bool) -> BNF_Util.fp_kind -> int list -> binding list -> typ list -> term list -> term list list list list -> BNF_FP_Def_Sugar.fp_sugar list -> local_theory -> (BNF_FP_Def_Sugar.fp_sugar list * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option)) * local_theory val nested_to_mutual_fps: (string -> bool) -> BNF_Util.fp_kind -> binding list -> typ list -> term list -> (term * term list list) list list -> local_theory -> (typ list * int list * BNF_FP_Def_Sugar.fp_sugar list * (BNF_FP_Def_Sugar.lfp_sugar_thms option * BNF_FP_Def_Sugar.gfp_sugar_thms option)) * local_theory end; structure BNF_FP_N2M_Sugar : BNF_FP_N2M_SUGAR = struct open Ctr_Sugar open BNF_Util open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_FP_N2M val n2mN = "n2m_"; type n2m_sugar = fp_sugar list * (lfp_sugar_thms option * gfp_sugar_thms option); structure Data = Generic_Data ( type T = n2m_sugar Typtab.table; val empty = Typtab.empty; fun merge data : T = Typtab.merge (K true) data; ); fun morph_n2m_sugar phi (fp_sugars, (lfp_sugar_thms_opt, gfp_sugar_thms_opt)) = (map (morph_fp_sugar phi) fp_sugars, (Option.map (morph_lfp_sugar_thms phi) lfp_sugar_thms_opt, Option.map (morph_gfp_sugar_thms phi) gfp_sugar_thms_opt)); val transfer_n2m_sugar = morph_n2m_sugar o Morphism.transfer_morphism; fun n2m_sugar_of ctxt = Typtab.lookup (Data.get (Context.Proof ctxt)) #> Option.map (transfer_n2m_sugar (Proof_Context.theory_of ctxt)); fun register_n2m_sugar key n2m_sugar = - Local_Theory.declaration {syntax = false, pervasive = false} + Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Data.map (Typtab.update (key, morph_n2m_sugar phi n2m_sugar))); fun unfold_lets_splits (Const (\<^const_name>\Let\, _) $ t $ u) = unfold_lets_splits (betapply (unfold_splits_lets u, t)) | unfold_lets_splits (t $ u) = betapply (unfold_lets_splits t, unfold_lets_splits u) | unfold_lets_splits (Abs (s, T, t)) = Abs (s, T, unfold_lets_splits t) | unfold_lets_splits t = t and unfold_splits_lets ((t as Const (\<^const_name>\case_prod\, _)) $ u) = (case unfold_splits_lets u of u' as Abs (s1, T1, Abs (s2, T2, _)) => let val v = Var ((s1 ^ s2, Term.maxidx_of_term u' + 1), HOLogic.mk_prodT (T1, T2)) in lambda v (incr_boundvars 1 (betapplys (u', [HOLogic.mk_fst v, HOLogic.mk_snd v]))) end | _ => t $ unfold_lets_splits u) | unfold_splits_lets (t as Const (\<^const_name>\Let\, _) $ _ $ _) = unfold_lets_splits t | unfold_splits_lets (t $ u) = betapply (unfold_splits_lets t, unfold_lets_splits u) | unfold_splits_lets (Abs (s, T, t)) = Abs (s, T, unfold_splits_lets t) | unfold_splits_lets t = unfold_lets_splits t; fun dest_either_map_or_pred map_or_pred_of_bnf ctxt T_name call = let val bnf = the (bnf_of ctxt T_name); val const0 = map_or_pred_of_bnf bnf; val live = live_of_bnf bnf; val (f_Ts, _) = strip_typeN live (fastype_of const0); val fs = map_index (fn (i, T) => Var (("f", i), T)) f_Ts; val pat = betapplys (const0, fs); val (_, tenv) = fo_match ctxt call pat; in (const0, Vartab.fold_rev (fn (_, (_, f)) => cons f) tenv []) end; val dest_map = dest_either_map_or_pred map_of_bnf; val dest_pred = dest_either_map_or_pred pred_of_bnf; fun dest_map_or_pred ctxt T_name call = (case try (dest_map ctxt T_name) call of SOME res => res | NONE => dest_pred ctxt T_name call); fun dest_abs_or_applied_map_or_pred _ _ (Abs (_, _, t)) = (Term.dummy, [t]) | dest_abs_or_applied_map_or_pred ctxt s (t1 $ _) = dest_map_or_pred ctxt s t1; fun map_partition f xs = fold_rev (fn x => fn (ys, (good, bad)) => case f x of SOME y => (y :: ys, (x :: good, bad)) | NONE => (ys, (good, x :: bad))) xs ([], ([], [])); fun key_of_fp_eqs fp As fpTs Xs ctrXs_repTs = Type (case_fp fp "l" "g", fpTs @ Xs @ ctrXs_repTs) |> Term.map_atyps (fn T as TFree (_, S) => (case find_index (curry (op =) T) As of ~1 => T | j => TFree ("'" ^ string_of_int j, S))); fun get_indices callers t = callers |> map_index (fn (i, v) => if exists_subterm (equal v) t then SOME i else NONE) |> map_filter I; fun mutualize_fp_sugars plugins fp mutual_cliques bs fpTs callers callssss fp_sugars0 no_defs_lthy = let val thy = Proof_Context.theory_of no_defs_lthy; val qsotm = quote o Syntax.string_of_term no_defs_lthy; fun incompatible_calls ts = error ("Incompatible " ^ co_prefix fp ^ "recursive calls: " ^ commas (map qsotm ts)); fun mutual_self_call caller t = error ("Unsupported mutual self-call " ^ qsotm t ^ " from " ^ qsotm caller); fun nested_self_call t = error ("Unsupported nested self-call " ^ qsotm t); val b_names = map Binding.name_of bs; val fp_b_names = map base_name_of_typ fpTs; val nn = length fpTs; val kks = 0 upto nn - 1; fun target_ctr_sugar_of_fp_sugar fpT ({T, fp_ctr_sugar = {ctr_sugar, ...}, ...} : fp_sugar) = let val rho = Vartab.fold (cons o apsnd snd) (Sign.typ_match thy (T, fpT) Vartab.empty) []; val phi = Morphism.term_morphism "BNF" (Term.subst_TVars rho); in morph_ctr_sugar phi ctr_sugar end; val ctor_iff_dtors = map (#ctor_iff_dtor o #fp_ctr_sugar) fp_sugars0; val ctr_defss = map (#ctr_defs o #fp_ctr_sugar) fp_sugars0; val mapss = map (#map_thms o #fp_bnf_sugar) fp_sugars0; val ctr_sugars = map2 target_ctr_sugar_of_fp_sugar fpTs fp_sugars0; val ctrss = map #ctrs ctr_sugars; val ctr_Tss = map (map fastype_of) ctrss; val As' = fold (fold Term.add_tfreesT) ctr_Tss []; val As = map TFree As'; val ((Cs, Xs), _) = no_defs_lthy |> fold Variable.declare_typ As |> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun check_call_dead live_call call = if null (get_indices callers call) then () else incompatible_calls [live_call, call]; fun freeze_fpTs_type_based_default (T as Type (s, Ts)) = (case filter (curry (op =) T o snd) (map_index I fpTs) of [(kk, _)] => nth Xs kk | _ => Type (s, map freeze_fpTs_type_based_default Ts)) | freeze_fpTs_type_based_default T = T; fun freeze_fpTs_mutual_call kk fpT calls T = (case fold (union (op =)) (map (get_indices callers) calls) [] of [] => if T = fpT then nth Xs kk else freeze_fpTs_type_based_default T | [kk'] => if T = fpT andalso kk' <> kk then mutual_self_call (nth callers kk) (the (find_first (not o null o get_indices callers) calls)) else if T = nth fpTs kk' then nth Xs kk' else freeze_fpTs_type_based_default T | _ => incompatible_calls calls); fun freeze_fpTs_map kk (fpT as Type (_, Ts')) (callss, (live_call :: _, dead_calls)) (Type (s, Ts)) = if Ts' = Ts then nested_self_call live_call else (List.app (check_call_dead live_call) dead_calls; Type (s, map2 (freeze_fpTs_call kk fpT) (flatten_type_args_of_bnf (the (bnf_of no_defs_lthy s)) [] (transpose callss)) Ts)) and freeze_fpTs_call kk fpT calls (T as Type (s, _)) = (case map_partition (try (snd o dest_map no_defs_lthy s)) calls of ([], _) => (case map_partition (try (snd o dest_abs_or_applied_map_or_pred no_defs_lthy s)) calls of ([], _) => freeze_fpTs_mutual_call kk fpT calls T | callsp => freeze_fpTs_map kk fpT callsp T) | callsp => freeze_fpTs_map kk fpT callsp T) | freeze_fpTs_call _ _ _ T = T; val ctr_Tsss = map (map binder_types) ctr_Tss; val ctrXs_Tsss = @{map 4} (map2 o map2 oo freeze_fpTs_call) kks fpTs callssss ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val key = key_of_fp_eqs fp As fpTs Xs ctrXs_repTs; in (case n2m_sugar_of no_defs_lthy key of SOME n2m_sugar => (n2m_sugar, no_defs_lthy) | NONE => let val base_fp_names = Name.variant_list [] fp_b_names; val fp_bs = map2 (fn b_name => fn base_fp_name => Binding.qualify true b_name (Binding.name (n2mN ^ base_fp_name))) b_names base_fp_names; val Type (s, Us) :: _ = fpTs; val killed_As' = (case bnf_of no_defs_lthy s of NONE => As' | SOME bnf => let val Type (_, Ts) = T_of_bnf bnf; val deads = deads_of_bnf bnf; val dead_Us = map_filter (fn (T, U) => if member (op =) deads T then SOME U else NONE) (Ts ~~ Us); in fold Term.add_tfreesT dead_Us [] end); val fp_absT_infos = map #absT_info fp_sugars0; val indexed_fp_ress = map (apsnd #fp_res o `(#fp_res_index)) fp_sugars0; val (((pre_bnfs, absT_infos), _), (fp_res as {xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_injects, dtor_ctors, xtor_co_rec_thms, ...}, lthy)) = fixpoint_bnf false I (construct_mutualized_fp fp mutual_cliques fpTs indexed_fp_ress) fp_bs As' killed_As' (map dest_TFree Xs) ctrXs_repTs empty_comp_cache no_defs_lthy; val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_abs_inverses = map #abs_inverse fp_absT_infos; val fp_type_definitions = map #type_definition fp_absT_infos; val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_inverses = map #abs_inverse absT_infos; val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun mk_binding b pre = Binding.prefix_name (pre ^ "_") b; val ((co_recs, co_rec_defs), lthy) = @{fold_map 2} (fn b => if fp = Least_FP then define_rec (the recs_args_types) (mk_binding b) fpTs Cs reps else define_corec (the corecs_args_types) (mk_binding b) fpTs Cs abss) fp_bs xtor_co_recs lthy |>> split_list; val timer = time (timer ("High-level " ^ co_prefix fp ^ "recursors")); val ((common_co_inducts, co_inductss', co_rec_thmss, co_rec_disc_thmss, co_rec_sel_thmsss), fp_sugar_thms) = if fp = Least_FP then derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss fp_abs_inverses fp_type_definitions abs_inverses ctrss ctr_defss co_recs co_rec_defs lthy |> `(fn ((inducts, induct, _), (rec_thmss, _)) => ([induct], [inducts], rec_thmss, replicate nn [], replicate nn [])) ||> (fn info => (SOME info, NONE)) else derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns fp_abs_inverses abs_inverses (fn thm => thm RS @{thm vimage2p_refl}) ctr_defss ctr_sugars co_recs co_rec_defs |> `(fn ((coinduct_thms_pairs, _), corec_thmss, corec_disc_thmss, _, (corec_sel_thmsss, _)) => (map snd coinduct_thms_pairs, map fst coinduct_thms_pairs, corec_thmss, corec_disc_thmss, corec_sel_thmsss)) ||> (fn info => (NONE, SOME info)); val timer = time (timer ("High-level " ^ co_prefix fp ^ "induction principles")); val names_lthy = lthy |> fold Variable.declare_typ (As @ Cs @ Xs); val phi = Proof_Context.export_morphism names_lthy lthy; fun mk_target_fp_sugar T X kk pre_bnf absT_info ctrXs_Tss ctor_iff_dtor ctr_defs ctr_sugar co_rec co_rec_def map_thms co_inducts co_rec_thms co_rec_disc_thms co_rec_sel_thmss ({fp_ctr_sugar = {ctr_transfers, case_transfers, disc_transfers, sel_transfers, ...}, fp_bnf_sugar = {map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases, ...}, fp_co_induct_sugar = SOME {co_rec_disc_iffs, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts, ...}, ...} : fp_sugar) = {T = T, BT = T (*wrong but harmless*), X = X, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = pre_bnf, fp_bnf = nth (#bnfs fp_res) kk, absT_info = absT_info, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = ctrXs_Tss, ctor_iff_dtor = ctor_iff_dtor, ctr_defs = ctr_defs, ctr_sugar = ctr_sugar, ctr_transfers = ctr_transfers, case_transfers = case_transfers, disc_transfers = disc_transfers, sel_transfers = sel_transfers}, fp_bnf_sugar = {map_thms = map_thms, map_disc_iffs = map_disc_iffs, map_selss = map_selss, rel_injects = rel_injects, rel_distincts = rel_distincts, rel_sels = rel_sels, rel_intros = rel_intros, rel_cases = rel_cases, pred_injects = pred_injects, set_thms = set_thms, set_selssss = set_selssss, set_introssss = set_introssss, set_cases = set_cases}, fp_co_induct_sugar = SOME {co_rec = co_rec, common_co_inducts = common_co_inducts, co_inducts = co_inducts, co_rec_def = co_rec_def, co_rec_thms = co_rec_thms, co_rec_discs = co_rec_disc_thms, co_rec_disc_iffs = co_rec_disc_iffs, co_rec_selss = co_rec_sel_thmss, co_rec_codes = co_rec_codes, co_rec_transfers = co_rec_transfers, co_rec_o_maps = co_rec_o_maps, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = rel_co_inducts, common_set_inducts = common_set_inducts, set_inducts = set_inducts}} |> morph_fp_sugar phi; val target_fp_sugars = @{map 17} mk_target_fp_sugar fpTs Xs kks pre_bnfs absT_infos ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs mapss (transpose co_inductss') co_rec_thmss co_rec_disc_thmss co_rec_sel_thmsss fp_sugars0; val n2m_sugar = (target_fp_sugars, fp_sugar_thms); in (n2m_sugar, lthy |> register_n2m_sugar key n2m_sugar) end) end; fun indexify_callsss ctrs callsss = let fun indexify_ctr ctr = (case AList.lookup Term.aconv_untyped callsss ctr of NONE => replicate (num_binder_types (fastype_of ctr)) [] | SOME callss => map (map (Envir.beta_eta_contract o unfold_lets_splits)) callss); in map indexify_ctr ctrs end; fun retypargs tyargs (Type (s, _)) = Type (s, tyargs); fun fold_subtype_pairs f (T as Type (s, Ts), U as Type (s', Us)) = f (T, U) #> (if s = s' then fold (fold_subtype_pairs f) (Ts ~~ Us) else I) | fold_subtype_pairs f TU = f TU; val impossible_caller = Bound ~1; fun nested_to_mutual_fps plugins fp actual_bs actual_Ts actual_callers actual_callssss0 lthy = let val qsoty = quote o Syntax.string_of_typ lthy; val qsotys = space_implode " or " o map qsoty; fun not_co_datatype0 T = error (qsoty T ^ " is not a " ^ co_prefix fp ^ "datatype"); fun not_co_datatype (T as Type (s, _)) = if fp = Least_FP andalso is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) s) then error (qsoty T ^ " is an old-style datatype") else not_co_datatype0 T | not_co_datatype T = not_co_datatype0 T; fun not_mutually_nested_rec Ts1 Ts2 = error (qsotys Ts1 ^ " is neither mutually " ^ co_prefix fp ^ "recursive with " ^ qsotys Ts2 ^ " nor nested " ^ co_prefix fp ^ "recursive through " ^ (if Ts1 = Ts2 andalso length Ts1 = 1 then "itself" else qsotys Ts2)); val sorted_actual_Ts = sort (prod_ord int_ord Term_Ord.typ_ord o apply2 (`Term.size_of_typ)) actual_Ts; fun the_ctrs_of (Type (s, Ts)) = map (mk_ctr Ts) (#ctrs (the (ctr_sugar_of lthy s))); fun gen_rhss_in gen_Ts rho (subTs as Type (_, sub_tyargs) :: _) = let fun maybe_insert (T, Type (_, gen_tyargs)) = member (op =) subTs T ? insert (op =) gen_tyargs | maybe_insert _ = I; val gen_ctrs = maps the_ctrs_of gen_Ts; val gen_ctr_Ts = maps (binder_types o fastype_of) gen_ctrs; val ctr_Ts = map (Term.typ_subst_atomic rho) gen_ctr_Ts; in (case fold (fold_subtype_pairs maybe_insert) (ctr_Ts ~~ gen_ctr_Ts) [] of [] => [map (typ_subst_nonatomic (map swap rho)) sub_tyargs] | gen_tyargss => gen_tyargss) end; fun the_fp_sugar_of (T as Type (T_name, _)) = (case fp_sugar_of lthy T_name of SOME (fp_sugar as {fp = fp', fp_co_induct_sugar = SOME _, ...}) => if fp = fp' then fp_sugar else not_co_datatype T | _ => not_co_datatype T); fun gather_types _ _ rev_seens gen_seen [] = (rev rev_seens, gen_seen) | gather_types lthy rho rev_seens gen_seen ((T as Type (_, tyargs)) :: Ts) = let val {T = Type (_, tyargs0), fp_res = {Ts = mutual_Ts0, ...}, ...} = the_fp_sugar_of T; val mutual_Ts = map (retypargs tyargs) mutual_Ts0; val rev_seen = flat rev_seens; val _ = null rev_seens orelse exists (exists_strict_subtype_in rev_seen) mutual_Ts orelse not_mutually_nested_rec mutual_Ts rev_seen; fun fresh_tyargs () = let val (unsorted_gen_tyargs, lthy') = variant_tfrees (replicate (length tyargs) "z") lthy |>> map Logic.varifyT_global; val gen_tyargs = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) tyargs0 unsorted_gen_tyargs; val rho' = (gen_tyargs ~~ tyargs) @ rho; in (rho', gen_tyargs, gen_seen, lthy') end; val (rho', gen_tyargs, gen_seen', lthy') = if exists (exists_subtype_in (flat rev_seens)) mutual_Ts then (case gen_rhss_in gen_seen rho mutual_Ts of [] => fresh_tyargs () | gen_tyargs :: gen_tyargss_tl => let val unify_pairs = split_list (maps (curry (op ~~) gen_tyargs) gen_tyargss_tl); val mgu = Type.raw_unifys unify_pairs Vartab.empty; val gen_tyargs' = map (Envir.norm_type mgu) gen_tyargs; val gen_seen' = map (Envir.norm_type mgu) gen_seen; in (rho, gen_tyargs', gen_seen', lthy) end) else fresh_tyargs (); val gen_mutual_Ts = map (retypargs gen_tyargs) mutual_Ts0; val other_mutual_Ts = remove1 (op =) T mutual_Ts; val Ts' = fold (remove1 (op =)) other_mutual_Ts Ts; in gather_types lthy' rho' (mutual_Ts :: rev_seens) (gen_seen' @ gen_mutual_Ts) Ts' end | gather_types _ _ _ _ (T :: _) = not_co_datatype T; val (perm_Tss, perm_gen_Ts) = gather_types lthy [] [] [] sorted_actual_Ts; val (perm_mutual_cliques, perm_Ts) = split_list (flat (map_index (fn (i, perm_Ts) => map (pair i) perm_Ts) perm_Tss)); val perm_frozen_gen_Ts = map Logic.unvarifyT_global perm_gen_Ts; val missing_Ts = fold (remove1 (op =)) actual_Ts perm_Ts; val Ts = actual_Ts @ missing_Ts; val nn = length Ts; val kks = 0 upto nn - 1; val callssss0 = pad_list [] nn actual_callssss0; val common_name = mk_common_name (map Binding.name_of actual_bs); val bs = pad_list (Binding.name common_name) nn actual_bs; val callers = pad_list impossible_caller nn actual_callers; fun permute xs = permute_like (op =) Ts perm_Ts xs; fun unpermute perm_xs = permute_like (op =) perm_Ts Ts perm_xs; (* The assignment of callers to mutual cliques is incorrect in general. This code would need to inspect the actual calls to discover the correct cliques in the presence of type duplicates. However, the naive scheme implemented here supports "prim(co)rec" specifications following reasonable ordering of the duplicates (e.g., keeping the cliques together). *) val perm_bs = permute bs; val perm_callers = permute callers; val perm_kks = permute kks; val perm_callssss0 = permute callssss0; val perm_fp_sugars0 = map (the o fp_sugar_of lthy o fst o dest_Type) perm_Ts; val perm_callssss = map2 (indexify_callsss o #ctrs o #ctr_sugar o #fp_ctr_sugar) perm_fp_sugars0 perm_callssss0; val ((perm_fp_sugars, fp_sugar_thms), lthy) = if length perm_Tss = 1 then ((perm_fp_sugars0, (NONE, NONE)), lthy) else mutualize_fp_sugars plugins fp perm_mutual_cliques perm_bs perm_frozen_gen_Ts perm_callers perm_callssss perm_fp_sugars0 lthy; val fp_sugars = unpermute perm_fp_sugars; in ((missing_Ts, perm_kks, fp_sugars, fp_sugar_thms), lthy) end; end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec.ML b/src/HOL/Tools/BNF/bnf_gfp_grec.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec.ML @@ -1,3228 +1,3228 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec.ML Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Aymeric Bouzy, Ecole polytechnique Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor construction. *) signature BNF_GFP_GREC = sig val Tsubst: typ -> typ -> typ -> typ val substT: typ -> typ -> term -> term val freeze_types: Proof.context -> (indexname * sort) list -> typ list -> typ list val dummify_atomic_types: term -> term val define_const: bool -> binding -> int -> string -> term -> local_theory -> (term * thm) * local_theory type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table} val map_buffer: (term -> term) -> buffer -> buffer val specialize_buffer_types: buffer -> buffer type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list} type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: BNF_FP_Def_Sugar.fp_sugar list, ssig_fp_sugar: BNF_FP_Def_Sugar.fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: BNF_Def.bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info} type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm} val not_codatatype: Proof.context -> typ -> 'a val mk_fp_binding: binding -> string -> binding val bnf_kill_all_but: int -> BNF_Def.bnf -> local_theory -> BNF_Def.bnf * local_theory val print_corec_infos: Proof.context -> unit val has_no_corec_info: Proof.context -> string -> bool val corec_info_of: typ -> local_theory -> corec_info * local_theory val maybe_corec_info_of: Proof.context -> typ -> corec_info option val corec_infos_of: Proof.context -> string -> corec_info list val corec_infos_of_generic: Context.generic -> Symtab.key -> corec_info list val prepare_friend_corec: string -> typ -> local_theory -> (corec_info * binding * int * typ * typ * typ * typ * typ * BNF_Def.bnf * BNF_Def.bnf * BNF_FP_Def_Sugar.fp_sugar * BNF_FP_Def_Sugar.fp_sugar * buffer) * local_theory val register_friend_corec: string -> binding -> int -> typ -> typ -> typ -> BNF_Def.bnf -> BNF_FP_Def_Sugar.fp_sugar -> BNF_FP_Def_Sugar.fp_sugar -> term -> term -> thm -> corec_info -> local_theory -> friend_info * local_theory end; structure BNF_GFP_Grec : BNF_GFP_GREC = struct open Ctr_Sugar open BNF_Util open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_LFP open BNF_FP_Def_Sugar open BNF_LFP_Rec_Sugar open BNF_GFP_Grec_Tactics val algLamN = "algLam"; val algLam_algLamN = "algLam_algLam"; val algLam_algrhoN = "algLam_algrho"; val algrhoN = "algrho"; val CLeafN = "CLeaf"; val congN = "congclp"; val cong_alg_introsN = "cong_alg_intros"; val cong_localeN = "cong_locale"; val corecUUN = "corecUU"; val corecUU_transferN = "corecUU_transfer"; val corecUU_uniqueN = "corecUU_unique"; val cutSsigN = "cutSsig"; val dtor_algLamN = "dtor_algLam"; val dtor_algrhoN = "dtor_algrho"; val dtor_coinductN = "dtor_coinduct"; val dtor_transferN = "dtor_transfer"; val embLN = "embL"; val embLLN = "embLL"; val embLRN = "embLR"; val embL_pointful_naturalN = "embL_pointful_natural"; val embL_transferN = "embL_transfer"; val equivp_RetrN = "equivp_Retr"; val evalN = "eval"; val eval_coreN = "eval_core"; val eval_core_pointful_naturalN = "eval_core_pointful_natural"; val eval_core_transferN = "eval_core_transfer"; val eval_flatN = "eval_flat"; val eval_simpsN = "eval_simps"; val flatN = "flat"; val flat_pointful_naturalN = "flat_pointful_natural"; val flat_transferN = "flat_transfer"; val k_as_ssig_naturalN = "k_as_ssig_natural"; val k_as_ssig_transferN = "k_as_ssig_transfer"; val LamN = "Lam"; val Lam_transferN = "Lam_transfer"; val Lam_pointful_naturalN = "Lam_pointful_natural"; val OperN = "Oper"; val proto_sctrN = "proto_sctr"; val proto_sctr_pointful_naturalN = "proto_sctr_pointful_natural"; val proto_sctr_transferN = "proto_sctr_transfer"; val rho_transferN = "rho_transfer"; val Retr_coinductN = "Retr_coinduct"; val sctrN = "sctr"; val sctr_transferN = "sctr_transfer"; val sctr_pointful_naturalN = "sctr_pointful_natural"; val sigN = "sig"; val SigN = "Sig"; val Sig_pointful_naturalN = "Sig_pointful_natural"; val corecUN = "corecU"; val corecU_ctorN = "corecU_ctor"; val corecU_uniqueN = "corecU_unique"; val unsigN = "unsig"; val VLeafN = "VLeaf"; val s_prefix = "s"; (* transforms "sig" into "ssig" *) fun not_codatatype ctxt T = error ("Not a codatatype: " ^ Syntax.string_of_typ ctxt T); fun mutual_codatatype () = error ("Mutually corecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\primcorec\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun noncorecursive_codatatype () = error ("Noncorecursive codatatypes are not supported (try " ^ quote (#1 \<^command_keyword>\definition\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")"); fun singleton_codatatype ctxt = error ("Singleton corecursive codatatypes are not supported (use " ^ quote (Syntax.string_of_typ ctxt \<^typ>\unit\) ^ " instead)"); fun merge_lists eq old1 old2 = (old1 |> subtract eq old2) @ old2; fun add_type_namesT (Type (s, Ts)) = insert (op =) s #> fold add_type_namesT Ts | add_type_namesT _ = I; fun Tsubst Y T = Term.typ_subst_atomic [(Y, T)]; fun substT Y T = Term.subst_atomic_types [(Y, T)]; fun freeze_types ctxt except_tvars Ts = let val As = fold Term.add_tvarsT Ts [] |> subtract (op =) except_tvars; val (Bs, _) = ctxt |> mk_TFrees' (map snd As); in map (Term.typ_subst_TVars (map fst As ~~ Bs)) Ts end; fun typ_unify_disjointly thy (T, T') = if T = T' then T else let val tvars = Term.add_tvar_namesT T []; val tvars' = Term.add_tvar_namesT T' []; val maxidx' = maxidx_of_typ T'; val T = T |> exists (member (op =) tvars') tvars ? Logic.incr_tvar (maxidx' + 1); val maxidx = Integer.max (maxidx_of_typ T) maxidx'; val (tyenv, _) = Sign.typ_unify thy (T, T') (Vartab.empty, maxidx); in Envir.subst_type tyenv T end; val dummify_atomic_types = Term.map_types (Term.map_atyps (K Term.dummyT)); fun mk_internal internal ctxt f = if internal andalso not (Config.get ctxt bnf_internals) then f else I fun mk_fp_binding fp_b pre = Binding.map_name (K pre) fp_b |> Binding.qualify true (Binding.name_of fp_b); fun mk_version_binding version = Binding.qualify false ("v" ^ string_of_int version); fun mk_version_fp_binding internal ctxt = mk_internal internal ctxt Binding.concealed ooo (mk_fp_binding oo mk_version_binding); (*FIXME: get rid of ugly names when typedef and primrec respect qualification*) fun mk_version_binding_ugly version = Binding.suffix_name ("_v" ^ string_of_int version); fun mk_version_fp_binding_ugly internal ctxt version fp_b pre = Binding.prefix_name (pre ^ "_") fp_b |> mk_version_binding_ugly version |> mk_internal internal ctxt Binding.concealed; fun mk_mapN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (TA --> TB) (map_of_bnf bnf) end; fun mk_relN ctxt live_AsBs TA bnf = let val TB = Term.typ_subst_atomic live_AsBs TA in enforce_type ctxt (snd o strip_typeN (length live_AsBs)) (mk_pred2T TA TB) (rel_of_bnf bnf) end; fun mk_map1 ctxt Y Z = mk_mapN ctxt [(Y, Z)]; fun mk_rel1 ctxt Y Z = mk_relN ctxt [(Y, Z)]; fun define_const internal fp_b version name rhs lthy = let val b = mk_version_fp_binding internal lthy version fp_b name; val ((free, (_, def_free)), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((b, NoSyn), ((Thm.def_binding b |> Binding.concealed, []), rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free), lthy) end; fun define_single_primrec b eqs lthy = let val (([free], [def_free], [simps_free]), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> primrec false [] [(b, NONE, NoSyn)] (map (fn eq => ((Binding.empty_atts, eq), [], [])) eqs) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val const = Morphism.term phi free; val const' = enforce_type lthy I (fastype_of free) const; in ((const', Morphism.thm phi def_free, map (Morphism.thm phi) simps_free), lthy) end; type buffer = {Oper: term, VLeaf: term, CLeaf: term, ctr_wrapper: term, friends: (typ * term) Symtab.table}; fun map_buffer f {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = {Oper = f Oper, VLeaf = f VLeaf, CLeaf = f CLeaf, ctr_wrapper = f ctr_wrapper, friends = Symtab.map (K (apsnd f)) friends}; fun morph_buffer phi = map_buffer (Morphism.term phi); fun specialize_buffer_types {Oper, VLeaf, CLeaf, ctr_wrapper, friends} = let val ssig_T as Type (_, Ts) = body_type (fastype_of VLeaf); val Y = List.last Ts; val ssigifyT = substT Y ssig_T; in {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ssigifyT ctr_wrapper, friends = Symtab.map (K (apsnd ssigifyT)) friends} end; type dtor_coinduct_info = {dtor_coinduct: thm, cong_def: thm, cong_locale: thm, cong_base: thm, cong_refl: thm, cong_sym: thm, cong_trans: thm, cong_alg_intros: thm list}; fun map_dtor_coinduct_info f {dtor_coinduct, cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros} = {dtor_coinduct = f dtor_coinduct, cong_def = f cong_def, cong_locale = f cong_locale, cong_base = f cong_base, cong_refl = f cong_refl, cong_sym = f cong_sym, cong_trans = f cong_trans, cong_alg_intros = map f cong_alg_intros}; fun morph_dtor_coinduct_info phi = map_dtor_coinduct_info (Morphism.thm phi); type corec_ad = {fpT: typ, friend_names: string list}; fun morph_corec_ad phi {fpT, friend_names} = {fpT = Morphism.typ phi fpT, friend_names = friend_names}; type corec_info = {fp_b: binding, version: int, fpT: typ, Y: typ, Z: typ, friend_names: string list, sig_fp_sugars: fp_sugar list, ssig_fp_sugar: fp_sugar, Lam: term, proto_sctr: term, flat: term, eval_core: term, eval: term, algLam: term, corecUU: term, dtor_transfer: thm, Lam_transfer: thm, Lam_pointful_natural: thm, proto_sctr_transfer: thm, flat_simps: thm list, eval_core_simps: thm list, eval_thm: thm, eval_simps: thm list, all_algLam_algs: thm list, algLam_thm: thm, dtor_algLam: thm, corecUU_thm: thm, corecUU_unique: thm, corecUU_transfer: thm, buffer: buffer, all_dead_k_bnfs: bnf list, Retr: term, equivp_Retr: thm, Retr_coinduct: thm, dtor_coinduct_info: dtor_coinduct_info}; fun morph_corec_info phi ({fp_b, version, fpT, Y, Z, friend_names, sig_fp_sugars, ssig_fp_sugar, Lam, proto_sctr, flat, eval_core, eval, algLam, corecUU, dtor_transfer, Lam_transfer, Lam_pointful_natural, proto_sctr_transfer, flat_simps, eval_core_simps, eval_thm, eval_simps, all_algLam_algs, algLam_thm, dtor_algLam, corecUU_thm, corecUU_unique, corecUU_transfer, buffer, all_dead_k_bnfs, Retr, equivp_Retr, Retr_coinduct, dtor_coinduct_info} : corec_info) = {fp_b = fp_b, version = version, fpT = Morphism.typ phi fpT, Y = Morphism.typ phi Y, Z = Morphism.typ phi Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugars (*no morph*), ssig_fp_sugar = ssig_fp_sugar (*no morph*), Lam = Morphism.term phi Lam, proto_sctr = Morphism.term phi proto_sctr, flat = Morphism.term phi flat, eval_core = Morphism.term phi eval_core, eval = Morphism.term phi eval, algLam = Morphism.term phi algLam, corecUU = Morphism.term phi corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Morphism.thm phi Lam_transfer, Lam_pointful_natural = Morphism.thm phi Lam_pointful_natural, proto_sctr_transfer = Morphism.thm phi proto_sctr_transfer, flat_simps = map (Morphism.thm phi) flat_simps, eval_core_simps = map (Morphism.thm phi) eval_core_simps, eval_thm = Morphism.thm phi eval_thm, eval_simps = map (Morphism.thm phi) eval_simps, all_algLam_algs = map (Morphism.thm phi) all_algLam_algs, algLam_thm = Morphism.thm phi algLam_thm, dtor_algLam = Morphism.thm phi dtor_algLam, corecUU_thm = Morphism.thm phi corecUU_thm, corecUU_unique = Morphism.thm phi corecUU_unique, corecUU_transfer = Morphism.thm phi corecUU_transfer, buffer = morph_buffer phi buffer, all_dead_k_bnfs = map (morph_bnf phi) all_dead_k_bnfs, Retr = Morphism.term phi Retr, equivp_Retr = Morphism.thm phi equivp_Retr, Retr_coinduct = Morphism.thm phi Retr_coinduct, dtor_coinduct_info = morph_dtor_coinduct_info phi dtor_coinduct_info}; datatype ('a, 'b) expr = Ad of 'a * (local_theory -> 'b * local_theory) | Info of 'b; fun is_Ad (Ad _) = true | is_Ad _ = false; fun is_Info (Info _) = true | is_Info _ = false; type corec_info_expr = (corec_ad, corec_info) expr; fun morph_corec_info_expr phi (Ad (ad, f)) = Ad (morph_corec_ad phi ad, f) | morph_corec_info_expr phi (Info info) = Info (morph_corec_info phi info); val transfer_corec_info_expr = morph_corec_info_expr o Morphism.transfer_morphism; type corec_data = int Symtab.table * corec_info_expr list Symtab.table list; structure Data = Generic_Data ( type T = corec_data; val empty = (Symtab.empty, [Symtab.empty]); fun merge ((version_tab1, info_tabs1), (version_tab2, info_tabs2)) : T = (Symtab.join (K Int.max) (version_tab1, version_tab2), info_tabs1 @ info_tabs2); ); fun corec_ad_of_expr (Ad (ad, _)) = ad | corec_ad_of_expr (Info {fpT, friend_names, ...}) = {fpT = fpT, friend_names = friend_names}; fun corec_info_exprs_of_generic context fpT_name = let val thy = Context.theory_of context; val info_tabs = snd (Data.get context); in maps (fn info_tab => these (Symtab.lookup info_tab fpT_name)) info_tabs |> map (transfer_corec_info_expr thy) end; val corec_info_exprs_of = corec_info_exprs_of_generic o Context.Proof; val keep_corec_infos = map_filter (fn Ad _ => NONE | Info info => SOME info); val corec_infos_of_generic = keep_corec_infos oo corec_info_exprs_of_generic; val corec_infos_of = keep_corec_infos oo corec_info_exprs_of; fun str_of_corec_ad ctxt {fpT, friend_names} = "[" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "]"; fun str_of_corec_info ctxt {fpT, version, friend_names, ...} = "{" ^ Syntax.string_of_typ ctxt fpT ^ "; " ^ commas friend_names ^ "; v" ^ string_of_int version ^ "}"; fun str_of_corec_info_expr ctxt (Ad (ad, _)) = str_of_corec_ad ctxt ad | str_of_corec_info_expr ctxt (Info info) = str_of_corec_info ctxt info; fun print_corec_infos ctxt = Symtab.fold (fn (fpT_name, exprs) => fn () => writeln (fpT_name ^ ":\n" ^ cat_lines (map (prefix " " o str_of_corec_info_expr ctxt) exprs))) (the_single (snd (Data.get (Context.Proof ctxt)))) (); val has_no_corec_info = null oo corec_info_exprs_of; fun get_name_next_version_of fpT_name ctxt = let val (version_tab, info_tabs) = Data.get (Context.Theory (Proof_Context.theory_of ctxt)); val fp_base = Long_Name.base_name fpT_name; val fp_b = Binding.name fp_base; val version_tab' = Symtab.map_default (fp_base, ~1) (Integer.add 1) version_tab; val SOME version = Symtab.lookup version_tab' fp_base; val ctxt' = ctxt |> Local_Theory.background_theory (Context.theory_map (Data.put (version_tab', info_tabs))); in ((fp_b, version), ctxt') end; type friend_info = {algrho: term, dtor_algrho: thm, algLam_algrho: thm}; fun morph_friend_info phi ({algrho, dtor_algrho, algLam_algrho} : friend_info) = {algrho = Morphism.term phi algrho, dtor_algrho = Morphism.thm phi dtor_algrho, algLam_algrho = Morphism.thm phi algLam_algrho}; fun checked_fp_sugar_of ctxt fpT_name = let val fp_sugar as {X, fp_res = {Ts = fpTs, ...}, fp_ctr_sugar = {ctrXs_Tss, ...}, ...} = (case fp_sugar_of ctxt fpT_name of SOME (fp_sugar as {fp = Greatest_FP, ...}) => fp_sugar | _ => not_codatatype ctxt (Type (fpT_name, [] (*yuck*)))); val _ = if length fpTs > 1 then mutual_codatatype () else if not (exists (exists (Term.exists_subtype (curry (op =) X))) ctrXs_Tss) then noncorecursive_codatatype () else if ctrXs_Tss = [[X]] then singleton_codatatype ctxt else (); in fp_sugar end; fun bnf_kill_all_but nn bnf lthy = ((empty_comp_cache, empty_unfolds), lthy) |> kill_bnf I (live_of_bnf bnf - nn) bnf ||> snd; fun bnf_with_deads_and_lives dead_Es live_As Y fpT T lthy = let val qsoty = quote o Syntax.string_of_typ lthy; val unfreeze_fp = Tsubst Y fpT; fun flatten_tyargs Ass = map dest_TFree live_As |> filter (fn T => exists (fn Ts => member (op =) Ts T) Ass); val ((bnf, _), (_, lthy)) = bnf_of_typ false Do_Inline I flatten_tyargs [Term.dest_TFree Y] (map Term.dest_TFree dead_Es) T ((empty_comp_cache, empty_unfolds), lthy) handle BAD_DEAD (Y, Y_backdrop) => (case Y_backdrop of Type (bad_tc, _) => let val T = qsoty (unfreeze_fp Y); val T_backdrop = qsoty (unfreeze_fp Y_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible occurrence of type " ^ T ^ " in type expression " ^ T_backdrop) else error ("Unsupported occurrence of type " ^ T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ T_backdrop ^ register_hint ()) end); val phi = Morphism.term_morphism "BNF" (Raw_Simplifier.rewrite_term (Proof_Context.theory_of lthy) @{thms BNF_Composition.id_bnf_def} []) $> Morphism.thm_morphism "BNF" (unfold_thms lthy @{thms BNF_Composition.id_bnf_def}); in (morph_bnf phi bnf, lthy) end; fun define_sig_type fp_b version fp_alives Es Y rhsT lthy = let val T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val ctr_b = mk_version_fp_binding false lthy version fp_b SigN; val sel_b = mk_version_fp_binding true lthy version fp_b unsigN; val lthy = (snd o Local_Theory.begin_nested) lthy; val T_name = Local_Theory.full_name lthy T_b; val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, ctr_b), [(sel_b, rhsT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = true; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.end_nested; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun define_ssig_type fp_b version fp_alives Es Y fpT lthy = let val sig_T_b = mk_version_fp_binding_ugly true lthy version fp_b sigN; val T_b = Binding.prefix_name s_prefix sig_T_b; val Oper_b = mk_version_fp_binding false lthy version fp_b OperN; val VLeaf_b = mk_version_fp_binding false lthy version fp_b VLeafN; val CLeaf_b = mk_version_fp_binding false lthy version fp_b CLeafN; val lthy = (snd o Local_Theory.begin_nested) lthy; val sig_T_name = Local_Theory.full_name lthy sig_T_b; val T_name = Long_Name.map_base_name (prefix s_prefix) sig_T_name; val As = Es @ [Y]; val ssig_sig_T = Type (sig_T_name, Es @ [Type (T_name, As)]); val tyargs = map2 (fn alive => fn T => (if alive then SOME Binding.empty else NONE, (T, Type.sort_of_atyp T))) (fp_alives @ [true]) (Es @ [Y]); val ctr_specs = [(((Binding.empty, Oper_b), [(Binding.empty, ssig_sig_T)]), NoSyn), (((Binding.empty, VLeaf_b), [(Binding.empty, Y)]), NoSyn), (((Binding.empty, CLeaf_b), [(Binding.empty, fpT)]), NoSyn)]; val spec = (((((tyargs, T_b), NoSyn), ctr_specs), (Binding.empty, Binding.empty, Binding.empty)), []); val plugins = Plugin_Name.make_filter lthy (K (curry (op =) transfer_plugin)); val discs_sels = false; val lthy = lthy |> Local_Theory.map_background_naming (mk_internal true lthy Name_Space.concealed) (*TODO check*) |> with_typedef_threshold ~1 (co_datatypes Least_FP construct_lfp ((plugins, discs_sels), [spec])) |> Local_Theory.end_nested; val SOME fp_sugar = fp_sugar_of lthy T_name; in (fp_sugar, lthy) end; fun embed_Sig ctxt Sig inl_or_r t = Library.foldl1 HOLogic.mk_comp [Sig, inl_or_r, dummify_atomic_types t] |> Syntax.check_term ctxt; fun mk_ctr_wrapper_friends ctxt friend_name friend_T old_sig_T k_T Sig old_buffer = let val embed_Sig_inl = embed_Sig ctxt Sig (Inl_const old_sig_T k_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old_buffer); val friends = Symtab.map (K (apsnd embed_Sig_inl)) (#friends old_buffer) |> Symtab.update_new (friend_name, (friend_T, HOLogic.mk_comp (Sig, Inr_const old_sig_T k_T))); in (ctr_wrapper, friends) end; fun pre_type_of_ctor Y ctor = let val (fp_preT, fpT) = dest_funT (fastype_of ctor); in typ_subst_nonatomic [(fpT, Y)] fp_preT end; fun mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf = let val inr' = Inr_const old_sig_T k_T; val dead_sig_map' = substT Z ssig_T dead_sig_map; in Library.foldl1 HOLogic.mk_comp [Oper, dead_sig_map' $ VLeaf, Sig, inr'] end; fun define_embL name fp_b version Y Z fpT old_sig_T old_ssig_T other_summand_T ssig_T Inl_or_r_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf lthy = let val embL_b = mk_version_fp_binding true lthy version fp_b name; val old_ssig_old_sig_T = Tsubst Y old_ssig_T old_sig_T; val ssig_old_sig_T = Tsubst Y ssig_T old_sig_T; val ssig_other_summand_T = Tsubst Y ssig_T other_summand_T; val sigx = Var (("s", 0), old_ssig_old_sig_T); val x = Var (("x", 0), Y); val j = Var (("j", 0), fpT); val embL = Free (Binding.name_of embL_b, old_ssig_T --> ssig_T); val dead_old_sig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_sig_map; val Sig' = substT Y ssig_T Sig; val inl' = Inl_or_r_const ssig_old_sig_T ssig_other_summand_T; val Oper_eq = mk_Trueprop_eq (embL $ (old_Oper $ sigx), Oper $ (Sig' $ (inl' $ (dead_old_sig_map' $ embL $ sigx)))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (embL $ (old_VLeaf $ x), VLeaf $ x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (embL $ (old_CLeaf $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec embL_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val snd' = snd_const YpreT; val dead_pre_map' = substT Z ssig_T dead_pre_map; val Sig' = substT Y ssig_T Sig; val unsig' = substT Y ssig_T unsig; val dead_sig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, ssig_T)] dead_sig_map; val rhs = HOLogic.mk_comp (unsig', dead_sig_map' $ Library.foldl1 HOLogic.mk_comp [Oper, Sig', dead_pre_map' $ VLeaf, snd']); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy = let val YpreT = HOLogic.mk_prodT (Y, preT); val unsig' = substT Y YpreT unsig; val rhs = HOLogic.mk_comp (mk_case_sum (left_case, right_case), unsig'); in define_const true fp_b version LamN rhs lthy end; fun define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case rho lthy end; fun define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam lthy = let val dead_pre_map' = Term.subst_atomic_types [(Y, old1_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, old2_ssig_T), (Z, ssig_T)] dead_pre_map; val left_case = HOLogic.mk_comp (dead_pre_map' $ embLL, old1_Lam); val right_case = HOLogic.mk_comp (dead_pre_map'' $ embLR, old2_Lam); in define_Lam_step_or_merge fp_b version Y preT unsig left_case right_case lthy end; fun define_proto_sctr_step_or_merge fp_b version old_sig_T right_T Sig old_proto_sctr = let val rhs = Library.foldl1 HOLogic.mk_comp [Sig, Inl_const old_sig_T right_T, old_proto_sctr]; in define_const true fp_b version proto_sctrN rhs end; fun define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map lthy = let val flat_b = mk_version_fp_binding true lthy version fp_b flatN; val ssig_sig_T = Tsubst Y ssig_T sig_T; val ssig_ssig_sig_T = Tsubst Y ssig_T ssig_sig_T; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), ssig_ssig_sig_T); val x = Var (("x", 0), ssig_T); val j = Var (("j", 0), fpT); val flat = Free (Binding.name_of flat_b, ssig_ssig_T --> ssig_T); val Oper' = substT Y ssig_T Oper; val VLeaf' = substT Y ssig_T VLeaf; val CLeaf' = substT Y ssig_T CLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_sig_map; val Oper_eq = mk_Trueprop_eq (flat $ (Oper' $ sigx), Oper $ (dead_sig_map' $ flat $ sigx)) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (flat $ (VLeaf' $ x), x) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (flat $ (CLeaf' $ j), CLeaf $ j) |> Logic.all j; in define_single_primrec flat_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam lthy = let val eval_core_b = mk_version_fp_binding true lthy version fp_b eval_coreN; val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_sig_T = Tsubst Y Ypre_ssig_T sig_T; val ssig_preT = Tsubst Y ssig_T preT; val ssig_YpreT = Tsubst Y ssig_T YpreT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val sigx = Var (("s", 0), Ypre_ssig_sig_T); val x = Var (("x", 0), YpreT); val j = Var (("j", 0), fpT); val eval_core = Free (Binding.name_of eval_core_b, Ypre_ssig_T --> ssig_preT); val Oper' = substT Y YpreT Oper; val VLeaf' = substT Y YpreT VLeaf; val CLeaf' = substT Y YpreT CLeaf; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_pre_map'' = substT Z ssig_T dead_pre_map; val dead_pre_map''' = Term.subst_atomic_types [(Y, fpT), (Z, ssig_T)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_sig_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val Lam' = substT Y ssig_T Lam; val fst' = fst_const YpreT; val snd' = snd_const YpreT; val Oper_eq = mk_Trueprop_eq (eval_core $ (Oper' $ sigx), dead_pre_map' $ flat $ (Lam' $ (dead_sig_map' $ (Abs (Name.uu, Ypre_ssig_T, HOLogic.mk_prod (dead_ssig_map' $ fst' $ Bound 0, eval_core $ Bound 0))) $ sigx))) |> Logic.all sigx; val VLeaf_eq = mk_Trueprop_eq (eval_core $ (VLeaf' $ x), dead_pre_map'' $ VLeaf $ (snd' $ x)) |> Logic.all x; val CLeaf_eq = mk_Trueprop_eq (eval_core $ (CLeaf' $ j), dead_pre_map''' $ CLeaf $ (dtor $ j)) |> Logic.all j; in define_single_primrec eval_core_b [Oper_eq, VLeaf_eq, CLeaf_eq] lthy end; fun define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core lthy = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_ssig_T = Tsubst Y fpT ssig_T; val dtor_unfold' = substT Z fp_ssig_T dtor_unfold; val dead_ssig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_ssig_map; val eval_core' = substT Y fpT eval_core; val id' = HOLogic.id_const fpT; val rhs = dtor_unfold' $ HOLogic.mk_comp (eval_core', dead_ssig_map' $ mk_convol (id', dtor)); in define_const true fp_b version evalN rhs lthy end; fun define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val h = Var (("h", 0), Y --> ssig_preT); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = substT Z ssig_ssig_preT dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val rhs = Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (VLeaf, h)] |> Term.lambda h; in define_const true fp_b version cutSsigN rhs lthy end; fun define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval lthy = let val fp_ssig_T = Tsubst Y fpT ssig_T; val Oper' = substT Y fpT Oper; val VLeaf' = substT Y fpT VLeaf; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_ssig_T)] dead_sig_map; val rhs = Library.foldl1 HOLogic.mk_comp [eval, Oper', dead_sig_map' $ VLeaf']; in define_const true fp_b version algLamN rhs lthy end; fun define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig lthy = let val ssig_preT = Tsubst Y ssig_T preT; val h = Var (("h", 0), Y --> ssig_preT); val dtor_unfold' = substT Z ssig_T dtor_unfold; val rhs = HOLogic.mk_comp (dtor_unfold' $ (cutSsig $ h), VLeaf) |> Term.lambda h; in define_const true fp_b version corecUN rhs lthy end; fun define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU lthy = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_ssig_T = Tsubst Y ssig_T ssig_T val ssig_ssig_preT = HOLogic.mk_prodT (ssig_T, ssig_preT); val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val h = Var (("h", 0), Y --> ssig_pre_ssig_T); val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val eval_core' = substT Y ssig_T eval_core; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_ssig_preT)] dead_ssig_map; val id' = HOLogic.id_const ssig_preT; val rhs = corecU $ Library.foldl1 HOLogic.mk_comp [dead_pre_map' $ flat, eval_core', dead_ssig_map' $ mk_convol (sctr, id'), h] |> Term.lambda h; in define_const true fp_b version corecUUN rhs lthy end; fun derive_sig_transfer maybe_swap ctxt live_AsBs pre_rel sig_rel Rs R const pre_rel_def preT_rel_eqs transfer_thm = let val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRsig_rel = list_comb (sig_rel, Rs) $ R; val constB = Term.subst_atomic_types live_AsBs const; val goal = uncurry mk_rel_fun (maybe_swap (RRpre_rel, RRsig_rel)) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_sig_transfer_tac ctxt pre_rel_def preT_rel_eqs transfer_thm)) |> Thm.close_derivation \<^here> end; fun derive_transfer_by_transfer_prover ctxt live_AsBs Rs R const const_defs rel_eqs transfers = let val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_parametricity_goal ctxt (Rs @ [R]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt (const_defs @ map (fn thm => thm RS sym) rel_eqs) rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_dtor_transfer ctxt live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm = let val Type (\<^type_name>\fun\, [fpT, Type (\<^type_name>\fun\, [fpTB, \<^typ>\bool\])]) = snd (strip_typeN (length live_EsFs) (fastype_of fp_rel)); val pre_rel' = Term.subst_atomic_types [(Y, fpT), (Z, fpTB)] pre_rel; val Rpre_rel = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val dtorB = Term.subst_atomic_types live_EsFs dtor; val goal = HOLogic.mk_Trueprop (mk_rel_fun Rfp_rel (Rpre_rel $ Rfp_rel) $ dtor $ dtorB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_transfer_tac ctxt dtor_rel_thm)) |> Thm.close_derivation \<^here> end; fun derive_Lam_or_eval_core_transfer ctxt live_AsBs Y Z preT ssig_T Rs R pre_rel sig_or_ssig_rel ssig_rel const const_def rel_eqs transfers = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreTB = typ_subst_atomic live_AsBs YpreT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val sig_or_ssig_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreTB)] sig_or_ssig_rel; val Rsig_or_ssig_rel' = list_comb (sig_or_ssig_rel', Rs); val RRpre_rel = list_comb (pre_rel, Rs) $ R; val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rpre_rel' = list_comb (pre_rel', Rs); val constB = subst_atomic_types live_AsBs const; val goal = mk_rel_fun (Rsig_or_ssig_rel' $ mk_rel_prod R RRpre_rel) (Rpre_rel' $ RRssig_rel) $ const $ constB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [const_def] rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_proto_sctr_transfer_step_or_merge ctxt Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs transfers = let val proto_sctrZ = substT Y Z proto_sctr; val goal = mk_rel_fun (dead_pre_rel $ R) (dead_sig_rel $ R) $ proto_sctr $ proto_sctrZ |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [proto_sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_sctr_transfer ctxt live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs transfers = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val sctrB = subst_atomic_types live_AsBs sctr; val goal = HOLogic.mk_Trueprop (mk_rel_fun (Rpre_rel' $ RRssig_rel) RRssig_rel $ sctr $ sctrB); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [sctr_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_transfer ctxt live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs transfers = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val ssig_preTB = typ_subst_atomic live_AsBs ssig_preT; val pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_rel; val ssig_rel' = Term.subst_atomic_types [(Y, ssig_preT), (Z, ssig_preTB)] ssig_rel; val Rpre_rel' = list_comb (pre_rel', Rs); val Rfp_rel = list_comb (fp_rel, Rs); val RRssig_rel = list_comb (ssig_rel, Rs) $ R; val Rssig_rel' = list_comb (ssig_rel', Rs); val corecUUB = subst_atomic_types live_AsBs corecUU; val goal = mk_rel_fun (mk_rel_fun R (Rssig_rel' $ (Rpre_rel' $ RRssig_rel))) (mk_rel_fun R Rfp_rel) $ corecUU $ corecUUB |> HOLogic.mk_Trueprop; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_transfer_by_transfer_prover_tac ctxt [cutSsig_def, corecU_def, corecUU_def] fp_k_T_rel_eqs transfers)) |> Thm.close_derivation \<^here> end; fun mk_natural_goal ctxt simple_T_mapfs fs t u = let fun build_simple (T, _) = (case AList.lookup (op =) simple_T_mapfs T of SOME mapf => mapf | NONE => the (find_first (fn f => domain_type (fastype_of f) = T) fs)); val simple_Ts = map fst simple_T_mapfs; val t_map = build_map ctxt simple_Ts [] build_simple (apply2 (range_type o fastype_of) (t, u)); val u_map = build_map ctxt simple_Ts [] build_simple (apply2 (domain_type o fastype_of) (t, u)); in mk_Trueprop_eq (HOLogic.mk_comp (u, u_map), HOLogic.mk_comp (t_map, t)) end; fun derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f const map_thms = let val ffpre_map = list_comb (pre_map, fs) $ f; val constB = subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt [(preT, ffpre_map)] (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_by_unfolding_tac ctxt map_thms)) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer ctxt live_AsBs simple_T_mapfs fs f const transfer bnfs subst_bnfs = let val m = length live_AsBs; val constB = Term.subst_atomic_types live_AsBs const; val goal = mk_natural_goal ctxt simple_T_mapfs (fs @ [f]) const constB; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_natural_from_transfer_tac ctxt m (replicate m true) transfer [] (map rel_Grp_of_bnf bnfs) (map rel_Grp_of_bnf subst_bnfs))) |> Thm.close_derivation \<^here> end; fun derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f = let val ssig_TB = typ_subst_atomic live_AsBs ssig_T; val preT' = Term.typ_subst_atomic [(Y, ssig_T), (Z, ssig_TB)] preT; val ffpre_map = list_comb (pre_map, fs) $ f; val pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssig_TB)] pre_map; val fpre_map' = list_comb (pre_map', fs); val ffssig_map = list_comb (ssig_map, fs) $ f; val preT_mapfs = [(preT, ffpre_map), (preT', fpre_map' $ ffssig_map)]; in derive_natural_from_transfer ctxt live_AsBs preT_mapfs fs f end; fun derive_Lam_Inl_Inr ctxt Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_sig_T = Tsubst Y YpreT old_sig_T; val Ypre_k_T = Tsubst Y YpreT k_T; val inl' = Inl_const Ypre_old_sig_T Ypre_k_T; val inr' = Inr_const Ypre_old_sig_T Ypre_k_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val Sig' = substT Y YpreT Sig; val Lam_o_Sig = HOLogic.mk_comp (Lam, Sig'); val inl_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inl'), HOLogic.mk_comp (dead_pre_map' $ embL, old_Lam)); val inr_goal = mk_Trueprop_eq (HOLogic.mk_comp (Lam_o_Sig, inr'), rho); val goals = [inl_goal, inr_goal]; val goal = Logic.mk_conjunction_balanced goals; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_Lam_Inl_Inr_tac ctxt unsig_thm Lam_def)) |> Conjunction.elim_balanced (length goals) |> map (Thm.close_derivation \<^here>) end; fun derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps = let val x' = substT Y ssig_T x; val dead_ssig_map' = substT Z ssig_T dead_ssig_map; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ VLeaf $ x'), x'); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (fp_map_id :: sig_map_ident :: sig_map_comp :: ssig_map_thms @ flat_simps @ @{thms o_apply id_apply id_def[symmetric]}))) |> Thm.close_derivation \<^here> end; fun derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps = let val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val ssig_ssig_ssig_T = Tsubst Y ssig_T ssig_ssig_T; val x' = substT Y ssig_ssig_ssig_T x; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_ssig_map; val flat' = substT Y ssig_T flat; val goal = mk_Trueprop_eq (flat $ (dead_ssig_map' $ flat $ x'), flat $ (flat' $ x')); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_VLeaf_or_flat_tac ctxt ssig_induct' sig_map_cong (o_apply :: fp_map_id :: sig_map_comp :: ssig_map_thms @ flat_simps))) |> Thm.close_derivation \<^here> end; fun derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val ssig_ssig_T = Tsubst Y ssig_T ssig_T; val Ypre_ssig_T = Tsubst Y YpreT ssig_T; val Ypre_ssig_ssig_T = Tsubst Y YpreT ssig_ssig_T; val ssig_YpreT = Tsubst Y ssig_T YpreT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_ssig_T), (Z, ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, Ypre_ssig_T), (Z, ssig_YpreT)] dead_ssig_map; val dead_ssig_map'' = Term.subst_atomic_types [(Y, YpreT), (Z, Y)] dead_ssig_map; val flat' = substT Y YpreT flat; val eval_core' = substT Y ssig_T eval_core; val x' = substT Y Ypre_ssig_ssig_T x; val fst' = fst_const YpreT; val goal = mk_Trueprop_eq (eval_core $ (flat' $ x'), dead_pre_map' $ flat $ (eval_core' $ (dead_ssig_map' $ mk_convol (dead_ssig_map'' $ fst', eval_core) $ x'))); val ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_flat_tac ctxt ssig_induct' dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def = (trans OF [iffD2 OF [dtor_inject, HOLogic.mk_obj_eq eval_def RS fun_cong], dtor_unfold_thm]) |> unfold_thms ctxt [o_apply, eval_def RS symmetric_thm]; fun derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_ssig_T = Tsubst Y fp_ssig_T ssig_T; val dead_ssig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_ssig_map; val flat' = substT Y fpT flat; val x' = substT Y fp_ssig_ssig_T x; val goal = mk_Trueprop_eq (eval $ (flat' $ x'), eval $ (dead_ssig_map' $ eval $ x')); val cond_eval_o_flat = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (HOLogic.mk_comp (eval, flat')))] (trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] RS fun_cong) OF [ext, ext]; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_flat_tac ctxt dead_pre_map_comp0 ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm cond_eval_o_flat)) |> Thm.close_derivation \<^here> end; fun derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val fp_ssig_sig_T = Tsubst Y fp_ssig_T sig_T; val dead_sig_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_sig_map; val Oper' = substT Y fpT Oper; val x' = substT Y fp_ssig_sig_T x; val goal = mk_Trueprop_eq (eval $ (Oper' $ x'), algLam $ (dead_sig_map' $ eval $ x')); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_Oper_tac ctxt live sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_eval_V_or_CLeaf ctxt Y fpT V_or_CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm = let val V_or_CLeaf' = substT Y fpT V_or_CLeaf; val x' = substT Y fpT x; val goal = mk_Trueprop_eq (eval $ (V_or_CLeaf' $ x'), x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_V_or_CLeaf_tac ctxt dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique V_or_CLeaf_map_thm eval_core_simps eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def = let val ssig_preT = Tsubst Y ssig_T preT; val dead_pre_map' = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)] dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, cutSsig $ g'), HOLogic.mk_comp (dtor, extdd_f)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_mor_tac ctxt dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def prem)) |> Thm.close_derivation \<^here> end; fun derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val dead_ssig_map' = substYZ dead_ssig_map; val f' = substYZ f; val g' = substT Z ssig_preT g; val cutSsig_g = cutSsig $ g'; val id' = HOLogic.id_const ssig_T; val convol' = mk_convol (id', cutSsig_g); val dead_ssig_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, range_type (fastype_of convol'))] dead_ssig_map; val eval_core' = substT Y ssig_T eval_core; val eval_core_o_map = HOLogic.mk_comp (eval_core', dead_ssig_map'' $ convol'); val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ f', cutSsig_g), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, dead_ssig_map' $ f'), HOLogic.mk_comp (f', flat)); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_mor_cutSsig_flat_tac ctxt eval_core_o_map dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm prem)) |> Thm.close_derivation \<^here> end; fun derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val f' = substT Z fpT f; val g' = substT Z ssig_preT g; val extdd_f = extdd $ f'; val prem = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ extdd_f, g'), HOLogic.mk_comp (dtor, f')); val goal = mk_Trueprop_eq (HOLogic.mk_comp (extdd_f, VLeaf), f'); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_extdd_o_VLeaf_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf prem)) |> Thm.close_derivation \<^here> end; fun derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def = let val ssig_preT = Tsubst Y ssig_T preT; val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val dead_pre_map' = substYZ dead_pre_map; val g' = substT Z ssig_preT g; val corecU_g = corecU $ g'; val goal = mk_Trueprop_eq (HOLogic.mk_comp (dead_pre_map' $ (extdd $ corecU_g), g'), HOLogic.mk_comp (dtor, corecU_g)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecU_pointfree_tac ctxt dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def)) |> Thm.close_derivation \<^here> end; fun derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def = let val corecU_pointfree = derive_corecU_pointfree ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd corecU g dead_pre_map_comp dtor_unfold_thm ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps cutSsig_def mor_cutSsig_flat corecU_def; val corecU_thm = corecU_pointfree RS @{thm comp_eq_dest}; val corecU_ctor = let val arg_cong' = infer_instantiate' ctxt [NONE, NONE, SOME (Thm.cterm_of ctxt ctor)] arg_cong; in unfold_thms ctxt [ctor_dtor] (corecU_thm RS arg_cong') end; val corecU_unique = let val substYZ = Term.subst_atomic_types [(Y, ssig_T), (Z, fpT)]; val f' = substYZ f; val abs_f_o_VLeaf = Term.lambda f' (HOLogic.mk_comp (f', VLeaf)); val inject_refine' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt abs_f_o_VLeaf), SOME (Thm.cterm_of ctxt extdd)] @{thm inject_refine}; in unfold_thms ctxt @{thms atomize_imp} (((inject_refine' OF [extdd_o_VLeaf, extdd_o_VLeaf] RS iffD1) OF [asm_rl, corecU_pointfree]) OF [asm_rl, trans OF [dtor_unfold_unique, dtor_unfold_unique RS sym] OF [extdd_mor, corecU_pointfree RS extdd_mor]]) RS @{thm obj_distinct_prems} end; in (corecU_ctor, corecU_unique) end; fun derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def = let val fp_preT = Tsubst Y fpT preT; val fppreT = HOLogic.mk_prodT (fpT, fp_preT); val fp_sig_T = Tsubst Y fpT sig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_sig_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_sig_map; val Lam' = substT Y fpT Lam; val x' = substT Y fp_sig_T x; val goal = mk_Trueprop_eq (dtor $ (algLam $ x'), dead_pre_map' $ eval $ (Lam' $ (dead_sig_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algLam_tac ctxt pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_base ctxt Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def = let val fp_preT = Tsubst Y fpT preT; val proto_sctr' = substT Y fpT proto_sctr; val dead_pre_map' = Term.subst_atomic_types [(Y, fpT), (Z, fp_preT)] dead_pre_map; val dead_pre_map_dtor = dead_pre_map' $ dtor; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_base_tac ctxt dead_pre_map_dtor dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def)) |> Thm.close_derivation \<^here> end; fun derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps = let val old_ssig_old_ssig_T = Tsubst Y old_ssig_T old_ssig_T; val dead_old_ssig_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_old_ssig_map; val embL' = substT Y ssig_T embL; val x' = substT Y old_ssig_old_ssig_T x; val goal = mk_Trueprop_eq (flat $ (embL' $ (dead_old_ssig_map' $ embL $ x')), embL $ (old_flat $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_flat_embL_tac ctxt old_ssig_induct' fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_old_ssig_T = Tsubst Y YpreT old_ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, old_ssig_T), (Z, ssig_T)] dead_pre_map; val embL' = substT Y YpreT embL; val x' = substT Y Ypre_old_ssig_T x; val goal = mk_Trueprop_eq (eval_core $ (embL' $ x'), dead_pre_map' $ embL $ (old_eval_core $ x')); val old_ssig_induct' = infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt x')] old_ssig_induct; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_embL_tac ctxt old_ssig_induct' dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL old_eval_core_simps eval_core_simps embL_simps embL_pointful_natural)) |> Thm.close_derivation \<^here> end; fun derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm = let val embL' = substT Y fpT embL; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, embL'), old_eval); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_embL_tac ctxt dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inx' = Inx_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inx'], old_algLam); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algLam_tac ctxt dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam)) |> Thm.close_derivation \<^here> end; fun derive_eval_core_k_as_ssig ctxt Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps = let val YpreT = HOLogic.mk_prodT (Y, preT); val Ypre_k_T = Tsubst Y YpreT k_T; val k_as_ssig' = substT Y YpreT k_as_ssig; val x' = substT Y Ypre_k_T x; val goal = mk_Trueprop_eq (eval_core $ (k_as_ssig' $ x'), rho $ x'); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_core_k_as_ssig_tac ctxt pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps)) |> Thm.close_derivation \<^here> end; fun derive_algLam_algrho ctxt Y fpT Sig algLam algrho algLam_def algrho_def = let val Sig' = substT Y fpT Sig; val (left_T, right_T) = dest_sumT (domain_type (fastype_of Sig')); val inr' = Inr_const left_T right_T; val goal = mk_Trueprop_eq (Library.foldl1 HOLogic.mk_comp [algLam, Sig', inr'], algrho); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_algrho_tac ctxt algLam_def algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_dtor_algrho ctxt Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def = let val YpreT = HOLogic.mk_prodT (Y, preT); val fppreT = Tsubst Y fpT YpreT; val fp_k_T = Tsubst Y fpT k_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val id' = HOLogic.id_const fpT; val convol' = mk_convol (id', dtor); val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_k_map' = Term.subst_atomic_types [(Y, fpT), (Z, fppreT)] dead_k_map; val rho' = substT Y fpT rho; val x' = substT Y fp_k_T x; val goal = mk_Trueprop_eq (dtor $ (algrho $ x'), dead_pre_map' $ eval $ (rho' $ (dead_k_map' $ convol' $ x'))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_dtor_algrho_tac ctxt eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def)) |> Thm.close_derivation \<^here> end; fun derive_algLam_step_or_merge ctxt Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam = let val proto_sctr' = substT Y fpT proto_sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (algLam, proto_sctr'), ctor); val algLam_algLam_pointful = mk_pointful ctxt algLam_algLam; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_algLam_step_tac ctxt proto_sctr_def old_algLam_pointful algLam_algLam_pointful)) |> Thm.close_derivation \<^here> end; fun derive_eval_sctr ctxt Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def = let val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val sctr' = substT Y fpT sctr; val goal = mk_Trueprop_eq (HOLogic.mk_comp (eval, sctr'), HOLogic.mk_comp (ctor, dead_pre_map' $ eval)); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_eval_sctr_tac ctxt proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def)) |> Thm.close_derivation \<^here> end; fun derive_corecUU_pointfree_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def = let val ssig_preT = Tsubst Y ssig_T preT; val ssig_pre_ssig_T = Tsubst Y ssig_preT ssig_T; val fp_ssig_T = Tsubst Y fpT ssig_T; val dead_pre_map' = Term.subst_atomic_types [(Y, fp_ssig_T), (Z, fpT)] dead_pre_map; val dead_pre_map'' = Term.subst_atomic_types [(Y, ssig_T), (Z, fp_ssig_T)] dead_pre_map; val dead_ssig_map' = Term.subst_atomic_types [(Y, ssig_preT), (Z, fpT)] dead_ssig_map; val dead_ssig_map'' = substT Z fpT dead_ssig_map; val f' = substT Z ssig_pre_ssig_T f; val g' = substT Z fpT g; val corecUU_f = corecUU $ f'; fun mk_eq fpf = mk_Trueprop_eq (fpf, Library.foldl1 HOLogic.mk_comp [eval, dead_ssig_map' $ Library.foldl1 HOLogic.mk_comp [ctor, dead_pre_map' $ eval, dead_pre_map'' $ (dead_ssig_map'' $ fpf)], f']); val corecUU_pointfree = let val goal = mk_eq corecUU_f; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_corecUU_pointfree_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor sctr_pointful_natural eval_sctr_pointful corecUU_def)) |> Thm.close_derivation \<^here> end; val corecUU_unique = let val prem = mk_eq g'; val goal = mk_Trueprop_eq (g', corecUU_f); in fold (Variable.add_free_names ctxt) [prem, goal] [] |> (fn vars => Goal.prove_sorry ctxt vars [prem] goal (fn {context = ctxt, prems = [prem]} => mk_corecUU_unique_tac ctxt dead_pre_map_comp0 dead_pre_map_comp dtor_ctor ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def prem)) |> Thm.close_derivation \<^here> end; in (corecUU_pointfree, corecUU_unique) end; fun define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer lthy = let val (flat_data as (flat, flat_def, _), lthy) = lthy |> define_flat fp_b version Y Z fpT sig_T ssig_T Oper VLeaf CLeaf dead_sig_map; val (eval_core_data as (eval_core, eval_core_def, _), lthy) = lthy |> define_eval_core fp_b version Y Z preT fpT sig_T ssig_T dtor Oper VLeaf CLeaf dead_pre_map dead_sig_map dead_ssig_map flat Lam; val ((eval_data as (eval, _), cutSsig_data as (cutSsig, _)), lthy) = lthy |> define_eval fp_b version Y Z preT fpT ssig_T dtor dtor_unfold dead_ssig_map eval_core ||>> define_cutSsig fp_b version Y Z preT ssig_T dead_pre_map VLeaf dead_ssig_map flat eval_core; val ((algLam_data, unfold_data), lthy) = lthy |> define_algLam fp_b version Y Z fpT ssig_T Oper VLeaf dead_sig_map eval ||>> define_corecU fp_b version Y Z preT ssig_T dtor_unfold VLeaf cutSsig; val flat_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R flat [flat_def] [] [sig_map_transfer]; val eval_core_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel ssig_rel ssig_rel eval_core eval_core_def fp_k_T_rel_eqs [pre_map_transfer, sig_map_transfer, ssig_map_transfer, flat_transfer, Lam_transfer, dtor_transfer]; in (((((((flat_data, flat_transfer), (eval_core_data, eval_core_transfer)), eval_data), cutSsig_data), algLam_data), unfold_data), lthy) end; fun derive_Sig_natural_etc ctxt live live_AsBs Y Z preT fpT k_or_fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf = let val SOME prod_bnf = bnf_of ctxt \<^type_name>\prod\; val f' = substT Z fpT f; val dead_ssig_map' = substT Z fpT dead_ssig_map; val extdd = Term.lambda f' (HOLogic.mk_comp (eval, dead_ssig_map' $ f')); val live_pre_map_def = map_def_of_bnf live_pre_bnf; val pre_map_comp = map_comp_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val dead_pre_map_cong = map_cong_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val sig_map_ident = map_ident_of_bnf sig_bnf; val sig_map_comp0 = map_comp0_of_bnf sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_cong = map_cong_of_bnf sig_bnf; val ssig_map_id = map_id_of_bnf ssig_bnf; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val dead_ssig_map_comp0 = map_comp0_of_bnf dead_ssig_bnf; val k_preT_map_id0s = map map_id0_of_bnf (map_filter (bnf_of ctxt) (fold add_type_namesT [preT, k_or_fpT] [])); val Sig_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Sig ([sig_map_thm, live_pre_map_def, @{thm BNF_Composition.id_bnf_def}] @ k_preT_map_id0s); val Oper_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f Oper [Oper_map_thm]; val VLeaf_natural = derive_natural_by_unfolding ctxt live_AsBs preT pre_map fs f VLeaf [VLeaf_map_thm]; val Lam_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f Lam Lam_transfer [prod_bnf, pre_bnf, sig_bnf, ssig_bnf] []; val flat_natural = derive_natural_from_transfer ctxt live_AsBs [] fs f flat flat_transfer [ssig_bnf] []; val eval_core_natural = derive_natural_from_transfer_with_pre_type ctxt live_AsBs Y Z preT ssig_T pre_map ssig_map fs f eval_core eval_core_transfer [prod_bnf, pre_bnf, ssig_bnf] []; val Sig_pointful_natural = mk_pointful ctxt Sig_natural RS sym; val Oper_natural_pointful = mk_pointful ctxt Oper_natural; val Oper_pointful_natural = Oper_natural_pointful RS sym; val flat_pointful_natural = mk_pointful ctxt flat_natural RS sym; val Lam_natural_pointful = mk_pointful ctxt Lam_natural; val Lam_pointful_natural = Lam_natural_pointful RS sym; val eval_core_pointful_natural = mk_pointful ctxt eval_core_natural RS sym; val cutSsig_def_pointful_natural = mk_pointful ctxt (HOLogic.mk_obj_eq cutSsig_def) RS sym; val flat_VLeaf = derive_flat_VLeaf ctxt Y Z ssig_T x VLeaf dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_ident sig_map_comp ssig_map_thms flat_simps; val flat_flat = derive_flat_flat ctxt Y Z ssig_T x dead_ssig_map flat ssig_induct fp_map_id sig_map_cong sig_map_comp ssig_map_thms flat_simps; val eval_core_flat = derive_eval_core_flat ctxt Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core x ssig_induct dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp fp_map_id sig_map_comp sig_map_cong ssig_map_thms ssig_map_comp flat_simps flat_pointful_natural flat_flat Lam_pointful_natural eval_core_simps; val eval_thm = derive_eval_thm ctxt dtor_inject dtor_unfold_thm eval_def; val eval_flat = derive_eval_flat ctxt Y Z fpT ssig_T dead_ssig_map flat eval x dead_pre_map_comp0 dtor_unfold_unique ssig_map_id ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_core_flat eval_thm; val eval_Oper = derive_eval_Oper ctxt live Y Z fpT sig_T ssig_T dead_sig_map Oper eval algLam x sig_map_ident sig_map_comp0 sig_map_comp Oper_natural_pointful VLeaf_natural flat_simps eval_flat algLam_def; val eval_VLeaf = derive_eval_V_or_CLeaf ctxt Y fpT VLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique VLeaf_map_thm eval_core_simps eval_thm; val eval_CLeaf = derive_eval_V_or_CLeaf ctxt Y fpT CLeaf eval x dead_pre_map_id dead_pre_map_comp fp_map_id dtor_unfold_unique CLeaf_map_thm eval_core_simps eval_thm; val extdd_mor = derive_extdd_mor ctxt Y Z preT fpT ssig_T dead_pre_map dtor extdd cutSsig f g dead_pre_map_comp0 dead_pre_map_comp VLeaf_map_thm ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat eval_VLeaf cutSsig_def; val mor_cutSsig_flat = derive_mor_cutSsig_flat ctxt Y Z preT fpT ssig_T dead_pre_map dead_ssig_map dtor flat eval_core eval cutSsig f g dead_pre_map_comp0 dead_pre_map_comp dead_pre_map_cong dtor_unfold_unique dead_ssig_map_comp0 ssig_map_comp flat_simps flat_pointful_natural eval_core_pointful_natural flat_flat flat_VLeaf eval_core_flat cutSsig_def cutSsig_def_pointful_natural eval_thm; val extdd_o_VLeaf = derive_extdd_o_VLeaf ctxt Y Z preT fpT ssig_T dead_pre_map dtor VLeaf extdd f g dead_pre_map_comp0 dead_pre_map_comp dtor_inject ssig_map_thms eval_core_simps eval_thm eval_VLeaf; val (corecU_ctor, corecU_unique) = derive_corecU_ctor_unique ctxt Y Z preT fpT ssig_T dead_pre_map ctor dtor VLeaf extdd corecU f g dead_pre_map_comp ctor_dtor dtor_unfold_thm dtor_unfold_unique ssig_map_thms dead_ssig_map_comp0 flat_simps flat_VLeaf eval_core_simps extdd_mor extdd_o_VLeaf cutSsig_def mor_cutSsig_flat corecU_def; val dtor_algLam = derive_dtor_algLam ctxt Y Z preT fpT sig_T ssig_T dead_pre_map dtor dead_sig_map Lam eval algLam x pre_map_comp dead_pre_map_id dead_pre_map_comp0 dead_pre_map_comp sig_map_comp Oper_pointful_natural ssig_map_thms dead_ssig_map_comp0 Lam_pointful_natural eval_core_simps eval_thm eval_flat eval_VLeaf algLam_def; in (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, [eval_Oper, eval_VLeaf, eval_CLeaf], corecU_ctor, corecU_unique, dtor_algLam) end; fun derive_embL_natural_etc ctxt Inx_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm = let val embL_natural = derive_natural_from_transfer ctxt [(Y, Z)] [] [] f embL embL_transfer [old_ssig_bnf, ssig_bnf] []; val embL_pointful_natural = mk_pointful ctxt embL_natural RS sym; val old_algLam_pointful = mk_pointful ctxt old_algLam_thm; val flat_embL = derive_flat_embL ctxt Y Z old_ssig_T ssig_T dead_old_ssig_map embL old_flat flat x old_ssig_induct fp_map_id Sig_pointful_natural old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_flat_simps flat_simps embL_simps; val eval_core_embL = derive_eval_core_embL ctxt Y Z preT old_ssig_T ssig_T dead_pre_map embL old_eval_core eval_core x old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp Sig_pointful_natural unsig_thm old_sig_map_comp old_sig_map_cong old_Lam_pointful_natural Lam_def flat_embL embL_simps embL_pointful_natural old_eval_core_simps eval_core_simps; val eval_embL = derive_eval_embL ctxt Y fpT embL old_eval eval dead_pre_map_comp0 dtor_unfold_unique embL_pointful_natural eval_core_embL old_eval_thm eval_thm; val algLam_algLam = derive_algLam_algLam ctxt Inx_const Y fpT Sig old_algLam algLam dead_pre_map_comp dtor_inject unsig_thm sig_map_thm Lam_def eval_embL old_dtor_algLam dtor_algLam; in (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) end; fun define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar lthy = let val ssig_bnf = #fp_bnf ssig_fp_sugar; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val ssig_map_comp = map_comp_of_bnf ssig_bnf; val sctr_rhs = HOLogic.mk_comp (Oper, substT Y ssig_T proto_sctr); val ((sctr, sctr_def), lthy) = lthy |> define_const true fp_b version sctrN sctr_rhs; val (corecUU_data as (corecUU, corecUU_def), lthy) = lthy |> define_corecUU fp_b version Y Z preT ssig_T dead_pre_map dead_ssig_map flat eval_core sctr corecU; val eval_sctr = derive_eval_sctr lthy Y Z fpT ssig_T dead_pre_map ctor eval sctr proto_sctr_pointful_natural eval_Oper algLam_thm sctr_def; val sctr_transfer = derive_sctr_transfer lthy live_AsBs Y Z ssig_T Rs R pre_rel ssig_rel sctr sctr_def fp_k_T_rel_eqs [proto_sctr_transfer]; val sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f sctr sctr_transfer [pre_bnf, ssig_bnf] []; val sctr_pointful_natural = mk_pointful lthy sctr_natural RS sym; val eval_sctr_pointful = mk_pointful lthy eval_sctr RS sym; val (corecUU_pointfree, corecUU_unique) = derive_corecUU_pointfree_unique lthy Y Z preT fpT ssig_T dead_pre_map ctor dead_ssig_map eval corecUU f g dead_pre_map_comp0 dead_pre_map_comp dtor_ctor dtor_inject ssig_map_comp flat_pointful_natural eval_core_pointful_natural eval_thm eval_flat corecU_ctor corecU_unique sctr_pointful_natural eval_sctr_pointful corecUU_def; val corecUU_thm = mk_pointful lthy corecUU_pointfree; val corecUU_transfer = derive_corecUU_transfer lthy live_AsBs Y Z Rs R preT ssig_T pre_rel fp_rel ssig_rel corecUU cutSsig_def corecU_def corecUU_def fp_k_T_rel_eqs [pre_map_transfer, dtor_unfold_transfer, dtor_transfer, ssig_map_transfer, flat_transfer, eval_core_transfer, sctr_transfer, @{thm convol_transfer} (*FIXME: needed?*)]; in ((corecUU_data, corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) end; fun mk_equivp T = Const (\<^const_name>\equivp\, mk_predT [mk_pred2T T T]); fun derive_equivp_Retr ctxt fpT Retr R dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm = let val prem = HOLogic.mk_Trueprop (mk_equivp fpT $ R); val goal = Logic.mk_implies (prem, HOLogic.mk_Trueprop (mk_equivp fpT $ (betapply (Retr, R)))); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_equivp_Retr_tac ctxt dead_pre_rel_refl_thm dead_pre_rel_flip_thm dead_pre_rel_mono_thm dead_pre_rel_compp_thm)) |> Thm.close_derivation \<^here> end; fun derive_Retr_coinduct ctxt fpT Retr R dtor_rel_coinduct_thm rel_eq_thm = let val goal = HOLogic.mk_Trueprop (list_all_free [R] (HOLogic.mk_imp (mk_leq R (Retr $ R), mk_leq R (HOLogic.eq_const fpT)))); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_Retr_coinduct_tac ctxt dtor_rel_coinduct_thm rel_eq_thm) |> Thm.close_derivation \<^here> end; fun derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy lthy = let val (R, _) = names_lthy |> yield_singleton (mk_Frees "R") (mk_pred2T fpT fpT); val pre_fpT = pre_type_of_ctor fpT ctor; val fp_pre_rel = mk_rel1 lthy fpT fpT pre_fpT dead_pre_bnf; val Retr = Abs ("R", fastype_of R, Abs ("a", fpT, Abs ("b", fpT, list_comb (fp_pre_rel, [Bound 2, dtor $ Bound 1, dtor $ Bound 0])))); val equivp_Retr = derive_equivp_Retr lthy fpT Retr R (rel_refl_of_bnf dead_pre_bnf) (rel_flip_of_bnf dead_pre_bnf) (rel_mono_of_bnf dead_pre_bnf) (rel_OO_of_bnf dead_pre_bnf); val Retr_coinduct = derive_Retr_coinduct lthy fpT Retr R (fp_sugar |> #fp_res |> #xtor_rel_co_induct) (fp_sugar |> #fp_bnf |> rel_eq_of_bnf); in (Retr, equivp_Retr, Retr_coinduct) end; fun mk_gen_cong fpT eval_domT = let val fp_relT = mk_pred2T fpT fpT in Const (\<^const_name>\cong.gen_cong\, [mk_predT [fp_relT, eval_domT, eval_domT], eval_domT --> fpT, fp_relT] ---> fp_relT) end; fun mk_cong_locale rel eval Retr = Const (\<^const_name>\cong\, mk_predT (map fastype_of [rel, eval, Retr])); fun derive_cong_locale ctxt rel eval Retr0 tac = let val Retr = enforce_type ctxt domain_type (domain_type (fastype_of rel)) Retr0; val goal = HOLogic.mk_Trueprop (list_comb (mk_cong_locale rel eval Retr, [rel, eval, Retr])); in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => tac ctxt)) |> Thm.close_derivation \<^here> end; fun derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy = let val eval_domT = domain_type (fastype_of eval); fun cong_locale_tac ctxt = mk_cong_locale_tac ctxt (rel_mono_of_bnf dead_pre_bnf) (rel_map_of_bnf dead_pre_bnf) equivp_Retr (rel_mono_of_bnf dead_ssig_bnf) (rel_map_of_bnf dead_ssig_bnf) eval_thm eval_core_transfer; val rel = mk_rel1 lthy fpT fpT eval_domT dead_ssig_bnf; val cong_rhs = list_comb (mk_gen_cong fpT eval_domT, [rel, eval]); val ((_, cong_def), lthy) = lthy |> define_const false fp_b version congN cong_rhs; val cong_locale = derive_cong_locale lthy rel eval Retr cong_locale_tac; val fold_cong_def = Local_Defs.fold lthy [cong_def]; fun instance_of_gen thm = fold_cong_def (thm OF [cong_locale]); val cong_base = instance_of_gen @{thm cong.imp_gen_cong}; val cong_refl = instance_of_gen @{thm cong.gen_cong_reflp}; val cong_sym = instance_of_gen @{thm cong.gen_cong_symp}; val cong_trans = instance_of_gen @{thm cong.gen_cong_transp}; fun mk_cong_rho thm = thm RS instance_of_gen @{thm cong.gen_cong_rho}; val dtor_coinduct = @{thm predicate2I_obj} RS (Retr_coinduct RS instance_of_gen @{thm cong.coinduction} RS @{thm predicate2D_obj}); in (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) end; fun derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct lthy = let val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val dead_pre_map_id0 = map_id0_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_cong0 = map_cong0_of_bnf dead_pre_bnf; val dead_pre_map_cong0' = @{thm box_equals[OF _ o_apply[symmetric] id_apply[symmetric]]} RS dead_pre_map_cong0 RS ext; val dead_pre_rel_map = rel_map_of_bnf dead_pre_bnf; val ctor_alt_thm = eval_VLeaf RS (@{thm eq_comp_compI} OF [eval_sctr, trans OF [dead_pre_map_comp0 RS sym, trans OF [dead_pre_map_cong0', dead_pre_map_id0]]]); val cong_ctor_intro = mk_cong_rho ctor_alt_thm |> unfold_thms lthy [o_apply] |> (fn thm => sctr_transfer RS rel_funD RS thm) |> unfold_thms lthy (id_apply :: dead_pre_rel_map @ #rel_injects ssig_fp_bnf_sugar); in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = [cong_ctor_intro]}, lthy) end; fun update_cong_alg_intros ctxt cong_def cong_locale old_cong_def old_cong_locale emb = let fun instance_of_gen thm = Local_Defs.fold ctxt [cong_def] (thm OF [cong_locale]); fun instance_of_old_gen thm = Local_Defs.fold ctxt [old_cong_def] (thm OF [old_cong_locale]); val emb_idem = @{thm ord_le_eq_trans} OF [emb, instance_of_gen @{thm cong.gen_cong_idem}]; fun mk_rel_mono bnf = instance_of_old_gen @{thm cong.leq_gen_cong} RS rel_mono_of_bnf bnf RS @{thm predicate2D}; fun mk_intro bnf thm = mk_rel_mono bnf RS (@{thm predicate2D} OF [emb_idem, thm]); in map2 mk_intro end fun derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs lthy = let val old_cong_def = #cong_def old_dtor_coinduct_info; val old_cong_locale = #cong_locale old_dtor_coinduct_info; val old_cong_alg_intros = #cong_alg_intros old_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, mk_cong_rho, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val cong_alg_intro = k_as_ssig_transfer RS rel_funD RS mk_cong_rho (HOLogic.mk_obj_eq algrho_def); val gen_cong_emb = (@{thm gen_cong_emb} OF [old_cong_locale, cong_locale, eval_embL, embL_transfer]) |> Local_Defs.fold lthy [old_cong_def, cong_def]; val cong_alg_intros = update_cong_alg_intros lthy cong_def cong_locale old_cong_def old_cong_locale gen_cong_emb old_all_dead_k_bnfs old_cong_alg_intros; in ({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_alg_intro :: cong_alg_intros}, lthy) end; fun derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs lthy = let val old1_cong_def = #cong_def old1_dtor_coinduct_info; val old1_cong_locale = #cong_locale old1_dtor_coinduct_info; val old1_cong_alg_intros = #cong_alg_intros old1_dtor_coinduct_info; val old2_cong_def = #cong_def old2_dtor_coinduct_info; val old2_cong_locale = #cong_locale old2_dtor_coinduct_info; val old2_cong_alg_intros = #cong_alg_intros old2_dtor_coinduct_info; val (cong_def, cong_locale, cong_base, cong_refl, cong_sym, cong_trans, dtor_coinduct, _, lthy) = derive_cong_general fp_b version fpT dead_ssig_bnf dead_pre_bnf eval Retr equivp_Retr Retr_coinduct eval_thm eval_core_transfer lthy; val emb1 = (@{thm gen_cong_emb} OF [old1_cong_locale, cong_locale, eval_embLL, embLL_transfer]) |> Local_Defs.fold lthy [old1_cong_def, cong_def]; val emb2 = (@{thm gen_cong_emb} OF [old2_cong_locale, cong_locale, eval_embLR, embLR_transfer]) |> Local_Defs.fold lthy [old2_cong_def, cong_def]; val cong_alg_intros1 = update_cong_alg_intros lthy cong_def cong_locale old1_cong_def old1_cong_locale emb1 old1_all_dead_k_bnfs old1_cong_alg_intros; val cong_alg_intros2 = update_cong_alg_intros lthy cong_def cong_locale old2_cong_def old2_cong_locale emb2 old2_all_dead_k_bnfs old2_cong_alg_intros; val (cong_algrho_intros1, cong_ctor_intro1) = split_last cong_alg_intros1; val (cong_algrho_intros2, _) = split_last cong_alg_intros2; val (old1_all_rho_k_bnfs, old1_Sig_bnf) = split_last old1_all_dead_k_bnfs; val (old2_all_rho_k_bnfs, _) = split_last old2_all_dead_k_bnfs; val (friend_names, (cong_algrho_intros, all_rho_k_bnfs)) = merge_lists (op = o apply2 fst) (old1_friend_names ~~ (cong_algrho_intros1 ~~ old1_all_rho_k_bnfs)) (old2_friend_names ~~ (cong_algrho_intros2 ~~ old2_all_rho_k_bnfs)) |> split_list ||> split_list; in (({cong_def = cong_def, cong_locale = cong_locale, cong_base = cong_base, cong_refl = cong_refl, cong_sym = cong_sym, cong_trans = cong_trans, dtor_coinduct = dtor_coinduct, cong_alg_intros = cong_algrho_intros @ [cong_ctor_intro1]}, all_rho_k_bnfs @ [old1_Sig_bnf], friend_names), lthy) end; fun derive_corecUU_base fpT_name lthy = let val fp_sugar as {T = Type (_, fpT_args0), pre_bnf, fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val fp_alives = liveness_of_fp_bnf (length fpT_args0) fp_bnf; val (((Es, Fs0), [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; val Fs = @{map 3} (fn alive => fn E => fn F => if alive then F else E) fp_alives Es Fs0; val As = Es @ [Y]; val Bs = Es @ [Z]; val live_EsFs = filter (op <>) (Es ~~ Fs); val live_AsBs = live_EsFs @ [(Y, Z)]; val fTs = map (op -->) live_EsFs; val RTs = map (uncurry mk_pred2T) live_EsFs; val live = length live_EsFs; val ((((((x, fs), f), g), Rs), R), names_lthy) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> mk_Frees "f" fTs ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> mk_Frees "R" RTs ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val ctor = mk_ctor Es (the_single (#ctors fp_res)); val dtor = mk_dtor Es (the_single (#dtors fp_res)); val fpT = Type (fpT_name, Es); val preT = pre_type_of_ctor Y ctor; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Es Y preT ||>> define_ssig_type fp_b version fp_alives Es Y fpT; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (((dead_pre_bnf, dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 pre_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val sig_T = Type (sig_T_name, As); val ssig_T = Type (ssig_T_name, As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_sig_map = mk_map 1 As Bs (map_of_bnf dead_sig_bnf); val [Oper, VLeaf, CLeaf] = map (mk_ctr As) (#ctrs ssig_ctr_sugar); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 As Bs (map_of_bnf dead_ssig_bnf); val ((Lam, Lam_def), lthy) = lthy |> define_Lam_base fp_b version Y Z preT ssig_T dead_pre_map Sig unsig dead_sig_map Oper VLeaf; val proto_sctr = Sig; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val pre_rel_def = rel_def_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_ctor] = #dtor_ctors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val [dtor_rel_thm] = #xtor_rels fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val dtor_transfer = derive_dtor_transfer lthy live_EsFs Y Z pre_rel fp_rel Rs dtor dtor_rel_thm; val preT_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (add_type_namesT preT [])); val Sig_transfer = derive_sig_transfer I lthy live_AsBs pre_rel sig_rel Rs R Sig pre_rel_def preT_rel_eqs (the_single (#ctr_transfers sig_fp_ctr_sugar)); val proto_sctr_transfer = Sig_transfer; val unsig_transfer = derive_sig_transfer swap lthy live_AsBs pre_rel sig_rel Rs R unsig pre_rel_def preT_rel_eqs (the_single (#sel_transfers sig_fp_ctr_sugar)); val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, sig_map_transfer, Sig_transfer, unsig_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, eval_VLeaf, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_pointful_natural = Sig_pointful_natural; val algLam_thm = derive_algLam_base lthy Y Z preT fpT dead_pre_map ctor dtor algLam proto_sctr dead_pre_map_id dead_pre_map_comp ctor_dtor dtor_ctor dtor_unfold_unique unsig_thm Sig_pointful_natural ssig_map_thms Lam_def flat_simps eval_core_simps eval_thm algLam_def; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, eval_sctr, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (Retr, equivp_Retr, Retr_coinduct) = lthy |> derive_Retr fp_sugar fpT dead_pre_bnf ctor dtor names_lthy; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_base fp_b version fpT dead_ssig_bnf ssig_fp_bnf_sugar dead_pre_bnf eval eval_thm eval_core_transfer eval_VLeaf eval_sctr sctr_transfer Retr equivp_Retr Retr_coinduct; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = Sig, friends = Symtab.empty}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (dtor_transferN, [dtor_transfer]), (equivp_RetrN, [equivp_Retr]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (Retr_coinductN, [Retr_coinduct]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = [], sig_fp_sugars = [sig_fp_sugar], ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = [algLam_thm], algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = [dead_pre_bnf], Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_step (fpT as Type (fpT_name, res_Ds)) ({friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugars as old_sig_fp_sugar :: _, ssig_fp_sugar = old_ssig_fp_sugar, Lam = old_Lam0, proto_sctr = old_proto_sctr0, flat = old_flat0, eval_core = old_eval_core0, eval = old_eval0, algLam = old_algLam0, dtor_transfer, Lam_transfer = old_Lam_transfer, Lam_pointful_natural = old_Lam_pointful_natural, proto_sctr_transfer = old_proto_sctr_transfer, flat_simps = old_flat_simps, eval_core_simps = old_eval_core_simps, eval_thm = old_eval_thm, all_algLam_algs = old_all_algLam_algs, algLam_thm = old_algLam_thm, dtor_algLam = old_dtor_algLam, buffer = old_buffer, all_dead_k_bnfs = old_all_dead_k_bnfs, Retr = old_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old_dtor_coinduct_info, ...} : corec_info) friend_name friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy = let val {pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val names_lthy = lthy |> fold Variable.declare_typ [Y, Z]; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val friend_names = friend_name :: old_friend_names; val old_sig_bnf = #fp_bnf old_sig_fp_sugar; val old_ssig_bnf = #fp_bnf old_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((dead_pre_bnf, dead_fp_bnf), dead_old_sig_bnf), dead_old_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old_sig_bnf ||>> bnf_kill_all_but 1 old_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old_ssig_fp_ctr_sugar = #fp_ctr_sugar old_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old_ssig_fp_bnf_sugar = #fp_bnf_sugar old_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old_ssig_ctr_sugar = #ctr_sugar old_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_ssig_T_name = fst (dest_Type (#T old_ssig_fp_sugar)); val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val old_sig_T = Type (old_sig_T_name, res_As); val old_ssig_T = Type (old_ssig_T_name, res_As); val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val old_Lam_domT = Tsubst Y YpreT old_sig_T; val old_eval_core_domT = Tsubst Y YpreT old_ssig_T; val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val dead_k_map = mk_map1 lthy Y Z k_T dead_k_bnf; val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old_Oper, old_VLeaf, old_CLeaf] = map (mk_ctr res_As) (#ctrs old_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val dead_old_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old_Lam = enforce_type lthy domain_type old_Lam_domT old_Lam0; val old_proto_sctr = enforce_type lthy domain_type preT old_proto_sctr0; val old_flat = enforce_type lthy range_type old_ssig_T old_flat0; val old_eval_core = enforce_type lthy domain_type old_eval_core_domT old_eval_core0; val old_eval = enforce_type lthy range_type fpT old_eval0; val old_algLam = enforce_type lthy range_type fpT old_algLam0; val ((embL, embL_def, embL_simps), lthy) = lthy |> define_embL embLN fp_b version Y Z fpT old_sig_T old_ssig_T k_T ssig_T Inl_const dead_old_sig_map Sig old_Oper old_VLeaf old_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_step fp_b version Y Z preT old_ssig_T ssig_T dead_pre_map unsig rho embL old_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old_sig_T k_T Sig old_proto_sctr; val pre_map_comp = map_comp_of_bnf pre_bnf; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_id = map_id_of_bnf dead_pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val fp_k_T_rel_eqs = map rel_eq_of_bnf (map_filter (bnf_of lthy) (fold add_type_namesT [fpT, k_T] [])); val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old_sig_map_comp = map_comp_of_bnf old_sig_bnf; val old_sig_map_cong = map_cong_of_bnf old_sig_bnf; val old_ssig_map_thms = #map_thms old_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old_sig_map_transfer = map_transfer_of_bnf old_sig_bnf; val sig_map_comp = map_comp_of_bnf sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old_ssig_induct = the_single (#co_inducts old_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def fp_k_T_rel_eqs [old_proto_sctr_transfer]; val embL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embL [embL_def] fp_k_T_rel_eqs [old_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def fp_k_T_rel_eqs [pre_map_transfer, old_Lam_transfer, embL_transfer, rho_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer fp_k_T_rel_eqs sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, Lam_natural_pointful, Lam_pointful_natural, flat_VLeaf, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT k_T sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embL_pointful_natural, old_algLam_pointful, eval_embL, algLam_algLam) = derive_embL_natural_etc lthy Inl_const old_ssig_bnf ssig_bnf Y Z preT fpT old_ssig_T ssig_T dead_pre_map Sig dead_old_ssig_map embL old_algLam algLam old_flat flat old_eval_core eval_core old_eval eval x f old_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old_sig_map_comp old_sig_map_cong old_ssig_map_thms old_Lam_pointful_natural Lam_def old_flat_simps flat_simps embL_simps embL_transfer old_eval_core_simps eval_core_simps old_eval_thm eval_thm old_dtor_algLam dtor_algLam old_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old_algLam_pointful algLam_algLam; val k_as_ssig = mk_k_as_ssig Z old_sig_T k_T ssig_T Sig dead_sig_map Oper VLeaf; val k_as_ssig' = substT Y fpT k_as_ssig; val algrho_rhs = HOLogic.mk_comp (eval, k_as_ssig'); val ((algrho, algrho_def), lthy) = lthy |> define_const true fp_b version algrhoN algrho_rhs; val k_as_ssig_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R k_as_ssig [] fp_k_T_rel_eqs [sig_map_transfer]; val k_as_ssig_natural = derive_natural_from_transfer lthy [(Y, Z)] [] [] f k_as_ssig k_as_ssig_transfer [ssig_bnf] [dead_k_bnf]; val k_as_ssig_natural_pointful = mk_pointful lthy k_as_ssig_natural; val [_, Lam_Inr] = derive_Lam_Inl_Inr lthy Y Z preT old_sig_T old_ssig_T k_T ssig_T dead_pre_map Sig embL old_Lam Lam rho unsig_thm Lam_def; val eval_core_k_as_ssig = derive_eval_core_k_as_ssig lthy Y preT k_T rho eval_core k_as_ssig x pre_map_comp dead_pre_map_id sig_map_comp ssig_map_thms Lam_natural_pointful Lam_Inr flat_VLeaf eval_core_simps; val algLam_algrho = derive_algLam_algrho lthy Y fpT Sig algLam algrho algLam_def algrho_def; val dtor_algrho = derive_dtor_algrho lthy Y Z preT fpT k_T ssig_T dead_pre_map dead_k_map dtor rho eval algrho x eval_thm k_as_ssig_natural_pointful eval_core_k_as_ssig algrho_def; val all_algLam_algs = algLam_algLam :: algLam_algrho :: old_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer fp_k_T_rel_eqs dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val Retr = enforce_type lthy (domain_type o domain_type) fpT old_Retr0; val (dtor_coinduct_info, lthy) = lthy |> derive_cong_step fp_b version fpT dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old_dtor_coinduct_info algrho_def k_as_ssig_transfer Retr equivp_Retr Retr_coinduct eval_embL embL_transfer old_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLam]), (algLam_algrhoN, [algLam_algrho]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_algrhoN, [dtor_algrho]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (embL_pointful_naturalN, [embL_pointful_natural]), (embL_transferN, [embL_transfer]), (evalN, [eval_thm]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (k_as_ssig_naturalN, [k_as_ssig_natural]), (k_as_ssig_transferN, [k_as_ssig_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (rho_transferN, [rho_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); val phi = Local_Theory.target_morphism lthy; in (({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_sig_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = dead_k_bnf :: old_all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info phi, ({algrho = algrho, dtor_algrho = dtor_algrho, algLam_algrho = algLam_algrho} |> morph_friend_info phi)), lthy |> Local_Theory.notes notes |> snd) end; fun derive_corecUU_merge (fpT as Type (fpT_name, res_Ds)) ({friend_names = old1_friend_names, sig_fp_sugars = old1_sig_fp_sugars as old1_sig_fp_sugar :: _, ssig_fp_sugar = old1_ssig_fp_sugar, Lam = old1_Lam0, proto_sctr = old1_proto_sctr0, flat = old1_flat0, eval_core = old1_eval_core0, eval = old1_eval0, algLam = old1_algLam0, dtor_transfer, Lam_transfer = old1_Lam_transfer, Lam_pointful_natural = old1_Lam_pointful_natural, proto_sctr_transfer = old1_proto_sctr_transfer, flat_simps = old1_flat_simps, eval_core_simps = old1_eval_core_simps, eval_thm = old1_eval_thm, all_algLam_algs = old1_all_algLam_algs, algLam_thm = old1_algLam_thm, dtor_algLam = old1_dtor_algLam, buffer = old1_buffer, all_dead_k_bnfs = old1_all_dead_k_bnfs, Retr = old1_Retr0, equivp_Retr, Retr_coinduct, dtor_coinduct_info = old1_dtor_coinduct_info, ...} : corec_info) ({friend_names = old2_friend_names, sig_fp_sugars = old2_sig_fp_sugars as old2_sig_fp_sugar :: _, ssig_fp_sugar = old2_ssig_fp_sugar, Lam = old2_Lam0, flat = old2_flat0, eval_core = old2_eval_core0, eval = old2_eval0, algLam = old2_algLam0, Lam_transfer = old2_Lam_transfer, Lam_pointful_natural = old2_Lam_pointful_natural, flat_simps = old2_flat_simps, eval_core_simps = old2_eval_core_simps, eval_thm = old2_eval_thm, all_algLam_algs = old2_all_algLam_algs, algLam_thm = old2_algLam_thm, dtor_algLam = old2_dtor_algLam, buffer = old2_buffer, all_dead_k_bnfs = old2_all_dead_k_bnfs, dtor_coinduct_info = old2_dtor_coinduct_info, ...} : corec_info) lthy = let val {T = Type (_, fpT_args0), pre_bnf = live_pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf (length fpT_args0) live_fp_bnf; val ((Ds, [Y, Z]), names_lthy) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val live_EsFs = []; val live_AsBs = live_EsFs @ [(Y, Z)]; val live = length live_EsFs; val ((((x, f), g), R), _) = names_lthy |> yield_singleton (mk_Frees "x") Y ||>> yield_singleton (mk_Frees "f") (Y --> Z) ||>> yield_singleton (mk_Frees "g") (Y --> Z) ||>> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); (* FIXME *) val fs = []; val Rs = []; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val dtor = mk_dtor res_Ds (the_single (#dtors fp_res)); val old1_sig_T_name = fst (dest_Type (#T old1_sig_fp_sugar)); val old2_sig_T_name = fst (dest_Type (#T old2_sig_fp_sugar)); val old1_ssig_T_name = fst (dest_Type (#T old1_ssig_fp_sugar)); val old2_ssig_T_name = fst (dest_Type (#T old2_ssig_fp_sugar)); val fp_alives = map (K false) live_fp_alives; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val res_Bs = res_Ds @ [Z]; val preT = pre_type_of_ctor Y ctor; val YpreT = HOLogic.mk_prodT (Y, preT); val fpT0 = Type (fpT_name, Ds); val old1_sig_T0 = Type (old1_sig_T_name, As); val old2_sig_T0 = Type (old2_sig_T_name, As); val old1_sig_T = Type (old1_sig_T_name, res_As); val old2_sig_T = Type (old2_sig_T_name, res_As); val old1_ssig_T = Type (old1_ssig_T_name, res_As); val old2_ssig_T = Type (old2_ssig_T_name, res_As); val old1_Lam_domT = Tsubst Y YpreT old1_sig_T; val old2_Lam_domT = Tsubst Y YpreT old2_sig_T; val old1_eval_core_domT = Tsubst Y YpreT old1_ssig_T; val old2_eval_core_domT = Tsubst Y YpreT old2_ssig_T; val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val ((sig_fp_sugar, ssig_fp_sugar), lthy) = lthy |> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old1_sig_T0, old2_sig_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val sig_T_name = fst (dest_Type (#T sig_fp_sugar)); val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val old1_sig_bnf = #fp_bnf old1_sig_fp_sugar; val old2_sig_bnf = #fp_bnf old2_sig_fp_sugar; val old1_ssig_bnf = #fp_bnf old1_ssig_fp_sugar; val old2_ssig_bnf = #fp_bnf old2_ssig_fp_sugar; val sig_bnf = #fp_bnf sig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ((((((((dead_pre_bnf, dead_fp_bnf), dead_old1_sig_bnf), dead_old2_sig_bnf), dead_old1_ssig_bnf), dead_old2_ssig_bnf), dead_sig_bnf), dead_ssig_bnf), lthy) = lthy |> bnf_kill_all_but 1 live_pre_bnf ||>> bnf_kill_all_but 0 live_fp_bnf ||>> bnf_kill_all_but 1 old1_sig_bnf ||>> bnf_kill_all_but 1 old2_sig_bnf ||>> bnf_kill_all_but 1 old1_ssig_bnf ||>> bnf_kill_all_but 1 old2_ssig_bnf ||>> bnf_kill_all_but 1 sig_bnf ||>> bnf_kill_all_but 1 ssig_bnf; (* FIXME *) val pre_bnf = dead_pre_bnf; val fp_bnf = dead_fp_bnf; val old1_ssig_fp_ctr_sugar = #fp_ctr_sugar old1_ssig_fp_sugar; val old2_ssig_fp_ctr_sugar = #fp_ctr_sugar old2_ssig_fp_sugar; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_fp_bnf_sugar = #fp_bnf_sugar sig_fp_sugar; val old1_ssig_fp_bnf_sugar = #fp_bnf_sugar old1_ssig_fp_sugar; val old2_ssig_fp_bnf_sugar = #fp_bnf_sugar old2_ssig_fp_sugar; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val old1_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old1_ssig_fp_sugar); val old2_ssig_fp_induct_sugar = the (#fp_co_induct_sugar old2_ssig_fp_sugar); val ssig_fp_induct_sugar = the (#fp_co_induct_sugar ssig_fp_sugar); val old1_ssig_ctr_sugar = #ctr_sugar old1_ssig_fp_ctr_sugar; val old2_ssig_ctr_sugar = #ctr_sugar old2_ssig_fp_ctr_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val sig_T = Type (sig_T_name, res_As); val ssig_T = Type (ssig_T_name, res_As); val pre_map = mk_mapN lthy live_AsBs preT pre_bnf; val pre_rel = mk_relN lthy live_AsBs preT pre_bnf; val dead_pre_map = mk_map1 lthy Y Z preT dead_pre_bnf; val dead_pre_rel = mk_rel1 lthy Y Z preT dead_pre_bnf; val fp_rel = mk_relN lthy live_EsFs fpT fp_bnf; val dtor_unfold = mk_co_rec (Proof_Context.theory_of lthy) Greatest_FP [Z] fpT (the_single (#xtor_un_folds fp_res)); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val unsig = mk_disc_or_sel res_As (the_single (the_single (#selss sig_ctr_sugar))); val sig_rel = mk_relN lthy live_AsBs sig_T sig_bnf; val dead_old1_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_sig_bnf); val dead_old2_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_sig_bnf); val dead_sig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_sig_bnf); val dead_sig_rel = mk_rel 1 res_As res_Bs (rel_of_bnf dead_sig_bnf); val [old1_Oper, old1_VLeaf, old1_CLeaf] = map (mk_ctr res_As) (#ctrs old1_ssig_ctr_sugar); val [old2_Oper, old2_VLeaf, old2_CLeaf] = map (mk_ctr res_As) (#ctrs old2_ssig_ctr_sugar); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val old1_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old1_ssig_bnf); val old2_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_old2_ssig_bnf); val ssig_map = mk_mapN lthy live_AsBs ssig_T ssig_bnf; val ssig_rel = mk_relN lthy live_AsBs ssig_T ssig_bnf; val dead_ssig_map = mk_map 1 res_As res_Bs (map_of_bnf dead_ssig_bnf); val old1_Lam = enforce_type lthy domain_type old1_Lam_domT old1_Lam0; val old2_Lam = enforce_type lthy domain_type old2_Lam_domT old2_Lam0; val old1_proto_sctr = enforce_type lthy domain_type preT old1_proto_sctr0; val old1_flat = enforce_type lthy range_type old1_ssig_T old1_flat0; val old2_flat = enforce_type lthy range_type old2_ssig_T old2_flat0; val old1_eval_core = enforce_type lthy domain_type old1_eval_core_domT old1_eval_core0; val old2_eval_core = enforce_type lthy domain_type old2_eval_core_domT old2_eval_core0; val old1_eval = enforce_type lthy range_type fpT old1_eval0; val old2_eval = enforce_type lthy range_type fpT old2_eval0; val old1_algLam = enforce_type lthy range_type fpT old1_algLam0; val old2_algLam = enforce_type lthy range_type fpT old2_algLam0; val ((embLL, embLL_def, embLL_simps), lthy) = lthy |> define_embL embLLN fp_b version Y Z fpT old1_sig_T old1_ssig_T old2_sig_T ssig_T Inl_const dead_old1_sig_map Sig old1_Oper old1_VLeaf old1_CLeaf Oper VLeaf CLeaf; val ((embLR, embLR_def, embLR_simps), lthy) = lthy |> define_embL embLRN fp_b version Y Z fpT old2_sig_T old2_ssig_T old1_sig_T ssig_T (fn T => fn U => Inr_const U T) dead_old2_sig_map Sig old2_Oper old2_VLeaf old2_CLeaf Oper VLeaf CLeaf; val ((Lam, Lam_def), lthy) = lthy |> define_Lam_merge fp_b version Y Z preT old1_ssig_T old2_ssig_T ssig_T dead_pre_map unsig embLL embLR old1_Lam old2_Lam; val ((proto_sctr, proto_sctr_def), lthy) = lthy |> define_proto_sctr_step_or_merge fp_b version old1_sig_T old2_sig_T Sig old1_proto_sctr; val pre_map_transfer = map_transfer_of_bnf pre_bnf; val dead_pre_map_comp0 = map_comp0_of_bnf dead_pre_bnf; val dead_pre_map_comp = map_comp_of_bnf dead_pre_bnf; val fp_map_id = map_id_of_bnf fp_bnf; val fp_rel_eq = rel_eq_of_bnf fp_bnf; val [ctor_dtor] = #ctor_dtors fp_res; val [dtor_inject] = #dtor_injects fp_res; val [dtor_unfold_thm] = #xtor_un_fold_thms fp_res; val dtor_unfold_unique = #xtor_un_fold_unique fp_res; val [dtor_unfold_transfer] = #xtor_un_fold_transfers fp_res; val unsig_thm = the_single (the_single (#sel_thmss sig_ctr_sugar)); val [sig_map_thm] = #map_thms sig_fp_bnf_sugar; val old1_sig_map_comp = map_comp_of_bnf old1_sig_bnf; val old2_sig_map_comp = map_comp_of_bnf old2_sig_bnf; val old1_sig_map_cong = map_cong_of_bnf old1_sig_bnf; val old2_sig_map_cong = map_cong_of_bnf old2_sig_bnf; val old1_ssig_map_thms = #map_thms old1_ssig_fp_bnf_sugar; val old2_ssig_map_thms = #map_thms old2_ssig_fp_bnf_sugar; val [Oper_map_thm, VLeaf_map_thm, CLeaf_map_thm] = #map_thms ssig_fp_bnf_sugar; val old1_sig_map_transfer = map_transfer_of_bnf old1_sig_bnf; val old2_sig_map_transfer = map_transfer_of_bnf old2_sig_bnf; val sig_map_transfer = map_transfer_of_bnf sig_bnf; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val ssig_map_transfer = map_transfer_of_bnf ssig_bnf; val old1_ssig_induct = the_single (#co_inducts old1_ssig_fp_induct_sugar); val old2_ssig_induct = the_single (#co_inducts old2_ssig_fp_induct_sugar); val ssig_induct = the_single (#co_inducts ssig_fp_induct_sugar); val proto_sctr_transfer = derive_proto_sctr_transfer_step_or_merge lthy Y Z R dead_pre_rel dead_sig_rel proto_sctr proto_sctr_def [] [old1_proto_sctr_transfer]; val embLL_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLL [embLL_def] [] [old1_sig_map_transfer]; val embLR_transfer = derive_transfer_by_transfer_prover lthy live_AsBs Rs R embLR [embLR_def] [] [old2_sig_map_transfer]; val Lam_transfer = derive_Lam_or_eval_core_transfer lthy live_AsBs Y Z preT ssig_T Rs R pre_rel sig_rel ssig_rel Lam Lam_def [] [pre_map_transfer, old1_Lam_transfer, old2_Lam_transfer, embLL_transfer, embLR_transfer]; val ((((((((flat, _, flat_simps), flat_transfer), ((eval_core, _, eval_core_simps), eval_core_transfer)), (eval, eval_def)), (cutSsig, cutSsig_def)), (algLam, algLam_def)), (corecU, corecU_def)), lthy) = lthy |> define_flat_etc fp_b version live_AsBs Y Z preT fpT sig_T ssig_T Oper VLeaf CLeaf pre_rel dead_pre_map dtor dtor_unfold dead_sig_map ssig_rel dead_ssig_map Lam Rs R pre_map_transfer [fp_rel_eq] sig_map_transfer ssig_map_transfer Lam_transfer dtor_transfer; val (Sig_pointful_natural, flat_pointful_natural, _, Lam_pointful_natural, _, eval_core_pointful_natural, eval_thm, eval_flat, eval_simps as [eval_Oper, _, _], corecU_ctor, corecU_unique, dtor_algLam) = derive_Sig_natural_etc lthy live live_AsBs Y Z preT fpT fpT sig_T ssig_T pre_map dead_pre_map ctor dtor Sig dead_sig_map Oper VLeaf CLeaf ssig_map dead_ssig_map Lam flat eval_core eval cutSsig algLam corecU x fs f g ctor_dtor dtor_inject dtor_unfold_thm dtor_unfold_unique sig_map_thm ssig_induct ssig_map_thms Oper_map_thm VLeaf_map_thm CLeaf_map_thm Lam_transfer flat_simps flat_transfer eval_core_simps eval_core_transfer eval_def cutSsig_def algLam_def corecU_def live_pre_bnf pre_bnf dead_pre_bnf fp_bnf sig_bnf ssig_bnf dead_ssig_bnf; val proto_sctr_natural = derive_natural_from_transfer_with_pre_type lthy live_AsBs Y Z preT ssig_T pre_map ssig_map fs f proto_sctr proto_sctr_transfer [pre_bnf, sig_bnf] []; val proto_sctr_pointful_natural = mk_pointful lthy proto_sctr_natural RS sym; val (embLL_pointful_natural, old1_algLam_pointful, eval_embLL, algLam_algLamL) = derive_embL_natural_etc lthy Inl_const old1_ssig_bnf ssig_bnf Y Z preT fpT old1_ssig_T ssig_T dead_pre_map Sig old1_ssig_map embLL old1_algLam algLam old1_flat flat old1_eval_core eval_core old1_eval eval x f old1_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old1_sig_map_comp old1_sig_map_cong old1_ssig_map_thms old1_Lam_pointful_natural Lam_def old1_flat_simps flat_simps embLL_simps embLL_transfer old1_eval_core_simps eval_core_simps old1_eval_thm eval_thm old1_dtor_algLam dtor_algLam old1_algLam_thm; val (embLR_pointful_natural, _, eval_embLR, algLam_algLamR) = derive_embL_natural_etc lthy Inr_const old2_ssig_bnf ssig_bnf Y Z preT fpT old2_ssig_T ssig_T dead_pre_map Sig old2_ssig_map embLR old2_algLam algLam old2_flat flat old2_eval_core eval_core old2_eval eval x f old2_ssig_induct dead_pre_map_comp0 dead_pre_map_comp fp_map_id dtor_inject dtor_unfold_unique Sig_pointful_natural unsig_thm sig_map_thm old2_sig_map_comp old2_sig_map_cong old2_ssig_map_thms old2_Lam_pointful_natural Lam_def old2_flat_simps flat_simps embLR_simps embLR_transfer old2_eval_core_simps eval_core_simps old2_eval_thm eval_thm old2_dtor_algLam dtor_algLam old2_algLam_thm; val algLam_thm = derive_algLam_step_or_merge lthy Y fpT ctor proto_sctr algLam proto_sctr_def old1_algLam_pointful algLam_algLamL; val all_algLam_algs = algLam_algLamL :: algLam_algLamR :: merge_lists (Thm.eq_thm_prop o apply2 zero_var_indexes) old1_all_algLam_algs old2_all_algLam_algs; val (((corecUU, _), corecUU_thm, corecUU_unique, corecUU_transfer, _, sctr_transfer, sctr_pointful_natural), lthy) = lthy |> define_corecUU_etc fp_b version live_AsBs Y Z preT fpT ssig_T pre_map dead_pre_map pre_rel fp_rel ctor Oper ssig_map dead_ssig_map ssig_rel proto_sctr flat eval_core eval corecU fs f g Rs R pre_map_transfer [] dtor_unfold_transfer dtor_transfer ssig_map_transfer proto_sctr_transfer proto_sctr_pointful_natural flat_transfer flat_pointful_natural eval_core_transfer eval_core_pointful_natural eval_thm eval_flat eval_Oper algLam_thm cutSsig_def corecU_def corecU_ctor corecU_unique pre_bnf dead_pre_bnf fp_res ssig_fp_sugar; val Retr = enforce_type lthy (domain_type o domain_type) fpT old1_Retr0; val embed_Sig_inl = embed_Sig lthy Sig (Inl_const old1_sig_T old2_sig_T); val embed_Sig_inr = embed_Sig lthy Sig (Inr_const old1_sig_T old2_sig_T); val ctr_wrapper = embed_Sig_inl (#ctr_wrapper old1_buffer); val friends = Symtab.merge (K true) (Symtab.map (K (apsnd embed_Sig_inl)) (#friends old1_buffer), Symtab.map (K (apsnd embed_Sig_inr)) (#friends old2_buffer)); val old_fp_sugars = merge_lists (op = o apply2 (fst o dest_Type o #T)) old1_sig_fp_sugars old2_sig_fp_sugars; val ((dtor_coinduct_info, all_dead_k_bnfs, friend_names), lthy) = lthy |> derive_cong_merge fp_b version fpT old1_friend_names old2_friend_names dead_ssig_bnf dead_pre_bnf eval eval_thm eval_core_transfer old1_dtor_coinduct_info old2_dtor_coinduct_info Retr equivp_Retr Retr_coinduct eval_embLL embLL_transfer eval_embLR embLR_transfer old1_all_dead_k_bnfs old2_all_dead_k_bnfs; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; val notes = [(corecUU_transferN, [corecUU_transfer])] @ (if Config.get lthy bnf_internals then [(algLamN, [algLam_thm]), (algLam_algLamN, [algLam_algLamL, algLam_algLamR]), (cong_alg_introsN, #cong_alg_intros dtor_coinduct_info), (cong_localeN, [#cong_locale dtor_coinduct_info]), (corecU_ctorN, [corecU_ctor]), (corecU_uniqueN, [corecU_unique]), (corecUUN, [corecUU_thm]), (corecUU_uniqueN, [corecUU_unique]), (dtor_algLamN, [dtor_algLam]), (dtor_coinductN, [#dtor_coinduct dtor_coinduct_info]), (eval_core_pointful_naturalN, [eval_core_pointful_natural]), (eval_core_transferN, [eval_core_transfer]), (embL_pointful_naturalN, [embLL_pointful_natural, embLR_pointful_natural]), (embL_transferN, [embLL_transfer, embLR_transfer]), (evalN, [eval_thm]), (eval_flatN, [eval_flat]), (eval_simpsN, eval_simps), (flat_pointful_naturalN, [flat_pointful_natural]), (flat_transferN, [flat_transfer]), (Lam_pointful_naturalN, [Lam_pointful_natural]), (Lam_transferN, [Lam_transfer]), (proto_sctr_pointful_naturalN, [proto_sctr_pointful_natural]), (proto_sctr_transferN, [proto_sctr_transfer]), (sctr_pointful_naturalN, [sctr_pointful_natural]), (sctr_transferN, [sctr_transfer]), (Sig_pointful_naturalN, [Sig_pointful_natural])] else []) |> map (fn (thmN, thms) => ((mk_version_fp_binding true lthy version fp_b thmN, []), [(thms, [])])); in ({fp_b = fp_b, version = version, fpT = fpT, Y = Y, Z = Z, friend_names = friend_names, sig_fp_sugars = sig_fp_sugar :: old_fp_sugars, ssig_fp_sugar = ssig_fp_sugar, Lam = Lam, proto_sctr = proto_sctr, flat = flat, eval_core = eval_core, eval = eval, algLam = algLam, corecUU = corecUU, dtor_transfer = dtor_transfer, Lam_transfer = Lam_transfer, Lam_pointful_natural = Lam_pointful_natural, proto_sctr_transfer = proto_sctr_transfer, flat_simps = flat_simps, eval_core_simps = eval_core_simps, eval_thm = eval_thm, eval_simps = eval_simps, all_algLam_algs = all_algLam_algs, algLam_thm = algLam_thm, dtor_algLam = dtor_algLam, corecUU_thm = corecUU_thm, corecUU_unique = corecUU_unique, corecUU_transfer = corecUU_transfer, buffer = buffer, all_dead_k_bnfs = all_dead_k_bnfs, Retr = Retr, equivp_Retr = equivp_Retr, Retr_coinduct = Retr_coinduct, dtor_coinduct_info = dtor_coinduct_info} |> morph_corec_info (Local_Theory.target_morphism lthy), lthy |> Local_Theory.notes notes |> snd) end; fun set_corec_info_exprs fpT_name f = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi => let val exprs = f phi in Data.map (apsnd (fn [info_tab] => [Symtab.map_default (fpT_name, exprs) (K exprs) info_tab])) end); fun subsume_corec_info_ad ctxt {fpT = fpT1, friend_names = friend_names1} {fpT = fpT2, friend_names = friend_names2} = Sign.typ_instance (Proof_Context.theory_of ctxt) (fpT1, fpT2) andalso subset (op =) (friend_names1, friend_names2); fun subsume_corec_info_expr ctxt expr1 expr2 = subsume_corec_info_ad ctxt (corec_ad_of_expr expr1) (corec_ad_of_expr expr2); fun instantiate_corec_info thy res_T (info as {fpT, ...}) = let val As_rho = tvar_subst thy [fpT] [res_T]; val substAT = Term.typ_subst_TVars As_rho; val substA = Term.subst_TVars As_rho; val phi = Morphism.typ_morphism "BNF" substAT $> Morphism.term_morphism "BNF" substA; in morph_corec_info phi info end; fun instantiate_corec_info_expr thy res_T (Ad ({friend_names, ...}, f)) = Ad ({fpT = res_T, friend_names = friend_names}, f #>> instantiate_corec_info thy res_T) | instantiate_corec_info_expr thy res_T (Info info) = Info (instantiate_corec_info thy res_T info); fun ensure_Info expr = corec_info_of_expr expr #>> Info and ensure_Info_if_Info old_expr (expr, lthy) = if is_Info old_expr then ensure_Info expr lthy else (expr, lthy) and merge_corec_info_exprs old_exprs expr1 expr2 lthy = if subsume_corec_info_expr lthy expr2 expr1 then if subsume_corec_info_expr lthy expr1 expr2 andalso is_Ad expr1 then (expr2, lthy) else ensure_Info_if_Info expr2 (expr1, lthy) else if subsume_corec_info_expr lthy expr1 expr2 then ensure_Info_if_Info expr1 (expr2, lthy) else let val thy = Proof_Context.theory_of lthy; val {fpT = fpT1, friend_names = friend_names1} = corec_ad_of_expr expr1; val {fpT = fpT2, friend_names = friend_names2} = corec_ad_of_expr expr2; val fpT0 = typ_unify_disjointly thy (fpT1, fpT2); val fpT = singleton (freeze_types lthy []) fpT0; val friend_names = merge_lists (op =) friend_names1 friend_names2; val expr = Ad ({fpT = fpT, friend_names = friend_names}, corec_info_of_expr expr1 ##>> corec_info_of_expr expr2 #-> uncurry (derive_corecUU_merge fpT)); val old_same_type_exprs = if old_exprs then [] |> Sign.typ_instance thy (fpT1, fpT0) ? cons expr1 |> Sign.typ_instance thy (fpT2, fpT0) ? cons expr2 else []; in (expr, lthy) |> fold ensure_Info_if_Info old_same_type_exprs end and insert_corec_info_expr expr exprs lthy = let val thy = Proof_Context.theory_of lthy; val {fpT = new_fpT, ...} = corec_ad_of_expr expr; val is_Tinst = curry (Sign.typ_instance thy); fun is_Tequiv T U = is_Tinst T U andalso is_Tinst U T; val (((equiv_exprs, sub_exprs), sup_exprs), incomp_exprs) = exprs |> List.partition ((fn {fpT, ...} => is_Tequiv fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst fpT new_fpT) o corec_ad_of_expr) ||>> List.partition ((fn {fpT, ...} => is_Tinst new_fpT fpT) o corec_ad_of_expr); fun add_instantiated_incomp_expr expr exprs = let val {fpT, ...} = corec_ad_of_expr expr in (case try (typ_unify_disjointly thy) (fpT, new_fpT) of SOME new_T => let val subsumes = (fn {fpT, ...} => is_Tinst new_T fpT) o corec_ad_of_expr in if exists (exists subsumes) [exprs, sub_exprs] then exprs else instantiate_corec_info_expr thy new_T expr :: exprs end | NONE => exprs) end; val unincomp_exprs = fold add_instantiated_incomp_expr incomp_exprs []; val ((merged_sub_exprs, merged_unincomp_exprs), lthy) = lthy |> fold_map (merge_corec_info_exprs true expr) sub_exprs ||>> fold_map (merge_corec_info_exprs false expr) unincomp_exprs; val (merged_equiv_expr, lthy) = (expr, lthy) |> fold (uncurry o merge_corec_info_exprs true) equiv_exprs; in (merged_unincomp_exprs @ merged_sub_exprs @ merged_equiv_expr :: sup_exprs @ incomp_exprs |> sort (rev_order o int_ord o apply2 (length o #friend_names o corec_ad_of_expr)), lthy) end and register_corec_info (info as {fpT = Type (fpT_name, _), ...}) lthy = let val (exprs, lthy) = insert_corec_info_expr (Info info) (corec_info_exprs_of lthy fpT_name) lthy; in lthy |> set_corec_info_exprs fpT_name (fn phi => map (morph_corec_info_expr phi) exprs) end and corec_info_of_expr (Ad (_, f)) lthy = f lthy | corec_info_of_expr (Info info) lthy = (info, lthy); fun nonempty_corec_info_exprs_of fpT_name lthy = (case corec_info_exprs_of lthy fpT_name of [] => derive_corecUU_base fpT_name lthy |> (fn (info, lthy) => ([Info info], lthy |> set_corec_info_exprs fpT_name (fn phi => [Info (morph_corec_info phi info)]))) | exprs => (exprs, lthy)); fun corec_info_of res_T lthy = (case res_T of Type (fpT_name, _) => let val (exprs, lthy) = nonempty_corec_info_exprs_of fpT_name lthy; val thy = Proof_Context.theory_of lthy; val expr = (case find_first ((fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) o corec_ad_of_expr) exprs of SOME expr => expr | NONE => error ("Invalid type: " ^ Syntax.string_of_typ lthy res_T)); val (info, lthy) = corec_info_of_expr expr lthy; in (instantiate_corec_info thy res_T info, lthy |> is_Ad expr ? register_corec_info info) end | _ => not_codatatype lthy res_T); fun maybe_corec_info_of ctxt res_T = (case res_T of Type (fpT_name, _) => let val thy = Proof_Context.theory_of ctxt; val infos = corec_infos_of ctxt fpT_name; in find_first (fn {fpT, ...} => Sign.typ_instance thy (res_T, fpT)) infos |> Option.map (instantiate_corec_info thy res_T) end | _ => not_codatatype ctxt res_T); fun prepare_friend_corec friend_name friend_T lthy = let val (arg_Ts, res_T) = strip_type friend_T; val Type (fpT_name, res_Ds) = (case res_T of T as Type _ => T | T => error (not_codatatype lthy T)); val _ = not (null arg_Ts) orelse error "Function with no argument cannot be registered as friend"; val {T = Type (fpT_name, fpT_args0), pre_bnf, fp_bnf = live_fp_bnf, fp_res, ...} = checked_fp_sugar_of lthy fpT_name; val num_fp_tyargs = length fpT_args0; val fpT_Ss = map Type.sort_of_atyp fpT_args0; val live_fp_alives = liveness_of_fp_bnf num_fp_tyargs live_fp_bnf; val (old_info as {friend_names = old_friend_names, sig_fp_sugars = old_sig_fp_sugar :: _, buffer = old_buffer, ...}, lthy) = corec_info_of res_T lthy; val old_sig_T_name = fst (dest_Type (#T old_sig_fp_sugar)); val old_sig_alives = liveness_of_fp_bnf (num_fp_tyargs + 1) (#fp_bnf old_sig_fp_sugar); (* FIXME: later *) val fp_alives = fst (split_last old_sig_alives); val fp_alives = map (K false) live_fp_alives; val _ = not (member (op =) old_friend_names friend_name) orelse error ("Function " ^ quote (Syntax.string_of_term lthy (Const (friend_name, friend_T))) ^ " already registered as friend"); val lthy = lthy |> Variable.declare_typ friend_T; val ((Ds, [Y, Z]), _) = lthy |> mk_TFrees' fpT_Ss ||>> mk_TFrees 2; (* FIXME *) val dead_Ds = Ds; val live_As = [Y]; val ctor = mk_ctor res_Ds (the_single (#ctors fp_res)); val fpT0 = Type (fpT_name, Ds); val k_Ts0 = map (typ_subst_nonatomic (res_Ds ~~ Ds) o typ_subst_nonatomic [(res_T, Y)]) arg_Ts; val k_T0 = mk_tupleT_balanced k_Ts0; val As = Ds @ [Y]; val res_As = res_Ds @ [Y]; val k_As = fold Term.add_tfreesT k_Ts0 []; val _ = (case filter_out (member (op =) As o TFree) k_As of [] => () | k_A :: _ => error ("Cannot have type variable " ^ quote (Syntax.string_of_typ lthy (TFree k_A)) ^ " in the argument types when it does not occur as an immediate argument of the result \ \type constructor")); val substDT = Term.typ_subst_atomic (Ds ~~ res_Ds); val old_sig_T0 = Type (old_sig_T_name, As); val ((fp_b, version), lthy) = lthy |> get_name_next_version_of fpT_name; val (((dead_k_bnf, sig_fp_sugar), ssig_fp_sugar), lthy) = lthy |> bnf_with_deads_and_lives dead_Ds live_As Y fpT0 k_T0 ||>> define_sig_type fp_b version fp_alives Ds Y (mk_sumT (old_sig_T0, k_T0)) ||>> define_ssig_type fp_b version fp_alives Ds Y fpT0; val _ = live_of_bnf dead_k_bnf = 1 orelse error "Impossible type for friend (the result codatatype must occur live in the arguments)"; val (dead_pre_bnf, lthy) = lthy |> bnf_kill_all_but 1 pre_bnf; val sig_fp_ctr_sugar = #fp_ctr_sugar sig_fp_sugar; val ssig_fp_ctr_sugar = #fp_ctr_sugar ssig_fp_sugar; val sig_ctr_sugar = #ctr_sugar sig_fp_ctr_sugar; val ssig_ctr_sugar = #ctr_sugar ssig_fp_ctr_sugar; val ssig_T_name = fst (dest_Type (#T ssig_fp_sugar)); val preT = pre_type_of_ctor Y ctor; val old_sig_T = substDT old_sig_T0; val k_T = substDT k_T0; val ssig_T = Type (ssig_T_name, res_As); val Sig = mk_ctr res_As (the_single (#ctrs sig_ctr_sugar)); val [Oper, VLeaf, CLeaf] = map (mk_ctr res_As) (#ctrs ssig_ctr_sugar); val (ctr_wrapper, friends) = mk_ctr_wrapper_friends lthy friend_name friend_T old_sig_T k_T Sig old_buffer; val buffer = {Oper = Oper, VLeaf = VLeaf, CLeaf = CLeaf, ctr_wrapper = ctr_wrapper, friends = friends}; in ((old_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy) end; fun register_friend_corec key fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar friend_const rho rho_transfer old_info lthy = let val friend_T = fastype_of friend_const; val res_T = body_type friend_T; in derive_corecUU_step res_T old_info key friend_T fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar rho rho_transfer lthy |> (fn ((info, friend_info), lthy) => (friend_info, register_corec_info info lthy)) end; fun merge_corec_info_exprss exprs1 exprs2 lthy = let fun all_friend_names_of exprs = fold (union (op =)) (map (#friend_names o corec_ad_of_expr) exprs) []; val friend_names1 = all_friend_names_of exprs1; val friend_names2 = all_friend_names_of exprs2; in if subset (op =) (friend_names2, friend_names1) then if subset (op =) (friend_names1, friend_names2) andalso length (filter is_Info exprs2) > length (filter is_Info exprs1) then (exprs2, lthy) else (exprs1, lthy) else if subset (op =) (friend_names1, friend_names2) then (exprs2, lthy) else fold_rev (uncurry o insert_corec_info_expr) exprs2 (exprs1, lthy) end; fun merge_corec_info_tabs info_tab1 info_tab2 lthy = let val fpT_names = union (op =) (Symtab.keys info_tab1) (Symtab.keys info_tab2); fun add_infos_of fpT_name (info_tab, lthy) = (case Symtab.lookup info_tab1 fpT_name of NONE => (Symtab.update_new (fpT_name, the (Symtab.lookup info_tab2 fpT_name)) info_tab, lthy) | SOME exprs1 => (case Symtab.lookup info_tab2 fpT_name of NONE => (Symtab.update_new (fpT_name, exprs1) info_tab, lthy) | SOME exprs2 => let val (exprs, lthy) = merge_corec_info_exprss exprs1 exprs2 lthy in (Symtab.update_new (fpT_name, exprs) info_tab, lthy) end)); in fold add_infos_of fpT_names (Symtab.empty, lthy) end; fun consolidate lthy = (case snd (Data.get (Context.Proof lthy)) of [_] => raise Same.SAME | info_tab :: info_tabs => let val (info_tab', lthy) = fold_rev (uncurry o merge_corec_info_tabs) info_tabs (info_tab, lthy); in - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} (fn phi => Data.map (apsnd (fn _ => [Symtab.map (K (map (morph_corec_info_expr phi))) info_tab']))) lthy end); fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML --- a/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_grec_sugar.ML @@ -1,2391 +1,2394 @@ (* Title: HOL/Tools/BNF/bnf_gfp_grec_sugar.ML Author: Aymeric Bouzy, Ecole polytechnique Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, ETH Zürich Copyright 2015, 2016 Generalized corecursor sugar ("corec" and friends). *) signature BNF_GFP_GREC_SUGAR = sig datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option val parse_corec_equation: Proof.context -> term list -> term -> term list * term val explore_corec_equation: Proof.context -> bool -> bool -> string -> term -> BNF_GFP_Grec_Sugar_Util.s_parse_info -> typ -> term list * term -> term list * term val build_corecUU_arg_and_goals: bool -> term -> term list * term -> local_theory -> (((thm list * thm list * thm list) * term list) * term) * local_theory val derive_eq_corecUU: Proof.context -> BNF_GFP_Grec.corec_info -> term -> term -> thm -> thm val derive_unique: Proof.context -> morphism -> term -> BNF_GFP_Grec.corec_info -> string -> thm -> thm val corec_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> local_theory val corecursive_cmd: bool -> corec_option list -> (binding * string option * mixfix) list * string -> local_theory -> Proof.state val friend_of_corec_cmd: (string * string option) * string -> local_theory -> Proof.state val coinduction_upto_cmd: string * string -> local_theory -> local_theory end; structure BNF_GFP_Grec_Sugar : BNF_GFP_GREC_SUGAR = struct open Ctr_Sugar open BNF_Util open BNF_Tactics open BNF_Def open BNF_Comp open BNF_FP_Util open BNF_FP_Def_Sugar open BNF_FP_N2M_Sugar open BNF_GFP_Util open BNF_GFP_Rec_Sugar open BNF_FP_Rec_Sugar_Transfer open BNF_GFP_Grec open BNF_GFP_Grec_Sugar_Util open BNF_GFP_Grec_Sugar_Tactics val cong_N = "cong_"; val baseN = "base"; val reflN = "refl"; val symN = "sym"; val transN = "trans"; val cong_introsN = prefix cong_N "intros"; val codeN = "code"; val coinductN = "coinduct"; val coinduct_uptoN = "coinduct_upto"; val corecN = "corec"; val ctrN = "ctr"; val discN = "disc"; val disc_iffN = "disc_iff"; val eq_algrhoN = "eq_algrho"; val eq_corecUUN = "eq_corecUU"; val friendN = "friend"; val inner_elimN = "inner_elim"; val inner_inductN = "inner_induct"; val inner_simpN = "inner_simp"; val rhoN = "rho"; val selN = "sel"; val uniqueN = "unique"; val inner_fp_suffix = "_inner_fp"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val unfold_id_thms1 = map (fn thm => thm RS eq_reflection) @{thms id_bnf_o o_id_bnf id_apply o_apply} @ @{thms fst_def[abs_def, symmetric] snd_def[abs_def, symmetric]}; fun unfold_id_bnf_etc lthy = let val thy = Proof_Context.theory_of lthy in Raw_Simplifier.rewrite_term thy unfold_id_thms1 [] #> Raw_Simplifier.rewrite_term thy @{thms BNF_Composition.id_bnf_def} [] end; datatype corec_option = Plugins_Option of Proof.context -> Plugin_Name.filter | Friend_Option | Transfer_Option; val corec_option_parser = Parse.group (K "option") (Plugin_Name.parse_filter >> Plugins_Option || Parse.reserved "friend" >> K Friend_Option || Parse.reserved "transfer" >> K Transfer_Option); type codatatype_extra = {case_dtor: thm, case_trivial: thm, abs_rep_transfers: thm list}; fun morph_codatatype_extra phi ({case_dtor, case_trivial, abs_rep_transfers} : codatatype_extra) = {case_dtor = Morphism.thm phi case_dtor, case_trivial = Morphism.thm phi case_trivial, abs_rep_transfers = map (Morphism.thm phi) abs_rep_transfers}; val transfer_codatatype_extra = morph_codatatype_extra o Morphism.transfer_morphism; type coinduct_extra = {coinduct: thm, coinduct_attrs: Token.src list, cong_intro_pairs: (string * thm) list}; fun morph_coinduct_extra phi ({coinduct, coinduct_attrs, cong_intro_pairs} : coinduct_extra) = {coinduct = Morphism.thm phi coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = map (apsnd (Morphism.thm phi)) cong_intro_pairs}; val transfer_coinduct_extra = morph_coinduct_extra o Morphism.transfer_morphism; type friend_extra = {eq_algrhos: thm list, algrho_eqs: thm list}; val empty_friend_extra = {eq_algrhos = [], algrho_eqs = []}; fun merge_friend_extras ({eq_algrhos = eq_algrhos1, algrho_eqs = algrho_eqs1}, {eq_algrhos = eq_algrhos2, algrho_eqs = algrho_eqs2}) = {eq_algrhos = union Thm.eq_thm_prop eq_algrhos1 eq_algrhos2, algrho_eqs = union Thm.eq_thm_prop algrho_eqs1 algrho_eqs2}; type corec_sugar_data = codatatype_extra Symtab.table * coinduct_extra Symtab.table * friend_extra Symtab.table; structure Data = Generic_Data ( type T = corec_sugar_data; val empty = (Symtab.empty, Symtab.empty, Symtab.empty); fun merge data : T = (Symtab.merge (K true) (apply2 #1 data), Symtab.merge (K true) (apply2 #2 data), Symtab.join (K merge_friend_extras) (apply2 #3 data)); ); fun register_codatatype_extra fpT_name extra = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra)))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(1)} (Symtab.update (fpT_name, morph_codatatype_extra phi extra)))); fun codatatype_extra_of ctxt = Symtab.lookup (#1 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_codatatype_extra (Proof_Context.theory_of ctxt)); fun all_codatatype_extras_of ctxt = Symtab.dest (#1 (Data.get (Context.Proof ctxt))); fun register_coinduct_extra fpT_name extra = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra)))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(2)} (Symtab.update (fpT_name, morph_coinduct_extra phi extra)))); fun coinduct_extra_of ctxt = Symtab.lookup (#2 (Data.get (Context.Proof ctxt))) #> Option.map (transfer_coinduct_extra (Proof_Context.theory_of ctxt)); fun register_friend_extra fun_name eq_algrho algrho_eq = - Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => - Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra) - (fn {eq_algrhos, algrho_eqs} => - {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos, - algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs})))); + Local_Theory.declaration {syntax = false, pervasive = true, pos = Position.thread_data ()} + (fn phi => + Data.map (@{apply 3(3)} (Symtab.map_default (fun_name, empty_friend_extra) + (fn {eq_algrhos, algrho_eqs} => + {eq_algrhos = Morphism.thm phi eq_algrho :: eq_algrhos, + algrho_eqs = Morphism.thm phi algrho_eq :: algrho_eqs})))); fun all_friend_extras_of ctxt = Symtab.dest (#3 (Data.get (Context.Proof ctxt))); fun coinduct_extras_of_generic context = corec_infos_of_generic context #> map (#corecUU #> dest_Const #> fst #> Symtab.lookup (#2 (Data.get context)) #> the #> transfer_coinduct_extra (Context.theory_of context)); fun get_coinduct_uptos fpT_name context = coinduct_extras_of_generic context fpT_name |> map #coinduct; fun get_cong_all_intros fpT_name context = coinduct_extras_of_generic context fpT_name |> maps (#cong_intro_pairs #> map snd); fun get_cong_intros fpT_name name context = coinduct_extras_of_generic context fpT_name |> map_filter (#cong_intro_pairs #> (fn ps => AList.lookup (op =) ps name)); fun ctr_names_of_fp_name lthy fpT_name = fpT_name |> fp_sugar_of lthy |> the |> #fp_ctr_sugar |> #ctr_sugar |> #ctrs |> map (Long_Name.base_name o name_of_ctr); fun register_coinduct_dynamic_base fpT_name lthy = let val fp_b = Binding.name (Long_Name.base_name fpT_name) in lthy |> fold Local_Theory.add_thms_dynamic ((mk_fp_binding fp_b coinduct_uptoN, get_coinduct_uptos fpT_name) :: map (fn N => let val N = cong_N ^ N in (mk_fp_binding fp_b N, get_cong_intros fpT_name N) end) ([baseN, reflN, symN, transN] @ ctr_names_of_fp_name lthy fpT_name)) |> Local_Theory.add_thms_dynamic (mk_fp_binding fp_b cong_introsN, get_cong_all_intros fpT_name) end; fun register_coinduct_dynamic_friend fpT_name friend_name = let val fp_b = Binding.name (Long_Name.base_name fpT_name); val friend_base_name = cong_N ^ Long_Name.base_name friend_name; in Local_Theory.add_thms_dynamic (mk_fp_binding fp_b friend_base_name, get_cong_intros fpT_name friend_base_name) end; fun derive_case_dtor ctxt fpT_name = let val thy = Proof_Context.theory_of ctxt; val SOME ({fp_res_index, fp_res = {dtors = dtors0, dtor_ctors, ...}, absT_info = {rep = rep0, abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {casex, exhaust, case_thms, ...}, ...}, ...}) = fp_sugar_of ctxt fpT_name; val (f_Ts, Type (_, [fpT as Type (_, As), _])) = strip_fun_type (fastype_of casex); val x_Tss = map binder_types f_Ts; val (((u, fs), xss), _) = ctxt |> yield_singleton (mk_Frees "y") fpT ||>> mk_Frees "f" f_Ts ||>> mk_Freess "x" x_Tss; val dtor0 = nth dtors0 fp_res_index; val dtor = mk_dtor As dtor0; val u' = dtor $ u; val absT = fastype_of u'; val rep = mk_rep absT rep0; val goal = mk_Trueprop_eq (list_comb (casex, fs) $ u, mk_case_absumprod absT rep fs xss xss $ u') |> Raw_Simplifier.rewrite_term thy @{thms comp_def[THEN eq_reflection]} []; val vars = map (fst o dest_Free) (u :: fs); val dtor_ctor = nth dtor_ctors fp_res_index; in Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_case_dtor_tac ctxt u abs_inverse dtor_ctor ctr_defs exhaust case_thms) |> Thm.close_derivation \<^here> end; fun derive_case_trivial ctxt fpT_name = let val SOME {casex, exhaust, case_thms, ...} = ctr_sugar_of ctxt fpT_name; val Type (_, As0) = domain_type (body_fun_type (fastype_of casex)); val (As, _) = ctxt |> mk_TFrees' (map Type.sort_of_atyp As0); val fpT = Type (fpT_name, As); val (var_name, ()) = singleton (Variable.variant_frees ctxt []) ("x", ()); val var = Free (var_name, fpT); val goal = mk_Trueprop_eq (expand_to_ctr_term ctxt fpT var, var); val exhaust' = infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt var)] exhaust; in Goal.prove_sorry ctxt [var_name] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (rtac ctxt exhaust') THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt case_thms THEN ALLGOALS (rtac ctxt refl)) |> Thm.close_derivation \<^here> end; fun mk_abs_rep_transfers ctxt fpT_name = [mk_abs_transfer ctxt fpT_name, mk_rep_transfer ctxt fpT_name] handle Fail _ => []; fun ensure_codatatype_extra fpT_name ctxt = (case codatatype_extra_of ctxt fpT_name of NONE => let val abs_rep_transfers = mk_abs_rep_transfers ctxt fpT_name in ctxt |> register_codatatype_extra fpT_name {case_dtor = derive_case_dtor ctxt fpT_name, case_trivial = derive_case_trivial ctxt fpT_name, abs_rep_transfers = abs_rep_transfers} |> set_transfer_rule_attrs abs_rep_transfers end | SOME {abs_rep_transfers, ...} => ctxt |> set_transfer_rule_attrs abs_rep_transfers); fun setup_base fpT_name = register_coinduct_dynamic_base fpT_name #> ensure_codatatype_extra fpT_name; fun is_set ctxt (const_name, T) = (case T of Type (\<^type_name>\fun\, [Type (fpT_name, _), Type (\<^type_name>\set\, [_])]) => (case bnf_of ctxt fpT_name of SOME bnf => exists (fn Const (s, _) => s = const_name | _ => false) (sets_of_bnf bnf) | NONE => false) | _ => false); fun case_eq_if_thms_of_term ctxt t = let val ctr_sugars = map_filter (ctr_sugar_of_case ctxt o fst) (Term.add_consts t []) in maps #case_eq_ifs ctr_sugars end; fun all_algrho_eqs_of ctxt = maps (#algrho_eqs o snd) (all_friend_extras_of ctxt); fun derive_code ctxt inner_fp_simps goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_thm, ...} fun_t fun_def = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_code_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs corecUU_thm all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps inner_fp_simps fun_def)) |> Thm.close_derivation \<^here> end; fun derive_unique ctxt phi code_goal {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} fpT_name eq_corecUU = let val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val \<^Const_>\Trueprop for \<^Const_>\HOL.eq _ for lhs rhs\\ = code_goal; val (fun_t, args) = strip_comb lhs; val closed_rhs = fold_rev lambda args rhs; val fun_T = fastype_of fun_t; val num_args = num_binder_types fun_T; val f = Free (singleton (Variable.variant_frees ctxt []) ("f", fun_T)); val is_self_call = curry (op aconv) fun_t; val has_self_call = exists_subterm is_self_call; fun fify args (t $ u) = fify (u :: args) t $ fify [] u | fify _ (Abs (s, T, t)) = Abs (s, T, fify [] t) | fify args t = if t = fun_t andalso not (exists has_self_call args) then f else t; val goal = Logic.mk_implies (mk_Trueprop_eq (f, fify [] closed_rhs), mk_Trueprop_eq (f, fun_t)) |> Morphism.term phi; in Goal.prove_sorry ctxt [fst (dest_Free f)] [] goal (fn {context = ctxt, prems = _} => mk_unique_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique eq_corecUU) |> Thm.close_derivation \<^here> end; fun derive_last_disc ctxt fcT_name = let val SOME {T = fcT, discs, exhaust, disc_thmss, ...} = ctr_sugar_of ctxt fcT_name; val (u, _) = ctxt |> yield_singleton (mk_Frees "x") fcT; val udiscs = map (rapp u) discs; val (not_udiscs, last_udisc) = split_last udiscs |>> map HOLogic.mk_not; val goal = mk_Trueprop_eq (last_udisc, foldr1 HOLogic.mk_conj not_udiscs); in Goal.prove_sorry ctxt [fst (dest_Free u)] [] goal (fn {context = ctxt, prems = _} => mk_last_disc_tac ctxt u exhaust (flat disc_thmss)) |> Thm.close_derivation \<^here> end; fun derive_eq_algrho ctxt {sig_fp_sugars, ssig_fp_sugar, eval, eval_simps, all_algLam_algs, corecUU_unique, ...} ({algrho = algrho0, dtor_algrho, ...} : friend_info) fun_t k_T code_goal const_transfers rho_def eq_corecUU = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, Ts)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {fp_res_index, fp_res, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_dtor, ...} = codatatype_extra_of ctxt fpT_name; val fp_nesting_Ts = map T_of_bnf fp_nesting_bnfs; fun is_nullary_disc_def (\<^Const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _))) = true | is_nullary_disc_def (Const (\<^const_name>\Pure.eq\, _) $ _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) = true | is_nullary_disc_def _ = false; val dtor_ctor = nth (#dtor_ctors fp_res) fp_res_index; val ctor_iff_dtor = #ctor_iff_dtor fp_ctr_sugar; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val nullary_disc_defs = filter (is_nullary_disc_def o Thm.prop_of) (#disc_defs ctr_sugar); val disc_sel_eq_cases = #disc_eq_cases ctr_sugar @ #sel_defs ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt code_goal; val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; fun add_tnameT (Type (s, Ts)) = insert (op =) s #> fold add_tnameT Ts | add_tnameT _ = I; fun map_disc_sels'_of s = (case fp_sugar_of ctxt s of SOME {fp_bnf_sugar = {map_disc_iffs, map_selss, ...}, ...} => let val map_selss' = if length map_selss <= 1 then map_selss else map (map (unfold_thms ctxt (no_refl [derive_last_disc ctxt s]))) map_selss; in map_disc_iffs @ flat map_selss' end | NONE => []); fun mk_const_pointful_natural const_transfer = SOME (mk_pointful_natural_from_transfer ctxt const_transfer) handle UNNATURAL () => NONE; val const_pointful_natural_opts = map mk_const_pointful_natural const_transfers; val const_pointful_naturals = map_filter I const_pointful_natural_opts; val fp_nesting_k_T_names = fold add_tnameT (k_T :: fp_nesting_Ts) []; val fp_nesting_k_map_disc_sels' = maps map_disc_sels'_of fp_nesting_k_T_names; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val ctor = nth (#ctors fp_res) fp_res_index; val abs = #abs absT_info; val rep = #rep absT_info; val algrho = mk_ctr Ts algrho0; val goal = mk_Trueprop_eq (fun_t, abs_curried_balanced arg_Ts algrho); fun const_of_transfer thm = (case Thm.prop_of thm of \<^Const>\Trueprop\ $ (_ $ cst $ _) => cst); val eq_algrho = Goal.prove (*no sorry*) ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_algrho_tac ctxt fpsig_nesting_maps abs rep ctor ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident dtor_ctor ctor_iff_dtor ctr_defs nullary_disc_defs disc_sel_eq_cases case_dtor case_eq_ifs const_pointful_naturals fp_nesting_k_map_disc_sels' rho_def dtor_algrho corecUU_unique eq_corecUU all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps) |> Thm.close_derivation \<^here> handle e as ERROR _ => (case filter (is_none o snd) (const_transfers ~~ const_pointful_natural_opts) of [] => Exn.reraise e | thm_nones => error ("Failed to state naturality property for " ^ commas (map (Syntax.string_of_term ctxt o const_of_transfer o fst) thm_nones))); val algrho_eq = eq_algrho RS (mk_curry_uncurryN_balanced ctxt num_args RS iffD2) RS sym; in (eq_algrho, algrho_eq) end; fun prime_rho_transfer_goal ctxt fpT_name rho_def goal = let val thy = Proof_Context.theory_of ctxt; val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; val simps = rel_def_of_bnf pre_bnf :: rho_transfer_simps; val fold_rho = unfold_thms ctxt [rho_def RS @{thm symmetric}]; fun derive_unprimed rho_transfer' = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => unfold_thms_tac ctxt simps THEN HEADGOAL (rtac ctxt rho_transfer'))) |> Thm.close_derivation \<^here>; val goal' = Raw_Simplifier.rewrite_term thy simps [] goal; in if null abs_rep_transfers then (goal', derive_unprimed #> fold_rho) else (goal, fold_rho) end; fun derive_rho_transfer_folded ctxt fpT_name const_transfers rho_def goal = let val SOME {pre_bnf, ...} = fp_sugar_of ctxt fpT_name; val SOME {abs_rep_transfers, ...} = codatatype_extra_of ctxt fpT_name; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_rho_transfer_tac ctxt (null abs_rep_transfers) (rel_def_of_bnf pre_bnf) const_transfers)) |> unfold_thms ctxt [rho_def RS @{thm symmetric}] |> Thm.close_derivation \<^here> end; fun mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong alg = let val xy_Ts = binder_types (fastype_of alg); val ((xs, ys), _) = ctxt |> mk_Frees "x" xy_Ts ||>> mk_Frees "y" xy_Ts; fun mk_prem xy_T x y = build_rel [] ctxt [fpT] [] (fn (T, _) => if T = fpT then Rcong else HOLogic.eq_const T) (xy_T, xy_T) $ x $ y; val prems = @{map 3} mk_prem xy_Ts xs ys; val concl = Rcong $ list_comb (alg, xs) $ list_comb (alg, ys); in Logic.list_implies (map HOLogic.mk_Trueprop prems, HOLogic.mk_Trueprop concl) end; fun derive_cong_ctr_intros ctxt cong_ctor_intro = let val \<^Const_>\Pure.imp\ $ _ $ (\<^Const_>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ (ctor $ _))) = Thm.prop_of cong_ctor_intro; val fpT as Type (fpT_name, fp_argTs) = range_type (fastype_of ctor); val SOME {pre_bnf, absT_info = {abs_inverse, ...}, fp_ctr_sugar = {ctr_defs, ctr_sugar = {ctrs = ctrs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val ctrs = map (mk_ctr fp_argTs) ctrs0; val pre_rel_def = rel_def_of_bnf pre_bnf; fun prove ctr_def goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt ctr_def [pre_rel_def, abs_inverse] cong_ctor_intro)) |> Thm.close_derivation \<^here>; val goals = map (mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong) ctrs; in map2 prove ctr_defs goals end; fun derive_cong_friend_intro ctxt cong_algrho_intro = let val \<^Const_>\Pure.imp\ $ _ $ (\<^Const_>\Trueprop\ $ ((Rcong as _ $ _) $ _ $ ((algrho as Const (algrho_name, _)) $ _))) = Thm.prop_of cong_algrho_intro; val fpT as Type (_, fp_argTs) = range_type (fastype_of algrho); fun has_algrho (\<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ rhs)) = fst (dest_Const (head_of (strip_abs_body rhs))) = algrho_name; val eq_algrho :: _ = maps (filter (has_algrho o Thm.prop_of) o #eq_algrhos o snd) (all_friend_extras_of ctxt); val \<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ friend0 $ _) = Thm.prop_of eq_algrho; val friend = mk_ctr fp_argTs friend0; val goal = mk_cong_intro_ctr_or_friend_goal ctxt fpT Rcong friend; in Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, prems = _} => mk_cong_intro_ctr_or_friend_tac ctxt eq_algrho [] cong_algrho_intro)) |> Thm.close_derivation \<^here> end; fun derive_cong_intros lthy ctr_names friend_names ({cong_base, cong_refl, cong_sym, cong_trans, cong_alg_intros, ...} : dtor_coinduct_info) = let val cong_ctor_intro :: cong_algrho_intros = rev cong_alg_intros; val names = map (prefix cong_N) ([baseN, reflN, symN, transN] @ ctr_names @ friend_names); val thms = [cong_base, cong_refl, cong_sym, cong_trans] @ derive_cong_ctr_intros lthy cong_ctor_intro @ map (derive_cong_friend_intro lthy) cong_algrho_intros; in names ~~ thms end; fun derive_coinduct ctxt (fpT as Type (fpT_name, fpT_args)) dtor_coinduct = let val thy = Proof_Context.theory_of ctxt; val \<^Const_>\Pure.imp\ $ (\<^Const_>\Trueprop\ $ (_ $ Abs (_, _, _ $ Abs (_, _, \<^Const_>\implies\ $ _ $ (_ $ (cong0 $ _) $ _ $ _))))) $ _ = Thm.prop_of dtor_coinduct; val SOME {X as TVar ((X_s, _), _), fp_res = {dtor_ctors, ...}, pre_bnf, absT_info = {abs_inverse, ...}, live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss, ctr_defs, ctr_sugar = ctr_sugar0 as {T = Type (_, T0_args), ctrs = ctrs0, discs = discs0, ...}, ...}, ...} = fp_sugar_of ctxt fpT_name; val n = length ctrXs_Tss; val ms = map length ctrXs_Tss; val X' = TVar ((X_s, maxidx_of_typ fpT + 1), \<^sort>\type\); val As_rho = tvar_subst thy T0_args fpT_args; val substXAT = Term.typ_subst_TVars As_rho o Tsubst X X'; val substXA = Term.subst_TVars As_rho o substT X X'; val phi = Morphism.typ_morphism "BNF" substXAT $> Morphism.term_morphism "BNF" substXA; fun mk_applied_cong arg = enforce_type ctxt domain_type (fastype_of arg) cong0 $ arg; val thm = derive_coinduct_thms_for_types ctxt false mk_applied_cong [pre_bnf] dtor_coinduct dtor_ctors live_nesting_bnfs [fpT] [substXAT X] [map (map substXAT) ctrXs_Tss] [n] [abs_inverse] [abs_inverse] I [ctr_defs] [morph_ctr_sugar phi ctr_sugar0] |> map snd |> the_single; val (attrs, _) = mk_coinduct_attrs [fpT] [ctrs0] [discs0] [ms]; in (thm, attrs) end; type explore_parameters = {bound_Us: typ list, bound_Ts: typ list, U: typ, T: typ}; fun update_UT {bound_Us, bound_Ts, ...} U T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = U, T = T}; fun explore_nested lthy explore {bound_Us, bound_Ts, U, T} t = let fun build_simple (T, U) = if T = U then \<^term>\%y. y\ else Bound 0 |> explore {bound_Us = T :: bound_Us, bound_Ts = T :: bound_Ts, U = U, T = T} |> (fn t => Abs (Name.uu, T, t)); in betapply (build_map lthy [] [] build_simple (T, U), t) end; fun add_boundvar t = betapply (incr_boundvars 1 t, Bound 0); fun explore_fun (arg_U :: arg_Us) explore {bound_Us, bound_Ts, U, T} t = let val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t) in add_boundvar t |> explore_fun arg_Us explore {bound_Us = arg_U :: bound_Us, bound_Ts = domain_type T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_U, t)) end | explore_fun [] explore params t = explore params t; fun massage_fun explore (params as {T, U, ...}) = if can dest_funT T then explore_fun [domain_type U] explore params else explore params; fun massage_star massages explore = let fun after_massage massages' t params t' = if t aconv t' then massage_any massages' params t else massage_any massages params t' and massage_any [] params t = explore params t | massage_any (massage :: massages') params t = massage (after_massage massages' t) params t; in massage_any massages end; fun massage_let explore params t = (case strip_comb t of (Const (\<^const_name>\Let\, _), [_, _]) => unfold_lets_splits t | _ => t) |> explore params; fun check_corec_equation ctxt fun_frees (lhs, rhs) = let val (fun_t, arg_ts) = strip_comb lhs; fun check_fun_name () = null fun_frees orelse member (op aconv) fun_frees fun_t orelse ill_formed_equation_head ctxt [] fun_t; fun check_no_other_frees () = (case Term.add_frees rhs [] |> map Free |> subtract (op =) (fun_frees @ arg_ts) |> find_first (not o Variable.is_fixed ctxt o fst o dest_Free) of NONE => () | SOME t => extra_variable_in_rhs ctxt [] t); in check_duplicate_variables_in_lhs ctxt [] arg_ts; check_fun_name (); check_all_fun_arg_frees ctxt [] (filter_out is_Var arg_ts); check_no_other_frees () end; fun parse_corec_equation ctxt fun_frees eq = let val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (drop_all eq)) handle TERM _ => ill_formed_equation_lhs_rhs ctxt [eq]; val _ = check_corec_equation ctxt fun_frees (lhs, rhs); val (fun_t, arg_ts) = strip_comb lhs; val (arg_Ts, _) = strip_type (fastype_of fun_t); val added_Ts = drop (length arg_ts) arg_Ts; val free_names = mk_names (length added_Ts) "x" ~~ added_Ts; val free_args = Variable.variant_frees ctxt [rhs, lhs] free_names |> map Free; in (arg_ts @ free_args, list_comb (rhs, free_args)) end; fun morph_views phi (code, ctrs, discs, disc_iffs, sels) = (Morphism.term phi code, map (Morphism.term phi) ctrs, map (Morphism.term phi) discs, map (Morphism.term phi) disc_iffs, map (Morphism.term phi) sels); fun generate_views ctxt eq fun_t (lhs_free_args, rhs) = let val lhs = list_comb (fun_t, lhs_free_args); val T as Type (T_name, Ts) = fastype_of rhs; val SOME {fp_ctr_sugar = {ctr_sugar = {ctrs = ctrs0, discs = discs0, selss = selss0, ...}, ...}, ...} = fp_sugar_of ctxt T_name; val ctrs = map (mk_ctr Ts) ctrs0; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val code_view = drop_all eq; fun can_case_expand t = not (can (dest_ctr ctxt T_name) t); fun generate_raw_views conds t raw_views = let fun analyse (ctr :: ctrs) (disc :: discs) ctr' = if ctr = ctr' then (conds, disc, ctr) else analyse ctrs discs ctr'; in (analyse ctrs discs (fst (strip_comb t))) :: raw_views end; fun generate_disc_views raw_views = if length discs = 1 then ([], []) else let fun collect_condss_disc condss [] _ = condss | collect_condss_disc condss ((conds, disc', _) :: raw_views) disc = collect_condss_disc (condss |> disc = disc' ? cons conds) raw_views disc; val grouped_disc_views = discs |> map (collect_condss_disc [] raw_views) |> curry (op ~~) (map (fn disc => disc $ lhs) discs); fun mk_disc_iff_props props [] = props | mk_disc_iff_props _ ((lhs, \<^Const_>\True\) :: _) = [lhs] | mk_disc_iff_props props ((lhs, rhs) :: views) = mk_disc_iff_props ((HOLogic.mk_eq (lhs, rhs)) :: props) views; in (grouped_disc_views |> map swap, grouped_disc_views |> map (apsnd (s_dnf #> mk_conjs)) |> mk_disc_iff_props [] |> map (fn eq => ([[]], eq))) end; fun generate_ctr_views raw_views = let fun collect_condss_ctr condss [] _ = condss | collect_condss_ctr condss ((conds, _, ctr') :: raw_views) ctr = collect_condss_ctr (condss |> ctr = ctr' ? cons conds) raw_views ctr; fun mk_ctr_eq ctr_sels ctr = let fun extract_arg n sel _(*bound_Ts*) fun_t arg_ts = if ctr = fun_t then nth arg_ts n else let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern (range_type (fastype_of sel)) end; in ctr_sels |> map_index (uncurry extract_arg) |> map (fn extract => massage_corec_code_rhs ctxt extract [] rhs) |> curry list_comb ctr |> curry HOLogic.mk_eq lhs end; fun remove_condss_if_alone [(_, concl)] = [([[]], concl)] | remove_condss_if_alone views = views; in ctrs |> `(map (collect_condss_ctr [] raw_views)) ||> map2 mk_ctr_eq selss |> op ~~ |> filter_out (null o fst) |> remove_condss_if_alone end; fun generate_sel_views raw_views only_one_ctr = let fun mk_sel_positions sel = let fun get_sel_position _ [] = NONE | get_sel_position i (sel' :: sels) = if sel = sel' then SOME i else get_sel_position (i + 1) sels; in ctrs ~~ map (get_sel_position 0) selss |> map_filter (fn (ctr, pos_opt) => if is_some pos_opt then SOME (ctr, the pos_opt) else NONE) end; fun collect_sel_condss0 condss [] _ = condss | collect_sel_condss0 condss ((conds, _, ctr) :: raw_views) sel_positions = let val condss' = condss |> is_some (AList.lookup (op =) sel_positions ctr) ? cons conds in collect_sel_condss0 condss' raw_views sel_positions end; val collect_sel_condss = if only_one_ctr then K [[]] else collect_sel_condss0 [] raw_views; fun mk_sel_rhs sel_positions sel = let val sel_T = range_type (fastype_of sel); fun extract_sel_value _(*bound_Ts*) fun_t arg_ts = (case AList.lookup (op =) sel_positions fun_t of SOME n => nth arg_ts n | NONE => let val t = list_comb (fun_t, arg_ts) in if can_case_expand t then sel $ t else Term.dummy_pattern sel_T end); in massage_corec_code_rhs ctxt extract_sel_value [] rhs end; val ordered_sels = distinct (op =) (flat selss); val sel_positionss = map mk_sel_positions ordered_sels; val sel_rhss = map2 mk_sel_rhs sel_positionss ordered_sels; val sel_lhss = map (rapp lhs o mk_disc_or_sel Ts) ordered_sels; val sel_condss = map collect_sel_condss sel_positionss; fun is_undefined (Const (\<^const_name>\undefined\, _)) = true | is_undefined _ = false; in sel_condss ~~ (sel_lhss ~~ sel_rhss) |> filter_out (is_undefined o snd o snd) |> map (apsnd HOLogic.mk_eq) end; fun mk_atomic_prop fun_args (condss, concl) = (Logic.list_all (map dest_Free fun_args, abstract_over_list fun_args (Logic.list_implies (map HOLogic.mk_Trueprop (s_dnf condss), HOLogic.mk_Trueprop concl)))); val raw_views = rhs |> massage_let_if_case ctxt (K false) (fn _(*bound_Ts*) => fn t => t |> can_case_expand t ? expand_to_ctr_term ctxt T) (K (K ())) (K I) [] |> (fn expanded_rhs => fold_rev_let_if_case ctxt generate_raw_views [] expanded_rhs []) |> rev; val (disc_views, disc_iff_views) = generate_disc_views raw_views; val ctr_views = generate_ctr_views raw_views; val sel_views = generate_sel_views raw_views (length ctr_views = 1); val mk_props = filter_out (null o fst) #> map (mk_atomic_prop lhs_free_args); in (code_view, mk_props ctr_views, mk_props disc_views, mk_props disc_iff_views, mk_props sel_views) end; fun find_all_associated_types [] _ = [] | find_all_associated_types ((Type (_, Ts1), Type (_, Ts2)) :: TTs) T = find_all_associated_types ((Ts1 ~~ Ts2) @ TTs) T | find_all_associated_types ((T1, T2) :: TTs) T = find_all_associated_types TTs T |> T1 = T ? cons T2; fun as_member_of tab = try dest_Const #> Option.mapPartial (fst #> Symtab.lookup tab); fun extract_rho_from_equation ({ctr_guards, inner_buffer = {Oper, VLeaf, CLeaf, ctr_wrapper, friends}, ...}, {pattern_ctrs, discs, sels, it, mk_case}) b version Y preT ssig_T friend_tm (lhs_args, rhs) lthy = let val thy = Proof_Context.theory_of lthy; val res_T = fastype_of rhs; val YpreT = HOLogic.mk_prodT (Y, preT); fun fpT_to new_T T = if T = res_T then new_T else (case T of Type (s, Ts) => Type (s, map (fpT_to new_T) Ts) | _ => T); fun build_params bound_Us bound_Ts T = {bound_Us = bound_Us, bound_Ts = bound_Ts, U = T, T = T}; fun typ_before explore {bound_Us, bound_Ts, ...} t = explore (build_params bound_Us bound_Ts (fastype_of1 (bound_Ts, t))) t; val is_self_call = curry (op aconv) friend_tm; val has_self_call = Term.exists_subterm is_self_call; fun has_res_T bound_Ts t = fastype_of1 (bound_Ts, t) = res_T; fun contains_res_T (Type (s, Ts)) = s = fst (dest_Type res_T) orelse exists contains_res_T Ts | contains_res_T _ = false; val res_T_lhs_args = filter (exists_type contains_res_T) lhs_args; val is_res_T_lhs_arg = member (op =) res_T_lhs_args; fun is_constant t = not (Term.exists_subterm is_res_T_lhs_arg t orelse has_self_call t orelse loose_bvar (t, 0)); fun is_nested_type T = T <> res_T andalso T <> YpreT andalso T <> ssig_T; val is_valid_case_argumentT = not o member (op =) [Y, ssig_T]; fun is_same_type_constr (Type (s, _)) (Type (s', _)) = (s = s') | is_same_type_constr _ _ = false; exception NO_ENCAPSULATION of unit; val parametric_consts = Unsynchronized.ref []; (* We are assuming that set functions are marked with "[transfer_rule]" (cf. the "transfer" plugin). Otherwise, the "eq_algrho" tactic might fail. *) fun is_special_parametric_const (x as (s, _)) = s = \<^const_name>\id\ orelse is_set lthy x; fun add_parametric_const s general_T T U = let fun tupleT_of_funT T = let val (Ts, T) = strip_type T in mk_tupleT_balanced (Ts @ [T]) end; fun funT_of_tupleT n = dest_tupleT_balanced (n + 1) #> split_last #> op --->; val m = num_binder_types general_T; val param1_T = Type_Infer.paramify_vars general_T; val param2_T = Type_Infer.paramify_vars param1_T; val deadfixed_T = build_map lthy [] [] (mk_undefined o op -->) (apply2 tupleT_of_funT (param1_T, param2_T)) |> singleton (Type_Infer_Context.infer_types lthy) |> singleton (Type_Infer.fixate lthy false) |> type_of |> dest_funT |-> generalize_types 1 |> funT_of_tupleT m; val j = maxidx_of_typ deadfixed_T + 1; fun varifyT (Type (s, Ts)) = Type (s, map varifyT Ts) | varifyT (TFree (s, T)) = TVar ((s, j), T) | varifyT T = T; val dedvarified_T = varifyT deadfixed_T; val new_vars = Sign.typ_match thy (dedvarified_T, T) Vartab.empty |> Vartab.dest |> filter (curry (op =) j o snd o fst) |> Vartab.make; val deadinstantiated_T = map_atyps (Type.devar new_vars) dedvarified_T; val final_T = if Sign.typ_instance thy (U, deadinstantiated_T) then deadfixed_T else general_T; in parametric_consts := insert (op =) (s, final_T) (!parametric_consts) end; fun encapsulate (params as {U, T, ...}) t = if U = T then t else if T = Y then VLeaf $ t else if T = res_T then CLeaf $ t else if T = YpreT then it $ t else if is_nested_type T andalso is_same_type_constr T U then explore_nested lthy encapsulate params t else raise NO_ENCAPSULATION (); fun build_function_after_encapsulation fun_t fun_t' (params as {bound_Us, ...}) arg_ts arg_ts' = let val arg_Us' = fst (strip_typeN (length arg_ts) (fastype_of1 (bound_Us, fun_t'))); fun the_or_error arg NONE = error ("Illegal argument " ^ quote (Syntax.string_of_term lthy arg) ^ " to " ^ quote (Syntax.string_of_term lthy fun_t)) | the_or_error _ (SOME arg') = arg'; in arg_ts' |> `(map (curry fastype_of1 bound_Us)) |>> map2 (update_UT params) arg_Us' |> op ~~ |> map (try (uncurry encapsulate)) |> map2 the_or_error arg_ts |> curry list_comb fun_t' end; fun rebuild_function_after_exploration old_fn new_fn explore params arg_ts = arg_ts |> map (typ_before explore params) |> build_function_after_encapsulation old_fn new_fn params arg_ts; fun update_case Us U casex = let val Type (T_name, _) = domain_type (snd (strip_fun_type (fastype_of casex))); val SOME {fp_ctr_sugar = {ctr_sugar = {T = Type (_, Ts), casex, ...}, ...}, ...} = fp_sugar_of lthy T_name; val T = body_type (fastype_of casex); in Term.subst_atomic_types ((T :: Ts) ~~ (U :: Us)) casex end; fun deduce_according_type default_T [] = default_T | deduce_according_type default_T Ts = (case distinct (op =) Ts of U :: [] => U | _ => fpT_to ssig_T default_T); fun massage_if explore_cond explore (params as {bound_Us, bound_Ts, ...}) t = (case strip_comb t of (const as Const (\<^const_name>\If\, _), obj :: (branches as [_, _])) => (case List.partition Term.is_dummy_pattern (map (explore params) branches) of (dummy_branch' :: _, []) => dummy_branch' | (_, [branch']) => branch' | (_, branches') => let val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (fastype_of1 (bound_Ts, hd branches)) brancheUs; val const_obj' = (If_const U, obj) ||> explore_cond (update_UT params \<^typ>\bool\ \<^typ>\bool\) |> op $; in build_function_after_encapsulation (const $ obj) const_obj' params branches branches' end) | _ => explore params t); fun massage_map explore (params as {bound_Us, bound_Ts, T = Type (T_name, Ts), ...}) (t as func $ mapped_arg) = if is_self_call (head_of func) then explore params t else (case try (dest_map lthy T_name) func of SOME (map_tm, fs) => let val n = length fs; val mapped_arg' = mapped_arg |> `(curry fastype_of1 bound_Ts) |>> build_params bound_Us bound_Ts |-> explore; in (case fastype_of1 (bound_Us, mapped_arg') of Type (U_name, Us0) => if U_name = T_name then let val Us = map (fpT_to ssig_T) Us0; val temporary_map = map_tm |> mk_map n Us Ts; val map_fn_Ts = fastype_of #> strip_fun_type #> fst; val binder_Uss = map_fn_Ts temporary_map |> map (map (fpT_to ssig_T) o binder_types); val fun_paramss = map_fn_Ts (head_of func) |> map (build_params bound_Us bound_Ts); val fs' = fs |> @{map 4} explore_fun binder_Uss (replicate n explore) fun_paramss; val SOME bnf = bnf_of lthy T_name; val Type (_, bnf_Ts) = T_of_bnf bnf; val typ_alist = lives_of_bnf bnf ~~ map (curry fastype_of1 bound_Us #> range_type) fs'; val Us' = map2 the_default Us (map (AList.lookup (op =) typ_alist) bnf_Ts); val map_tm' = map_tm |> mk_map n Us Us'; in build_function_after_encapsulation func (list_comb (map_tm', fs')) params [mapped_arg] [mapped_arg'] end else explore params t | _ => explore params t) end | NONE => explore params t) | massage_map explore params t = explore params t; fun massage_comp explore (params as {bound_Us, ...}) t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => let val args' = map (typ_before explore params) args; val f2' = typ_before (explore_fun (map (curry fastype_of1 bound_Us) args') explore) params f2; val f1' = typ_before (explore_fun [range_type (fastype_of1 (bound_Us, f2'))] explore) params f1; in betapply (f1', list_comb (f2', args')) end | _ => explore params t); fun massage_ctr explore (params as {T = T as Type (s, Ts), bound_Us, ...}) t = if T <> res_T then (case try (dest_ctr lthy s) t of SOME (ctr, args) => let val args' = map (typ_before explore params) args; val SOME {T = Type (_, ctr_Ts), ...} = ctr_sugar_of lthy s; val temp_ctr = mk_ctr ctr_Ts ctr; val argUs = map (curry fastype_of1 bound_Us) args'; val typ_alist = binder_types (fastype_of temp_ctr) ~~ argUs; val Us = ctr_Ts |> map (find_all_associated_types typ_alist) |> map2 deduce_according_type Ts; val ctr' = mk_ctr Us ctr; in build_function_after_encapsulation ctr ctr' params args args' end | NONE => explore params t) else explore params t | massage_ctr explore params t = explore params t; fun const_of [] _ = NONE | const_of ((sel as Const (s1, _)) :: r) (const as Const (s2, _)) = if s1 = s2 then SOME sel else const_of r const | const_of _ _ = NONE; fun massage_disc explore (params as {T, bound_Us, bound_Ts, ...}) t = (case (strip_comb t, T = \<^typ>\bool\) of ((fun_t, arg :: []), true) => let val arg_T = fastype_of1 (bound_Ts, arg) in if arg_T <> res_T then (case arg_T |> try (fst o dest_Type) |> Option.mapPartial (ctr_sugar_of lthy) of SOME {discs, T = Type (_, Ts), ...} => (case const_of discs fun_t of SOME disc => let val arg' = arg |> typ_before explore params; val Type (_, Us) = fastype_of1 (bound_Us, arg'); val disc' = disc |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in disc' $ arg' end | NONE => explore params t) | NONE => explore params t) else explore params t end | _ => explore params t); fun massage_sel explore (params as {bound_Us, bound_Ts, ...}) t = let val (fun_t, args) = strip_comb t in if args = [] then explore params t else let val T = fastype_of1 (bound_Ts, hd args) in (case (Option.mapPartial (ctr_sugar_of lthy) (try (fst o dest_Type) T), T <> res_T) of (SOME {selss, T = Type (_, Ts), ...}, true) => (case const_of (flat selss) fun_t of SOME sel => let val args' = args |> map (typ_before explore params); val Type (_, Us) = fastype_of1 (bound_Us, hd args'); val sel' = sel |> Term.subst_TVars (map (fst o dest_TVar) Ts ~~ Us); in build_function_after_encapsulation sel sel' params args args' end | NONE => explore params t) | _ => explore params t) end end; fun massage_equality explore (params as {bound_Us, bound_Ts, ...}) (t as Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = let val check_is_VLeaf = not o (Term.exists_subterm (fn t => t aconv CLeaf orelse t aconv Oper)); fun try_pattern_matching (fun_t, arg_ts) t = (case as_member_of pattern_ctrs fun_t of SOME (disc, sels) => let val t' = typ_before explore params t in if fastype_of1 (bound_Us, t') = YpreT then let val arg_ts' = map (typ_before explore params) arg_ts; val sels_t' = map (fn sel => betapply (sel, t')) sels; val Ts = map (curry fastype_of1 bound_Us) arg_ts'; val Us = map (curry fastype_of1 bound_Us) sels_t'; val arg_ts' = map2 encapsulate (map2 (update_UT params) Us Ts) arg_ts'; in if forall check_is_VLeaf arg_ts' then SOME (Library.foldl1 HOLogic.mk_conj (betapply (disc, t') :: (map HOLogic.mk_eq (arg_ts' ~~ sels_t')))) else NONE end else NONE end | NONE => NONE); in (case try_pattern_matching (strip_comb t1) t2 of SOME cond => cond | NONE => (case try_pattern_matching (strip_comb t2) t1 of SOME cond => cond | NONE => let val T = fastype_of1 (bound_Ts, t1); val params' = build_params bound_Us bound_Ts T; val t1' = explore params' t1; val t2' = explore params' t2; in if fastype_of1 (bound_Us, t1') = T andalso fastype_of1 (bound_Us, t2') = T then HOLogic.mk_eq (t1', t2') else error ("Unsupported condition: " ^ quote (Syntax.string_of_term lthy t)) end)) end | massage_equality explore params t = explore params t; fun infer_types (TVar _) (TVar _) = [] | infer_types (U as TVar _) T = [(U, T)] | infer_types (Type (s', Us)) (Type (s, Ts)) = if s' = s then flat (map2 infer_types Us Ts) else [] | infer_types _ _ = []; fun group_by_fst associations [] = associations | group_by_fst associations ((a, b) :: r) = group_by_fst (add_association a b associations) r and add_association a b [] = [(a, [b])] | add_association a b ((c, d) :: r) = if a = c then (c, b :: d) :: r else (c, d) :: (add_association a b r); fun new_TVar known_TVars = Name.invent_list (map (fst o fst o dest_TVar) known_TVars) "x" 1 |> (fn [s] => TVar ((s, 0), [])); fun instantiate_type inferred_types = Term.typ_subst_TVars (map (apfst (fst o dest_TVar)) inferred_types); fun chose_unknown_TVar (T as TVar _) = SOME T | chose_unknown_TVar (Type (_, Ts)) = fold (curry merge_options) (map chose_unknown_TVar Ts) NONE | chose_unknown_TVar _ = NONE; (* The function under definition might not be defined yet when this is queried. *) fun maybe_const_type ctxt (s, T) = Sign.const_type (Proof_Context.theory_of ctxt) s |> the_default T; fun massage_const polymorphic explore (params as {bound_Us, ...}) t = let val (fun_t, arg_ts) = strip_comb t in (case fun_t of Const (fun_x as (s, fun_T)) => let val general_T = if polymorphic then maybe_const_type lthy fun_x else fun_T in if fun_t aconv friend_tm orelse contains_res_T (body_type general_T) orelse is_constant t then explore params t else let val inferred_types = infer_types general_T fun_T; fun prepare_skeleton [] _ = [] | prepare_skeleton ((T, U) :: inferred_types) As = let fun schematize_res_T U As = if U = res_T then let val A = new_TVar As in (A, A :: As) end else (case U of Type (s, Us) => fold_map schematize_res_T Us As |>> curry Type s | _ => (U, As)); val (U', As') = schematize_res_T U As; in (T, U') :: (prepare_skeleton inferred_types As') end; val inferred_types' = prepare_skeleton inferred_types (map fst inferred_types); val skeleton_T = instantiate_type inferred_types' general_T; fun explore_if_possible (exp_arg as (_, true)) _ = exp_arg | explore_if_possible (exp_arg as (arg, false)) arg_T = if exists (exists_subtype is_TVar) (binder_types arg_T) then exp_arg else (typ_before (explore_fun (binder_types arg_T) explore) params arg, true); fun collect_inferred_types [] _ = [] | collect_inferred_types ((arg, explored) :: exp_arg_ts) (arg_T :: arg_Ts) = (if explored then infer_types arg_T (fastype_of1 (bound_Us, arg)) else []) @ collect_inferred_types exp_arg_ts arg_Ts; fun propagate exp_arg_ts skeleton_T = let val arg_gen_Ts = binder_types skeleton_T; val exp_arg_ts = map2 explore_if_possible exp_arg_ts arg_gen_Ts; val inferred_types = collect_inferred_types exp_arg_ts arg_gen_Ts |> group_by_fst [] |> map (apsnd (deduce_according_type ssig_T)); in (exp_arg_ts, instantiate_type inferred_types skeleton_T) end; val remaining_to_be_explored = filter_out snd #> length; fun try_exploring_args exp_arg_ts skeleton_T = let val n = remaining_to_be_explored exp_arg_ts; val (exp_arg_ts', skeleton_T') = propagate exp_arg_ts skeleton_T; val n' = remaining_to_be_explored exp_arg_ts'; fun try_instantiating A T = try (try_exploring_args exp_arg_ts') (instantiate_type [(A, T)] skeleton_T'); in if n' = 0 then SOME (exp_arg_ts', skeleton_T') else if n = n' then if exists_subtype is_TVar skeleton_T' then let val SOME A = chose_unknown_TVar skeleton_T' in (case try_instantiating A ssig_T of SOME result => result | NONE => (case try_instantiating A YpreT of SOME result => result | NONE => (case try_instantiating A res_T of SOME result => result | NONE => NONE))) end else NONE else try_exploring_args exp_arg_ts' skeleton_T' end; in (case try_exploring_args (map (fn arg => (arg, false)) arg_ts) skeleton_T of SOME (exp_arg_ts, fun_U) => let val arg_ts' = map fst exp_arg_ts; val fun_t' = Const (s, fun_U); fun finish_off () = let val t' = build_function_after_encapsulation fun_t fun_t' params arg_ts arg_ts'; in if can type_of1 (bound_Us, t') then (if fun_T = fun_U orelse is_special_parametric_const (s, fun_T) then () else add_parametric_const s general_T fun_T fun_U; t') else explore params t end; in if polymorphic then finish_off () else (case try finish_off () of SOME t' => t' | NONE => explore params t) end | NONE => explore params t) end end | _ => explore params t) end; fun massage_rho explore = massage_star [massage_let, massage_if explore_cond, massage_case, massage_fun, massage_comp, massage_map, massage_ctr, massage_sel, massage_disc, massage_equality, massage_const false, massage_const true] explore and massage_case explore (params as {bound_Ts, bound_Us, ...}) t = (case strip_comb t of (casex as Const (case_x as (c, _)), args as _ :: _ :: _) => (case try strip_fun_type (maybe_const_type lthy case_x) of SOME (gen_branch_Ts, gen_body_fun_T) => let val gen_branch_ms = map num_binder_types gen_branch_Ts; val n = length gen_branch_ms; val (branches, obj_leftovers) = chop n args; in if n < length args then (case gen_body_fun_T of Type (_, [Type (T_name, _), _]) => if case_of lthy T_name = SOME (c, true) then let val brancheTs = binder_fun_types (fastype_of1 (bound_Ts, casex)); val obj_leftover_Ts = map (curry fastype_of1 bound_Ts) obj_leftovers; val obj_leftovers' = if is_constant (hd obj_leftovers) then obj_leftovers else (obj_leftover_Ts, obj_leftovers) |>> map (build_params bound_Us bound_Ts) |> op ~~ |> map (uncurry explore_inner); val obj_leftoverUs = obj_leftovers' |> map (curry fastype_of1 bound_Us); val _ = is_valid_case_argumentT (hd obj_leftoverUs) orelse error (quote (Syntax.string_of_term lthy (hd obj_leftovers)) ^ " is not a valid case argument"); val Us = obj_leftoverUs |> hd |> dest_Type |> snd; val branche_binderUss = (if hd obj_leftoverUs = YpreT then mk_case HOLogic.boolT else update_case Us HOLogic.boolT casex) |> fastype_of |> binder_fun_types |> map binder_types; val b_params = map (build_params bound_Us bound_Ts) brancheTs; val branches' = branches |> @{map 4} explore_fun branche_binderUss (replicate n explore) b_params; val brancheUs = map (curry fastype_of1 bound_Us) branches'; val U = deduce_according_type (body_type (hd brancheTs)) (map body_type brancheUs); val casex' = if hd obj_leftoverUs = YpreT then mk_case U else update_case Us U casex; in build_function_after_encapsulation casex casex' params (branches @ obj_leftovers) (branches' @ obj_leftovers') end else explore params t | _ => explore params t) else explore params t end | NONE => explore params t) | _ => explore params t) and explore_cond params t = if has_self_call t then unexpected_rec_call_in lthy [] t else explore_inner params t and explore_inner params t = massage_rho explore_inner_general params t and explore_inner_general (params as {bound_Us, bound_Ts, T, ...}) t = let val (fun_t, arg_ts) = strip_comb t in if is_constant t then t else (case (as_member_of discs fun_t, length arg_ts = 1 andalso has_res_T bound_Ts (the_single arg_ts)) of (SOME disc', true) => let val arg' = explore_inner params (the_single arg_ts); val arg_U = fastype_of1 (bound_Us, arg'); in if arg_U = res_T then fun_t $ arg' else if arg_U = YpreT then disc' $ arg' else error ("Discriminator " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | _ => (case as_member_of sels fun_t of SOME sel' => let val arg_ts' = map (explore_inner params) arg_ts; val arg_U = fastype_of1 (bound_Us, hd arg_ts'); in if arg_U = res_T then build_function_after_encapsulation fun_t fun_t params arg_ts arg_ts' else if arg_U = YpreT then build_function_after_encapsulation fun_t sel' params arg_ts arg_ts' else error ("Selector " ^ quote (Syntax.string_of_term lthy fun_t) ^ " cannot be applied to non-variable " ^ quote (Syntax.string_of_term lthy (hd arg_ts))) end | NONE => (case as_member_of friends fun_t of SOME (_, friend') => rebuild_function_after_exploration fun_t friend' explore_inner params arg_ts |> curry (op $) Oper | NONE => (case as_member_of ctr_guards fun_t of SOME ctr_guard' => rebuild_function_after_exploration fun_t ctr_guard' explore_inner params arg_ts |> curry (op $) ctr_wrapper |> curry (op $) Oper | NONE => if is_Bound fun_t then rebuild_function_after_exploration fun_t fun_t explore_inner params arg_ts else if is_Free fun_t then let val fun_t' = map_types (fpT_to YpreT) fun_t in rebuild_function_after_exploration fun_t fun_t' explore_inner params arg_ts end else if T = res_T then error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this and no friend") else error (quote (Syntax.string_of_term lthy fun_t) ^ " not polymorphic enough to be applied like this"))))) end; fun explore_ctr params t = massage_rho explore_ctr_general params t and explore_ctr_general params t = let val (fun_t, arg_ts) = strip_comb t; val ctr_opt = as_member_of ctr_guards fun_t; in if is_some ctr_opt then rebuild_function_after_exploration fun_t (the ctr_opt) explore_inner params arg_ts else not_constructor_in_rhs lthy [] fun_t end; val rho_rhs = rhs |> explore_ctr (build_params [] [] (fastype_of rhs)) |> abs_tuple_balanced (map (map_types (fpT_to YpreT)) lhs_args) |> unfold_id_bnf_etc lthy; in lthy |> define_const false b version rhoN rho_rhs |>> pair (!parametric_consts, rho_rhs) end; fun mk_rho_parametricity_goal ctxt Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs = let val YpreT = HOLogic.mk_prodT (Y, preT); val ZpreT = Tsubst Y Z YpreT; val ssigZ_T = Tsubst Y Z ssig_T; val dead_pre_rel' = Term.subst_atomic_types [(Y, ssig_T), (Z, ssigZ_T)] dead_pre_rel; val dead_k_rel' = Term.subst_atomic_types [(Y, YpreT), (Z, ZpreT)] dead_k_rel; val (R, _) = ctxt |> yield_singleton (mk_Frees "R") (mk_pred2T Y Z); val rho_rel = mk_rel_fun (dead_k_rel' $ mk_rel_prod R (dead_pre_rel $ R)) (dead_pre_rel' $ (dead_ssig_rel $ R)); val rho_rhsZ = substT Y Z rho_rhs; in HOLogic.mk_Trueprop (rho_rel $ rho_rhs $ rho_rhsZ) end; fun extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy = let fun mk_rel T bnf = let val ZT = Tsubst Y Z T; val rel_T = mk_predT [mk_pred2T Y Z, T, ZT]; in enforce_type lthy I rel_T (rel_of_bnf bnf) end; val ssig_bnf = #fp_bnf ssig_fp_sugar; val (dead_ssig_bnf, lthy') = bnf_kill_all_but 1 ssig_bnf lthy; val dead_pre_rel = mk_rel preT dead_pre_bnf; val dead_k_rel = mk_rel k_T dead_k_bnf; val dead_ssig_rel = mk_rel ssig_T dead_ssig_bnf; val (((parametric_consts, rho_rhs), rho_data), lthy'') = extract_rho_from_equation friend_parse_info fun_b version Y preT ssig_T fun_t parsed_eq lthy'; val const_transfer_goals = map (mk_const_transfer_goal lthy'') parametric_consts; val rho_transfer_goal = mk_rho_parametricity_goal lthy'' Y Z preT ssig_T dead_pre_rel dead_k_rel dead_ssig_rel rho_rhs; in ((rho_data, (const_transfer_goals, rho_transfer_goal)), lthy'') end; fun explore_corec_equation ctxt could_be_friend friend fun_name fun_free {outer_buffer, ctr_guards, inner_buffer} res_T (args, rhs) = let val is_self_call = curry (op aconv) fun_free; val has_self_call = Term.exists_subterm is_self_call; val outer_ssig_T = body_type (fastype_of (#Oper outer_buffer)); fun inner_fp_of (Free (s, _)) = Free (s ^ inner_fp_suffix, mk_tupleT_balanced (map fastype_of args) --> outer_ssig_T); fun build_params bound_Ts U T = {bound_Us = bound_Ts, bound_Ts = bound_Ts, U = U, T = T}; fun rebuild_function_after_exploration new_fn explore {bound_Ts, ...} arg_ts = let val binder_types_old_fn = map (curry fastype_of1 bound_Ts) arg_ts; val binder_types_new_fn = new_fn |> binder_types o (curry fastype_of1 bound_Ts) |> take (length binder_types_old_fn); val paramss = map2 (build_params bound_Ts) binder_types_new_fn binder_types_old_fn; in map2 explore paramss arg_ts |> curry list_comb new_fn end; fun massage_map_corec explore {bound_Ts, U, T, ...} t = let val explore' = explore ooo build_params in massage_nested_corec_call ctxt has_self_call explore' explore' bound_Ts U T t end; fun massage_comp explore params t = (case strip_comb t of (Const (\<^const_name>\comp\, _), f1 :: f2 :: args) => explore params (betapply (f1, (betapplys (f2, args)))) | _ => explore params t); fun massage_fun explore (params as {bound_Us, bound_Ts, U, T}) t = if can dest_funT T then let val arg_T = domain_type T; val arg_name = the_default Name.uu (try (fn (Abs (s, _, _)) => s) t); in add_boundvar t |> explore {bound_Us = arg_T :: bound_Us, bound_Ts = arg_T :: bound_Ts, U = range_type U, T = range_type T} |> (fn t => Abs (arg_name, arg_T, t)) end else explore params t fun massage_let_if_case_corec explore {bound_Ts, U, T, ...} t = massage_let_if_case ctxt has_self_call (fn bound_Ts => explore (build_params bound_Ts U T)) (K (unexpected_corec_call_in ctxt [t])) (K (unsupported_case_around_corec_call ctxt [t])) bound_Ts t; val massage_map_let_if_case = massage_star [massage_map_corec, massage_fun, massage_comp, massage_let_if_case_corec]; fun explore_arg _ t = if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a nested corecursive call" ^ (if could_be_friend then " (try specifying \"(friend)\")" else "")) else t; fun explore_inner params t = massage_map_let_if_case explore_inner_general params t and explore_inner_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in if has_self_call t then (case as_member_of (#friends inner_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer) | NONE => (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#ctr_wrapper inner_buffer) |> curry (op $) (#Oper inner_buffer) | NONE => if is_self_call f_t then if friend andalso exists has_self_call arg_ts then (case Symtab.lookup (#friends inner_buffer) fun_name of SOME (_, friend') => rebuild_function_after_exploration friend' explore_inner params arg_ts |> curry (op $) (#Oper inner_buffer)) else let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (#VLeaf inner_buffer) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) else #CLeaf inner_buffer $ t end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_inner_general params t; fun explore_outer params t = massage_map_let_if_case explore_outer_general params t and explore_outer_general (params as {bound_Ts, T, ...}) t = if T = res_T then let val (f_t, arg_ts) = strip_comb t in (case as_member_of ctr_guards f_t of SOME ctr_guard' => rebuild_function_after_exploration ctr_guard' explore_inner params arg_ts |> curry (op $) (#VLeaf outer_buffer) | NONE => if not (has_self_call t) then t |> expand_to_ctr_term ctxt T |> massage_let_if_case_corec explore_outer_general params else (case as_member_of (#friends outer_buffer) f_t of SOME (_, friend') => rebuild_function_after_exploration friend' explore_outer params arg_ts |> curry (op $) (#Oper outer_buffer) | NONE => if is_self_call f_t then let val arg_Ts = binder_types (fastype_of1 (bound_Ts, f_t)) in map2 explore_arg (map2 (update_UT params) arg_Ts arg_Ts) arg_ts |> mk_tuple1_balanced bound_Ts |> curry (op $) (inner_fp_of f_t) end else error (quote (Syntax.string_of_term ctxt f_t) ^ " not registered as friend"))) end else if has_self_call t then error (quote (Syntax.string_of_term ctxt t) ^ " contains a corecursive call but has type " ^ quote (Syntax.string_of_typ ctxt T)) else explore_nested ctxt explore_outer_general params t; in (args, rhs |> explore_outer (build_params [] outer_ssig_T res_T) |> abs_tuple_balanced args) end; fun mk_corec_fun_def_rhs ctxt arg_Ts corecUU0 corecUU_arg = let val corecUU = enforce_type ctxt domain_type (fastype_of corecUU_arg) corecUU0 in abs_curried_balanced arg_Ts (corecUU $ unfold_id_bnf_etc ctxt corecUU_arg) end; fun get_options ctxt opts = let val plugins = get_first (fn Plugins_Option f => SOME (f ctxt) | _ => NONE) (rev opts) |> the_default Plugin_Name.default_filter; val friend = exists (can (fn Friend_Option => ())) opts; val transfer = exists (can (fn Transfer_Option => ())) opts; in (plugins, friend, transfer) end; fun add_function binding parsed_eq lthy = let fun pat_completeness_auto ctxt = Pat_Completeness.pat_completeness_tac ctxt 1 THEN auto_tac ctxt; val ({defname, pelims = [[pelim]], pinducts = [pinduct], psimps = [psimp], ...}, lthy') = Function.add_function [(Binding.concealed binding, NONE, NoSyn)] [(((Binding.concealed Binding.empty, []), parsed_eq), [], [])] Function_Common.default_config pat_completeness_auto lthy; in ((defname, (pelim, pinduct, psimp)), lthy') end; fun build_corecUU_arg_and_goals prove_termin (Free (fun_base_name, _)) (arg_ts, explored_rhs) lthy = let val inner_fp_name0 = fun_base_name ^ inner_fp_suffix; val inner_fp_free = Free (inner_fp_name0, fastype_of explored_rhs); in if Term.exists_subterm (curry (op aconv) inner_fp_free) explored_rhs then let val arg = mk_tuple_balanced arg_ts; val inner_fp_eq = mk_Trueprop_eq (betapply (inner_fp_free, arg), betapply (explored_rhs, arg)); val ((inner_fp_name, (pelim, pinduct, psimp)), lthy') = add_function (Binding.name inner_fp_name0) inner_fp_eq lthy; fun mk_triple elim induct simp = ([elim], [induct], [simp]); fun prepare_termin () = let val {goal, ...} = Proof.goal (Function.termination NONE lthy'); val termin_goal = goal |> Thm.concl_of |> Logic.unprotect |> Envir.beta_eta_contract; in (lthy', (mk_triple pelim pinduct psimp, [termin_goal])) end; val (lthy'', (inner_fp_triple, termin_goals)) = if prove_termin then (case try (Function.prove_termination NONE (Function_Common.termination_prover_tac true lthy')) lthy' of NONE => prepare_termin () | SOME ({elims = SOME [[elim]], inducts = SOME [induct], simps = SOME [simp], ...}, lthy'') => (lthy'', (mk_triple elim induct simp, []))) else prepare_termin (); val inner_fp_const = (Binding.name_of inner_fp_name, fastype_of explored_rhs) |>> Proof_Context.read_const {proper = true, strict = false} lthy' |> (fn (Const (s, _), T) => Const (s, T)); in (((inner_fp_triple, termin_goals), inner_fp_const), lthy'') end else (((([], [], []), []), explored_rhs), lthy) end; fun derive_eq_corecUU ctxt {sig_fp_sugars, ssig_fp_sugar, eval, corecUU, eval_simps, all_algLam_algs, corecUU_unique, ...} fun_t corecUU_arg fun_code = let val fun_T = fastype_of fun_t; val (arg_Ts, Type (fpT_name, _)) = strip_type fun_T; val num_args = length arg_Ts; val SOME {pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, ...} = fp_sugar_of ctxt fpT_name; val SOME {case_trivial, ...} = codatatype_extra_of ctxt fpT_name; val ctr_sugar = #ctr_sugar fp_ctr_sugar; val pre_map_def = map_def_of_bnf pre_bnf; val abs_inverse = #abs_inverse absT_info; val ctr_defs = #ctr_defs fp_ctr_sugar; val case_eq_ifs = #case_eq_ifs ctr_sugar @ case_eq_if_thms_of_term ctxt (Thm.prop_of fun_code); val all_sig_map_thms = maps (#map_thms o #fp_bnf_sugar) sig_fp_sugars; val fp_map_ident = map_ident_of_bnf fp_bnf; val fpsig_nesting_bnfs = fp_nesting_bnfs @ maps #live_nesting_bnfs sig_fp_sugars; val fpsig_nesting_T_names = map (fst o dest_Type o T_of_bnf) fpsig_nesting_bnfs; val fpsig_nesting_fp_sugars = map_filter (fp_sugar_of ctxt) fpsig_nesting_T_names; val fpsig_nesting_fp_bnf_sugars = map #fp_bnf_sugar fpsig_nesting_fp_sugars; val ssig_fp_bnf_sugar = #fp_bnf_sugar ssig_fp_sugar; val ssig_bnf = #fp_bnf ssig_fp_sugar; val ssig_map = map_of_bnf ssig_bnf; val fpsig_nesting_maps = map map_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_ident0s = map map_ident0_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_comps = map map_comp_of_bnf fpsig_nesting_bnfs; val fpsig_nesting_map_thms = maps #map_thms fpsig_nesting_fp_bnf_sugars; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val ssig_map_thms = #map_thms ssig_fp_bnf_sugar; val all_algLam_alg_pointfuls = map (mk_pointful ctxt) all_algLam_algs; val def_rhs = mk_corec_fun_def_rhs ctxt arg_Ts corecUU corecUU_arg; val goal = mk_Trueprop_eq (fun_t, def_rhs); in Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => mk_eq_corecUU_tac ctxt num_args fpsig_nesting_maps ssig_map eval pre_map_def abs_inverse fpsig_nesting_map_ident0s fpsig_nesting_map_comps fpsig_nesting_map_thms live_nesting_map_ident0s fp_map_ident case_trivial ctr_defs case_eq_ifs all_sig_map_thms ssig_map_thms all_algLam_alg_pointfuls (all_algrho_eqs_of ctxt) eval_simps corecUU_unique fun_code) |> Thm.close_derivation \<^here> end; fun derive_coinduct_cong_intros ({fpT = fpT0 as Type (fpT_name, _), friend_names = friend_names0, corecUU = Const (corecUU_name, _), dtor_coinduct_info as {dtor_coinduct, ...}, ...}) lthy = let val thy = Proof_Context.theory_of lthy; val phi = Proof_Context.export_morphism lthy (Local_Theory.target_of lthy); val fpT = Morphism.typ phi fpT0; val general_fpT = body_type (Sign.the_const_type thy corecUU_name); val most_general = Sign.typ_instance thy (general_fpT, fpT); in (case (most_general, coinduct_extra_of lthy corecUU_name) of (true, SOME extra) => ((false, extra), lthy) | _ => let val ctr_names = ctr_names_of_fp_name lthy fpT_name; val friend_names = friend_names0 |> map Long_Name.base_name |> rev; val cong_intro_pairs = derive_cong_intros lthy ctr_names friend_names dtor_coinduct_info; val (coinduct, coinduct_attrs) = derive_coinduct lthy fpT0 dtor_coinduct; val ((_, [coinduct]), lthy) = (* TODO check: only if most_general?*) Local_Theory.note ((Binding.empty, coinduct_attrs), [coinduct]) lthy; val extra = {coinduct = coinduct, coinduct_attrs = coinduct_attrs, cong_intro_pairs = cong_intro_pairs}; in ((most_general, extra), lthy |> most_general ? register_coinduct_extra corecUU_name extra) end) end; fun update_coinduct_cong_intross_dynamic fpT_name lthy = let val all_corec_infos = corec_infos_of lthy fpT_name in lthy |> fold_map (apfst snd oo derive_coinduct_cong_intros) all_corec_infos |> snd end; fun derive_and_update_coinduct_cong_intross [] = pair (false, []) | derive_and_update_coinduct_cong_intross (corec_infos as {fpT = Type (fpT_name, _), ...} :: _) = fold_map derive_coinduct_cong_intros corec_infos #>> split_list #> (fn ((changeds, extras), lthy) => if exists I changeds then ((true, extras), lthy |> update_coinduct_cong_intross_dynamic fpT_name) else ((false, extras), lthy)); fun prepare_corec_ursive_cmd int long_cmd opts (raw_fixes, raw_eq) lthy = let val _ = can the_single raw_fixes orelse error "Mutually corecursive functions not supported"; val (plugins, friend, transfer) = get_options lthy opts; val ([((b, fun_T), mx)], [(_, eq)]) = fst (Specification.read_multi_specs raw_fixes [((Binding.empty_atts, raw_eq), [], [])] lthy); val _ = check_top_sort lthy b fun_T; val (arg_Ts, res_T) = strip_type fun_T; val fpT_name = (case res_T of Type (s, _) => s | _ => not_codatatype lthy res_T); val fun_free = Free (Binding.name_of b, fun_T); val parsed_eq = parse_corec_equation lthy [fun_free] eq; val fun_name = Local_Theory.full_name lthy b; val fun_t = Const (fun_name, fun_T); (* FIXME: does this work with locales that fix variables? *) val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; fun extract_rho lthy' = let val lthy'' = lthy' |> Variable.declare_typ fun_T; val (prepared as (_, _, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, _, ssig_fp_sugar, buffer), lthy''') = prepare_friend_corec fun_name fun_T lthy''; val friend_parse_info = friend_parse_info_of lthy''' arg_Ts res_T buffer; val parsed_eq' = parsed_eq ||> subst_atomic [(fun_free, fun_t)]; in lthy''' |> extract_rho_return_transfer_goals b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq' |>> pair prepared end; val ((prepareds, (rho_datas, transfer_goal_datas)), lthy2) = if friend then extract_rho lthy1 |>> (apfst single ##> (apfst single #> apsnd single)) else (([], ([], [])), lthy1); val ((buffer, corec_infos), lthy3) = if friend then ((#13 (the_single prepareds), []), lthy2) else corec_info_of res_T lthy2 ||> no_base ? update_coinduct_cong_intross_dynamic fpT_name |>> (fn info as {buffer, ...} => (buffer, [info])); val corec_parse_info = corec_parse_info_of lthy3 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy3 true friend fun_name fun_free corec_parse_info res_T parsed_eq; val (((inner_fp_triple, termin_goals), corecUU_arg), lthy4) = build_corecUU_arg_and_goals (not long_cmd) fun_free explored_eq lthy3; fun def_fun (inner_fp_elims0, inner_fp_inducts0, inner_fp_simps0) const_transfers rho_transfers_foldeds lthy5 = let fun register_friend lthy' = let val [(old_corec_info, fp_b, version, Y, Z, _, k_T, _, _, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, _)] = prepareds; val [(rho, rho_def)] = rho_datas; val [(_, rho_transfer_goal)] = transfer_goal_datas; val Type (fpT_name, _) = res_T; val rho_transfer_folded = (case rho_transfers_foldeds of [] => derive_rho_transfer_folded lthy' fpT_name const_transfers rho_def rho_transfer_goal | [thm] => thm); in lthy' |> register_coinduct_dynamic_friend fpT_name fun_name |> register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho rho_transfer_folded old_corec_info end; val (friend_infos, lthy6) = lthy5 |> (if friend then register_friend #>> single else pair []); val (corec_info as {corecUU = corecUU0, ...}, lthy7) = (case corec_infos of [] => corec_info_of res_T lthy6 | [info] => (info, lthy6)); val def_rhs = mk_corec_fun_def_rhs lthy7 arg_Ts corecUU0 corecUU_arg; val def = ((b, mx), ((Binding.concealed (Thm.def_binding b), []), def_rhs)); val ((fun_lhs0, (_, fun_def0)), (lthy9, lthy8')) = lthy7 |> (snd o Local_Theory.begin_nested) |> Local_Theory.define def |> tap (fn (def, lthy') => print_def_consts int [def] lthy') ||> `Local_Theory.end_nested; val parsed_eq = parse_corec_equation lthy9 [fun_free] eq; val views0 = generate_views lthy9 eq fun_free parsed_eq; val lthy9' = lthy9 |> fold Variable.declare_typ (res_T :: arg_Ts); val phi = Proof_Context.export_morphism lthy8' lthy9'; val fun_lhs = Morphism.term phi fun_lhs0; val fun_def = Morphism.thm phi fun_def0; val inner_fp_elims = map (Morphism.thm phi) inner_fp_elims0; val inner_fp_inducts = map (Morphism.thm phi) inner_fp_inducts0; val inner_fp_simps = map (Morphism.thm phi) inner_fp_simps0; val (code_goal, _, _, _, _) = morph_views phi views0; fun derive_and_note_friend_extra_theorems lthy' = let val k_T = #7 (the_single prepareds); val rho_def = snd (the_single rho_datas); val (eq_algrho, algrho_eq) = derive_eq_algrho lthy' corec_info (the_single friend_infos) fun_lhs k_T code_goal const_transfers rho_def fun_def; val notes = (if Config.get lthy' bnf_internals then [(eq_algrhoN, [eq_algrho])] else []) |> map (fn (thmN, thms) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false friendN (Binding.name thmN)), []), [(thms, [])])); in lthy' |> register_friend_extra fun_name eq_algrho algrho_eq |> Local_Theory.notes notes |> snd end; val lthy10 = lthy9 |> friend ? derive_and_note_friend_extra_theorems; val code_thm = derive_code lthy10 inner_fp_simps code_goal corec_info fun_lhs fun_def; (* TODO: val ctr_thmss = map mk_thm (#2 views); val disc_thmss = map mk_thm (#3 views); val disc_iff_thmss = map mk_thm (#4 views); val sel_thmss = map mk_thm (#5 views); *) val uniques = if null inner_fp_simps then [derive_unique lthy10 phi (#1 views0) corec_info fpT_name fun_def] else []; (* TODO: val disc_iff_or_disc_thmss = map2 (fn [] => I | disc_iffs => K disc_iffs) disc_iff_thmss disc_thmss; val simp_thmss = map2 append disc_iff_or_disc_thmss sel_thmss; *) val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy11) = lthy10 |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val anonymous_notes = []; (* TODO: [(flat disc_iff_or_disc_thmss, simp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); *) val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (codeN, [code_thm], nitpicksimp_attrs), (coinductN, [coinduct], coinduct_attrs), (inner_inductN, inner_fp_inducts, []), (uniqueN, uniques, [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy11 bnf_internals then [(inner_elimN, inner_fp_elims, []), (inner_simpN, inner_fp_simps, [])] else []) (* TODO: (ctrN, ctr_thms, []), (discN, disc_thms, []), (disc_iffN, disc_iff_thms, []), (selN, sel_thms, simp_attrs), (simpsN, simp_thms, []), *) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of b) (Binding.qualify false corecN (Binding.name thmN)), attrs), [(thms, [])])) |> filter_out (null o fst o hd o snd); in lthy11 (* TODO: |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat sel_thmss) |> Spec_Rules.add Spec_Rules.equational ([fun_lhs], flat ctr_thmss) *) |> Spec_Rules.add Binding.empty Spec_Rules.equational [fun_lhs] [code_thm] |> plugins code_plugin ? Code.declare_default_eqns [(code_thm, true)] |> Local_Theory.notes (anonymous_notes @ notes) |> snd end; fun prove_transfer_goal ctxt goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove (*no sorry*) ctxt vars [] goal (fn {context = ctxt, prems = _} => HEADGOAL (Transfer.transfer_prover_tac ctxt))) |> Thm.close_derivation \<^here>; fun maybe_prove_transfer_goal ctxt goal = (case try (prove_transfer_goal ctxt) goal of SOME thm => apfst (cons thm) | NONE => apsnd (cons goal)); val const_transfer_goals = fold (union (op aconv) o fst) transfer_goal_datas []; val (const_transfers, const_transfer_goals') = if long_cmd then ([], const_transfer_goals) else fold (maybe_prove_transfer_goal lthy4) const_transfer_goals ([], []); in ((def_fun, (([res_T], prepareds, rho_datas, map snd transfer_goal_datas), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals'))), lthy4) end; fun corec_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (_, (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int false opts (raw_fixes, raw_eq) lthy; in if not (null termin_goals) then error ("Termination prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else if not (null const_transfer_goals) then error ("Transfer prover failed (try " ^ quote (#1 \<^command_keyword>\corecursive\) ^ " instead of " ^ quote (#1 \<^command_keyword>\corec\) ^ ")") else def_fun inner_fp_triple const_transfers [] lthy' end; fun corecursive_cmd int opts (raw_fixes, raw_eq) lthy = let val ((def_fun, (([Type (fpT_name, _)], prepareds, rho_datas, rho_transfer_goals), (inner_fp_triple, termin_goals), (const_transfers, const_transfer_goals))), lthy') = prepare_corec_ursive_cmd int true opts (raw_fixes, raw_eq) lthy; val (rho_transfer_goals', unprime_rho_transfer_and_folds) = @{map 3} (fn (_, _, _, _, _, _, _, _, _, _, _, _, _) => fn (_, rho_def) => prime_rho_transfer_goal lthy' fpT_name rho_def) prepareds rho_datas rho_transfer_goals |> split_list; in Proof.theorem NONE (fn [termin_thms, const_transfers', rho_transfers'] => let val remove_domain_condition = full_simplify (put_simpset HOL_basic_ss lthy' addsimps (@{thm True_implies_equals} :: termin_thms)); in def_fun (@{apply 3} (map remove_domain_condition) inner_fp_triple) (const_transfers @ const_transfers') (map2 (fn f => f) unprime_rho_transfer_and_folds rho_transfers') end) (map (map (rpair [])) [termin_goals, const_transfer_goals, rho_transfer_goals']) lthy' end; fun friend_of_corec_cmd ((raw_fun_name, raw_fun_T_opt), raw_eq) lthy = let val Const (fun_name, _) = Proof_Context.read_const {proper = true, strict = false} lthy raw_fun_name; val fake_lthy = lthy |> (case raw_fun_T_opt of SOME raw_T => Proof_Context.add_const_constraint (fun_name, SOME (Syntax.read_typ lthy raw_T)) | NONE => I) handle TYPE (s, _, _) => error s; val fun_b = Binding.name (Long_Name.base_name fun_name); val code_goal = Syntax.read_prop fake_lthy raw_eq; val fun_T = (case code_goal of \<^Const_>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _) => fastype_of (head_of t) | _ => ill_formed_equation_lhs_rhs lthy [code_goal]); val fun_t = Const (fun_name, fun_T); val (arg_Ts, res_T as Type (fpT_name, _)) = strip_type fun_T; val no_base = has_no_corec_info lthy fpT_name; val lthy1 = lthy |> no_base ? setup_base fpT_name; val lthy2 = lthy1 |> Variable.declare_typ fun_T; val ((old_corec_info, fp_b, version, Y, Z, preT, k_T, ssig_T, dead_pre_bnf, dead_k_bnf, sig_fp_sugar, ssig_fp_sugar, buffer), lthy3) = prepare_friend_corec fun_name fun_T lthy2; val friend_parse_info = friend_parse_info_of lthy3 arg_Ts res_T buffer; val parsed_eq = parse_corec_equation lthy3 [] code_goal; val (((rho, rho_def), (const_transfer_goals, rho_transfer_goal)), lthy4) = extract_rho_return_transfer_goals fun_b version dead_pre_bnf dead_k_bnf Y Z preT k_T ssig_T ssig_fp_sugar friend_parse_info fun_t parsed_eq lthy3; fun register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T friend_info lthy5 = let val (corec_info, lthy6) = corec_info_of res_T lthy5; val fun_free = Free (Binding.name_of fun_b, fun_T); fun freeze_fun (t as Const (s, T)) = if s = fun_name andalso T = fun_T then fun_free else t | freeze_fun t = t; val eq = Term.map_aterms freeze_fun code_goal; val parsed_eq = parse_corec_equation lthy6 [fun_free] eq; val corec_parse_info = corec_parse_info_of lthy6 arg_Ts res_T buffer; val explored_eq = explore_corec_equation lthy6 false false fun_name fun_free corec_parse_info res_T parsed_eq; val ((_, corecUU_arg), _) = build_corecUU_arg_and_goals false fun_free explored_eq lthy6; val eq_corecUU = derive_eq_corecUU lthy6 corec_info fun_t corecUU_arg code_thm; val (eq_algrho, algrho_eq) = derive_eq_algrho lthy6 corec_info friend_info fun_t k_T code_goal const_transfers rho_def eq_corecUU; val ((_, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy7) = lthy6 |> register_friend_extra fun_name eq_algrho algrho_eq |> register_coinduct_dynamic_friend fpT_name fun_name |> derive_and_update_coinduct_cong_intross [corec_info]; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val unique = derive_unique lthy7 Morphism.identity code_goal corec_info fpT_name eq_corecUU; val notes = [(codeN, [code_thm], []), (coinductN, [coinduct], coinduct_attrs), (cong_introsN, maps snd cong_intros_pairs, []), (uniqueN, [unique], [])] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs @ (if Config.get lthy7 bnf_internals then [(eq_algrhoN, [eq_algrho], []), (eq_corecUUN, [eq_corecUU], [])] else []) |> map (fn (thmN, thms, attrs) => ((Binding.qualify true (Binding.name_of fun_b) (Binding.qualify false friendN (Binding.name thmN)), attrs), [(thms, [])])); in lthy7 |> Local_Theory.notes notes |> snd end; val (rho_transfer_goal', unprime_rho_transfer_and_fold) = prime_rho_transfer_goal lthy4 fpT_name rho_def rho_transfer_goal; in lthy4 |> Proof.theorem NONE (fn [[code_thm], const_transfers, [rho_transfer']] => register_friend_corec fun_name fp_b version Y Z k_T dead_k_bnf sig_fp_sugar ssig_fp_sugar fun_t rho (unprime_rho_transfer_and_fold rho_transfer') old_corec_info #-> register_friend_extra_and_note_thms code_goal code_thm const_transfers k_T) (map (map (rpair [])) [[code_goal], const_transfer_goals, [rho_transfer_goal']]) |> Proof.refine_singleton (Method.primitive_text (K I)) end; fun coinduction_upto_cmd (base_name, raw_fpT) lthy = let val fpT as Type (fpT_name, _) = Syntax.read_typ lthy raw_fpT; val no_base = has_no_corec_info lthy fpT_name; val (corec_info as {version, ...}, lthy1) = lthy |> corec_info_of fpT; val lthy2 = lthy1 |> no_base ? setup_base fpT_name; val ((changed, [{cong_intro_pairs, coinduct, coinduct_attrs}]), lthy3) = lthy2 |> derive_and_update_coinduct_cong_intross [corec_info]; val lthy4 = lthy3 |> (changed orelse no_base) ? update_coinduct_cong_intross_dynamic fpT_name; val cong_intros_pairs = AList.group (op =) cong_intro_pairs; val notes = [(cong_introsN, maps snd cong_intros_pairs, []), (coinduct_uptoN, [coinduct], coinduct_attrs)] @ map (fn (thmN, thms) => (thmN, thms, [])) cong_intros_pairs |> map (fn (thmN, thms, attrs) => (((Binding.qualify true base_name (Binding.qualify false ("v" ^ string_of_int version) (Binding.name thmN))), attrs), [(thms, [])])); in lthy4 |> Local_Theory.notes notes |> snd end; fun consolidate lthy = let val corec_infoss = map (corec_infos_of lthy o fst) (all_codatatype_extras_of lthy); val (changeds, lthy') = lthy |> fold_map (apfst fst oo derive_and_update_coinduct_cong_intross) corec_infoss; in if exists I changeds then lthy' else raise Same.SAME end; fun consolidate_global thy = SOME (Named_Target.theory_map consolidate thy) handle Same.SAME => NONE; val _ = Outer_Syntax.local_theory \<^command_keyword>\corec\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corec_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\corecursive\ "define nonprimitive corecursive functions" ((Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Parse.list1 corec_option_parser) --| \<^keyword>\)\) []) -- (Parse.vars --| Parse.where_ -- Parse.prop) >> uncurry (corecursive_cmd true)); val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\friend_of_corec\ "register a function as a legal context for nonprimitive corecursion" (Parse.const -- Scan.option (\<^keyword>\::\ |-- Parse.typ) --| Parse.where_ -- Parse.prop >> friend_of_corec_cmd); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinduction_upto\ "derive a coinduction up-to principle and a corresponding congruence closure" (Parse.name --| \<^keyword>\:\ -- Parse.typ >> coinduction_upto_cmd); val _ = Theory.setup (Theory.at_begin consolidate_global); end; diff --git a/src/HOL/Tools/BNF/bnf_lfp_size.ML b/src/HOL/Tools/BNF/bnf_lfp_size.ML --- a/src/HOL/Tools/BNF/bnf_lfp_size.ML +++ b/src/HOL/Tools/BNF/bnf_lfp_size.ML @@ -1,399 +1,399 @@ (* Title: HOL/Tools/BNF/bnf_lfp_size.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2014 Generation of size functions for datatypes. *) signature BNF_LFP_SIZE = sig val register_size: string -> string -> thm -> thm list -> thm list -> local_theory -> local_theory val register_size_global: string -> string -> thm -> thm list -> thm list -> theory -> theory val size_of: Proof.context -> string -> (string * (thm * thm list * thm list)) option val size_of_global: theory -> string -> (string * (thm * thm list * thm list)) option end; structure BNF_LFP_Size : BNF_LFP_SIZE = struct open BNF_Util open BNF_Tactics open BNF_Def open BNF_FP_Def_Sugar val size_N = "size_"; val sizeN = "size"; val size_genN = "size_gen"; val size_gen_o_mapN = "size_gen_o_map"; val size_neqN = "size_neq"; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun mk_plus_nat (t1, t2) = Const (\<^const_name>\Groups.plus\, HOLogic.natT --> HOLogic.natT --> HOLogic.natT) $ t1 $ t2; fun mk_to_natT T = T --> HOLogic.natT; fun mk_abs_zero_nat T = Term.absdummy T HOLogic.zero; fun mk_unabs_def_unused_0 n = funpow n (fn thm => thm RS @{thm fun_cong_unused_0} handle THM _ => thm RS fun_cong); structure Data = Generic_Data ( type T = (string * (thm * thm list * thm list)) Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ); fun check_size_type thy T_name size_name = let val n = Sign.arity_number thy T_name; val As = map (fn s => TFree (s, \<^sort>\type\)) (Name.invent Name.context Name.aT n); val T = Type (T_name, As); val size_T = map mk_to_natT As ---> mk_to_natT T; val size_const = Const (size_name, size_T); in can (Thm.global_cterm_of thy) size_const orelse error ("Constant " ^ quote size_name ^ " registered as size function for " ^ quote T_name ^ " must have type\n" ^ quote (Syntax.string_of_typ_global thy size_T)) end; fun register_size T_name size_name overloaded_size_def size_simps size_gen_o_maps lthy = (check_size_type (Proof_Context.theory_of lthy) T_name size_name; Context.proof_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) lthy); fun register_size_global T_name size_name overloaded_size_def size_simps size_gen_o_maps thy = (check_size_type thy T_name size_name; Context.theory_map (Data.map (Symtab.update (T_name, (size_name, (overloaded_size_def, size_simps, size_gen_o_maps))))) thy); val size_of = Symtab.lookup o Data.get o Context.Proof; val size_of_global = Symtab.lookup o Data.get o Context.Theory; fun all_overloaded_size_defs_of ctxt = Symtab.fold (fn (_, (_, (overloaded_size_def, _, _))) => can (Logic.dest_equals o Thm.prop_of) overloaded_size_def ? cons overloaded_size_def) (Data.get (Context.Proof ctxt)) []; val size_gen_o_map_simps = @{thms inj_on_id snd_comp_apfst[simplified apfst_def]}; fun mk_size_gen_o_map_tac ctxt size_def rec_o_map inj_maps size_maps = unfold_thms_tac ctxt [size_def] THEN HEADGOAL (rtac ctxt (rec_o_map RS trans) THEN' asm_simp_tac (ss_only (inj_maps @ size_maps @ size_gen_o_map_simps) ctxt)) THEN IF_UNSOLVED (unfold_thms_tac ctxt @{thms id_def o_def} THEN HEADGOAL (rtac ctxt refl)); fun mk_size_neq ctxt cts exhaust sizes = HEADGOAL (rtac ctxt (infer_instantiate' ctxt (map SOME cts) exhaust)) THEN ALLGOALS (hyp_subst_tac ctxt) THEN unfold_thms_tac ctxt (@{thm neq0_conv} :: sizes) THEN ALLGOALS (REPEAT_DETERM o (rtac ctxt @{thm zero_less_Suc} ORELSE' rtac ctxt @{thm trans_less_add2})); fun generate_datatype_size (fp_sugars as ({T = Type (_, As), BT = Type (_, Bs), fp = Least_FP, fp_res = {bnfs = fp_bnfs, ...}, fp_nesting_bnfs, live_nesting_bnfs, fp_co_induct_sugar = SOME _, ...} : fp_sugar) :: _) lthy0 = let val data = Data.get (Context.Proof lthy0); val Ts = map #T fp_sugars val T_names = map (fst o dest_Type) Ts; val nn = length Ts; val B_ify = Term.typ_subst_atomic (As ~~ Bs); val recs = map (#co_rec o the o #fp_co_induct_sugar) fp_sugars; val rec_thmss = map (#co_rec_thms o the o #fp_co_induct_sugar) fp_sugars; val rec_Ts as rec_T1 :: _ = map fastype_of recs; val rec_arg_Ts = binder_fun_types rec_T1; val Cs = map body_type rec_Ts; val Cs_rho = map (rpair HOLogic.natT) Cs; val substCnatT = Term.subst_atomic_types Cs_rho; val f_Ts = map mk_to_natT As; val f_TsB = map mk_to_natT Bs; val num_As = length As; fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names f_Ts; val fsB = map2 (curry Free) f_names f_TsB; val As_fs = As ~~ fs; val size_bs = map ((fn base => Binding.qualify false base (Binding.name (prefix size_N base))) o Long_Name.base_name) T_names; fun is_prod_C \<^type_name>\prod\ [_, T'] = member (op =) Cs T' | is_prod_C _ _ = false; fun mk_size_of_typ (T as TFree _) = pair (case AList.lookup (op =) As_fs T of SOME f => f | NONE => if member (op =) Cs T then Term.absdummy T (Bound 0) else mk_abs_zero_nat T) | mk_size_of_typ (T as Type (s, Ts)) = if is_prod_C s Ts then pair (snd_const T) else if exists (exists_subtype_in (As @ Cs)) Ts then (case Symtab.lookup data s of SOME (size_name, (_, _, size_gen_o_maps)) => let val (args, size_gen_o_mapss') = fold_map mk_size_of_typ Ts []; val size_T = map fastype_of args ---> mk_to_natT T; val size_const = Const (size_name, size_T); in append (size_gen_o_maps :: size_gen_o_mapss') #> pair (Term.list_comb (size_const, args)) end | _ => pair (mk_abs_zero_nat T)) else pair (mk_abs_zero_nat T); fun mk_size_of_arg t = mk_size_of_typ (fastype_of t) #>> (fn s => substCnatT (betapply (s, t))); fun is_recursive_or_plain_case Ts = exists (exists_subtype_in Cs) Ts orelse forall (not o exists_subtype_in As) Ts; (* We want the size function to enjoy the following properties: 1. The size of a list should coincide with its length. 2. All the nonrecursive constructors of a type should have the same size. 3. Each constructor through which nested recursion takes place should count as at least one in the generic size function. 4. The "size" function should be definable as "size_t (%_. 0) ... (%_. 0)", where "size_t" is the generic function. This explains the somewhat convoluted logic ahead. *) val base_case = if forall (is_recursive_or_plain_case o binder_types) rec_arg_Ts then HOLogic.zero else HOLogic.Suc_zero; fun mk_size_arg rec_arg_T = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val (summands, size_gen_o_mapss) = fold_map mk_size_of_arg xs [] |>> remove (op =) HOLogic.zero; val sum = if null summands then base_case else foldl1 mk_plus_nat (summands @ [HOLogic.Suc_zero]); in append size_gen_o_mapss #> pair (fold_rev Term.lambda (map substCnatT xs) sum) end; fun mk_size_rhs recx = fold_map mk_size_arg rec_arg_Ts #>> (fn args => fold_rev Term.lambda fs (Term.list_comb (substCnatT recx, args))); val maybe_conceal_def_binding = Thm.def_binding #> not (Config.get lthy0 bnf_internals) ? Binding.concealed; val (size_rhss, nested_size_gen_o_mapss) = fold_map mk_size_rhs recs []; val size_Ts = map fastype_of size_rhss; val nested_size_gen_o_maps_complete = forall (not o null) nested_size_gen_o_mapss; val nested_size_gen_o_maps = fold (union Thm.eq_thm_prop) nested_size_gen_o_mapss []; val ((raw_size_consts, raw_size_defs), (lthy1, lthy1_old)) = lthy0 |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 2} (fn b => fn rhs => Local_Theory.define ((b, NoSyn), ((maybe_conceal_def_binding b, []), rhs)) #>> apsnd snd) size_bs size_rhss ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy1_old lthy1; val size_defs = map (Morphism.thm phi) raw_size_defs; val size_consts0 = map (Morphism.term phi) raw_size_consts; val size_consts = map2 retype_const_or_free size_Ts size_consts0; val size_constsB = map (Term.map_types B_ify) size_consts; val zeros = map mk_abs_zero_nat As; val overloaded_size_rhss = map (fn c => Term.list_comb (c, zeros)) size_consts; val overloaded_size_Ts = map fastype_of overloaded_size_rhss; val overloaded_size_consts = map (curry Const \<^const_name>\size\) overloaded_size_Ts; val overloaded_size_def_bs = map (maybe_conceal_def_binding o Binding.suffix_name "_overloaded") size_bs; fun define_overloaded_size def_b lhs0 rhs lthy = let val Free (c, _) = Syntax.check_term lthy lhs0; val ((_, (_, thm)), lthy') = lthy |> Local_Theory.define ((Binding.name c, NoSyn), ((def_b, []), rhs)); val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy'); val thm' = singleton (Proof_Context.export lthy' thy_ctxt) thm; in (thm', lthy') end; val (overloaded_size_defs, lthy2) = lthy1 |> Local_Theory.background_theory_result (Class.instantiation (T_names, map dest_TFree As, [HOLogic.class_size]) #> @{fold_map 3} define_overloaded_size overloaded_size_def_bs overloaded_size_consts overloaded_size_rhss ##> Class.prove_instantiation_instance (fn ctxt => Class.intro_classes_tac ctxt []) ##> Local_Theory.exit_global); val size_defs' = map (mk_unabs_def (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val size_defs_unused_0 = map (mk_unabs_def_unused_0 (num_As + 1) o HOLogic.mk_obj_eq) size_defs; val overloaded_size_defs' = map (mk_unabs_def 1 o HOLogic.mk_obj_eq) overloaded_size_defs; val nested_size_maps = map (mk_pointful lthy2) nested_size_gen_o_maps @ nested_size_gen_o_maps; val all_inj_maps = @{thm prod.inj_map} :: map inj_map_of_bnf (fp_bnfs @ fp_nesting_bnfs @ live_nesting_bnfs) |> distinct Thm.eq_thm_prop; fun derive_size_simp size_def' simp0 = (trans OF [size_def', simp0]) |> Simplifier.asm_full_simplify (ss_only (@{thms inj_on_convol_ident id_def o_def snd_conv} @ all_inj_maps @ nested_size_maps) lthy2) |> Local_Defs.fold lthy2 size_defs_unused_0; fun derive_overloaded_size_simp overloaded_size_def' simp0 = (trans OF [overloaded_size_def', simp0]) |> unfold_thms lthy2 @{thms add_0_left add_0_right} |> Local_Defs.fold lthy2 (overloaded_size_defs @ all_overloaded_size_defs_of lthy2); val size_simpss = map2 (map o derive_size_simp) size_defs' rec_thmss; val size_simps = flat size_simpss; val overloaded_size_simpss = map2 (map o derive_overloaded_size_simp) overloaded_size_defs' size_simpss; val overloaded_size_simps = flat overloaded_size_simpss; val size_thmss = map2 append size_simpss overloaded_size_simpss; val size_gen_thmss = size_simpss; fun rhs_is_zero thm = let val Const (trueprop, _) $ (Const (eq, _) $ _ $ rhs) = Thm.prop_of thm in trueprop = \<^const_name>\Trueprop\ andalso eq = \<^const_name>\HOL.eq\ andalso rhs = HOLogic.zero end; val size_neq_thmss = @{map 3} (fn fp_sugar => fn size => fn size_thms => if exists rhs_is_zero size_thms then [] else let val (xs, _) = mk_Frees "x" (binder_types (fastype_of size)) lthy2; val goal = HOLogic.mk_Trueprop (BNF_LFP_Util.mk_not_eq (list_comb (size, xs)) HOLogic.zero); val vars = Variable.add_free_names lthy2 goal []; val thm = Goal.prove_sorry lthy2 vars [] goal (fn {context = ctxt, ...} => mk_size_neq ctxt (map (Thm.cterm_of lthy2) xs) (#exhaust (#ctr_sugar (#fp_ctr_sugar fp_sugar))) size_thms) |> single |> map (Thm.close_derivation \<^here>); in thm end) fp_sugars overloaded_size_consts overloaded_size_simpss; val ABs = As ~~ Bs; val g_names = variant_names num_As "g"; val gs = map2 (curry Free) g_names (map (op -->) ABs); val liveness = map (op <>) ABs; val live_gs = AList.find (op =) (gs ~~ liveness) true; val live = length live_gs; val maps0 = map map_of_bnf fp_bnfs; val size_gen_o_map_thmss = if live = 0 then replicate nn [] else let val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val size_gen_o_map_conds = if exists (can Logic.dest_implies o Thm.prop_of) nested_size_gen_o_maps then map (HOLogic.mk_Trueprop o mk_inj) live_gs else []; val fsizes = map (fn size_constB => Term.list_comb (size_constB, fsB)) size_constsB; val size_gen_o_map_lhss = map2 (curry HOLogic.mk_comp) fsizes gmaps; val fgs = map2 (fn fB => fn g as Free (_, Type (_, [A, B])) => if A = B then fB else HOLogic.mk_comp (fB, g)) fsB gs; val size_gen_o_map_rhss = map (fn c => Term.list_comb (c, fgs)) size_consts; val size_gen_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fsB, gs] o curry Logic.list_implies size_gen_o_map_conds o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) size_gen_o_map_lhss size_gen_o_map_rhss; val rec_o_maps = fold_rev (curry (op @) o #co_rec_o_maps o the o #fp_co_induct_sugar) fp_sugars []; val size_gen_o_map_thmss = if nested_size_gen_o_maps_complete andalso forall (fn TFree (_, S) => S = \<^sort>\type\) As then @{map 3} (fn goal => fn size_def => fn rec_o_map => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_size_gen_o_map_tac ctxt size_def rec_o_map all_inj_maps nested_size_maps) |> Thm.close_derivation \<^here> |> single) size_gen_o_map_goals size_defs rec_o_maps else replicate nn []; in size_gen_o_map_thmss end; val massage_multi_notes = maps (fn (thmN, thmss, attrs) => map2 (fn T_name => fn thms => ((Binding.qualify true (Long_Name.base_name T_name) (Binding.name thmN), attrs), [(thms, [])])) T_names thmss) #> filter_out (null o fst o hd o snd); val notes = [(sizeN, size_thmss, nitpicksimp_attrs @ simp_attrs), (size_genN, size_gen_thmss, []), (size_gen_o_mapN, size_gen_o_map_thmss, []), (size_neqN, size_neq_thmss, [])] |> massage_multi_notes; val (noted, lthy3) = lthy2 |> Spec_Rules.add Binding.empty Spec_Rules.equational size_consts size_simps |> Spec_Rules.add Binding.empty Spec_Rules.equational overloaded_size_consts overloaded_size_simps |> Code.declare_default_eqns (map (rpair true) (flat size_thmss)) (*Ideally, this would be issued only if the "code" plugin is enabled.*) |> Local_Theory.notes notes; val phi0 = substitute_noted_thm noted; in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (@{fold 3} (fn T_name => fn Const (size_name, _) => fn overloaded_size_def => let val morph = Morphism.thm (phi0 $> phi) in Symtab.update (T_name, (size_name, (morph overloaded_size_def, map morph overloaded_size_simps, maps (map morph) size_gen_o_map_thmss))) end) T_names size_consts overloaded_size_defs)) end | generate_datatype_size _ lthy = lthy; val size_plugin = Plugin_Name.declare_setup \<^binding>\size\; val _ = Theory.setup (fp_sugars_interpretation size_plugin generate_datatype_size); end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1272 +1,1272 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => let val pos = Position.thread_data () in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = let val tab = Data.get (Context.Theory thy); val pos = Position.thread_data (); in if Symtab.defined tab name then thy else thy |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^Const_>\Not for \t' $ Bound 0\\) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, \<^Const_>\HOL.eq _ for \Bound 0\ t'\) => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^Const_>\Not for \<^Const_>\HOL.eq _ for \Bound 0\ t'\\) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] (Thm.instantiate (TVars.make rho_As, Vars.empty) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') (Thm.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^Const>\True\ else \<^Const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; - val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); - val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); + val exhaust_case_names_attr = Attrib.internal \<^here> (K (Rule_Cases.case_names exhaust_cases)); + val cases_type_attr = Attrib.internal \<^here> (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) - #> Local_Theory.declaration {syntax = false, pervasive = false} + #> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); (** external views **) (* document antiquotations *) local fun antiquote_setup binding co = Document_Output.antiquotation_pretty_source_embedded binding ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- Args.type_name {proper = true, strict = true}) (fn ctxt => fn (pos, type_name) => let fun err () = error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); in (case ctr_sugar_of ctxt type_name of NONE => err () | SOME {kind, T = T0, ctrs = ctrs0, ...} => let val _ = if co = (kind = Codatatype) then () else err (); val T = Logic.unvarifyT_global T0; val ctrs = map Logic.unvarify_global ctrs0; val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); fun pretty_ctr ctr = Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: map pretty_typ_bracket (binder_types (fastype_of ctr)))); in Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) end) end); in val _ = Theory.setup (antiquote_setup \<^binding>\datatype\ false #> antiquote_setup \<^binding>\codatatype\ true); end; (* theory export *) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); val datatypes = (Data.get (Context.Theory thy), []) |-> Symtab.fold (fn (name, (pos, {kind, T, ctrs, ...})) => if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I else let val pos_properties = Thy_Info.adjust_pos_properties context pos; val typ = Logic.unvarifyT_global T; val constrs = map Logic.unvarify_global ctrs; val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); val constructors = map (fn t => (t, Term.type_of t)) constrs; in cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) end); in if null datatypes then () else Export_Theory.export_body thy "datatypes" let open XML.Encode Term_XML.Encode in list (pair properties (pair string (pair bool (pair (list (pair string sort)) (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes end end else ()); end; diff --git a/src/HOL/Tools/Function/function.ML b/src/HOL/Tools/Function/function.ML --- a/src/HOL/Tools/Function/function.ML +++ b/src/HOL/Tools/Function/function.ML @@ -1,288 +1,288 @@ (* Title: HOL/Tools/Function/function.ML Author: Alexander Krauss, TU Muenchen Main entry points to the function package. *) signature FUNCTION = sig type info = Function_Common.info val add_function: (binding * typ option * mixfix) list -> Specification.multi_specs -> Function_Common.function_config -> (Proof.context -> tactic) -> local_theory -> info * local_theory val add_function_cmd: (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> Function_Common.function_config -> (Proof.context -> tactic) -> bool -> local_theory -> info * local_theory val function: (binding * typ option * mixfix) list -> Specification.multi_specs -> Function_Common.function_config -> local_theory -> Proof.state val function_cmd: (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> Function_Common.function_config -> bool -> local_theory -> Proof.state val prove_termination: term option -> tactic -> local_theory -> info * local_theory val prove_termination_cmd: string option -> tactic -> local_theory -> info * local_theory val termination : term option -> local_theory -> Proof.state val termination_cmd : string option -> local_theory -> Proof.state val get_congs : Proof.context -> thm list val get_info : Proof.context -> term -> info end structure Function : FUNCTION = struct open Function_Lib open Function_Common val simp_attribs = @{attributes [simp, nitpick_simp]} val psimp_attribs = @{attributes [nitpick_psimp]} fun note_derived (a, atts) (fname, thms) = Local_Theory.note ((derived_name fname a, atts), thms) #> apfst snd fun add_simps fnames post sort extra_qualify label mod_binding moreatts simps lthy = let val spec = post simps |> map (apfst (apsnd (fn ats => moreatts @ ats))) |> map (apfst (apfst extra_qualify)) val (saved_spec_simps, lthy') = fold_map Local_Theory.note spec lthy val saved_simps = maps snd saved_spec_simps val simps_by_f = sort saved_simps fun note fname simps = Local_Theory.note ((mod_binding (derived_name fname label), []), simps) #> snd in (saved_simps, fold2 note fnames simps_by_f lthy') end fun prepare_function do_print prep fixspec eqns config lthy = let val ((fixes0, spec0), ctxt') = prep fixspec eqns lthy val fixes = map (apfst (apfst Binding.name_of)) fixes0 val spec = map (fn (bnd, prop) => (bnd, [prop])) spec0 val (eqs, post, sort_cont, cnames) = get_preproc lthy config ctxt' fixes spec val fnames = map (fst o fst) fixes0 val defname = Binding.conglomerate fnames; val FunctionConfig {partials, default, ...} = config val _ = if is_some default then legacy_feature "\"function (default)\" -- use 'partial_function' instead" else () val ((goal_state, cont), lthy') = Function_Mutual.prepare_function_mutual config defname fixes0 eqs lthy fun afterqed [[proof]] lthy1 = let val result = cont lthy1 (Thm.close_derivation \<^here> proof) val FunctionResult {fs, R, dom, psimps, simple_pinducts, termination, domintros, cases, ...} = result val pelims = Function_Elims.mk_partial_elim_rules lthy1 result val concealed_partial = if partials then I else Binding.concealed val addsmps = add_simps fnames post sort_cont val (((((psimps', [pinducts']), [termination']), cases'), pelims'), lthy2) = lthy1 |> addsmps (concealed_partial o Binding.qualify false "partial") "psimps" concealed_partial psimp_attribs psimps ||>> Local_Theory.notes [((concealed_partial (derived_name defname "pinduct"), []), simple_pinducts |> map (fn th => ([th], [Attrib.case_names cnames, Attrib.consumes (1 - Thm.nprems_of th)] @ @{attributes [induct pred]})))] ||>> (apfst snd o Local_Theory.note ((Binding.concealed (derived_name defname "termination"), []), [termination])) ||>> fold_map (note_derived ("cases", [Attrib.case_names cnames])) (fnames ~~ map single cases) ||>> fold_map (note_derived ("pelims", [Attrib.consumes 1, Attrib.constraints 1])) (fnames ~~ pelims) ||> (case domintros of NONE => I | SOME thms => Local_Theory.note ((derived_name defname "domintros", []), thms) #> snd) val info = { add_simps=addsmps, fnames=fnames, case_names=cnames, psimps=psimps', pinducts=snd pinducts', simps=NONE, inducts=NONE, termination=termination', totality=NONE, fs=fs, R=R, dom=dom, defname=defname, is_partial=true, cases=flat cases', pelims=pelims',elims=NONE} val _ = Proof_Display.print_consts do_print (Position.thread_data ()) lthy2 (K false) (map fst fixes) in (info, - lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false} + lthy2 |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => add_function_data (transform_function_data phi info))) end in ((goal_state, afterqed), lthy') end fun gen_add_function do_print prep fixspec eqns config tac lthy = let val ((goal_state, afterqed), lthy') = prepare_function do_print prep fixspec eqns config lthy val pattern_thm = case SINGLE (tac lthy') goal_state of NONE => error "pattern completeness and compatibility proof failed" | SOME st => Goal.finish lthy' st in lthy' |> afterqed [[pattern_thm]] end val add_function = gen_add_function false Specification.check_multi_specs fun add_function_cmd a b c d int = gen_add_function int Specification.read_multi_specs a b c d fun gen_function do_print prep fixspec eqns config lthy = let val ((goal_state, afterqed), lthy') = prepare_function do_print prep fixspec eqns config lthy in lthy' |> Proof.theorem NONE (snd oo afterqed) [[(Logic.unprotect (Thm.concl_of goal_state), [])]] |> Proof.refine_singleton (Method.primitive_text (K (K goal_state))) end val function = gen_function false Specification.check_multi_specs fun function_cmd a b c int = gen_function int Specification.read_multi_specs a b c fun prepare_termination_proof prep_binding prep_term raw_term_opt lthy = let val term_opt = Option.map (prep_term lthy) raw_term_opt val info = (case term_opt of SOME t => (case import_function_data t lthy of SOME info => info | NONE => error ("Not a function: " ^ quote (Syntax.string_of_term lthy t))) | NONE => (case import_last_function lthy of SOME info => info | NONE => error "Not a function")) val { termination, fs, R, add_simps, case_names, psimps, pinducts, defname, fnames, cases, dom, pelims, ...} = info val domT = domain_type (fastype_of R) val goal = HOLogic.mk_Trueprop (HOLogic.mk_all ("x", domT, mk_acc domT R $ Free ("x", domT))) fun afterqed [[raw_totality]] lthy1 = let val totality = Thm.close_derivation \<^here> raw_totality val remove_domain_condition = full_simplify (put_simpset HOL_basic_ss lthy1 addsimps [totality, @{thm True_implies_equals}]) val tsimps = map remove_domain_condition psimps val tinduct = map remove_domain_condition pinducts val telims = map (map remove_domain_condition) pelims in lthy1 |> add_simps prep_binding "simps" prep_binding simp_attribs tsimps ||> Code.declare_default_eqns (map (rpair true) tsimps) ||>> Local_Theory.note ((prep_binding (derived_name defname "induct"), [Attrib.case_names case_names]), tinduct) ||>> fold_map (note_derived ("elims", [Attrib.consumes 1, Attrib.constraints 1])) (map prep_binding fnames ~~ telims) |-> (fn ((simps,(_,inducts)), elims) => fn lthy2 => let val info' = { is_partial=false, defname=defname, fnames=fnames, add_simps=add_simps, case_names=case_names, fs=fs, R=R, dom=dom, psimps=psimps, pinducts=pinducts, simps=SOME simps, inducts=SOME inducts, termination=termination, totality=SOME totality, cases=cases, pelims=pelims, elims=SOME elims} |> transform_function_data (Morphism.binding_morphism "" prep_binding) in (info', lthy2 - |> Local_Theory.declaration {syntax = false, pervasive = false} + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => add_function_data (transform_function_data phi info')) |> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef fs tsimps) end) end in (goal, afterqed, termination) end fun gen_prove_termination prep_term raw_term_opt tac lthy = let val (goal, afterqed, termination) = prepare_termination_proof I prep_term raw_term_opt lthy val totality = Goal.prove lthy [] [] goal (K tac) in afterqed [[totality]] lthy end val prove_termination = gen_prove_termination Syntax.check_term val prove_termination_cmd = gen_prove_termination Syntax.read_term fun gen_termination prep_term raw_term_opt lthy = let val (goal, afterqed, termination) = prepare_termination_proof Binding.reset_pos prep_term raw_term_opt lthy in lthy |> Proof_Context.note_thms "" ((Binding.empty, [Context_Rules.rule_del]), [([allI], [])]) |> snd |> Proof_Context.note_thms "" ((Binding.empty, [Context_Rules.intro_bang (SOME 1)]), [([allI], [])]) |> snd |> Proof_Context.note_thms "" ((Binding.name "termination", [Context_Rules.intro_bang (SOME 0)]), [([Goal.norm_result lthy termination], [])]) |> snd |> Proof.theorem NONE (snd oo afterqed) [[(goal, [])]] end val termination = gen_termination Syntax.check_term val termination_cmd = gen_termination Syntax.read_term (* Datatype hook to declare datatype congs as "function_congs" *) fun add_case_cong n thy = let val cong = #case_cong (Old_Datatype_Data.the_info thy n) |> safe_mk_meta_eq in Context.theory_map (Function_Context_Tree.add_function_cong cong) thy end val _ = Theory.setup (Old_Datatype_Data.interpretation (K (fold add_case_cong))) (* get info *) val get_congs = Function_Context_Tree.get_function_congs fun get_info ctxt t = Function_Common.retrieve_function_data ctxt t |> the_single |> snd (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof' \<^command_keyword>\function\ "define general recursive functions" (function_parser default_config >> (fn (config, (fixes, specs)) => function_cmd fixes specs config)) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\termination\ "prove termination of a recursive function" (Scan.option Parse.term >> termination_cmd) end diff --git a/src/HOL/Tools/Lifting/lifting_def.ML b/src/HOL/Tools/Lifting/lifting_def.ML --- a/src/HOL/Tools/Lifting/lifting_def.ML +++ b/src/HOL/Tools/Lifting/lifting_def.ML @@ -1,675 +1,675 @@ (* Title: HOL/Tools/Lifting/lifting_def.ML Author: Ondrej Kuncar Definitions for constants on quotient types. *) signature LIFTING_DEF = sig datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ type lift_def val rty_of_lift_def: lift_def -> typ val qty_of_lift_def: lift_def -> typ val rhs_of_lift_def: lift_def -> term val lift_const_of_lift_def: lift_def -> term val def_thm_of_lift_def: lift_def -> thm val rsp_thm_of_lift_def: lift_def -> thm val abs_eq_of_lift_def: lift_def -> thm val rep_eq_of_lift_def: lift_def -> thm option val code_eq_of_lift_def: lift_def -> code_eq val transfer_rules_of_lift_def: lift_def -> thm list val morph_lift_def: morphism -> lift_def -> lift_def val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def val mk_lift_const_of_lift_def: typ -> lift_def -> term type config = { notes: bool } val map_config: (bool -> bool) -> config -> config val default_config: config val generate_parametric_transfer_rule: Proof.context -> thm -> thm -> thm val add_lift_def: config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory val prepare_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> thm list -> local_theory -> term option * (thm -> Proof.context -> lift_def * local_theory) val gen_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val lift_def: config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val can_generate_code_cert: thm -> bool end structure Lifting_Def: LIFTING_DEF = struct open Lifting_Util infix 0 MRSL datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ datatype lift_def = LIFT_DEF of { rty: typ, qty: typ, rhs: term, lift_const: term, def_thm: thm, rsp_thm: thm, abs_eq: thm, rep_eq: thm option, code_eq: code_eq, transfer_rules: thm list }; fun rep_lift_def (LIFT_DEF lift_def) = lift_def; val rty_of_lift_def = #rty o rep_lift_def; val qty_of_lift_def = #qty o rep_lift_def; val rhs_of_lift_def = #rhs o rep_lift_def; val lift_const_of_lift_def = #lift_const o rep_lift_def; val def_thm_of_lift_def = #def_thm o rep_lift_def; val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def; val abs_eq_of_lift_def = #abs_eq o rep_lift_def; val rep_eq_of_lift_def = #rep_eq o rep_lift_def; val code_eq_of_lift_def = #code_eq o rep_lift_def; val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def; fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules = LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }; fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }) = LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const, def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq, code_eq = f9 code_eq, transfer_rules = f10 transfer_rules } fun morph_lift_def phi = let val mtyp = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm) end fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty) fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def) (lift_const_of_lift_def lift_def) fun inst_of_lift_def ctxt qty lift_def = let val instT = Vartab.fold (fn (a, (S, T)) => cons ((a, S), Thm.ctyp_of ctxt T)) (mk_inst_of_lift_def qty lift_def) [] val phi = Morphism.instantiate_morphism (TVars.make instT, Vars.empty) in morph_lift_def phi lift_def end (* Config *) type config = { notes: bool }; fun map_config f1 { notes = notes } = { notes = f1 notes } val default_config = { notes = true }; (* Reflexivity prover *) fun mono_eq_prover ctxt prop = let val refl_rules = Lifting_Info.get_reflexivity_rules ctxt val transfer_rules = Transfer.get_transfer_raw ctxt fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i in SOME (Goal.prove ctxt [] [] prop (K (main_tac 1))) handle ERROR _ => NONE end fun try_prove_refl_rel ctxt rel = let fun mk_ge_eq x = let val T = fastype_of x in Const (\<^const_name>\less_eq\, T --> T --> HOLogic.boolT) $ (Const (\<^const_name>\HOL.eq\, T)) $ x end; val goal = HOLogic.mk_Trueprop (mk_ge_eq rel); in mono_eq_prover ctxt goal end; fun try_prove_reflexivity ctxt prop = let val cprop = Thm.cterm_of ctxt prop val rule = @{thm ge_eq_refl} val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule) val insts = Thm.first_order_match (concl_pat, cprop) val rule = Drule.instantiate_normalize insts rule val prop = hd (Thm.prems_of rule) in case mono_eq_prover ctxt prop of SOME thm => SOME (thm RS rule) | NONE => NONE end (* Generates a parametrized transfer rule. transfer_rule - of the form T t f parametric_transfer_rule - of the form par_R t' t Result: par_T t' f, after substituing op= for relations in par_R that relate a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f using Lifting_Term.merge_transfer_relations *) fun generate_parametric_transfer_rule ctxt0 transfer_rule parametric_transfer_rule = let fun preprocess ctxt thm = let val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm; val param_rel = (snd o dest_comb o fst o dest_comb) tm; val free_vars = Term.add_vars param_rel []; fun make_subst (xi, typ) subst = let val [rty, rty'] = binder_types typ in if Term.is_TVar rty andalso Term.is_Type rty' then (xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst else subst end; val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm; in Conv.fconv_rule ((Conv.concl_conv (Thm.nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv) (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm end fun inst_relcomppI ctxt ant1 ant2 = let val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1 val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2 val fun1 = Thm.cterm_of ctxt (strip_args 2 t1) val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1) val fun2 = Thm.cterm_of ctxt (strip_args 2 t2) val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2) val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI} val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) [])) in infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI end fun zip_transfer_rules ctxt thm = let fun mk_POS ty = Const (\<^const_name>\POS\, ty --> ty --> HOLogic.boolT) val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm val typ = Thm.typ_of_cterm rel val POS_const = Thm.cterm_of ctxt (mk_POS typ) val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) val goal = Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var) in [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply} end val thm = inst_relcomppI ctxt0 parametric_transfer_rule transfer_rule OF [parametric_transfer_rule, transfer_rule] val preprocessed_thm = preprocess ctxt0 thm val (fixed_thm, ctxt1) = ctxt0 |> yield_singleton (apfst snd oo Variable.import true) preprocessed_thm val assms = cprems_of fixed_thm val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add val (prems, ctxt2) = ctxt1 |> fold_map Thm.assume_hyps assms val ctxt3 = ctxt2 |> Context.proof_map (fold add_transfer_rule prems) val zipped_thm = fixed_thm |> undisch_all |> zip_transfer_rules ctxt3 |> implies_intr_list assms |> singleton (Variable.export ctxt3 ctxt0) in zipped_thm end fun print_generate_transfer_info msg = let val error_msg = cat_lines ["Generation of a parametric transfer rule failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in error error_msg end fun map_ter _ x [] = x | map_ter f _ xs = map f xs fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms = let val transfer_rule = ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer}) |> Lifting_Term.parametrize_transfer_rule lthy in (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule])) end (* Generation of the code certificate from the rsp theorem *) fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V) | get_body_types (U, V) = (U, V) fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W) | get_binder_types _ = [] fun get_binder_types_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types_by_rel S (U, W) | get_binder_types_by_rel _ _ = [] fun get_body_type_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_type_by_rel S (U, V) | get_body_type_by_rel _ (U, V) = (U, V) fun get_binder_rels (Const (\<^const_name>\rel_fun\, _) $ R $ S) = R :: get_binder_rels S | get_binder_rels _ = [] fun force_rty_type ctxt rty rhs = let val thy = Proof_Context.theory_of ctxt val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs val rty_schematic = fastype_of rhs_schematic val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty in Envir.subst_term_types match rhs_schematic end fun unabs_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) fun dest_abs (Abs (var_name, T, _)) = (var_name, T) | dest_abs tm = raise TERM("get_abs_var",[tm]) val (var_name, T) = dest_abs (Thm.term_of rhs) val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T))) in Thm.combination def refl_thm |> singleton (Proof_Context.export ctxt' ctxt) end fun unabs_all_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) val xs = strip_abs_vars (Thm.term_of rhs) in fold (K (unabs_def ctxt)) xs def end val map_fun_unfolded = @{thm map_fun_def[abs_def]} |> unabs_def \<^context> |> unabs_def \<^context> |> Local_Defs.unfold0 \<^context> [@{thm comp_def}] fun unfold_fun_maps ctm = let fun unfold_conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\map_fun\, _) $ _ $ _ => (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm | _ => Conv.all_conv ctm in (Conv.fun_conv unfold_conv) ctm end fun unfold_fun_maps_beta ctm = let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false) in (unfold_fun_maps then_conv try_beta_conv) ctm end fun prove_rel ctxt rsp_thm (rty, qty) = let val ty_args = get_binder_types (rty, qty) fun disch_arg args_ty thm = let val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty in [quot_thm, thm] MRSL @{thm apply_rsp''} end in fold disch_arg ty_args rsp_thm end exception CODE_CERT_GEN of string fun simplify_code_eq ctxt def_thm = Local_Defs.unfold0 ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm (* quot_thm - quotient theorem (Quotient R Abs Rep T). returns: whether the Lifting package is capable to generate code for the abstract type represented by quot_thm *) fun can_generate_code_cert quot_thm = case quot_thm_rel quot_thm of Const (\<^const_name>\HOL.eq\, _) => true | Const (\<^const_name>\eq_onp\, _) $ _ => true | _ => false fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) = let val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm val unabs_def = unabs_all_def ctxt unfolded_def in if body_type rty = body_type qty then SOME (simplify_code_eq ctxt (HOLogic.mk_obj_eq unabs_def)) else let val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty)) val rel_fun = prove_rel ctxt rsp_thm (rty, qty) val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq} in case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of SOME mono_eq_thm => let val rep_abs_eq = mono_eq_thm RS rep_abs_thm val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm) val rep_refl = HOLogic.mk_obj_eq (Thm.reflexive rep) val repped_eq = [rep_refl, HOLogic.mk_obj_eq unabs_def] MRSL @{thm cong} val code_cert = [repped_eq, rep_abs_eq] MRSL trans in SOME (simplify_code_eq ctxt code_cert) end | NONE => NONE end end fun generate_abs_eq ctxt def_thm rsp_thm quot_thm = let val abs_eq_with_assms = let val (rty, qty) = quot_thm_rty_qty quot_thm val rel = quot_thm_rel quot_thm val ty_args = get_binder_types_by_rel rel (rty, qty) val body_type = get_body_type_by_rel rel (rty, qty) val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type val rep_abs_folded_unmapped_thm = let val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq} val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id) val unfolded_maps_eq = unfold_fun_maps ctm val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap} val prems_pat = (hd o Drule.cprems_of) t1 val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq) in unfolded_maps_eq RS (Drule.instantiate_normalize insts t1) end in rep_abs_folded_unmapped_thm |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm])) end val prem_rels = get_binder_rels (quot_thm_rel quot_thm); val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt) |> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the) |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl})) val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms in simplify_code_eq ctxt abs_eq end fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy = let fun no_abstr (t $ u) = no_abstr t andalso no_abstr u | no_abstr (Abs (_, _, t)) = no_abstr t | no_abstr (Const (name, _)) = not (Code.is_abstr thy name) | no_abstr _ = true fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) andalso no_abstr (Thm.prop_of eqn) fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq) in if is_valid_eq abs_eq_thm then (ABS_EQ, Code.declare_default_eqns_global [(abs_eq_thm, true)] thy) else let val (rty_body, qty_body) = get_body_types (rty, qty) in if rty_body = qty_body then (REP_EQ, Code.declare_default_eqns_global [(the opt_rep_eq_thm, true)] thy) else if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm) then (REP_EQ, Code.declare_abstract_eqn_global (the opt_rep_eq_thm) thy) else (NONE_EQ, thy) end end local fun no_no_code ctxt (rty, qty) = if same_type_constrs (rty, qty) then forall (no_no_code ctxt) (Targs rty ~~ Targs qty) else if Term.is_Type qty then if Lifting_Info.is_no_code_type ctxt (Tname qty) then false else let val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty) val (rty's, rtyqs) = (Targs rty', Targs rtyq) in forall (no_no_code ctxt) (rty's ~~ rtyqs) end else true fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = let fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term in Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty] end exception DECODE fun decode_code_eq thm = if Thm.nprems_of thm > 0 then raise DECODE else let val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type in (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) end structure Data = Generic_Data ( type T = code_eq option val empty = NONE fun merge _ = NONE ); fun register_encoded_code_eq thm thy = let val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy in Context.theory_map (Data.put (SOME code_eq)) thy end handle DECODE => thy val register_code_eq_attribute = Thm.declaration_attribute (fn thm => Context.mapping (register_encoded_code_eq thm) I) - val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute) + val register_code_eq_attrib = Attrib.internal \<^here> (K register_code_eq_attribute) in fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy = let val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty) in if no_no_code lthy (rty, qty) then let val lthy' = lthy |> (#2 oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy')) val code_eq = if is_some opt_code_eq then the opt_code_eq else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know which code equation is going to be used. This is going to be resolved at the point when an interpretation of the locale is executed. *) val lthy'' = lthy' - |> Local_Theory.declaration {syntax = false, pervasive = true} (K (Data.put NONE)) + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (K (Data.put NONE)) in (code_eq, lthy'') end else (NONE_EQ, lthy) end end (* Defines an operation on an abstract type in terms of a corresponding operation on a representation type. var - a binding and a mixfix of the new constant being defined qty - an abstract type of the new constant rhs - a term representing the new constant on the raw level rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'), i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs" par_thms - a parametricity theorem for rhs *) fun add_lift_def (config: config) (binding, mx) qty rhs rsp_thm par_thms lthy0 = let val rty = fastype_of rhs val quot_thm = Lifting_Term.prove_quot_thm lthy0 (rty, qty) val absrep_trm = quot_thm_abs quot_thm val rty_forced = (domain_type o fastype_of) absrep_trm val forced_rhs = force_rty_type lthy0 rty_forced rhs val lhs = Free (Binding.name_of binding, qty) val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs) val (_, prop') = Local_Defs.cert_def lthy0 (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val var = ((#notes config = false ? Binding.concealed) binding, mx) val def_name = Thm.make_def_binding (#notes config) (#1 var) val ((lift_const, (_ , def_thm)), lthy1) = lthy0 |> Local_Theory.define (var, ((def_name, []), newrhs)) val transfer_rules = generate_transfer_rules lthy1 quot_thm rsp_thm def_thm par_thms val abs_eq_thm = generate_abs_eq lthy1 def_thm rsp_thm quot_thm val opt_rep_eq_thm = generate_rep_eq lthy1 def_thm rsp_thm (rty_forced, qty) fun notes names = let val lhs_name = Binding.reset_pos (#1 var) val rsp_thmN = Binding.qualify_name true lhs_name "rsp" val abs_eq_thmN = Binding.qualify_name true lhs_name "abs_eq" val rep_eq_thmN = Binding.qualify_name true lhs_name "rep_eq" val transfer_ruleN = Binding.qualify_name true lhs_name "transfer" val notes = [(rsp_thmN, [], [rsp_thm]), (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules), (abs_eq_thmN, [], [abs_eq_thm])] @ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => []) in if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes else map_filter (fn (_, attrs, thms) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) notes end val (code_eq, lthy2) = lthy1 |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm opt_rep_eq_thm code_eq transfer_rules in lthy2 |> (snd o Local_Theory.begin_nested) |> Local_Theory.notes (notes (#notes config)) |> snd |> `(fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def) ||> Local_Theory.end_nested end (* This is not very cheap way of getting the rules but we have only few active liftings in the current setting *) fun get_cr_pcr_eqs ctxt = let fun collect (data : Lifting_Info.quotient) l = if is_some (#pcr_info data) then ((Thm.symmetric o safe_mk_meta_eq o Thm.transfer' ctxt o #pcr_cr_eq o the o #pcr_info) data :: l) else l val table = Lifting_Info.get_quotients ctxt in Symtab.fold (fn (_, data) => fn l => collect data l) table [] end fun prepare_lift_def add_lift_def var qty rhs par_thms ctxt = let val rsp_rel = Lifting_Term.equiv_relation ctxt (fastype_of rhs, qty) val rty_forced = (domain_type o fastype_of) rsp_rel; val forced_rhs = force_rty_type ctxt rty_forced rhs; val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (Raw_Simplifier.rewrite ctxt false (get_cr_pcr_eqs ctxt))) val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs) |> Thm.cterm_of ctxt |> cr_to_pcr_conv |> `Thm.concl_of |>> Logic.dest_equals |>> snd val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2 val opt_proven_rsp_thm = try_prove_reflexivity ctxt prsp_tm fun after_qed internal_rsp_thm = add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms in case opt_proven_rsp_thm of SOME thm => (NONE, K (after_qed thm)) | NONE => (SOME prsp_tm, after_qed) end fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy = let val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy in case goal of SOME goal => let val rsp_thm = Goal.prove_sorry lthy [] [] goal (tac o #context) |> Thm.close_derivation \<^here> in after_qed rsp_thm lthy end | NONE => after_qed Drule.dummy_thm lthy end val lift_def = gen_lift_def o add_lift_def end (* structure *) diff --git a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML --- a/src/HOL/Tools/Lifting/lifting_def_code_dt.ML +++ b/src/HOL/Tools/Lifting/lifting_def_code_dt.ML @@ -1,833 +1,833 @@ (* Title: HOL/Tools/Lifting/lifting_def_code_dt.ML Author: Ondrej Kuncar Workaround that allows us to execute lifted constants that have as a return type a datatype containing a subtype; lift_definition command *) signature LIFTING_DEF_CODE_DT = sig type rep_isom_data val isom_of_rep_isom_data: rep_isom_data -> term val transfer_of_rep_isom_data: rep_isom_data -> thm val bundle_name_of_rep_isom_data: rep_isom_data -> string val pointer_of_rep_isom_data: rep_isom_data -> string type code_dt val rty_of_code_dt: code_dt -> typ val qty_of_code_dt: code_dt -> typ val wit_of_code_dt: code_dt -> term val wit_thm_of_code_dt: code_dt -> thm val rep_isom_data_of_code_dt: code_dt -> rep_isom_data option val morph_code_dt: morphism -> code_dt -> code_dt val mk_witness_of_code_dt: typ -> code_dt -> term val mk_rep_isom_of_code_dt: typ -> code_dt -> term option val code_dt_of: Proof.context -> typ * typ -> code_dt option val code_dt_of_global: theory -> typ * typ -> code_dt option val all_code_dt_of: Proof.context -> code_dt list val all_code_dt_of_global: theory -> code_dt list type config_code_dt = { code_dt: bool, lift_config: Lifting_Def.config } val default_config_code_dt: config_code_dt val add_lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_code_dt: config_code_dt -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> Lifting_Def.lift_def * local_theory val lift_def_cmd: string list * (binding * string option * mixfix) * string * (Facts.ref * Token.src list) list -> local_theory -> Proof.state end structure Lifting_Def_Code_Dt: LIFTING_DEF_CODE_DT = struct open Ctr_Sugar_Util BNF_Util BNF_FP_Util BNF_FP_Def_Sugar Lifting_Def Lifting_Util infix 0 MRSL (** data structures **) (* all type variables in qty are in rty *) datatype rep_isom_data = REP_ISOM of { isom: term, transfer: thm, bundle_name: string, pointer: string } fun isom_of_rep_isom_data (REP_ISOM rep_isom) = #isom rep_isom; fun transfer_of_rep_isom_data (REP_ISOM rep_isom) = #transfer rep_isom; fun bundle_name_of_rep_isom_data (REP_ISOM rep_isom) = #bundle_name rep_isom; fun pointer_of_rep_isom_data (REP_ISOM rep_isom) = #pointer rep_isom; datatype code_dt = CODE_DT of { rty: typ, qty: typ, wit: term, wit_thm: thm, rep_isom_data: rep_isom_data option }; fun rty_of_code_dt (CODE_DT code_dt) = #rty code_dt; fun qty_of_code_dt (CODE_DT code_dt) = #qty code_dt; fun wit_of_code_dt (CODE_DT code_dt) = #wit code_dt; fun wit_thm_of_code_dt (CODE_DT code_dt) = #wit_thm code_dt; fun rep_isom_data_of_code_dt (CODE_DT code_dt) = #rep_isom_data code_dt; fun ty_alpha_equiv (T, U) = Type.raw_instance (T, U) andalso Type.raw_instance (U, T); fun code_dt_eq c = (ty_alpha_equiv o apply2 rty_of_code_dt) c andalso (ty_alpha_equiv o apply2 qty_of_code_dt) c; fun term_of_code_dt code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> Net.encode_type |> single; (* modulo renaming, typ must contain TVars *) fun is_code_dt_of_type (rty, qty) code_dt = code_dt |> `rty_of_code_dt ||> qty_of_code_dt |> HOLogic.mk_prodT |> curry ty_alpha_equiv (HOLogic.mk_prodT (rty, qty)); fun mk_rep_isom_data isom transfer bundle_name pointer = REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer} fun mk_code_dt rty qty wit wit_thm rep_isom_data = CODE_DT { rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data }; fun map_rep_isom_data f1 f2 f3 f4 (REP_ISOM { isom = isom, transfer = transfer, bundle_name = bundle_name, pointer = pointer }) = REP_ISOM { isom = f1 isom, transfer = f2 transfer, bundle_name = f3 bundle_name, pointer = f4 pointer }; fun map_code_dt f1 f2 f3 f4 f5 f6 f7 f8 (CODE_DT {rty = rty, qty = qty, wit = wit, wit_thm = wit_thm, rep_isom_data = rep_isom_data}) = CODE_DT {rty = f1 rty, qty = f2 qty, wit = f3 wit, wit_thm = f4 wit_thm, rep_isom_data = Option.map (map_rep_isom_data f5 f6 f7 f8) rep_isom_data}; fun update_rep_isom isom transfer binding pointer i = mk_code_dt (rty_of_code_dt i) (qty_of_code_dt i) (wit_of_code_dt i) (wit_thm_of_code_dt i) (SOME (mk_rep_isom_data isom transfer binding pointer)) fun morph_code_dt phi = let val mty = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_code_dt mty mty mterm mthm mterm mthm I I end val transfer_code_dt = morph_code_dt o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = code_dt Item_Net.T val empty = Item_Net.init code_dt_eq term_of_code_dt val merge = Item_Net.merge ); fun code_dt_of_generic context (rty, qty) = let val typ = HOLogic.mk_prodT (rty, qty) val prefiltred = Item_Net.retrieve_matching (Data.get context) (Net.encode_type typ) in prefiltred |> filter (is_code_dt_of_type (rty, qty)) |> map (transfer_code_dt (Context.theory_of context)) |> find_first (fn _ => true) end; fun code_dt_of ctxt (rty, qty) = let val sch_rty = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty val sch_qty = Logic.type_map (singleton (Variable.polymorphic ctxt)) qty in code_dt_of_generic (Context.Proof ctxt) (sch_rty, sch_qty) end; fun code_dt_of_global thy (rty, qty) = let val sch_rty = Logic.varifyT_global rty val sch_qty = Logic.varifyT_global qty in code_dt_of_generic (Context.Theory thy) (sch_rty, sch_qty) end; fun all_code_dt_of_generic context = Item_Net.content (Data.get context) |> map (transfer_code_dt (Context.theory_of context)); val all_code_dt_of = all_code_dt_of_generic o Context.Proof; val all_code_dt_of_global = all_code_dt_of_generic o Context.Theory; fun update_code_dt code_dt = (snd o Local_Theory.begin_nested) - #> Local_Theory.declaration {syntax = false, pervasive = true} + #> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Data.map (Item_Net.update (morph_code_dt phi code_dt))) #> Local_Theory.end_nested fun mk_match_of_code_dt qty code_dt = Vartab.empty |> Type.raw_match (qty_of_code_dt code_dt, qty) |> Vartab.dest |> map (fn (x, (S, T)) => (TVar (x, S), T)); fun mk_witness_of_code_dt qty code_dt = Term.subst_atomic_types (mk_match_of_code_dt qty code_dt) (wit_of_code_dt code_dt) fun mk_rep_isom_of_code_dt qty code_dt = Option.map (isom_of_rep_isom_data #> Term.subst_atomic_types (mk_match_of_code_dt qty code_dt)) (rep_isom_data_of_code_dt code_dt) (** unique name for a type **) fun var_name name sort = if sort = \<^sort>\{type}\ orelse sort = [] then ["x" ^ name] else "x" ^ name :: "x_" :: sort @ ["x_"]; fun concat_Tnames (Type (name, ts)) = name :: maps concat_Tnames ts | concat_Tnames (TFree (name, sort)) = var_name name sort | concat_Tnames (TVar ((name, _), sort)) = var_name name sort; fun unique_Tname (rty, qty) = let val Tnames = map Long_Name.base_name (concat_Tnames rty @ ["x_x"] @ concat_Tnames qty); in fold (Binding.qualify false) (tl Tnames) (Binding.name (hd Tnames)) end; (** witnesses **) fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_witness quot_thm = let val wit_thm = quot_thm RS @{thm type_definition_Quotient_not_empty_witness} val wit = quot_thm_rep quot_thm $ mk_undefined (quot_thm_rty_qty quot_thm |> snd) in (wit, wit_thm) end (** config **) type config_code_dt = { code_dt: bool, lift_config: config } val default_config_code_dt = { code_dt = false, lift_config = default_config } (** Main code **) val ld_no_notes = { notes = false } fun comp_lift_error _ _ = error "Composition of abstract types has not been implemented yet." fun lift qty (quot_thm, (lthy, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm; in if is_none (code_dt_of lthy (rty, qty)) then let val (wit, wit_thm) = (mk_witness quot_thm handle THM _ => error ("code_dt: " ^ quote (Tname qty) ^ " was not defined as a subtype.")) val code_dt = mk_code_dt rty qty wit wit_thm NONE in (quot_thm, (update_code_dt code_dt lthy, rel_eq_onps)) end else (quot_thm, (lthy, rel_eq_onps)) end; fun case_tac rule = Subgoal.FOCUS_PARAMS (fn {context = ctxt, params, ...} => HEADGOAL (rtac ctxt (infer_instantiate' ctxt [SOME (snd (hd params))] rule))); fun bundle_name_of_bundle_binding binding phi context = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding phi binding); fun prove_schematic_quot_thm actions ctxt = Lifting_Term.prove_schematic_quot_thm actions (Lifting_Info.get_quotients ctxt) ctxt fun prove_code_dt (rty, qty) lthy = let val (fold_quot_thm: (local_theory * thm list) Lifting_Term.fold_quot_thm) = { constr = constr, lift = lift, comp_lift = comp_lift_error }; in prove_schematic_quot_thm fold_quot_thm lthy (rty, qty) (lthy, []) |> snd end and add_lift_def_code_dt config var qty rhs rsp_thm par_thms lthy = let fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun ret_rel_conv conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => binop_conv2 Conv.all_conv conv ctm | _ => conv ctm fun R_conv rel_eq_onps ctxt = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} ctxt then_conv Conv.bottom_rewrs_conv rel_eq_onps ctxt val (ret_lift_def, lthy1) = add_lift_def (#lift_config config) var qty rhs rsp_thm par_thms lthy in if (not (#code_dt config) orelse (code_eq_of_lift_def ret_lift_def <> NONE_EQ) andalso (code_eq_of_lift_def ret_lift_def <> UNKNOWN_EQ)) (* Let us try even in case of UNKNOWN_EQ. If this leads to problems, the user can always say that they do not want this workaround. *) then (ret_lift_def, lthy1) else let val lift_def = inst_of_lift_def lthy1 qty ret_lift_def val rty = rty_of_lift_def lift_def val rty_ret = body_type rty val qty_ret = body_type qty val (lthy2, rel_eq_onps) = prove_code_dt (rty_ret, qty_ret) lthy1 val code_dt = code_dt_of lthy2 (rty_ret, qty_ret) in if is_none code_dt orelse is_none (rep_isom_data_of_code_dt (the code_dt)) then (ret_lift_def, lthy2) else let val code_dt = the code_dt val rhs = dest_comb (rhs_of_lift_def lift_def) |> snd val rep_isom_data = code_dt |> rep_isom_data_of_code_dt |> the val pointer = pointer_of_rep_isom_data rep_isom_data val quot_active = Lifting_Info.lookup_restore_data lthy2 pointer |> the |> #quotient |> #quot_thm |> Lifting_Info.lookup_quot_thm_quotients lthy2 |> is_some val qty_code_dt_bundle_name = bundle_name_of_rep_isom_data rep_isom_data val rep_isom = mk_rep_isom_of_code_dt qty_ret code_dt |> the val lthy3 = if quot_active then lthy2 else Bundle.includes [qty_code_dt_bundle_name] lthy2 fun qty_isom_of_rep_isom rep = rep |> dest_Const |> snd |> domain_type val qty_isom = qty_isom_of_rep_isom rep_isom val f'_var = (Binding.suffix_name "_aux" (fst var), NoSyn); val f'_qty = strip_type qty |> fst |> rpair qty_isom |> op ---> val f'_rsp_rel = Lifting_Term.equiv_relation lthy3 (rty, f'_qty); val rsp = rsp_thm_of_lift_def lift_def val rel_eq_onps_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (ret_rel_conv (R_conv rel_eq_onps lthy3))) val rsp_norm = Conv.fconv_rule rel_eq_onps_conv rsp val f'_rsp_goal = HOLogic.mk_Trueprop (f'_rsp_rel $ rhs $ rhs); val f'_rsp = Goal.prove_sorry lthy3 [] [] f'_rsp_goal (fn {context = ctxt, prems = _} => HEADGOAL (CONVERSION (rel_eq_onps_conv) THEN' rtac ctxt rsp_norm)) |> Thm.close_derivation \<^here> val (f'_lift_def, lthy4) = add_lift_def ld_no_notes f'_var f'_qty rhs f'_rsp [] lthy3 val f'_lift_def = inst_of_lift_def lthy4 f'_qty f'_lift_def val f'_lift_const = mk_lift_const_of_lift_def f'_qty f'_lift_def val (args, args_ctxt) = mk_Frees "x" (binder_types qty) lthy4 val f_alt_def_goal_lhs = list_comb (lift_const_of_lift_def lift_def, args); val f_alt_def_goal_rhs = rep_isom $ list_comb (f'_lift_const, args); val f_alt_def_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (f_alt_def_goal_lhs, f_alt_def_goal_rhs)); fun f_alt_def_tac ctxt i = EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [id_apply]), rtac ctxt refl] i; val rep_isom_transfer = transfer_of_rep_isom_data rep_isom_data val (_, transfer_ctxt) = args_ctxt |> Proof_Context.note_thms "" (Binding.empty_atts, [([rep_isom_transfer], [Transfer.transfer_add])]) val f_alt_def = Goal.prove_sorry transfer_ctxt [] [] f_alt_def_goal (fn {context = goal_ctxt, ...} => HEADGOAL (f_alt_def_tac goal_ctxt)) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy4) val lthy5 = lthy4 |> Local_Theory.note ((Binding.empty, @{attributes [code]}), [f_alt_def]) |> snd (* if processing a mutual datatype (there is a cycle!) the corresponding quotient will be needed later and will be forgotten later *) |> (if quot_active then I else Lifting_Setup.lifting_forget pointer) in (ret_lift_def, lthy5) end end end and mk_rep_isom qty_isom_bundle (rty, qty, qty_isom) lthy0 = let (* logical definition of qty qty_isom isomorphism *) val uTname = unique_Tname (rty, qty) fun eq_onp_to_top_tac ctxt = SELECT_GOAL (Local_Defs.unfold0_tac ctxt (@{thm eq_onp_top_eq_eq[symmetric]} :: Lifting_Info.get_relator_eq_onp_rules ctxt)) fun lift_isom_tac ctxt = HEADGOAL (eq_onp_to_top_tac ctxt THEN' (rtac ctxt @{thm id_transfer})); val (rep_isom_lift_def, lthy1) = lthy0 |> (snd o Local_Theory.begin_nested) |> lift_def ld_no_notes (Binding.qualify_name true uTname "Rep_isom", NoSyn) (qty_isom --> qty) (HOLogic.id_const rty) lift_isom_tac [] |>> inst_of_lift_def lthy0 (qty_isom --> qty); val (abs_isom, lthy2) = lthy1 |> lift_def ld_no_notes (Binding.qualify_name true uTname "Abs_isom", NoSyn) (qty --> qty_isom) (HOLogic.id_const rty) lift_isom_tac [] |>> mk_lift_const_of_lift_def (qty --> qty_isom); val rep_isom = lift_const_of_lift_def rep_isom_lift_def val pointer = Lifting_Setup.pointer_of_bundle_binding lthy2 qty_isom_bundle fun code_dt phi context = code_dt_of lthy2 (rty, qty) |> the |> update_rep_isom rep_isom (transfer_rules_of_lift_def rep_isom_lift_def |> hd) (bundle_name_of_bundle_binding qty_isom_bundle phi context) pointer; val lthy3 = lthy2 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => Data.map (Item_Net.update (morph_code_dt phi (code_dt phi context))) context) |> Local_Theory.end_nested (* in order to make the qty qty_isom isomorphism executable we have to define discriminators and selectors for qty_isom *) val (rty_name, typs) = dest_Type rty val (_, qty_typs) = dest_Type qty val fp = BNF_FP_Def_Sugar.fp_sugar_of lthy3 rty_name val fp = if is_some fp then the fp else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val ctr_sugar = fp |> #fp_ctr_sugar |> #ctr_sugar val ctrs = map (Ctr_Sugar.mk_ctr typs) (#ctrs ctr_sugar); val qty_ctrs = map (Ctr_Sugar.mk_ctr qty_typs) (#ctrs ctr_sugar); val ctr_Tss = map (dest_Const #> snd #> binder_types) ctrs; val qty_ctr_Tss = map (dest_Const #> snd #> binder_types) qty_ctrs; val n = length ctrs; val ks = 1 upto n; val (xss, _) = mk_Freess "x" ctr_Tss lthy3; fun sel_retT (rty' as Type (s, rtys'), qty' as Type (s', qtys')) = if (rty', qty') = (rty, qty) then qty_isom else (if s = s' then Type (s, map sel_retT (rtys' ~~ qtys')) else qty') | sel_retT (_, qty') = qty'; val sel_retTs = map2 (map2 (sel_retT oo pair)) ctr_Tss qty_ctr_Tss fun lazy_prove_code_dt (rty, qty) lthy = if is_none (code_dt_of lthy (rty, qty)) then prove_code_dt (rty, qty) lthy |> fst else lthy; val lthy4 = fold2 (fold2 (lazy_prove_code_dt oo pair)) ctr_Tss sel_retTs lthy3 val sel_argss = @{map 4} (fn k => fn xs => @{map 2} (fn x => fn qty_ret => (k, qty_ret, (xs, x)))) ks xss xss sel_retTs; fun mk_sel_casex (_, _, (_, x)) = Ctr_Sugar.mk_case typs (x |> dest_Free |> snd) (#casex ctr_sugar); val dis_casex = Ctr_Sugar.mk_case typs HOLogic.boolT (#casex ctr_sugar); fun mk_sel_case_args lthy ctr_Tss ks (k, qty_ret, (xs, x)) = let val T = x |> dest_Free |> snd; fun gen_undef_wit Ts wits = case code_dt_of lthy (T, qty_ret) of SOME code_dt => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_witness_of_code_dt qty_ret code_dt), wit_thm_of_code_dt code_dt :: wits) | NONE => (fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T), wits) in @{fold_map 2} (fn Ts => fn k' => fn wits => (if k = k' then (fold_rev Term.lambda xs x, wits) else gen_undef_wit Ts wits)) ctr_Tss ks [] end; fun mk_sel_rhs arg = let val (sel_rhs, wits) = mk_sel_case_args lthy4 ctr_Tss ks arg in (arg |> #2, wits, list_comb (mk_sel_casex arg, sel_rhs)) end; fun mk_dis_case_args args k = map (fn (k', arg) => (if k = k' then fold_rev Term.lambda arg \<^Const>\True\ else fold_rev Term.lambda arg \<^Const>\False\)) args; val sel_rhs = map (map mk_sel_rhs) sel_argss val dis_rhs = map (fn k => list_comb (dis_casex, mk_dis_case_args (ks ~~ xss) k)) ks val dis_qty = qty_isom --> HOLogic.boolT; val dis_names = map (fn k => Binding.qualify_name true uTname ("dis" ^ string_of_int k)) ks; val (diss, lthy5) = @{fold_map 2} (fn b => fn rhs => fn lthy => lift_def ld_no_notes (b, NoSyn) dis_qty rhs (K all_tac) [] lthy |>> mk_lift_const_of_lift_def dis_qty) dis_names dis_rhs lthy4 val unfold_lift_sel_rsp = @{lemma "(\x. P1 x \ P2 (f x)) \ (rel_fun (eq_onp P1) (eq_onp P2)) f f" by (simp add: eq_onp_same_args rel_fun_eq_onp_rel)} fun lift_sel_tac exhaust_rule dt_rules wits ctxt i = (Method.insert_tac ctxt wits THEN' eq_onp_to_top_tac ctxt THEN' (* normalize *) rtac ctxt unfold_lift_sel_rsp THEN' case_tac exhaust_rule ctxt THEN_ALL_NEW ( EVERY' [hyp_subst_tac ctxt, (* does not kill wits because = was rewritten to eq_onp top *) Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq dt_rules), REPEAT_DETERM o etac ctxt conjE, assume_tac ctxt])) i val pred_simps = Transfer.lookup_pred_data lthy5 (Tname rty) |> the |> Transfer.pred_simps val sel_tac = lift_sel_tac (#exhaust ctr_sugar) (#case_thms ctr_sugar @ pred_simps) val sel_names = map (fn (k, xs) => map (fn k' => Binding.qualify_name true uTname ("sel" ^ string_of_int k ^ string_of_int k')) (1 upto length xs)) (ks ~~ ctr_Tss); val (selss, lthy6) = @{fold_map 2} (@{fold_map 2} (fn b => fn (qty_ret, wits, rhs) => fn lthy => lift_def_code_dt { code_dt = true, lift_config = ld_no_notes } (b, NoSyn) (qty_isom --> qty_ret) rhs (HEADGOAL o sel_tac wits) [] lthy |>> mk_lift_const_of_lift_def (qty_isom --> qty_ret))) sel_names sel_rhs lthy5 (* now we can execute the qty qty_isom isomorphism *) fun mk_type_definition newT oldT RepC AbsC A = let val typedefC = Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in typedefC $ RepC $ AbsC $ A end; val typedef_goal = mk_type_definition qty_isom qty rep_isom abs_isom (HOLogic.mk_UNIV qty) |> HOLogic.mk_Trueprop; fun typ_isom_tac ctxt i = EVERY' [ SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms type_definition_def}), DETERM o Transfer.transfer_tac true ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms eq_onp_top_eq_eq}) (* normalize *), Raw_Simplifier.rewrite_goal_tac ctxt (map safe_mk_meta_eq @{thms id_apply simp_thms Ball_def}), rtac ctxt TrueI] i; val (_, transfer_ctxt) = Proof_Context.note_thms "" (Binding.empty_atts, [(@{thms right_total_UNIV_transfer}, [Transfer.transfer_add]), (@{thms Domain_eq_top}, [Transfer.transfer_domain_add])]) lthy6; val quot_thm_isom = Goal.prove_sorry transfer_ctxt [] [] typedef_goal (fn {context = goal_ctxt, ...} => typ_isom_tac goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export transfer_ctxt lthy6) |> (fn thm => @{thm UNIV_typedef_to_Quotient} OF [thm, @{thm reflexive}]) val qty_isom_name = Tname qty_isom; val quot_isom_rep = let val (quotients : Lifting_Term.quotients) = Symtab.insert (Lifting_Info.quotient_eq) (qty_isom_name, {quot_thm = quot_thm_isom, pcr_info = NONE}) Symtab.empty val id_actions = { constr = K I, lift = K I, comp_lift = K I } in fn ctxt => fn (rty, qty) => Lifting_Term.prove_schematic_quot_thm id_actions quotients ctxt (rty, qty) () |> fst |> Lifting_Term.force_qty_type ctxt qty |> quot_thm_rep end; val (x, x_ctxt) = yield_singleton (mk_Frees "x") qty_isom lthy6; fun mk_ctr ctr ctr_Ts sels = let val sel_ret_Ts = map (dest_Const #> snd #> body_type) sels; fun rep_isom lthy t (rty, qty) = let val rep = quot_isom_rep lthy (rty, qty) in if is_Const rep andalso (rep |> dest_Const |> fst) = \<^const_name>\id\ then t else rep $ t end; in @{fold 3} (fn sel => fn ctr_T => fn sel_ret_T => fn ctr => ctr $ rep_isom x_ctxt (sel $ x) (ctr_T, sel_ret_T)) sels ctr_Ts sel_ret_Ts ctr end; (* stolen from Metis *) exception BREAK_LIST fun break_list (x :: xs) = (x, xs) | break_list _ = raise BREAK_LIST val (ctr, ctrs) = qty_ctrs |> rev |> break_list; val (ctr_Ts, ctr_Tss) = qty_ctr_Tss |> rev |> break_list; val (sel, rselss) = selss |> rev |> break_list; val rdiss = rev diss |> tl; val first_ctr = mk_ctr ctr ctr_Ts sel; fun mk_If_ctr dis ctr ctr_Ts sel elsex = mk_If (dis$x) (mk_ctr ctr ctr_Ts sel) elsex; val rhs = @{fold 4} mk_If_ctr rdiss ctrs ctr_Tss rselss first_ctr; val rep_isom_code_goal = HOLogic.mk_Trueprop (HOLogic.mk_eq (rep_isom$x, rhs)); local val rep_isom_code_tac_rules = map safe_mk_meta_eq @{thms refl id_apply if_splits simp_thms} in fun rep_isom_code_tac (ctr_sugar:Ctr_Sugar.ctr_sugar) ctxt i = let val exhaust = ctr_sugar |> #exhaust val cases = ctr_sugar |> #case_thms val map_ids = fp |> #fp_nesting_bnfs |> map BNF_Def.map_id0_of_bnf val simp_rules = map safe_mk_meta_eq (cases @ map_ids) @ rep_isom_code_tac_rules in EVERY' [Transfer.gen_frees_tac [] ctxt, DETERM o (Transfer.transfer_tac true ctxt), case_tac exhaust ctxt THEN_ALL_NEW EVERY' [hyp_subst_tac ctxt, Raw_Simplifier.rewrite_goal_tac ctxt simp_rules, rtac ctxt TrueI ]] i end end (* stolen from bnf_fp_n2m.ML *) fun force_typ ctxt T = Term.map_types Type_Infer.paramify_vars #> Type.constraint T #> singleton (Type_Infer_Context.infer_types ctxt); (* The following tests that types in rty have corresponding arities imposed by constraints of the datatype fp. Otherwise rep_isom_code_tac could fail (especially transfer in it) is such a way that it is not easy to infer the problem with sorts. *) val _ = yield_singleton (mk_Frees "x") (#T fp) x_ctxt |> fst |> force_typ x_ctxt qty val rep_isom_code = Goal.prove_sorry x_ctxt [] [] rep_isom_code_goal (fn {context = goal_ctxt, ...} => rep_isom_code_tac ctr_sugar goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton(Variable.export x_ctxt lthy6) in lthy6 |> snd o Local_Theory.note ((Binding.empty, @{attributes [code]}), [rep_isom_code]) |> Lifting_Setup.lifting_forget pointer |> pair (selss, diss, rep_isom_code) end and constr qty (quot_thm, (lthy0, rel_eq_onps)) = let val quot_thm = Lifting_Term.force_qty_type lthy0 qty quot_thm val (rty, qty) = quot_thm_rty_qty quot_thm val rty_name = Tname rty; val pred_data = Transfer.lookup_pred_data lthy0 rty_name val pred_data = if is_some pred_data then the pred_data else error ("code_dt: " ^ quote rty_name ^ " is not a datatype.") val rel_eq_onp = safe_mk_meta_eq (Transfer.rel_eq_onp pred_data); val rel_eq_onps = insert Thm.eq_thm rel_eq_onp rel_eq_onps fun R_conv ctxt = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} ctxt then_conv Conv.rewr_conv rel_eq_onp val quot_thm = Conv.fconv_rule (HOLogic.Trueprop_conv (Quotient_R_conv (R_conv lthy0))) quot_thm; in if is_none (code_dt_of lthy0 (rty, qty)) then let val non_empty_pred = quot_thm RS @{thm type_definition_Quotient_not_empty} val pred = quot_thm_rel quot_thm |> dest_comb |> snd; val (pred, lthy1) = lthy0 |> (snd o Local_Theory.begin_nested) |> yield_singleton (Variable.import_terms true) pred; val TFrees = Term.add_tfreesT qty [] fun non_empty_typedef_tac non_empty_pred ctxt i = (Method.insert_tac ctxt [non_empty_pred] THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt [mem_Collect_eq]) THEN' assume_tac ctxt) i val uTname = unique_Tname (rty, qty) val Tdef_set = HOLogic.mk_Collect ("x", rty, pred $ Free("x", rty)); val ((_, tcode_dt), lthy2) = lthy1 |> conceal_naming_result (typedef (Binding.concealed uTname, TFrees, NoSyn) Tdef_set NONE (fn lthy => HEADGOAL (non_empty_typedef_tac non_empty_pred lthy))); val type_definition_thm = tcode_dt |> snd |> #type_definition; val qty_isom = tcode_dt |> fst |> #abs_type; val (binding, lthy3) = lthy2 |> conceal_naming_result (Lifting_Setup.setup_by_typedef_thm {notes = false} type_definition_thm) ||> Local_Theory.end_nested val (wit, wit_thm) = mk_witness quot_thm; val code_dt = mk_code_dt rty qty wit wit_thm NONE; val lthy4 = lthy3 |> update_code_dt code_dt |> mk_rep_isom binding (rty, qty, qty_isom) |> snd in (quot_thm, (lthy4, rel_eq_onps)) end else (quot_thm, (lthy0, rel_eq_onps)) end and lift_def_code_dt config = gen_lift_def (add_lift_def_code_dt config) (** from parsed parameters to the config record **) fun map_config_code_dt f1 f2 ({code_dt = code_dt, lift_config = lift_config}: config_code_dt) = {code_dt = f1 code_dt, lift_config = f2 lift_config} fun update_config_code_dt nval = map_config_code_dt (K nval) I val config_flags = [("code_dt", update_config_code_dt true)] fun evaluate_params params = let fun eval_param param config = case AList.lookup (op =) config_flags param of SOME update => update config | NONE => error ("Unknown parameter: " ^ (quote param)) in fold eval_param params default_config_code_dt end (** lift_definition command. It opens a proof of a corresponding respectfulness theorem in a user-friendly, readable form. Then add_lift_def_code_dt is called internally. **) local val eq_onp_assms_tac_fixed_rules = map (Transfer.prep_transfer_domain_thm \<^context>) [@{thm pcr_Domainp_total}, @{thm pcr_Domainp_par_left_total}, @{thm pcr_Domainp_par}, @{thm pcr_Domainp}] in fun mk_readable_rsp_thm_eq tm ctxt = let val ctm = Thm.cterm_of ctxt tm fun assms_rewr_conv tactic rule ct = let fun prove_extra_assms thm = let val assms = cprems_of thm fun finish thm = if Thm.no_prems thm then SOME (Goal.conclude thm) else NONE fun prove ctm = Option.mapPartial finish (SINGLE tactic (Goal.init ctm)) in map_interrupt prove assms end fun cconl_of thm = Drule.strip_imp_concl (Thm.cprop_of thm) fun lhs_of thm = fst (Thm.dest_equals (cconl_of thm)) fun rhs_of thm = snd (Thm.dest_equals (cconl_of thm)) val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule; val lhs = lhs_of rule1; val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1; val rule3 = Thm.instantiate (Thm.match (lhs, ct)) rule2 handle Pattern.MATCH => raise CTERM ("assms_rewr_conv", [lhs, ct]); val proved_assms = prove_extra_assms rule3 in case proved_assms of SOME proved_assms => let val rule3 = proved_assms MRSL rule3 val rule4 = if lhs_of rule3 aconvc ct then rule3 else let val ceq = Thm.dest_fun2 (Thm.cprop_of rule3) in rule3 COMP Thm.trivial (Thm.mk_binop ceq ct (rhs_of rule3)) end in Thm.transitive rule4 (Thm.beta_conversion true (rhs_of rule4)) end | NONE => Conv.no_conv ct end fun assms_rewrs_conv tactic rules = Conv.first_conv (map (assms_rewr_conv tactic) rules) fun simp_arrows_conv ctm = let val unfold_conv = Conv.rewrs_conv [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_onp_rel[THEN eq_reflection]}, @{thm rel_fun_eq[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, @{thm rel_fun_def[THEN eq_reflection]}] fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 val eq_onp_assms_tac_rules = @{thm left_unique_OO} :: eq_onp_assms_tac_fixed_rules @ (Transfer.get_transfer_raw ctxt) val intro_top_rule = @{thm eq_onp_top_eq_eq[symmetric, THEN eq_reflection]} val kill_tops = Conv.top_sweep_rewrs_conv @{thms eq_onp_top_eq_eq[THEN eq_reflection]} ctxt val eq_onp_assms_tac = (CONVERSION kill_tops THEN' TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_onp_assms_tac_rules) THEN_ALL_NEW (DETERM o Transfer.eq_tac ctxt)) 1 val relator_eq_onp_conv = Conv.bottom_conv (K (Conv.try_conv (assms_rewrs_conv eq_onp_assms_tac (intro_top_rule :: Lifting_Info.get_relator_eq_onp_rules ctxt)))) ctxt then_conv kill_tops val relator_eq_conv = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv (Transfer.get_relator_eq ctxt)))) ctxt in case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => (binop_conv2 simp_arrows_conv simp_arrows_conv then_conv unfold_conv) ctm | _ => (relator_eq_onp_conv then_conv relator_eq_conv) ctm end val unfold_ret_val_invs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt val unfold_inv_conv = Conv.top_sweep_rewrs_conv @{thms eq_onp_def[THEN eq_reflection]} ctxt val simp_conv = HOLogic.Trueprop_conv (Conv.fun2_conv simp_arrows_conv) val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt val beta_conv = Thm.beta_conversion true val eq_thm = (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs then_conv unfold_inv_conv) ctm in Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) end end fun rename_to_tnames ctxt term = let fun all_typs (Const (\<^const_name>\Pure.all\, _) $ Abs (_, T, t)) = T :: all_typs t | all_typs _ = [] fun rename (Const (\<^const_name>\Pure.all\, T1) $ Abs (_, T2, t)) (new_name :: names) = (Const (\<^const_name>\Pure.all\, T1) $ Abs (new_name, T2, rename t names)) | rename t _ = t val (fixed_def_t, _) = yield_singleton (Variable.importT_terms) term ctxt val new_names = Old_Datatype_Prop.make_tnames (all_typs fixed_def_t) in rename term new_names end fun quot_thm_err ctxt (rty, qty) pretty_msg = let val error_msg = cat_lines ["Lifting failed for the following types:", Pretty.string_of (Pretty.block [Pretty.str "Raw type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty]), Pretty.string_of (Pretty.block [Pretty.str "Abstract type:", Pretty.brk 2, Syntax.pretty_typ ctxt qty]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, pretty_msg]))] in error error_msg end fun check_rty_err ctxt (rty_schematic, rty_forced) (raw_var, rhs_raw) = let val (_, ctxt') = Proof_Context.read_var raw_var ctxt val rhs = Syntax.read_term ctxt' rhs_raw val error_msg = cat_lines ["Lifting failed for the following term:", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt rhs]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt rty_schematic]), "", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, Pretty.str "The type of the term cannot be instantiated to", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt rty_forced), Pretty.str "."]))] in error error_msg end fun lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy0 = let val config = evaluate_params params val ((binding, SOME qty, mx), lthy1) = Proof_Context.read_var raw_var lthy0 val var = (binding, mx) val rhs = Syntax.read_term lthy1 rhs_raw val par_thms = Attrib.eval_thms lthy1 par_xthms val (goal, after_qed) = lthy1 |> prepare_lift_def (add_lift_def_code_dt config) var qty rhs par_thms val (goal, after_qed) = case goal of NONE => (goal, K (after_qed Drule.dummy_thm)) | SOME prsp_tm => let val readable_rsp_thm_eq = mk_readable_rsp_thm_eq prsp_tm lthy1 val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) val readable_rsp_tm_tnames = rename_to_tnames lthy1 readable_rsp_tm fun after_qed' [[thm]] lthy = let val internal_rsp_thm = Goal.prove lthy [] [] prsp_tm (fn {context = goal_ctxt, ...} => rtac goal_ctxt readable_rsp_thm_eq 1 THEN Proof_Context.fact_tac goal_ctxt [thm] 1) in after_qed internal_rsp_thm lthy end in (SOME readable_rsp_tm_tnames, after_qed') end fun after_qed_with_err_handling thmss ctxt = (after_qed thmss ctxt handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy1 (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy1 (rty_schematic, rty_forced) (raw_var, rhs_raw); in lthy1 |> Proof.theorem NONE (snd oo after_qed_with_err_handling) [map (rpair []) (the_list goal)] end fun lift_def_cmd_with_err_handling (params, (raw_var, rhs_raw, par_xthms)) lthy = (lift_def_cmd (params, raw_var, rhs_raw, par_xthms) lthy handle Lifting_Term.QUOT_THM (rty, qty, msg) => quot_thm_err lthy (rty, qty) msg) handle Lifting_Term.CHECK_RTY (rty_schematic, rty_forced) => check_rty_err lthy (rty_schematic, rty_forced) (raw_var, rhs_raw); val parse_param = Parse.name val parse_params = Scan.optional (Args.parens (Parse.list parse_param)) []; (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\lift_definition\ "definition for constants over the quotient type" (parse_params -- (((Parse.binding -- (\<^keyword>\::\ |-- (Parse.typ >> SOME) -- Parse.opt_mixfix') >> Scan.triple2) -- (\<^keyword>\is\ |-- Parse.term) -- Scan.optional (\<^keyword>\parametric\ |-- Parse.!!! Parse.thms1) []) >> Scan.triple1) >> lift_def_cmd_with_err_handling); end diff --git a/src/HOL/Tools/Lifting/lifting_setup.ML b/src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML +++ b/src/HOL/Tools/Lifting/lifting_setup.ML @@ -1,1046 +1,1046 @@ (* Title: HOL/Tools/Lifting/lifting_setup.ML Author: Ondrej Kuncar Setting up the lifting infrastructure. *) signature LIFTING_SETUP = sig exception SETUP_LIFTING_INFR of string type config = { notes: bool }; val default_config: config; val setup_by_quotient: config -> thm -> thm option -> thm option -> local_theory -> binding * local_theory val setup_by_typedef_thm: config -> thm -> local_theory -> binding * local_theory val lifting_restore: Lifting_Info.quotient -> Context.generic -> Context.generic val lifting_forget: string -> local_theory -> local_theory val update_transfer_rules: string -> local_theory -> local_theory val pointer_of_bundle_binding: Proof.context -> binding -> string end structure Lifting_Setup: LIFTING_SETUP = struct open Lifting_Util infix 0 MRSL exception SETUP_LIFTING_INFR of string (* Config *) type config = { notes: bool }; val default_config = { notes = true }; fun define_crel (config: config) rep_fun lthy = let val (qty, rty) = (dest_funT o fastype_of) rep_fun val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0) val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph)) val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val crel_name = Binding.prefix_name "cr_" qty_name val (fixed_def_term, lthy1) = lthy |> yield_singleton (Variable.importT_terms) def_term val ((_, (_ , def_thm)), lthy2) = if #notes config then Local_Theory.define ((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy1 else Local_Theory.define ((Binding.concealed crel_name, NoSyn), (Binding.empty_atts, fixed_def_term)) lthy1 in (def_thm, lthy2) end fun print_define_pcrel_warning msg = let val warning_msg = cat_lines ["Generation of a parametrized correspondence relation failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in warning warning_msg end fun define_pcrel (config: config) crel lthy0 = let val (fixed_crel, lthy1) = yield_singleton Variable.importT_terms crel lthy0 val [rty', qty] = (binder_types o fastype_of) fixed_crel val (param_rel, args) = Lifting_Term.generate_parametrized_relator lthy1 rty' val rty_raw = (domain_type o range_type o fastype_of) param_rel val tyenv_match = Sign.typ_match (Proof_Context.theory_of lthy1) (rty_raw, rty') Vartab.empty val param_rel_subst = Envir.subst_term (tyenv_match,Vartab.empty) param_rel val args_subst = map (Envir.subst_term (tyenv_match,Vartab.empty)) args val (instT, lthy2) = lthy1 |> Variable.declare_names fixed_crel |> Variable.importT_inst (param_rel_subst :: args_subst) val args_fixed = (map (Term_Subst.instantiate (instT, Vars.empty))) args_subst val param_rel_fixed = Term_Subst.instantiate (instT, Vars.empty) param_rel_subst val rty = (domain_type o fastype_of) param_rel_fixed val relcomp_op = Const (\<^const_name>\relcompp\, (rty --> rty' --> HOLogic.boolT) --> (rty' --> qty --> HOLogic.boolT) --> rty --> qty --> HOLogic.boolT) val qty_name = (fst o dest_Type) qty val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name) val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT]) val lhs = Library.foldl (op $) ((Free (Binding.name_of pcrel_name, relator_type)), args_fixed) val rhs = relcomp_op $ param_rel_fixed $ fixed_crel val definition_term = Logic.mk_equals (lhs, rhs) fun note_def lthy = Specification.definition (SOME (pcrel_name, SOME relator_type, NoSyn)) [] [] (Binding.empty_atts, definition_term) lthy |>> (snd #> snd); fun raw_def lthy = let val ((_, rhs), prove) = Local_Defs.derived_def lthy (K []) {conditional = true} definition_term; val ((_, (_, raw_th)), lthy') = Local_Theory.define ((Binding.concealed pcrel_name, NoSyn), (Binding.empty_atts, rhs)) lthy; val th = prove lthy' raw_th; in (th, lthy') end val (def_thm, lthy3) = if #notes config then note_def lthy2 else raw_def lthy2 in (SOME def_thm, lthy3) end handle Lifting_Term.PARAM_QUOT_THM (_, msg) => (print_define_pcrel_warning msg; (NONE, lthy0)) local val eq_OO_meta = mk_meta_eq @{thm eq_OO} fun print_generate_pcr_cr_eq_error ctxt term = let val goal = Const (\<^const_name>\HOL.eq\, dummyT) $ term $ Const (\<^const_name>\HOL.eq\, dummyT) val error_msg = cat_lines ["Generation of a pcr_cr_eq failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])), "Most probably a relator_eq rule for one of the involved types is missing."] in error error_msg end in fun define_pcr_cr_eq (config: config) lthy pcr_rel_def = let val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type o List.last o binder_types o fastype_of) lhs val args = (snd o strip_comb) lhs fun make_inst var ctxt = let val typ = snd (relation_types (#2 (dest_Var var))) val sort = Type.sort_of_atyp typ val (fresh_var, ctxt') = yield_singleton Variable.invent_types sort ctxt val inst = (#1 (dest_Var var), Thm.cterm_of ctxt' (HOLogic.eq_const (TFree fresh_var))) in (inst, ctxt') end val (args_inst, args_ctxt) = fold_map make_inst args lthy val pcr_cr_eq = pcr_rel_def |> infer_instantiate args_ctxt args_inst |> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.bottom_rewrs_conv (Transfer.get_relator_eq args_ctxt) args_ctxt))) in case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of Const (\<^const_name>\relcompp\, _) $ Const (\<^const_name>\HOL.eq\, _) $ _ => let val thm = pcr_cr_eq |> Conv.fconv_rule (Conv.arg_conv (Conv.rewr_conv eq_OO_meta)) |> HOLogic.mk_obj_eq |> singleton (Variable.export args_ctxt lthy) val lthy' = lthy |> #notes config ? (Local_Theory.note ((Binding.qualify_name true qty_name "pcr_cr_eq", []), [thm]) #> #2) in (thm, lthy') end | Const (\<^const_name>\relcompp\, _) $ t $ _ => print_generate_pcr_cr_eq_error args_ctxt t | _ => error "generate_pcr_cr_eq: implementation error" end end fun define_code_constr quot_thm lthy = let val abs = quot_thm_abs quot_thm in if is_Const abs then let val (fixed_abs, lthy') = yield_singleton Variable.importT_terms abs lthy in Local_Theory.background_theory (Code.declare_datatype_global [dest_Const fixed_abs]) lthy' end else lthy end fun define_abs_type quot_thm = Lifting_Def.can_generate_code_cert quot_thm ? Code.declare_abstype (quot_thm RS @{thm Quotient_abs_rep}); local exception QUOT_ERROR of Pretty.T list in fun quot_thm_sanity_check ctxt quot_thm = let val _ = if (Thm.nprems_of quot_thm > 0) then raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem has extra assumptions:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] else () val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient handle TERM _ => raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem is not of the right form:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt val (rty, qty) = quot_thm_rty_qty quot_thm_fixed val rty_tfreesT = Term.add_tfree_namesT rty [] val qty_tfreesT = Term.add_tfree_namesT qty [] val extra_rty_tfrees = case subtract (op =) qty_tfreesT rty_tfreesT of [] => [] | extras => [Pretty.block ([Pretty.str "Extra variables in the raw type:", Pretty.brk 1] @ ((Pretty.commas o map (Pretty.str o quote)) extras) @ [Pretty.str "."])] val not_type_constr = case qty of Type _ => [] | _ => [Pretty.block [Pretty.str "The quotient type ", Pretty.quote (Syntax.pretty_typ ctxt' qty), Pretty.brk 1, Pretty.str "is not a type constructor."]] val errs = extra_rty_tfrees @ not_type_constr in if null errs then () else raise QUOT_ERROR errs end handle QUOT_ERROR errs => error (cat_lines (["Sanity check of the quotient theorem failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end fun lifting_bundle qty_full_name qinfo lthy = let val thy = Proof_Context.theory_of lthy val binding = Binding.qualify_name true (qty_full_name |> Long_Name.base_name |> Binding.name) "lifting" val morphed_binding = Morphism.binding (Local_Theory.target_morphism lthy) binding val bundle_name = Name_Space.full_name (Name_Space.naming_of (Context.Theory thy)) morphed_binding fun phi_qinfo phi = Lifting_Info.transform_quotient phi qinfo val dummy_thm = Thm.transfer thy Drule.dummy_thm val restore_lifting_att = ([dummy_thm], [map Token.make_string0 ["Lifting.lifting_restore_internal", bundle_name]]) in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi)) |> Bundle.bundle ((binding, [restore_lifting_att])) [] |> pair binding end fun setup_lifting_infr config quot_thm opt_reflp_thm lthy = let val _ = quot_thm_sanity_check lthy quot_thm val (_, qty) = quot_thm_rty_qty quot_thm val (pcrel_def, lthy1) = define_pcrel config (quot_thm_crel quot_thm) lthy (**) val pcrel_def = Option.map (Morphism.thm (Local_Theory.target_morphism lthy1)) pcrel_def (**) val (pcr_cr_eq, lthy2) = case pcrel_def of SOME pcrel_def => apfst SOME (define_pcr_cr_eq config lthy1 pcrel_def) | NONE => (NONE, lthy1) val pcr_info = case pcrel_def of SOME pcrel_def => SOME { pcrel_def = pcrel_def, pcr_cr_eq = the pcr_cr_eq } | NONE => NONE val quotients = { quot_thm = quot_thm, pcr_info = pcr_info } val qty_full_name = (fst o dest_Type) qty fun quot_info phi = Lifting_Info.transform_quotient phi quotients - val reflexivity_rule_attr = Attrib.internal (K Lifting_Info.add_reflexivity_rule_attribute) + val reflexivity_rule_attr = Attrib.internal \<^here> (K Lifting_Info.add_reflexivity_rule_attribute) val lthy3 = case opt_reflp_thm of SOME reflp_thm => lthy2 |> (#2 oo Local_Theory.note) ((Binding.empty, [reflexivity_rule_attr]), [reflp_thm RS @{thm reflp_ge_eq}]) |> define_code_constr quot_thm | NONE => lthy2 |> define_abs_type quot_thm in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) |> lifting_bundle qty_full_name quotients end local fun importT_inst_exclude exclude ts ctxt = let val tvars = rev (subtract op= exclude (fold Term.add_tvars ts [])) val (tfrees, ctxt') = Variable.invent_types (map #2 tvars) ctxt in (TVars.make (tvars ~~ map TFree tfrees), ctxt') end fun import_inst_exclude exclude ts ctxt = let val excludeT = fold (Term.add_tvarsT o snd) exclude [] val (instT, ctxt') = importT_inst_exclude excludeT ts ctxt val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (subtract op= exclude (fold Term.add_vars ts []))) val (xs, ctxt'') = Variable.variant_fixes (map (#1 o #1) vars) ctxt' val inst = Vars.make (vars ~~ map Free (xs ~~ map #2 vars)) in ((instT, inst), ctxt'') end fun import_terms_exclude exclude ts ctxt = let val (inst, ctxt') = import_inst_exclude exclude ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end in fun reduce_goal not_fix goal tac ctxt = let val (fixed_goal, ctxt') = yield_singleton (import_terms_exclude not_fix) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE tac init_goal)) end end local val OO_rules = @{thms left_total_OO left_unique_OO right_total_OO right_unique_OO bi_total_OO bi_unique_OO} in fun parametrize_class_constraint ctxt0 pcr_def constraint = let fun generate_transfer_rule pcr_def constraint goal ctxt = let val (fixed_goal, ctxt') = yield_singleton (Variable.import_terms true) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) val rules = Transfer.get_transfer_raw ctxt' val rules = constraint :: OO_rules @ rules val tac = K (Local_Defs.unfold0_tac ctxt' [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt' rules) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal)) end fun make_goal pcr_def constr = let val pred_name = (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.prop_of) constr val arg = (fst o Logic.dest_equals o Thm.prop_of) pcr_def in HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg) end val check_assms = let val right_names = ["right_total", "right_unique", "left_total", "left_unique", "bi_total", "bi_unique"] fun is_right_name name = member op= right_names (Long_Name.base_name name) fun is_trivial_assm (Const (name, _) $ Var (_, _)) = is_right_name name | is_trivial_assm (Const (name, _) $ Free (_, _)) = is_right_name name | is_trivial_assm _ = false in fn thm => let val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm) val thm_name = (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm val non_trivial_assms = filter_out is_trivial_assm prems in if null non_trivial_assms then () else Pretty.block ([Pretty.str "Non-trivial assumptions in ", Pretty.str thm_name, Pretty.str " transfer rule found:", Pretty.brk 1] @ Pretty.commas (map (Syntax.pretty_term ctxt0) non_trivial_assms)) |> Pretty.string_of |> warning end end val goal = make_goal pcr_def constraint val thm = generate_transfer_rule pcr_def constraint goal ctxt0 val _ = check_assms thm in thm end end local val id_unfold = (Conv.rewr_conv (mk_meta_eq @{thm id_def})) in fun generate_parametric_id lthy rty id_transfer_rule = let (* it doesn't raise an exception because it would have already raised it in define_pcrel *) val (quot_thm, _, ctxt') = Lifting_Term.prove_param_quot_thm lthy rty val parametrized_relator = singleton (Variable.export_terms ctxt' lthy) (quot_thm_crel quot_thm) val id_transfer = @{thm id_transfer} |> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1) |> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold)) val var = hd (Term.add_vars (Thm.prop_of id_transfer) []) val inst = [(#1 var, Thm.cterm_of lthy parametrized_relator)] val id_par_thm = infer_instantiate lthy inst id_transfer in Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm end handle Lifting_Term.MERGE_TRANSFER_REL msg => let val error_msg = cat_lines ["Generation of a parametric transfer rule for the abs. or the rep. function failed.", "A non-parametric version will be used.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in (warning error_msg; id_transfer_rule) end end local fun rewrite_first_Domainp_arg rewr_thm thm = Conv.fconv_rule (Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv rewr_thm))))) thm fun fold_Domainp_pcrel pcrel_def thm = let val ct = thm |> Thm.cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg val pcrel_def = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) pcrel_def val thm' = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def]) in rewrite_first_Domainp_arg (Thm.symmetric pcrel_def) thm' end fun reduce_Domainp ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var val reduced_assm = reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end in fun parametrize_domain dom_thm (pcr_info : Lifting_Info.pcr) ctxt0 = let fun reduce_first_assm ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val reduced_assm = reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end val pcr_cr_met_eq = #pcr_cr_eq pcr_info RS @{thm eq_reflection} val pcr_Domainp_eq = rewrite_first_Domainp_arg (Thm.symmetric pcr_cr_met_eq) dom_thm val pcrel_def = #pcrel_def pcr_info val pcr_Domainp_par_left_total = (dom_thm RS @{thm pcr_Domainp_par_left_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_first_assm ctxt0 (Lifting_Info.get_reflexivity_rules ctxt0) val pcr_Domainp_par = (dom_thm RS @{thm pcr_Domainp_par}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt0 (Transfer.get_relator_domain ctxt0) val pcr_Domainp = (dom_thm RS @{thm pcr_Domainp}) |> fold_Domainp_pcrel pcrel_def val thms = [("domain", [pcr_Domainp], @{attributes [transfer_domain_rule]}), ("domain_par", [pcr_Domainp_par], @{attributes [transfer_domain_rule]}), ("domain_par_left_total", [pcr_Domainp_par_left_total], @{attributes [transfer_domain_rule]}), ("domain_eq", [pcr_Domainp_eq], @{attributes [transfer_domain_rule]})] in thms end fun parametrize_total_domain left_total pcrel_def ctxt = let val thm = (left_total RS @{thm pcr_Domainp_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt) in [("domain", [thm], @{attributes [transfer_domain_rule]})] end end fun get_pcrel_info ctxt qty_full_name = #pcr_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name)) fun get_Domainp_thm quot_thm = the (get_first (try(curry op RS quot_thm)) [@{thm eq_onp_to_Domainp}, @{thm Quotient_to_Domainp}]) fun notes names thms = let val notes = if names then map (fn (name, thms, attrs) => ((name, []), [(thms, attrs)])) thms else map_filter (fn (_, thms, attrs) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) thms in Local_Theory.notes notes #> snd end fun map_thms map_name map_thm thms = map (fn (name, thms, attr) => (map_name name, map map_thm thms, attr)) thms (* Sets up the Lifting package by a quotient theorem. quot_thm - a quotient theorem (Quotient R Abs Rep T) opt_reflp_thm - a theorem saying that a relation from quot_thm is reflexive (in the form "reflp R") opt_par_thm - a parametricity theorem for R *) fun setup_by_quotient (config: config) quot_thm opt_reflp_thm opt_par_thm lthy0 = let (**) val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy0) quot_thm (**) val (rty, qty) = quot_thm_rty_qty quot_thm - val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty)))) + val induct_attr = Attrib.internal \<^here> (K (Induct.induct_type (fst (dest_Type qty)))) val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("abs_induct", @{thms Quotient_total_abs_induct}, [induct_attr]), ("abs_eq_iff", @{thms Quotient_total_abs_eq_iff}, [])] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => let val thms = [("abs_induct", @{thms Quotient_abs_induct}, [induct_attr])] in map_thms qualify (fn thm => quot_thm RS thm) thms end val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val notes2 = map_thms qualify (fn thm => quot_thm RS thm) [("rel_eq_transfer", @{thms Quotient_rel_eq_transfer}, @{attributes [transfer_rule]}), ("right_unique", @{thms Quotient_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms Quotient_right_total}, @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun generate_parametric_rel_eq ctxt transfer_rule opt_param_thm = (case opt_param_thm of NONE => transfer_rule | SOME param_thm => (Lifting_Def.generate_parametric_transfer_rule ctxt transfer_rule param_thm handle Lifting_Term.MERGE_TRANSFER_REL msg => error ("Generation of a parametric transfer rule for the quotient relation failed:\n" ^ Pretty.string_of msg))) fun setup_transfer_rules_par ctxt notes = let val pcrel_info = the (get_pcrel_info ctxt qty_full_name) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val thms = [("id_abs_transfer", [id_abs_transfer], @{attributes [transfer_rule]}), ("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val rel_eq_transfer = generate_parametric_rel_eq ctxt (Lifting_Term.parametrize_transfer_rule ctxt (quot_thm RS @{thm Quotient_rel_eq_transfer})) opt_par_thm val right_unique = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_unique}) val right_total = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_total}) val notes2 = map_thms qualify I [("rel_eq_transfer", [rel_eq_transfer], @{attributes [transfer_rule]}), ("right_unique", [right_unique], @{attributes [transfer_rule]}), ("right_total", [right_total], @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy0 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end (* Sets up the Lifting package by a typedef theorem. gen_code - flag if an abstract type given by typedef_thm should be registred as an abstract type in the code generator typedef_thm - a typedef theorem (type_definition Rep Abs S) *) fun setup_by_typedef_thm config typedef_thm lthy0 = let val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm val (T_def, lthy1) = define_crel config rep_fun lthy0 (**) val T_def = Morphism.thm (Local_Theory.target_morphism lthy1) T_def (**) val quot_thm = case typedef_set of Const (\<^const_name>\top\, _) => [typedef_thm, T_def] MRSL @{thm UNIV_typedef_to_Quotient} | Const (\<^const_name>\Collect\, _) $ Abs (_, _, _) => [typedef_thm, T_def] MRSL @{thm open_typedef_to_Quotient} | _ => [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient} val (rty, qty) = quot_thm_rty_qty quot_thm val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val opt_reflp_thm = case typedef_set of Const (\<^const_name>\top\, _) => SOME ((typedef_thm RS @{thm UNIV_typedef_to_equivp}) RS @{thm equivp_reflp2}) | _ => NONE val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val thms = [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]}), ("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [typedef_thm, T_def] MRSL thm) thms @ notes1 @ notes end fun setup_transfer_rules_par ctxt notes = let val pcrel_info = (the (get_pcrel_info ctxt qty_full_name)) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val thms = [("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]}), ("id_abs_transfer",[id_abs_transfer], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val notes2 = map_thms qualify (fn thm => generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([typedef_thm, T_def] MRSL thm))) [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]})]; val notes3 = map_thms qualify (fn thm => parametrize_class_constraint ctxt pcrel_def ([typedef_thm, T_def] MRSL thm)) [("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique},@{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]})] in notes3 @ notes2 @ notes1 @ notes end val notes1 = [(Binding.prefix_name "Quotient_" qty_name, [quot_thm], [])] fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy1 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy = let val input_thm = singleton (Attrib.eval_thms lthy) xthm val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." fun sanity_check_reflp_thm reflp_thm = let val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." in case reflp_tm of \<^Const_>\reflp_on _ for \<^Const>\top_class.top _\ _\ => () | _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." end fun check_qty qty = if not (is_Type qty) then error "The abstract type must be a type constructor." else () fun setup_quotient () = let val opt_reflp_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_reflp_xthm val _ = if is_some opt_reflp_thm then sanity_check_reflp_thm (the opt_reflp_thm) else () val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm val _ = check_qty (snd (quot_thm_rty_qty input_thm)) in setup_by_quotient default_config input_thm opt_reflp_thm opt_par_thm lthy |> snd end fun setup_typedef () = let val qty = (range_type o fastype_of o hd o get_args 2) input_term val _ = check_qty qty in case opt_reflp_xthm of SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." | NONE => ( case opt_par_xthm of SOME _ => error "The parametricity theorem cannot be specified if the type_definition theorem is used." | NONE => setup_by_typedef_thm default_config input_thm lthy |> snd ) end in case input_term of (Const (\<^const_name>\Quotient\, _) $ _ $ _ $ _ $ _) => setup_quotient () | (Const (\<^const_name>\type_definition\, _) $ _ $ _ $ _) => setup_typedef () | _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." end val _ = Outer_Syntax.local_theory \<^command_keyword>\setup_lifting\ "setup lifting infrastructure" (Parse.thm -- Scan.option Parse.thm -- Scan.option (\<^keyword>\parametric\ |-- Parse.!!! Parse.thm) >> (fn ((xthm, opt_reflp_xthm), opt_par_xthm) => setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm)) (* restoring lifting infrastructure *) local exception PCR_ERROR of Pretty.T list in fun lifting_restore_sanity_check ctxt (qinfo:Lifting_Info.quotient) = let val quot_thm = (#quot_thm qinfo) val _ = quot_thm_sanity_check ctxt quot_thm val pcr_info_err = (case #pcr_info qinfo of SOME pcr => let val pcrel_def = #pcrel_def pcr val pcr_cr_eq = #pcr_cr_eq pcr val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr definiton theorem is not a plain meta equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcrel_def]] val pcr_const_def = head_of def_lhs val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq)) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr_cr equation theorem is not a plain equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] val (pcr_const_eq, eqs) = strip_comb eq_lhs fun is_eq (Const (\<^const_name>\HOL.eq\, _)) = true | is_eq _ = false fun eq_Const (Const (name1, _)) (Const (name2, _)) = (name1 = name2) | eq_Const _ _ = false val all_eqs = if not (forall is_eq eqs) then [Pretty.block [Pretty.str "Arguments of the lhs of the pcr_cr equation theorem are not only equalities:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] else [] val pcr_consts_not_equal = if not (eq_Const pcr_const_def pcr_const_eq) then [Pretty.block [Pretty.str "Parametrized correspondence relation constants in pcr_def and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_def, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_eq]] else [] val crel = quot_thm_crel quot_thm val cr_consts_not_equal = if not (eq_Const crel eq_rhs) then [Pretty.block [Pretty.str "Correspondence relation constants in the Quotient theorem and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt crel, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt eq_rhs]] else [] in all_eqs @ pcr_consts_not_equal @ cr_consts_not_equal end | NONE => []) val errs = pcr_info_err in if null errs then () else raise PCR_ERROR errs end handle PCR_ERROR errs => error (cat_lines (["Sanity check failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end (* Registers the data in qinfo in the Lifting infrastructure. *) fun lifting_restore qinfo ctxt = let val _ = lifting_restore_sanity_check (Context.proof_of ctxt) qinfo val (_, qty) = quot_thm_rty_qty (#quot_thm qinfo) val qty_full_name = (fst o dest_Type) qty val stored_qinfo = Lifting_Info.lookup_quotients (Context.proof_of ctxt) qty_full_name in if is_some (stored_qinfo) andalso not (Lifting_Info.quotient_eq (qinfo, (the stored_qinfo))) then error (Pretty.string_of (Pretty.block [Pretty.str "Lifting is already setup for the type", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Context.proof_of ctxt) qty)])) else Lifting_Info.update_quotients qty_full_name qinfo ctxt end val parse_opt_pcr = Scan.optional (Attrib.thm -- Attrib.thm >> (fn (pcrel_def, pcr_cr_eq) => SOME {pcrel_def = pcrel_def, pcr_cr_eq = pcr_cr_eq})) NONE val lifting_restore_attribute_setup = Attrib.setup \<^binding>\lifting_restore\ ((Attrib.thm -- parse_opt_pcr) >> (fn (quot_thm, opt_pcr) => let val qinfo = { quot_thm = quot_thm, pcr_info = opt_pcr} in Thm.declaration_attribute (K (lifting_restore qinfo)) end)) "restoring lifting infrastructure" val _ = Theory.setup lifting_restore_attribute_setup fun lifting_restore_internal bundle_name ctxt = let val restore_info = Lifting_Info.lookup_restore_data (Context.proof_of ctxt) bundle_name in case restore_info of SOME restore_info => ctxt |> lifting_restore (#quotient restore_info) |> fold_rev Transfer.transfer_raw_add (Item_Net.content (#transfer_rules restore_info)) | NONE => ctxt end val lifting_restore_internal_attribute_setup = Attrib.setup \<^binding>\lifting_restore_internal\ (Scan.lift Parse.string >> (fn name => Thm.declaration_attribute (K (lifting_restore_internal name)))) "restoring lifting infrastructure; internal attribute; not meant to be used directly by regular users" val _ = Theory.setup lifting_restore_internal_attribute_setup (* lifting_forget *) val monotonicity_names = [\<^const_name>\right_unique\, \<^const_name>\left_unique\, \<^const_name>\right_total\, \<^const_name>\left_total\, \<^const_name>\bi_unique\, \<^const_name>\bi_total\] fun fold_transfer_rel f (Const (\<^const_name>\Transfer.Rel\, _) $ rel $ _ $ _) = f rel | fold_transfer_rel f (Const (\<^const_name>\HOL.eq\, _) $ (Const (\<^const_name>\Domainp\, _) $ rel) $ _) = f rel | fold_transfer_rel f (Const (name, _) $ rel) = if member op= monotonicity_names name then f rel else f \<^term>\undefined\ | fold_transfer_rel f _ = f \<^term>\undefined\ fun filter_transfer_rules_by_rel transfer_rel transfer_rules = let val transfer_rel_name = transfer_rel |> dest_Const |> fst; fun has_transfer_rel thm = let val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop in member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name end handle TERM _ => false in filter has_transfer_rel transfer_rules end type restore_data = {quotient : Lifting_Info.quotient, transfer_rules: thm Item_Net.T} fun get_transfer_rel (qinfo : Lifting_Info.quotient) = let fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of in if is_some (#pcr_info qinfo) then get_pcrel (#pcrel_def (the (#pcr_info qinfo))) else quot_thm_crel (#quot_thm qinfo) end fun pointer_of_bundle_name bundle_name ctxt = let val bundle = Bundle.read ctxt bundle_name fun err () = error "The provided bundle is not a lifting bundle" in (case bundle of [(_, [arg_src])] => let val (name, _) = Token.syntax (Scan.lift Parse.string) arg_src ctxt handle ERROR _ => err () in name end | _ => err ()) end fun pointer_of_bundle_binding ctxt binding = Name_Space.full_name (Name_Space.naming_of (Context.Theory (Proof_Context.theory_of ctxt))) binding fun lifting_forget pointer lthy = let fun get_transfer_rules_to_delete qinfo ctxt = let val transfer_rel = get_transfer_rel qinfo in filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw ctxt) end in case Lifting_Info.lookup_restore_data lthy pointer of SOME restore_info => let val qinfo = #quotient restore_info val quot_thm = #quot_thm qinfo val transfer_rules = get_transfer_rules_to_delete qinfo lthy in - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm)) lthy end | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_forget_cmd bundle_name lthy = lifting_forget (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_forget\ "unsetup Lifting and Transfer for the given lifting bundle" (Parse.name_position >> lifting_forget_cmd) (* lifting_update *) fun update_transfer_rules pointer lthy = let fun new_transfer_rules ({ quotient = qinfo, ... }:Lifting_Info.restore_data) lthy = let val transfer_rel = get_transfer_rel qinfo val transfer_rules = filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw lthy) in fn phi => fold_rev (Item_Net.update o Morphism.thm phi) transfer_rules Thm.item_net end in case Lifting_Info.lookup_restore_data lthy pointer of SOME refresh_data => - Local_Theory.declaration {syntax = false, pervasive = true} + Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer (new_transfer_rules refresh_data lthy phi)) lthy | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_update_cmd bundle_name lthy = update_transfer_rules (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_update\ "add newly introduced transfer rules to a bundle storing the state of Lifting and Transfer" (Parse.name_position >> lifting_update_cmd) end diff --git a/src/HOL/Tools/Quotient/quotient_def.ML b/src/HOL/Tools/Quotient/quotient_def.ML --- a/src/HOL/Tools/Quotient/quotient_def.ML +++ b/src/HOL/Tools/Quotient/quotient_def.ML @@ -1,210 +1,210 @@ (* Title: HOL/Tools/Quotient/quotient_def.ML Author: Cezary Kaliszyk and Christian Urban Definitions for constants on quotient types. *) signature QUOTIENT_DEF = sig val add_quotient_def: ((binding * mixfix) * Attrib.binding) * (term * term) -> thm -> local_theory -> Quotient_Info.quotconsts * local_theory val quotient_def: (binding * typ option * mixfix) option * (Attrib.binding * (term * term)) -> local_theory -> Proof.state val quotient_def_cmd: (binding * string option * mixfix) option * (Attrib.binding * (string * string)) -> local_theory -> Proof.state end; structure Quotient_Def: QUOTIENT_DEF = struct (** Interface and Syntax Setup **) (* Generation of the code certificate from the rsp theorem *) open Lifting_Util infix 0 MRSL (* The ML-interface for a quotient definition takes as argument: - an optional binding and mixfix annotation - attributes - the new constant as term - the rhs of the definition as term - respectfulness theorem for the rhs It stores the qconst_info in the quotconsts data slot. Restriction: At the moment the left- and right-hand side of the definition must be a constant. *) fun error_msg bind str = let val name = Binding.name_of bind val pos = Position.here (Binding.pos_of bind) in error ("Head of quotient_definition " ^ quote str ^ " differs from declaration " ^ name ^ pos) end fun add_quotient_def ((var, (name, atts)), (lhs, rhs)) rsp_thm lthy = let val rty = fastype_of rhs val qty = fastype_of lhs val absrep_trm = Quotient_Term.absrep_fun lthy Quotient_Term.AbsF (rty, qty) $ rhs val prop = Syntax.check_term lthy (Logic.mk_equals (lhs, absrep_trm)) val (_, prop') = Local_Defs.cert_def lthy (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val ((qconst, (_ , def)), lthy') = Local_Theory.define (var, ((Thm.def_binding_optional (#1 var) name, atts), newrhs)) lthy fun qconst_data phi = Quotient_Info.transform_quotconsts phi {qconst = qconst, rconst = rhs, def = def} fun qualify defname suffix = Binding.name suffix |> Binding.qualify true defname val lhs_name = Binding.name_of (#1 var) val rsp_thm_name = qualify lhs_name "rsp" val lthy'' = lthy' - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => (case qconst_data phi of qcinfo as {qconst = Const (c, _), ...} => Quotient_Info.update_quotconsts (c, qcinfo) | _ => I)) |> (snd oo Local_Theory.note) ((rsp_thm_name, @{attributes [quot_respect]}), [rsp_thm]) in (qconst_data Morphism.identity, lthy'') end fun mk_readable_rsp_thm_eq tm ctxt = let val ctm = Thm.cterm_of ctxt tm fun abs_conv2 cv = Conv.abs_conv (Conv.abs_conv (cv o #2) o #2) ctxt fun erase_quants ctxt' ctm' = case (Thm.term_of ctm') of Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Conv.all_conv ctm' | _ => (Conv.binder_conv (erase_quants o #2) ctxt' then_conv Conv.rewr_conv @{thm fun_eq_iff[symmetric, THEN eq_reflection]}) ctm' val norm_fun_eq = abs_conv2 erase_quants then_conv Thm.eta_conversion fun simp_arrows_conv ctm = let val unfold_conv = Conv.rewrs_conv [@{thm rel_fun_eq_eq_onp[THEN eq_reflection]}, @{thm rel_fun_eq_rel[THEN eq_reflection]}, @{thm rel_fun_def[THEN eq_reflection]}] val left_conv = simp_arrows_conv then_conv Conv.try_conv norm_fun_eq fun binop_conv2 cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 in case (Thm.term_of ctm) of Const (\<^const_name>\rel_fun\, _) $ _ $ _ => (binop_conv2 left_conv simp_arrows_conv then_conv unfold_conv) ctm | _ => Conv.all_conv ctm end val unfold_ret_val_invs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewr_conv @{thm eq_onp_same_args[THEN eq_reflection]}))) ctxt val simp_conv = Conv.arg_conv (Conv.fun2_conv simp_arrows_conv) val univq_conv = Conv.rewr_conv @{thm HOL.all_simps(6)[symmetric, THEN eq_reflection]} val univq_prenex_conv = Conv.top_conv (K (Conv.try_conv univq_conv)) ctxt val beta_conv = Thm.beta_conversion true val eq_thm = (simp_conv then_conv univq_prenex_conv then_conv beta_conv then_conv unfold_ret_val_invs) ctm in Object_Logic.rulify ctxt (eq_thm RS Drule.equal_elim_rule2) end fun gen_quotient_def prep_var parse_term (raw_var, (attr, (raw_lhs, raw_rhs))) lthy = let val (opt_var, ctxt) = (case raw_var of NONE => (NONE, lthy) | SOME var => prep_var var lthy |>> SOME) val lhs_constraints = (case opt_var of SOME (_, SOME T, _) => [T] | _ => []) fun prep_term Ts = parse_term ctxt #> fold Type.constraint Ts #> Syntax.check_term ctxt; val lhs = prep_term lhs_constraints raw_lhs val rhs = prep_term [] raw_rhs val (lhs_str, lhs_ty) = dest_Free lhs handle TERM _ => error "Constant already defined" val _ = if null (strip_abs_vars rhs) then () else error "The definiens cannot be an abstraction" val _ = if is_Const rhs then () else warning "The definiens is not a constant" val var = (case opt_var of NONE => (Binding.name lhs_str, NoSyn) | SOME (binding, _, mx) => if Variable.check_name binding = lhs_str then (binding, mx) else error_msg binding lhs_str); fun try_to_prove_refl thm = let val lhs_eq = thm |> Thm.prop_of |> Logic.dest_implies |> fst |> strip_all_body |> try HOLogic.dest_Trueprop in case lhs_eq of SOME (Const (\<^const_name>\HOL.eq\, _) $ _ $ _) => SOME (@{thm refl} RS thm) | SOME _ => (case body_type (fastype_of lhs) of Type (typ_name, _) => \<^try>\ #equiv_thm (the (Quotient_Info.lookup_quotients lthy typ_name)) RS @{thm Equiv_Relations.equivp_reflp} RS thm\ | _ => NONE ) | _ => NONE end val rsp_rel = Quotient_Term.equiv_relation lthy (fastype_of rhs, lhs_ty) val internal_rsp_tm = HOLogic.mk_Trueprop (Syntax.check_term lthy (rsp_rel $ rhs $ rhs)) val readable_rsp_thm_eq = mk_readable_rsp_thm_eq internal_rsp_tm lthy val maybe_proven_rsp_thm = try_to_prove_refl readable_rsp_thm_eq val (readable_rsp_tm, _) = Logic.dest_implies (Thm.prop_of readable_rsp_thm_eq) fun after_qed thm_list lthy = let val internal_rsp_thm = case thm_list of [] => the maybe_proven_rsp_thm | [[thm]] => Goal.prove ctxt [] [] internal_rsp_tm (fn _ => resolve_tac ctxt [readable_rsp_thm_eq] 1 THEN Proof_Context.fact_tac ctxt [thm] 1) in snd (add_quotient_def ((var, attr), (lhs, rhs)) internal_rsp_thm lthy) end in case maybe_proven_rsp_thm of SOME _ => Proof.theorem NONE after_qed [] lthy | NONE => Proof.theorem NONE after_qed [[(readable_rsp_tm,[])]] lthy end val quotient_def = gen_quotient_def Proof_Context.cert_var (K I) val quotient_def_cmd = gen_quotient_def Proof_Context.read_var Syntax.parse_term (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\quotient_definition\ "definition for constants over the quotient type" (Scan.option Parse_Spec.constdecl -- Parse.!!! (Parse_Spec.opt_thm_name ":" -- (Parse.term -- (\<^keyword>\is\ |-- Parse.term))) >> quotient_def_cmd); end; diff --git a/src/HOL/Tools/Quotient/quotient_type.ML b/src/HOL/Tools/Quotient/quotient_type.ML --- a/src/HOL/Tools/Quotient/quotient_type.ML +++ b/src/HOL/Tools/Quotient/quotient_type.ML @@ -1,355 +1,355 @@ (* Title: HOL/Tools/Quotient/quotient_type.ML Author: Cezary Kaliszyk and Christian Urban Definition of a quotient type. *) signature QUOTIENT_TYPE = sig val add_quotient_type: {overloaded: bool} -> ((string list * binding * mixfix) * (typ * term * bool) * ((binding * binding) option * thm option)) * thm -> local_theory -> Quotient_Info.quotients * local_theory val quotient_type: {overloaded: bool} -> (string list * binding * mixfix) * (typ * term * bool) * ((binding * binding) option * thm option) -> Proof.context -> Proof.state val quotient_type_cmd: {overloaded: bool} -> (((((string list * binding) * mixfix) * string) * (bool * string)) * (binding * binding) option) * (Facts.ref * Token.src list) option -> Proof.context -> Proof.state end; structure Quotient_Type: QUOTIENT_TYPE = struct (*** definition of quotient types ***) val mem_def1 = @{lemma "y \ Collect S \ S y" by simp} val mem_def2 = @{lemma "S y \ y \ Collect S" by simp} (* constructs the term {c. EX (x::rty). rel x x \ c = Collect (rel x)} *) fun typedef_term rel rty lthy = let val [x, c] = [("x", rty), ("c", HOLogic.mk_setT rty)] |> Variable.variant_frees lthy [rel] |> map Free in HOLogic.Collect_const (HOLogic.mk_setT rty) $ (lambda c (HOLogic.exists_const rty $ lambda x (HOLogic.mk_conj (rel $ x $ x, HOLogic.mk_eq (c, HOLogic.Collect_const rty $ (rel $ x)))))) end (* makes the new type definitions and proves non-emptyness *) fun typedef_make overloaded (vs, qty_name, mx, rel, rty) equiv_thm lthy = let fun typedef_tac ctxt = EVERY1 (map (resolve_tac ctxt o single) [@{thm part_equivp_typedef}, equiv_thm]) in Typedef.add_typedef overloaded (qty_name, map (rpair dummyS) vs, mx) (typedef_term rel rty lthy) NONE typedef_tac lthy end (* tactic to prove the quot_type theorem for the new type *) fun typedef_quot_type_tac ctxt equiv_thm ((_, typedef_info): Typedef.info) = let val rep_thm = #Rep typedef_info RS mem_def1 val rep_inv = #Rep_inverse typedef_info val abs_inv = #Abs_inverse typedef_info val rep_inj = #Rep_inject typedef_info in (resolve_tac ctxt @{thms quot_type.intro} THEN' RANGE [ resolve_tac ctxt [equiv_thm], resolve_tac ctxt [rep_thm], resolve_tac ctxt [rep_inv], resolve_tac ctxt [abs_inv] THEN' resolve_tac ctxt [mem_def2] THEN' assume_tac ctxt, resolve_tac ctxt [rep_inj]]) 1 end (* proves the quot_type theorem for the new type *) fun typedef_quot_type_thm (rel, abs, rep, equiv_thm, typedef_info) lthy = let val quot_type_const = Const (\<^const_name>\quot_type\, fastype_of rel --> fastype_of abs --> fastype_of rep --> \<^typ>\bool\) val goal = HOLogic.mk_Trueprop (quot_type_const $ rel $ abs $ rep) in Goal.prove lthy [] [] goal (fn {context = ctxt, ...} => typedef_quot_type_tac ctxt equiv_thm typedef_info) end open Lifting_Util infix 0 MRSL fun define_cr_rel equiv_thm abs_fun lthy = let fun force_type_of_rel rel forced_ty = let val thy = Proof_Context.theory_of lthy val rel_ty = (domain_type o fastype_of) rel val ty_inst = Sign.typ_match thy (rel_ty, forced_ty) Vartab.empty in Envir.subst_term_types ty_inst rel end val (rty, qty) = (dest_funT o fastype_of) abs_fun val abs_fun_graph = HOLogic.mk_eq(abs_fun $ Bound 1, Bound 0) val Abs_body = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of Const (\<^const_name>\equivp\, _) $ _ => abs_fun_graph | Const (\<^const_name>\part_equivp\, _) $ rel => HOLogic.mk_conj (force_type_of_rel rel rty $ Bound 1 $ Bound 1, abs_fun_graph) | _ => error "unsupported equivalence theorem" ) val def_term = Abs ("x", rty, Abs ("y", qty, Abs_body)); val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val cr_rel_name = Binding.prefix_name "cr_" qty_name val (fixed_def_term, lthy') = yield_singleton (Variable.importT_terms) def_term lthy val ((_, (_ , def_thm)), lthy'') = Local_Theory.define ((cr_rel_name, NoSyn), ((Thm.def_binding cr_rel_name, []), fixed_def_term)) lthy' in (def_thm, lthy'') end; fun setup_lifting_package quot3_thm equiv_thm opt_par_thm lthy = let val (_ $ _ $ abs_fun $ _) = (HOLogic.dest_Trueprop o Thm.prop_of) quot3_thm val (T_def, lthy') = define_cr_rel equiv_thm abs_fun lthy val (rty, qty) = (dest_funT o fastype_of) abs_fun val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val quotient_thm_name = Binding.prefix_name "Quotient_" qty_name val (reflp_thm, quot_thm) = (case (HOLogic.dest_Trueprop o Thm.prop_of) equiv_thm of Const (\<^const_name>\equivp\, _) $ _ => (SOME (equiv_thm RS @{thm equivp_reflp2}), [quot3_thm, T_def, equiv_thm] MRSL @{thm Quotient3_to_Quotient_equivp}) | Const (\<^const_name>\part_equivp\, _) $ _ => (NONE, [quot3_thm, T_def] MRSL @{thm Quotient3_to_Quotient}) | _ => error "unsupported equivalence theorem") val config = { notes = true } in lthy' |> Lifting_Setup.setup_by_quotient config quot_thm reflp_thm opt_par_thm |> snd |> (snd oo Local_Theory.note) ((quotient_thm_name, []), [quot_thm]) end fun init_quotient_infr quot_thm equiv_thm opt_par_thm lthy = let val (_ $ rel $ abs $ rep) = (HOLogic.dest_Trueprop o Thm.prop_of) quot_thm val (qtyp, rtyp) = (dest_funT o fastype_of) rep val qty_full_name = (fst o dest_Type) qtyp fun quotients phi = Quotient_Info.transform_quotients phi {qtyp = qtyp, rtyp = rtyp, equiv_rel = rel, equiv_thm = equiv_thm, quot_thm = quot_thm} fun abs_rep phi = Quotient_Info.transform_abs_rep phi {abs = abs, rep = rep} in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Quotient_Info.update_quotients (qty_full_name, quotients phi) #> Quotient_Info.update_abs_rep (qty_full_name, abs_rep phi)) |> setup_lifting_package quot_thm equiv_thm opt_par_thm end (* main function for constructing a quotient type *) fun add_quotient_type overloaded (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), equiv_thm) lthy = let val part_equiv = if partial then equiv_thm else equiv_thm RS @{thm equivp_implies_part_equivp} (* generates the typedef *) val ((_, typedef_info), lthy1) = typedef_make overloaded (vs, qty_name, mx, rel, rty) part_equiv lthy (* abs and rep functions from the typedef *) val Abs_ty = #abs_type (#1 typedef_info) val Rep_ty = #rep_type (#1 typedef_info) val Abs_name = #Abs_name (#1 typedef_info) val Rep_name = #Rep_name (#1 typedef_info) val Abs_const = Const (Abs_name, Rep_ty --> Abs_ty) val Rep_const = Const (Rep_name, Abs_ty --> Rep_ty) (* more useful abs and rep definitions *) val abs_const = Const (\<^const_name>\quot_type.abs\, (rty --> rty --> \<^typ>\bool\) --> (Rep_ty --> Abs_ty) --> rty --> Abs_ty) val rep_const = Const (\<^const_name>\quot_type.rep\, (Abs_ty --> Rep_ty) --> Abs_ty --> rty) val abs_trm = abs_const $ rel $ Abs_const val rep_trm = rep_const $ Rep_const val (rep_name, abs_name) = (case opt_morphs of NONE => (Binding.prefix_name "rep_" qty_name, Binding.prefix_name "abs_" qty_name) | SOME morphs => morphs) val ((_, (_, abs_def)), lthy2) = lthy1 |> Local_Theory.define ((abs_name, NoSyn), ((Thm.def_binding abs_name, []), abs_trm)) val ((_, (_, rep_def)), lthy3) = lthy2 |> Local_Theory.define ((rep_name, NoSyn), ((Thm.def_binding rep_name, []), rep_trm)) (* quot_type theorem *) val quot_thm = typedef_quot_type_thm (rel, Abs_const, Rep_const, part_equiv, typedef_info) lthy3 (* quotient theorem *) val quotient_thm_name = Binding.prefix_name "Quotient3_" qty_name val quotient_thm = (quot_thm RS @{thm quot_type.Quotient}) |> fold_rule lthy3 [abs_def, rep_def] (* name equivalence theorem *) val equiv_thm_name = Binding.suffix_name "_equivp" qty_name (* storing the quotients *) val quotients = {qtyp = Abs_ty, rtyp = rty, equiv_rel = rel, equiv_thm = equiv_thm, quot_thm = quotient_thm} val lthy4 = lthy3 |> init_quotient_infr quotient_thm equiv_thm opt_par_thm |> (snd oo Local_Theory.note) ((equiv_thm_name, if partial then [] else @{attributes [quot_equiv]}), [equiv_thm]) |> (snd oo Local_Theory.note) ((quotient_thm_name, @{attributes [quot_thm]}), [quotient_thm]) in (quotients, lthy4) end (* sanity checks for the quotient type specifications *) fun sanity_check ((vs, qty_name, _), (rty, rel, _), _) = let val rty_tfreesT = map fst (Term.add_tfreesT rty []) val rel_tfrees = map fst (Term.add_tfrees rel []) val rel_frees = map fst (Term.add_frees rel []) val rel_vars = Term.add_vars rel [] val rel_tvars = Term.add_tvars rel [] val qty_str = Binding.print qty_name ^ ": " val illegal_rel_vars = if null rel_vars andalso null rel_tvars then [] else [qty_str ^ "illegal schematic variable(s) in the relation."] val dup_vs = (case duplicates (op =) vs of [] => [] | dups => [qty_str ^ "duplicate type variable(s) on the lhs: " ^ commas_quote dups]) val extra_rty_tfrees = (case subtract (op =) vs rty_tfreesT of [] => [] | extras => [qty_str ^ "extra type variable(s) on the lhs: " ^ commas_quote extras]) val extra_rel_tfrees = (case subtract (op =) vs rel_tfrees of [] => [] | extras => [qty_str ^ "extra type variable(s) in the relation: " ^ commas_quote extras]) val illegal_rel_frees = (case rel_frees of [] => [] | xs => [qty_str ^ "illegal variable(s) in the relation: " ^ commas_quote xs]) val errs = illegal_rel_vars @ dup_vs @ extra_rty_tfrees @ extra_rel_tfrees @ illegal_rel_frees in if null errs then () else error (cat_lines errs) end (* check for existence of map functions *) fun map_check ctxt (_, (rty, _, _), _) = let fun map_check_aux rty warns = (case rty of Type (_, []) => warns | Type (s, _) => if Symtab.defined (Functor.entries ctxt) s then warns else s :: warns | _ => warns) val warns = map_check_aux rty [] in if null warns then () else warning ("No map function defined for " ^ commas warns ^ ". This will cause problems later on.") end (*** interface and syntax setup ***) (* the ML-interface takes a list of tuples consisting of: - the name of the quotient type - its free type variables (first argument) - its mixfix annotation - the type to be quotient - the partial flag (a boolean) - the relation according to which the type is quotient - optional names of morphisms (rep/abs) - flag if code should be generated by Lifting package it opens a proof-state in which one has to show that the relations are equivalence relations *) fun quotient_type overloaded quot lthy = let (* sanity check *) val _ = sanity_check quot val _ = map_check lthy quot fun mk_goal (rty, rel, partial) = let val equivp_ty = ([rty, rty] ---> \<^typ>\bool\) --> \<^typ>\bool\ val const = if partial then \<^const_name>\part_equivp\ else \<^const_name>\equivp\ in HOLogic.mk_Trueprop (Const (const, equivp_ty) $ rel) end val goal = (mk_goal o #2) quot fun after_qed [[thm]] = (snd oo add_quotient_type overloaded) (quot, thm) in Proof.theorem NONE after_qed [[(goal, [])]] lthy end fun quotient_type_cmd overloaded spec lthy = let fun parse_spec ((((((vs, qty_name), mx), rty_str), (partial, rel_str)), opt_morphs), opt_par_xthm) lthy = let val rty = Syntax.read_typ lthy rty_str val tmp_lthy1 = Variable.declare_typ rty lthy val rel = Syntax.parse_term tmp_lthy1 rel_str |> Type.constraint (rty --> rty --> \<^typ>\bool\) |> Syntax.check_term tmp_lthy1 val tmp_lthy2 = Variable.declare_term rel tmp_lthy1 val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm in (((vs, qty_name, mx), (rty, rel, partial), (opt_morphs, opt_par_thm)), tmp_lthy2) end val (spec', _) = parse_spec spec lthy in quotient_type overloaded spec' lthy end (* command syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\quotient_type\ "quotient type definitions (require equivalence proofs)" (* FIXME Parse.type_args_constrained and standard treatment of sort constraints *) (Parse_Spec.overloaded -- (Parse.type_args -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- Parse.typ) -- (\<^keyword>\/\ |-- Scan.optional (Parse.reserved "partial" -- \<^keyword>\:\ >> K true) false -- Parse.term) -- Scan.option (\<^keyword>\morphisms\ |-- Parse.!!! (Parse.binding -- Parse.binding)) -- Scan.option (\<^keyword>\parametric\ |-- Parse.!!! Parse.thm)) >> (fn (overloaded, spec) => quotient_type_cmd {overloaded = overloaded} spec)) end diff --git a/src/HOL/Tools/Transfer/transfer_bnf.ML b/src/HOL/Tools/Transfer/transfer_bnf.ML --- a/src/HOL/Tools/Transfer/transfer_bnf.ML +++ b/src/HOL/Tools/Transfer/transfer_bnf.ML @@ -1,281 +1,281 @@ (* Title: HOL/Tools/Transfer/transfer_bnf.ML Author: Ondrej Kuncar, TU Muenchen Setup for Transfer for types that are BNF. *) signature TRANSFER_BNF = sig exception NO_PRED_DATA of unit val base_name_of_bnf: BNF_Def.bnf -> binding val type_name_of_bnf: BNF_Def.bnf -> string val lookup_defined_pred_data: Proof.context -> string -> Transfer.pred_data val bnf_only_type_ctr: (BNF_Def.bnf -> 'a -> 'a) -> BNF_Def.bnf -> 'a -> 'a end structure Transfer_BNF : TRANSFER_BNF = struct open BNF_Util open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar exception NO_PRED_DATA of unit (* util functions *) fun base_name_of_bnf bnf = Binding.name (Binding.name_of (name_of_bnf bnf)) fun mk_Domainp P = let val PT = fastype_of P val argT = hd (binder_types PT) in Const (\<^const_name>\Domainp\, PT --> argT --> HOLogic.boolT) $ P end fun type_name_of_bnf bnf = T_of_bnf bnf |> dest_Type |> fst fun bnf_only_type_ctr f bnf = if is_Type (T_of_bnf bnf) then f bnf else I fun bnf_of_fp_sugar (fp_sugar:fp_sugar) = nth (#bnfs (#fp_res fp_sugar)) (#fp_res_index fp_sugar) fun fp_sugar_only_type_ctr f fp_sugars = (case filter (is_Type o T_of_bnf o bnf_of_fp_sugar) fp_sugars of [] => I | fp_sugars' => f fp_sugars') (* relation constraints - bi_total & co. *) fun mk_relation_constraint name arg = (Const (name, fastype_of arg --> HOLogic.boolT)) $ arg fun side_constraint_tac bnf constr_defs ctxt = let val thms = constr_defs @ map mk_sym [rel_eq_of_bnf bnf, rel_conversep_of_bnf bnf, rel_OO_of_bnf bnf] in SELECT_GOAL (Local_Defs.unfold0_tac ctxt thms) THEN' rtac ctxt (rel_mono_of_bnf bnf) THEN_ALL_NEW assume_tac ctxt end fun bi_constraint_tac constr_iff sided_constr_intros ctxt = SELECT_GOAL (Local_Defs.unfold0_tac ctxt [constr_iff]) THEN' CONJ_WRAP' (fn thm => rtac ctxt thm THEN_ALL_NEW (REPEAT_DETERM o etac ctxt conjE THEN' assume_tac ctxt)) sided_constr_intros fun generate_relation_constraint_goal ctxt bnf constraint_def = let val constr_name = constraint_def |> Thm.prop_of |> HOLogic.dest_Trueprop |> fst o HOLogic.dest_eq |> head_of |> fst o dest_Const val live = live_of_bnf bnf val (((As, Bs), Ds), ctxt') = ctxt |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees' (map Type.sort_of_atyp (deads_of_bnf bnf)) val relator = mk_rel_of_bnf Ds As Bs bnf val relsT = map2 mk_pred2T As Bs val (args, ctxt'') = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt' val concl = HOLogic.mk_Trueprop (mk_relation_constraint constr_name (list_comb (relator, args))) val assms = map (HOLogic.mk_Trueprop o (mk_relation_constraint constr_name)) args val goal = Logic.list_implies (assms, concl) in (goal, ctxt'') end fun prove_relation_side_constraint ctxt bnf constraint_def = let val (goal, ctxt') = generate_relation_constraint_goal ctxt bnf constraint_def in Goal.prove_sorry ctxt' [] [] goal (fn {context = goal_ctxt, ...} => side_constraint_tac bnf [constraint_def] goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt' ctxt) |> Drule.zero_var_indexes end fun prove_relation_bi_constraint ctxt bnf constraint_def side_constraints = let val (goal, ctxt') = generate_relation_constraint_goal ctxt bnf constraint_def in Goal.prove_sorry ctxt' [] [] goal (fn {context = goal_ctxt, ...} => bi_constraint_tac constraint_def side_constraints goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt' ctxt) |> Drule.zero_var_indexes end val defs = [("left_total_rel", @{thm left_total_alt_def}), ("right_total_rel", @{thm right_total_alt_def}), ("left_unique_rel", @{thm left_unique_alt_def}), ("right_unique_rel", @{thm right_unique_alt_def})] fun prove_relation_constraints bnf ctxt = let val transfer_attr = @{attributes [transfer_rule]} val Tname = base_name_of_bnf bnf val defs = map (apsnd (prove_relation_side_constraint ctxt bnf)) defs val bi_total = prove_relation_bi_constraint ctxt bnf @{thm bi_total_alt_def} [snd (nth defs 0), snd (nth defs 1)] val bi_unique = prove_relation_bi_constraint ctxt bnf @{thm bi_unique_alt_def} [snd (nth defs 2), snd (nth defs 3)] val defs = ("bi_total_rel", bi_total) :: ("bi_unique_rel", bi_unique) :: defs in maps (fn (a, thm) => [((Binding.qualify_name true Tname a, []), [([thm], transfer_attr)])]) defs end (* relator_eq *) fun relator_eq bnf = [(Binding.empty_atts, [([rel_eq_of_bnf bnf], @{attributes [relator_eq]})])] (* transfer rules *) fun bnf_transfer_rules bnf = let val transfer_rules = map_transfer_of_bnf bnf :: pred_transfer_of_bnf bnf :: rel_transfer_of_bnf bnf :: set_transfer_of_bnf bnf val transfer_attr = @{attributes [transfer_rule]} in map (fn thm => (Binding.empty_atts, [([thm],transfer_attr)])) transfer_rules end (* Domainp theorem for predicator *) fun Domainp_tac bnf pred_def ctxt = let val n = live_of_bnf bnf val set_map's = set_map_of_bnf bnf in EVERY' [rtac ctxt ext, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [@{thm Domainp.simps}, in_rel_of_bnf bnf, pred_def]), rtac ctxt iffI, REPEAT_DETERM o eresolve_tac ctxt [exE, conjE, CollectE], hyp_subst_tac ctxt, CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt ballI, dtac ctxt (set_map RS equalityD1 RS set_mp), etac ctxt imageE, dtac ctxt set_rev_mp, assume_tac ctxt, REPEAT_DETERM o eresolve_tac ctxt [CollectE, @{thm case_prodE}], hyp_subst_tac ctxt, rtac ctxt @{thm iffD2[OF arg_cong2[of _ _ _ _ Domainp, OF refl fst_conv]]}, etac ctxt @{thm DomainPI}]) set_map's, REPEAT_DETERM o etac ctxt conjE, REPEAT_DETERM o resolve_tac ctxt [exI, (refl RS conjI), rotate_prems 1 conjI], rtac ctxt refl, rtac ctxt (box_equals OF [map_cong0_of_bnf bnf, map_comp_of_bnf bnf RS sym, map_id_of_bnf bnf]), REPEAT_DETERM_N n o (EVERY' [rtac ctxt @{thm box_equals[OF _ sym[OF o_apply] sym[OF id_apply]]}, rtac ctxt @{thm fst_conv}]), rtac ctxt CollectI, CONJ_WRAP' (fn set_map => EVERY' [rtac ctxt (set_map RS @{thm ord_eq_le_trans}), REPEAT_DETERM o resolve_tac ctxt [@{thm image_subsetI}, CollectI, @{thm case_prodI}], dtac ctxt (rotate_prems 1 bspec), assume_tac ctxt, etac ctxt @{thm DomainpE}, etac ctxt @{thm someI}]) set_map's ] end fun prove_Domainp_rel ctxt bnf pred_def = let val live = live_of_bnf bnf val (((As, Bs), Ds), ctxt') = ctxt |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees' (map Type.sort_of_atyp (deads_of_bnf bnf)) val relator = mk_rel_of_bnf Ds As Bs bnf val relsT = map2 mk_pred2T As Bs val (args, ctxt'') = Ctr_Sugar_Util.mk_Frees "R" relsT ctxt' val lhs = mk_Domainp (list_comb (relator, args)) val rhs = Term.list_comb (mk_pred_of_bnf Ds As bnf, map mk_Domainp args) val goal = HOLogic.mk_eq (lhs, rhs) |> HOLogic.mk_Trueprop in Goal.prove_sorry ctxt'' [] [] goal (fn {context = goal_ctxt, ...} => Domainp_tac bnf pred_def goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt'' ctxt) |> Drule.zero_var_indexes end fun predicator bnf lthy = let val pred_def = pred_set_of_bnf bnf val Domainp_rel = prove_Domainp_rel lthy bnf pred_def val rel_eq_onp = rel_eq_onp_of_bnf bnf val Domainp_rel_thm_name = Binding.qualify_name true (base_name_of_bnf bnf) "Domainp_rel" val pred_data = Transfer.mk_pred_data pred_def rel_eq_onp [] val type_name = type_name_of_bnf bnf val relator_domain_attr = @{attributes [relator_domain]} val notes = [((Domainp_rel_thm_name, []), [([Domainp_rel], relator_domain_attr)])] in lthy |> Local_Theory.notes notes |> snd - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data)) end (* BNF interpretation *) fun transfer_bnf_interpretation bnf lthy = let val dead = dead_of_bnf bnf val constr_notes = if dead > 0 then [] else prove_relation_constraints bnf lthy val relator_eq_notes = if dead > 0 then [] else relator_eq bnf val transfer_rule_notes = if dead > 0 then [] else bnf_transfer_rules bnf in lthy |> Local_Theory.notes (constr_notes @ relator_eq_notes @ transfer_rule_notes) |> snd |> predicator bnf end val _ = Theory.setup (bnf_interpretation transfer_plugin (bnf_only_type_ctr transfer_bnf_interpretation)) (* simplification rules for the predicator *) fun lookup_defined_pred_data ctxt name = case Transfer.lookup_pred_data ctxt name of SOME data => data | NONE => raise NO_PRED_DATA () (* fp_sugar interpretation *) fun fp_sugar_transfer_rules (fp_sugar:fp_sugar) = let val fp_ctr_sugar = #fp_ctr_sugar fp_sugar val transfer_rules = #ctr_transfers fp_ctr_sugar @ #case_transfers fp_ctr_sugar @ #disc_transfers fp_ctr_sugar @ #sel_transfers fp_ctr_sugar @ these (Option.map #co_rec_transfers (#fp_co_induct_sugar fp_sugar)) val transfer_attr = @{attributes [transfer_rule]} in map (fn thm => (Binding.empty_atts, [([thm], transfer_attr)])) transfer_rules end fun register_pred_injects fp_sugar lthy = let val pred_injects = #pred_injects (#fp_bnf_sugar fp_sugar) val type_name = type_name_of_bnf (#fp_bnf fp_sugar) val pred_data = lookup_defined_pred_data lthy type_name |> Transfer.update_pred_simps pred_injects in lthy - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => Transfer.update_pred_data type_name (Transfer.morph_pred_data phi pred_data)) end fun transfer_fp_sugars_interpretation fp_sugar lthy = let val lthy = register_pred_injects fp_sugar lthy val transfer_rules_notes = fp_sugar_transfer_rules fp_sugar in lthy |> Local_Theory.notes transfer_rules_notes |> snd end val _ = Theory.setup (fp_sugars_interpretation transfer_plugin (fp_sugar_only_type_ctr (fold transfer_fp_sugars_interpretation))) end diff --git a/src/HOL/Tools/functor.ML b/src/HOL/Tools/functor.ML --- a/src/HOL/Tools/functor.ML +++ b/src/HOL/Tools/functor.ML @@ -1,281 +1,281 @@ (* Title: HOL/Tools/functor.ML Author: Florian Haftmann, TU Muenchen Functorial structure of types. *) signature FUNCTOR = sig val find_atomic: Proof.context -> typ -> (typ * (bool * bool)) list val construct_mapper: Proof.context -> (string * bool -> term) -> bool -> typ -> typ -> term val functor_: string option -> term -> local_theory -> Proof.state val functor_cmd: string option -> string -> Proof.context -> Proof.state type entry val entries: Proof.context -> entry list Symtab.table end; structure Functor : FUNCTOR = struct (* bookkeeping *) val compN = "comp"; val idN = "id"; val compositionalityN = "compositionality"; val identityN = "identity"; type entry = { mapper: term, variances: (sort * (bool * bool)) list, comp: thm, id: thm }; structure Data = Generic_Data ( type T = entry list Symtab.table val empty = Symtab.empty fun merge data = Symtab.merge (K true) data ); val entries = Data.get o Context.Proof; (* type analysis *) fun term_with_typ ctxt T t = Envir.subst_term_types (Sign.typ_match (Proof_Context.theory_of ctxt) (fastype_of t, T) Vartab.empty) t; fun find_atomic ctxt T = let val variances_of = Option.map #variances o try hd o Symtab.lookup_list (entries ctxt); fun add_variance is_contra T = AList.map_default (op =) (T, (false, false)) ((if is_contra then apsnd else apfst) (K true)); fun analyze' is_contra (_, (co, contra)) T = (if co then analyze is_contra T else I) #> (if contra then analyze (not is_contra) T else I) and analyze is_contra (T as Type (tyco, Ts)) = (case variances_of tyco of NONE => add_variance is_contra T | SOME variances => fold2 (analyze' is_contra) variances Ts) | analyze is_contra T = add_variance is_contra T; in analyze false T [] end; fun construct_mapper ctxt atomic = let val lookup = hd o Symtab.lookup_list (entries ctxt); fun constructs is_contra (_, (co, contra)) T T' = (if co then [construct is_contra T T'] else []) @ (if contra then [construct (not is_contra) T T'] else []) and construct is_contra (T as Type (tyco, Ts)) (T' as Type (_, Ts')) = let val { mapper = raw_mapper, variances, ... } = lookup tyco; val args = maps (fn (arg_pattern, (T, T')) => constructs is_contra arg_pattern T T') (variances ~~ (Ts ~~ Ts')); val (U, U') = if is_contra then (T', T) else (T, T'); val mapper = term_with_typ ctxt (map fastype_of args ---> U --> U') raw_mapper; in list_comb (mapper, args) end | construct is_contra (TFree (v, _)) (TFree _) = atomic (v, is_contra); in construct end; (* mapper properties *) val compositionality_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [Simpdata.mk_eq @{thm comp_def}]); fun make_comp_prop ctxt variances (tyco, mapper) = let val sorts = map fst variances val (((vs3, vs2), vs1), _) = ctxt |> Variable.invent_types sorts ||>> Variable.invent_types sorts ||>> Variable.invent_types sorts val (Ts1, Ts2, Ts3) = (map TFree vs1, map TFree vs2, map TFree vs3); fun mk_argT ((T, T'), (_, (co, contra))) = (if co then [(T --> T')] else []) @ (if contra then [(T' --> T)] else []); val contras = maps (fn (_, (co, contra)) => (if co then [false] else []) @ (if contra then [true] else [])) variances; val Ts21 = maps mk_argT ((Ts2 ~~ Ts1) ~~ variances); val Ts32 = maps mk_argT ((Ts3 ~~ Ts2) ~~ variances); fun invents n k nctxt = let val names = Name.invent nctxt n k; in (names, fold Name.declare names nctxt) end; val ((names21, names32), nctxt) = Variable.names_of ctxt |> invents "f" (length Ts21) ||>> invents "f" (length Ts32); val T1 = Type (tyco, Ts1); val T2 = Type (tyco, Ts2); val T3 = Type (tyco, Ts3); val (args21, args32) = (names21 ~~ Ts21, names32 ~~ Ts32); val args31 = map2 (fn is_contra => fn ((f21, T21), (f32, T32)) => if not is_contra then HOLogic.mk_comp (Free (f21, T21), Free (f32, T32)) else HOLogic.mk_comp (Free (f32, T32), Free (f21, T21)) ) contras (args21 ~~ args32) fun mk_mapper T T' args = list_comb (term_with_typ ctxt (map fastype_of args ---> T --> T') mapper, args); val mapper21 = mk_mapper T2 T1 (map Free args21); val mapper32 = mk_mapper T3 T2 (map Free args32); val mapper31 = mk_mapper T3 T1 args31; val eq1 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (HOLogic.mk_comp (mapper21, mapper32), mapper31); val x = Free (the_single (Name.invent nctxt (Long_Name.base_name tyco) 1), T3) val eq2 = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (mapper21 $ (mapper32 $ x), mapper31 $ x); val comp_prop = fold_rev Logic.all (map Free (args21 @ args32)) eq1; val compositionality_prop = fold_rev Logic.all (map Free (args21 @ args32) @ [x]) eq2; fun prove_compositionality ctxt comp_thm = Goal.prove_sorry ctxt [] [] compositionality_prop (K (ALLGOALS (Method.insert_tac ctxt [@{thm fun_cong} OF [comp_thm]] THEN' Simplifier.asm_lr_simp_tac (put_simpset compositionality_ss ctxt) THEN_ALL_NEW (Goal.assume_rule_tac ctxt)))); in (comp_prop, prove_compositionality) end; val identity_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [Simpdata.mk_eq @{thm id_def}]); fun make_id_prop ctxt variances (tyco, mapper) = let val (vs, _) = Variable.invent_types (map fst variances) ctxt; val Ts = map TFree vs; fun bool_num b = if b then 1 else 0; fun mk_argT (T, (_, (co, contra))) = replicate (bool_num co + bool_num contra) T val arg_Ts = maps mk_argT (Ts ~~ variances) val T = Type (tyco, Ts); val head = term_with_typ ctxt (map (fn T => T --> T) arg_Ts ---> T --> T) mapper; val lhs1 = list_comb (head, map (HOLogic.id_const) arg_Ts); val lhs2 = list_comb (head, map (fn arg_T => Abs ("x", arg_T, Bound 0)) arg_Ts); val rhs = HOLogic.id_const T; val (id_prop, identity_prop) = apply2 (HOLogic.mk_Trueprop o HOLogic.mk_eq o rpair rhs) (lhs1, lhs2); fun prove_identity ctxt id_thm = Goal.prove_sorry ctxt [] [] identity_prop (K (ALLGOALS (Method.insert_tac ctxt [id_thm] THEN' Simplifier.asm_lr_simp_tac (put_simpset identity_ss ctxt)))); in (id_prop, prove_identity) end; (* analyzing and registering mappers *) fun consume _ _ [] = (false, []) | consume eq x (ys as z :: zs) = if eq (x, z) then (true, zs) else (false, ys); fun split_mapper_typ "fun" T = let val (Ts', T') = strip_type T; val (Ts'', T'') = split_last Ts'; val (Ts''', T''') = split_last Ts''; in (Ts''', T''', T'' --> T') end | split_mapper_typ _ T = let val (Ts', T') = strip_type T; val (Ts'', T'') = split_last Ts'; in (Ts'', T'', T') end; fun analyze_mapper ctxt input_mapper = let val T = fastype_of input_mapper; val _ = Type.no_tvars T; val _ = if null (subtract (op =) (Term.add_tfreesT T []) (Term.add_tfrees input_mapper [])) then () else error ("Illegal additional type variable(s) in term: " ^ Syntax.string_of_term ctxt input_mapper); val _ = if null (Term.add_vars (singleton (Variable.export_terms (Proof_Context.augment input_mapper ctxt) ctxt) input_mapper) []) then () else error ("Illegal locally free variable(s) in term: " ^ Syntax.string_of_term ctxt input_mapper); val mapper = singleton (Variable.polymorphic ctxt) input_mapper; val _ = if null (Term.add_tfreesT (fastype_of mapper) []) then () else error ("Illegal locally fixed type variable(s) in type: " ^ Syntax.string_of_typ ctxt T); fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts | add_tycos _ = I; val tycos = add_tycos T []; val tyco = if tycos = ["fun"] then "fun" else case remove (op =) "fun" tycos of [tyco] => tyco | _ => error ("Bad number of type constructors: " ^ Syntax.string_of_typ ctxt T); in (mapper, T, tyco) end; fun analyze_variances ctxt tyco T = let fun bad_typ () = error ("Bad mapper type: " ^ Syntax.string_of_typ ctxt T); val (Ts, T1, T2) = split_mapper_typ tyco T handle List.Empty => bad_typ (); val _ = apply2 ((fn tyco' => if tyco' = tyco then () else bad_typ ()) o fst o dest_Type) (T1, T2) handle TYPE _ => bad_typ (); val (vs1, vs2) = apply2 (map dest_TFree o snd o dest_Type) (T1, T2) handle TYPE _ => bad_typ (); val _ = if has_duplicates (eq_fst (op =)) (vs1 @ vs2) then bad_typ () else (); fun check_variance_pair (var1 as (_, sort1), var2 as (_, sort2)) = let val coT = TFree var1 --> TFree var2; val contraT = TFree var2 --> TFree var1; val sort = Sign.inter_sort (Proof_Context.theory_of ctxt) (sort1, sort2); in consume (op =) coT ##>> consume (op =) contraT #>> pair sort end; val (variances, left_variances) = fold_map check_variance_pair (vs1 ~~ vs2) Ts; val _ = if null left_variances then () else bad_typ (); in variances end; fun gen_functor prep_term some_prfx raw_mapper lthy = let val (mapper, T, tyco) = analyze_mapper lthy (prep_term lthy raw_mapper); val prfx = the_default (Long_Name.base_name tyco) some_prfx; val variances = analyze_variances lthy tyco T; val (comp_prop, prove_compositionality) = make_comp_prop lthy variances (tyco, mapper); val (id_prop, prove_identity) = make_id_prop lthy variances (tyco, mapper); val qualify = Binding.qualify true prfx o Binding.name; fun mapper_declaration comp_thm id_thm phi context = let val typ_instance = Sign.typ_instance (Context.theory_of context); val mapper' = Morphism.term phi mapper; val T_T' = apply2 fastype_of (mapper, mapper'); val vars = Term.add_vars mapper' []; in if null vars andalso typ_instance T_T' andalso typ_instance (swap T_T') then (Data.map o Symtab.cons_list) (tyco, { mapper = mapper', variances = variances, comp = Morphism.thm phi comp_thm, id = Morphism.thm phi id_thm }) context else context end; fun after_qed [single_comp_thm, single_id_thm] lthy = lthy |> Local_Theory.note ((qualify compN, []), single_comp_thm) ||>> Local_Theory.note ((qualify idN, []), single_id_thm) |-> (fn ((_, [comp_thm]), (_, [id_thm])) => fn lthy => lthy |> Local_Theory.note ((qualify compositionalityN, []), [prove_compositionality lthy comp_thm]) |> snd |> Local_Theory.note ((qualify identityN, []), [prove_identity lthy id_thm]) |> snd - |> Local_Theory.declaration {syntax = false, pervasive = false} + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (mapper_declaration comp_thm id_thm)) in lthy |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) [comp_prop, id_prop]) end val functor_ = gen_functor Syntax.check_term; val functor_cmd = gen_functor Syntax.read_term; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\functor\ "register operations managing the functorial structure of a type" (Scan.option (Parse.name --| \<^keyword>\:\) -- Parse.term >> uncurry functor_cmd); end; diff --git a/src/HOL/Tools/inductive.ML b/src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML +++ b/src/HOL/Tools/inductive.ML @@ -1,1318 +1,1318 @@ (* Title: HOL/Tools/inductive.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer and Markus Wenzel, TU Muenchen (Co)Inductive Definition module for HOL. Features: * least or greatest fixedpoints * mutually recursive definitions * definitions involving arbitrary monotone operators * automatically proves introduction and elimination rules Introduction rules have the form [| M Pj ti, ..., Q x, ... |] ==> Pk t where M is some monotone operator (usually the identity) Q x is any side condition on the free variables ti, t are any terms Pj, Pk are two of the predicates being defined in mutual recursion *) signature INDUCTIVE = sig type result = {preds: term list, elims: thm list, raw_induct: thm, induct: thm, inducts: thm list, intrs: thm list, eqs: thm list} val transform_result: morphism -> result -> result type info = {names: string list, coind: bool} * result val the_inductive: Proof.context -> term -> info val the_inductive_global: Proof.context -> string -> info val print_inductives: bool -> Proof.context -> unit val get_monos: Proof.context -> thm list val mono_add: attribute val mono_del: attribute val mk_cases_tac: Proof.context -> tactic val mk_cases: Proof.context -> term -> thm val inductive_forall_def: thm val rulify: Proof.context -> thm -> thm val inductive_cases: (Attrib.binding * term list) list -> local_theory -> (string * thm list) list * local_theory val inductive_cases_cmd: (Attrib.binding * string list) list -> local_theory -> (string * thm list) list * local_theory val ind_cases_rules: Proof.context -> string list -> (binding * string option * mixfix) list -> thm list val inductive_simps: (Attrib.binding * term list) list -> local_theory -> (string * thm list) list * local_theory val inductive_simps_cmd: (Attrib.binding * string list) list -> local_theory -> (string * thm list) list * local_theory type flags = {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool} val add_inductive: flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> result * local_theory val add_inductive_cmd: bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> result * local_theory val arities_of: thm -> (string * int) list val params_of: thm -> term list val partition_rules: thm -> thm list -> (string * thm list) list val partition_rules': thm -> (thm * 'a) list -> (string * (thm * 'a) list) list val unpartition_rules: thm list -> (string * 'a list) list -> 'a list val infer_intro_vars: theory -> thm -> int -> thm list -> term list list val inductive_internals: bool Config.T val select_disj_tac: Proof.context -> int -> int -> int -> tactic type add_ind_def = flags -> term list -> (Attrib.binding * term) list -> thm list -> term list -> (binding * mixfix) list -> local_theory -> result * local_theory val declare_rules: binding -> bool -> bool -> binding -> string list -> term list -> thm list -> binding list -> Token.src list list -> (thm * string list * int) list -> thm list -> thm -> local_theory -> thm list * thm list * thm list * thm * thm list * local_theory val add_ind_def: add_ind_def val gen_add_inductive: add_ind_def -> flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> result * local_theory val gen_add_inductive_cmd: add_ind_def -> bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> result * local_theory val gen_ind_decl: add_ind_def -> bool -> (local_theory -> local_theory) parser end; structure Inductive: INDUCTIVE = struct (** theory context references **) val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_conj_def = @{thm HOL.induct_conj_def}; val inductive_conj = @{thms induct_conj}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; val inductive_rulify_fallback = @{thms induct_rulify_fallback}; val simp_thms1 = map mk_meta_eq @{lemma "(\ True) = False" "(\ False) = True" "(True \ P) = P" "(False \ P) = True" "(P \ True) = P" "(True \ P) = P" by (fact simp_thms)+}; val simp_thms2 = map mk_meta_eq [@{thm inf_fun_def}, @{thm inf_bool_def}] @ simp_thms1; val simp_thms3 = @{thms le_rel_bool_arg_iff if_False if_True conj_ac le_fun_def le_bool_def sup_fun_def sup_bool_def simp_thms if_bool_eq_disj all_simps ex_simps imp_conjL}; (** misc utilities **) val inductive_internals = Attrib.setup_config_bool \<^binding>\inductive_internals\ (K false); fun message quiet_mode s = if quiet_mode then () else writeln s; fun clean_message ctxt quiet_mode s = if Config.get ctxt quick_and_dirty then () else message quiet_mode s; fun coind_prefix true = "co" | coind_prefix false = ""; fun log (b: int) m n = if m >= n then 0 else 1 + log b (b * m) n; fun make_bool_args f g [] i = [] | make_bool_args f g (x :: xs) i = (if i mod 2 = 0 then f x else g x) :: make_bool_args f g xs (i div 2); fun make_bool_args' xs = make_bool_args (K \<^term>\False\) (K \<^term>\True\) xs; fun arg_types_of k c = drop k (binder_types (fastype_of c)); fun find_arg T x [] = raise Fail "find_arg" | find_arg T x ((p as (_, (SOME _, _))) :: ps) = apsnd (cons p) (find_arg T x ps) | find_arg T x ((p as (U, (NONE, y))) :: ps) = if (T: typ) = U then (y, (U, (SOME x, y)) :: ps) else apsnd (cons p) (find_arg T x ps); fun make_args Ts xs = map (fn (T, (NONE, ())) => Const (\<^const_name>\undefined\, T) | (_, (SOME t, ())) => t) (fold (fn (t, T) => snd o find_arg T t) xs (map (rpair (NONE, ())) Ts)); fun make_args' Ts xs Us = fst (fold_map (fn T => find_arg T ()) Us (Ts ~~ map (pair NONE) xs)); fun dest_predicate cs params t = let val k = length params; val (c, ts) = strip_comb t; val (xs, ys) = chop k ts; val i = find_index (fn c' => c' = c) cs; in if xs = params andalso i >= 0 then SOME (c, i, ys, chop (length ys) (arg_types_of k c)) else NONE end; fun mk_names a 0 = [] | mk_names a 1 = [a] | mk_names a n = map (fn i => a ^ string_of_int i) (1 upto n); fun select_disj_tac ctxt = let fun tacs 1 1 = [] | tacs _ 1 = [resolve_tac ctxt @{thms disjI1}] | tacs n i = resolve_tac ctxt @{thms disjI2} :: tacs (n - 1) (i - 1); in fn n => fn i => EVERY' (tacs n i) end; (** context data **) type result = {preds: term list, elims: thm list, raw_induct: thm, induct: thm, inducts: thm list, intrs: thm list, eqs: thm list}; fun transform_result phi {preds, elims, raw_induct: thm, induct, inducts, intrs, eqs} = let val term = Morphism.term phi; val thm = Morphism.thm phi; val fact = Morphism.fact phi; in {preds = map term preds, elims = fact elims, raw_induct = thm raw_induct, induct = thm induct, inducts = fact inducts, intrs = fact intrs, eqs = fact eqs} end; type info = {names: string list, coind: bool} * result; val empty_infos = Item_Net.init (op = o apply2 (#names o fst)) (#preds o snd) val empty_equations = Item_Net.init Thm.eq_thm_prop (single o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of); datatype data = Data of {infos: info Item_Net.T, monos: thm list, equations: thm Item_Net.T}; fun make_data (infos, monos, equations) = Data {infos = infos, monos = monos, equations = equations}; structure Data = Generic_Data ( type T = data; val empty = make_data (empty_infos, [], empty_equations); fun merge (Data {infos = infos1, monos = monos1, equations = equations1}, Data {infos = infos2, monos = monos2, equations = equations2}) = make_data (Item_Net.merge (infos1, infos2), Thm.merge_thms (monos1, monos2), Item_Net.merge (equations1, equations2)); ); fun map_data f = Data.map (fn Data {infos, monos, equations} => make_data (f (infos, monos, equations))); fun rep_data ctxt = Data.get (Context.Proof ctxt) |> (fn Data rep => rep); fun print_inductives verbose ctxt = let val {infos, monos, ...} = rep_data ctxt; val space = Consts.space_of (Proof_Context.consts_of ctxt); val consts = Item_Net.content infos |> maps (fn ({names, ...}, result) => map (rpair result) names) in [Pretty.block (Pretty.breaks (Pretty.str "(co)inductives:" :: map (Pretty.mark_str o #1) (Name_Space.markup_entries verbose ctxt space consts))), Pretty.big_list "monotonicity rules:" (map (Thm.pretty_thm_item ctxt) monos)] end |> Pretty.writeln_chunks; (* inductive info *) fun the_inductive ctxt term = Item_Net.retrieve (#infos (rep_data ctxt)) term |> the_single |> apsnd (transform_result (Morphism.transfer_morphism' ctxt)) fun the_inductive_global ctxt name = #infos (rep_data ctxt) |> Item_Net.content |> filter (fn ({names, ...}, _) => member op = names name) |> the_single |> apsnd (transform_result (Morphism.transfer_morphism' ctxt)) fun put_inductives info = map_data (fn (infos, monos, equations) => (Item_Net.update (apsnd (transform_result Morphism.trim_context_morphism) info) infos, monos, equations)); (* monotonicity rules *) fun get_monos ctxt = #monos (rep_data ctxt) |> map (Thm.transfer' ctxt); fun mk_mono ctxt thm = let fun eq_to_mono thm' = thm' RS (thm' RS @{thm eq_to_mono}); fun dest_less_concl thm = dest_less_concl (thm RS @{thm le_funD}) handle THM _ => thm RS @{thm le_boolD} in (case Thm.concl_of thm of Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => eq_to_mono (HOLogic.mk_obj_eq thm) | _ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _) => eq_to_mono thm | _ $ (Const (\<^const_name>\Orderings.less_eq\, _) $ _ $ _) => dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}])) thm)) | _ => thm) end handle THM _ => error ("Bad monotonicity theorem:\n" ^ Thm.string_of_thm ctxt thm); val mono_add = Thm.declaration_attribute (fn thm => fn context => map_data (fn (infos, monos, equations) => (infos, Thm.add_thm (Thm.trim_context (mk_mono (Context.proof_of context) thm)) monos, equations)) context); val mono_del = Thm.declaration_attribute (fn thm => fn context => map_data (fn (infos, monos, equations) => (infos, Thm.del_thm (mk_mono (Context.proof_of context) thm) monos, equations)) context); val _ = Theory.setup (Attrib.setup \<^binding>\mono\ (Attrib.add_del mono_add mono_del) "declaration of monotonicity rule"); (* equations *) fun retrieve_equations ctxt = Item_Net.retrieve (#equations (rep_data ctxt)) #> map (Thm.transfer' ctxt); val equation_add_permissive = Thm.declaration_attribute (fn thm => map_data (fn (infos, monos, equations) => (infos, monos, perhaps (try (Item_Net.update (Thm.trim_context thm))) equations))); (** process rules **) local fun err_in_rule ctxt name t msg = error (cat_lines ["Ill-formed introduction rule " ^ Binding.print name, Syntax.string_of_term ctxt t, msg]); fun err_in_prem ctxt name t p msg = error (cat_lines ["Ill-formed premise", Syntax.string_of_term ctxt p, "in introduction rule " ^ Binding.print name, Syntax.string_of_term ctxt t, msg]); val bad_concl = "Conclusion of introduction rule must be an inductive predicate"; val bad_ind_occ = "Inductive predicate occurs in argument of inductive predicate"; val bad_app = "Inductive predicate must be applied to parameter(s) "; fun atomize_term thy = Raw_Simplifier.rewrite_term thy inductive_atomize []; in fun check_rule ctxt cs params ((binding, att), rule) = let val params' = Term.variant_frees rule (Logic.strip_params rule); val frees = rev (map Free params'); val concl = subst_bounds (frees, Logic.strip_assums_concl rule); val prems = map (curry subst_bounds frees) (Logic.strip_assums_hyp rule); val rule' = Logic.list_implies (prems, concl); val aprems = map (atomize_term (Proof_Context.theory_of ctxt)) prems; val arule = fold_rev (Logic.all o Free) params' (Logic.list_implies (aprems, concl)); fun check_ind err t = (case dest_predicate cs params t of NONE => err (bad_app ^ commas (map (Syntax.string_of_term ctxt) params)) | SOME (_, _, ys, _) => if exists (fn c => exists (fn t => Logic.occs (c, t)) ys) cs then err bad_ind_occ else ()); fun check_prem' prem t = if member (op =) cs (head_of t) then check_ind (err_in_prem ctxt binding rule prem) t else (case t of Abs (_, _, t) => check_prem' prem t | t $ u => (check_prem' prem t; check_prem' prem u) | _ => ()); fun check_prem (prem, aprem) = if can HOLogic.dest_Trueprop aprem then check_prem' prem prem else err_in_prem ctxt binding rule prem "Non-atomic premise"; val _ = (case concl of Const (\<^const_name>\Trueprop\, _) $ t => if member (op =) cs (head_of t) then (check_ind (err_in_rule ctxt binding rule') t; List.app check_prem (prems ~~ aprems)) else err_in_rule ctxt binding rule' bad_concl | _ => err_in_rule ctxt binding rule' bad_concl); in ((binding, att), arule) end; fun rulify ctxt = hol_simplify ctxt inductive_conj #> hol_simplify ctxt inductive_rulify #> hol_simplify ctxt inductive_rulify_fallback #> Simplifier.norm_hhf ctxt; end; (** proofs for (co)inductive predicates **) (* prove monotonicity *) fun prove_mono quiet_mode skip_mono predT fp_fun monos ctxt = (message (quiet_mode orelse skip_mono andalso Config.get ctxt quick_and_dirty) " Proving monotonicity ..."; (if skip_mono then Goal.prove_sorry else Goal.prove_future) ctxt [] [] (HOLogic.mk_Trueprop (\<^Const>\monotone_on predT predT for \<^Const>\top \<^Type>\set predT\\ \<^Const>\less_eq predT\ \<^Const>\less_eq predT\ fp_fun\)) (fn _ => EVERY [resolve_tac ctxt @{thms monoI} 1, REPEAT (resolve_tac ctxt [@{thm le_funI}, @{thm le_boolI'}] 1), REPEAT (FIRST [assume_tac ctxt 1, resolve_tac ctxt (map (mk_mono ctxt) monos @ get_monos ctxt) 1, eresolve_tac ctxt @{thms le_funE} 1, dresolve_tac ctxt @{thms le_boolD} 1])])); (* prove introduction rules *) fun prove_intrs quiet_mode coind mono fp_def k intr_ts rec_preds_defs ctxt ctxt' = let val _ = clean_message ctxt quiet_mode " Proving the introduction rules ..."; val unfold = funpow k (fn th => th RS fun_cong) (mono RS (fp_def RS (if coind then @{thm def_gfp_unfold} else @{thm def_lfp_unfold}))); val rules = [refl, TrueI, @{lemma "\ False" by (rule notI)}, exI, conjI]; val intrs = map_index (fn (i, intr) => Goal.prove_sorry ctxt [] [] intr (fn _ => EVERY [rewrite_goals_tac ctxt rec_preds_defs, resolve_tac ctxt [unfold RS iffD2] 1, select_disj_tac ctxt (length intr_ts) (i + 1) 1, (*Not ares_tac, since refl must be tried before any equality assumptions; backtracking may occur if the premises have extra variables!*) DEPTH_SOLVE_1 (resolve_tac ctxt rules 1 APPEND assume_tac ctxt 1)]) |> singleton (Proof_Context.export ctxt ctxt')) intr_ts in (intrs, unfold) end; (* prove elimination rules *) fun prove_elims quiet_mode cs params intr_ts intr_names unfold rec_preds_defs ctxt ctxt''' = let val _ = clean_message ctxt quiet_mode " Proving the elimination rules ..."; val ([pname], ctxt') = Variable.variant_fixes ["P"] ctxt; val P = HOLogic.mk_Trueprop (Free (pname, HOLogic.boolT)); fun dest_intr r = (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))), Logic.strip_assums_hyp r, Logic.strip_params r); val intrs = map dest_intr intr_ts ~~ intr_names; val rules1 = [disjE, exE, FalseE]; val rules2 = [conjE, FalseE, @{lemma "\ True \ R" by (rule notE [OF _ TrueI])}]; fun prove_elim c = let val Ts = arg_types_of (length params) c; val (anames, ctxt'') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt'; val frees = map Free (anames ~~ Ts); fun mk_elim_prem ((_, _, us, _), ts, params') = Logic.list_all (params', Logic.list_implies (map (HOLogic.mk_Trueprop o HOLogic.mk_eq) (frees ~~ us) @ ts, P)); val c_intrs = filter (equal c o #1 o #1 o #1) intrs; val prems = HOLogic.mk_Trueprop (list_comb (c, params @ frees)) :: map mk_elim_prem (map #1 c_intrs) in (Goal.prove_sorry ctxt'' [] prems P (fn {context = ctxt4, prems} => EVERY [cut_tac (hd prems) 1, rewrite_goals_tac ctxt4 rec_preds_defs, dresolve_tac ctxt4 [unfold RS iffD1] 1, REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules1)), REPEAT (FIRSTGOAL (eresolve_tac ctxt4 rules2)), EVERY (map (fn prem => DEPTH_SOLVE_1 (assume_tac ctxt4 1 ORELSE resolve_tac ctxt [rewrite_rule ctxt4 rec_preds_defs prem, conjI] 1)) (tl prems))]) |> singleton (Proof_Context.export ctxt'' ctxt'''), map #2 c_intrs, length Ts) end in map prove_elim cs end; (* prove simplification equations *) fun prove_eqs quiet_mode cs params intr_ts intrs (elims: (thm * bstring list * int) list) ctxt ctxt'' = (* FIXME ctxt'' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the simplification rules ..."; fun dest_intr r = (the (dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r))), Logic.strip_assums_hyp r, Logic.strip_params r); val intr_ts' = map dest_intr intr_ts; fun prove_eq c (elim: thm * 'a * 'b) = let val Ts = arg_types_of (length params) c; val (anames, ctxt') = Variable.variant_fixes (mk_names "a" (length Ts)) ctxt; val frees = map Free (anames ~~ Ts); val c_intrs = filter (equal c o #1 o #1 o #1) (intr_ts' ~~ intrs); fun mk_intr_conj (((_, _, us, _), ts, params'), _) = let fun list_ex ([], t) = t | list_ex ((a, T) :: vars, t) = HOLogic.exists_const T $ Abs (a, T, list_ex (vars, t)); val conjs = map2 (curry HOLogic.mk_eq) frees us @ map HOLogic.dest_Trueprop ts; in list_ex (params', if null conjs then \<^term>\True\ else foldr1 HOLogic.mk_conj conjs) end; val lhs = list_comb (c, params @ frees); val rhs = if null c_intrs then \<^term>\False\ else foldr1 HOLogic.mk_disj (map mk_intr_conj c_intrs); val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, rhs)); fun prove_intr1 (i, _) = Subgoal.FOCUS_PREMS (fn {context = ctxt'', params, prems, ...} => select_disj_tac ctxt'' (length c_intrs) (i + 1) 1 THEN EVERY (replicate (length params) (resolve_tac ctxt'' @{thms exI} 1)) THEN (if null prems then resolve_tac ctxt'' @{thms TrueI} 1 else let val (prems', last_prem) = split_last prems; in EVERY (map (fn prem => (resolve_tac ctxt'' @{thms conjI} 1 THEN resolve_tac ctxt'' [prem] 1)) prems') THEN resolve_tac ctxt'' [last_prem] 1 end)) ctxt' 1; fun prove_intr2 (((_, _, us, _), ts, params'), intr) = EVERY (replicate (length params') (eresolve_tac ctxt' @{thms exE} 1)) THEN (if null ts andalso null us then resolve_tac ctxt' [intr] 1 else EVERY (replicate (length ts + length us - 1) (eresolve_tac ctxt' @{thms conjE} 1)) THEN Subgoal.FOCUS_PREMS (fn {context = ctxt'', prems, ...} => let val (eqs, prems') = chop (length us) prems; val rew_thms = map (fn th => th RS @{thm eq_reflection}) eqs; in rewrite_goal_tac ctxt'' rew_thms 1 THEN resolve_tac ctxt'' [intr] 1 THEN EVERY (map (fn p => resolve_tac ctxt'' [p] 1) prems') end) ctxt' 1); in Goal.prove_sorry ctxt' [] [] eq (fn _ => resolve_tac ctxt' @{thms iffI} 1 THEN eresolve_tac ctxt' [#1 elim] 1 THEN EVERY (map_index prove_intr1 c_intrs) THEN (if null c_intrs then eresolve_tac ctxt' @{thms FalseE} 1 else let val (c_intrs', last_c_intr) = split_last c_intrs in EVERY (map (fn ci => eresolve_tac ctxt' @{thms disjE} 1 THEN prove_intr2 ci) c_intrs') THEN prove_intr2 last_c_intr end)) |> rulify ctxt' |> singleton (Proof_Context.export ctxt' ctxt'') end; in map2 prove_eq cs elims end; (* derivation of simplified elimination rules *) local (*delete needless equality assumptions*) val refl_thin = Goal.prove_global \<^theory>\HOL\ [] [] \<^prop>\\P. a = a \ P \ P\ (fn {context = ctxt, ...} => assume_tac ctxt 1); val elim_rls = [asm_rl, FalseE, refl_thin, conjE, exE]; fun elim_tac ctxt = REPEAT o eresolve_tac ctxt elim_rls; fun simp_case_tac ctxt i = EVERY' [elim_tac ctxt, asm_full_simp_tac ctxt, elim_tac ctxt, REPEAT o bound_hyp_subst_tac ctxt] i; in fun mk_cases_tac ctxt = ALLGOALS (simp_case_tac ctxt) THEN prune_params_tac ctxt; fun mk_cases ctxt prop = let fun err msg = error (Pretty.string_of (Pretty.block [Pretty.str msg, Pretty.fbrk, Syntax.pretty_term ctxt prop])); val elims = Induct.find_casesP ctxt prop; val cprop = Thm.cterm_of ctxt prop; fun mk_elim rl = Thm.implies_intr cprop (Tactic.rule_by_tactic ctxt (mk_cases_tac ctxt) (Thm.assume cprop RS rl)) |> singleton (Proof_Context.export (Proof_Context.augment prop ctxt) ctxt); in (case get_first (try mk_elim) elims of SOME r => r | NONE => err "Proposition not an inductive predicate:") end; end; (* inductive_cases *) fun gen_inductive_cases prep_att prep_prop args lthy = let val thmss = map snd args |> burrow (grouped 10 Par_List.map_independent (mk_cases lthy o prep_prop lthy)); val facts = map2 (fn ((a, atts), _) => fn thms => ((a, map (prep_att lthy) atts), [(thms, [])])) args thmss; in lthy |> Local_Theory.notes facts end; val inductive_cases = gen_inductive_cases (K I) Syntax.check_prop; val inductive_cases_cmd = gen_inductive_cases Attrib.check_src Syntax.read_prop; (* ind_cases *) fun ind_cases_rules ctxt raw_props raw_fixes = let val (props, ctxt') = Specification.read_props raw_props raw_fixes ctxt; val rules = Proof_Context.export ctxt' ctxt (map (mk_cases ctxt') props); in rules end; val _ = Theory.setup (Method.setup \<^binding>\ind_cases\ (Scan.lift (Scan.repeat1 Parse.prop -- Parse.for_fixes) >> (fn (props, fixes) => fn ctxt => Method.erule ctxt 0 (ind_cases_rules ctxt props fixes))) "case analysis for inductive definitions, based on simplified elimination rule"); (* derivation of simplified equation *) fun mk_simp_eq ctxt prop = let val thy = Proof_Context.theory_of ctxt; val ctxt' = Proof_Context.augment prop ctxt; val lhs_of = fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of; val substs = retrieve_equations ctxt (HOLogic.dest_Trueprop prop) |> map_filter (fn eq => SOME (Pattern.match thy (lhs_of eq, HOLogic.dest_Trueprop prop) (Vartab.empty, Vartab.empty), eq) handle Pattern.MATCH => NONE); val (subst, eq) = (case substs of [s] => s | _ => error ("equations matching pattern " ^ Syntax.string_of_term ctxt prop ^ " is not unique")); val inst = map (fn v => (fst v, Thm.cterm_of ctxt' (Envir.subst_term subst (Var v)))) (Term.add_vars (lhs_of eq) []); in infer_instantiate ctxt' inst eq |> Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.full_rewrite ctxt'))) |> singleton (Proof_Context.export ctxt' ctxt) end (* inductive simps *) fun gen_inductive_simps prep_att prep_prop args lthy = let val facts = args |> map (fn ((a, atts), props) => ((a, map (prep_att lthy) atts), map (Thm.no_attributes o single o mk_simp_eq lthy o prep_prop lthy) props)); in lthy |> Local_Theory.notes facts end; val inductive_simps = gen_inductive_simps (K I) Syntax.check_prop; val inductive_simps_cmd = gen_inductive_simps Attrib.check_src Syntax.read_prop; (* prove induction rule *) fun prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def rec_preds_defs ctxt ctxt''' = (* FIXME ctxt''' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the induction rule ..."; (* predicates for induction rule *) val (pnames, ctxt') = Variable.variant_fixes (mk_names "P" (length cs)) ctxt; val preds = map2 (curry Free) pnames (map (fn c => arg_types_of (length params) c ---> HOLogic.boolT) cs); (* transform an introduction rule into a premise for induction rule *) fun mk_ind_prem r = let fun subst s = (case dest_predicate cs params s of SOME (_, i, ys, (_, Ts)) => let val k = length Ts; val bs = map Bound (k - 1 downto 0); val P = list_comb (nth preds i, map (incr_boundvars k) ys @ bs); val Q = fold_rev Term.abs (mk_names "x" k ~~ Ts) (HOLogic.mk_binop \<^const_name>\HOL.induct_conj\ (list_comb (incr_boundvars k s, bs), P)); in (Q, case Ts of [] => SOME (s, P) | _ => NONE) end | NONE => (case s of t $ u => (fst (subst t) $ fst (subst u), NONE) | Abs (a, T, t) => (Abs (a, T, fst (subst t)), NONE) | _ => (s, NONE))); fun mk_prem s prems = (case subst s of (_, SOME (t, u)) => t :: u :: prems | (t, _) => t :: prems); val SOME (_, i, ys, _) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); in fold_rev (Logic.all o Free) (Logic.strip_params r) (Logic.list_implies (map HOLogic.mk_Trueprop (fold_rev mk_prem (map HOLogic.dest_Trueprop (Logic.strip_assums_hyp r)) []), HOLogic.mk_Trueprop (list_comb (nth preds i, ys)))) end; val ind_prems = map mk_ind_prem intr_ts; (* make conclusions for induction rules *) val Tss = map (binder_types o fastype_of) preds; val (xnames, ctxt'') = Variable.variant_fixes (mk_names "x" (length (flat Tss))) ctxt'; val mutual_ind_concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (((xnames, Ts), c), P) => let val frees = map Free (xnames ~~ Ts) in HOLogic.mk_imp (list_comb (c, params @ frees), list_comb (P, frees)) end) (unflat Tss xnames ~~ Tss ~~ cs ~~ preds))); (* make predicate for instantiation of abstract induction rule *) val ind_pred = fold_rev lambda (bs @ xs) (foldr1 HOLogic.mk_conj (map_index (fn (i, P) => fold_rev (curry HOLogic.mk_imp) (make_bool_args HOLogic.mk_not I bs i) (list_comb (P, make_args' argTs xs (binder_types (fastype_of P))))) preds)); val ind_concl = HOLogic.mk_Trueprop (HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ (rec_const, ind_pred)); val raw_fp_induct = mono RS (fp_def RS @{thm def_lfp_induct}); val induct = Goal.prove_sorry ctxt'' [] ind_prems ind_concl (fn {context = ctxt3, prems} => EVERY [rewrite_goals_tac ctxt3 [inductive_conj_def], DETERM (resolve_tac ctxt3 [raw_fp_induct] 1), REPEAT (resolve_tac ctxt3 [@{thm le_funI}, @{thm le_boolI}] 1), rewrite_goals_tac ctxt3 simp_thms2, (*This disjE separates out the introduction rules*) REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [disjE, exE, FalseE])), (*Now break down the individual cases. No disjE here in case some premise involves disjunction.*) REPEAT (FIRSTGOAL (eresolve_tac ctxt3 [conjE] ORELSE' bound_hyp_subst_tac ctxt3)), REPEAT (FIRSTGOAL (resolve_tac ctxt3 [conjI, impI] ORELSE' (eresolve_tac ctxt3 [notE] THEN' assume_tac ctxt3))), EVERY (map (fn prem => DEPTH_SOLVE_1 (assume_tac ctxt3 1 ORELSE resolve_tac ctxt3 [rewrite_rule ctxt3 (inductive_conj_def :: rec_preds_defs @ simp_thms2) prem, conjI, refl] 1)) prems)]); val lemma = Goal.prove_sorry ctxt'' [] [] (Logic.mk_implies (ind_concl, mutual_ind_concl)) (fn {context = ctxt3, ...} => EVERY [rewrite_goals_tac ctxt3 rec_preds_defs, REPEAT (EVERY [REPEAT (resolve_tac ctxt3 [conjI, impI] 1), REPEAT (eresolve_tac ctxt3 [@{thm le_funE}, @{thm le_boolE}] 1), assume_tac ctxt3 1, rewrite_goals_tac ctxt3 simp_thms1, assume_tac ctxt3 1])]); in singleton (Proof_Context.export ctxt'' ctxt''') (induct RS lemma) end; (* prove coinduction rule *) fun If_const T = Const (\<^const_name>\If\, HOLogic.boolT --> T --> T --> T); fun mk_If p t f = let val T = fastype_of t in If_const T $ p $ t $ f end; fun prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono fp_def rec_preds_defs ctxt ctxt''' = (* FIXME ctxt''' ?? *) let val _ = clean_message ctxt quiet_mode " Proving the coinduction rule ..."; val n = length cs; val (ns, xss) = map_split (fn pred => make_args' argTs xs (arg_types_of (length params) pred) |> `length) preds; val xTss = map (map fastype_of) xss; val (Rs_names, names_ctxt) = Variable.variant_fixes (mk_names "X" n) ctxt; val Rs = map2 (fn name => fn Ts => Free (name, Ts ---> \<^typ>\bool\)) Rs_names xTss; val Rs_applied = map2 (curry list_comb) Rs xss; val preds_applied = map2 (curry list_comb) (map (fn p => list_comb (p, params)) preds) xss; val abstract_list = fold_rev (absfree o dest_Free); val bss = map (make_bool_args (fn b => HOLogic.mk_eq (b, \<^term>\False\)) (fn b => HOLogic.mk_eq (b, \<^term>\True\)) bs) (0 upto n - 1); val eq_undefinedss = map (fn ys => map (fn x => HOLogic.mk_eq (x, Const (\<^const_name>\undefined\, fastype_of x))) (subtract (op =) ys xs)) xss; val R = @{fold 3} (fn bs => fn eqs => fn R => fn t => if null bs andalso null eqs then R else mk_If (Library.foldr1 HOLogic.mk_conj (bs @ eqs)) R t) bss eq_undefinedss Rs_applied \<^term>\False\ |> abstract_list (bs @ xs); fun subst t = (case dest_predicate cs params t of SOME (_, i, ts, (_, Us)) => let val l = length Us; val bs = map Bound (l - 1 downto 0); val args = map (incr_boundvars l) ts @ bs in HOLogic.mk_disj (list_comb (nth Rs i, args), list_comb (nth preds i, params @ args)) |> fold_rev absdummy Us end | NONE => (case t of t1 $ t2 => subst t1 $ subst t2 | Abs (x, T, u) => Abs (x, T, subst u) | _ => t)); fun mk_coind_prem r = let val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); val ps = map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @ map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r); in (i, fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps)) end; fun mk_prem i Ps = Logic.mk_implies ((nth Rs_applied i, Library.foldr1 HOLogic.mk_disj Ps) |> @{apply 2} HOLogic.mk_Trueprop) |> fold_rev Logic.all (nth xss i); val prems = map mk_coind_prem intr_ts |> AList.group (op =) |> sort (int_ord o apply2 fst) |> map (uncurry mk_prem); val concl = @{map 3} (fn xs => Ctr_Sugar_Util.list_all_free xs oo curry HOLogic.mk_imp) xss Rs_applied preds_applied |> Library.foldr1 HOLogic.mk_conj |> HOLogic.mk_Trueprop; val pred_defs_sym = if null rec_preds_defs then [] else map2 (fn n => fn thm => funpow n (fn thm => thm RS @{thm meta_fun_cong}) thm RS @{thm Pure.symmetric}) ns rec_preds_defs; val simps = simp_thms3 @ pred_defs_sym; val simprocs = [Simplifier.the_simproc ctxt "HOL.defined_All"]; val simplify = asm_full_simplify (Ctr_Sugar_Util.ss_only simps ctxt addsimprocs simprocs); val coind = (mono RS (fp_def RS @{thm def_coinduct})) |> infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt R)] |> simplify; fun idx_of t = find_index (fn R => R = the_single (subtract (op =) (preds @ params) (map Free (Term.add_frees t [])))) Rs; val coind_concls = HOLogic.dest_Trueprop (Thm.concl_of coind) |> HOLogic.dest_conj |> map (fn t => (idx_of t, t)) |> sort (int_ord o @{apply 2} fst) |> map snd; val reorder_bound_goals = map_filter (fn (t, u) => if t aconv u then NONE else SOME (HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)))) ((HOLogic.dest_Trueprop concl |> HOLogic.dest_conj) ~~ coind_concls); val reorder_bound_thms = map (fn goal => Goal.prove_sorry ctxt [] [] goal (fn {context = ctxt, prems = _} => HEADGOAL (EVERY' [resolve_tac ctxt [iffI], REPEAT_DETERM o resolve_tac ctxt [allI, impI], REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt, REPEAT_DETERM o resolve_tac ctxt [allI, impI], REPEAT_DETERM o dresolve_tac ctxt [spec], eresolve_tac ctxt [mp], assume_tac ctxt]))) reorder_bound_goals; val coinduction = Goal.prove_sorry ctxt [] prems concl (fn {context = ctxt, prems = CIH} => HEADGOAL (full_simp_tac (Ctr_Sugar_Util.ss_only (simps @ reorder_bound_thms) ctxt addsimprocs simprocs) THEN' resolve_tac ctxt [coind]) THEN ALLGOALS (REPEAT_ALL_NEW (REPEAT_DETERM o resolve_tac ctxt [allI, impI, conjI] THEN' REPEAT_DETERM o eresolve_tac ctxt [exE, conjE] THEN' dresolve_tac ctxt (map simplify CIH) THEN' REPEAT_DETERM o (assume_tac ctxt ORELSE' eresolve_tac ctxt [conjE] ORELSE' dresolve_tac ctxt [spec, mp])))) in coinduction |> length cs = 1 ? (Object_Logic.rulify ctxt #> rotate_prems ~1) |> singleton (Proof_Context.export names_ctxt ctxt''') end (** specification of (co)inductive predicates **) fun mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy = let val fp_name = if coind then \<^const_name>\Inductive.gfp\ else \<^const_name>\Inductive.lfp\; val argTs = fold (combine (op =) o arg_types_of (length params)) cs []; val k = log 2 1 (length cs); val predT = replicate k HOLogic.boolT ---> argTs ---> HOLogic.boolT; val p :: xs = map Free (Variable.variant_frees lthy intr_ts (("p", predT) :: (mk_names "x" (length argTs) ~~ argTs))); val bs = map Free (Variable.variant_frees lthy (p :: xs @ intr_ts) (map (rpair HOLogic.boolT) (mk_names "b" k))); fun subst t = (case dest_predicate cs params t of SOME (_, i, ts, (Ts, Us)) => let val l = length Us; val zs = map Bound (l - 1 downto 0); in fold_rev (Term.abs o pair "z") Us (list_comb (p, make_bool_args' bs i @ make_args argTs ((map (incr_boundvars l) ts ~~ Ts) @ (zs ~~ Us)))) end | NONE => (case t of t1 $ t2 => subst t1 $ subst t2 | Abs (x, T, u) => Abs (x, T, subst u) | _ => t)); (* transform an introduction rule into a conjunction *) (* [| p_i t; ... |] ==> p_j u *) (* is transformed into *) (* b_j & x_j = u & p b_j t & ... *) fun transform_rule r = let val SOME (_, i, ts, (Ts, _)) = dest_predicate cs params (HOLogic.dest_Trueprop (Logic.strip_assums_concl r)); val ps = make_bool_args HOLogic.mk_not I bs i @ map HOLogic.mk_eq (make_args' argTs xs Ts ~~ ts) @ map (subst o HOLogic.dest_Trueprop) (Logic.strip_assums_hyp r); in fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) (Logic.strip_params r) (if null ps then \<^term>\True\ else foldr1 HOLogic.mk_conj ps) end; (* make a disjunction of all introduction rules *) val fp_fun = fold_rev lambda (p :: bs @ xs) (if null intr_ts then \<^term>\False\ else foldr1 HOLogic.mk_disj (map transform_rule intr_ts)); (* add definition of recursive predicates to theory *) val is_auxiliary = length cs > 1; val rec_binding = if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name; val rec_name = Binding.name_of rec_binding; val internals = Config.get lthy inductive_internals; val ((rec_const, (_, fp_def)), lthy') = lthy |> is_auxiliary ? Proof_Context.concealed |> Local_Theory.define ((rec_binding, case cnames_syn of [(_, mx)] => mx | _ => NoSyn), ((Thm.make_def_binding internals rec_binding, @{attributes [nitpick_unfold]}), fold_rev lambda params (Const (fp_name, (predT --> predT) --> predT) $ fp_fun))) ||> Proof_Context.restore_naming lthy; val fp_def' = Simplifier.rewrite (put_simpset HOL_basic_ss lthy' addsimps [fp_def]) (Thm.cterm_of lthy' (list_comb (rec_const, params))); val specs = if is_auxiliary then map_index (fn (i, ((b, mx), c)) => let val Ts = arg_types_of (length params) c; val xs = map Free (Variable.variant_frees lthy' intr_ts (mk_names "x" (length Ts) ~~ Ts)); in ((b, mx), ((Thm.make_def_binding internals b, []), fold_rev lambda (params @ xs) (list_comb (rec_const, params @ make_bool_args' bs i @ make_args argTs (xs ~~ Ts))))) end) (cnames_syn ~~ cs) else []; val (consts_defs, lthy'') = lthy' |> fold_map Local_Theory.define specs; val preds = (case cs of [_] => [rec_const] | _ => map #1 consts_defs); val (_, ctxt'') = Variable.add_fixes (map (fst o dest_Free) params) lthy''; val mono = prove_mono quiet_mode skip_mono predT fp_fun monos ctxt''; val (_, lthy''') = lthy'' |> Local_Theory.note ((if internals then Binding.qualify true rec_name (Binding.name "mono") else Binding.empty, []), Proof_Context.export ctxt'' lthy'' [mono]); in (lthy''', Proof_Context.transfer (Proof_Context.theory_of lthy''') ctxt'', rec_binding, mono, fp_def', map (#2 o #2) consts_defs, list_comb (rec_const, params), preds, argTs, bs, xs) end; fun declare_rules rec_binding coind no_ind spec_name cnames preds intrs intr_bindings intr_atts elims eqs raw_induct lthy = let val rec_name = Binding.name_of rec_binding; fun rec_qualified qualified = Binding.qualify qualified rec_name; val intr_names = map Binding.name_of intr_bindings; val ind_case_names = if forall (equal "") intr_names then [] else [Attrib.case_names intr_names]; val induct = if coind then (raw_induct, [Attrib.case_names [rec_name], Attrib.case_conclusion (rec_name, intr_names), Attrib.consumes (1 - Thm.nprems_of raw_induct), - Attrib.internal (K (Induct.coinduct_pred (hd cnames)))]) + Attrib.internal \<^here> (K (Induct.coinduct_pred (hd cnames)))]) else if no_ind orelse length cnames > 1 then (raw_induct, ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))]) else (raw_induct RSN (2, rev_mp), ind_case_names @ [Attrib.consumes (~ (Thm.nprems_of raw_induct))]); val (intrs', lthy1) = lthy |> Spec_Rules.add spec_name (if coind then Spec_Rules.Co_Inductive else Spec_Rules.Inductive) preds intrs |> Local_Theory.notes (map (rec_qualified false) intr_bindings ~~ intr_atts ~~ map (fn th => [([th], @{attributes [Pure.intro?]})]) intrs) |>> map (hd o snd); val (((_, elims'), (_, [induct'])), lthy2) = lthy1 |> Local_Theory.note ((rec_qualified true (Binding.name "intros"), []), intrs') ||>> fold_map (fn (name, (elim, cases, k)) => Local_Theory.note ((Binding.qualify true (Long_Name.base_name name) (Binding.name "cases"), ((if forall (equal "") cases then [] else [Attrib.case_names cases]) @ [Attrib.consumes (1 - Thm.nprems_of elim), Attrib.constraints k, - Attrib.internal (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})), + Attrib.internal \<^here> (K (Induct.cases_pred name))] @ @{attributes [Pure.elim?]})), [elim]) #> apfst (hd o snd)) (if null elims then [] else cnames ~~ elims) ||>> Local_Theory.note ((rec_qualified true (Binding.name (coind_prefix coind ^ "induct")), #2 induct), [rulify lthy1 (#1 induct)]); val (eqs', lthy3) = lthy2 |> fold_map (fn (name, eq) => Local_Theory.note ((Binding.qualify true (Long_Name.base_name name) (Binding.name "simps"), - [Attrib.internal (K equation_add_permissive)]), [eq]) + [Attrib.internal \<^here> (K equation_add_permissive)]), [eq]) #> apfst (hd o snd)) (if null eqs then [] else (cnames ~~ eqs)) val (inducts, lthy4) = if no_ind orelse coind then ([], lthy3) else let val inducts = cnames ~~ Project_Rule.projects lthy3 (1 upto length cnames) induct' in lthy3 |> Local_Theory.notes [((rec_qualified true (Binding.name "inducts"), []), inducts |> map (fn (name, th) => ([th], ind_case_names @ [Attrib.consumes (1 - Thm.nprems_of th), - Attrib.internal (K (Induct.induct_pred name))])))] |>> snd o hd + Attrib.internal \<^here> (K (Induct.induct_pred name))])))] |>> snd o hd end; in (intrs', elims', eqs', induct', inducts, lthy4) end; type flags = {quiet_mode: bool, verbose: bool, alt_name: binding, coind: bool, no_elim: bool, no_ind: bool, skip_mono: bool}; type add_ind_def = flags -> term list -> (Attrib.binding * term) list -> thm list -> term list -> (binding * mixfix) list -> local_theory -> result * local_theory; fun add_ind_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono} cs intros monos params cnames_syn lthy = let val _ = null cnames_syn andalso error "No inductive predicates given"; val names = map (Binding.name_of o fst) cnames_syn; val _ = message (quiet_mode andalso not verbose) ("Proofs for " ^ coind_prefix coind ^ "inductive predicate(s) " ^ commas_quote names); val spec_name = Binding.conglomerate (map #1 cnames_syn); val cnames = map (Local_Theory.full_name lthy o #1) cnames_syn; (* FIXME *) val ((intr_names, intr_atts), intr_ts) = apfst split_list (split_list (map (check_rule lthy cs params) intros)); val (lthy1, lthy2, rec_binding, mono, fp_def, rec_preds_defs, rec_const, preds, argTs, bs, xs) = mk_ind_def quiet_mode skip_mono alt_name coind cs intr_ts monos params cnames_syn lthy; val (intrs, unfold) = prove_intrs quiet_mode coind mono fp_def (length bs + length xs) intr_ts rec_preds_defs lthy2 lthy1; val elims = if no_elim then [] else prove_elims quiet_mode cs params intr_ts (map Binding.name_of intr_names) unfold rec_preds_defs lthy2 lthy1; val raw_induct = zero_var_indexes (if no_ind then Drule.asm_rl else if coind then prove_coindrule quiet_mode preds cs argTs bs xs params intr_ts mono fp_def rec_preds_defs lthy2 lthy1 else prove_indrule quiet_mode cs argTs bs xs rec_const params intr_ts mono fp_def rec_preds_defs lthy2 lthy1); val eqs = if no_elim then [] else prove_eqs quiet_mode cs params intr_ts intrs elims lthy2 lthy1; val elims' = map (fn (th, ns, i) => (rulify lthy1 th, ns, i)) elims; val intrs' = map (rulify lthy1) intrs; val (intrs'', elims'', eqs', induct, inducts, lthy3) = declare_rules rec_binding coind no_ind spec_name cnames preds intrs' intr_names intr_atts elims' eqs raw_induct lthy1; val result = {preds = preds, intrs = intrs'', elims = elims'', raw_induct = rulify lthy3 raw_induct, induct = induct, inducts = inducts, eqs = eqs'}; val lthy4 = lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => + |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => let val result' = transform_result phi result; in put_inductives ({names = cnames, coind = coind}, result') end); in (result, lthy4) end; (* external interfaces *) fun gen_add_inductive mk_def flags cnames_syn pnames spec monos lthy = let (* abbrevs *) val (_, ctxt1) = Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn) lthy; fun get_abbrev ((name, atts), t) = if can (Logic.strip_assums_concl #> Logic.dest_equals) t then let val _ = Binding.is_empty name andalso null atts orelse error "Abbreviations may not have names or attributes"; val ((x, T), rhs) = Local_Defs.abs_def (snd (Local_Defs.cert_def ctxt1 (K []) t)); val var = (case find_first (fn ((c, _), _) => Binding.name_of c = x) cnames_syn of NONE => error ("Undeclared head of abbreviation " ^ quote x) | SOME ((b, T'), mx) => if T <> T' then error ("Bad type specification for abbreviation " ^ quote x) else (b, mx)); in SOME (var, rhs) end else NONE; val abbrevs = map_filter get_abbrev spec; val bs = map (Binding.name_of o fst o fst) abbrevs; (* predicates *) val pre_intros = filter_out (is_some o get_abbrev) spec; val cnames_syn' = filter_out (member (op =) bs o Binding.name_of o fst o fst) cnames_syn; val cs = map (Free o apfst Binding.name_of o fst) cnames_syn'; val ps = map Free pnames; val (_, ctxt2) = lthy |> Variable.add_fixes (map (Binding.name_of o fst o fst) cnames_syn'); val ctxt3 = ctxt2 |> fold (snd oo Local_Defs.fixed_abbrev) abbrevs; val expand = Assumption.export_term ctxt3 lthy #> Proof_Context.cert_term lthy; fun close_rule r = fold (Logic.all o Free) (fold_aterms (fn t as Free (v as (s, _)) => if Variable.is_fixed ctxt1 s orelse member (op =) ps t then I else insert (op =) v | _ => I) r []) r; val intros = map (apsnd (Syntax.check_term lthy #> close_rule #> expand)) pre_intros; val preds = map (fn ((c, _), mx) => (c, mx)) cnames_syn'; in lthy |> mk_def flags cs intros monos ps preds ||> fold (snd oo Local_Theory.abbrev Syntax.mode_default) abbrevs end; fun gen_add_inductive_cmd mk_def verbose coind cnames_syn pnames_syn intro_srcs raw_monos lthy = let val ((vars, intrs), _) = lthy |> Proof_Context.set_mode Proof_Context.mode_abbrev |> Specification.read_multi_specs (cnames_syn @ pnames_syn) intro_srcs; val (cs, ps) = chop (length cnames_syn) vars; val monos = Attrib.eval_thms lthy raw_monos; val flags = {quiet_mode = false, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = false, no_ind = false, skip_mono = false}; in lthy |> gen_add_inductive mk_def flags cs (map (apfst Binding.name_of o fst) ps) intrs monos end; val add_inductive = gen_add_inductive add_ind_def; val add_inductive_cmd = gen_add_inductive_cmd add_ind_def; (* read off arities of inductive predicates from raw induction rule *) fun arities_of induct = map (fn (_ $ t $ u) => (fst (dest_Const (head_of t)), length (snd (strip_comb u)))) (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct))); (* read off parameters of inductive predicate from raw induction rule *) fun params_of induct = let val (_ $ t $ u :: _) = HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct)); val (_, ts) = strip_comb t; val (_, us) = strip_comb u; in List.take (ts, length ts - length us) end; val pname_of_intr = Thm.concl_of #> HOLogic.dest_Trueprop #> head_of #> dest_Const #> fst; (* partition introduction rules according to predicate name *) fun gen_partition_rules f induct intros = fold_rev (fn r => AList.map_entry op = (pname_of_intr (f r)) (cons r)) intros (map (rpair [] o fst) (arities_of induct)); val partition_rules = gen_partition_rules I; fun partition_rules' induct = gen_partition_rules fst induct; fun unpartition_rules intros xs = fold_map (fn r => AList.map_entry_yield op = (pname_of_intr r) (fn x :: xs => (x, xs)) #>> the) intros xs |> fst; (* infer order of variables in intro rules from order of quantifiers in elim rule *) fun infer_intro_vars thy elim arity intros = let val _ :: cases = Thm.prems_of elim; val used = map (fst o fst) (Term.add_vars (Thm.prop_of elim) []); fun mtch (t, u) = let val params = Logic.strip_params t; val vars = map (Var o apfst (rpair 0)) (Name.variant_list used (map fst params) ~~ map snd params); val ts = map (curry subst_bounds (rev vars)) (List.drop (Logic.strip_assums_hyp t, arity)); val us = Logic.strip_imp_prems u; val tab = fold (Pattern.first_order_match thy) (ts ~~ us) (Vartab.empty, Vartab.empty); in map (Envir.subst_term tab) vars end in map (mtch o apsnd Thm.prop_of) (cases ~~ intros) end; (** outer syntax **) fun gen_ind_decl mk_def coind = Parse.vars -- Parse.for_fixes -- Scan.optional Parse_Spec.where_multi_specs [] -- Scan.optional (\<^keyword>\monos\ |-- Parse.!!! Parse.thms1) [] >> (fn (((preds, params), specs), monos) => (snd o gen_add_inductive_cmd mk_def true coind preds params specs monos)); val ind_decl = gen_ind_decl add_ind_def; val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive\ "define inductive predicates" (ind_decl false); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinductive\ "define coinductive predicates" (ind_decl true); val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_cases\ "create simplified instances of elimination rules" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_cases_cmd)); val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_simps\ "create simplification rules for inductive predicates" (Parse.and_list1 Parse_Spec.simple_specs >> (snd oo inductive_simps_cmd)); val _ = Outer_Syntax.command \<^command_keyword>\print_inductives\ "print (co)inductive definitions and monotonicity rules" (Parse.opt_bang >> (fn b => Toplevel.keep (print_inductives b o Toplevel.context_of))); end; diff --git a/src/HOL/Tools/inductive_set.ML b/src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML +++ b/src/HOL/Tools/inductive_set.ML @@ -1,560 +1,560 @@ (* Title: HOL/Tools/inductive_set.ML Author: Stefan Berghofer, TU Muenchen Wrapper for defining inductive sets using package for inductive predicates, including infrastructure for converting between predicates and sets. *) signature INDUCTIVE_SET = sig val to_set_att: thm list -> attribute val to_pred_att: thm list -> attribute val to_pred : thm list -> Context.generic -> thm -> thm val pred_set_conv_att: attribute val add_inductive: Inductive.flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> Inductive.result * local_theory val add_inductive_cmd: bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> Inductive.result * local_theory val mono_add: attribute val mono_del: attribute end; structure Inductive_Set: INDUCTIVE_SET = struct (***********************************************************************************) (* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *) (* and (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y}) *) (* used for converting "strong" (co)induction rules *) (***********************************************************************************) val anyt = Free ("t", TFree ("'t", [])); fun strong_ind_simproc tab = Simplifier.make_simproc \<^context> "strong_ind" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let fun close p t f = let val vs = Term.add_vars t [] in Thm.instantiate' [] (rev (map (SOME o Thm.cterm_of ctxt o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; fun mkop \<^const_name>\HOL.conj\ T x = SOME (Const (\<^const_name>\Lattices.inf\, T --> T --> T), x) | mkop \<^const_name>\HOL.disj\ T x = SOME (Const (\<^const_name>\Lattices.sup\, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T in HOLogic.Collect_const U $ HOLogic.mk_ptupleabs (HOLogic.flat_tuple_paths p) U HOLogic.boolT t end; fun decomp (Const (s, _) $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) = mkop s T (m, p, S, mk_collect p T (head_of u)) | decomp (Const (s, _) $ u $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S)) = mkop s T (m, p, mk_collect p T (head_of u), S) | decomp _ = NONE; val simp = full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv}) 1; fun mk_rew t = (case strip_abs_vars t of [] => NONE | xs => (case decomp (strip_abs_body t) of NONE => NONE | SOME (bop, (m, p, S, S')) => SOME (close (Goal.prove ctxt [] []) (Logic.mk_equals (t, fold_rev Term.abs xs (m $ p $ (bop $ S $ S')))) (K (EVERY [resolve_tac ctxt [eq_reflection] 1, REPEAT (resolve_tac ctxt @{thms ext} 1), resolve_tac ctxt @{thms iffI} 1, EVERY [eresolve_tac ctxt @{thms conjE} 1, resolve_tac ctxt @{thms IntI} 1, simp, simp, eresolve_tac ctxt @{thms IntE} 1, resolve_tac ctxt @{thms conjI} 1, simp, simp] ORELSE EVERY [eresolve_tac ctxt @{thms disjE} 1, resolve_tac ctxt @{thms UnI1} 1, simp, resolve_tac ctxt @{thms UnI2} 1, simp, eresolve_tac ctxt @{thms UnE} 1, resolve_tac ctxt @{thms disjI1} 1, simp, resolve_tac ctxt @{thms disjI2} 1, simp]]))) handle ERROR _ => NONE)) in (case strip_comb (Thm.term_of ct) of (h as Const (name, _), ts) => if Symtab.defined tab name then let val rews = map mk_rew ts in if forall is_none rews then NONE else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1) (map2 (fn SOME r => K r | NONE => Thm.reflexive o Thm.cterm_of ctxt) rews ts) (Thm.reflexive (Thm.cterm_of ctxt h))) end else NONE | _ => NONE) end}; (* only eta contract terms occurring as arguments of functions satisfying p *) fun eta_contract p = let fun eta b (Abs (a, T, body)) = (case eta b body of body' as (f $ Bound 0) => if Term.is_dependent f orelse not b then Abs (a, T, body') else incr_boundvars ~1 f | body' => Abs (a, T, body')) | eta b (t $ u) = eta b t $ eta (p (head_of t)) u | eta b t = t in eta false end; fun eta_contract_thm ctxt p = Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct => Thm.transitive (Thm.eta_conversion ct) (Thm.symmetric (Thm.eta_conversion (Thm.cterm_of ctxt (eta_contract p (Thm.term_of ct))))))); (***********************************************************) (* rules for converting between predicate and set notation *) (* *) (* rules for converting predicates to sets have the form *) (* P (%x y. (x, y) : s) = (%x y. (x, y) : S s) *) (* *) (* rules for converting sets to predicates have the form *) (* S {(x, y). p x y} = {(x, y). P p x y} *) (* *) (* where s and p are parameters *) (***********************************************************) structure Data = Generic_Data ( type T = {(* rules for converting predicates to sets *) to_set_simps: thm list, (* rules for converting sets to predicates *) to_pred_simps: thm list, (* arities of functions of type t set => ... => u set *) set_arities: (typ * (int list list option list * int list list option)) list Symtab.table, (* arities of functions of type (t => ... => bool) => u => ... => bool *) pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table}; val empty = {to_set_simps = [], to_pred_simps = [], set_arities = Symtab.empty, pred_arities = Symtab.empty}; fun merge ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1, set_arities = set_arities1, pred_arities = pred_arities1}, {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2, set_arities = set_arities2, pred_arities = pred_arities2}) : T = {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2), to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2), set_arities = Symtab.merge_list (op =) (set_arities1, set_arities2), pred_arities = Symtab.merge_list (op =) (pred_arities1, pred_arities2)}; ); fun name_type_of (Free p) = SOME p | name_type_of (Const p) = SOME p | name_type_of _ = NONE; fun map_type f (Free (s, T)) = Free (s, f T) | map_type f (Var (ixn, T)) = Var (ixn, f T) | map_type f _ = error "map_type"; fun find_most_specific is_inst f eq xs T = find_first (fn U => is_inst (T, f U) andalso forall (fn U' => eq (f U, f U') orelse not (is_inst (T, f U') andalso is_inst (f U', f U))) xs) xs; fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of NONE => NONE | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T; fun lookup_rule thy f rules = find_most_specific (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules; fun infer_arities thy arities (optf, t) fs = case strip_comb t of (Abs (_, _, u), []) => infer_arities thy arities (NONE, u) fs | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of SOME (SOME (_, (arity, _))) => (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs handle General.Subscript => error "infer_arities: bad term") | _ => fold (infer_arities thy arities) (map (pair NONE) ts) (case optf of NONE => fs | SOME f => AList.update op = (u, the_default f (Option.map (fn g => inter (op =) g f) (AList.lookup op = fs u))) fs)); (**************************************************************) (* derive the to_pred equation from the to_set equation *) (* *) (* 1. instantiate each set parameter with {(x, y). p x y} *) (* 2. apply %P. {(x, y). P x y} to both sides of the equation *) (* 3. simplify *) (**************************************************************) fun mk_to_pred_inst ctxt fs = map (fn (x, ps) => let val (Ts, T) = strip_type (fastype_of x); val U = HOLogic.dest_setT T; val x' = map_type (K (Ts @ HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x; in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs ps U HOLogic.boolT (list_comb (x', map Bound (length Ts - 1 downto 0)))))) end) fs; fun mk_to_pred_eq ctxt p fs optfs' T thm = let val insts = mk_to_pred_inst ctxt fs; val thm' = Thm.instantiate (TVars.empty, Vars.make insts) thm; val thm'' = (case optfs' of NONE => thm' RS sym | SOME fs' => let val U = HOLogic.dest_setT (body_type T); val Ts = HOLogic.strip_ptupleT fs' U; val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong; val (Var (arg_cong_f, _), _) = arg_cong' |> Thm.concl_of |> dest_comb |> snd |> strip_comb |> snd |> hd |> dest_comb; in thm' RS (infer_instantiate ctxt [(arg_cong_f, Thm.cterm_of ctxt (Abs ("P", Ts ---> HOLogic.boolT, HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs' U HOLogic.boolT (Bound 0))))] arg_cong' RS sym) end) in Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv} addsimprocs [\<^simproc>\Collect_mem\]) thm'' |> zero_var_indexes |> eta_contract_thm ctxt (equal p) end; (**** declare rules for converting predicates to sets ****) exception Malformed of string; fun add context thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) = (case Thm.prop_of thm of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ lhs $ rhs) => (case body_type T of \<^typ>\bool\ => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; fun factors_of t fs = case strip_abs_body t of Const (\<^const_name>\Set.member\, _) $ u $ S => if is_Free S orelse is_Var S then let val ps = HOLogic.flat_tuple_paths u in (SOME ps, (S, ps) :: fs) end else (NONE, fs) | _ => (NONE, fs); val (h, ts) = strip_comb lhs val (pfs, fs) = fold_map factors_of ts []; val ((h', ts'), fs') = (case rhs of Abs _ => (case strip_abs_body rhs of Const (\<^const_name>\Set.member\, _) $ u $ S => (strip_comb S, SOME (HOLogic.flat_tuple_paths u)) | _ => raise Malformed "member symbol on right-hand side expected") | _ => (strip_comb rhs, NONE)) in case (name_type_of h, name_type_of h') of (SOME (s, T), SOME (s', T')) => if exists (fn (U, _) => Sign.typ_instance thy (T', U) andalso Sign.typ_instance thy (U, T')) (Symtab.lookup_list set_arities s') then (if Context_Position.is_really_visible ctxt then warning ("Ignoring conversion rule for operator " ^ s') else (); tab) else {to_set_simps = Thm.trim_context thm :: to_set_simps, to_pred_simps = Thm.trim_context (mk_to_pred_eq ctxt h fs fs' T' thm) :: to_pred_simps, set_arities = Symtab.insert_list op = (s', (T', (map (AList.lookup op = fs) ts', fs'))) set_arities, pred_arities = Symtab.insert_list op = (s, (T, (pfs, fs'))) pred_arities} | _ => raise Malformed "set / predicate constant expected" end | _ => raise Malformed "equation between predicates expected") | _ => raise Malformed "equation expected") handle Malformed msg => let val ctxt = Context.proof_of context val _ = if Context_Position.is_really_visible ctxt then warning ("Ignoring malformed set / predicate conversion rule: " ^ msg ^ "\n" ^ Thm.string_of_thm ctxt thm) else (); in tab end; val pred_set_conv_att = Thm.declaration_attribute (fn thm => fn ctxt => Data.map (add ctxt thm) ctxt); (**** convert theorem in set notation to predicate notation ****) fun is_pred tab t = case Option.map (Symtab.lookup tab o fst) (name_type_of t) of SOME (SOME _) => true | _ => false; fun to_pred_simproc rules = let val rules' = map mk_meta_eq rules in Simplifier.make_simproc \<^context> "to_pred" {lhss = [anyt], proc = fn _ => fn ctxt => fn ct => lookup_rule (Proof_Context.theory_of ctxt) (Thm.prop_of #> Logic.dest_equals) rules' (Thm.term_of ct)} end; fun to_pred_proc thy rules t = case lookup_rule thy I rules t of NONE => NONE | SOME (lhs, rhs) => SOME (Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs); fun to_pred thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_pred_simps, set_arities, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy set_arities (NONE, Thm.prop_of thm) []); (* instantiate each set parameter with {(x, y). p x y} *) val insts = mk_to_pred_inst ctxt fs in thm |> Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimprocs [to_pred_simproc (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: map (Thm.transfer thy) to_pred_simps)]) |> eta_contract_thm ctxt (is_pred pred_arities) |> Rule_Cases.save thm end; val to_pred_att = Thm.rule_attribute [] o to_pred; (**** convert theorem in predicate notation to set notation ****) fun to_set thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_set_simps, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy pred_arities (NONE, Thm.prop_of thm) []); (* instantiate each predicate parameter with %x y. (x, y) : s *) val insts = map (fn (x, ps) => let val Ts = binder_types (fastype_of x); val l = length Ts; val k = length ps; val (Rs, Us) = chop (l - k - 1) Ts; val T = HOLogic.mk_ptupleT ps Us; val x' = map_type (K (Rs ---> HOLogic.mk_setT T)) x in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple ps T (map Bound (k downto 0)), list_comb (x', map Bound (l - 1 downto k + 1)))))) end) fs; in thm |> Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps to_set_simps addsimprocs [strong_ind_simproc pred_arities, \<^simproc>\Collect_mem\]) |> Rule_Cases.save thm end; val to_set_att = Thm.rule_attribute [] o to_set; (**** definition of inductive sets ****) fun add_ind_set_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono} cs intros monos params cnames_syn lthy = let val thy = Proof_Context.theory_of lthy; val {set_arities, pred_arities, to_pred_simps, ...} = Data.get (Context.Proof lthy); fun infer (Abs (_, _, t)) = infer t | infer (Const (\<^const_name>\Set.member\, _) $ t $ u) = infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u) | infer (t $ u) = infer t #> infer u | infer _ = I; val new_arities = filter_out (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 0 | _ => false) (fold (snd #> infer) intros []); val params' = map (fn x => (case AList.lookup op = new_arities x of SOME fs => let val T = HOLogic.dest_setT (fastype_of x); val Ts = HOLogic.strip_ptupleT fs T; val x' = map_type (K (Ts ---> HOLogic.boolT)) x in (x, (x', (HOLogic.Collect_const T $ HOLogic.mk_ptupleabs fs T HOLogic.boolT x', fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs T (map Bound (length fs downto 0)), x))))) end | NONE => (x, (x, (x, x))))) params; val (params1, (params2, params3)) = params' |> map snd |> split_list ||> split_list; val paramTs = map fastype_of params; (* equations for converting sets to predicates *) val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) => let val fs = the_default [] (AList.lookup op = new_arities c); val (Us, U) = strip_type T |> apsnd HOLogic.dest_setT; val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks [Pretty.str "Argument types", Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)), Pretty.str ("of " ^ s ^ " do not agree with types"), Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)), Pretty.str "of declared parameters"])); val Ts = HOLogic.strip_ptupleT fs U; val c' = Free (s ^ "p", map fastype_of params1 @ Ts ---> HOLogic.boolT) in ((c', (fs, U, Ts)), (list_comb (c, params2), HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (c', params1)))) end) |> split_list |>> split_list; val eqns' = eqns @ map (Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: to_pred_simps); (* predicate version of the introduction rules *) val intros' = map (fn (name_atts, t) => (name_atts, t |> map_aterms (fn u => (case AList.lookup op = params' u of SOME (_, (u', _)) => u' | NONE => u)) |> Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |> eta_contract (member op = cs' orf is_pred pred_arities))) intros; val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn; val monos' = map (to_pred [] (Context.Proof lthy)) monos; val ({preds, intrs, elims, raw_induct, eqs, ...}, lthy1) = Inductive.add_ind_def {quiet_mode = quiet_mode, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono} cs' intros' monos' params1 cnames_syn' lthy; (* define inductive sets using previously defined predicates *) val (defs, lthy2) = lthy1 |> fold_map Local_Theory.define (map (fn (((b, mx), (fs, U, _)), p) => ((b, mx), ((Thm.def_binding b, []), fold_rev lambda params (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (p, params3)))))) (cnames_syn ~~ cs_info ~~ preds)); (* prove theorems for converting predicate to set notation *) val lthy3 = fold (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn lthy => let val conv_thm = Goal.prove lthy (map (fst o dest_Free) params) [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (p, params3), fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs U (map Bound (length fs downto 0)), list_comb (c, params)))))) (K (REPEAT (resolve_tac lthy @{thms ext} 1) THEN simp_tac (put_simpset HOL_basic_ss lthy addsimps [def, @{thm mem_Collect_eq}, @{thm case_prod_conv}]) 1)) in lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"), - [Attrib.internal (K pred_set_conv_att)]), + [Attrib.internal \<^here> (K pred_set_conv_att)]), [conv_thm]) |> snd end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2; (* convert theorems to set notation *) val rec_name = if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name; val cnames = map (Local_Theory.full_name lthy3 o #1) cnames_syn; (* FIXME *) val spec_name = Binding.conglomerate (map #1 cnames_syn); val (intr_names, intr_atts) = split_list (map fst intros); val raw_induct' = to_set [] (Context.Proof lthy3) raw_induct; val (intrs', elims', eqs', induct, inducts, lthy4) = Inductive.declare_rules rec_name coind no_ind spec_name cnames (map fst defs) (map (to_set [] (Context.Proof lthy3)) intrs) intr_names intr_atts (map (fn th => (to_set [] (Context.Proof lthy3) th, map (fst o fst) (fst (Rule_Cases.get th)), Rule_Cases.get_constraints th)) elims) (map (to_set [] (Context.Proof lthy3)) eqs) raw_induct' lthy3; in ({intrs = intrs', elims = elims', induct = induct, inducts = inducts, raw_induct = raw_induct', preds = map fst defs, eqs = eqs'}, lthy4) end; val add_inductive = Inductive.gen_add_inductive add_ind_set_def; val add_inductive_cmd = Inductive.gen_add_inductive_cmd add_ind_set_def; fun mono_att att = Thm.declaration_attribute (fn thm => fn context => Thm.attribute_declaration att (to_pred [] context thm) context); val mono_add = mono_att Inductive.mono_add; val mono_del = mono_att Inductive.mono_del; (** package setup **) (* attributes *) val _ = Theory.setup (Attrib.setup \<^binding>\pred_set_conv\ (Scan.succeed pred_set_conv_att) "declare rules for converting between predicate and set notation" #> Attrib.setup \<^binding>\to_set\ (Attrib.thms >> to_set_att) "convert rule to set notation" #> Attrib.setup \<^binding>\to_pred\ (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #> Attrib.setup \<^binding>\mono_set\ (Attrib.add_del mono_add mono_del) "declare of monotonicity rule for set operators"); (* commands *) val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def; val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_set\ "define inductive sets" (ind_set_decl false); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinductive_set\ "define coinductive sets" (ind_set_decl true); end; diff --git a/src/HOL/Tools/semiring_normalizer.ML b/src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML +++ b/src/HOL/Tools/semiring_normalizer.ML @@ -1,885 +1,885 @@ (* Title: HOL/Tools/semiring_normalizer.ML Author: Amine Chaieb, TU Muenchen Normalization of expressions in semirings. *) signature SEMIRING_NORMALIZER = sig type entry val match: Proof.context -> cterm -> entry option val the_semiring: Proof.context -> thm -> cterm list * thm list val the_ring: Proof.context -> thm -> cterm list * thm list val the_field: Proof.context -> thm -> cterm list * thm list val the_idom: Proof.context -> thm -> thm list val the_ideal: Proof.context -> thm -> thm list val declare: thm -> {semiring: term list * thm list, ring: term list * thm list, field: term list * thm list, idom: thm list, ideal: thm list} -> local_theory -> local_theory val semiring_normalize_conv: Proof.context -> conv val semiring_normalize_ord_conv: Proof.context -> cterm ord -> conv val semiring_normalize_wrapper: Proof.context -> entry -> conv val semiring_normalize_ord_wrapper: Proof.context -> entry -> cterm ord -> conv val semiring_normalizers_conv: cterm list -> cterm list * thm list -> cterm list * thm list -> cterm list * thm list -> (cterm -> bool) * conv * conv * conv -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} val semiring_normalizers_ord_wrapper: Proof.context -> entry -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} end structure Semiring_Normalizer: SEMIRING_NORMALIZER = struct (** data **) type entry = {vars: cterm list, semiring: cterm list * thm list, ring: cterm list * thm list, field: cterm list * thm list, idom: thm list, ideal: thm list} * {is_const: cterm -> bool, dest_const: cterm -> Rat.rat, mk_const: ctyp -> Rat.rat -> cterm, conv: Proof.context -> cterm -> thm}; structure Data = Generic_Data ( type T = (thm * entry) list; val empty = []; fun merge data = AList.merge Thm.eq_thm (K true) data; ); fun the_rules ctxt = fst o the o AList.lookup Thm.eq_thm (Data.get (Context.Proof ctxt)) val the_semiring = #semiring oo the_rules val the_ring = #ring oo the_rules val the_field = #field oo the_rules val the_idom = #idom oo the_rules val the_ideal = #ideal oo the_rules fun match ctxt tm = let fun match_inst ({vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal}, fns) pat = let fun h instT = let val substT = Thm.instantiate (instT, Vars.empty); val substT_cterm = Drule.cterm_rule substT; val vars' = map substT_cterm vars; val semiring' = (map substT_cterm sr_ops, map substT sr_rules); val ring' = (map substT_cterm r_ops, map substT r_rules); val field' = (map substT_cterm f_ops, map substT f_rules); val idom' = map substT idom; val ideal' = map substT ideal; val result = ({vars = vars', semiring = semiring', ring = ring', field = field', idom = idom', ideal = ideal'}, fns); in SOME result end in (case try Thm.match (pat, tm) of NONE => NONE | SOME (instT, _) => h instT) end; fun match_struct (_, entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) = get_first (match_inst entry) (sr_ops @ r_ops @ f_ops); in get_first match_struct (Data.get (Context.Proof ctxt)) end; (* extra-logical functions *) val semiring_norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms semiring_norm}); val semiring_funs = {is_const = can HOLogic.dest_number o Thm.term_of, dest_const = (fn ct => Rat.of_int (snd (HOLogic.dest_number (Thm.term_of ct) handle TERM _ => error "ring_dest_const"))), mk_const = (fn cT => fn x => Numeral.mk_cnumber cT (case Rat.dest x of (i, 1) => i | _ => error "int_of_rat: bad int")), conv = (fn ctxt => Simplifier.rewrite (put_simpset semiring_norm_ss ctxt) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms numeral_One}))}; val divide_const = Thm.cterm_of \<^context> (Logic.varify_global \<^term>\(/)\); val [divide_tvar] = Term.add_tvars (Thm.term_of divide_const) []; val field_funs = let fun numeral_is_const ct = case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b => can HOLogic.dest_number a andalso can HOLogic.dest_number b | Const (\<^const_name>\Fields.inverse\,_)$t => can HOLogic.dest_number t | t => can HOLogic.dest_number t fun dest_const ct = ((case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b=> Rat.make (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | Const (\<^const_name>\Fields.inverse\,_)$t => Rat.inv (Rat.of_int (snd (HOLogic.dest_number t))) | t => Rat.of_int (snd (HOLogic.dest_number t))) handle TERM _ => error "ring_dest_const") fun mk_const cT x = let val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber cT a else Thm.apply (Thm.apply (Thm.instantiate_cterm (TVars.make1 (divide_tvar, cT), Vars.empty) divide_const) (Numeral.mk_cnumber cT a)) (Numeral.mk_cnumber cT b) end in {is_const = numeral_is_const, dest_const = dest_const, mk_const = mk_const, conv = Numeral_Simprocs.field_comp_conv} end; (* logical content *) val semiringN = "semiring"; val ringN = "ring"; val fieldN = "field"; val idomN = "idom"; fun declare raw_key {semiring = raw_semiring0, ring = raw_ring0, field = raw_field0, idom = raw_idom, ideal = raw_ideal} lthy = let val ctxt' = fold Proof_Context.augment (fst raw_semiring0 @ fst raw_ring0 @ fst raw_field0) lthy; val prepare_ops = apfst (Variable.export_terms ctxt' lthy #> map (Thm.cterm_of lthy)); val raw_semiring = prepare_ops raw_semiring0; val raw_ring = prepare_ops raw_ring0; val raw_field = prepare_ops raw_field0; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} (fn phi => fn context => let val ctxt = Context.proof_of context; val key = Morphism.thm phi raw_key; fun transform_ops_rules (ops, rules) = (map (Morphism.cterm phi) ops, Morphism.fact phi rules); val (sr_ops, sr_rules) = transform_ops_rules raw_semiring; val (r_ops, r_rules) = transform_ops_rules raw_ring; val (f_ops, f_rules) = transform_ops_rules raw_field; val idom = Morphism.fact phi raw_idom; val ideal = Morphism.fact phi raw_ideal; fun check kind name xs n = null xs orelse length xs = n orelse error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name); val check_ops = check "operations"; val check_rules = check "rules"; val _ = check_ops semiringN sr_ops 5 andalso check_rules semiringN sr_rules 36 andalso check_ops ringN r_ops 2 andalso check_rules ringN r_rules 2 andalso check_ops fieldN f_ops 2 andalso check_rules fieldN f_rules 2 andalso check_rules idomN idom 2; val mk_meta = Local_Defs.meta_rewrite_rule ctxt; val sr_rules' = map mk_meta sr_rules; val r_rules' = map mk_meta r_rules; val f_rules' = map mk_meta f_rules; fun rule i = nth sr_rules' (i - 1); val (cx, cy) = Thm.dest_binop (hd sr_ops); val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val ((clx, crx), (cly, cry)) = rule 13 |> Thm.rhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val ((ca, cb), (cc, cd)) = rule 20 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg; val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_arg; val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry]; val semiring = (sr_ops, sr_rules'); val ring = (r_ops, r_rules'); val field = (f_ops, f_rules'); val ideal' = map (Thm.symmetric o mk_meta) ideal in context |> Data.map (AList.update Thm.eq_thm (key, ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal'}, (if null f_ops then semiring_funs else field_funs)))) end) end; (** auxiliary **) fun is_comb ct = (case Thm.term_of ct of _ $ _ => true | _ => false); val concl = Thm.cprop_of #> Thm.dest_arg; fun is_binop ct ct' = (case Thm.term_of ct' of c $ _ $ _ => Thm.term_of ct aconv c | _ => false); fun dest_binop ct ct' = if is_binop ct ct' then Thm.dest_binop ct' else raise CTERM ("dest_binop: bad binop", [ct, ct']) fun inst_thm inst = Thm.instantiate (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) inst)); val dest_number = Thm.term_of #> HOLogic.dest_number #> snd; val is_number = can dest_number; fun numeral01_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One}]); fun zero1_numeral_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One} RS sym]); fun zerone_conv ctxt cv = zero1_numeral_conv ctxt then_conv cv then_conv numeral01_conv ctxt; val nat_add_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms arith_simps} @ @{thms diff_nat_numeral} @ @{thms rel_simps} @ @{thms if_False if_True Nat.add_0 add_Suc add_numeral_left Suc_eq_plus1} @ map (fn th => th RS sym) @{thms numerals}); fun nat_add_conv ctxt = zerone_conv ctxt (Simplifier.rewrite (put_simpset nat_add_ss ctxt)); val zeron_tm = \<^cterm>\0::nat\; val onen_tm = \<^cterm>\1::nat\; val true_tm = \<^cterm>\True\; (** normalizing conversions **) (* core conversion *) fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) (f_ops, f_rules) (is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) = let val [pthm_02, pthm_03, pthm_04, pthm_05, pthm_07, pthm_08, pthm_09, pthm_10, pthm_11, pthm_12, pthm_13, pthm_14, pthm_15, pthm_16, pthm_17, pthm_18, pthm_19, pthm_21, pthm_22, pthm_23, pthm_24, pthm_25, pthm_26, pthm_27, pthm_28, pthm_29, pthm_30, pthm_31, pthm_32, pthm_33, pthm_34, pthm_35, pthm_36, pthm_37, pthm_38, _] = sr_rules; val [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry] = vars; val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; val [add_tm, mul_tm, pow_tm] = map (Thm.dest_fun o Thm.dest_fun) [add_pat, mul_pat, pow_pat]; val dest_add = dest_binop add_tm val dest_mul = dest_binop mul_tm fun dest_pow tm = let val (l,r) = dest_binop pow_tm tm in if is_number r then (l,r) else raise CTERM ("dest_pow",[tm]) end; val is_add = is_binop add_tm val is_mul = is_binop mul_tm val (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, cx', cy') = (case (r_ops, r_rules) of ([sub_pat, neg_pat], [neg_mul, sub_add]) => let val sub_tm = Thm.dest_fun (Thm.dest_fun sub_pat) val neg_tm = Thm.dest_fun neg_pat val dest_sub = dest_binop sub_tm in (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, neg_mul |> concl |> Thm.dest_arg, sub_add |> concl |> Thm.dest_arg |> Thm.dest_arg) end | _ => (TrueI, TrueI, true_tm, true_tm, (fn t => (t,t)), true_tm, true_tm)); val (divide_inverse, divide_tm, inverse_tm) = (case (f_ops, f_rules) of ([divide_pat, inverse_pat], [div_inv, _]) => let val div_tm = funpow 2 Thm.dest_fun divide_pat val inv_tm = Thm.dest_fun inverse_pat in (div_inv, div_tm, inv_tm) end | _ => (TrueI, true_tm, true_tm)); in fn variable_ord => let (* Conversion for "x^n * x^m", with either x^n = x and/or x^m = x possible. *) (* Also deals with "const * const", but both terms must involve powers of *) (* the same variable, or both be constants, or behaviour may be incorrect. *) fun powvar_mul_conv ctxt tm = let val (l,r) = dest_mul tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_mul_conv tm else ((let val (lx,ln) = dest_pow l in ((let val (_, rn) = dest_pow r val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => (let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end)) end) handle CTERM _ => ((let val (rx,rn) = dest_pow r val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => inst_thm [(cx,l)] pthm_32 )) end; (* Remove "1 * m" from a monomial, and just leave m. *) fun monomial_deone th = (let val (l,r) = dest_mul(concl th) in if l aconvc one_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_13) else th end) handle CTERM _ => th; (* Conversion for "(monomial)^n", where n is a numeral. *) fun monomial_pow_conv ctxt = let fun monomial_pow tm bod ntm = if not(is_comb bod) then Thm.reflexive tm else if is_semiring_constant bod then semiring_pow_conv tm else let val (lopr,r) = Thm.dest_comb bod in if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 val (l,r) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv ctxt r)) end else if opr aconvc mul_tm then let val th1 = inst_thm [(cx,l),(cy,r),(cq,ntm)] pthm_33 val (xy,z) = Thm.dest_comb(concl th1) val (x,y) = Thm.dest_comb xy val thl = monomial_pow y l ntm val thr = monomial_pow z r ntm in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule x thl) thr) end else Thm.reflexive tm end end in fn tm => let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if not (opr aconvc pow_tm) orelse not(is_number r) then raise CTERM ("monomial_pow_conv", [tm]) else if r aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if r aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else monomial_deone(monomial_pow tm l r) end end; (* Multiplication of canonical monomials. *) fun monomial_mul_conv ctxt = let fun powvar tm = if is_semiring_constant tm then one_tm else ((let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then l else raise CTERM ("monomial_mul_conv",[tm]) end) handle CTERM _ => tm) (* FIXME !? *) fun vorder x y = if x aconvc y then 0 else if x aconvc one_tm then ~1 else if y aconvc one_tm then 1 else if is_less (variable_ord (x, y)) then ~1 else 1 fun monomial_mul tm l r = ((let val (lx,ly) = dest_mul l val vl = powvar lx in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 val th3 = Thm.transitive th1 th2 val (tm5,tm6) = Thm.dest_comb(concl th3) val (tm7,tm8) = Thm.dest_comb tm6 val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8 in Thm.transitive th3 (Drule.arg_cong_rule tm5 th4) end else let val th0 = if ord < 0 then pthm_16 else pthm_17 val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 in Thm.transitive th1 th2 end else if ord < 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else inst_thm [(ca,l),(cb,r)] pthm_09 end)) end) handle CTERM _ => (let val vl = powvar l in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2) end else if ord > 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else Thm.reflexive tm end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then powvar_mul_conv ctxt tm else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 else Thm.reflexive tm end)) end)) in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r) end end; (* Multiplication by monomial of a polynomial. *) fun polynomial_monomial_mul_conv ctxt = let fun pmm_conv tm = let val (l,r) = dest_mul tm in ((let val (y,z) = dest_add r val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv ctxt tm4)) (pmm_conv tm2) in Thm.transitive th1 th2 end) handle CTERM _ => monomial_mul_conv ctxt tm) end in pmm_conv end; (* Addition of two monomials identical except for constant multiples. *) fun monomial_add_conv tm = let val (l,r) = dest_add tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_add_conv tm else let val th1 = if is_mul l andalso is_semiring_constant(Thm.dest_arg1 l) then if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(ca,Thm.dest_arg1 l),(cm,Thm.dest_arg r), (cb,Thm.dest_arg1 r)] pthm_02 else inst_thm [(ca,Thm.dest_arg1 l),(cm,r)] pthm_03 else if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(cm,l),(ca,Thm.dest_arg1 r)] pthm_04 else inst_thm [(cm,r)] pthm_05 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4) val th3 = Thm.transitive th1 (Drule.fun_cong_rule th2 tm2) val tm5 = concl th3 in if (Thm.dest_arg1 tm5) aconvc zero_tm then Thm.transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11) else monomial_deone th3 end end; (* Ordering on monomials. *) fun striplist dest = let fun strip x acc = ((let val (l,r) = dest x in strip l (strip r acc) end) handle CTERM _ => x::acc) (* FIXME !? *) in fn x => strip x [] end; fun powervars tm = let val ptms = striplist dest_mul tm in if is_semiring_constant (hd ptms) then tl ptms else ptms end; val num_0 = 0; val num_1 = 1; fun dest_varpow tm = ((let val (x,n) = dest_pow tm in (x,dest_number n) end) handle CTERM _ => (tm,(if is_semiring_constant tm then num_0 else num_1))); val morder = let fun lexorder ls = case ls of ([],[]) => 0 | (_ ,[]) => ~1 | ([], _) => 1 | (((x1,n1)::vs1),((x2,n2)::vs2)) => (case variable_ord (x1, x2) of LESS => 1 | GREATER => ~1 | EQUAL => if n1 < n2 then ~1 else if n2 < n1 then 1 else lexorder (vs1, vs2)) in fn tm1 => fn tm2 => let val vdegs1 = map dest_varpow (powervars tm1) val vdegs2 = map dest_varpow (powervars tm2) val deg1 = fold (Integer.add o snd) vdegs1 num_0 val deg2 = fold (Integer.add o snd) vdegs2 num_0 in if deg1 < deg2 then ~1 else if deg1 > deg2 then 1 else lexorder (vdegs1, vdegs2) end end; (* Addition of two polynomials. *) fun polynomial_add_conv ctxt = let fun dezero_rule th = let val tm = concl th in if not(is_add tm) then th else let val (lopr,r) = Thm.dest_comb tm val l = Thm.dest_arg lopr in if l aconvc zero_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_07) else if r aconvc zero_tm then Thm.transitive th (inst_thm [(ca,l)] pthm_08) else th end end fun padd tm = let val (l,r) = dest_add tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_07 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_08 else if is_add l then let val (a,b) = dest_add l in if is_add r then let val (c,d) = dest_add r val ord = morder a c in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,c),(cd,d)] pthm_23 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4) in dezero_rule (Thm.transitive th1 (Thm.combination th2 (padd tm2))) end else (* ord <> 0*) let val th1 = if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else (* not (is_add r)*) let val ord = morder a r in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_26 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else (* ord <> 0*) if ord > 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end else (* not (is_add l)*) if is_add r then let val (c,d) = dest_add r val ord = morder l c in if ord = 0 then let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_28 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else if ord > 0 then Thm.reflexive tm else let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else let val ord = morder l r in if ord = 0 then monomial_add_conv tm else if ord > 0 then dezero_rule(Thm.reflexive tm) else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end in padd end; (* Multiplication of two polynomials. *) fun polynomial_mul_conv ctxt = let fun pmul tm = let val (l,r) = dest_mul tm in if not(is_add l) then polynomial_monomial_mul_conv ctxt tm else if not(is_add r) then let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 in Thm.transitive th1 (polynomial_monomial_mul_conv ctxt (concl th1)) end else let val (a,b) = dest_add l val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv ctxt tm4) val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2)) in Thm.transitive th3 (polynomial_add_conv ctxt (concl th3)) end end in fn tm => let val (l,r) = dest_mul tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_11 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_12 else if l aconvc one_tm then inst_thm [(ca,r)] pthm_13 else if r aconvc one_tm then inst_thm [(ca,l)] pthm_14 else pmul tm end end; (* Power of polynomial (optimized for the monomial and trivial cases). *) fun num_conv ctxt n = nat_add_conv ctxt (Thm.apply \<^cterm>\Suc\ (Numeral.mk_cnumber \<^ctyp>\nat\ (dest_number n - 1))) |> Thm.symmetric; fun polynomial_pow_conv ctxt = let fun ppow tm = let val (l,n) = dest_pow tm in if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else let val th1 = num_conv ctxt n val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 val (tm1,tm2) = Thm.dest_comb(concl th2) val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 in Thm.transitive th4 (polynomial_mul_conv ctxt (concl th4)) end end in fn tm => if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv ctxt tm end; (* Negation. *) fun polynomial_neg_conv ctxt tm = let val (l,r) = Thm.dest_comb tm in if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else let val th1 = inst_thm [(cx', r)] neg_mul val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1)) in Thm.transitive th2 (polynomial_monomial_mul_conv ctxt (concl th2)) end end; (* Subtraction. *) fun polynomial_sub_conv ctxt tm = let val (l,r) = dest_sub tm val th1 = inst_thm [(cx', l), (cy', r)] sub_add val (tm1,tm2) = Thm.dest_comb(concl th1) val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv ctxt tm2) in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv ctxt (concl th2))) end; (* Conversion from HOL term. *) fun polynomial_conv ctxt tm = if is_semiring_constant tm then semiring_add_conv tm else if not(is_comb tm) then Thm.reflexive tm else let val (lopr,r) = Thm.dest_comb tm in if lopr aconvc neg_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (polynomial_neg_conv ctxt (concl th1)) end else if lopr aconvc inverse_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (semiring_mul_conv (concl th1)) end else if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) r in Thm.transitive th1 (polynomial_pow_conv ctxt (concl th1)) end else if opr aconvc divide_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv ctxt) (Thm.rhs_of th1) in Thm.transitive th1 th2 end else if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val f = if opr aconvc add_tm then polynomial_add_conv ctxt else if opr aconvc mul_tm then polynomial_mul_conv ctxt else polynomial_sub_conv ctxt in Thm.transitive th1 (f (concl th1)) end else Thm.reflexive tm end end; in {main = polynomial_conv, add = polynomial_add_conv, mul = polynomial_mul_conv, pow = polynomial_pow_conv, neg = polynomial_neg_conv, sub = polynomial_sub_conv} end end; val nat_exp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps}) addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]); (* various normalizing conversions *) fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = let val pow_conv = Conv.arg_conv (Simplifier.rewrite (put_simpset nat_exp_ss ctxt)) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) then_conv conv ctxt val dat = (is_const, conv ctxt, conv ctxt, pow_conv) in semiring_normalizers_conv vars semiring ring field dat term_ord end; fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = #main (semiring_normalizers_ord_wrapper ctxt ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal}, {conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) term_ord) ctxt; fun semiring_normalize_wrapper ctxt data = semiring_normalize_ord_wrapper ctxt data Thm.term_ord; fun semiring_normalize_ord_conv ctxt ord tm = (case match ctxt tm of NONE => Thm.reflexive tm | SOME res => semiring_normalize_ord_wrapper ctxt res ord tm); fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt Thm.term_ord; end; diff --git a/src/HOL/Tools/typedef.ML b/src/HOL/Tools/typedef.ML --- a/src/HOL/Tools/typedef.ML +++ b/src/HOL/Tools/typedef.ML @@ -1,373 +1,373 @@ (* Title: HOL/Tools/typedef.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen Gordon/HOL-style type definitions: create a new syntactic type represented by a non-empty set. *) signature TYPEDEF = sig type info = {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} * {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, Rep_induct: thm, Abs_induct: thm} val transform_info: morphism -> info -> info val get_info: Proof.context -> string -> info list val get_info_global: theory -> string -> info list val interpretation: (string -> local_theory -> local_theory) -> theory -> theory type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding} val default_bindings: binding -> bindings val make_bindings: binding -> bindings option -> bindings val make_morphisms: binding -> (binding * binding) option -> bindings val overloaded: bool Config.T val add_typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> (Proof.context -> tactic) -> local_theory -> (string * info) * local_theory val add_typedef_global: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> (Proof.context -> tactic) -> theory -> (string * info) * theory val typedef: {overloaded: bool} -> binding * (string * sort) list * mixfix -> term -> bindings option -> local_theory -> Proof.state val typedef_cmd: {overloaded: bool} -> binding * (string * string option) list * mixfix -> string -> bindings option -> local_theory -> Proof.state end; structure Typedef: TYPEDEF = struct (** type definitions **) (* theory data *) type info = (*global part*) {rep_type: typ, abs_type: typ, Rep_name: string, Abs_name: string, axiom_name: string} * (*local part*) {inhabited: thm, type_definition: thm, Rep: thm, Rep_inverse: thm, Abs_inverse: thm, Rep_inject: thm, Abs_inject: thm, Rep_cases: thm, Abs_cases: thm, Rep_induct: thm, Abs_induct: thm}; fun transform_info phi (info: info) = let val thm = Morphism.thm phi; val (global_info, {inhabited, type_definition, Rep, Rep_inverse, Abs_inverse, Rep_inject, Abs_inject, Rep_cases, Abs_cases, Rep_induct, Abs_induct}) = info; in (global_info, {inhabited = thm inhabited, type_definition = thm type_definition, Rep = thm Rep, Rep_inverse = thm Rep_inverse, Abs_inverse = thm Abs_inverse, Rep_inject = thm Rep_inject, Abs_inject = thm Abs_inject, Rep_cases = thm Rep_cases, Abs_cases = thm Abs_cases, Rep_induct = thm Rep_induct, Abs_induct = thm Abs_induct}) end; structure Data = Generic_Data ( type T = info list Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge_list (K true) data; ); fun get_info_generic context = Symtab.lookup_list (Data.get context) #> map (transform_info (Morphism.transfer_morphism'' context)); val get_info = get_info_generic o Context.Proof; val get_info_global = get_info_generic o Context.Theory; fun put_info name info = Data.map (Symtab.cons_list (name, transform_info Morphism.trim_context_morphism info)); (* global interpretation *) structure Typedef_Plugin = Plugin(type T = string); val typedef_plugin = Plugin_Name.declare_setup \<^binding>\typedef\; fun interpretation f = Typedef_Plugin.interpretation typedef_plugin (fn name => fn lthy => lthy |> Local_Theory.map_background_naming (Name_Space.root_path #> Name_Space.add_path (Long_Name.qualifier name)) |> f name |> Local_Theory.restore_background_naming lthy); (* primitive typedef axiomatization -- for fresh typedecl *) val typedef_overloaded = Attrib.setup_config_bool \<^binding>\typedef_overloaded\ (K false); fun mk_inhabited A = let val T = HOLogic.dest_setT (Term.fastype_of A) in HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, HOLogic.mk_mem (Bound 0, A))) end; fun mk_typedef newT oldT RepC AbsC A = let val typedefC = Const (\<^const_name>\type_definition\, (newT --> oldT) --> (oldT --> newT) --> HOLogic.mk_setT oldT --> HOLogic.boolT); in Logic.mk_implies (mk_inhabited A, HOLogic.mk_Trueprop (typedefC $ RepC $ AbsC $ A)) end; fun primitive_typedef {overloaded} type_definition_name newT oldT Rep_name Abs_name A lthy = let (* errors *) fun show_names pairs = commas_quote (map fst pairs); val lhs_tfrees = Term.add_tfreesT newT []; val rhs_tfrees = Term.add_tfreesT oldT []; val _ = (case fold (remove (op =)) lhs_tfrees rhs_tfrees of [] => () | extras => error ("Extra type variables in representing set: " ^ show_names extras)); val _ = (case Term.add_frees A [] of [] => [] | xs => error ("Illegal variables in representing set: " ^ show_names xs)); (* axiomatization *) val ((RepC, AbsC), consts_lthy) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((Rep_name, newT --> oldT), NoSyn) ##>> Sign.declare_const lthy ((Abs_name, oldT --> newT), NoSyn)); val const_dep = Theory.const_dep (Proof_Context.theory_of consts_lthy); val defs_context = Proof_Context.defs_context consts_lthy; val A_consts = fold_aterms (fn Const c => insert (op =) (const_dep c) | _ => I) A []; val A_types = (fold_types o fold_subtypes) (fn Type t => insert (op =) (Theory.type_dep t) | _ => I) A []; val typedef_deps = A_consts @ A_types; val newT_dep = Theory.type_dep (dest_Type newT); val ((axiom_name, axiom), axiom_lthy) = consts_lthy |> Local_Theory.background_theory_result (Thm.add_axiom consts_lthy (type_definition_name, mk_typedef newT oldT RepC AbsC A) ##> Theory.add_deps defs_context "" newT_dep typedef_deps ##> Theory.add_deps defs_context "" (const_dep (dest_Const RepC)) [newT_dep] ##> Theory.add_deps defs_context "" (const_dep (dest_Const AbsC)) [newT_dep]); val axiom_defs = Theory.defs_of (Proof_Context.theory_of axiom_lthy); val newT_deps = maps #2 (Defs.get_deps axiom_defs (#1 newT_dep)); val _ = if null newT_deps orelse overloaded orelse Config.get lthy typedef_overloaded then () else error (Pretty.string_of (Pretty.chunks [Pretty.paragraph (Pretty.text "Type definition with open dependencies, use" @ [Pretty.brk 1, Pretty.str "\"", Pretty.keyword1 "typedef", Pretty.brk 1, Pretty.str "(", Pretty.keyword2 "overloaded", Pretty.str ")\"", Pretty.brk 1] @ Pretty.text "or enable configuration option \"typedef_overloaded\" in the context."), Pretty.block [Pretty.str " Type:", Pretty.brk 2, Syntax.pretty_typ axiom_lthy newT], Pretty.block (Pretty.str " Deps:" :: Pretty.brk 2 :: Pretty.commas (map (Defs.pretty_entry (Proof_Context.defs_context axiom_lthy)) newT_deps))])) in ((RepC, AbsC, axiom_name, axiom), axiom_lthy) end; (* derived bindings *) type bindings = {Rep_name: binding, Abs_name: binding, type_definition_name: binding}; fun prefix_binding prfx name = Binding.reset_pos (Binding.qualify_name false name (prfx ^ Binding.name_of name)); fun qualify_binding name = Binding.qualify false (Binding.name_of name); fun default_bindings name = {Rep_name = prefix_binding "Rep_" name, Abs_name = prefix_binding "Abs_" name, type_definition_name = prefix_binding "type_definition_" name}; fun make_bindings name NONE = default_bindings name | make_bindings _ (SOME bindings) = bindings; fun make_morphisms name NONE = default_bindings name | make_morphisms name (SOME (Rep_name, Abs_name)) = {Rep_name = qualify_binding name Rep_name, Abs_name = qualify_binding name Abs_name, type_definition_name = #type_definition_name (default_bindings name)}; (* prepare_typedef *) fun prepare_typedef prep_term overloaded (name, raw_args, mx) raw_set opt_bindings lthy = let (* rhs *) val tmp_ctxt = lthy |> fold (Variable.declare_typ o TFree) raw_args; val set = prep_term tmp_ctxt raw_set; val tmp_ctxt' = tmp_ctxt |> Variable.declare_term set; val setT = Term.fastype_of set; val oldT = HOLogic.dest_setT setT handle TYPE _ => error ("Not a set type: " ^ quote (Syntax.string_of_typ lthy setT)); val bname = Binding.name_of name; val goal = mk_inhabited set; val goal_pat = mk_inhabited (Var (the_default (bname, 0) (Lexicon.read_variable bname), setT)); (* lhs *) val args = map (Proof_Context.check_tfree tmp_ctxt') raw_args; val (newT, typedecl_lthy) = lthy |> Typedecl.typedecl {final = false} (name, args, mx) ||> Variable.declare_term set; val Type (full_name, _) = newT; (* axiomatization *) val {Rep_name, Abs_name, type_definition_name} = make_bindings name opt_bindings; val ((RepC, AbsC, axiom_name, typedef), typedef_lthy) = typedecl_lthy |> primitive_typedef overloaded type_definition_name newT oldT Rep_name Abs_name set; val alias_lthy = typedef_lthy |> Local_Theory.const_alias Rep_name (#1 (Term.dest_Const RepC)) |> Local_Theory.const_alias Abs_name (#1 (Term.dest_Const AbsC)); (* result *) fun note ((b, atts), th) = Local_Theory.note ((b, atts), [th]) #>> (fn (_, [th']) => th'); fun typedef_result inhabited lthy1 = let val ((_, [type_definition]), lthy2) = lthy1 |> Local_Theory.note ((type_definition_name, []), [inhabited RS typedef]); fun make th = Goal.norm_result lthy2 (type_definition RS th); val (((((((((Rep, Rep_inverse), Abs_inverse), Rep_inject), Abs_inject), Rep_cases), Abs_cases), Rep_induct), Abs_induct), lthy3) = lthy2 |> note ((Rep_name, []), make @{thm type_definition.Rep}) ||>> note ((Binding.suffix_name "_inverse" Rep_name, []), make @{thm type_definition.Rep_inverse}) ||>> note ((Binding.suffix_name "_inverse" Abs_name, []), make @{thm type_definition.Abs_inverse}) ||>> note ((Binding.suffix_name "_inject" Rep_name, []), make @{thm type_definition.Rep_inject}) ||>> note ((Binding.suffix_name "_inject" Abs_name, []), make @{thm type_definition.Abs_inject}) ||>> note ((Binding.suffix_name "_cases" Rep_name, [Attrib.case_names [Binding.name_of Rep_name], - Attrib.internal (K (Induct.cases_pred full_name))]), + Attrib.internal \<^here> (K (Induct.cases_pred full_name))]), make @{thm type_definition.Rep_cases}) ||>> note ((Binding.suffix_name "_cases" Abs_name, [Attrib.case_names [Binding.name_of Abs_name], - Attrib.internal (K (Induct.cases_type full_name))]), + Attrib.internal \<^here> (K (Induct.cases_type full_name))]), make @{thm type_definition.Abs_cases}) ||>> note ((Binding.suffix_name "_induct" Rep_name, [Attrib.case_names [Binding.name_of Rep_name], - Attrib.internal (K (Induct.induct_pred full_name))]), + Attrib.internal \<^here> (K (Induct.induct_pred full_name))]), make @{thm type_definition.Rep_induct}) ||>> note ((Binding.suffix_name "_induct" Abs_name, [Attrib.case_names [Binding.name_of Abs_name], - Attrib.internal (K (Induct.induct_type full_name))]), + Attrib.internal \<^here> (K (Induct.induct_type full_name))]), make @{thm type_definition.Abs_induct}); val info = ({rep_type = oldT, abs_type = newT, Rep_name = #1 (Term.dest_Const RepC), Abs_name = #1 (Term.dest_Const AbsC), axiom_name = axiom_name}, {inhabited = inhabited, type_definition = type_definition, Rep = Rep, Rep_inverse = Rep_inverse, Abs_inverse = Abs_inverse, Rep_inject = Rep_inject, Abs_inject = Abs_inject, Rep_cases = Rep_cases, Abs_cases = Abs_cases, Rep_induct = Rep_induct, Abs_induct = Abs_induct}); in lthy3 - |> Local_Theory.declaration {syntax = false, pervasive = true} + |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => put_info full_name (transform_info phi info)) |> Typedef_Plugin.data Plugin_Name.default_filter full_name |> pair (full_name, info) end; in ((goal, goal_pat, typedef_result), alias_lthy) end handle ERROR msg => cat_error msg ("The error(s) above occurred in typedef " ^ Binding.print name); (* add_typedef: tactic interface *) fun add_typedef overloaded typ set opt_bindings tac lthy = let val ((goal, _, typedef_result), lthy') = prepare_typedef Syntax.check_term overloaded typ set opt_bindings lthy; val inhabited = Goal.prove lthy' [] [] goal (tac o #context) |> Goal.norm_result lthy'; in typedef_result inhabited lthy' end; fun add_typedef_global overloaded typ set opt_bindings tac = Named_Target.theory_map_result (apsnd o transform_info) (add_typedef overloaded typ set opt_bindings tac) (* typedef: proof interface *) local fun gen_typedef prep_term prep_constraint overloaded (b, raw_args, mx) set opt_bindings lthy = let val args = map (apsnd (prep_constraint lthy)) raw_args; val ((goal, goal_pat, typedef_result), lthy') = prepare_typedef prep_term overloaded (b, args, mx) set opt_bindings lthy; fun after_qed [[th]] = snd o typedef_result th; in Proof.theorem NONE after_qed [[(goal, [goal_pat])]] lthy' end; in val typedef = gen_typedef Syntax.check_term (K I); val typedef_cmd = gen_typedef Syntax.read_term Typedecl.read_constraint; end; (** outer syntax **) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\typedef\ "HOL type definition (requires non-emptiness proof)" (Parse_Spec.overloaded -- Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- Parse.term) -- Scan.option (\<^keyword>\morphisms\ |-- Parse.!!! (Parse.binding -- Parse.binding)) >> (fn (((((overloaded, vs), t), mx), A), opt_morphs) => fn lthy => typedef_cmd {overloaded = overloaded} (t, vs, mx) A (SOME (make_morphisms t opt_morphs)) lthy)); val overloaded = typedef_overloaded; (** theory export **) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parent_spaces = map Sign.type_space (Theory.parents_of thy); val typedefs = Name_Space.dest_table (#types (Type.rep_tsig (Sign.tsig_of thy))) |> maps (fn (name, _) => if exists (fn space => Name_Space.declared space name) parent_spaces then [] else get_info_global thy name |> map (fn ({rep_type, abs_type, Rep_name, Abs_name, axiom_name}, _) => (name, (rep_type, (abs_type, (Rep_name, (Abs_name, axiom_name))))))); val encode = let open XML.Encode Term_XML.Encode in list (pair string (pair typ (pair typ (pair string (pair string string))))) end; in if null typedefs then () else Export_Theory.export_body thy "typedefs" (encode typedefs) end else ()); end; diff --git a/src/Provers/order_tac.ML b/src/Provers/order_tac.ML --- a/src/Provers/order_tac.ML +++ b/src/Provers/order_tac.ML @@ -1,434 +1,435 @@ signature REIFY_TABLE = sig type table val empty : table val get_var : term -> table -> (int * table) val get_term : int -> table -> term option end structure Reifytab: REIFY_TABLE = struct type table = (int * (term * int) list) val empty = (0, []) fun get_var t (max_var, tis) = (case AList.lookup Envir.aeconv tis t of SOME v => (v, (max_var, tis)) | NONE => (max_var, (max_var + 1, (t, max_var) :: tis)) ) fun get_term v (_, tis) = Library.find_first (fn (_, v2) => v = v2) tis |> Option.map fst end signature LOGIC_SIGNATURE = sig val mk_Trueprop : term -> term val dest_Trueprop : term -> term val Trueprop_conv : conv -> conv val Not : term val conj : term val disj : term val notI : thm (* (P \ False) \ \ P *) val ccontr : thm (* (\ P \ False) \ P *) val conjI : thm (* P \ Q \ P \ Q *) val conjE : thm (* P \ Q \ (P \ Q \ R) \ R *) val disjE : thm (* P \ Q \ (P \ R) \ (Q \ R) \ R *) val not_not_conv : conv (* \ (\ P) \ P *) val de_Morgan_conj_conv : conv (* \ (P \ Q) \ \ P \ \ Q *) val de_Morgan_disj_conv : conv (* \ (P \ Q) \ \ P \ \ Q *) val conj_disj_distribL_conv : conv (* P \ (Q \ R) \ (P \ Q) \ (P \ R) *) val conj_disj_distribR_conv : conv (* (Q \ R) \ P \ (Q \ P) \ (R \ P) *) end (* Control tracing output of the solver. *) val order_trace_cfg = Attrib.setup_config_bool @{binding "order_trace"} (K false) (* In partial orders, literals of the form \ x < y will force the order solver to perform case distinctions, which leads to an exponential blowup of the runtime. The split limit controls the number of literals of this form that are passed to the solver. *) val order_split_limit_cfg = Attrib.setup_config_int @{binding "order_split_limit"} (K 8) datatype order_kind = Order | Linorder type order_literal = (bool * Order_Procedure.order_atom) type order_context = { kind : order_kind, ops : term list, thms : (string * thm) list, conv_thms : (string * thm) list } signature BASE_ORDER_TAC = sig val tac : (order_literal Order_Procedure.fm -> Order_Procedure.prf_trm option) -> order_context -> thm list -> Proof.context -> int -> tactic end functor Base_Order_Tac( structure Logic_Sig : LOGIC_SIGNATURE; val excluded_types : typ list) : BASE_ORDER_TAC = struct open Order_Procedure fun expect _ (SOME x) = x | expect f NONE = f () fun list_curry0 f = (fn [] => f, 0) fun list_curry1 f = (fn [x] => f x, 1) fun list_curry2 f = (fn [x, y] => f x y, 2) fun dereify_term consts reifytab t = let fun dereify_term' (App (t1, t2)) = (dereify_term' t1) $ (dereify_term' t2) | dereify_term' (Const s) = AList.lookup (op =) consts s |> expect (fn () => raise TERM ("Const " ^ s ^ " not in", map snd consts)) | dereify_term' (Var v) = Reifytab.get_term (integer_of_int v) reifytab |> the in dereify_term' t end fun dereify_order_fm (eq, le, lt) reifytab t = let val consts = [ ("eq", eq), ("le", le), ("lt", lt), ("Not", Logic_Sig.Not), ("disj", Logic_Sig.disj), ("conj", Logic_Sig.conj) ] in dereify_term consts reifytab t end fun strip_AppP t = let fun strip (AppP (f, s), ss) = strip (f, s::ss) | strip x = x in strip (t, []) end fun replay_conv convs cvp = let val convs = convs @ [("all_conv", list_curry0 Conv.all_conv)] @ map (apsnd list_curry1) [ ("atom_conv", I), ("neg_atom_conv", I), ("arg_conv", Conv.arg_conv)] @ map (apsnd list_curry2) [ ("combination_conv", Conv.combination_conv), ("then_conv", curry (op then_conv))] fun lookup_conv convs c = AList.lookup (op =) convs c |> expect (fn () => error ("Can't replay conversion: " ^ c)) fun rp_conv t = (case strip_AppP t ||> map rp_conv of (PThm c, cvs) => let val (conv, arity) = lookup_conv convs c in if arity = length cvs then conv cvs else error ("Expected " ^ Int.toString arity ^ " arguments for conversion " ^ c ^ " but got " ^ (length cvs |> Int.toString) ^ " arguments") end | _ => error "Unexpected constructor in conversion proof") in rp_conv cvp end fun replay_prf_trm replay_conv dereify ctxt thmtab assmtab p = let fun replay_prf_trm' _ (PThm s) = AList.lookup (op =) thmtab s |> expect (fn () => error ("Cannot replay theorem: " ^ s)) | replay_prf_trm' assmtab (Appt (p, t)) = replay_prf_trm' assmtab p |> Drule.infer_instantiate' ctxt [SOME (Thm.cterm_of ctxt (dereify t))] | replay_prf_trm' assmtab (AppP (p1, p2)) = apply2 (replay_prf_trm' assmtab) (p2, p1) |> op COMP | replay_prf_trm' assmtab (AbsP (reified_t, p)) = let val t = dereify reified_t val t_thm = Logic_Sig.mk_Trueprop t |> Thm.cterm_of ctxt |> Assumption.assume ctxt val rp = replay_prf_trm' (Termtab.update (Thm.prop_of t_thm, t_thm) assmtab) p in Thm.implies_intr (Thm.cprop_of t_thm) rp end | replay_prf_trm' assmtab (Bound reified_t) = let val t = dereify reified_t |> Logic_Sig.mk_Trueprop in Termtab.lookup assmtab t |> expect (fn () => raise TERM ("Assumption not found:", t::Termtab.keys assmtab)) end | replay_prf_trm' assmtab (Conv (t, cp, p)) = let val thm = replay_prf_trm' assmtab (Bound t) val conv = Logic_Sig.Trueprop_conv (replay_conv cp) val conv_thm = Conv.fconv_rule conv thm val conv_term = Thm.prop_of conv_thm in replay_prf_trm' (Termtab.update (conv_term, conv_thm) assmtab) p end in replay_prf_trm' assmtab p end fun replay_order_prf_trm ord_ops {thms = thms, conv_thms = conv_thms, ...} ctxt reifytab assmtab = let val thmtab = thms @ [ ("conjE", Logic_Sig.conjE), ("conjI", Logic_Sig.conjI), ("disjE", Logic_Sig.disjE) ] val convs = map (apsnd list_curry0) ( map (apsnd Conv.rewr_conv) conv_thms @ [ ("not_not_conv", Logic_Sig.not_not_conv), ("de_Morgan_conj_conv", Logic_Sig.de_Morgan_conj_conv), ("de_Morgan_disj_conv", Logic_Sig.de_Morgan_disj_conv), ("conj_disj_distribR_conv", Logic_Sig.conj_disj_distribR_conv), ("conj_disj_distribL_conv", Logic_Sig.conj_disj_distribL_conv) ]) val dereify = dereify_order_fm ord_ops reifytab in replay_prf_trm (replay_conv convs) dereify ctxt thmtab assmtab end fun strip_Not (nt $ t) = if nt = Logic_Sig.Not then t else nt $ t | strip_Not t = t fun limit_not_less [_, _, lt] ctxt decomp_prems = let val thy = Proof_Context.theory_of ctxt val trace = Config.get ctxt order_trace_cfg val limit = Config.get ctxt order_split_limit_cfg fun is_not_less_term t = case try (strip_Not o Logic_Sig.dest_Trueprop) t of SOME (binop $ _ $ _) => Pattern.matches thy (lt, binop) | NONE => false val not_less_prems = filter (is_not_less_term o Thm.prop_of o fst) decomp_prems val _ = if trace andalso length not_less_prems > limit then tracing "order split limit exceeded" else () in filter_out (is_not_less_term o Thm.prop_of o fst) decomp_prems @ take limit not_less_prems end fun decomp [eq, le, lt] ctxt t = let fun is_excluded t = exists (fn ty => ty = fastype_of t) excluded_types fun decomp'' (binop $ t1 $ t2) = let open Order_Procedure val thy = Proof_Context.theory_of ctxt fun try_match pat = try (Pattern.match thy (pat, binop)) (Vartab.empty, Vartab.empty) in if is_excluded t1 then NONE else case (try_match eq, try_match le, try_match lt) of (SOME env, _, _) => SOME (true, EQ, (t1, t2), env) | (_, SOME env, _) => SOME (true, LEQ, (t1, t2), env) | (_, _, SOME env) => SOME (true, LESS, (t1, t2), env) | _ => NONE end | decomp'' _ = NONE fun decomp' (nt $ t) = if nt = Logic_Sig.Not then decomp'' t |> Option.map (fn (b, c, p, e) => (not b, c, p, e)) else decomp'' (nt $ t) | decomp' t = decomp'' t in try Logic_Sig.dest_Trueprop t |> Option.mapPartial decomp' end fun maximal_envs envs = let fun test_opt p (SOME x) = p x | test_opt _ NONE = false fun leq_env (tyenv1, tenv1) (tyenv2, tenv2) = Vartab.forall (fn (v, ty) => Vartab.lookup tyenv2 v |> test_opt (fn ty2 => ty2 = ty)) tyenv1 andalso Vartab.forall (fn (v, (ty, t)) => Vartab.lookup tenv2 v |> test_opt (fn (ty2, t2) => ty2 = ty andalso t2 aconv t)) tenv1 fun fold_env (i, env) es = fold_index (fn (i2, env2) => fn es => if i = i2 then es else if leq_env env env2 then (i, i2) :: es else es) envs es val env_order = fold_index fold_env envs [] val graph = fold_index (fn (i, env) => fn g => Int_Graph.new_node (i, env) g) envs Int_Graph.empty val graph = fold Int_Graph.add_edge env_order graph val strong_conns = Int_Graph.strong_conn graph val maximals = filter (fn comp => length comp = length (Int_Graph.all_succs graph comp)) strong_conns in map (Int_Graph.all_preds graph) maximals end fun order_tac raw_order_proc octxt simp_prems = Subgoal.FOCUS (fn {prems=prems, context=ctxt, ...} => let val trace = Config.get ctxt order_trace_cfg fun these' _ [] = [] | these' f (x :: xs) = case f x of NONE => these' f xs | SOME y => (x, y) :: these' f xs val prems = simp_prems @ prems |> filter (fn p => null (Term.add_vars (Thm.prop_of p) [])) |> map (Conv.fconv_rule Thm.eta_conversion) val decomp_prems = these' (decomp (#ops octxt) ctxt o Thm.prop_of) prems fun env_of (_, (_, _, _, env)) = env val env_groups = maximal_envs (map env_of decomp_prems) fun order_tac' (_, []) = no_tac | order_tac' (env, decomp_prems) = let val [eq, le, lt] = #ops octxt |> map (Envir.eta_contract o Envir.subst_term env) val decomp_prems = case #kind octxt of Order => limit_not_less (#ops octxt) ctxt decomp_prems | _ => decomp_prems fun reify_prem (_, (b, ctor, (x, y), _)) (ps, reifytab) = (Reifytab.get_var x ##>> Reifytab.get_var y) reifytab |>> (fn vp => (b, ctor (apply2 Int_of_integer vp)) :: ps) val (reified_prems, reifytab) = fold_rev reify_prem decomp_prems ([], Reifytab.empty) val reified_prems_conj = foldl1 (fn (x, a) => And (x, a)) (map Atom reified_prems) val prems_conj_thm = map fst decomp_prems |> foldl1 (fn (x, a) => Logic_Sig.conjI OF [x, a]) |> Conv.fconv_rule Thm.eta_conversion val prems_conj = prems_conj_thm |> Thm.prop_of val proof = raw_order_proc reified_prems_conj val pretty_term_list = Pretty.list "" "" o map (Syntax.pretty_term (Config.put show_types true ctxt)) val pretty_thm_list = Pretty.list "" "" o map (Thm.pretty_thm ctxt) fun pretty_type_of t = Pretty.block [ Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt (Term.fastype_of t)) ] fun pretty_trace () = [ ("order kind:", Pretty.str (@{make_string} (#kind octxt))) , ("order operators:", Pretty.block [ pretty_term_list [eq, le, lt], Pretty.brk 1 , pretty_type_of le ]) , ("premises:", pretty_thm_list prems) , ("selected premises:", pretty_thm_list (map fst decomp_prems)) , ("reified premises:", Pretty.str (@{make_string} reified_prems)) , ("contradiction:", Pretty.str (@{make_string} (Option.isSome proof))) ] |> map (fn (t, pp) => Pretty.block [Pretty.str t, Pretty.brk 1, pp]) |> Pretty.big_list "order solver called with the parameters" val _ = if trace then tracing (Pretty.string_of (pretty_trace ())) else () val assmtab = Termtab.make [(prems_conj, prems_conj_thm)] val replay = replay_order_prf_trm (eq, le, lt) octxt ctxt reifytab assmtab in case proof of NONE => no_tac | SOME p => SOLVED' (resolve_tac ctxt [replay p]) 1 end in map (fn is => ` (env_of o hd) (map (nth decomp_prems) is) |> order_tac') env_groups |> FIRST end) val ad_absurdum_tac = SUBGOAL (fn (A, i) => case try (Logic_Sig.dest_Trueprop o Logic.strip_assums_concl) A of SOME (nt $ _) => if nt = Logic_Sig.Not then resolve0_tac [Logic_Sig.notI] i else resolve0_tac [Logic_Sig.ccontr] i | _ => resolve0_tac [Logic_Sig.ccontr] i) fun tac raw_order_proc octxt simp_prems ctxt = ad_absurdum_tac THEN' order_tac raw_order_proc octxt simp_prems ctxt end functor Order_Tac(structure Base_Tac : BASE_ORDER_TAC) = struct fun order_context_eq ({kind = kind1, ops = ops1, ...}, {kind = kind2, ops = ops2, ...}) = kind1 = kind2 andalso eq_list (op aconv) (ops1, ops2) fun order_data_eq (x, y) = order_context_eq (fst x, fst y) structure Data = Generic_Data( type T = (order_context * (order_context -> thm list -> Proof.context -> int -> tactic)) list val empty = [] fun merge data = Library.merge order_data_eq data ) fun declare (octxt as {kind = kind, raw_proc = raw_proc, ...}) lthy = - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val ops = map (Morphism.term phi) (#ops octxt) - val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt) - val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt) - val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms} - in - context |> Data.map (Library.insert order_data_eq (octxt', raw_proc)) - end) + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} + (fn phi => fn context => + let + val ops = map (Morphism.term phi) (#ops octxt) + val thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#thms octxt) + val conv_thms = map (fn (s, thm) => (s, Morphism.thm phi thm)) (#conv_thms octxt) + val octxt' = {kind = kind, ops = ops, thms = thms, conv_thms = conv_thms} + in + context |> Data.map (Library.insert order_data_eq (octxt', raw_proc)) + end) fun declare_order { ops = {eq = eq, le = le, lt = lt}, thms = { trans = trans, (* x \ y \ y \ z \ x \ z *) refl = refl, (* x \ x *) eqD1 = eqD1, (* x = y \ x \ y *) eqD2 = eqD2, (* x = y \ y \ x *) antisym = antisym, (* x \ y \ y \ x \ x = y *) contr = contr (* \ P \ P \ R *) }, conv_thms = { less_le = less_le, (* x < y \ x \ y \ x \ y *) nless_le = nless_le (* \ a < b \ \ a \ b \ a = b *) } } = declare { kind = Order, ops = [eq, le, lt], thms = [("trans", trans), ("refl", refl), ("eqD1", eqD1), ("eqD2", eqD2), ("antisym", antisym), ("contr", contr)], conv_thms = [("less_le", less_le), ("nless_le", nless_le)], raw_proc = Base_Tac.tac Order_Procedure.po_contr_prf } fun declare_linorder { ops = {eq = eq, le = le, lt = lt}, thms = { trans = trans, (* x \ y \ y \ z \ x \ z *) refl = refl, (* x \ x *) eqD1 = eqD1, (* x = y \ x \ y *) eqD2 = eqD2, (* x = y \ y \ x *) antisym = antisym, (* x \ y \ y \ x \ x = y *) contr = contr (* \ P \ P \ R *) }, conv_thms = { less_le = less_le, (* x < y \ x \ y \ x \ y *) nless_le = nless_le, (* \ x < y \ y \ x *) nle_le = nle_le (* \ a \ b \ b \ a \ b \ a *) } } = declare { kind = Linorder, ops = [eq, le, lt], thms = [("trans", trans), ("refl", refl), ("eqD1", eqD1), ("eqD2", eqD2), ("antisym", antisym), ("contr", contr)], conv_thms = [("less_le", less_le), ("nless_le", nless_le), ("nle_le", nle_le)], raw_proc = Base_Tac.tac Order_Procedure.lo_contr_prf } (* Try to solve the goal by calling the order solver with each of the declared orders. *) fun tac simp_prems ctxt = let fun app_tac (octxt, tac0) = CHANGED o tac0 octxt simp_prems ctxt in FIRST' (map app_tac (Data.get (Context.Proof ctxt))) end end diff --git a/src/Pure/Admin/isabelle_cronjob.scala b/src/Pure/Admin/isabelle_cronjob.scala --- a/src/Pure/Admin/isabelle_cronjob.scala +++ b/src/Pure/Admin/isabelle_cronjob.scala @@ -1,645 +1,646 @@ /* Title: Pure/Admin/isabelle_cronjob.scala Author: Makarius Main entry point for administrative cronjob at TUM. */ package isabelle import java.nio.file.Files import scala.annotation.tailrec object Isabelle_Cronjob { /* global resources: owned by main cronjob */ val backup = "lxbroy10:cronjob" val main_dir: Path = Path.explode("~/cronjob") val main_state_file: Path = main_dir + Path.explode("run/main.state") val build_release_log: Path = main_dir + Path.explode("run/build_release.log") val current_log: Path = main_dir + Path.explode("run/main.log") // owned by log service val cumulative_log: Path = main_dir + Path.explode("log/main.log") // owned by log service val isabelle_repos: Path = main_dir + Path.explode("isabelle") val afp_repos: Path = main_dir + Path.explode("AFP") val mailman_archives_dir = Path.explode("~/cronjob/Mailman") val build_log_dirs = List(Path.explode("~/log"), Path.explode("~/afp/log"), Path.explode("~/cronjob/log")) /** logger tasks **/ sealed case class Logger_Task(name: String = "", body: Logger => Unit) /* init and exit */ def get_rev(): String = Mercurial.repository(isabelle_repos).id() def get_afp_rev(): String = Mercurial.repository(afp_repos).id() val init: Logger_Task = Logger_Task("init", { logger => Isabelle_Devel.make_index() Mercurial.setup_repository(Isabelle_System.afp_repository.root, afp_repos) File.write(logger.log_dir + Build_Log.log_filename("isabelle_identify", logger.start_date), Build_Log.Identify.content(logger.start_date, Some(get_rev()), Some(get_afp_rev()))) Isabelle_System.bash( File.bash_path(Component_Rsync.local_program) + """ -a --include="*/" --include="plain_identify*" --exclude="*" """ + Bash.string(backup + "/log/.") + " " + File.bash_path(main_dir) + "/log/.").check - if (!Isabelle_Devel.cronjob_log.is_file) + if (!Isabelle_Devel.cronjob_log.is_file) { Files.createSymbolicLink(Isabelle_Devel.cronjob_log.java_path, current_log.java_path) + } }) val exit: Logger_Task = Logger_Task("exit", { logger => Isabelle_System.bash( File.bash_path(Component_Rsync.local_program) + " -a " + File.bash_path(main_dir) + "/log/." + " " + Bash.string(backup) + "/log/.") .check }) /* Mailman archives */ val mailman_archives: Logger_Task = Logger_Task("mailman_archives", { logger => Mailman.Isabelle_Users.download(mailman_archives_dir) Mailman.Isabelle_Dev.download(mailman_archives_dir) }) /* build release */ val build_release: Logger_Task = Logger_Task("build_release", { logger => build_release_log.file.delete Isabelle_Devel.release_snapshot(logger.options, get_rev(), get_afp_rev(), progress = new File_Progress(build_release_log)) }) /* remote build_history */ sealed case class Item( known: Boolean, isabelle_version: String, afp_version: Option[String], pull_date: Date ) { def unknown: Boolean = !known def versions: (String, Option[String]) = (isabelle_version, afp_version) def known_versions(rev: String, afp_rev: Option[String]): Boolean = known && rev != "" && isabelle_version == rev && (afp_rev.isEmpty || afp_rev.get != "" && afp_version == afp_rev) } def recent_items( db: SQL.Database, days: Int, rev: String, afp_rev: Option[String], sql: PostgreSQL.Source ): List[Item] = { val afp = afp_rev.isDefined db.execute_query_statement( Build_Log.Data.select_recent_versions( days = days, rev = rev, afp_rev = afp_rev, sql = SQL.where(sql)), List.from[Item], { res => val known = res.bool(Build_Log.Data.known) val isabelle_version = res.string(Build_Log.Prop.isabelle_version) val afp_version = if (afp) proper_string(res.string(Build_Log.Prop.afp_version)) else None val pull_date = res.date(Build_Log.Data.pull_date(afp)) Item(known, isabelle_version, afp_version, pull_date) }) } def unknown_runs(items: List[Item]): List[List[Item]] = { val (run, rest) = Library.take_prefix[Item](_.unknown, items.dropWhile(_.known)) if (run.nonEmpty) run :: unknown_runs(rest) else Nil } sealed case class Remote_Build( description: String, host: String, user: String = "", port: Int = 0, historic: Boolean = false, history: Int = 0, history_base: String = "build_history_base", components_base: String = Components.dynamic_components_base, clean_components: Boolean = true, java_heap: String = "", options: String = "", args: String = "", afp: Boolean = false, bulky: Boolean = false, more_hosts: List[String] = Nil, detect: PostgreSQL.Source = "", active: () => Boolean = () => true ) { def open_session(options: Options): SSH.Session = SSH.open_session(options, host = host, user = user, port = port) def sql: PostgreSQL.Source = SQL.and( Build_Log.Prop.build_engine.equal(Build_History.engine), Build_Log.Prop.build_host.member(host :: more_hosts), if_proper(detect, SQL.enclose(detect))) def profile: Build_Status.Profile = Build_Status.Profile(description, history = history, afp = afp, bulky = bulky, sql = sql) def pick( options: Options, rev: String = "", filter: Item => Boolean = _ => true ): Option[(String, Option[String])] = { val afp_rev = if (afp) Some(get_afp_rev()) else None val store = Build_Log.store(options) using(store.open_database()) { db => def pick_days(days: Int, gap: Int): Option[(String, Option[String])] = { val items = recent_items(db, days, rev, afp_rev, sql).filter(filter) def runs = unknown_runs(items).filter(run => run.length >= gap) if (historic || items.exists(_.known_versions(rev, afp_rev))) { val longest_run = runs.foldLeft(List.empty[Item]) { case (item1, item2) => if (item1.length >= item2.length) item1 else item2 } if (longest_run.isEmpty) None else Some(longest_run(longest_run.length / 2).versions) } else if (rev != "") Some((rev, afp_rev)) else runs.flatten.headOption.map(_.versions) } pick_days(options.int("build_log_history") max history, 2) orElse pick_days(200, 5) orElse pick_days(2000, 1) } } def build_history_options: String = " -h " + Bash.string(host) + " " + if_proper(java_heap, "-e 'ISABELLE_TOOL_JAVA_OPTIONS=\"$ISABELLE_TOOL_JAVA_OPTIONS -Xmx" + java_heap + "\"' ") + options } val remote_builds_old: List[Remote_Build] = List( Remote_Build("macOS 10.15 Catalina", "laramac01", user = "makarius", options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true -p pide_session=false", args = "-a -d '~~/src/Benchmarks'"), Remote_Build("Linux A", "i21of4", user = "i21isatest", options = "-m32 -M1x4,2,4" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=mlton" + " -e ISABELLE_SMLNJ=sml" + " -e ISABELLE_SWIPL=swipl", args = "-a -d '~~/src/Benchmarks'"), Remote_Build("Linux A", "lxbroy9", java_heap = "2g", options = "-m32 -B -M1x2,2", args = "-N -g timing"), Remote_Build("Linux Benchmarks", "lxbroy5", historic = true, history = 90, java_heap = "2g", options = "-m32 -B -M1x2,2 -t Benchmarks" + " -e ISABELLE_GHC=ghc -e ISABELLE_MLTON=mlton -e ISABELLE_OCAML=ocaml" + " -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind -e ISABELLE_SMLNJ=sml" + " -e ISABELLE_SWIPL=swipl", args = "-N -a -d '~~/src/Benchmarks'", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("Benchmarks")), Remote_Build("macOS 10.14 Mojave (Old)", "lapnipkow3", options = "-m32 -M1,2 -e ISABELLE_GHC_SETUP=true -p pide_session=false", args = "-a -d '~~/src/Benchmarks'"), Remote_Build("AFP old bulky", "lrzcloud1", options = "-m64 -M6 -U30000 -s10 -t AFP", args = "-g large -g slow", afp = true, bulky = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")), Remote_Build("AFP old", "lxbroy7", args = "-N -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")), Remote_Build("Poly/ML 5.7 Linux", "lxbroy8", history_base = "37074e22e8be", options = "-m32 -B -M1x2,2 -t polyml-5.7 -i 'init_component /home/isabelle/contrib/polyml-5.7'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7") + " AND " + Build_Log.Settings.ML_OPTIONS + " <> " + SQL.string("-H 500")), Remote_Build("Poly/ML 5.7.1 Linux", "lxbroy8", history_base = "a9d5b59c3e12", options = "-m32 -B -M1x2,2 -t polyml-5.7.1-pre2 -i 'init_component /home/isabelle/contrib/polyml-test-905dae2ebfda'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7.1-pre1") + " OR " + Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7.1-pre2")), Remote_Build("Poly/ML 5.7 macOS", "macbroy2", history_base = "37074e22e8be", options = "-m32 -B -M1x4,4 -t polyml-5.7 -i 'init_component /home/isabelle/contrib/polyml-5.7'", args = "-a", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7")), Remote_Build("Poly/ML 5.7.1 macOS", "macbroy2", history_base = "a9d5b59c3e12", options = "-m32 -B -M1x4,4 -t polyml-5.7.1-pre2 -i 'init_component /home/isabelle/contrib/polyml-test-905dae2ebfda'", args = "-a", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-5.7.1-pre1") + " OR " + Build_Log.Prop.build_tags + " = " + SQL.string("polyml-5.7.1-pre2")), Remote_Build("macOS", "macbroy2", options = "-m32 -M8" + " -e ISABELLE_GHC=ghc -e ISABELLE_MLTON=mlton -e ISABELLE_OCAML=ocaml" + " -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_OPAM_ROOT=\"$ISABELLE_HOME/opam\"" + " -e ISABELLE_SMLNJ=/home/isabelle/smlnj/110.85/bin/sml" + " -p pide_session=false", args = "-a", detect = Build_Log.Prop.build_tags.undefined, history_base = "2c0f24e927dd"), Remote_Build("macOS, quick_and_dirty", "macbroy2", options = "-m32 -M8 -t quick_and_dirty -p pide_session=false", args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("quick_and_dirty"), history_base = "2c0f24e927dd"), Remote_Build("macOS, skip_proofs", "macbroy2", options = "-m32 -M8 -t skip_proofs -p pide_session=false", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("skip_proofs"), history_base = "2c0f24e927dd"), Remote_Build("Poly/ML test", "lxbroy8", options = "-m32 -B -M1x2,2 -t polyml-test -i 'init_component /home/isabelle/contrib/polyml-5.7-20170217'", args = "-N -g timing", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("polyml-test")), Remote_Build("macOS 10.12 Sierra", "macbroy30", options = "-m32 -M2 -p pide_session=false", args = "-a", detect = Build_Log.Prop.build_start.toString + " > date '2017-03-03'"), Remote_Build("macOS 10.10 Yosemite", "macbroy31", options = "-m32 -M2 -p pide_session=false", args = "-a"), Remote_Build("macOS 10.8 Mountain Lion", "macbroy30", options = "-m32 -M2", args = "-a", detect = Build_Log.Prop.build_start.toString + " < date '2017-03-03'")) ::: { - for { (n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5")) } + for ((n, hosts) <- List(1 -> List("lxbroy6"), 2 -> List("lxbroy8", "lxbroy5"))) yield { Remote_Build("AFP old", host = hosts.head, more_hosts = hosts.tail, options = "-m32 -M1x2 -t AFP -P" + n + " -e ISABELLE_GHC=ghc" + " -e ISABELLE_MLTON=mlton" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind" + " -e ISABELLE_SMLNJ=sml", args = "-N -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP")) } } val remote_builds1: List[List[Remote_Build]] = { List( List(Remote_Build("Linux A", "augsburg1", options = "-m32 -B -M4" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=mlton -e ISABELLE_MLTON_OPTIONS=" + " -e ISABELLE_SMLNJ=sml" + " -e ISABELLE_SWIPL=swipl", args = "-a -d '~~/src/Benchmarks'")), List(Remote_Build("Linux B", "lxbroy10", historic = true, history = 90, options = "-m32 -B -M1x4,2,4,6", args = "-N -g timing")), List(Remote_Build("macOS 10.13 High Sierra", "lapbroy68", options = "-m32 -B -M1,2,4 -e ISABELLE_GHC_SETUP=true -p pide_session=false", args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("macOS 11 Big Sur", "mini1", options = "-m32 -B -M1x2,2,4 -p pide_session=false" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=/usr/local/bin/mlton -e ISABELLE_MLTON_OPTIONS=" + " -e ISABELLE_SMLNJ=/usr/local/smlnj/bin/sml" + " -e ISABELLE_SWIPL=/usr/local/bin/swipl", args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("macOS 10.14 Mojave", "mini2", options = "-m32 -B -M1x2,2,4 -p pide_session=false" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_MLTON=/usr/local/bin/mlton -e ISABELLE_MLTON_OPTIONS=" + " -e ISABELLE_SMLNJ=/usr/local/smlnj/bin/sml" + " -e ISABELLE_SWIPL=/usr/local/bin/swipl", args = "-a -d '~~/src/Benchmarks'"), Remote_Build("macOS, quick_and_dirty", "mini2", options = "-m32 -M4 -t quick_and_dirty -p pide_session=false", args = "-a -o quick_and_dirty", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("quick_and_dirty")), Remote_Build("macOS, skip_proofs", "mini2", options = "-m32 -M4 -t skip_proofs -p pide_session=false", args = "-a -o skip_proofs", detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("skip_proofs"))), List( Remote_Build("macOS 12 Monterey", "monterey", user = "makarius", options = "-m32 -M4 -e ISABELLE_GHC_SETUP=true -p pide_session=false", args = "-a -d '~~/src/Benchmarks'")), List( Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, components_base = "/cygdrive/d/isatest/contrib", options = "-m32 -M4" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_SMLNJ=/usr/local/smlnj-110.81/bin/sml", args = "-a", detect = Build_Log.Settings.ML_PLATFORM.toString + " = " + SQL.string("x86-windows") + " OR " + Build_Log.Settings.ML_PLATFORM + " = " + SQL.string("x86_64_32-windows")), Remote_Build("Windows", "vmnipkow9", historic = true, history = 90, components_base = "/cygdrive/d/isatest/contrib", options = "-m64 -M4" + " -C /cygdrive/d/isatest/contrib" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAML_SETUP=true" + " -e ISABELLE_GHC_SETUP=true" + " -e ISABELLE_SMLNJ=/usr/local/smlnj-110.81/bin/sml", args = "-a", detect = Build_Log.Settings.ML_PLATFORM.toString + " = " + SQL.string("x86_64-windows")))) } val remote_builds2: List[List[Remote_Build]] = List( List( Remote_Build("AFP", "lrzcloud2", java_heap = "8g", options = "-m32 -M1x5 -t AFP" + " -e ISABELLE_GHC=ghc" + " -e ISABELLE_MLTON=mlton -e ISABELLE_MLTON_OPTIONS=" + " -e ISABELLE_OCAML=ocaml -e ISABELLE_OCAMLC=ocamlc -e ISABELLE_OCAMLFIND=ocamlfind" + " -e ISABELLE_SMLNJ=sml", args = "-a -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP"), active = () => Date.now().unix_epoch_day % 2 == 0), Remote_Build("AFP", "lrzcloud2", java_heap = "8g", options = "-m64 -M8 -U30000 -s10 -t AFP", args = "-g large -g slow", afp = true, bulky = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP"), active = () => Date.now().unix_epoch_day % 2 == 1))) def remote_build_history( rev: String, afp_rev: Option[String], i: Int, r: Remote_Build ) : Logger_Task = { val task_name = "build_history-" + r.host Logger_Task(task_name, { logger => using(r.open_session(logger.options)) { ssh => val results = Build_History.remote_build(ssh, isabelle_repos, isabelle_repos.ext(r.host), isabelle_identifier = "cronjob_build_history", components_base = r.components_base, clean_platform = r.clean_components, clean_archives = r.clean_components, rev = rev, afp_repos = if (afp_rev.isDefined) Some(afp_repos) else None, afp_rev = afp_rev.getOrElse(""), options = " -N " + Bash.string(task_name) + (if (i < 0) "" else "_" + (i + 1).toString) + " -f " + r.build_history_options, args = "-o timeout=10800 " + r.args) for ((log_name, bytes) <- results) { logger.log(Date.now(), log_name) Bytes.write(logger.log_dir + Path.explode(log_name), bytes) } } }) } val build_status_profiles: List[Build_Status.Profile] = (remote_builds_old :: remote_builds1 ::: remote_builds2).flatten.map(_.profile) /** task logging **/ object Log_Service { def apply(options: Options, progress: Progress = new Progress): Log_Service = new Log_Service(options, progress) } class Log_Service private(val options: Options, progress: Progress) { current_log.file.delete private val thread: Consumer_Thread[String] = Consumer_Thread.fork("cronjob: logger", daemon = true)( consume = { (text: String) => // critical File.append(current_log, text + "\n") File.append(cumulative_log, text + "\n") progress.echo(text) true }) def shutdown(): Unit = { thread.shutdown() } val hostname: String = Isabelle_System.hostname() def log(date: Date, task_name: String, msg: String): Unit = if (task_name != "") { thread.send( "[" + Build_Log.print_date(date) + ", " + hostname + ", " + task_name + "]: " + msg) } def start_logger(start_date: Date, task_name: String): Logger = new Logger(this, start_date, task_name) def run_task(start_date: Date, task: Logger_Task): Unit = { val logger = start_logger(start_date, task.name) val res = Exn.capture { task.body(logger) } val end_date = Date.now() val err = res match { case Exn.Res(_) => None case Exn.Exn(exn) => Output.writeln("Exception trace for " + quote(task.name) + ":") exn.printStackTrace() val first_line = split_lines(Exn.message(exn)).headOption getOrElse "exception" Some(first_line) } logger.log_end(end_date, err) } def fork_task(start_date: Date, task: Logger_Task): Task = new Task(task.name, run_task(start_date, task)) } class Logger private[Isabelle_Cronjob]( log_service: Log_Service, val start_date: Date, val task_name: String ) { def options: Options = log_service.options def log(date: Date, msg: String): Unit = log_service.log(date, task_name, msg) def log_end(end_date: Date, err: Option[String]): Unit = { val elapsed_time = end_date.time - start_date.time val msg = (if (err.isEmpty) "finished" else "ERROR " + err.get) + (if (elapsed_time.seconds < 3.0) "" else " (" + elapsed_time.message_hms + " elapsed time)") log(end_date, msg) } val log_dir = Isabelle_System.make_directory(main_dir + Build_Log.log_subdir(start_date)) log(start_date, "started") } class Task private[Isabelle_Cronjob](name: String, body: => Unit) { private val future: Future[Unit] = Future.thread("cronjob: " + name) { body } def is_finished: Boolean = future.is_finished } /** cronjob **/ def cronjob(progress: Progress, exclude_task: Set[String]): Unit = { /* soft lock */ val still_running = try { Some(File.read(main_state_file)) } catch { case ERROR(_) => None } still_running match { case None | Some("") => case Some(running) => error("Isabelle cronjob appears to be still running: " + running) } /* log service */ val log_service = Log_Service(Options.init(), progress = progress) def run(start_date: Date, task: Logger_Task): Unit = log_service.run_task(start_date, task) def run_now(task: Logger_Task): Unit = run(Date.now(), task) /* structured tasks */ def SEQ(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = _ => for (task <- tasks.iterator if !exclude_task(task.name) || task.name == "") run_now(task)) def PAR(tasks: List[Logger_Task]): Logger_Task = Logger_Task(body = { _ => @tailrec def join(running: List[Task]): Unit = { running.partition(_.is_finished) match { case (Nil, Nil) => case (Nil, _ :: _) => Time.seconds(0.5).sleep(); join(running) case (_ :: _, remaining) => join(remaining) } } val start_date = Date.now() val running = for (task <- tasks if !exclude_task(task.name)) yield log_service.fork_task(start_date, task) join(running) }) /* repository structure */ val hg = Mercurial.repository(isabelle_repos) val hg_graph = hg.graph() def history_base_filter(r: Remote_Build): Item => Boolean = { val base_rev = hg.id(r.history_base) val nodes = hg_graph.all_succs(List(base_rev)).toSet (item: Item) => nodes(item.isabelle_version) } /* main */ val main_start_date = Date.now() File.write(main_state_file, main_start_date.toString + " " + log_service.hostname) run(main_start_date, Logger_Task("isabelle_cronjob", logger => run_now( SEQ(List( init, PAR(List(mailman_archives, build_release)), PAR( List(remote_builds1, remote_builds2).map(remote_builds => SEQ(List( PAR(remote_builds.map(_.filter(_.active())).map(seq => SEQ( for { (r, i) <- (if (seq.length <= 1) seq.map((_, -1)) else seq.zipWithIndex) (rev, afp_rev) <- r.pick(logger.options, hg.id(), history_base_filter(r)) } yield remote_build_history(rev, afp_rev, i, r)))), Logger_Task("build_log_database", logger => Build_Log.build_log_database(logger.options, build_log_dirs, vacuum = true, ml_statistics = true, snapshot = Some(Isabelle_Devel.build_log_snapshot))), Logger_Task("build_status", logger => Isabelle_Devel.build_status(logger.options)))))), exit))))) log_service.shutdown() main_state_file.file.delete } /** command line entry point **/ def main(args: Array[String]): Unit = { Command_Line.tool { var force = false var verbose = false var exclude_task = Set.empty[String] val getopts = Getopts(""" Usage: Admin/cronjob/main [OPTIONS] Options are: -f apply force to do anything -v verbose -x NAME exclude tasks with this name """, "f" -> (_ => force = true), "v" -> (_ => verbose = true), "x:" -> (arg => exclude_task += arg)) val more_args = getopts(args) if (more_args.nonEmpty) getopts.usage() val progress = if (verbose) new Console_Progress() else new Progress if (force) cronjob(progress, exclude_task) else error("Need to apply force to do anything") } } } diff --git a/src/Pure/Isar/attrib.ML b/src/Pure/Isar/attrib.ML --- a/src/Pure/Isar/attrib.ML +++ b/src/Pure/Isar/attrib.ML @@ -1,658 +1,660 @@ (* Title: Pure/Isar/attrib.ML Author: Markus Wenzel, TU Muenchen Symbolic representation of attributes -- with name and syntax. *) signature ATTRIB = sig type thms = Attrib.thms type fact = Attrib.fact val print_attributes: bool -> Proof.context -> unit val attribute_space: Context.generic -> Name_Space.T val define_global: binding -> (Token.src -> attribute) -> string -> theory -> string * theory val define: binding -> (Token.src -> attribute) -> string -> local_theory -> string * local_theory val check_name_generic: Context.generic -> xstring * Position.T -> string val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val attribs: Token.src list context_parser val opt_attribs: Token.src list context_parser val pretty_attribs: Proof.context -> Token.src list -> Pretty.T list val pretty_binding: Proof.context -> Attrib.binding -> string -> Pretty.T list val attribute: Proof.context -> Token.src -> attribute val attribute_global: theory -> Token.src -> attribute val attribute_cmd: Proof.context -> Token.src -> attribute val attribute_cmd_global: theory -> Token.src -> attribute val map_specs: ('a list -> 'att list) -> (('c * 'a list) * 'b) list -> (('c * 'att list) * 'b) list val map_facts: ('a list -> 'att list) -> (('c * 'a list) * ('d * 'a list) list) list -> (('c * 'att list) * ('d * 'att list) list) list val map_facts_refs: ('a list -> 'att list) -> ('b -> 'fact) -> (('c * 'a list) * ('b * 'a list) list) list -> (('c * 'att list) * ('fact * 'att list) list) list val trim_context_binding: Attrib.binding -> Attrib.binding val trim_context_thms: thms -> thms val trim_context_fact: fact -> fact val global_notes: string -> fact list -> theory -> (string * thm list) list * theory val local_notes: string -> fact list -> Proof.context -> (string * thm list) list * Proof.context val generic_notes: string -> fact list -> Context.generic -> (string * thm list) list * Context.generic val lazy_notes: string -> binding * thm list lazy -> Context.generic -> Context.generic val eval_thms: Proof.context -> (Facts.ref * Token.src list) list -> thm list val attribute_syntax: attribute context_parser -> Token.src -> attribute val setup: binding -> attribute context_parser -> string -> theory -> theory val local_setup: binding -> attribute context_parser -> string -> local_theory -> local_theory val attribute_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory - val internal: (morphism -> attribute) -> Token.src - val internal_declaration: Morphism.declaration_entity -> thms + val internal: Position.T -> (morphism -> attribute) -> Token.src + val internal_declaration: Position.T -> Morphism.declaration_entity -> thms val add_del: attribute -> attribute -> attribute context_parser val thm: thm context_parser val thms: thm list context_parser val multi_thm: thm list context_parser val transform_facts: morphism -> fact list -> fact list val partial_evaluation: Proof.context -> fact list -> fact list val print_options: bool -> Proof.context -> unit val config_bool: binding -> (Context.generic -> bool) -> bool Config.T * (theory -> theory) val config_int: binding -> (Context.generic -> int) -> int Config.T * (theory -> theory) val config_real: binding -> (Context.generic -> real) -> real Config.T * (theory -> theory) val config_string: binding -> (Context.generic -> string) -> string Config.T * (theory -> theory) val setup_config_bool: binding -> (Context.generic -> bool) -> bool Config.T val setup_config_int: binding -> (Context.generic -> int) -> int Config.T val setup_config_real: binding -> (Context.generic -> real) -> real Config.T val setup_config_string: binding -> (Context.generic -> string) -> string Config.T val option_bool: string * Position.T -> bool Config.T * (theory -> theory) val option_int: string * Position.T -> int Config.T * (theory -> theory) val option_real: string * Position.T -> real Config.T * (theory -> theory) val option_string: string * Position.T -> string Config.T * (theory -> theory) val setup_option_bool: string * Position.T -> bool Config.T val setup_option_int: string * Position.T -> int Config.T val setup_option_real: string * Position.T -> real Config.T val setup_option_string: string * Position.T -> string Config.T val consumes: int -> Token.src val constraints: int -> Token.src val cases_open: Token.src val case_names: string list -> Token.src val case_conclusion: string * string list -> Token.src end; structure Attrib: sig type binding = Attrib.binding include ATTRIB end = struct open Attrib; (** named attributes **) (* theory data *) structure Attributes = Generic_Data ( type T = ((Token.src -> attribute) * string) Name_Space.table; val empty : T = Name_Space.empty_table Markup.attributeN; fun merge data : T = Name_Space.merge_tables data; ); val ops_attributes = {get_data = Attributes.get, put_data = Attributes.put}; val get_attributes = Attributes.get o Context.Proof; fun print_attributes verbose ctxt = let val attribs = get_attributes ctxt; fun prt_attr (name, (_, "")) = Pretty.mark_str name | prt_attr (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "attributes:" (map prt_attr (Name_Space.markup_table verbose ctxt attribs))] |> Pretty.writeln_chunks end; val attribute_space = Name_Space.space_of_table o Attributes.get; (* define *) fun define_global binding att comment = Entity.define_global ops_attributes binding (att, comment); fun define binding att comment = Entity.define ops_attributes binding (att, comment); (* check *) fun check_name_generic context = #1 o Name_Space.check context (Attributes.get context); val check_name = check_name_generic o Context.Proof; fun check_src ctxt src = let val _ = if Token.checked_src src then () else Context_Position.report ctxt (#1 (Token.range_of src)) Markup.language_attribute; in #1 (Token.check_src ctxt get_attributes src) end; val attribs = Args.context -- Scan.lift Parse.attribs >> (fn (ctxt, srcs) => map (check_src ctxt) srcs); val opt_attribs = Scan.optional attribs []; (* pretty printing *) fun pretty_attribs _ [] = [] | pretty_attribs ctxt srcs = [Pretty.enum "," "[" "]" (map (Token.pretty_src ctxt) srcs)]; fun pretty_binding ctxt (b, atts) sep = (case (Binding.is_empty b, null atts) of (true, true) => [] | (false, true) => [Pretty.block [Binding.pretty b, Pretty.str sep]] | (true, false) => [Pretty.block (pretty_attribs ctxt atts @ [Pretty.str sep])] | (false, false) => [Pretty.block (Binding.pretty b :: Pretty.brk 1 :: pretty_attribs ctxt atts @ [Pretty.str sep])]); (* get attributes *) fun attribute_generic context = let val table = Attributes.get context in fn src => let val name = #1 (Token.name_of_src src); val label = Long_Name.qualify Markup.attributeN name; val att = #1 (Name_Space.get table name) src; in Position.setmp_thread_data_label label att : attribute end end; val attribute = attribute_generic o Context.Proof; val attribute_global = attribute_generic o Context.Theory; fun attribute_cmd ctxt = attribute ctxt o check_src ctxt; fun attribute_cmd_global thy = attribute_global thy o check_src (Proof_Context.init_global thy); (* attributed declarations *) fun map_specs f = map (apfst (apsnd f)); fun map_facts f = map (apfst (apsnd f) o apsnd (map (apsnd f))); fun map_facts_refs f g = map_facts f #> map (apsnd (map (apfst g))); (* implicit context *) val trim_context_binding: Attrib.binding -> Attrib.binding = apsnd ((map o map) Token.trim_context); val trim_context_thms: thms -> thms = map (fn (thms, atts) => (map Thm.trim_context thms, (map o map) Token.trim_context atts)); fun trim_context_fact (binding, thms) = (trim_context_binding binding, trim_context_thms thms); (* fact expressions *) fun global_notes kind facts thy = thy |> Global_Theory.note_thmss kind (map_facts (map (attribute_global thy)) facts); fun local_notes kind facts ctxt = ctxt |> Proof_Context.note_thmss kind (map_facts (map (attribute ctxt)) facts); fun generic_notes kind facts context = context |> Context.mapping_result (global_notes kind facts) (local_notes kind facts); fun lazy_notes kind arg = Context.mapping (Global_Theory.add_thms_lazy kind arg) (Proof_Context.add_thms_lazy kind arg); fun eval_thms ctxt srcs = ctxt |> Proof_Context.note_thmss "" (map_facts_refs (map (attribute_cmd ctxt)) (Proof_Context.get_fact ctxt) [(Binding.empty_atts, srcs)]) |> fst |> maps snd; (* attribute setup *) fun attribute_syntax scan src (context, th) = let val (a, context') = Token.syntax_generic scan src context in a (context', th) end; fun setup binding scan comment = define_global binding (attribute_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (attribute_syntax scan) comment #> snd; fun attribute_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Attrib.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (") " ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* internal attribute *) fun make_name ctxt name = Token.make_src (name, Position.none) [] |> check_src ctxt |> hd; local val internal_binding = Binding.make ("attribute", \<^here>); val _ = Theory.setup (setup internal_binding (Scan.lift Args.internal_attribute >> Morphism.form || Scan.lift Args.internal_declaration >> (Thm.declaration_attribute o K o Morphism.form)) "internal attribute"); val internal_name = make_name (Context.the_local_context ()) (Binding.name_of internal_binding); -val internal_arg = Token.make_string0 ""; -fun internal_source value = [internal_name, Token.assign (SOME value) internal_arg]; +fun internal_source name value = [internal_name, Token.assign (SOME value) (Token.make_string name)]; in -fun internal arg = internal_source (Token.Attribute (Morphism.entity arg)); -fun internal_declaration arg = [([Drule.dummy_thm], [internal_source (Token.Declaration arg)])]; +fun internal pos arg = + internal_source ("internal", pos) (Token.Attribute (Morphism.entity arg)); + +fun internal_declaration pos arg = + [([Drule.dummy_thm], [internal_source ("declaration", pos) (Token.Declaration arg)])]; end; (* add/del syntax *) fun add_del add del = Scan.lift (Args.add >> K add || Args.del >> K del || Scan.succeed add); (** parsing attributed theorems **) local val fact_name = Parse.position Args.internal_fact >> (fn (_, pos) => ("", pos)) || Args.name_position; fun gen_thm pick = Scan.depend (fn context => let val get = Proof_Context.get_fact_generic context; val get_fact = get o Facts.Fact; fun get_named is_sel pos name = let val (a, ths) = get (Facts.Named ((name, pos), NONE)) in (if is_sel then NONE else a, ths) end; in Parse.$$$ "[" |-- Scan.pass context attribs --| Parse.$$$ "]" >> (fn srcs => let val atts = map (attribute_generic context) srcs; val (th', context') = fold (uncurry o Thm.apply_attribute) atts (Drule.dummy_thm, context); in (context', pick ("", Position.none) [th']) end) || (Scan.ahead Args.alt_name -- Args.named_fact get_fact >> (fn (s, fact) => ("", Facts.Fact s, fact)) || Scan.ahead (fact_name -- Scan.option Parse.thm_sel) :|-- (fn ((name, pos), sel) => Args.named_fact (get_named (is_some sel) pos) --| Scan.option Parse.thm_sel >> (fn fact => (name, Facts.Named ((name, pos), sel), fact)))) -- Scan.pass context opt_attribs >> (fn ((name, thmref, fact), srcs) => let val ths = Facts.select thmref fact; val atts = map (attribute_generic context) srcs; val (ths', context') = fold_map (curry (fold (uncurry o Thm.apply_attribute) atts)) ths context; in (context', pick (name, Facts.ref_pos thmref) ths') end) end); in val thm = gen_thm Facts.the_single; val multi_thm = gen_thm (K I); val thms = Scan.repeats multi_thm; end; (* transform fact expressions *) fun transform_facts phi = map (fn ((a, atts), bs) => ((Morphism.binding phi a, (map o map) (Token.transform phi) atts), bs |> map (fn (ths, btts) => (Morphism.fact phi ths, (map o map) (Token.transform phi) btts)))); (** partial evaluation -- observing rule/declaration/mixed attributes **) (*NB: result length may change due to rearrangement of symbolic expression*) local fun apply_att src (context, th) = let val src1 = map Token.init_assignable src; val result = attribute_generic context src1 (context, th); val src2 = map Token.closure src1; in (src2, result) end; fun err msg src = let val (name, pos) = Token.name_of_src src in error (msg ^ " " ^ quote name ^ Position.here pos) end; fun eval src ((th, dyn), (decls, context)) = (case (apply_att src (context, th), dyn) of ((_, (NONE, SOME th')), NONE) => ((th', NONE), (decls, context)) | ((_, (NONE, SOME _)), SOME _) => err "Mixed dynamic attribute followed by static rule" src | ((src', (SOME context', NONE)), NONE) => let val decls' = (case decls of [] => [(th, [src'])] | (th2, srcs2) :: rest => if Thm.eq_thm_strict (th, th2) then ((th2, src' :: srcs2) :: rest) else (th, [src']) :: (th2, srcs2) :: rest); in ((th, NONE), (decls', context')) end | ((src', (opt_context', opt_th')), _) => let val context' = the_default context opt_context'; val th' = the_default th opt_th'; val dyn' = (case dyn of NONE => SOME (th, [src']) | SOME (dyn_th, srcs) => SOME (dyn_th, src' :: srcs)); in ((th', dyn'), (decls, context')) end); in fun partial_evaluation ctxt facts = (facts, Context.Proof (Context_Position.not_really ctxt)) |-> fold_map (fn ((b, more_atts), fact) => fn context => let val (fact', (decls, context')) = (fact, ([], context)) |-> fold_map (fn (ths, atts) => fn res1 => (ths, res1) |-> fold_map (fn th => fn res2 => let val ((th', dyn'), res3) = fold eval (atts @ more_atts) ((th, NONE), res2); val th_atts' = (case dyn' of NONE => (th', []) | SOME (dyn_th', atts') => (dyn_th', rev atts')); in (th_atts', res3) end)) |>> flat; val decls' = rev (map (apsnd rev) decls); val facts' = if eq_list (eq_fst Thm.eq_thm_strict) (decls', fact') then [((b, []), map2 (fn (th, atts1) => fn (_, atts2) => (th, atts1 @ atts2)) decls' fact')] else if null decls' then [((b, []), fact')] else [(Binding.empty_atts, decls'), ((b, []), fact')]; in (facts', context') end) |> fst |> flat |> map (apsnd (map (apfst single))) |> filter_out (fn (b, fact) => Binding.is_empty_atts b andalso forall (null o #2) fact); end; (** configuration options **) (* naming *) structure Configs = Theory_Data ( type T = Config.value Config.T Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ); fun print_options verbose ctxt = let fun prt (name, config) = let val value = Config.get ctxt config in Pretty.block [Pretty.mark_str name, Pretty.str (": " ^ Config.print_type value ^ " ="), Pretty.brk 1, Pretty.str (Config.print_value value)] end; val space = attribute_space (Context.Proof ctxt); val configs = Name_Space.markup_entries verbose ctxt space (Symtab.dest (Configs.get (Proof_Context.theory_of ctxt))); in Pretty.writeln (Pretty.big_list "configuration options" (map prt configs)) end; (* concrete syntax *) local val equals = Parse.$$$ "="; fun scan_value (Config.Bool _) = equals -- Args.$$$ "false" >> K (Config.Bool false) || equals -- Args.$$$ "true" >> K (Config.Bool true) || Scan.succeed (Config.Bool true) | scan_value (Config.Int _) = equals |-- Parse.int >> Config.Int | scan_value (Config.Real _) = equals |-- Parse.real >> Config.Real | scan_value (Config.String _) = equals |-- Args.name >> Config.String; fun scan_config thy config = let val config_type = Config.get_global thy config in scan_value config_type >> (K o Thm.declaration_attribute o K o Config.put_generic config) end; fun register binding config thy = let val name = Sign.full_name thy binding in thy |> setup binding (Scan.lift (scan_config thy config) >> Morphism.form_entity) "configuration option" |> Configs.map (Symtab.update (name, config)) end; fun declare make coerce binding default = let val name = Binding.name_of binding; val pos = Binding.pos_of binding; val config_value = Config.declare (name, pos) (make o default); val config = coerce config_value; in (config, register binding config_value) end; in fun register_config config = register (Binding.make (Config.name_of config, Config.pos_of config)) config; val register_config_bool = register_config o Config.bool_value; val register_config_int = register_config o Config.int_value; val register_config_real = register_config o Config.real_value; val register_config_string = register_config o Config.string_value; val config_bool = declare Config.Bool Config.bool; val config_int = declare Config.Int Config.int; val config_real = declare Config.Real Config.real; val config_string = declare Config.String Config.string; end; (* implicit setup *) local fun setup_config declare_config binding default = let val (config, setup) = declare_config binding default; val _ = Theory.setup setup; in config end; in val setup_config_bool = setup_config config_bool; val setup_config_int = setup_config config_int; val setup_config_string = setup_config config_string; val setup_config_real = setup_config config_real; end; (* system options *) local fun declare_option coerce (name, pos) = let val config = Config.declare_option (name, pos); in (coerce config, register_config config) end; fun setup_option coerce (name, pos) = let val config = Config.declare_option (name, pos); val _ = Theory.setup (register_config config); in coerce config end; in val option_bool = declare_option Config.bool; val option_int = declare_option Config.int; val option_real = declare_option Config.real; val option_string = declare_option Config.string; val setup_option_bool = setup_option Config.bool; val setup_option_int = setup_option Config.int; val setup_option_real = setup_option Config.real; val setup_option_string = setup_option Config.string; end; (* theory setup *) val _ = Theory.setup (setup \<^binding>\tagged\ (Scan.lift (Args.name -- Args.name) >> Thm.tag) "tagged theorem" #> setup \<^binding>\untagged\ (Scan.lift Args.name >> Thm.untag) "untagged theorem" #> setup \<^binding>\kind\ (Scan.lift Args.name >> Thm.kind) "theorem kind" #> setup \<^binding>\THEN\ (Scan.lift (Scan.optional (Args.bracks Parse.nat) 1) -- thm >> (fn (i, B) => Thm.rule_attribute [B] (fn _ => fn A => A RSN (i, B)))) "resolution with rule" #> setup \<^binding>\OF\ (thms >> (fn Bs => Thm.rule_attribute Bs (fn _ => fn A => A OF Bs))) "rule resolved with facts" #> setup \<^binding>\rename_abs\ (Scan.lift (Scan.repeat (Args.maybe Args.name)) >> (fn vs => Thm.rule_attribute [] (K (Drule.rename_bvars' vs)))) "rename bound variables in abstractions" #> setup \<^binding>\unfolded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.unfold (Context.proof_of context) ths))) "unfolded definitions" #> setup \<^binding>\folded\ (thms >> (fn ths => Thm.rule_attribute ths (fn context => Local_Defs.fold (Context.proof_of context) ths))) "folded definitions" #> setup \<^binding>\consumes\ (Scan.lift (Scan.optional Parse.int 1) >> Rule_Cases.consumes) "number of consumed facts" #> setup \<^binding>\constraints\ (Scan.lift Parse.nat >> Rule_Cases.constraints) "number of equality constraints" #> setup \<^binding>\cases_open\ (Scan.succeed Rule_Cases.cases_open) "rule with open parameters" #> setup \<^binding>\case_names\ (Scan.lift (Scan.repeat (Args.name -- Scan.optional (Parse.$$$ "[" |-- Scan.repeat1 (Args.maybe Args.name) --| Parse.$$$ "]") [])) >> (fn cs => Rule_Cases.cases_hyp_names (map #1 cs) (map (map (the_default Rule_Cases.case_hypsN) o #2) cs))) "named rule cases" #> setup \<^binding>\case_conclusion\ (Scan.lift (Args.name -- Scan.repeat Args.name) >> Rule_Cases.case_conclusion) "named conclusion of rule cases" #> setup \<^binding>\params\ (Scan.lift (Parse.and_list1 (Scan.repeat Args.name)) >> Rule_Cases.params) "named rule parameters" #> setup \<^binding>\rule_format\ (Scan.lift (Args.mode "no_asm") >> (fn true => Object_Logic.rule_format_no_asm | false => Object_Logic.rule_format)) "result put into canonical rule format" #> setup \<^binding>\elim_format\ (Scan.succeed (Thm.rule_attribute [] (K Tactic.make_elim))) "destruct rule turned into elimination rule format" #> setup \<^binding>\no_vars\ (Scan.succeed (Thm.rule_attribute [] (Variable.import_vars o Context.proof_of))) "imported schematic variables" #> setup \<^binding>\atomize\ (Scan.succeed Object_Logic.declare_atomize) "declaration of atomize rule" #> setup \<^binding>\rulify\ (Scan.succeed Object_Logic.declare_rulify) "declaration of rulify rule" #> setup \<^binding>\rotated\ (Scan.lift (Scan.optional Parse.int 1 >> (fn n => Thm.rule_attribute [] (fn _ => rotate_prems n)))) "rotated theorem premises" #> setup \<^binding>\defn\ (add_del Local_Defs.defn_add Local_Defs.defn_del) "declaration of definitional transformations" #> setup \<^binding>\abs_def\ (Scan.succeed (Thm.rule_attribute [] (Local_Defs.abs_def_rule o Context.proof_of))) "abstract over free variables of definitional theorem" #> register_config_bool Goal.quick_and_dirty #> register_config_bool Ast.trace #> register_config_bool Ast.stats #> register_config_bool Printer.show_brackets #> register_config_bool Printer.show_sorts #> register_config_bool Printer.show_types #> register_config_bool Printer.show_markup #> register_config_bool Printer.show_structs #> register_config_bool Printer.show_question_marks #> register_config_bool Syntax.ambiguity_warning #> register_config_int Syntax.ambiguity_limit #> register_config_bool Syntax_Trans.eta_contract #> register_config_bool Name_Space.names_long #> register_config_bool Name_Space.names_short #> register_config_bool Name_Space.names_unique #> register_config_int ML_Print_Depth.print_depth #> register_config_string ML_Env.ML_environment #> register_config_bool ML_Env.ML_read_global #> register_config_bool ML_Env.ML_write_global #> register_config_bool ML_Options.source_trace #> register_config_bool ML_Options.exception_trace #> register_config_bool ML_Options.exception_debugger #> register_config_bool ML_Options.debugger #> register_config_bool Proof_Context.show_abbrevs #> register_config_int Goal_Display.goals_limit #> register_config_bool Goal_Display.show_main_goal #> register_config_bool Thm.show_consts #> register_config_bool Thm.show_hyps #> register_config_bool Thm.show_tags #> register_config_bool Pattern.unify_trace_failure #> register_config_int Unify.trace_bound #> register_config_int Unify.search_bound #> register_config_bool Unify.trace_simp #> register_config_bool Unify.trace_types #> register_config_int Raw_Simplifier.simp_depth_limit #> register_config_int Raw_Simplifier.simp_trace_depth_limit #> register_config_bool Raw_Simplifier.simp_debug #> register_config_bool Raw_Simplifier.simp_trace #> register_config_bool Local_Defs.unfold_abs_def); (* internal source *) local val make_name0 = make_name (Context.the_local_context ()); val consumes_name = make_name0 "consumes"; val constraints_name = make_name0 "constraints"; val cases_open_name = make_name0 "cases_open"; val case_names_name = make_name0 "case_names"; val case_conclusion_name = make_name0 "case_conclusion"; in fun consumes i = consumes_name :: Token.make_int i; fun constraints i = constraints_name :: Token.make_int i; val cases_open = [cases_open_name]; fun case_names names = case_names_name :: map Token.make_string0 names; fun case_conclusion (name, names) = case_conclusion_name :: map Token.make_string0 (name :: names); end; end; \ No newline at end of file diff --git a/src/Pure/Isar/bundle.ML b/src/Pure/Isar/bundle.ML --- a/src/Pure/Isar/bundle.ML +++ b/src/Pure/Isar/bundle.ML @@ -1,251 +1,251 @@ (* Title: Pure/Isar/bundle.ML Author: Makarius Bundled declarations (notes etc.). *) signature BUNDLE = sig type name = string val bundle_space: Context.generic -> Name_Space.T val check: Proof.context -> xstring * Position.T -> string val get: Proof.context -> name -> Attrib.thms val read: Proof.context -> xstring * Position.T -> Attrib.thms val extern: Proof.context -> string -> xstring val print_bundles: bool -> Proof.context -> unit val bundle: binding * Attrib.thms -> (binding * typ option * mixfix) list -> local_theory -> local_theory val bundle_cmd: binding * (Facts.ref * Token.src list) list -> (binding * string option * mixfix) list -> local_theory -> local_theory val init: binding -> theory -> local_theory val unbundle: name list -> local_theory -> local_theory val unbundle_cmd: (xstring * Position.T) list -> local_theory -> local_theory val includes: name list -> Proof.context -> Proof.context val includes_cmd: (xstring * Position.T) list -> Proof.context -> Proof.context val include_: name list -> Proof.state -> Proof.state val include_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state val including: name list -> Proof.state -> Proof.state val including_cmd: (xstring * Position.T) list -> Proof.state -> Proof.state end; structure Bundle: BUNDLE = struct (** context data **) structure Data = Generic_Data ( type T = Attrib.thms Name_Space.table * Attrib.thms option; val empty : T = (Name_Space.empty_table Markup.bundleN, NONE); fun merge ((tab1, target1), (tab2, target2)) = (Name_Space.merge_tables (tab1, tab2), merge_options (target1, target2)); ); (* bundles *) type name = string val get_all_generic = #1 o Data.get; val get_all = get_all_generic o Context.Proof; val bundle_space = Name_Space.space_of_table o #1 o Data.get; fun check ctxt = #1 o Name_Space.check (Context.Proof ctxt) (get_all ctxt); val get = Name_Space.get o get_all_generic o Context.Proof; fun read ctxt = get ctxt o check ctxt; fun extern ctxt = Name_Space.extern ctxt (Name_Space.space_of_table (get_all ctxt)); (* target -- bundle under construction *) fun the_target thy = #2 (Data.get (Context.Theory thy)) |> \<^if_none>\error "Missing bundle target"\; val reset_target = (Context.theory_map o Data.map o apsnd o K) NONE; val set_target = Context.theory_map o Data.map o apsnd o K o SOME o Attrib.trim_context_thms; fun augment_target thms = Local_Theory.background_theory (fn thy => set_target (the_target thy @ thms) thy); (* print bundles *) fun pretty_bundle ctxt (markup_name, bundle) = let val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt; fun prt_thm_attribs atts th = Pretty.block (Pretty.breaks (prt_thm th :: Attrib.pretty_attribs ctxt atts)); fun prt_thms (ths, []) = map prt_thm ths | prt_thms (ths, atts) = map (prt_thm_attribs atts) ths; in Pretty.block ([Pretty.keyword1 "bundle", Pretty.str " ", Pretty.mark_str markup_name] @ (if null bundle then [] else Pretty.fbreaks (Pretty.str " =" :: maps prt_thms bundle))) end; fun print_bundles verbose ctxt = Pretty.writeln_chunks (map (pretty_bundle ctxt) (Name_Space.markup_table verbose ctxt (get_all ctxt))); (** define bundle **) (* context and morphisms *) val trim_context_bundle = map (fn (fact, atts) => (map Thm.trim_context fact, (map o map) Token.trim_context atts)); fun transfer_bundle thy = map (fn (fact, atts) => (map (Thm.transfer thy) fact, (map o map) (Token.transfer thy) atts)); fun transform_bundle phi = map (fn (fact, atts) => (Morphism.fact phi fact, (map o map) (Token.transform phi) atts)); fun define_bundle (b, bundle) context = let val (name, bundles') = get_all_generic context |> Name_Space.define context true (b, trim_context_bundle bundle); val context' = (Data.map o apfst o K) bundles' context; in (name, context') end; (* command *) local fun gen_bundle prep_fact prep_att add_fixes (binding, raw_bundle) raw_fixes lthy = let val (_, ctxt') = add_fixes raw_fixes lthy; val bundle0 = raw_bundle |> map (fn (fact, atts) => (prep_fact ctxt' fact, map (prep_att ctxt') atts)); val bundle = Attrib.partial_evaluation ctxt' [(Binding.empty_atts, bundle0)] |> maps #2 |> transform_bundle (Proof_Context.export_morphism ctxt' lthy) |> trim_context_bundle; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of binding} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi in #2 (define_bundle (Morphism.binding psi binding, transform_bundle psi bundle) context) end) end; in val bundle = gen_bundle (K I) (K I) Proof_Context.add_fixes; val bundle_cmd = gen_bundle Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; end; (* target *) local fun bad_operation _ = error "Not possible in bundle target"; fun conclude invisible binding = Local_Theory.background_theory_result (fn thy => thy |> invisible ? Context_Position.set_visible_global false |> Context.Theory |> define_bundle (binding, the_target thy) ||> (Context.the_theory #> invisible ? Context_Position.restore_visible_global thy #> reset_target)); fun pretty binding lthy = let val bundle = the_target (Proof_Context.theory_of lthy); val (name, lthy') = lthy |> Local_Theory.raw_theory (Context_Position.set_visible_global false) |> conclude true binding; val thy_ctxt' = Proof_Context.init_global (Proof_Context.theory_of lthy'); val markup_name = Name_Space.markup_extern thy_ctxt' (Name_Space.space_of_table (get_all thy_ctxt')) name; in [pretty_bundle lthy' (markup_name, bundle)] end; fun bundle_notes kind facts lthy = let val bundle = facts |> maps (fn ((_, more_atts), thms) => map (fn (ths, atts) => (ths, atts @ more_atts)) thms); in lthy |> augment_target (transform_bundle (Local_Theory.standard_morphism_theory lthy) bundle) |> Generic_Target.standard_notes (op <>) kind facts |> Attrib.local_notes kind facts end; -fun bundle_declaration decl lthy = +fun bundle_declaration pos decl lthy = lthy - |> (augment_target o Attrib.internal_declaration) + |> (augment_target o Attrib.internal_declaration pos) (Morphism.transform (Local_Theory.standard_morphism_theory lthy) decl) |> Generic_Target.standard_declaration (K true) decl; in fun init binding thy = thy |> Local_Theory.init {background_naming = Sign.naming_of thy, setup = set_target [] #> Proof_Context.init_global, conclude = conclude false binding #> #2} {define = bad_operation, notes = bundle_notes, abbrev = bad_operation, - declaration = K bundle_declaration, + declaration = fn {pos, ...} => bundle_declaration pos, theory_registration = bad_operation, locale_dependency = bad_operation, - pretty = pretty binding} + pretty = pretty binding}; end; (** activate bundles **) local fun gen_activate notes prep_bundle args ctxt = let val thy = Proof_Context.theory_of ctxt; val decls = maps (prep_bundle ctxt) args |> transfer_bundle thy; in ctxt |> Context_Position.set_visible false |> notes [(Binding.empty_atts, decls)] |> #2 |> Context_Position.restore_visible ctxt end; fun gen_unbundle prep_bundle = gen_activate Local_Theory.notes prep_bundle; fun gen_includes prep_bundle = gen_activate (Attrib.local_notes "") prep_bundle; fun gen_include prep_bundle bs = Proof.assert_forward #> Proof.map_context (gen_includes prep_bundle bs) #> Proof.reset_facts; fun gen_including prep_bundle bs = Proof.assert_backward #> Proof.map_context (gen_includes prep_bundle bs) in val unbundle = gen_unbundle get; val unbundle_cmd = gen_unbundle read; val includes = gen_includes get; val includes_cmd = gen_includes read; val include_ = gen_include get; val include_cmd = gen_include read; val including = gen_including get; val including_cmd = gen_including read; end; end; diff --git a/src/Pure/Isar/code.ML b/src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML +++ b/src/Pure/Isar/code.ML @@ -1,1570 +1,1570 @@ (* Title: Pure/Isar/code.ML Author: Florian Haftmann, TU Muenchen Abstract executable ingredients of theory. Management of data dependent on executable ingredients as synchronized cache; purged on any change of underlying executable ingredients. *) signature CODE = sig (*constants*) val check_const: theory -> term -> string val read_const: theory -> string -> string val string_of_const: theory -> string -> string val args_number: theory -> string -> int (*constructor sets*) val constrset_of_consts: theory -> (string * typ) list -> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list) (*code equations and certificates*) val assert_eqn: theory -> thm * bool -> thm * bool val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string) type cert val constrain_cert: theory -> sort list -> cert -> cert val conclude_cert: cert -> cert val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list val equations_of_cert: theory -> cert -> ((string * sort) list * typ) * (((string option * term) list * (string option * term)) * (thm option * bool)) list option val pretty_cert: theory -> cert -> Pretty.T list (*executable code*) type constructors type abs_type val type_interpretation: (string -> theory -> theory) -> theory -> theory val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory val declare_datatype_global: (string * typ) list -> theory -> theory val declare_datatype_cmd: string list -> theory -> theory val declare_abstype: thm -> local_theory -> local_theory val declare_abstype_global: thm -> theory -> theory val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory val declare_default_eqns_global: (thm * bool) list -> theory -> theory val declare_eqns: (thm * bool) list -> local_theory -> local_theory val declare_eqns_global: (thm * bool) list -> theory -> theory val add_eqn_global: thm * bool -> theory -> theory val del_eqn_global: thm -> theory -> theory val declare_abstract_eqn: thm -> local_theory -> local_theory val declare_abstract_eqn_global: thm -> theory -> theory val declare_aborting_global: string -> theory -> theory val declare_unimplemented_global: string -> theory -> theory val declare_case_global: thm -> theory -> theory val declare_undefined_global: string -> theory -> theory val get_type: theory -> string -> constructors * bool val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option val is_constr: theory -> string -> bool val is_abstr: theory -> string -> bool val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list -> string -> cert type case_schema val get_case_schema: theory -> string -> case_schema option val get_case_cong: theory -> string -> thm option val is_undefined: theory -> string -> bool val print_codesetup: theory -> unit end; signature CODE_DATA_ARGS = sig type T val empty: T end; signature CODE_DATA = sig type T val change: theory option -> (T -> T) -> T val change_yield: theory option -> (T -> 'a * T) -> 'a * T end; signature PRIVATE_CODE = sig include CODE val declare_data: Any.T -> serial val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a) -> theory -> ('a -> 'b * 'a) -> 'b * 'a end; structure Code : PRIVATE_CODE = struct (** auxiliary **) (* printing *) fun string_of_typ thy = Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); fun string_of_const thy c = let val ctxt = Proof_Context.init_global thy in case Axclass.inst_of_param thy c of SOME (c, tyco) => Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]" (Proof_Context.extern_type ctxt tyco) | NONE => Proof_Context.extern_const ctxt c end; (* constants *) fun const_typ thy = Type.strip_sorts o Sign.the_const_type thy; fun args_number thy = length o binder_types o const_typ thy; fun devarify ty = let val tys = build (fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty); val vs = Name.invent Name.context Name.aT (length tys); val mapping = map2 (fn v => fn (vi, sort) => (vi, TFree (v, sort))) vs tys; in Term.typ_subst_TVars mapping ty end; fun typscheme thy (c, ty) = (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); fun typscheme_equiv (ty1, ty2) = Type.raw_instance (devarify ty1, ty2) andalso Type.raw_instance (devarify ty2, ty1); fun check_bare_const thy t = case try dest_Const t of SOME c_ty => c_ty | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); fun check_unoverload thy (c, ty) = let val c' = Axclass.unoverload_const thy (c, ty); val ty_decl = const_typ thy c'; in if typscheme_equiv (ty_decl, Logic.varifyT_global ty) then c' else error ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_const thy = check_unoverload thy o check_bare_const thy; fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; fun read_const thy = check_unoverload thy o read_bare_const thy; (** executable specifications **) (* types *) datatype type_spec = Constructors of { constructors: (string * ((string * sort) list * typ list)) list, case_combinators: string list} | Abstractor of { abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string, more_abstract_functions: string list}; fun concrete_constructors_of (Constructors {constructors, ...}) = constructors | concrete_constructors_of _ = []; fun constructors_of (Constructors {constructors, ...}) = (constructors, false) | constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) = ([(co, (vs, [ty]))], true); fun case_combinators_of (Constructors {case_combinators, ...}) = case_combinators | case_combinators_of (Abstractor _) = []; fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) = (vs, Constructors {constructors = constructors, case_combinators = insert (op =) c case_combinators}); fun projection_of (Constructors _) = NONE | projection_of (Abstractor {projection, ...}) = SOME projection; fun abstract_functions_of (Constructors _) = [] | abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) = projection :: more_abstract_functions; fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) = (vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = insert (op =) c more_abstract_functions}); fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1}, Constructors {case_combinators = case_combinators2, ...}) = Constructors {constructors = constructors, case_combinators = merge (op =) (case_combinators1, case_combinators2)} | join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1}, Abstractor {more_abstract_functions = more_abstract_functions2, ...}) = Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)}; fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2)); (* functions *) datatype fun_spec = Eqns of bool * (thm * bool) list | Proj of term * (string * string) | Abstr of thm * (string * string); val unimplemented = Eqns (true, []); fun is_unimplemented (Eqns (true, [])) = true | is_unimplemented _ = false; fun is_default (Eqns (true, _)) = true | is_default _ = false; val aborting = Eqns (false, []); fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs | associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs | associated_abstype _ = NONE; (* cases *) type case_schema = int * (int * (string * int) option list); datatype case_spec = No_Case | Case of {schema: case_schema, tycos: string list, cong: thm} | Undefined; fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map fst (map_filter I raw_cos)) | associated_datatypes _ = ([], []); (** background theory data store **) (* historized declaration data *) structure History = struct type 'a T = { entry: 'a, suppressed: bool, (*incompatible entries are merely suppressed after theory merge but sustain*) history: serial list (*explicit trace of declaration history supports non-monotonic declarations*) } Symtab.table; fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry | some_entry _ = NONE; fun lookup table = Symtab.lookup table #> some_entry; fun register key entry table = if is_some (Symtab.lookup table key) then Symtab.map_entry key (fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table; fun modify_entry key f = Symtab.map_entry key (fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history}); fun all table = Symtab.dest table |> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE); local fun merge_history join_same ({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) = let val history = merge (op =) (history1, history2); val entry = if hd history1 = hd history2 then join_same (entry1, entry2) else if hd history = hd history1 then entry1 else entry2; in {entry = entry, suppressed = false, history = history} end; in fun join join_same tables = Symtab.join (K (merge_history join_same)) tables; fun suppress key = Symtab.map_entry key (fn {entry, history, ...} => {entry = entry, suppressed = true, history = history}); fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} => {entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history}); end; end; datatype specs = Specs of { types: ((string * sort) list * type_spec) History.T, pending_eqns: (thm * bool) list Symtab.table, functions: fun_spec History.T, cases: case_spec History.T }; fun types_of (Specs {types, ...}) = types; fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns; fun functions_of (Specs {functions, ...}) = functions; fun cases_of (Specs {cases, ...}) = cases; fun make_specs (types, ((pending_eqns, functions), cases)) = Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}; val empty_specs = make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty)); fun map_specs f (Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}) = make_specs (f (types, ((pending_eqns, functions), cases))); fun merge_specs (Specs {types = types1, pending_eqns = _, functions = functions1, cases = cases1}, Specs {types = types2, pending_eqns = _, functions = functions2, cases = cases2}) = let val types = History.join join_same_types (types1, types2); val all_types = map (snd o snd) (History.all types); fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of NONE => true | SOME (tyco, abs) => (case History.lookup types tyco of NONE => false | SOME (_, Constructors _) => false | SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) => abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c)); fun check_datatypes (_, case_spec) = let val (tycos, required_constructors) = associated_datatypes case_spec; val allowed_constructors = tycos |> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types) |> map fst; in subset (op =) (required_constructors, allowed_constructors) end; val all_constructors = maps (fst o constructors_of) all_types; val functions = History.join fst (functions1, functions2) |> fold (History.suppress o fst) all_constructors |> History.suppress_except check_abstype; val cases = History.join fst (cases1, cases2) |> History.suppress_except check_datatypes; in make_specs (types, ((Symtab.empty, functions), cases)) end; val map_types = map_specs o apfst; val map_pending_eqns = map_specs o apsnd o apfst o apfst; val map_functions = map_specs o apsnd o apfst o apsnd; val map_cases = map_specs o apsnd o apsnd; (* data slots dependent on executable code *) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); local type kind = {empty: Any.T}; val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table); fun invoke f k = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => f kind | NONE => raise Fail "Invalid code data identifier"); in fun declare_data empty = let val k = serial (); val kind = {empty = empty}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; fun invoke_init k = invoke (fn kind => #empty kind) k; end; (*local*) (* global theory store *) local type data = Any.T Datatab.table; fun make_dataref () = Synchronized.var "code data" (NONE : (data * Context.theory_id) option); structure Code_Data = Theory_Data ( type T = specs * (string * (data * Context.theory_id) option Synchronized.var); val empty = (empty_specs, (Context.theory_long_name (Context.the_global_context ()), make_dataref ())); fun merge ((specs1, dataref), (specs2, _)) = (merge_specs (specs1, specs2), dataref); ); fun init_dataref thy = let val thy_name = Context.theory_long_name thy in if #1 (#2 (Code_Data.get thy)) = thy_name then NONE else SOME ((Code_Data.map o apsnd) (K (thy_name, make_dataref ())) thy) end; in val _ = Theory.setup (Theory.at_begin init_dataref); (* access to executable specifications *) val specs_of : theory -> specs = fst o Code_Data.get; fun modify_specs f thy = let val thy_name = Context.theory_long_name thy in Code_Data.map (fn (specs, _) => (f specs, (thy_name, make_dataref ()))) thy end; (* access to data dependent on executable specifications *) fun change_yield_data (kind, mk, dest) theory f = let val dataref = #2 (#2 (Code_Data.get theory)); val (datatab, thy_id) = case Synchronized.value dataref of SOME (datatab, thy_id) => if Context.eq_thy_id (Context.theory_id theory, thy_id) then (datatab, thy_id) else (Datatab.empty, Context.theory_id theory) | NONE => (Datatab.empty, Context.theory_id theory) val data = case Datatab.lookup datatab kind of SOME data => data | NONE => invoke_init kind; val result as (_, data') = f (dest data); val _ = Synchronized.change dataref ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id)); in result end; end; (*local*) (* pending function equations *) (* Ideally, *all* equations implementing a functions would be treated as *one* atomic declaration; unfortunately, we cannot implement this: the too-well-established declaration interface are Isar attributes which operate on *one* single theorem. Hence we treat such Isar declarations as "pending" and historize them as proper declarations at the end of each theory. *) fun modify_pending_eqns c f specs = let val existing_eqns = case History.lookup (functions_of specs) c of SOME (Eqns (false, eqns)) => eqns | _ => []; in specs |> map_pending_eqns (Symtab.map_default (c, existing_eqns) f) end; fun register_fun_spec c spec = map_pending_eqns (Symtab.delete_safe c) #> map_functions (History.register c spec); fun lookup_fun_spec specs c = case Symtab.lookup (pending_eqns_of specs) c of SOME eqns => Eqns (false, eqns) | NONE => (case History.lookup (functions_of specs) c of SOME spec => spec | NONE => unimplemented); fun lookup_proper_fun_spec specs c = let val spec = lookup_fun_spec specs c in if is_unimplemented spec then NONE else SOME spec end; fun all_fun_specs specs = map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c)) (union (op =) ((Symtab.keys o pending_eqns_of) specs) ((Symtab.keys o functions_of) specs)); fun historize_pending_fun_specs thy = let val pending_eqns = (pending_eqns_of o specs_of) thy; in if Symtab.is_empty pending_eqns then NONE else thy |> modify_specs (map_functions (Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns) #> map_pending_eqns (K Symtab.empty)) |> SOME end; val _ = Theory.setup (Theory.at_end historize_pending_fun_specs); (** foundation **) (* types *) fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s); fun analyze_constructor thy (c, ty) = let val _ = Thm.global_cterm_of thy (Const (c, ty)); val ty_decl = devarify (const_typ thy c); fun last_typ c_ty ty = let val tfrees = Term.add_tfreesT ty []; val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty)) handle TYPE _ => no_constr thy "bad type" c_ty val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else (); val _ = if has_duplicates (eq_fst (op =)) vs then no_constr thy "duplicate type variables in datatype" c_ty else (); val _ = if length tfrees <> length vs then no_constr thy "type variables missing in datatype" c_ty else (); in (tyco, vs) end; val (tyco, _) = last_typ (c, ty) ty_decl; val (_, vs) = last_typ (c, ty) ty; in ((tyco, map snd vs), (c, (map fst vs, ty))) end; fun constrset_of_consts thy consts = let val _ = map (fn (c, _) => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) consts; val raw_constructors = map (analyze_constructor thy) consts; val tyco = case distinct (op =) (map (fst o fst) raw_constructors) of [tyco] => tyco | [] => error "Empty constructor set" | tycos => error ("Different type constructors in constructor set: " ^ commas_quote tycos) val vs = Name.invent Name.context Name.aT (Sign.arity_number thy tyco); fun inst vs' (c, (vs, ty)) = let val the_v = the o AList.lookup (op =) (vs ~~ vs'); val ty' = map_type_tfree (fn (v, _) => TFree (the_v v, [])) ty; val (vs'', ty'') = typscheme thy (c, ty'); in (c, (vs'', binder_types ty'')) end; val constructors = map (inst vs o snd) raw_constructors; in (tyco, (map (rpair []) vs, constructors)) end; fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy); type constructors = (string * sort) list * (string * ((string * sort) list * typ list)) list; fun get_type thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec) | NONE => Sign.arity_number thy tyco |> Name.invent Name.context Name.aT |> map (rpair []) |> rpair [] |> rpair false; type abs_type = (string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string}; fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) => (vs, {abs_rep = Thm.transfer thy abs_rep, abstractor = abstractor, projection = projection}) | _ => error ("Not an abstract type: " ^ tyco); fun get_type_of_constr_or_abstr thy c = case (body_type o const_typ thy) c of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end | _ => NONE; fun is_constr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, false) => true | _ => false; fun is_abstr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, true) => true | _ => false; (* bare code equations *) (* convention for variables: ?x ?'a for free-floating theorems (e.g. in the data store) ?x 'a for certificates x 'a for final representation of equations *) exception BAD_THM of string; fun bad_thm msg = raise BAD_THM msg; datatype strictness = Silent | Liberal | Strict fun handle_strictness thm_of f strictness thy x = SOME (f x) handle BAD_THM msg => case strictness of Silent => NONE | Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE) | Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); fun is_linear thm = let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm in not (has_duplicates (op =) ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])) end; fun check_decl_ty thy (c, ty) = let val ty_decl = const_typ thy c; in if typscheme_equiv (ty_decl, ty) then () else bad_thm ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) = let fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | Free _ => bad_thm "Illegal free variable" | _ => I) t []; fun tvars_of t = fold_term_types (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v | TFree _ => bad_thm "Illegal free type variable")) t []; val lhs_vs = vars_of lhs; val rhs_vs = vars_of rhs; val lhs_tvs = tvars_of lhs; val rhs_tvs = tvars_of rhs; val _ = if null (subtract (op =) lhs_vs rhs_vs) then () else bad_thm "Free variables on right hand side of equation"; val _ = if null (subtract (op =) lhs_tvs rhs_tvs) then () else bad_thm "Free type variables on right hand side of equation"; val (head, args) = strip_comb lhs; val (c, ty) = case head of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty) | _ => bad_thm "Equation not headed by constant"; fun check _ (Abs _) = bad_thm "Abstraction on left hand side of equation" | check 0 (Var _) = () | check _ (Var _) = bad_thm "Variable with application on left hand side of equation" | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | check n (Const (c_ty as (c, ty))) = if allow_pats then let val c' = Axclass.unoverload_const thy c_ty in if n = (length o binder_types) ty then if allow_consts orelse is_constr thy c' then () else bad_thm (quote c ^ " is not a constructor, on left hand side of equation") else bad_thm ("Partially applied constant " ^ quote c ^ " on left hand side of equation") end else bad_thm ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation") val _ = map (check 0) args; val _ = if allow_nonlinear orelse is_linear thm then () else bad_thm "Duplicate variables on left hand side of equation"; val _ = if (is_none o Axclass.class_of_param thy) c then () else bad_thm "Overloaded constant as head in equation"; val _ = if not (is_constr thy c) then () else bad_thm "Constructor as head in equation"; val _ = if not (is_abstr thy c) then () else bad_thm "Abstractor as head in equation"; val _ = check_decl_ty thy (c, ty); val _ = case strip_type ty of (Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of SOME (_, type_spec) => (case projection_of type_spec of SOME proj => if c = proj then bad_thm "Projection as head in equation" else () | _ => ()) | _ => ()) | _ => (); in () end; local fun raw_assert_eqn thy check_patterns (thm, proper) = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val _ = check_eqn thy {allow_nonlinear = not proper, allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs); in (thm, proper) end; fun raw_assert_abs_eqn thy some_tyco thm = let val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val (proj_t, lhs) = dest_comb full_lhs handle TERM _ => bad_thm "Not an abstract equation"; val (proj, ty) = dest_Const proj_t handle TERM _ => bad_thm "Not an abstract equation"; val (tyco, Ts) = (dest_Type o domain_type) ty handle TERM _ => bad_thm "Not an abstract equation" | TYPE _ => bad_thm "Not an abstract equation"; val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () else bad_thm ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco') | NONE => (); val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec) | _ => bad_thm ("Not an abstract type: " ^ tyco); val _ = if proj = proj' then () else bad_thm ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj'); val _ = check_eqn thy {allow_nonlinear = false, allow_consts = false, allow_pats = false} thm (lhs, rhs); val _ = if ListPair.all (fn (T, (_, sort)) => Sign.of_sort thy (T, sort)) (Ts, vs) then () else error ("Type arguments do not satisfy sort constraints of abstype certificate."); in (thm, (tyco, abs')) end; in fun generic_assert_eqn strictness thy check_patterns eqn = handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn; fun generic_assert_abs_eqn strictness thy check_patterns thm = handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm; end; fun assert_eqn thy = the o generic_assert_eqn Strict thy true; fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco; val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun const_typ_eqn thy thm = let val (c, ty) = head_eqn thm; val c' = Axclass.unoverload_const thy (c, ty); (*permissive wrt. to overloaded constants!*) in (c', ty) end; fun const_eqn thy = fst o const_typ_eqn thy; fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun mk_proj tyco vs ty abs rep = let val ty_abs = Type (tyco, map TFree vs); val xarg = Var (("x", 0), ty); in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; (* technical transformations of code equations *) fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); fun same_arity thy thms = let val lhs_rhss = map (Logic.dest_equals o Thm.plain_prop_of) thms; val k = fold (Integer.max o length o snd o strip_comb o fst) lhs_rhss 0; fun expand_eta (lhs, rhs) thm = let val l = k - length (snd (strip_comb lhs)); val (raw_vars, _) = Term.strip_abs_eta l rhs; val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) raw_vars; fun expand (v, ty) thm = Drule.fun_cong_rule thm (Thm.global_cterm_of thy (Var ((v, 0), ty))); in thm |> fold expand vars |> Conv.fconv_rule Drule.beta_eta_conversion end; in map2 expand_eta lhs_rhss thms end; fun mk_desymbolization pre post mk vs = let val names = map (pre o fst o fst) vs |> map (Name.desymbolize (SOME false)) |> Name.variant_list [] |> map post; in map_filter (fn (((v, i), x), v') => if v = v' andalso i = 0 then NONE else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) end; fun desymbolize_tvars thy thms = let val tvs = build (fold (Term.add_tvars o Thm.prop_of) thms); val instT = mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs; in map (Thm.instantiate (TVars.make instT, Vars.empty)) thms end; fun desymbolize_vars thy thm = let val vs = Term.add_vars (Thm.prop_of thm) []; val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs; in Thm.instantiate (TVars.empty, Vars.make inst) thm end; fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); (* preparation and classification of code equations *) fun prep_eqn strictness thy = apfst (meta_rewrite thy) #> generic_assert_eqn strictness thy false #> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn)); fun prep_eqns strictness thy = map_filter (prep_eqn strictness thy) #> AList.group (op =); fun prep_abs_eqn strictness thy = meta_rewrite thy #> generic_assert_abs_eqn strictness thy NONE #> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn)); fun prep_maybe_abs_eqn thy raw_thm = let val thm = meta_rewrite thy raw_thm; val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm; in case some_abs_thm of SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco)) | NONE => generic_assert_eqn Liberal thy false (thm, false) |> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE))) end; (* abstype certificates *) local fun raw_abstype_cert thy proto_thm = let val thm = (Axclass.unoverload (Proof_Context.init_global thy) o meta_rewrite thy) proto_thm; val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb) o apfst dest_Const o dest_comb) lhs handle TERM _ => bad_thm "Not an abstype certificate"; val _ = apply2 (fn c => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep); val _ = check_decl_ty thy (abs, raw_ty); val _ = check_decl_ty thy (rep, rep_ty); val _ = if length (binder_types raw_ty) = 1 then () else bad_thm "Bad type for abstract constructor"; val _ = (fst o dest_Var) param handle TERM _ => bad_thm "Not an abstype certificate"; val _ = if param = rhs then () else bad_thm "Not an abstype certificate"; val ((tyco, sorts), (abs, (vs, ty'))) = analyze_constructor thy (abs, devarify raw_ty); val ty = domain_type ty'; val (vs', _) = typscheme thy (abs, ty'); in (tyco, (vs ~~ sorts, ((abs, (vs', ty)), (rep, thm)))) end; in fun check_abstype_cert strictness thy proto_thm = handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm; end; (* code equation certificates *) fun build_head thy (c, ty) = Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); fun get_head thy cert_thm = let val [head] = Thm.chyps_of cert_thm; val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; in (typscheme thy (c, ty), head) end; fun typscheme_projection thy = typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; fun typscheme_abs thy = typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of; fun constrain_thm thy vs sorts thm = let val mapping = map2 (fn (v, sort) => fn sort' => (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; val instT = TVars.build (fold2 (fn (v, sort) => fn (_, sort') => TVars.add (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping); val subst = (Term.map_types o map_type_tfree) (fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); in thm |> Thm.varifyT_global |> Thm.instantiate (instT, Vars.empty) |> pair subst end; fun concretify_abs thy tyco abs_thm = let val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco; val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm val ty = fastype_of lhs; val ty_abs = (fastype_of o snd o dest_comb) lhs; val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs)); val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm]; in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; fun add_rhss_of_eqn thy t = let val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t; fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) | add_const _ = I val add_consts = fold_aterms add_const in add_consts rhs o fold add_consts args end; val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global; abstype cert = Nothing of thm | Equations of thm * bool list | Projection of term * string | Abstract of thm * string with fun dummy_thm ctxt c = let val thy = Proof_Context.theory_of ctxt; val raw_ty = devarify (const_typ thy c); val (vs, _) = typscheme thy (c, raw_ty); val sortargs = case Axclass.class_of_param thy c of SOME class => [[class]] | NONE => (case get_type_of_constr_or_abstr thy c of SOME (tyco, _) => (map snd o fst o the) (AList.lookup (op =) ((snd o fst o get_type thy) tyco) c) | NONE => replicate (length vs) []); val the_sort = the o AList.lookup (op =) (map fst vs ~~ sortargs); val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty val chead = build_head thy (c, ty); in Thm.weaken chead Drule.dummy_thm end; fun nothing_cert ctxt c = Nothing (dummy_thm ctxt c); fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, []) | cert_of_eqns ctxt c raw_eqns = let val thy = Proof_Context.theory_of ctxt; val eqns = burrow_fst (canonize_thms thy) raw_eqns; val _ = map (assert_eqn thy) eqns; val (thms, propers) = split_list eqns; val _ = map (fn thm => if c = const_eqn thy thm then () else error ("Wrong head of code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms; val tvars_of = build_rev o Term.add_tvarsT; val vss = map (tvars_of o snd o head_eqn) thms; val inter_sorts = build o fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd); val sorts = map_transpose inter_sorts vss; val vts = Name.invent_names Name.context Name.aT sorts; fun instantiate vs = Thm.instantiate (TVars.make (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts), Vars.empty); val thms' = map2 instantiate vss thms; val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms')))); fun head_conv ct = if can Thm.dest_comb ct then Conv.fun_conv head_conv ct else Conv.rewr_conv head_thm ct; val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv); val cert_thm = Conjunction.intr_balanced (map rewrite_head thms'); in Equations (cert_thm, propers) end; fun cert_of_proj ctxt proj tyco = let val thy = Proof_Context.theory_of ctxt val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco; val _ = if proj = proj' then () else error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj); in Projection (mk_proj tyco vs ty abs proj, tyco) end; fun cert_of_abs ctxt tyco c raw_abs_thm = let val thy = Proof_Context.theory_of ctxt; val abs_thm = singleton (canonize_thms thy) raw_abs_thm; val _ = assert_abs_eqn thy (SOME tyco) abs_thm; val _ = if c = const_abs_eqn thy abs_thm then () else error ("Wrong head of abstract code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy abs_thm); in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; fun constrain_cert_thm thy sorts cert_thm = let val ((vs, _), head) = get_head thy cert_thm; val (subst, cert_thm') = cert_thm |> Thm.implies_intr head |> constrain_thm thy vs sorts; val head' = Thm.term_of head |> subst |> Thm.global_cterm_of thy; val cert_thm'' = cert_thm' |> Thm.elim_implies (Thm.assume head'); in cert_thm'' end; fun constrain_cert thy sorts (Nothing cert_thm) = Nothing (constrain_cert_thm thy sorts cert_thm) | constrain_cert thy sorts (Equations (cert_thm, propers)) = Equations (constrain_cert_thm thy sorts cert_thm, propers) | constrain_cert _ _ (cert as Projection _) = cert | constrain_cert thy sorts (Abstract (abs_thm, tyco)) = Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); fun conclude_cert (Nothing cert_thm) = Nothing (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context) | conclude_cert (Equations (cert_thm, propers)) = Equations (Thm.close_derivation \<^here> cert_thm |> Thm.trim_context, propers) | conclude_cert (cert as Projection _) = cert | conclude_cert (Abstract (abs_thm, tyco)) = Abstract (Thm.close_derivation \<^here> abs_thm |> Thm.trim_context, tyco); fun typscheme_of_cert thy (Nothing cert_thm) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Equations (cert_thm, _)) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Projection (proj, _)) = typscheme_projection thy proj | typscheme_of_cert thy (Abstract (abs_thm, _)) = typscheme_abs thy abs_thm; fun typargs_deps_of_cert thy (Nothing cert_thm) = let val vs = (fst o fst) (get_head thy cert_thm); in (vs, []) end | typargs_deps_of_cert thy (Equations (cert_thm, propers)) = let val vs = (fst o fst) (get_head thy cert_thm); val equations = if null propers then [] else Thm.prop_of cert_thm |> Logic.dest_conjunction_balanced (length propers); in (vs, build (fold (add_rhss_of_eqn thy) equations)) end | typargs_deps_of_cert thy (Projection (t, _)) = (fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = let val vs = fst (typscheme_abs thy abs_thm); val (_, concrete_thm) = concretify_abs thy tyco abs_thm; in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end; fun equations_of_cert thy (cert as Nothing _) = (typscheme_of_cert thy cert, NONE) | equations_of_cert thy (cert as Equations (cert_thm, propers)) = let val tyscm = typscheme_of_cert thy cert; val thms = if null propers then [] else cert_thm |> Thm.transfer thy |> Local_Defs.expand [snd (get_head thy cert_thm)] |> Thm.varifyT_global |> Conjunction.elim_balanced (length propers); fun abstractions (args, rhs) = (map (pair NONE) args, (NONE, rhs)); in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end | equations_of_cert thy (Projection (t, tyco)) = let val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco; val tyscm = typscheme_projection thy t; val t' = Logic.varify_types_global t; fun abstractions (args, rhs) = (map (pair (SOME abs)) args, (NONE, rhs)); in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end | equations_of_cert thy (Abstract (abs_thm, tyco)) = let val tyscm = typscheme_abs thy abs_thm; val (abs, concrete_thm) = concretify_abs thy tyco (Thm.transfer thy abs_thm); fun abstractions (args, rhs) = (map (pair NONE) args, (SOME abs, rhs)); in (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm, (SOME (Thm.varifyT_global abs_thm), true))]) end; fun pretty_cert _ (Nothing _) = [] | pretty_cert thy (cert as Equations _) = (map_filter (Option.map (Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy)) o fst o snd) o these o snd o equations_of_cert thy) cert | pretty_cert thy (Projection (t, _)) = [Syntax.pretty_term_global thy (Logic.varify_types_global t)] | pretty_cert thy (Abstract (abs_thm, _)) = [(Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm]; end; (* code certificate access with preprocessing *) fun eqn_conv conv ct = let fun lhs_conv ct = if can Thm.dest_comb ct then Conv.combination_conv lhs_conv conv ct else Conv.all_conv ct; in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end; fun rewrite_eqn conv ctxt = singleton (Variable.trade (K (map (Conv.fconv_rule (conv (Simplifier.rewrite ctxt))))) ctxt) fun apply_functrans ctxt functrans = let fun trace_eqns s eqns = (Pretty.writeln o Pretty.chunks) (Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns); val tracing = if Config.get ctxt simp_trace then trace_eqns else (K o K) (); in tap (tracing "before function transformation") #> (perhaps o perhaps_loop o perhaps_apply) functrans #> tap (tracing "after function transformation") end; fun preprocess conv ctxt = rewrite_eqn conv ctxt #> Axclass.unoverload ctxt; fun get_cert ctxt functrans c = case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of NONE => nothing_cert ctxt c | SOME (Eqns (_, eqns)) => eqns |> (map o apfst) (Thm.transfer' ctxt) |> apply_functrans ctxt functrans |> (map o apfst) (preprocess eqn_conv ctxt) |> cert_of_eqns ctxt c | SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco | SOME (Abstr (abs_thm, (tyco, _))) => abs_thm |> Thm.transfer' ctxt |> preprocess Conv.arg_conv ctxt |> cert_of_abs ctxt tyco c; (* case certificates *) local fun raw_case_cert thm = let val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; val _ = case head of Free _ => () | Var _ => () | _ => raise TERM ("case_cert", []); val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; val (Const (case_const, _), raw_params) = strip_comb case_expr; val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; val _ = if n = ~1 then raise TERM ("case_cert", []) else (); val params = map (fst o dest_Var) (nth_drop n raw_params); fun dest_case t = let val (head' $ t_co, rhs) = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val (Const (co, _), args) = strip_comb t_co; val (Var (param, _), args') = strip_comb rhs; val _ = if args' = args then () else raise TERM ("case_cert", []); in (param, co) end; fun analyze_cases cases = let val co_list = build (fold (AList.update (op =) o dest_case) cases); in map (AList.lookup (op =) co_list) params end; fun analyze_let t = let val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val _ = if arg' = arg then () else raise TERM ("case_cert", []); val _ = if [param'] = params then () else raise TERM ("case_cert", []); in [] end; fun analyze (cases as [let_case]) = (analyze_cases cases handle Bind => analyze_let let_case) | analyze cases = analyze_cases cases; in (case_const, (n, analyze cases)) end; in fun case_cert thm = raw_case_cert thm handle Bind => error "bad case certificate" | TERM _ => error "bad case certificate"; end; fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy); fun get_case_schema thy c = case lookup_case_spec thy c of SOME (Case {schema, ...}) => SOME schema | _ => NONE; fun get_case_cong thy c = case lookup_case_spec thy c of SOME (Case {cong, ...}) => SOME cong | _ => NONE; fun is_undefined thy c = case lookup_case_spec thy c of SOME Undefined => true | _ => false; (* diagnostic *) fun print_codesetup thy = let val ctxt = Proof_Context.init_global thy; val specs = specs_of thy; fun pretty_equations const thms = (Pretty.block o Pretty.fbreaks) (Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms); fun pretty_function (const, Eqns (_, eqns)) = pretty_equations const (map fst eqns) | pretty_function (const, Proj (proj, _)) = Pretty.block [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; fun pretty_typ (tyco, vs) = Pretty.str (string_of_typ thy (Type (tyco, map TFree vs))); fun pretty_type_spec (typ, (cos, abstract)) = if null cos then pretty_typ typ else (Pretty.block o Pretty.breaks) ( pretty_typ typ :: Pretty.str "=" :: (if abstract then [Pretty.str "(abstract)"] else []) @ separate (Pretty.str "|") (map (fn (c, (_, [])) => Pretty.str (string_of_const thy c) | (c, (_, tys)) => (Pretty.block o Pretty.breaks) (Pretty.str (string_of_const thy c) :: Pretty.str "of" :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) ); fun pretty_case_param NONE = "" | pretty_case_param (SOME (c, _)) = string_of_const thy c fun pretty_case (const, Case {schema = (_, (_, [])), ...}) = Pretty.str (string_of_const thy const) | pretty_case (const, Case {schema = (_, (_, cos)), ...}) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str "with", (Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos] | pretty_case (const, Undefined) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str ""]; val functions = all_fun_specs specs |> sort (string_ord o apply2 fst); val types = History.all (types_of specs) |> map (fn (tyco, (vs, spec)) => ((tyco, vs), constructors_of spec)) |> sort (string_ord o apply2 (fst o fst)); val cases = History.all (cases_of specs) |> filter (fn (_, No_Case) => false | _ => true) |> sort (string_ord o apply2 fst); in Pretty.writeln_chunks [ Pretty.block ( Pretty.str "types:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_type_spec) types ), Pretty.block ( Pretty.str "functions:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_function) functions ), Pretty.block ( Pretty.str "cases:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_case) cases ) ] end; (** declaration of executable ingredients **) (* plugins for dependent applications *) structure Codetype_Plugin = Plugin(type T = string); val codetype_plugin = Plugin_Name.declare_setup \<^binding>\codetype\; fun type_interpretation f = Codetype_Plugin.interpretation codetype_plugin (fn tyco => Local_Theory.background_theory (fn thy => thy |> Sign.root_path |> Sign.add_path (Long_Name.qualifier tyco) |> f tyco |> Sign.restore_naming thy)); fun datatype_interpretation f = type_interpretation (fn tyco => fn thy => case get_type thy tyco of (spec, false) => f (tyco, spec) thy | (_, true) => thy ); fun abstype_interpretation f = type_interpretation (fn tyco => fn thy => case try (get_abstype_spec thy) tyco of SOME spec => f (tyco, spec) thy | NONE => thy ); fun register_tyco_for_plugin tyco = Named_Target.theory_map (Codetype_Plugin.data_default tyco); (* abstract code declarations *) fun code_declaration strictness lift_phi f x = - Local_Theory.declaration {syntax = false, pervasive = false} + Local_Theory.declaration {syntax = false, pervasive = false, pos = Position.thread_data ()} (fn phi => Context.mapping (f strictness (lift_phi phi x)) I); (* types *) fun invalidate_constructors_of (_, type_spec) = fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec)); fun invalidate_abstract_functions_of (_, type_spec) = fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec); fun invalidate_case_combinators_of (_, type_spec) = fold (fn c => History.register c No_Case) (case_combinators_of type_spec); fun register_type (tyco, vs_typ_spec) specs = let val olds = the_list (History.lookup (types_of specs) tyco); in specs |> map_functions (fold invalidate_abstract_functions_of olds #> invalidate_constructors_of vs_typ_spec) |> map_cases (fold invalidate_case_combinators_of olds) |> map_types (History.register tyco vs_typ_spec) end; fun declare_datatype_global proto_constrs thy = let fun unoverload_const_typ (c, ty) = (Axclass.unoverload_const thy (c, ty), ty); val constrs = map unoverload_const_typ proto_constrs; val (tyco, (vs, cos)) = constrset_of_consts thy constrs; in thy |> modify_specs (register_type (tyco, (vs, Constructors {constructors = cos, case_combinators = []}))) |> register_tyco_for_plugin tyco end; fun declare_datatype_cmd raw_constrs thy = declare_datatype_global (map (read_bare_const thy) raw_constrs) thy; fun generic_declare_abstype strictness proto_thm thy = case check_abstype_cert strictness thy proto_thm of SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) => thy |> modify_specs (register_type (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, abs_rep = Thm.trim_context abs_rep, more_abstract_functions = []})) #> register_fun_spec proj (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs)))) |> register_tyco_for_plugin tyco | NONE => thy; val declare_abstype_global = generic_declare_abstype Strict; val declare_abstype = code_declaration Liberal Morphism.thm generic_declare_abstype; (* functions *) (* strictness wrt. shape of theorem propositions: * default equations: silent * using declarations and attributes: warnings (after morphism application!) * using global declarations (... -> thy -> thy): strict * internal processing after storage: strict *) local fun subsumptive_add thy verbose (thm, proper) eqns = let val args_of = drop_prefix is_Var o rev o snd o strip_comb o Term.map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of o Thm.transfer thy; val args = args_of thm; val incr_idx = Logic.incr_indexes ([], [], Thm.maxidx_of thm + 1); fun matches_args args' = let val k = length args' - length args in if k >= 0 then Pattern.matchess thy (args, (map incr_idx o drop k) args') else false end; fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then (if verbose then warning ("Code generator: dropping subsumed code equation\n" ^ Thm.string_of_thm_global thy thm') else (); true) else false; in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end; fun add_eqn_for (c, eqn) thy = thy |> modify_specs (modify_pending_eqns c (subsumptive_add thy true eqn)); fun add_eqns_for default (c, proto_eqns) thy = thy |> modify_specs (fn specs => if is_default (lookup_fun_spec specs c) orelse not default then let val eqns = [] |> fold_rev (subsumptive_add thy (not default)) proto_eqns; in specs |> register_fun_spec c (Eqns (default, eqns)) end else specs); fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) = modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm |> Thm.trim_context, tyco_abs)) #> map_types (History.modify_entry tyco (add_abstract_function c))) in fun generic_declare_eqns default strictness raw_eqns thy = fold (add_eqns_for default) (prep_eqns strictness thy raw_eqns) thy; fun generic_add_eqn strictness raw_eqn thy = fold add_eqn_for (the_list (prep_eqn strictness thy raw_eqn)) thy; fun generic_declare_abstract_eqn strictness raw_abs_eqn thy = fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy; fun add_maybe_abs_eqn_liberal thm thy = case prep_maybe_abs_eqn thy thm of SOME (c, (eqn, NONE)) => add_eqn_for (c, eqn) thy | SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy | NONE => thy; end; val declare_default_eqns_global = generic_declare_eqns true Silent; val declare_default_eqns = code_declaration Silent (map o apfst o Morphism.thm) (generic_declare_eqns true); val declare_eqns_global = generic_declare_eqns false Strict; val declare_eqns = code_declaration Liberal (map o apfst o Morphism.thm) (generic_declare_eqns false); val add_eqn_global = generic_add_eqn Strict; fun del_eqn_global thm thy = case prep_eqn Liberal thy (thm, false) of SOME (c, (thm, _)) => modify_specs (modify_pending_eqns c (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')))) thy | NONE => thy; val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict; val declare_abstract_eqn = code_declaration Liberal Morphism.thm generic_declare_abstract_eqn; fun declare_aborting_global c = modify_specs (register_fun_spec c aborting); fun declare_unimplemented_global c = modify_specs (register_fun_spec c unimplemented); (* cases *) fun case_cong thy case_const (num_args, (pos, _)) = let val ([x, y], ctxt) = fold_map Name.variant ["A", "A'"] Name.context; val (zs, _) = fold_map Name.variant (replicate (num_args - 1) "") ctxt; val (ws, vs) = chop pos zs; val T = devarify (const_typ thy case_const); val Ts = binder_types T; val T_cong = nth Ts pos; fun mk_prem z = Free (z, T_cong); fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts); val (prem, concl) = apply2 Logic.mk_equals (apply2 mk_prem (x, y), apply2 mk_concl (x, y)); in Goal.prove_sorry_global thy (x :: y :: zs) [prem] concl (fn {context = ctxt', prems} => Simplifier.rewrite_goals_tac ctxt' prems THEN ALLGOALS (Proof_Context.fact_tac ctxt' [Drule.reflexive_thm])) end; fun declare_case_global thm thy = let val (case_const, (k, cos)) = case_cert thm; fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of SOME (c, false) => SOME c | _ => NONE; val cos_with_tycos = (map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos; val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of [] => () | cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs); val tycos = distinct (op =) (map_filter snd cos_with_tycos); val schema = (1 + Int.max (1, length cos), (k, (map o Option.map) (fn c => (c, args_number thy c)) cos)); val cong = case_cong thy case_const schema; in thy |> modify_specs (map_cases (History.register case_const (Case {schema = schema, tycos = tycos, cong = cong})) #> map_types (fold (fn tyco => History.modify_entry tyco (add_case_combinator case_const)) tycos)) end; fun declare_undefined_global c = (modify_specs o map_cases) (History.register c Undefined); (* attributes *) fun code_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); fun code_thm_attribute g f = Scan.lift (g |-- Scan.succeed (code_attribute f)); fun code_const_attribute g f = Scan.lift (g -- Args.colon) |-- Scan.repeat1 Args.term >> (fn ts => code_attribute (K (fold (fn t => fn thy => f ((check_const thy o Logic.unvarify_types_global) t) thy) ts))); val _ = Theory.setup (let val code_attribute_parser = code_thm_attribute (Args.$$$ "equation") (fn thm => generic_add_eqn Liberal (thm, true)) || code_thm_attribute (Args.$$$ "nbe") (fn thm => generic_add_eqn Liberal (thm, false)) || code_thm_attribute (Args.$$$ "abstract") (generic_declare_abstract_eqn Liberal) || code_thm_attribute (Args.$$$ "abstype") (generic_declare_abstype Liberal) || code_thm_attribute Args.del del_eqn_global || code_const_attribute (Args.$$$ "abort") declare_aborting_global || code_const_attribute (Args.$$$ "drop") declare_unimplemented_global || Scan.succeed (code_attribute add_maybe_abs_eqn_liberal); in Attrib.setup \<^binding>\code\ code_attribute_parser "declare theorems for code generation" end); end; (*struct*) (* type-safe interfaces for data dependent on executable code *) functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = struct type T = Data.T; exception Data of T; fun dest (Data x) = x val kind = Code.declare_data (Data Data.empty); val data_op = (kind, Data, dest); fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f | change_yield NONE f = f Data.empty fun change some_thy f = snd (change_yield some_thy (pair () o f)); end; structure Code : CODE = struct open Code; end; diff --git a/src/Pure/Isar/entity.ML b/src/Pure/Isar/entity.ML --- a/src/Pure/Isar/entity.ML +++ b/src/Pure/Isar/entity.ML @@ -1,58 +1,59 @@ (* Title: Pure/Isar/entity.ML Author: Makarius Entity definitions within a global or local theory context. *) signature ENTITY = sig type 'a data_ops = {get_data: Context.generic -> 'a Name_Space.table, put_data: 'a Name_Space.table -> Context.generic -> Context.generic} val define_global: 'a data_ops -> binding -> 'a -> theory -> string * theory val define: 'a data_ops -> binding -> 'a -> local_theory -> string * local_theory end; structure Entity: ENTITY = struct (* context data *) type 'a data_ops = {get_data: Context.generic -> 'a Name_Space.table, put_data: 'a Name_Space.table -> Context.generic -> Context.generic}; (* global definition (foundation) *) fun define_global {get_data, put_data} b x thy = let val context = Context.Theory thy; val (name, data') = Name_Space.define context true (b, x) (get_data context); in (name, Context.the_theory (put_data data' context)) end; (* local definition *) fun alias {get_data, put_data} binding name = - Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val naming = Name_Space.naming_of context; - val binding' = Morphism.binding phi binding; - val data' = Name_Space.alias_table naming binding' name (get_data context); - in put_data data' context end); + Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} + (fn phi => fn context => + let + val naming = Name_Space.naming_of context; + val binding' = Morphism.binding phi binding; + val data' = Name_Space.alias_table naming binding' name (get_data context); + in put_data data' context end); fun transfer {get_data, put_data} ctxt = let val data0 = get_data (Context.Theory (Proof_Context.theory_of ctxt)); val data' = Name_Space.merge_tables (data0, get_data (Context.Proof ctxt)); in Context.proof_map (put_data data') ctxt end; fun define ops binding x = Local_Theory.background_theory_result (define_global ops binding x) #-> (fn name => Local_Theory.map_contexts (K (transfer ops)) #> alias ops binding name #> pair name); end; diff --git a/src/Pure/Isar/expression.ML b/src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML +++ b/src/Pure/Isar/expression.ML @@ -1,881 +1,881 @@ (* Title: Pure/Isar/expression.ML Author: Clemens Ballarin, TU Muenchen Locale expressions and user interface layer of locales. *) signature EXPRESSION = sig (* Locale expressions *) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list type 'term rewrites = (Attrib.binding * 'term) list type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list type expression_i = (string, term) expr * (binding * typ option * mixfix) list type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list (* Processing of context statements *) val cert_statement: Element.context_i list -> Element.statement_i -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context val read_statement: Element.context list -> Element.statement -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context (* Declaring locales *) val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) (*FIXME*) val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val add_locale: binding -> binding -> Bundle.name list -> expression_i -> Element.context_i list -> theory -> string * local_theory val add_locale_cmd: binding -> binding -> (xstring * Position.T) list -> expression -> Element.context list -> theory -> string * local_theory (* Processing of locale expressions *) val cert_goal_expression: expression_i -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context val read_goal_expression: expression -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context end; structure Expression : EXPRESSION = struct datatype ctxt = datatype Element.ctxt; (*** Expressions ***) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list; type 'term rewrites = (Attrib.binding * 'term) list; type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list; type expression_i = (string, term) expr * (binding * typ option * mixfix) list; type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; (** Internalise locale names in expr **) fun check_expr thy instances = map (apfst (Locale.check thy)) instances; (** Parameters of expression **) (*Sanity check of instantiations and extraction of implicit parameters. The latter only occurs iff strict = false. Positional instantiations are extended to match full length of parameter list of instantiated locale.*) fun parameters_of thy strict (expr, fixed) = let val ctxt = Proof_Context.init_global thy; fun reject_dups message xs = (case duplicates (op =) xs of [] => () | dups => error (message ^ commas dups)); fun parm_eq ((p1, mx1), (p2, mx2)) = p1 = p2 andalso (Mixfix.equal (mx1, mx2) orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^ Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); fun params_inst (loc, (prfx, (Positional insts, eqns))) = let val ps = params_loc loc; val d = length ps - length insts; val insts' = if d < 0 then error ("More arguments than parameters in instantiation of locale " ^ quote (Locale.markup_name ctxt loc)) else insts @ replicate d NONE; val ps' = (ps ~~ insts') |> map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); in (ps', (loc, (prfx, (Positional insts', eqns)))) end | params_inst (loc, (prfx, (Named insts, eqns))) = let val _ = reject_dups "Duplicate instantiation of the following parameter(s): " (map fst insts); val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => if AList.defined (op =) ps p then AList.delete (op =) p ps else error (quote p ^ " not a parameter of instantiated expression")); in (ps', (loc, (prfx, (Named insts, eqns)))) end; fun params_expr is = let val (is', ps') = fold_map (fn i => fn ps => let val (ps', i') = params_inst i; val ps'' = distinct parm_eq (ps @ ps'); in (i', ps'') end) is [] in (ps', is') end; val (implicit, expr') = params_expr expr; val implicit' = map #1 implicit; val fixed' = map (Variable.check_name o #1) fixed; val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; val implicit'' = if strict then [] else let val _ = reject_dups "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed'); in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; in (expr', implicit'' @ fixed) end; (** Read instantiation **) (* Parse positional or named instantiation *) local fun prep_inst prep_term ctxt parms (Positional insts) = (insts ~~ parms) |> map (fn (NONE, p) => Free (p, dummyT) | (SOME t, _) => prep_term ctxt t) | prep_inst prep_term ctxt parms (Named insts) = parms |> map (fn p => (case AList.lookup (op =) insts p of SOME t => prep_term ctxt t | NONE => Free (p, dummyT))); in fun parse_inst x = prep_inst Syntax.parse_term x; fun make_inst x = prep_inst (K I) x; end; (* Instantiation morphism *) fun inst_morphism params ((prfx, mandatory), insts') ctxt = let (* parameters *) val parm_types = map #2 params; val type_parms = fold Term.add_tfreesT parm_types []; (* type inference *) val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; val type_parms' = fold Term.add_tvarsT parm_types' []; val checked = (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) val (type_parms'', insts'') = chop (length type_parms') checked; (* context *) val ctxt' = fold Proof_Context.augment checked ctxt; val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; (* instantiation *) val instT = TFrees.build (fold2 (fn v => fn T => not (TFree v = T) ? TFrees.add (v, T)) type_parms (map Logic.dest_type type_parms'')); val cert_inst = Frees.build (fold2 (fn v => fn t => not (Free v = t) ? Frees.add (v, cert t)) (map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) insts''); in (Element.instantiate_normalize_morphism (TFrees.map (K certT) instT, cert_inst) $> Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') end; (*** Locale processing ***) (** Parsing **) fun parse_elem prep_typ prep_term ctxt = Element.map_ctxt {binding = I, typ = prep_typ ctxt, term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), fact = I, attrib = I}; fun prepare_stmt prep_prop prep_obtains ctxt stmt = (case stmt of Element.Shows raw_shows => raw_shows |> (map o apsnd o map) (fn (t, ps) => (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) | Element.Obtains raw_obtains => let val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); (** Simultaneous type inference: instantiations + elements + statement **) local fun mk_type T = (Logic.mk_type T, []); fun mk_term t = (t, []); fun mk_propp (p, pats) = (Type.constraint propT p, pats); fun dest_type (T, []) = Logic.dest_type T; fun dest_term (t, []) = t; fun dest_propp (p, pats) = (p, pats); fun extract_inst (_, (_, ts)) = map mk_term ts; fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)); fun extract_eqns es = map (mk_term o snd) es; fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs; fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | extract_elem (Notes _) = [] | extract_elem (Lazy_Notes _) = []; fun restore_elem (Fixes fixes, css) = (fixes ~~ css) |> map (fn ((x, _, mx), cs) => (x, cs |> map dest_type |> try hd, mx)) |> Fixes | restore_elem (Constrains csts, css) = (csts ~~ css) |> map (fn ((x, _), cs) => (x, cs |> map dest_type |> hd)) |> Constrains | restore_elem (Assumes asms, css) = (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes | restore_elem (Defines defs, css) = (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | restore_elem (elem as Notes _, _) = elem | restore_elem (elem as Lazy_Notes _, _) = elem; fun prep (_, pats) (ctxt, t :: ts) = let val ctxt' = Proof_Context.augment t ctxt in ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), (ctxt', ts)) end; fun check cs ctxt = let val (cs', (ctxt', _)) = fold_map prep cs (ctxt, Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); in (cs', ctxt') end; in fun check_autofix insts eqnss elems concl ctxt = let val inst_cs = map extract_inst insts; val eqns_cs = map extract_eqns eqnss; val elem_css = map extract_elem elems; val concl_cs = (map o map) mk_propp (map snd concl); (* Type inference *) val (inst_cs' :: eqns_cs' :: css', ctxt') = (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; val (elem_css', [concl_cs']) = chop (length elem_css) css'; in ((map restore_inst (insts ~~ inst_cs'), map restore_eqns (eqnss ~~ eqns_cs'), map restore_elem (elems ~~ elem_css'), map fst concl ~~ concl_cs'), ctxt') end; end; (** Prepare locale elements **) fun declare_elem prep_var (Fixes fixes) ctxt = let val (vars, _) = fold_map prep_var fixes ctxt in ctxt |> Proof_Context.add_fixes vars |> snd end | declare_elem prep_var (Constrains csts) ctxt = ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd | declare_elem _ (Assumes _) ctxt = ctxt | declare_elem _ (Defines _) ctxt = ctxt | declare_elem _ (Notes _) ctxt = ctxt | declare_elem _ (Lazy_Notes _) ctxt = ctxt; (** Finish locale elements **) fun finish_inst ctxt (loc, (prfx, inst)) = let val thy = Proof_Context.theory_of ctxt; val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; in (loc, morph) end; fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => let val x = Binding.name_of binding in (binding, AList.lookup (op =) parms x, mx) end); local fun closeup _ _ false elem = elem | closeup (outer_ctxt, ctxt) parms true elem = let (* FIXME consider closing in syntactic phase -- before type checking *) fun close_frees t = let val rev_frees = Term.fold_aterms (fn Free (x, T) => if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t []; in fold (Logic.all o Free) rev_frees t end; fun no_binds [] = [] | no_binds _ = error "Illegal term bindings in context element"; in (case elem of Assumes asms => Assumes (asms |> map (fn (a, propps) => (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | e => e) end; in fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | finish_elem _ _ _ (Constrains _) = Constrains [] | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | finish_elem _ _ _ (elem as Notes _) = elem | finish_elem _ _ _ (elem as Lazy_Notes _) = elem; end; (** Process full context statement: instantiations + elements + statement **) (* Interleave incremental parsing and type inference over entire parsed stretch. *) local fun abs_def ctxt = Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; fun prep_full_context_statement parse_typ parse_prop prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt ctxt1 = let val thy = Proof_Context.theory_of ctxt1; val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = let val params = map #1 (Locale.params_of thy loc); val inst' = prep_inst ctxt (map #1 params) inst; val parm_types' = params |> map (#2 #> Logic.varifyT_global #> Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> Type_Infer.paramify_vars); val inst'' = map2 Type.constraint parm_types' inst'; val insts' = insts @ [(loc, (prfx, inst''))]; val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt; val inst''' = insts'' |> List.last |> snd |> snd; val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt; val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 handle ERROR msg => if null eqns then error msg else (Locale.tracing ctxt1 (msg ^ "\nFalling back to reading rewrites clause before activation."); ctxt2); val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns; val eqns' = (prep_eqns ctxt' o map snd) eqns; val eqnss' = [attrss ~~ eqns']; val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt'; val rewrite_morph = eqns' |> map (abs_def ctxt') |> Variable.export_terms ctxt' ctxt |> Element.eq_term_morphism |> Morphism.default; val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt; val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns']; in (i + 1, insts', eqnss', ctxt'') end; fun prep_elem raw_elem ctxt = let val ctxt' = ctxt |> Context_Position.set_visible false |> declare_elem prep_var_elem raw_elem |> Context_Position.restore_visible ctxt; val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem; in (elems', ctxt') end; val fors = fold_map prep_var_inst fixed ctxt1 |> fst; val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2); fun prep_stmt elems ctxt = check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt; val _ = if fixed_frees then () else (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of [] => () | frees => error ("Illegal free variables in expression: " ^ commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))); val ((insts, _, elems', concl), ctxt4) = ctxt3 |> init_body |> fold_map prep_elem raw_elems |-> prep_stmt; (* parameters from expression and elements *) val xs = maps (fn Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) (Fixes fors :: elems'); val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4; val fors' = finish_fixes parms fors; val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; val deps = map (finish_inst ctxt5) insts; val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems'; in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end; in fun cert_full_context_statement x = prep_full_context_statement (K I) (K I) Obtain.cert_obtains Proof_Context.cert_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun cert_read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src Proof_Context.read_var check_expr x; end; (* Context statement: elements + statement *) local fun prep_statement prep activate raw_elems raw_stmt ctxt = let val ((_, _, _, elems, concl), _) = prep {strict = true, do_close = false, fixed_frees = true} ([], []) I raw_elems raw_stmt ctxt; val ctxt' = ctxt |> Proof_Context.set_stmt true |> fold_map activate elems |> #2 |> Proof_Context.restore_stmt ctxt; in (concl, ctxt') end; in fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x; fun read_statement x = prep_statement read_full_context_statement Element.activate x; end; (* Locale declaration: import + elements *) fun fix_params params = Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd; local fun prep_declaration prep activate raw_import init_body raw_elems ctxt = let val ((fixed, deps, eqnss, elems, _), (parms, ctxt0)) = prep {strict = false, do_close = true, fixed_frees = false} raw_import init_body raw_elems (Element.Shows []) ctxt; val _ = null (flat eqnss) orelse error "Illegal rewrites clause(s) in declaration of locale"; (* Declare parameters and imported facts *) val ctxt' = ctxt |> fix_params fixed |> fold (Context.proof_map o Locale.activate_facts NONE) deps; val (elems', ctxt'') = ctxt' |> Proof_Context.set_stmt true |> fold_map activate elems ||> Proof_Context.restore_stmt ctxt'; in ((fixed, deps, elems', ctxt''), (parms, ctxt0)) end; in fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x; fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x; fun read_declaration x = prep_declaration read_full_context_statement Element.activate x; end; (* Locale expression to set up a goal *) local fun props_of thy (name, morph) = let val (asm, defs) = Locale.specification_of thy name in map (Morphism.term morph) (the_list asm @ defs) end; fun prep_goal_expression prep expression ctxt = let val thy = Proof_Context.theory_of ctxt; val ((fixed, deps, eqnss, _, _), _) = prep {strict = true, do_close = true, fixed_frees = true} expression I [] (Element.Shows []) ctxt; (* proof obligations *) val propss = map (props_of thy) deps; val eq_propss = (map o map) snd eqnss; val goal_ctxt = ctxt |> fix_params fixed |> (fold o fold) Proof_Context.augment (propss @ eq_propss); val export = Proof_Context.export_morphism goal_ctxt ctxt; val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; val exp_term = Term_Subst.zero_var_indexes o Morphism.term export; val exp_typ = Logic.type_map exp_term; val export' = Morphism.morphism "Expression.prep_goal" {binding = [], typ = [K exp_typ], term = [K exp_term], fact = [K exp_fact]}; in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end; in fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x; fun read_goal_expression x = prep_goal_expression read_full_context_statement x; end; (*** Locale declarations ***) (* extract specification text *) val norm_term = Envir.beta_norm oo Term.subst_atomic; fun bind_def ctxt eq (env, eqs) = let val _ = Local_Defs.cert_def ctxt (K []) eq; val ((y, T), b) = Local_Defs.abs_def eq; val b' = norm_term env b; fun err msg = error (msg ^ ": " ^ quote y); in (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of [] => ((Free (y, T), b') :: env, eq :: eqs) | dups => if forall (fn (_, b'') => b' aconv b'') dups then (env, eqs) else err "Attempt to redefine variable") end; (* text has the following structure: (((exts, exts'), (ints, ints')), (xs, env, defs)) where exts: external assumptions (terms in assumes elements) exts': dito, normalised wrt. env ints: internal assumptions (terms in assumptions from insts) ints': dito, normalised wrt. env xs: the free variables in exts' and ints' and rhss of definitions, this includes parameters except defined parameters env: list of term pairs encoding substitutions, where the first term is a free variable; substitutions represent defines elements and the rhs is normalised wrt. the previous env defs: the equations from the defines elements *) fun eval_text _ _ (Fixes _) text = text | eval_text _ _ (Constrains _) text = text | eval_text _ is_ext (Assumes asms) (((exts, exts'), (ints, ints')), (env, defs)) = let val ts = maps (map #1 o #2) asms; val ts' = map (norm_term env) ts; val spec' = if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) else ((exts, exts'), (ints @ ts, ints' @ ts')); in (spec', (env, defs)) end | eval_text ctxt _ (Defines defs) (spec, binds) = (spec, fold (bind_def ctxt o #1 o #2) defs binds) | eval_text _ _ (Notes _) text = text | eval_text _ _ (Lazy_Notes _) text = text; fun eval_inst ctxt (loc, morph) text = let val thy = Proof_Context.theory_of ctxt; val (asm, defs) = Locale.specification_of thy loc; val asm' = Option.map (Morphism.term morph) asm; val defs' = map (Morphism.term morph) defs; val text' = text |> (if is_some asm then eval_text ctxt false (Assumes [(Binding.empty_atts, [(the asm', [])])]) else I) |> (if not (null defs) then eval_text ctxt false (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs')) else I) (* FIXME clone from locale.ML *) in text' end; fun eval_elem ctxt elem text = eval_text ctxt true elem text; fun eval ctxt deps elems = let val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [])); val ((spec, (_, defs))) = fold (eval_elem ctxt) elems text'; in (spec, defs) end; (* axiomsN: name of theorem set with destruct rules for locale predicates, also name suffix of delta predicates and assumptions. *) val axiomsN = "axioms"; local (* introN: name of theorems for introduction rules of locale and delta predicates *) val introN = "intro"; fun atomize_spec ctxt ts = let val t = Logic.mk_conjunction_balanced ts; val body = Object_Logic.atomize_term ctxt t; val bodyT = Term.fastype_of body; in if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of ctxt t)) else (body, bodyT, Object_Logic.atomize ctxt (Thm.cterm_of ctxt t)) end; (* achieve plain syntax for locale predicates (without "PROP") *) fun aprop_tr' n c = let val c' = Lexicon.mark_const c; fun tr' (_: Proof.context) T args = if T <> dummyT andalso length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.const c', args) else raise Match; in (c', tr') end; (* define one predicate including its intro rule and axioms - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) - norm_ts: terms representing locale assumptions (normalised wrt. defs) - thy: the theory *) fun def_pred binding parms defs ts norm_ts thy = let val name = Sign.full_name thy binding; val thy_ctxt = Proof_Context.init_global thy; val (body, bodyT, body_eq) = atomize_spec thy_ctxt norm_ts; val env = Names.build (Names.add_free_names body); val xs = filter (Names.defined env o #1) parms; val Ts = map #2 xs; val type_tfrees = TFrees.build (fold TFrees.add_tfreesT Ts); val extra_tfrees = TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) body) |> TFrees.keys |> map TFree; val predT = map Term.itselfT extra_tfrees ---> Ts ---> bodyT; val args = map Logic.mk_type extra_tfrees @ map Free xs; val head = Term.list_comb (Const (name, predT), args); val statement = Object_Logic.ensure_propT thy_ctxt head; val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.typed_print_translation [aprop_tr' (length args) name] |> Sign.declare_const_global ((binding, predT), NoSyn) |> snd |> Global_Theory.add_defs false [((Thm.def_binding binding, Logic.mk_equals (head, body)), [])]; val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; val intro = Goal.prove_global defs_thy [] norm_ts statement (fn {context = ctxt, ...} => rewrite_goals_tac ctxt [pred_def] THEN compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN compose_tac defs_ctxt (false, Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1); val conjuncts = (Drule.equal_elim_rule2 OF [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_ctxt statement))]) |> Conjunction.elim_balanced (length ts); val (_, axioms_ctxt) = defs_ctxt |> Assumption.add_assumes (maps Thm.chyps_of (defs @ conjuncts)); val axioms = ts ~~ conjuncts |> map (fn (t, ax) => Element.prove_witness axioms_ctxt t (rewrite_goals_tac axioms_ctxt defs THEN compose_tac axioms_ctxt (false, ax, 0) 1)); in ((statement, intro, axioms), defs_thy) end; in (* main predicate definition function *) fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val ctxt = Proof_Context.init_global thy; val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs; val (a_pred, a_intro, a_axioms, thy2) = if null exts then (NONE, NONE, [], thy) else let val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy1) = thy |> def_pred abinding parms defs' exts exts'; val ((_, [intro']), thy2) = thy1 |> Sign.qualified_path true abinding |> Global_Theory.note_thms "" ((Binding.name introN, []), [([intro], [Locale.unfold_add])]) ||> Sign.restore_naming thy1; in (SOME statement, SOME intro', axioms, thy2) end; val (b_pred, b_intro, b_axioms, thy4) = if null ints then (NONE, NONE, [], thy2) else let val ((statement, intro, axioms), thy3) = thy2 |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val conclude_witness = Drule.export_without_context o Element.conclude_witness (Proof_Context.init_global thy3); val ([(_, [intro']), _], thy4) = thy3 |> Sign.qualified_path true binding |> Global_Theory.note_thmss "" [((Binding.name introN, []), [([intro], [Locale.intro_add])]), ((Binding.name axiomsN, []), [(map conclude_witness axioms, [])])] ||> Sign.restore_naming thy3; in (SOME statement, SOME intro', axioms, thy4) end; in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy4) end; end; local fun assumes_to_notes (Assumes asms) axms = fold_map (fn (a, spec) => fn axs => let val (ps, qs) = chop (length spec) axs in ((a, [(ps, [])]), qs) end) asms axms |> apfst (curry Notes "") | assumes_to_notes e axms = (e, axms); fun defines_to_notes ctxt (Defines defs) = Notes ("", map (fn (a, (def, _)) => (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], - [Attrib.internal (K Locale.witness_add)])])) defs) + [Attrib.internal @{here} (K Locale.witness_add)])])) defs) | defines_to_notes _ e = e; val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false; fun gen_add_locale prep_include prep_decl binding raw_predicate_binding raw_includes raw_import raw_body thy = let val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); val ctxt = Proof_Context.init_global thy; val includes = map (prep_include ctxt) raw_includes; val ((fixed, deps, body_elems, _), (parms, ctxt')) = ctxt |> Bundle.includes includes |> prep_decl raw_import I raw_body; val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; val type_tfrees = TFrees.build (fold (TFrees.add_tfreesT o #2) parms); val extra_tfrees = TFrees.build (fold (TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) exts') |> TFrees.keys; val _ = if null extra_tfrees then () else warning ("Additional type variable(s) in locale specification " ^ Binding.print binding ^ ": " ^ commas (map (Syntax.string_of_typ ctxt' o TFree) extra_tfrees)); val predicate_binding = if Binding.is_empty raw_predicate_binding then binding else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = define_preds predicate_binding parms text thy; val pred_ctxt = Proof_Context.init_global thy'; val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; val params = fixed @ maps (fn Fixes fixes => map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems; val asm = if is_some b_statement then b_statement else a_statement; val hyp_spec = filter is_hyp body_elems; val notes = if is_some asm then [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []), [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))], - [Attrib.internal (K Locale.witness_add)])])])] + [Attrib.internal @{here} (K Locale.witness_add)])])])] else []; val notes' = body_elems |> map (Element.transfer_ctxt thy') |> map (defines_to_notes pred_ctxt) |> map (Element.transform_ctxt a_satisfy) |> (fn elems => fold_map assumes_to_notes elems (map (Element.conclude_witness pred_ctxt) a_axioms)) |> fst |> map (Element.transform_ctxt b_satisfy) |> map_filter (fn Notes notes => SOME notes | _ => NONE); val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; val axioms = map (Element.conclude_witness pred_ctxt) b_axioms; val loc_ctxt = thy' |> Locale.register_locale binding (extra_tfrees, params) (asm, rev defs) (a_intro, b_intro) axioms hyp_spec [] (rev notes) (rev deps') |> Named_Target.init includes name |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; in val add_locale = gen_add_locale (K I) cert_declaration; val add_locale_cmd = gen_add_locale Bundle.check read_declaration; end; end; diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,480 +1,481 @@ (* Title: Pure/Isar/generic_target.ML Author: Makarius Author: Florian Haftmann, TU Muenchen Common target infrastructure. *) signature GENERIC_TARGET = sig (*auxiliary*) val export_abbrev: Proof.context -> (term -> term) -> term -> term * ((string * sort) list * (term list * term list)) val check_mixfix: Proof.context -> binding * (string * sort) list -> mixfix -> mixfix val check_mixfix_global: binding * bool -> mixfix -> mixfix (*background primitives*) val background_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val background_declaration: Morphism.declaration_entity -> local_theory -> local_theory val background_abbrev: binding * term -> term list -> local_theory -> (term * term) * local_theory val add_foundation_interpretation: (binding * (term * term list) -> Context.generic -> Context.generic) -> theory -> theory (*nested local theories primitives*) val standard_facts: local_theory -> Proof.context -> Attrib.fact list -> Attrib.fact list val standard_notes: (int * int -> bool) -> string -> Attrib.fact list -> local_theory -> local_theory val standard_declaration: (int * int -> bool) -> Morphism.declaration_entity -> local_theory -> local_theory val standard_const: (int * int -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val local_interpretation: Locale.registration -> local_theory -> local_theory (*lifting target primitives to local theory operations*) val define: (((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory) -> bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val notes: (string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory) -> string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: (Syntax.mode -> binding * mixfix -> term -> term list * term list -> local_theory -> local_theory) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory (*theory target primitives*) val theory_target_foundation: ((binding * typ) * mixfix) * (binding * term) -> term list * term list -> local_theory -> (term * thm) * local_theory val theory_target_notes: string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val theory_target_abbrev: Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory (*theory target operations*) val theory_abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory val theory_declaration: Morphism.declaration_entity -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory (*locale target primitives*) val locale_target_notes: string -> string -> Attrib.fact list -> Attrib.fact list -> local_theory -> local_theory val locale_target_abbrev: string -> Syntax.mode -> (binding * mixfix) -> term -> term list * term list -> local_theory -> local_theory - val locale_target_declaration: string -> bool -> Morphism.declaration_entity -> - local_theory -> local_theory + val locale_target_declaration: string -> {syntax: bool, pos: Position.T} -> + Morphism.declaration_entity -> local_theory -> local_theory val locale_target_const: string -> (morphism -> bool) -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory (*locale operations*) val locale_abbrev: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory - val locale_declaration: string -> {syntax: bool, pervasive: bool} -> Morphism.declaration_entity -> - local_theory -> local_theory + val locale_declaration: string -> {syntax: bool, pervasive: bool, pos: Position.T} -> + Morphism.declaration_entity -> local_theory -> local_theory val locale_const: string -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory val locale_dependency: string -> Locale.registration -> local_theory -> local_theory end structure Generic_Target: GENERIC_TARGET = struct (** consts **) fun export_abbrev lthy preprocess rhs = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); val rhs' = rhs |> Assumption.export_term lthy (Local_Theory.target_of lthy) |> preprocess; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) (Variable.add_fixed lthy rhs' [])); val u = fold_rev lambda term_params rhs'; val global_rhs = singleton (Variable.polymorphic thy_ctxt) u; val type_tfrees = TFrees.build (TFrees.add_tfreesT (Term.fastype_of u)); val extra_tfrees = TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) u) |> TFrees.keys; val type_params = map (Logic.mk_type o TFree) extra_tfrees; in (global_rhs, (extra_tfrees, (type_params, term_params))) end; fun check_mixfix ctxt (b, extra_tfrees) mx = if null extra_tfrees then mx else (if Context_Position.is_visible ctxt then warning ("Additional type variable(s) in specification of " ^ Binding.print b ^ ": " ^ commas (map (Syntax.string_of_typ ctxt o TFree) extra_tfrees) ^ (if Mixfix.is_empty mx then "" else "\nDropping mixfix syntax " ^ Pretty.string_of (Mixfix.pretty_mixfix mx))) else (); NoSyn); fun check_mixfix_global (b, no_params) mx = if no_params orelse Mixfix.is_empty mx then mx else (warning ("Dropping global mixfix syntax: " ^ Binding.print b ^ " " ^ Pretty.string_of (Mixfix.pretty_mixfix mx)); NoSyn); fun same_const (Const (c, _), Const (c', _)) = c = c' | same_const (t $ _, t' $ _) = same_const (t, t') | same_const (_, _) = false; fun const_decl phi_pred prmode ((b, mx), rhs) = Morphism.entity (fn phi => fn context => if phi_pred phi then let val b' = Morphism.binding phi b; val rhs' = Morphism.term phi rhs; val same_shape = Term.aconv_untyped (rhs, rhs'); val same_stem = same_shape orelse same_const (rhs, rhs'); val const_alias = if same_shape then (case rhs' of Const (c, T) => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; in (case Type_Infer_Context.const_type ctxt c of SOME T' => if Sign.typ_equiv thy (T, T') then SOME c else NONE | NONE => NONE) end | _ => NONE) else NONE; in (case const_alias of SOME c => context |> Context.mapping (Sign.const_alias b' c) (Proof_Context.const_alias b' c) |> Morphism.form_entity (Proof_Context.generic_notation true prmode [(rhs', mx)]) | NONE => context |> Proof_Context.generic_add_abbrev Print_Mode.internal (b', Term.close_schematic_term rhs') |-> (fn (const as Const (c, _), _) => same_stem ? (Proof_Context.generic_revert_abbrev (#1 prmode) c #> same_shape ? Morphism.form_entity (Proof_Context.generic_notation true prmode [(const, mx)])))) end else context); (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = ((binding * (term * term list) -> Context.generic -> Context.generic) * stamp) list val empty = []; val merge = Library.merge (eq_snd (op =)); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (cons (f, stamp ())); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = fold (fn (f, _) => f binding_const_params) interps; in lthy |> Local_Theory.background_theory (Context.theory_map interp) |> Local_Theory.map_contexts (K (Context.proof_map interp)) end; fun background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) lthy = let val params = type_params @ term_params; val target_params = type_params @ take_prefix is_Free (Variable.export_terms lthy (Local_Theory.target_of lthy) term_params); val mx' = check_mixfix_global (b, null params) mx; val (const, lthy2) = lthy |> Local_Theory.background_theory_result (Sign.declare_const lthy ((b, U), mx')); val lhs = Term.list_comb (const, params); val ((_, def), lthy3) = lthy2 |> Local_Theory.background_theory_result (Thm.add_def (Proof_Context.defs_context lthy2) false false (Thm.def_binding_optional b b_def, Logic.mk_equals (lhs, rhs))) ||> foundation_interpretation (b, (const, target_params)); in ((lhs, def), lthy3) end; fun background_declaration decl lthy = let fun theory_decl context = Local_Theory.standard_form lthy (Proof_Context.init_global (Context.theory_of context)) decl context; in Local_Theory.background_theory (Context.theory_map theory_decl) lthy end; fun background_abbrev (b, global_rhs) params = Local_Theory.background_theory_result (Sign.add_abbrev Print_Mode.internal (b, global_rhs)) #>> apply2 (fn t => Term.list_comb (Logic.unvarify_global t, params)) (** nested local theories primitives **) fun standard_facts lthy ctxt = Attrib.transform_facts (Local_Theory.standard_morphism lthy ctxt); fun standard_notes pred kind facts lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Attrib.local_notes kind (standard_facts lthy ctxt facts) ctxt |> snd else ctxt) lthy; fun standard_declaration pred decl lthy = Local_Theory.map_contexts (fn level => fn ctxt => if pred (Local_Theory.level lthy, level) then Context.proof_map (Local_Theory.standard_form lthy ctxt decl) ctxt else ctxt) lthy; fun standard_const pred prmode ((b, mx), rhs) = standard_declaration pred (const_decl (K true) prmode ((b, mx), rhs)); fun standard_registration pred registration lthy = Local_Theory.map_contexts (fn level => if pred (Local_Theory.level lthy, level) then Context.proof_map (Locale.add_registration registration) else I) lthy; val local_interpretation = standard_registration (fn (n, level) => level >= n - 1); (** lifting target primitives to local theory operations **) (* define *) fun define foundation internal ((b, mx), ((b_def, atts), rhs)) lthy = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy); (*term and type parameters*) val ((defs, _), rhs') = Thm.cterm_of lthy rhs |> Local_Defs.export_cterm lthy thy_ctxt ||> Thm.term_of; val xs = Variable.add_fixed lthy rhs' []; val T = Term.fastype_of rhs; val type_tfrees = TFrees.build (TFrees.add_tfreesT T #> fold (TFrees.add_tfreesT o #2) xs); val extra_tfrees = TFrees.build (rhs |> TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) |> TFrees.keys; val mx' = check_mixfix lthy (b, extra_tfrees) mx; val type_params = map (Logic.mk_type o TFree) extra_tfrees; val term_params = map Free (sort (Variable.fixed_ord lthy o apply2 #1) xs); val params = type_params @ term_params; val U = map Term.fastype_of params ---> T; (*foundation*) val ((lhs', global_def), lthy2) = lthy |> foundation (((b, U), mx'), (b_def, rhs')) (type_params, term_params); (*local definition*) val ([(lhs, (_, local_def))], lthy3) = lthy2 |> Context_Position.set_visible false |> Local_Defs.define [((b, NoSyn), (Binding.empty_atts, lhs'))] ||> Context_Position.restore_visible lthy2; (*result*) val def = Thm.transitive local_def global_def |> Local_Defs.contract lthy3 defs (Thm.cterm_of lthy3 (Logic.mk_equals (lhs, rhs))); val ([(res_name, [res])], lthy4) = lthy3 |> Local_Theory.notes [((if internal then Binding.empty else b_def, atts), [([def], [])])]; in ((lhs, (res_name, res)), lthy4) end; (* notes *) local fun import_export_proof ctxt (name, raw_th) = let val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of ctxt); (*export assumes/defines*) val th = Goal.norm_result ctxt raw_th; val ((defs, asms), th') = Local_Defs.export ctxt thy_ctxt th; val asms' = map (rewrite_rule ctxt (Drule.norm_hhf_eqs @ defs)) asms; (*export fixes*) val tfrees = TFrees.build (Thm.fold_terms {hyps = true} TFrees.add_tfrees th') |> TFrees.keys |> map TFree; val frees = Frees.build (Thm.fold_terms {hyps = true} Frees.add_frees th') |> Frees.list_set_rev |> map Free; val (th'' :: vs) = (th' :: map (Drule.mk_term o Thm.cterm_of ctxt) (map Logic.mk_type tfrees @ frees)) |> Variable.export ctxt thy_ctxt |> Drule.zero_var_indexes_list; (*thm definition*) val result = Global_Theory.name_thm Global_Theory.official1 name th''; (*import fixes*) val (tvars, vars) = chop (length tfrees) (map (Thm.term_of o Drule.dest_term) vs) |>> map Logic.dest_type; val instT = TVars.build (fold2 (fn a => fn b => (case a of TVar v => TVars.add (v, b) | _ => I)) tvars tfrees); val cinstT = TVars.map (K (Thm.ctyp_of ctxt)) instT; val cinst = Vars.build (fold2 (fn v => fn t => (case v of Var (xi, T) => Vars.add ((xi, Term_Subst.instantiateT instT T), Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) | _ => I)) vars frees); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; -fun locale_target_declaration locale syntax decl lthy = lthy +fun locale_target_declaration locale params decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> - Locale.add_declaration locale syntax + Locale.add_declaration locale params (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = - locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) + locale_target_declaration locale {syntax = true, pos = Binding.pos_of b} + (const_decl phi_pred prmode ((b, mx), rhs)); (** locale operations **) -fun locale_declaration locale {syntax, pervasive} decl = +fun locale_declaration locale {syntax, pervasive, pos} decl = pervasive ? background_declaration decl - #> locale_target_declaration locale syntax decl + #> locale_target_declaration locale {syntax = syntax, pos = pos} decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/Isar/isar_cmd.ML b/src/Pure/Isar/isar_cmd.ML --- a/src/Pure/Isar/isar_cmd.ML +++ b/src/Pure/Isar/isar_cmd.ML @@ -1,301 +1,302 @@ (* Title: Pure/Isar/isar_cmd.ML Author: Markus Wenzel, TU Muenchen Miscellaneous Isar commands. *) signature ISAR_CMD = sig val setup: Input.source -> theory -> theory val local_setup: Input.source -> Proof.context -> Proof.context val parse_ast_translation: Input.source -> theory -> theory val parse_translation: Input.source -> theory -> theory val print_translation: Input.source -> theory -> theory val typed_print_translation: Input.source -> theory -> theory val print_ast_translation: Input.source -> theory -> theory val translations: (xstring * string) Syntax.trrule list -> theory -> theory val no_translations: (xstring * string) Syntax.trrule list -> theory -> theory val oracle: bstring * Position.range -> Input.source -> theory -> theory val declaration: {syntax: bool, pervasive: bool} -> Input.source -> local_theory -> local_theory val simproc_setup: string * Position.T -> string list -> Input.source -> local_theory -> local_theory val qed: Method.text_range option -> Toplevel.transition -> Toplevel.transition val terminal_proof: Method.text_range * Method.text_range option -> Toplevel.transition -> Toplevel.transition val default_proof: Toplevel.transition -> Toplevel.transition val immediate_proof: Toplevel.transition -> Toplevel.transition val done_proof: Toplevel.transition -> Toplevel.transition val skip_proof: Toplevel.transition -> Toplevel.transition val ml_diag: bool -> Input.source -> Toplevel.transition -> Toplevel.transition val diag_state: Proof.context -> Toplevel.state val diag_goal: Proof.context -> {context: Proof.context, facts: thm list, goal: thm} val pretty_theorems: bool -> Toplevel.state -> Pretty.T list val print_stmts: string list * (Facts.ref * Token.src list) list -> Toplevel.transition -> Toplevel.transition val print_thms: string list * (Facts.ref * Token.src list) list -> Toplevel.transition -> Toplevel.transition val print_prfs: bool -> string list * (Facts.ref * Token.src list) list option -> Toplevel.transition -> Toplevel.transition val print_prop: (string list * string) -> Toplevel.transition -> Toplevel.transition val print_term: (string list * string) -> Toplevel.transition -> Toplevel.transition val print_type: (string list * (string * string option)) -> Toplevel.transition -> Toplevel.transition end; structure Isar_Cmd: ISAR_CMD = struct (** theory declarations **) (* generic setup *) fun setup source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (" @ ML_Lex.read_source source @ ML_Lex.read ")") |> Context.theory_map; fun local_setup source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.local_setup (" @ ML_Lex.read_source source @ ML_Lex.read ")") |> Context.proof_map; (* translation functions *) fun parse_ast_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.parse_ast_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun parse_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.parse_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun print_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.print_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun typed_print_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.typed_print_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; fun print_ast_translation source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "Theory.setup (Sign.print_ast_translation (" @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.theory_map; (* translation rules *) fun read_trrules thy raw_rules = let val ctxt = Proof_Context.init_global thy; val read_root = #1 o dest_Type o Proof_Context.read_type_name {proper = true, strict = false} ctxt; in raw_rules |> map (Syntax.map_trrule (fn (r, s) => Syntax_Phases.parse_ast_pattern ctxt (read_root r, s))) end; fun translations args thy = Sign.add_trrules (read_trrules thy args) thy; fun no_translations args thy = Sign.del_trrules (read_trrules thy args) thy; (* oracles *) fun oracle (name, range) source = ML_Context.expression (Input.pos_of source) (ML_Lex.read "val " @ ML_Lex.read_range range name @ ML_Lex.read (" = snd (Context.>>> (Context.map_theory_result (Thm.add_oracle (" ^ ML_Syntax.make_binding (name, #1 range) ^ ", ") @ ML_Lex.read_source source @ ML_Lex.read "))))") |> Context.theory_map; (* declarations *) fun declaration {syntax, pervasive} source = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Local_Theory.declaration {syntax = " ^ - Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ "} (") @ + Bool.toString syntax ^ ", pervasive = " ^ Bool.toString pervasive ^ + ", pos = " ^ ML_Syntax.print_position (Position.thread_data ()) ^ "} (") @ ML_Lex.read_source source @ ML_Lex.read "))") |> Context.proof_map; (* simprocs *) fun simproc_setup name lhss source = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Simplifier.define_simproc_cmd (" ^ ML_Syntax.make_binding name ^ ") {lhss = " ^ ML_Syntax.print_strings lhss ^ ", proc = (") @ ML_Lex.read_source source @ ML_Lex.read ")})") |> Context.proof_map; (* local endings *) fun local_qed m = Toplevel.proof (Proof.local_qed (m, true)); val local_terminal_proof = Toplevel.proof o Proof.local_future_terminal_proof; val local_default_proof = Toplevel.proof Proof.local_default_proof; val local_immediate_proof = Toplevel.proof Proof.local_immediate_proof; val local_done_proof = Toplevel.proof Proof.local_done_proof; val local_skip_proof = Toplevel.proof' Proof.local_skip_proof; (* global endings *) fun global_qed m = Toplevel.end_proof (K (Proof.global_qed (m, true))); val global_terminal_proof = Toplevel.end_proof o K o Proof.global_future_terminal_proof; val global_default_proof = Toplevel.end_proof (K Proof.global_default_proof); val global_immediate_proof = Toplevel.end_proof (K Proof.global_immediate_proof); val global_skip_proof = Toplevel.end_proof Proof.global_skip_proof; val global_done_proof = Toplevel.end_proof (K Proof.global_done_proof); (* common endings *) fun qed m = local_qed m o global_qed m; fun terminal_proof m = local_terminal_proof m o global_terminal_proof m; val default_proof = local_default_proof o global_default_proof; val immediate_proof = local_immediate_proof o global_immediate_proof; val done_proof = local_done_proof o global_done_proof; val skip_proof = local_skip_proof o global_skip_proof; (* diagnostic ML evaluation *) structure Diag_State = Proof_Data ( type T = Toplevel.state option; fun init _ = NONE; ); fun ml_diag verbose source = Toplevel.keep (fn state => let val opt_ctxt = try Toplevel.generic_theory_of state |> Option.map (Context.proof_of #> Diag_State.put (SOME state)); val flags = ML_Compiler.verbose verbose ML_Compiler.flags; in ML_Context.eval_source_in opt_ctxt flags source end); fun diag_state ctxt = (case Diag_State.get ctxt of SOME st => st | NONE => Toplevel.make_state NONE); val diag_goal = Proof.goal o Toplevel.proof_of o diag_state; val _ = Theory.setup (ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\state\) (Scan.succeed "Isar_Cmd.diag_state ML_context") #> ML_Antiquotation.value (Binding.qualify true "Isar" \<^binding>\goal\) (Scan.succeed "Isar_Cmd.diag_goal ML_context")); (* theorems of theory or proof context *) fun pretty_theorems verbose st = if Toplevel.is_proof st then Proof_Context.pretty_local_facts verbose (Toplevel.context_of st) else let val ctxt = Toplevel.context_of st; val prev_thys = (case Toplevel.previous_theory_of st of SOME thy => [thy] | NONE => Theory.parents_of (Proof_Context.theory_of ctxt)); in Proof_Display.pretty_theorems_diff verbose prev_thys ctxt end; (* print theorems, terms, types etc. *) local fun string_of_stmts ctxt args = Attrib.eval_thms ctxt args |> map (Element.pretty_statement ctxt Thm.theoremK) |> Pretty.chunks2 |> Pretty.string_of; fun string_of_thms ctxt args = Pretty.string_of (Proof_Context.pretty_fact ctxt ("", Attrib.eval_thms ctxt args)); fun string_of_prfs full state arg = Pretty.string_of (case arg of NONE => let val {context = ctxt, goal = thm} = Proof.simple_goal (Toplevel.proof_of state); val thy = Proof_Context.theory_of ctxt; val prf = Thm.proof_of thm; val prop = Thm.full_prop_of thm; val prf' = Proofterm.rewrite_proof_notypes ([], []) prf; in Proof_Syntax.pretty_proof ctxt (if full then Proofterm.reconstruct_proof thy prop prf' else prf') end | SOME srcs => let val ctxt = Toplevel.context_of state; val pretty_proof = Proof_Syntax.pretty_standard_proof_of ctxt full; in Pretty.chunks (map pretty_proof (Attrib.eval_thms ctxt srcs)) end); fun string_of_prop ctxt s = let val prop = Syntax.read_prop ctxt s; val ctxt' = Proof_Context.augment prop ctxt; in Pretty.string_of (Pretty.quote (Syntax.pretty_term ctxt' prop)) end; fun string_of_term ctxt s = let val t = Syntax.read_term ctxt s; val T = Term.type_of t; val ctxt' = Proof_Context.augment t ctxt; in Pretty.string_of (Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t), Pretty.fbrk, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' T)]) end; fun string_of_type ctxt (s, NONE) = let val T = Syntax.read_typ ctxt s in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt T)) end | string_of_type ctxt (s1, SOME s2) = let val ctxt' = Config.put show_sorts true ctxt; val raw_T = Syntax.parse_typ ctxt' s1; val S = Syntax.read_sort ctxt' s2; val T = Syntax.check_term ctxt' (Logic.mk_type raw_T |> Type.constraint (Term.itselfT (Type_Infer.anyT S))) |> Logic.dest_type; in Pretty.string_of (Pretty.quote (Syntax.pretty_typ ctxt' T)) end; fun print_item string_of (modes, arg) = Toplevel.keep (fn state => Print_Mode.with_modes modes (fn () => writeln (string_of state arg)) ()); in val print_stmts = print_item (string_of_stmts o Toplevel.context_of); val print_thms = print_item (string_of_thms o Toplevel.context_of); val print_prfs = print_item o string_of_prfs; val print_prop = print_item (string_of_prop o Toplevel.context_of); val print_term = print_item (string_of_term o Toplevel.context_of); val print_type = print_item (string_of_type o Toplevel.context_of); end; end; diff --git a/src/Pure/Isar/local_theory.ML b/src/Pure/Isar/local_theory.ML --- a/src/Pure/Isar/local_theory.ML +++ b/src/Pure/Isar/local_theory.ML @@ -1,440 +1,440 @@ (* Title: Pure/Isar/local_theory.ML Author: Makarius Local theory operations, with abstract target context. *) type local_theory = Proof.context; type generic_theory = Context.generic; structure Attrib = struct type binding = binding * Token.src list; type thms = (thm list * Token.src list) list; type fact = binding * thms; end; structure Locale = struct type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism}; end; signature LOCAL_THEORY = sig type operations val assert: local_theory -> local_theory val level: Proof.context -> int val map_contexts: (int -> Proof.context -> Proof.context) -> local_theory -> local_theory val background_naming_of: local_theory -> Name_Space.naming val map_background_naming: (Name_Space.naming -> Name_Space.naming) -> local_theory -> local_theory val restore_background_naming: local_theory -> local_theory -> local_theory val full_name: local_theory -> binding -> string val new_group: local_theory -> local_theory val reset_group: local_theory -> local_theory val standard_morphism: local_theory -> Proof.context -> morphism val standard_morphism_theory: local_theory -> morphism val standard_form: local_theory -> Proof.context -> 'a Morphism.entity -> 'a val raw_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val raw_theory: (theory -> theory) -> local_theory -> local_theory val background_theory_result: (theory -> 'a * theory) -> local_theory -> 'a * local_theory val background_theory: (theory -> theory) -> local_theory -> local_theory val target_of: local_theory -> Proof.context val target: (Proof.context -> Proof.context) -> local_theory -> local_theory val target_morphism: local_theory -> morphism val propagate_ml_env: generic_theory -> generic_theory val touch_ml_env: generic_theory -> generic_theory val operations_of: local_theory -> operations val define: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val define_internal: (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory val note: Attrib.binding * thm list -> local_theory -> (string * thm list) * local_theory val notes: Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val notes_kind: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory val abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory - val declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration -> + val declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration -> local_theory -> local_theory val theory_registration: Locale.registration -> local_theory -> local_theory val locale_dependency: Locale.registration -> local_theory -> local_theory val pretty: local_theory -> Pretty.T list val add_thms_dynamic: binding * (Context.generic -> thm list) -> local_theory -> local_theory val set_defsort: sort -> local_theory -> local_theory val syntax: bool -> Syntax.mode -> (string * typ * mixfix) list -> local_theory -> local_theory val syntax_cmd: bool -> Syntax.mode -> (string * string * mixfix) list -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val type_alias: binding -> string -> local_theory -> local_theory val const_alias: binding -> string -> local_theory -> local_theory val init: {background_naming: Name_Space.naming, setup: theory -> Proof.context, conclude: local_theory -> Proof.context} -> operations -> theory -> local_theory val exit: local_theory -> Proof.context val exit_global: local_theory -> theory val exit_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * Proof.context val exit_result_global: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * theory val begin_nested: local_theory -> Binding.scope * local_theory val end_nested: local_theory -> local_theory val end_nested_result: (morphism -> 'a -> 'b) -> 'a * local_theory -> 'b * local_theory end; signature PRIVATE_LOCAL_THEORY = sig include LOCAL_THEORY val reset: local_theory -> local_theory end structure Local_Theory: PRIVATE_LOCAL_THEORY = struct (** local theory data **) (* type lthy *) type operations = {define: bool -> (binding * mixfix) * (Attrib.binding * term) -> local_theory -> (term * (string * thm)) * local_theory, notes: string -> Attrib.fact list -> local_theory -> (string * thm list) list * local_theory, abbrev: Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * local_theory, - declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration_entity -> + declaration: {syntax: bool, pervasive: bool, pos: Position.T} -> Morphism.declaration_entity -> local_theory -> local_theory, theory_registration: Locale.registration -> local_theory -> local_theory, locale_dependency: Locale.registration -> local_theory -> local_theory, pretty: local_theory -> Pretty.T list}; type lthy = {background_naming: Name_Space.naming, operations: operations, conclude: Proof.context -> Proof.context, target: Proof.context}; fun make_lthy (background_naming, operations, conclude, target) : lthy = {background_naming = background_naming, operations = operations, conclude = conclude, target = target}; (* context data *) structure Data = Proof_Data ( type T = lthy list; fun init _ = []; ); (* nested structure *) val level = length o Data.get; (*1: main target at bottom, >= 2: nested target context*) fun assert lthy = if level lthy = 0 then error "Missing local theory context" else lthy; fun assert_bottom lthy = let val _ = assert lthy; in if level lthy > 1 then error "Not at bottom of local theory nesting" else lthy end; fun assert_not_bottom lthy = let val _ = assert lthy; in if level lthy = 1 then error "Already at bottom of local theory nesting" else lthy end; val bottom_of = List.last o Data.get o assert; val top_of = hd o Data.get o assert; fun map_top f = assert #> Data.map (fn {background_naming, operations, conclude, target} :: parents => make_lthy (f (background_naming, operations, conclude, target)) :: parents); fun reset lthy = #target (top_of lthy) |> Data.put (Data.get lthy); fun map_contexts f lthy = let val n = level lthy in lthy |> (Data.map o map_index) (fn (i, {background_naming, operations, conclude, target}) => make_lthy (background_naming, operations, conclude, target |> Context_Position.set_visible false |> f (n - i - 1) |> Context_Position.restore_visible target)) |> f n end; (* naming for background theory *) val background_naming_of = #background_naming o top_of; fun map_background_naming f = map_top (fn (background_naming, operations, conclude, target) => (f background_naming, operations, conclude, target)); val restore_background_naming = map_background_naming o K o background_naming_of; val full_name = Name_Space.full_name o background_naming_of; val new_group = map_background_naming Name_Space.new_group; val reset_group = map_background_naming Name_Space.reset_group; (* standard morphisms *) fun standard_morphism lthy ctxt = Morphism.set_context' lthy (Proof_Context.export_morphism lthy ctxt $> Morphism.thm_morphism' "Local_Theory.standard" (Goal.norm_result o Proof_Context.init_global) $> Morphism.binding_morphism "Local_Theory.standard_binding" (Name_Space.transform_binding (Proof_Context.naming_of lthy))); fun standard_morphism_theory lthy = standard_morphism lthy (Proof_Context.init_global (Proof_Context.theory_of lthy)); fun standard_form lthy ctxt x = Morphism.form (Morphism.transform (standard_morphism lthy ctxt) x); (* background theory *) fun raw_theory_result f lthy = let val (res, thy') = f (Proof_Context.theory_of lthy); val lthy' = map_contexts (K (Proof_Context.transfer thy')) lthy; in (res, lthy') end; fun raw_theory f = #2 o raw_theory_result (f #> pair ()); fun background_theory_result f lthy = let val naming = background_naming_of lthy |> Name_Space.transform_naming (Proof_Context.naming_of lthy); in lthy |> raw_theory_result (fn thy => thy |> Sign.map_naming (K naming) |> f ||> Sign.restore_naming thy) end; fun background_theory f = #2 o background_theory_result (f #> pair ()); (* target contexts *) val target_of = #target o bottom_of; fun target f lthy = let val ctxt = target_of lthy; val ctxt' = ctxt |> Context_Position.set_visible false |> f |> Context_Position.restore_visible ctxt; val thy' = Proof_Context.theory_of ctxt'; in map_contexts (fn 0 => K ctxt' | _ => Proof_Context.transfer thy') lthy end; fun target_morphism lthy = standard_morphism lthy (target_of lthy); fun propagate_ml_env (context as Context.Proof lthy) = let val inherit = ML_Env.inherit [context] in lthy |> background_theory (Context.theory_map inherit) |> map_contexts (K (Context.proof_map inherit)) |> Context.Proof end | propagate_ml_env context = context; fun touch_ml_env context = if Context.enabled_tracing () then (case context of Context.Theory _ => ML_Env.touch context | Context.Proof _ => context) else context; (** operations **) val operations_of = #operations o top_of; fun operation f lthy = f (operations_of lthy) lthy; fun operation1 f x = operation (fn ops => f ops x); fun operation2 f x y = operation (fn ops => f ops x y); (* primitives *) val pretty = operation #pretty; val abbrev = operation2 #abbrev; val define = operation2 #define false; val define_internal = operation2 #define true; val notes_kind = operation2 #notes; fun declaration args = operation2 #declaration args o Morphism.entity; val theory_registration = operation1 #theory_registration; fun locale_dependency registration = assert_bottom #> operation1 #locale_dependency registration; (* theorems *) val notes = notes_kind ""; fun note (a, ths) = notes [(a, [(ths, [])])] #>> the_single; fun add_thms_dynamic (binding, f) lthy = lthy |> background_theory_result (fn thy => thy |> Global_Theory.add_thms_dynamic' (Sign.inherit_naming thy lthy) (binding, f)) |-> (fn name => map_contexts (fn _ => fn ctxt => Proof_Context.transfer_facts (Proof_Context.theory_of ctxt) ctxt) #> - declaration {syntax = false, pervasive = false} (fn phi => + declaration {syntax = false, pervasive = false, pos = Binding.pos_of binding} (fn phi => let val binding' = Morphism.binding phi binding in Context.mapping (Global_Theory.alias_fact binding' name) (Proof_Context.alias_fact binding' name) end)); (* default sort *) fun set_defsort S = - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (K (Context.mapping (Sign.set_defsort S) (Proof_Context.set_defsort S))); (* syntax *) fun gen_syntax prep_type add mode raw_args lthy = let val args = map (fn (c, T, mx) => (c, prep_type lthy T, mx)) raw_args; val args' = map (fn (c, T, mx) => (c, T, Mixfix.reset_pos mx)) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.syntax add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn _ => Proof_Context.generic_syntax add mode args') lthy end; val syntax = gen_syntax (K I); val syntax_cmd = gen_syntax Proof_Context.read_typ_syntax; (* notation *) local fun gen_type_notation prep_type add mode raw_args lthy = let val prepare = prep_type lthy #> Logic.type_map (Assumption.export_term lthy (target_of lthy)); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.type_notation add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (Proof_Context.generic_type_notation add mode args') lthy end; fun gen_notation prep_const add mode raw_args lthy = let val prepare = prep_const lthy #> Assumption.export_term lthy (target_of lthy); val args = map (apfst prepare) raw_args; val args' = map (apsnd Mixfix.reset_pos) args; val _ = lthy |> Context_Position.is_visible lthy ? Proof_Context.notation add mode args; in - declaration {syntax = true, pervasive = false} + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (Proof_Context.generic_notation add mode args') lthy end; in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* name space aliases *) fun syntax_alias global_alias local_alias b name = - declaration {syntax = true, pervasive = false} (fn phi => + declaration {syntax = true, pervasive = false, pos = Position.thread_data ()} (fn phi => let val b' = Morphism.binding phi b in Context.mapping (global_alias b' name) (local_alias b' name) end); val type_alias = syntax_alias Sign.type_alias Proof_Context.type_alias; val const_alias = syntax_alias Sign.const_alias Proof_Context.const_alias; (** manage targets **) (* main target *) fun init_target background_naming conclude operations target = Data.map (K [make_lthy (background_naming, operations, conclude, target)]) target fun init {background_naming, setup, conclude} operations thy = thy |> Sign.change_begin |> setup |> init_target background_naming (conclude #> target_of #> Sign.change_end_local) operations; val exit_of = #conclude o bottom_of; fun exit lthy = exit_of lthy (assert_bottom lthy); val exit_global = Proof_Context.theory_of o exit; fun exit_result decl (x, lthy) = let val ctxt = exit lthy; val phi = standard_morphism lthy ctxt; in (decl phi x, ctxt) end; fun exit_result_global decl (x, lthy) = let val thy = exit_global lthy; val thy_ctxt = Proof_Context.init_global thy; val phi = standard_morphism lthy thy_ctxt; in (decl phi x, thy) end; (* nested targets *) fun begin_nested lthy = let val _ = assert lthy; val (scope, target) = Proof_Context.new_scope lthy; val entry = make_lthy (background_naming_of lthy, operations_of lthy, Proof_Context.restore_naming lthy, target); val lthy' = Data.map (cons entry) target; in (scope, lthy') end; fun end_nested lthy = let val _ = assert_not_bottom lthy; val ({conclude, ...} :: rest) = Data.get lthy; in lthy |> Data.put rest |> reset |> conclude end; fun end_nested_result decl (x, lthy) = let val outer_lthy = end_nested lthy; val phi = Proof_Context.export_morphism lthy outer_lthy; in (decl phi x, outer_lthy) end; end; diff --git a/src/Pure/Isar/locale.ML b/src/Pure/Isar/locale.ML --- a/src/Pure/Isar/locale.ML +++ b/src/Pure/Isar/locale.ML @@ -1,803 +1,804 @@ (* Title: Pure/Isar/locale.ML Author: Clemens Ballarin, TU Muenchen Locales -- managed Isar proof contexts, based on Pure predicates. Draws basic ideas from Florian Kammueller's original version of locales, but uses the richer infrastructure of Isar instead of the raw meta-logic. Furthermore, structured composition of contexts (with merge and instantiation) is provided, as well as type-inference of the signature parts and predicate definitions of the specification text. Interpretation enables the transfer of declarations and theorems to other contexts, namely those defined by theories, structured proofs and locales themselves. A comprehensive account of locales is available: [1] Clemens Ballarin. Locales: a module system for mathematical theories. Journal of Automated Reasoning, 52(2):123-153, 2014. See also: [2] Clemens Ballarin. Locales and Locale Expressions in Isabelle/Isar. In Stefano Berardi et al., Types for Proofs and Programs: International Workshop, TYPES 2003, Torino, Italy, LNCS 3085, pages 34-50, 2004. [3] Clemens Ballarin. Interpretation of Locales in Isabelle: Managing Dependencies between Locales. Technical Report TUM-I0607, Technische Universitaet Muenchen, 2006. [4] Clemens Ballarin. Interpretation of Locales in Isabelle: Theories and Proof Contexts. In J.M. Borwein and W.M. Farmer, MKM 2006, LNAI 4108, pages 31-43, 2006. *) signature LOCALE = sig (* Locale specification *) val register_locale: binding -> (string * sort) list * ((string * typ) * mixfix) list -> term option * term list -> thm option * thm option -> thm list -> Element.context_i list -> Morphism.declaration_entity list -> (string * Attrib.fact list) list -> (string * morphism) list -> theory -> theory val locale_space: theory -> Name_Space.T val intern: theory -> xstring -> string val check: theory -> xstring * Position.T -> string val extern: theory -> string -> xstring val markup_name: Proof.context -> string -> string val pretty_name: Proof.context -> string -> Pretty.T val defined: theory -> string -> bool val parameters_of: theory -> string -> (string * sort) list * ((string * typ) * mixfix) list val params_of: theory -> string -> ((string * typ) * mixfix) list val intros_of: theory -> string -> thm option * thm option val axioms_of: theory -> string -> thm list val instance_of: theory -> string -> morphism -> term list val specification_of: theory -> string -> term option * term list val hyp_spec_of: theory -> string -> Element.context_i list (* Storing results *) val add_facts: string -> string -> Attrib.fact list -> Proof.context -> Proof.context - val add_declaration: string -> bool -> Morphism.declaration_entity -> Proof.context -> Proof.context + val add_declaration: string -> {syntax: bool, pos: Position.T} -> + Morphism.declaration_entity -> Proof.context -> Proof.context (* Activation *) val activate_facts: morphism option -> string * morphism -> Context.generic -> Context.generic val activate_declarations: string * morphism -> Proof.context -> Proof.context val init: string -> theory -> Proof.context (* Reasoning about locales *) val get_witnesses: Proof.context -> thm list val get_intros: Proof.context -> thm list val get_unfolds: Proof.context -> thm list val witness_add: attribute val intro_add: attribute val unfold_add: attribute val intro_locales_tac: {strict: bool, eager: bool} -> Proof.context -> thm list -> tactic (* Registrations and dependencies *) type registration = {inst: string * morphism, mixin: (morphism * bool) option, export: morphism} val amend_registration: registration -> Context.generic -> Context.generic val add_registration: registration -> Context.generic -> Context.generic val registrations_of: Context.generic -> string -> (string * morphism) list val add_dependency: string -> registration -> theory -> theory (* Diagnostic *) val get_locales: theory -> string list val locale_notes: theory -> string -> (string * Attrib.fact list) list val pretty_locales: theory -> bool -> Pretty.T val pretty_locale: theory -> bool -> string -> Pretty.T val pretty_registrations: Proof.context -> string -> Pretty.T val pretty_locale_deps: theory -> {name: string, parents: string list, body: Pretty.T} list type locale_dependency = {source: string, target: string, prefix: (string * bool) list, morphism: morphism, pos: Position.T, serial: serial} val dest_dependencies: theory list -> theory -> locale_dependency list val tracing : Proof.context -> string -> unit end; structure Locale: LOCALE = struct datatype ctxt = datatype Element.ctxt; (*** Locales ***) type dep = {name: string, morphisms: morphism * morphism, pos: Position.T, serial: serial}; fun eq_dep (dep1: dep, dep2: dep) = #serial dep1 = #serial dep2; fun transfer_dep thy ({name, morphisms, pos, serial}: dep) : dep = {name = name, morphisms = apply2 (Morphism.set_context thy) morphisms, pos = pos, serial = serial}; fun make_dep (name, morphisms) : dep = {name = name, morphisms = apply2 Morphism.reset_context morphisms, pos = Position.thread_data (), serial = serial ()}; (*table of mixin lists, per list mixins in reverse order of declaration; lists indexed by registration/dependency serial, entries for empty lists may be omitted*) type mixin = (morphism * bool) * serial; type mixins = mixin list Inttab.table; fun lookup_mixins (mixins: mixins) serial' = Inttab.lookup_list mixins serial'; val merge_mixins: mixins * mixins -> mixins = Inttab.merge_list (eq_snd op =); fun insert_mixin serial' (morph, b) : mixins -> mixins = Inttab.cons_list (serial', ((Morphism.reset_context morph, b), serial ())); fun rename_mixin (old, new) (mixins: mixins) = (case Inttab.lookup mixins old of NONE => mixins | SOME mixin => Inttab.delete old mixins |> Inttab.update_new (new, mixin)); fun compose_mixins (mixins: mixin list) = fold_rev Morphism.compose (map (fst o fst) mixins) Morphism.identity; datatype locale = Loc of { (* static part *) (*type and term parameters*) parameters: (string * sort) list * ((string * typ) * mixfix) list, (*assumptions (as a single predicate expression) and defines*) spec: term option * term list, intros: thm option * thm option, axioms: thm list, (*diagnostic device: theorem part of hypothetical body as specified by the user*) hyp_spec: Element.context_i list, (* dynamic part *) (*syntax declarations*) syntax_decls: (Morphism.declaration_entity * serial) list, (*theorem declarations*) notes: ((string * Attrib.fact list) * serial) list, (*locale dependencies (sublocale relation) in reverse order*) dependencies: dep list, (*mixin part of dependencies*) mixins: mixins }; fun mk_locale ((parameters, spec, intros, axioms, hyp_spec), ((syntax_decls, notes), (dependencies, mixins))) = Loc {parameters = parameters, spec = spec, intros = intros, axioms = axioms, hyp_spec = hyp_spec, syntax_decls = syntax_decls, notes = notes, dependencies = dependencies, mixins = mixins}; fun map_locale f (Loc {parameters, spec, intros, axioms, hyp_spec, syntax_decls, notes, dependencies, mixins}) = mk_locale (f ((parameters, spec, intros, axioms, hyp_spec), ((syntax_decls, notes), (dependencies, mixins)))); fun merge_locale (Loc {parameters, spec, intros, axioms, hyp_spec, syntax_decls, notes, dependencies, mixins}, Loc {syntax_decls = syntax_decls', notes = notes', dependencies = dependencies', mixins = mixins', ...}) = mk_locale ((parameters, spec, intros, axioms, hyp_spec), ((merge (eq_snd op =) (syntax_decls, syntax_decls'), merge (eq_snd op =) (notes, notes')), (merge eq_dep (dependencies, dependencies'), (merge_mixins (mixins, mixins'))))); structure Locales = Theory_Data ( type T = locale Name_Space.table; val empty : T = Name_Space.empty_table Markup.localeN; val merge = Name_Space.join_tables (K merge_locale); ); val locale_space = Name_Space.space_of_table o Locales.get; val intern = Name_Space.intern o locale_space; fun check thy = #1 o Name_Space.check (Context.Theory thy) (Locales.get thy); val _ = Theory.setup (ML_Antiquotation.inline_embedded \<^binding>\locale\ (Args.theory -- Scan.lift Parse.embedded_position >> (ML_Syntax.print_string o uncurry check))); fun extern thy = Name_Space.extern (Proof_Context.init_global thy) (locale_space thy); fun markup_extern ctxt = Name_Space.markup_extern ctxt (locale_space (Proof_Context.theory_of ctxt)); fun markup_name ctxt name = markup_extern ctxt name |-> Markup.markup; fun pretty_name ctxt name = markup_extern ctxt name |> Pretty.mark_str; val get_locale = Name_Space.lookup o Locales.get; val defined = is_some oo get_locale; fun the_locale thy name = (case get_locale thy name of SOME (Loc loc) => loc | NONE => error ("Unknown locale " ^ quote name)); fun register_locale binding parameters spec intros axioms hyp_spec syntax_decls notes dependencies thy = thy |> Locales.map (Name_Space.define (Context.Theory thy) true (binding, mk_locale ((parameters, spec, (apply2 o Option.map) Thm.trim_context intros, map Thm.trim_context axioms, map Element.trim_context_ctxt hyp_spec), ((map (fn decl => (Morphism.entity_reset_context decl, serial ())) syntax_decls, map (fn (a, facts) => ((a, map Attrib.trim_context_fact facts), serial ())) notes), (map (fn (name, morph) => make_dep (name, (morph, Morphism.identity))) dependencies, Inttab.empty)))) #> snd); (* FIXME Morphism.identity *) fun change_locale name = Locales.map o Name_Space.map_table_entry name o map_locale o apsnd; (** Primitive operations **) fun parameters_of thy = #parameters o the_locale thy; val params_of = #2 oo parameters_of; fun intros_of thy = (apply2 o Option.map) (Thm.transfer thy) o #intros o the_locale thy; fun axioms_of thy = map (Thm.transfer thy) o #axioms o the_locale thy; fun instance_of thy name morph = params_of thy name |> map (Morphism.term (Morphism.set_context thy morph) o Free o #1); fun specification_of thy = #spec o the_locale thy; fun hyp_spec_of thy = map (Element.transfer_ctxt thy) o #hyp_spec o the_locale thy; fun dependencies_of thy = map (transfer_dep thy) o #dependencies o the_locale thy; fun mixins_of thy name serial = lookup_mixins (#mixins (the_locale thy name)) serial |> (map o apfst o apfst) (Morphism.set_context thy); (* Print instance and qualifiers *) fun pretty_reg_inst ctxt qs (name, ts) = let fun print_qual (qual, mandatory) = qual ^ (if mandatory then "" else "?"); fun prt_quals qs = Pretty.str (space_implode "." (map print_qual qs)); val prt_term = Pretty.quote o Syntax.pretty_term ctxt; fun prt_term' t = if Config.get ctxt show_types then Pretty.block [prt_term t, Pretty.brk 1, Pretty.str "::", Pretty.brk 1, (Pretty.quote o Syntax.pretty_typ ctxt) (type_of t)] else prt_term t; fun prt_inst ts = Pretty.block (Pretty.breaks (pretty_name ctxt name :: map prt_term' ts)); in (case qs of [] => prt_inst ts | qs => Pretty.block [prt_quals qs, Pretty.brk 1, Pretty.str ":", Pretty.brk 1, prt_inst ts]) end; fun pretty_reg ctxt export (name, morph) = let val thy = Proof_Context.theory_of ctxt; val morph' = morph $> export; val qs = Morphism.binding_prefix morph'; val ts = instance_of thy name morph'; in pretty_reg_inst ctxt qs (name, ts) end; (*** Identifiers: activated locales in theory or proof context ***) type idents = term list list Symtab.table; (* name ~> instance (grouped by name) *) val empty_idents : idents = Symtab.empty; val insert_idents = Symtab.insert_list (eq_list (op aconv)); val merge_idents = Symtab.merge_list (eq_list (op aconv)); fun redundant_ident thy idents (name, instance) = exists (fn pat => Pattern.matchess thy (pat, instance)) (Symtab.lookup_list idents name); structure Idents = Generic_Data ( type T = idents; val empty = empty_idents; val merge = merge_idents; ); (** Resolve locale dependencies in a depth-first fashion **) local val roundup_bound = 120; fun add thy depth stem export (name, morph) (deps, marked) = if depth > roundup_bound then error "Roundup bound exceeded (sublocale relation probably not terminating)." else let val instance = instance_of thy name (morph $> stem $> export); in if redundant_ident thy marked (name, instance) then (deps, marked) else let (*no inheritance of mixins, regardless of requests by clients*) val dependencies = dependencies_of thy name |> map (fn dep as {morphisms = (morph', export'), ...} => (#name dep, morph' $> export' $> compose_mixins (mixins_of thy name (#serial dep)))); val marked' = insert_idents (name, instance) marked; val (deps', marked'') = fold_rev (add thy (depth + 1) (morph $> stem) export) dependencies ([], marked'); in ((name, morph $> stem) :: deps' @ deps, marked'') end end; in (* Note that while identifiers always have the external (exported) view, activate_dep is presented with the internal view. *) fun roundup thy activate_dep export (name, morph) (marked, input) = let (* Find all dependencies including new ones (which are dependencies enriching existing registrations). *) val (dependencies, marked') = add thy 0 Morphism.identity export (name, morph) ([], empty_idents); (* Filter out fragments from marked; these won't be activated. *) val dependencies' = filter_out (fn (name, morph) => redundant_ident thy marked (name, instance_of thy name (morph $> export))) dependencies; in (merge_idents (marked, marked'), input |> fold_rev activate_dep dependencies') end; end; (*** Registrations: interpretations in theories or proof contexts ***) val total_ident_ord = prod_ord fast_string_ord (list_ord Term_Ord.fast_term_ord); structure Idtab = Table(type key = string * term list val ord = total_ident_ord); type reg = {morphisms: morphism * morphism, pos: Position.T, serial: serial}; val eq_reg: reg * reg -> bool = op = o apply2 #serial; (* FIXME consolidate with locale dependencies, consider one data slot only *) structure Global_Registrations = Theory_Data' ( (*registrations, indexed by locale name and instance; unique registration serial points to mixin list*) type T = reg Idtab.table * mixins; val empty: T = (Idtab.empty, Inttab.empty); fun merge args = let val ctxt0 = Syntax.init_pretty_global (#1 (hd args)); fun recursive_merge ((regs1, mixins1), (regs2, mixins2)) : T = (Idtab.merge eq_reg (regs1, regs2), merge_mixins (mixins1, mixins2)) handle Idtab.DUP id => (*distinct interpretations with same base: merge their mixins*) let val reg1 = Idtab.lookup regs1 id |> the; val reg2 = Idtab.lookup regs2 id |> the; val reg2' = {morphisms = #morphisms reg2, pos = Position.thread_data (), serial = #serial reg1}; val regs2' = Idtab.update (id, reg2') regs2; val mixins2' = rename_mixin (#serial reg2, #serial reg1) mixins2; val _ = warning ("Removed duplicate interpretation after retrieving its mixins" ^ Position.here_list [#pos reg1, #pos reg2] ^ ":\n " ^ Pretty.string_of (pretty_reg_inst ctxt0 [] id)); in recursive_merge ((regs1, mixins1), (regs2', mixins2')) end; in Library.foldl1 recursive_merge (map #2 args) end; ); structure Local_Registrations = Proof_Data ( type T = Global_Registrations.T; val init = Global_Registrations.get; ); val get_registrations = Context.cases Global_Registrations.get Local_Registrations.get; fun map_registrations f (Context.Theory thy) = Context.Theory (Global_Registrations.map f thy) | map_registrations f (Context.Proof ctxt) = Context.Proof (Local_Registrations.map f ctxt); (* Primitive operations *) fun add_reg thy export (name, morph) = let val reg = {morphisms = (Morphism.reset_context morph, Morphism.reset_context export), pos = Position.thread_data (), serial = serial ()}; val id = (name, instance_of thy name (morph $> export)); in (map_registrations o apfst) (Idtab.insert (K false) (id, reg)) end; fun add_mixin serial' mixin = (* registration to be amended identified by its serial id *) (map_registrations o apsnd) (insert_mixin serial' mixin); val get_regs = #1 o get_registrations; fun get_mixins context (name, morph) = let val thy = Context.theory_of context; val (regs, mixins) = get_registrations context; in (case Idtab.lookup regs (name, instance_of thy name morph) of NONE => [] | SOME {serial, ...} => lookup_mixins mixins serial) end; fun collect_mixins context (name, morph) = let val thy = Context.theory_of context; in roundup thy (fn dep => fn mixins => merge (eq_snd op =) (mixins, get_mixins context dep)) Morphism.identity (name, morph) (insert_idents (name, instance_of thy name morph) empty_idents, []) |> snd |> filter (snd o fst) (* only inheritable mixins *) |> (fn x => merge (eq_snd op =) (x, get_mixins context (name, morph))) |> compose_mixins end; (*** Activate context elements of locale ***) fun activate_err msg kind (name, morph) context = cat_error msg ("The above error(s) occurred while activating " ^ kind ^ " of locale instance\n" ^ (pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |> Pretty.string_of)); fun init_element elem context = context |> Context.mapping I (Thm.unchecked_hyps #> Context_Position.not_really) |> Element.init elem |> Context.mapping I (fn ctxt => let val ctxt0 = Context.proof_of context in ctxt |> Context_Position.restore_visible ctxt0 |> Thm.restore_hyps ctxt0 end); (* Potentially lazy notes *) fun make_notes kind = map (fn ((b, atts), facts) => if null atts andalso forall (null o #2) facts then Lazy_Notes (kind, (b, Lazy.value (maps #1 facts))) else Notes (kind, [((b, atts), facts)])); fun locale_notes thy loc = fold (cons o #1) (#notes (the_locale thy loc)) []; fun lazy_notes thy loc = locale_notes thy loc |> maps (fn (kind, notes) => make_notes kind notes); fun consolidate_notes elems = elems |> map_filter (fn Lazy_Notes (_, (_, ths)) => SOME ths | _ => NONE) |> Lazy.consolidate |> ignore; fun force_notes (Lazy_Notes (kind, (b, ths))) = Notes (kind, [((b, []), [(Lazy.force ths, [])])]) | force_notes elem = elem; (* Declarations, facts and entire locale content *) val trace_locales = Attrib.setup_config_bool (Binding.make ("trace_locales", \<^here>)) (K false); fun tracing context msg = if Config.get context trace_locales then Output.tracing msg else (); fun trace kind (name, morph) context = tracing (Context.proof_of context) ("Activating " ^ kind ^ " of " ^ (pretty_reg (Context.proof_of context) Morphism.identity (name, morph) |> Pretty.string_of)); fun activate_syntax_decls (name, morph) context = let val _ = trace "syntax" (name, morph) context; val thy = Context.theory_of context; val {syntax_decls, ...} = the_locale thy name; val form_syntax_decl = Morphism.form o Morphism.transform morph o Morphism.entity_set_context thy; in fold_rev (form_syntax_decl o #1) syntax_decls context handle ERROR msg => activate_err msg "syntax" (name, morph) context end; fun activate_notes activ_elem context export' (name, morph) input = let val thy = Context.theory_of context; val mixin = (case export' of NONE => Morphism.identity | SOME export => collect_mixins context (name, morph $> export) $> export); val morph' = Morphism.set_context thy (morph $> mixin); val notes' = map (Element.transform_ctxt morph') (lazy_notes thy name); in (notes', input) |-> fold (fn elem => fn res => activ_elem (Element.transfer_ctxt thy elem) res) end handle ERROR msg => activate_err msg "facts" (name, morph) context; fun activate_notes_trace activ_elem context export' (name, morph) context' = let val _ = trace "facts" (name, morph) context'; in activate_notes activ_elem context export' (name, morph) context' end; fun activate_all name thy activ_elem (marked, input) = let val {parameters = (_, params), spec = (asm, defs), ...} = the_locale thy name; val input' = input |> (not (null params) ? activ_elem (Fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params))) |> (* FIXME type parameters *) (case asm of SOME A => activ_elem (Assumes [(Binding.empty_atts, [(A, [])])]) | _ => I) |> (not (null defs) ? activ_elem (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs))); val activate = activate_notes activ_elem (Context.Theory thy) NONE; in roundup thy activate Morphism.identity (name, Morphism.identity) (marked, input') end; (** Public activation functions **) fun activate_facts export dep context = context |> Context_Position.set_visible_generic false |> pair (Idents.get context) |> roundup (Context.theory_of context) (activate_notes_trace init_element context export) (Morphism.default export) dep |-> Idents.put |> Context_Position.restore_visible_generic context; fun activate_declarations dep = Context.proof_map (fn context => context |> Context_Position.set_visible_generic false |> pair (Idents.get context) |> roundup (Context.theory_of context) activate_syntax_decls Morphism.identity dep |-> Idents.put |> Context_Position.restore_visible_generic context); fun init name thy = let val context = Context.Proof (Proof_Context.init_global thy); val marked = Idents.get context; in context |> Context_Position.set_visible_generic false |> pair empty_idents |> activate_all name thy init_element |-> (fn marked' => Idents.put (merge_idents (marked, marked'))) |> Context_Position.restore_visible_generic context |> Context.proof_of end; (*** Add and extend registrations ***) type registration = Locale.registration; fun amend_registration {mixin = NONE, ...} context = context | amend_registration {inst = (name, morph), mixin = SOME mixin, export} context = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val regs = get_regs context; val base = instance_of thy name (morph $> export); val serial' = (case Idtab.lookup regs (name, base) of NONE => error ("No interpretation of locale " ^ quote (markup_name ctxt name) ^ " with\nparameter instantiation " ^ space_implode " " (map (quote o Syntax.string_of_term_global thy) base) ^ " available") | SOME {serial = serial', ...} => serial'); in add_mixin serial' mixin context end; (* Note that a registration that would be subsumed by an existing one will not be generated, and it will not be possible to amend it. *) fun add_registration {inst = (name, base_morph), mixin, export} context = let val thy = Context.theory_of context; val pos_morph = Morphism.binding_morphism "position" (Binding.set_pos (Position.thread_data ())); val mix_morph = (case mixin of NONE => base_morph | SOME (mix, _) => base_morph $> mix); val inst = instance_of thy name mix_morph; val idents = Idents.get context; in if redundant_ident thy idents (name, inst) then context (* FIXME amend mixins? *) else (idents, context) (* add new registrations with inherited mixins *) |> roundup thy (add_reg thy export) export (name, mix_morph) |> #2 (* add mixin *) |> amend_registration {inst = (name, mix_morph), mixin = mixin, export = export} (* activate import hierarchy as far as not already active *) |> activate_facts (SOME export) (name, mix_morph $> pos_morph) end; (*** Dependencies ***) fun registrations_of context loc = Idtab.fold_rev (fn ((name, _), {morphisms, ...}) => name = loc ? cons (name, morphisms)) (get_regs context) [] (*with inherited mixins*) |> map (fn (name, (base, export)) => (name, base $> (collect_mixins context (name, base $> export)) $> export)); fun add_dependency loc {inst = (name, morph), mixin, export} thy = let val dep = make_dep (name, (morph, export)); val add_dep = apfst (cons dep) #> apsnd (case mixin of NONE => I | SOME mixin => insert_mixin (#serial dep) mixin); val thy' = change_locale loc (apsnd add_dep) thy; val context' = Context.Theory thy'; val (_, regs) = fold_rev (roundup thy' cons export) (registrations_of context' loc) (Idents.get context', []); in fold_rev (fn inst => Context.theory_map (add_registration {inst = inst, mixin = NONE, export = export})) regs thy' end; (*** Storing results ***) fun add_facts loc kind facts ctxt = if null facts then ctxt else let val stored_notes = ((kind, map Attrib.trim_context_fact facts), serial ()); val applied_notes = make_notes kind facts; fun apply_notes morph = applied_notes |> fold (fn elem => fn thy => let val elem' = Element.transform_ctxt (Morphism.set_context thy morph) elem in Context.theory_map (Element.init elem') thy end); fun apply_registrations thy = fold_rev (apply_notes o #2) (registrations_of (Context.Theory thy) loc) thy; in ctxt |> Attrib.local_notes kind facts |> #2 |> Proof_Context.background_theory ((change_locale loc o apfst o apsnd) (cons stored_notes) #> apply_registrations) end; -fun add_declaration loc syntax decl = +fun add_declaration loc {syntax, pos} decl = let val decl0 = Morphism.entity_reset_context decl in syntax ? Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl0, serial ()))) - #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration decl0)] + #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration pos decl0)] end; (*** Reasoning about locales ***) (* Storage for witnesses, intro and unfold rules *) structure Thms = Generic_Data ( type T = thm Item_Net.T * thm Item_Net.T * thm Item_Net.T; val empty = (Thm.item_net, Thm.item_net, Thm.item_net); fun merge ((witnesses1, intros1, unfolds1), (witnesses2, intros2, unfolds2)) = (Item_Net.merge (witnesses1, witnesses2), Item_Net.merge (intros1, intros2), Item_Net.merge (unfolds1, unfolds2)); ); fun get_thms which ctxt = map (Thm.transfer' ctxt) (which (Thms.get (Context.Proof ctxt))); val get_witnesses = get_thms (Item_Net.content o #1); val get_intros = get_thms (Item_Net.content o #2); val get_unfolds = get_thms (Item_Net.content o #3); val witness_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (Item_Net.update (Thm.trim_context th) x, y, z))); val intro_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, Item_Net.update (Thm.trim_context th) y, z))); val unfold_add = Thm.declaration_attribute (fn th => Thms.map (fn (x, y, z) => (x, y, Item_Net.update (Thm.trim_context th) z))); (* Tactics *) fun intro_locales_tac {strict, eager} ctxt = (if strict then Method.intros_tac else Method.try_intros_tac) ctxt (get_witnesses ctxt @ get_intros ctxt @ (if eager then get_unfolds ctxt else [])); val _ = Theory.setup (Method.setup \<^binding>\intro_locales\ (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = false})) "back-chain introduction rules of locales without unfolding predicates" #> Method.setup \<^binding>\unfold_locales\ (Scan.succeed (METHOD o intro_locales_tac {strict = false, eager = true})) "back-chain all introduction rules of locales"); (*** diagnostic commands and interfaces ***) fun get_locales thy = map #1 (Name_Space.dest_table (Locales.get thy)); fun pretty_locales thy verbose = Pretty.block (Pretty.breaks (Pretty.str "locales:" :: map (Pretty.mark_str o #1) (Name_Space.markup_table verbose (Proof_Context.init_global thy) (Locales.get thy)))); fun pretty_locale thy show_facts name = let val locale_ctxt = init name thy; fun cons_elem (elem as Notes _) = show_facts ? cons elem | cons_elem (elem as Lazy_Notes _) = show_facts ? cons elem | cons_elem elem = cons elem; val elems = activate_all name thy cons_elem (empty_idents, []) |> snd |> rev |> tap consolidate_notes |> map force_notes; in Pretty.block (Pretty.keyword1 "locale" :: Pretty.brk 1 :: pretty_name locale_ctxt name :: maps (fn elem => [Pretty.fbrk, Pretty.chunks (Element.pretty_ctxt locale_ctxt elem)]) elems) end; fun pretty_registrations ctxt name = (case registrations_of (Context.Proof ctxt) name of [] => Pretty.str "no interpretations" | regs => Pretty.big_list "interpretations:" (map (pretty_reg ctxt Morphism.identity) (rev regs))); fun pretty_locale_deps thy = let fun make_node name = {name = name, parents = map #name (dependencies_of thy name), body = pretty_locale thy false name}; val names = sort_strings (Name_Space.fold_table (cons o #1) (Locales.get thy) []); in map make_node names end; type locale_dependency = {source: string, target: string, prefix: (string * bool) list, morphism: morphism, pos: Position.T, serial: serial}; fun dest_dependencies prev_thys thy = let fun remove_prev loc prev_thy deps = (case get_locale prev_thy loc of NONE => deps | SOME (Loc {dependencies = prev_deps, ...}) => if eq_list eq_dep (prev_deps, deps) then [] else subtract eq_dep prev_deps deps); fun result loc (dep: dep) = let val morphism = op $> (#morphisms dep) in {source = #name dep, target = loc, prefix = Morphism.binding_prefix morphism, morphism = morphism, pos = #pos dep, serial = #serial dep} end; fun add (loc, Loc {dependencies = deps, ...}) = fold (cons o result loc) (fold (remove_prev loc) prev_thys deps); in Name_Space.fold_table add (Locales.get thy) [] |> sort (int_ord o apply2 #serial) end; end; diff --git a/src/Pure/Isar/method.ML b/src/Pure/Isar/method.ML --- a/src/Pure/Isar/method.ML +++ b/src/Pure/Isar/method.ML @@ -1,836 +1,837 @@ (* Title: Pure/Isar/method.ML Author: Markus Wenzel, TU Muenchen Isar proof methods. *) signature METHOD = sig type method = thm list -> context_tactic val CONTEXT_METHOD: (thm list -> context_tactic) -> method val METHOD: (thm list -> tactic) -> method val fail: method val succeed: method val insert_tac: Proof.context -> thm list -> int -> tactic val insert: thm list -> method val SIMPLE_METHOD: tactic -> method val SIMPLE_METHOD': (int -> tactic) -> method val SIMPLE_METHOD'': ((int -> tactic) -> tactic) -> (int -> tactic) -> method val goal_cases_tac: string list -> context_tactic val cheating: bool -> method val intro: Proof.context -> thm list -> method val elim: Proof.context -> thm list -> method val unfold: thm list -> Proof.context -> method val fold: thm list -> Proof.context -> method val atomize: bool -> Proof.context -> method val this: Proof.context -> method val fact: thm list -> Proof.context -> method val assm_tac: Proof.context -> int -> tactic val all_assm_tac: Proof.context -> tactic val assumption: Proof.context -> method val rule_trace: bool Config.T val trace: Proof.context -> thm list -> unit val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val some_rule_tac: Proof.context -> thm list -> thm list -> int -> tactic val intros_tac: Proof.context -> thm list -> thm list -> tactic val try_intros_tac: Proof.context -> thm list -> thm list -> tactic val rule: Proof.context -> thm list -> method val erule: Proof.context -> int -> thm list -> method val drule: Proof.context -> int -> thm list -> method val frule: Proof.context -> int -> thm list -> method val method_space: Context.generic -> Name_Space.T val set_tactic: (morphism -> thm list -> tactic) -> Context.generic -> Context.generic val clean_facts: thm list -> thm list val set_facts: thm list -> Proof.context -> Proof.context val get_facts: Proof.context -> thm list type combinator_info val no_combinator_info: combinator_info datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list val map_source: (Token.src -> Token.src) -> text -> text val primitive_text: (Proof.context -> thm -> thm) -> text val succeed_text: text val standard_text: text val this_text: text val done_text: text val sorry_text: bool -> text val finish_text: text option * bool -> text val print_methods: bool -> Proof.context -> unit val check_name: Proof.context -> xstring * Position.T -> string val check_src: Proof.context -> Token.src -> Token.src val check_text: Proof.context -> text -> text val checked_text: text -> bool val method_syntax: (Proof.context -> method) context_parser -> Token.src -> Proof.context -> method val setup: binding -> (Proof.context -> method) context_parser -> string -> theory -> theory val local_setup: binding -> (Proof.context -> method) context_parser -> string -> local_theory -> local_theory val method_setup: bstring * Position.T -> Input.source -> string -> local_theory -> local_theory val method: Proof.context -> Token.src -> Proof.context -> method val method_closure: Proof.context -> Token.src -> Token.src val closure: bool Config.T val method_cmd: Proof.context -> Token.src -> Proof.context -> method val detect_closure_state: thm -> bool val STATIC: (unit -> unit) -> context_tactic val RUNTIME: context_tactic -> context_tactic val sleep: Time.time -> context_tactic val evaluate: text -> Proof.context -> method val evaluate_runtime: text -> Proof.context -> method type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T} val modifier: attribute -> Position.T -> modifier val old_section_parser: bool Config.T val sections: modifier parser list -> unit context_parser type text_range = text * Position.range val text: text_range option -> text option val position: text_range option -> Position.T val reports_of: text_range -> Position.report list val report: text_range -> unit val parser: int -> text_range parser val parse: text_range parser val parse_by: ((text_range * text_range option) * Position.report list) parser val read: Proof.context -> Token.src -> text val read_closure: Proof.context -> Token.src -> text * Token.src val read_closure_input: Proof.context -> Input.source -> text * Token.src val text_closure: text context_parser end; structure Method: METHOD = struct (** proof methods **) (* type method *) type method = thm list -> context_tactic; fun CONTEXT_METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac) #> Seq.maps_results (tac facts); fun METHOD tac : method = fn facts => CONTEXT_TACTIC (ALLGOALS Goal.conjunction_tac THEN tac facts); val fail = METHOD (K no_tac); val succeed = METHOD (K all_tac); (* insert facts *) fun insert_tac _ [] _ = all_tac | insert_tac ctxt facts i = EVERY (map (fn r => resolve_tac ctxt [Thm.forall_intr_vars r COMP_INCR revcut_rl] i) facts); fun insert thms = CONTEXT_METHOD (fn _ => fn (ctxt, st) => st |> ALLGOALS (insert_tac ctxt thms) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> (ALLGOALS (insert_tac ctxt facts) THEN tac) |> TACTIC_CONTEXT ctxt); fun SIMPLE_METHOD'' quant tac = CONTEXT_METHOD (fn facts => fn (ctxt, st) => st |> quant (insert_tac ctxt facts THEN' tac) |> TACTIC_CONTEXT ctxt); val SIMPLE_METHOD' = SIMPLE_METHOD'' HEADGOAL; (* goals as cases *) fun goal_cases_tac case_names : context_tactic = fn (ctxt, st) => let val cases = (if null case_names then map string_of_int (1 upto Thm.nprems_of st) else case_names) |> map (rpair [] o rpair []) |> Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params st)); in CONTEXT_CASES cases all_tac (ctxt, st) end; (* cheating *) fun cheating int = CONTEXT_METHOD (fn _ => fn (ctxt, st) => if int orelse Config.get ctxt quick_and_dirty then TACTIC_CONTEXT ctxt (ALLGOALS (Skip_Proof.cheat_tac ctxt) st) else error "Cheating requires quick_and_dirty mode!"); (* unfold intro/elim rules *) fun intro ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (match_tac ctxt ths)); fun elim ctxt ths = SIMPLE_METHOD' (CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt ths)); (* unfold/fold definitions *) fun unfold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.unfold_tac ctxt ths)); fun fold_meth ths ctxt = SIMPLE_METHOD (CHANGED_PROP (Local_Defs.fold_tac ctxt ths)); (* atomize rule statements *) fun atomize false ctxt = SIMPLE_METHOD' (CHANGED_PROP o Object_Logic.atomize_prems_tac ctxt) | atomize true ctxt = Context_Tactic.CONTEXT_TACTIC o K (HEADGOAL (CHANGED_PROP o Object_Logic.full_atomize_tac ctxt)); (* this -- resolve facts directly *) fun this ctxt = METHOD (EVERY o map (HEADGOAL o resolve_tac ctxt o single)); (* fact -- composition by facts from context *) fun fact [] ctxt = SIMPLE_METHOD' (Proof_Context.some_fact_tac ctxt) | fact rules ctxt = SIMPLE_METHOD' (Proof_Context.fact_tac ctxt rules); (* assumption *) local fun cond_rtac ctxt cond rule = SUBGOAL (fn (prop, i) => if cond (Logic.strip_assums_concl prop) then resolve_tac ctxt [rule] i else no_tac); in fun assm_tac ctxt = assume_tac ctxt APPEND' Goal.assume_rule_tac ctxt APPEND' cond_rtac ctxt (can Logic.dest_equals) Drule.reflexive_thm APPEND' cond_rtac ctxt (can Logic.dest_term) Drule.termI; fun all_assm_tac ctxt = let fun tac i st = if i > Thm.nprems_of st then all_tac st else ((assm_tac ctxt i THEN tac i) ORELSE tac (i + 1)) st; in tac 1 end; fun assumption ctxt = METHOD (HEADGOAL o (fn [] => assm_tac ctxt | [fact] => solve_tac ctxt [fact] | _ => K no_tac)); fun finish immed ctxt = METHOD (K ((if immed then all_assm_tac ctxt else all_tac) THEN flexflex_tac ctxt)); end; (* rule etc. -- single-step refinements *) val rule_trace = Attrib.setup_config_bool \<^binding>\rule_trace\ (fn _ => false); fun trace ctxt rules = if Config.get ctxt rule_trace andalso not (null rules) then Pretty.big_list "rules:" (map (Thm.pretty_thm_item ctxt) rules) |> Pretty.string_of |> tracing else (); local fun gen_rule_tac tac ctxt rules facts = (fn i => fn st => if null facts then tac ctxt rules i st else Seq.maps (fn rule => (tac ctxt o single) rule i st) (Drule.multi_resolves (SOME ctxt) facts rules)) THEN_ALL_NEW Goal.norm_hhf_tac ctxt; fun gen_arule_tac tac ctxt j rules facts = EVERY' (gen_rule_tac tac ctxt rules facts :: replicate j (assume_tac ctxt)); fun gen_some_rule_tac tac ctxt arg_rules facts = SUBGOAL (fn (goal, i) => let val rules = if not (null arg_rules) then arg_rules else flat (Context_Rules.find_rules ctxt false facts goal); in trace ctxt rules; tac ctxt rules facts i end); fun meth tac x y = METHOD (HEADGOAL o tac x y); fun meth' tac x y z = METHOD (HEADGOAL o tac x y z); in val rule_tac = gen_rule_tac resolve_tac; val rule = meth rule_tac; val some_rule_tac = gen_some_rule_tac rule_tac; val some_rule = meth some_rule_tac; val erule = meth' (gen_arule_tac eresolve_tac); val drule = meth' (gen_arule_tac dresolve_tac); val frule = meth' (gen_arule_tac forward_tac); end; (* intros_tac -- pervasive search spanned by intro rules *) fun gen_intros_tac goals ctxt intros facts = goals (insert_tac ctxt facts THEN' REPEAT_ALL_NEW (resolve_tac ctxt intros)) THEN Tactic.distinct_subgoals_tac; val intros_tac = gen_intros_tac ALLGOALS; val try_intros_tac = gen_intros_tac TRYALL; (** method syntax **) (* context data *) structure Data = Generic_Data ( type T = {methods: ((Token.src -> Proof.context -> method) * string) Name_Space.table, ml_tactic: (morphism -> thm list -> tactic) option, facts: thm list option}; val empty : T = {methods = Name_Space.empty_table Markup.methodN, ml_tactic = NONE, facts = NONE}; fun merge ({methods = methods1, ml_tactic = ml_tactic1, facts = facts1}, {methods = methods2, ml_tactic = ml_tactic2, facts = facts2}) : T = {methods = Name_Space.merge_tables (methods1, methods2), ml_tactic = merge_options (ml_tactic1, ml_tactic2), facts = merge_options (facts1, facts2)}; ); fun map_data f = Data.map (fn {methods, ml_tactic, facts} => let val (methods', ml_tactic', facts') = f (methods, ml_tactic, facts) in {methods = methods', ml_tactic = ml_tactic', facts = facts'} end); val get_methods = #methods o Data.get; val ops_methods = {get_data = get_methods, put_data = fn methods => map_data (fn (_, ml_tactic, facts) => (methods, ml_tactic, facts))}; val method_space = Name_Space.space_of_table o get_methods; (* ML tactic *) fun set_tactic ml_tactic = map_data (fn (methods, _, facts) => (methods, SOME ml_tactic, facts)); fun the_tactic context = #ml_tactic (Data.get context) |> \<^if_none>\raise Fail "Undefined ML tactic"\; val parse_tactic = Scan.state :|-- (fn context => Scan.lift (Args.embedded_declaration (fn source => let val tac = context |> ML_Context.expression (Input.pos_of source) (ML_Lex.read "Context.>> (Method.set_tactic (fn morphism: Morphism.morphism => fn facts: thm list => (" @ ML_Lex.read_source source @ ML_Lex.read ")))") |> the_tactic; in Morphism.entity (fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi))) end)) >> (fn decl => Morphism.form_entity (the_tactic (Morphism.form decl context)))); (* method facts *) val clean_facts = filter_out Thm.is_dummy; fun set_facts facts = (Context.proof_map o map_data) (fn (methods, ml_tactic, _) => (methods, ml_tactic, SOME (clean_facts facts))); val get_facts_generic = these o #facts o Data.get; val get_facts = get_facts_generic o Context.Proof; val _ = Theory.setup (Global_Theory.add_thms_dynamic (Binding.make ("method_facts", \<^here>), get_facts_generic)); (* method text *) datatype combinator_info = Combinator_Info of {keywords: Position.T list}; fun combinator_info keywords = Combinator_Info {keywords = keywords}; val no_combinator_info = combinator_info []; datatype combinator = Then | Then_All_New | Orelse | Try | Repeat1 | Select_Goals of int; datatype text = Source of Token.src | Basic of Proof.context -> method | Combinator of combinator_info * combinator * text list; fun map_source f (Source src) = Source (f src) | map_source _ (Basic meth) = Basic meth | map_source f (Combinator (info, comb, txts)) = Combinator (info, comb, map (map_source f) txts); fun primitive_text r = Basic (SIMPLE_METHOD o PRIMITIVE o r); val succeed_text = Basic (K succeed); val standard_text = Source (Token.make_src ("standard", Position.none) []); val this_text = Basic this; val done_text = Basic (K (SIMPLE_METHOD all_tac)); fun sorry_text int = Basic (fn _ => cheating int); fun finish_text (NONE, immed) = Basic (finish immed) | finish_text (SOME txt, immed) = Combinator (no_combinator_info, Then, [txt, Basic (finish immed)]); (* method definitions *) fun print_methods verbose ctxt = let val meths = get_methods (Context.Proof ctxt); fun prt_meth (name, (_, "")) = Pretty.mark_str name | prt_meth (name, (_, comment)) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.brk 2 :: Pretty.text comment); in [Pretty.big_list "methods:" (map prt_meth (Name_Space.markup_table verbose ctxt meths))] |> Pretty.writeln_chunks end; (* define *) fun define_global binding meth comment = Entity.define_global ops_methods binding (meth, comment); fun define binding meth comment = Entity.define ops_methods binding (meth, comment); (* check *) fun check_name ctxt = let val context = Context.Proof ctxt in #1 o Name_Space.check context (get_methods context) end; fun check_src ctxt = #1 o Token.check_src ctxt (get_methods o Context.Proof); fun check_text ctxt (Source src) = Source (check_src ctxt src) | check_text _ (Basic m) = Basic m | check_text ctxt (Combinator (x, y, body)) = Combinator (x, y, map (check_text ctxt) body); fun checked_text (Source src) = Token.checked_src src | checked_text (Basic _) = true | checked_text (Combinator (_, _, body)) = forall checked_text body; val _ = Theory.setup (ML_Antiquotation.inline_embedded \<^binding>\method\ (Args.context -- Scan.lift Parse.embedded_position >> (ML_Syntax.print_string o uncurry check_name))); (* method setup *) fun method_syntax scan src ctxt : method = let val (m, ctxt') = Token.syntax scan src ctxt in m ctxt' end; fun setup binding scan comment = define_global binding (method_syntax scan) comment #> snd; fun local_setup binding scan comment = define binding (method_syntax scan) comment #> snd; fun method_setup binding source comment = ML_Context.expression (Input.pos_of source) (ML_Lex.read ("Theory.local_setup (Method.local_setup (" ^ ML_Syntax.make_binding binding ^ ") (") @ ML_Lex.read_source source @ ML_Lex.read (")" ^ ML_Syntax.print_string comment ^ ")")) |> Context.proof_map; (* prepare methods *) fun method ctxt = let val table = get_methods (Context.Proof ctxt) in fn src => #1 (Name_Space.get table (#1 (Token.name_of_src src))) src end; fun method_closure ctxt src = let val src' = map Token.init_assignable src; val ctxt' = Context_Position.not_really ctxt; val _ = Seq.pull (method ctxt' src' ctxt' [] (ctxt', Goal.protect 0 Drule.dummy_thm)); in map Token.closure src' end; val closure = Config.declare_bool ("Method.closure", \<^here>) (K true); fun method_cmd ctxt = check_src ctxt #> Config.get ctxt closure ? method_closure ctxt #> method ctxt; (* static vs. runtime state *) fun detect_closure_state st = (case try Logic.dest_term (Thm.concl_of (perhaps (try Goal.conclude) st)) of NONE => false | SOME t => Term.is_dummy_pattern t); fun STATIC test : context_tactic = fn (ctxt, st) => if detect_closure_state st then (test (); Seq.single (Seq.Result (ctxt, st))) else Seq.empty; fun RUNTIME (tac: context_tactic) (ctxt, st) = if detect_closure_state st then Seq.empty else tac (ctxt, st); fun sleep t = RUNTIME (fn ctxt_st => (OS.Process.sleep t; Seq.single (Seq.Result ctxt_st))); (* evaluate method text *) local val op THEN = Seq.THEN; fun BYPASS_CONTEXT (tac: tactic) = fn result => (case result of Seq.Error _ => Seq.single result | Seq.Result (ctxt, st) => tac st |> TACTIC_CONTEXT ctxt); val preparation = BYPASS_CONTEXT (ALLGOALS Goal.conjunction_tac); fun RESTRICT_GOAL i n method = BYPASS_CONTEXT (PRIMITIVE (Goal.restrict i n)) THEN method THEN BYPASS_CONTEXT (PRIMITIVE (Goal.unrestrict i)); fun SELECT_GOAL method i = RESTRICT_GOAL i 1 method; fun (method1 THEN_ALL_NEW method2) i (result : context_state Seq.result) = (case result of Seq.Error _ => Seq.single result | Seq.Result (_, st) => result |> method1 i |> Seq.maps (fn result' => (case result' of Seq.Error _ => Seq.single result' | Seq.Result (_, st') => result' |> Seq.INTERVAL method2 i (i + Thm.nprems_of st' - Thm.nprems_of st)))) fun COMBINATOR1 comb [meth] = comb meth | COMBINATOR1 _ _ = raise Fail "Method combinator requires exactly one argument"; fun combinator Then = Seq.EVERY | combinator Then_All_New = (fn [] => Seq.single | methods => preparation THEN (foldl1 (op THEN_ALL_NEW) (map SELECT_GOAL methods) 1)) | combinator Orelse = Seq.FIRST | combinator Try = COMBINATOR1 Seq.TRY | combinator Repeat1 = COMBINATOR1 Seq.REPEAT1 | combinator (Select_Goals n) = COMBINATOR1 (fn method => preparation THEN RESTRICT_GOAL 1 n method); in fun evaluate text ctxt0 facts = let val ctxt = set_facts facts ctxt0; fun eval0 m = Seq.single #> Seq.maps_results (m facts); fun eval (Basic m) = eval0 (m ctxt) | eval (Source src) = eval0 (method_cmd ctxt src ctxt) | eval (Combinator (_, c, txts)) = combinator c (map eval txts); in eval text o Seq.Result end; end; fun evaluate_runtime text ctxt = let val text' = text |> (map_source o map o Token.map_facts) (fn SOME name => (case Proof_Context.lookup_fact ctxt name of SOME {dynamic = true, thms} => K thms | _ => I) | NONE => I); val ctxt' = Config.put closure false ctxt; in fn facts => RUNTIME (fn st => evaluate text' ctxt' facts st) end; (** concrete syntax **) (* type modifier *) type modifier = {init: Proof.context -> Proof.context, attribute: attribute, pos: Position.T}; fun modifier attribute pos : modifier = {init = I, attribute = attribute, pos = pos}; (* sections *) val old_section_parser = Config.declare_bool ("Method.old_section_parser", \<^here>) (K false); local fun thms ss = Scan.repeats (Scan.unless (Scan.lift (Scan.first ss)) Attrib.multi_thm); fun app {init, attribute, pos = _} ths context = fold_map (Thm.apply_attribute attribute) ths (Context.map_proof init context); fun section ss = Scan.depend (fn context => (Scan.first ss -- Scan.pass context (thms ss)) :|-- (fn (m, ths) => Scan.succeed (swap (app m ths context)))); in fun old_sections ss = Scan.repeat (section ss) >> K (); end; local fun sect (modifier : modifier parser) = Scan.depend (fn context => Scan.ahead Parse.not_eof -- Scan.trace modifier -- Scan.repeat (Scan.unless modifier Parse.thm) >> (fn ((tok0, ({init, attribute, pos}, modifier_toks)), xthms) => let val decl = (case Token.get_value tok0 of SOME (Token.Declaration decl) => decl | _ => let val ctxt = Context.proof_of context; val prep_att = Attrib.check_src ctxt #> map (Token.assign NONE); val thms = map (fn (a, bs) => (Proof_Context.get_fact ctxt a, map prep_att bs)) xthms; val facts = Attrib.partial_evaluation ctxt [((Binding.name "dummy", []), thms)] |> map (fn (_, bs) => - ((Binding.empty, [Attrib.internal (K attribute)]), Attrib.trim_context_thms bs)); + ((Binding.empty, [Attrib.internal pos (K attribute)]), + Attrib.trim_context_thms bs)); val decl = Morphism.entity (fn phi => fn context => let val psi = Morphism.set_context'' context phi in context |> Context.mapping I init |> Attrib.generic_notes "" (Attrib.transform_facts psi facts) |> snd end); val modifier_report = (#1 (Token.range_of modifier_toks), Position.entity_markup Markup.method_modifierN ("", pos)); val _ = Context_Position.reports ctxt (modifier_report :: Token.reports_of_value tok0); val _ = Token.assign (SOME (Token.Declaration decl)) tok0; in decl end); in (Morphism.form decl context, decl) end)); in fun sections ss = Args.context :|-- (fn ctxt => if Config.get ctxt old_section_parser then old_sections ss else Scan.repeat (sect (Scan.first ss)) >> K ()); end; (* extra rule methods *) fun xrule_meth meth = Scan.lift (Scan.optional (Args.parens Parse.nat) 0) -- Attrib.thms >> (fn (n, ths) => fn ctxt => meth ctxt n ths); (* text range *) type text_range = text * Position.range; fun text NONE = NONE | text (SOME (txt, _)) = SOME txt; fun position NONE = Position.none | position (SOME (_, (pos, _))) = pos; (* reports *) local fun keyword_positions (Source _) = [] | keyword_positions (Basic _) = [] | keyword_positions (Combinator (Combinator_Info {keywords}, _, texts)) = keywords @ maps keyword_positions texts; in fun reports_of ((text, (pos, _)): text_range) = (pos, Markup.language_method) :: maps (fn p => map (pair p) (Markup.keyword3 :: Completion.suppress_abbrevs "")) (keyword_positions text); fun report text_range = if Context_Position.reports_enabled0 () then Position.reports (reports_of text_range) else (); end; (* parser *) local fun is_symid_meth s = s <> "|" andalso s <> "?" andalso s <> "+" andalso Token.ident_or_symbolic s; in fun parser pri = let val meth_name = Parse.token Parse.name; fun meth5 x = (meth_name >> (Source o single) || Scan.ahead Parse.cartouche |-- Parse.not_eof >> (fn tok => Source (Token.make_src ("cartouche", Position.none) [tok])) || Parse.$$$ "(" |-- Parse.!!! (meth0 --| Parse.$$$ ")")) x and meth4 x = (meth5 -- Parse.position (Parse.$$$ "?") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Try, [m])) || meth5 -- Parse.position (Parse.$$$ "+") >> (fn (m, (_, pos)) => Combinator (combinator_info [pos], Repeat1, [m])) || meth5 -- (Parse.position (Parse.$$$ "[") -- Scan.optional Parse.nat 1 -- Parse.position (Parse.$$$ "]")) >> (fn (m, (((_, pos1), n), (_, pos2))) => Combinator (combinator_info [pos1, pos2], Select_Goals n, [m])) || meth5) x and meth3 x = (meth_name ::: Parse.args1 is_symid_meth >> Source || meth4) x and meth2 x = (Parse.enum1_positions "," meth3 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then, ms))) x and meth1 x = (Parse.enum1_positions ";" meth2 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Then_All_New, ms))) x and meth0 x = (Parse.enum1_positions "|" meth1 >> (fn ([m], _) => m | (ms, ps) => Combinator (combinator_info ps, Orelse, ms))) x; val meth = nth [meth0, meth1, meth2, meth3, meth4, meth5] pri handle General.Subscript => raise Fail ("Bad method parser priority " ^ string_of_int pri); in Scan.trace meth >> (fn (m, toks) => (m, Token.range_of toks)) end; val parse = parser 4; end; val parse_by = Parse.$$$ "by" |-- parse -- Scan.option parse >> (fn (m1, m2) => ((m1, m2), maps reports_of (m1 :: the_list m2))); (* read method text *) fun read ctxt src = (case Scan.read Token.stopper (Parse.!!! (parser 0 --| Scan.ahead Parse.eof)) src of SOME (text, range) => if checked_text text then text else (report (text, range); check_text ctxt text) | NONE => error ("Failed to parse method" ^ Position.here (#1 (Token.range_of src)))); fun read_closure ctxt src0 = let val src1 = map Token.init_assignable src0; val text = read ctxt src1 |> map_source (method_closure ctxt); val src2 = map Token.closure src1; in (text, src2) end; fun read_closure_input ctxt = let val keywords = Keyword.no_major_keywords (Thy_Header.get_keywords' ctxt) in Parse.read_embedded ctxt keywords (Scan.many Token.not_eof) #> read_closure ctxt end; val text_closure = Args.context -- Scan.lift (Parse.token Parse.embedded) >> (fn (ctxt, tok) => (case Token.get_value tok of SOME (Token.Source src) => read ctxt src | _ => let val (text, src) = read_closure_input ctxt (Token.input_of tok); val _ = Token.assign (SOME (Token.Source src)) tok; in text end)); (* theory setup *) val _ = Theory.setup (setup \<^binding>\fail\ (Scan.succeed (K fail)) "force failure" #> setup \<^binding>\succeed\ (Scan.succeed (K succeed)) "succeed" #> setup \<^binding>\sleep\ (Scan.lift Parse.real >> (fn s => fn _ => fn _ => sleep (seconds s))) "succeed after delay (in seconds)" #> setup \<^binding>\-\ (Scan.succeed (K (SIMPLE_METHOD all_tac))) "insert current facts, nothing else" #> setup \<^binding>\goal_cases\ (Scan.lift (Scan.repeat Args.name_token) >> (fn names => fn _ => CONTEXT_METHOD (fn _ => fn (ctxt, st) => (case drop (Thm.nprems_of st) names of [] => NONE | bad => if detect_closure_state st then NONE else SOME (fn () => ("Excessive case name(s): " ^ commas_quote (map Token.content_of bad) ^ Position.here (#1 (Token.range_of bad))))) |> (fn SOME msg => Seq.single (Seq.Error msg) | NONE => goal_cases_tac (map Token.content_of names) (ctxt, st))))) "bind cases for goals" #> setup \<^binding>\subproofs\ (text_closure >> (Context_Tactic.SUBPROOFS ooo evaluate_runtime)) "apply proof method to subproofs with closed derivation" #> setup \<^binding>\insert\ (Attrib.thms >> (K o insert)) "insert theorems, ignoring facts" #> setup \<^binding>\intro\ (Attrib.thms >> (fn ths => fn ctxt => intro ctxt ths)) "repeatedly apply introduction rules" #> setup \<^binding>\elim\ (Attrib.thms >> (fn ths => fn ctxt => elim ctxt ths)) "repeatedly apply elimination rules" #> setup \<^binding>\unfold\ (Attrib.thms >> unfold_meth) "unfold definitions" #> setup \<^binding>\fold\ (Attrib.thms >> fold_meth) "fold definitions" #> setup \<^binding>\atomize\ (Scan.lift (Args.mode "full") >> atomize) "present local premises as object-level statements" #> setup \<^binding>\rule\ (Attrib.thms >> (fn ths => fn ctxt => some_rule ctxt ths)) "apply some intro/elim rule" #> setup \<^binding>\erule\ (xrule_meth erule) "apply rule in elimination manner (improper)" #> setup \<^binding>\drule\ (xrule_meth drule) "apply rule in destruct manner (improper)" #> setup \<^binding>\frule\ (xrule_meth frule) "apply rule in forward manner (improper)" #> setup \<^binding>\this\ (Scan.succeed this) "apply current facts as rules" #> setup \<^binding>\fact\ (Attrib.thms >> fact) "composition by facts from context" #> setup \<^binding>\assumption\ (Scan.succeed assumption) "proof by assumption, preferring facts" #> setup \<^binding>\rename_tac\ (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.name) >> (fn (quant, xs) => K (SIMPLE_METHOD'' quant (rename_tac xs)))) "rename parameters of goal" #> setup \<^binding>\rotate_tac\ (Args.goal_spec -- Scan.lift (Scan.optional Parse.int 1) >> (fn (quant, i) => K (SIMPLE_METHOD'' quant (rotate_tac i)))) "rotate assumptions of goal" #> setup \<^binding>\tactic\ (parse_tactic >> (K o METHOD)) "ML tactic as proof method" #> setup \<^binding>\raw_tactic\ (parse_tactic >> (fn tac => fn _ => Context_Tactic.CONTEXT_TACTIC o tac)) "ML tactic as raw proof method" #> setup \<^binding>\use\ (Attrib.thms -- (Scan.lift (Parse.$$$ "in") |-- text_closure) >> (fn (thms, text) => fn ctxt => fn _ => evaluate_runtime text ctxt thms)) "indicate method facts and context for method expression"); (*final declarations of this structure!*) val unfold = unfold_meth; val fold = fold_meth; end; val CONTEXT_METHOD = Method.CONTEXT_METHOD; val METHOD = Method.METHOD; val SIMPLE_METHOD = Method.SIMPLE_METHOD; val SIMPLE_METHOD' = Method.SIMPLE_METHOD'; val SIMPLE_METHOD'' = Method.SIMPLE_METHOD''; diff --git a/src/Pure/Isar/spec_rules.ML b/src/Pure/Isar/spec_rules.ML --- a/src/Pure/Isar/spec_rules.ML +++ b/src/Pure/Isar/spec_rules.ML @@ -1,189 +1,189 @@ (* Title: Pure/Isar/spec_rules.ML Author: Makarius Rules that characterize specifications, with optional name and rough classification. NB: In the face of arbitrary morphisms, the original shape of specifications may get lost. *) signature SPEC_RULES = sig datatype recursion = Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion val recursion_ord: recursion ord val encode_recursion: recursion XML.Encode.T datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown val rough_classification_ord: rough_classification ord val equational_primrec: string list -> rough_classification val equational_recdef: rough_classification val equational_primcorec: string list -> rough_classification val equational_corec: rough_classification val equational: rough_classification val is_equational: rough_classification -> bool val is_inductive: rough_classification -> bool val is_co_inductive: rough_classification -> bool val is_relational: rough_classification -> bool val is_unknown: rough_classification -> bool val encode_rough_classification: rough_classification XML.Encode.T type spec_rule = {pos: Position.T, name: string, rough_classification: rough_classification, terms: term list, rules: thm list} val get: Proof.context -> spec_rule list val get_global: theory -> spec_rule list val dest_theory: theory -> spec_rule list val retrieve: Proof.context -> term -> spec_rule list val retrieve_global: theory -> term -> spec_rule list val add: binding -> rough_classification -> term list -> thm list -> local_theory -> local_theory val add_global: binding -> rough_classification -> term list -> thm list -> theory -> theory end; structure Spec_Rules: SPEC_RULES = struct (* recursion *) datatype recursion = Primrec of string list | Recdef | Primcorec of string list | Corec | Unknown_Recursion; val recursion_index = fn Primrec _ => 0 | Recdef => 1 | Primcorec _ => 2 | Corec => 3 | Unknown_Recursion => 4; fun recursion_ord (Primrec Ts1, Primrec Ts2) = list_ord fast_string_ord (Ts1, Ts2) | recursion_ord (Primcorec Ts1, Primcorec Ts2) = list_ord fast_string_ord (Ts1, Ts2) | recursion_ord rs = int_ord (apply2 recursion_index rs); val encode_recursion = let open XML.Encode in variant [fn Primrec a => ([], list string a), fn Recdef => ([], []), fn Primcorec a => ([], list string a), fn Corec => ([], []), fn Unknown_Recursion => ([], [])] end; (* rough classification *) datatype rough_classification = Equational of recursion | Inductive | Co_Inductive | Unknown; fun rough_classification_ord (Equational r1, Equational r2) = recursion_ord (r1, r2) | rough_classification_ord cs = int_ord (apply2 (fn Equational _ => 0 | Inductive => 1 | Co_Inductive => 2 | Unknown => 3) cs); val equational_primrec = Equational o Primrec; val equational_recdef = Equational Recdef; val equational_primcorec = Equational o Primcorec; val equational_corec = Equational Corec; val equational = Equational Unknown_Recursion; val is_equational = fn Equational _ => true | _ => false; val is_inductive = fn Inductive => true | _ => false; val is_co_inductive = fn Co_Inductive => true | _ => false; val is_relational = is_inductive orf is_co_inductive; val is_unknown = fn Unknown => true | _ => false; val encode_rough_classification = let open XML.Encode in variant [fn Equational r => ([], encode_recursion r), fn Inductive => ([], []), fn Co_Inductive => ([], []), fn Unknown => ([], [])] end; (* rules *) type spec_rule = {pos: Position.T, name: string, rough_classification: rough_classification, terms: term list, rules: thm list}; fun eq_spec (specs: spec_rule * spec_rule) = (op =) (apply2 #name specs) andalso is_equal (rough_classification_ord (apply2 #rough_classification specs)) andalso eq_list (op aconv) (apply2 #terms specs) andalso eq_list Thm.eq_thm_prop (apply2 #rules specs); fun map_spec_rules f ({pos, name, rough_classification, terms, rules}: spec_rule) : spec_rule = {pos = pos, name = name, rough_classification = rough_classification, terms = terms, rules = map f rules}; structure Data = Generic_Data ( type T = spec_rule Item_Net.T; val empty : T = Item_Net.init eq_spec #terms; val merge = Item_Net.merge; ); (* get *) fun get_generic imports context = let val thy = Context.theory_of context; val transfer = Global_Theory.transfer_theories thy; fun imported spec = imports |> exists (fn thy => Item_Net.member (Data.get (Context.Theory thy)) spec); in Item_Net.content (Data.get context) |> filter_out imported |> (map o map_spec_rules) transfer end; val get = get_generic [] o Context.Proof; val get_global = get_generic [] o Context.Theory; fun dest_theory thy = rev (get_generic (Theory.parents_of thy) (Context.Theory thy)); (* retrieve *) fun retrieve_generic context = Item_Net.retrieve (Data.get context) #> (map o map_spec_rules) (Thm.transfer'' context); val retrieve = retrieve_generic o Context.Proof; val retrieve_global = retrieve_generic o Context.Theory; (* add *) fun add b rough_classification terms rules lthy = let val n = length terms; val thms0 = map Thm.trim_context (map (Drule.mk_term o Thm.cterm_of lthy) terms @ rules); in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = Binding.pos_of b} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi; val pos = Position.thread_data (); val name = Name_Space.full_name (Name_Space.naming_of context) (Morphism.binding psi b); val (terms', rules') = chop n (Morphism.fact psi thms0) |>> map (Thm.term_of o Drule.dest_term); in context |> (Data.map o Item_Net.update) {pos = pos, name = name, rough_classification = rough_classification, terms = terms', rules = rules'} end) end; fun add_global b rough_classification terms rules thy = thy |> (Context.theory_map o Data.map o Item_Net.update) {pos = Position.thread_data (), name = Sign.full_name thy b, rough_classification = rough_classification, terms = terms, rules = map Thm.trim_context rules}; end; diff --git a/src/Pure/Isar/token.ML b/src/Pure/Isar/token.ML --- a/src/Pure/Isar/token.ML +++ b/src/Pure/Isar/token.ML @@ -1,833 +1,839 @@ (* Title: Pure/Isar/token.ML Author: Markus Wenzel, TU Muenchen Outer token syntax for Isabelle/Isar. *) signature TOKEN = sig datatype kind = (*immediate source*) Command | Keyword | Ident | Long_Ident | Sym_Ident | Var | Type_Ident | Type_Var | Nat | Float | Space | (*delimited content*) String | Alt_String | Cartouche | Control of Antiquote.control | Comment of Comment.kind option | (*special content*) Error of string | EOF val control_kind: kind val str_of_kind: kind -> string type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T} type T type src = T list type name_value = {name: string, kind: string, print: Proof.context -> Markup.T * xstring} datatype value = Source of src | Literal of bool * Markup.T | Name of name_value * morphism | Typ of typ | Term of term | Fact of string option * thm list | Attribute of attribute Morphism.entity | Declaration of Morphism.declaration_entity | Files of file Exn.result list | Output of XML.body option val pos_of: T -> Position.T val adjust_offsets: (int -> int option) -> T -> T val eof: T val is_eof: T -> bool val not_eof: T -> bool val stopper: T Scan.stopper val kind_of: T -> kind val is_kind: kind -> T -> bool val get_control: T -> Antiquote.control option val is_command: T -> bool val keyword_with: (string -> bool) -> T -> bool val is_command_modifier: T -> bool val ident_with: (string -> bool) -> T -> bool val is_proper: T -> bool val is_comment: T -> bool val is_informal_comment: T -> bool val is_formal_comment: T -> bool val is_document_marker: T -> bool val is_ignored: T -> bool val is_begin_ignore: T -> bool val is_end_ignore: T -> bool val is_error: T -> bool val is_space: T -> bool val is_blank: T -> bool val is_newline: T -> bool val range_of: T list -> Position.range val core_range_of: T list -> Position.range val content_of: T -> string val source_of: T -> string val input_of: T -> Input.source val inner_syntax_of: T -> string val keyword_markup: bool * Markup.T -> string -> Markup.T val completion_report: T -> Position.report_text list val reports: Keyword.keywords -> T -> Position.report_text list val markups: Keyword.keywords -> T -> Markup.T list val unparse: T -> string val print: T -> string val text_of: T -> string * string val file_source: file -> Input.source val get_files: T -> file Exn.result list val put_files: file Exn.result list -> T -> T val get_output: T -> XML.body option val put_output: XML.body -> T -> T val get_value: T -> value option val reports_of_value: T -> Position.report list val name_value: name_value -> value val get_name: T -> name_value option val declare_maxidx: T -> Proof.context -> Proof.context val map_facts: (string option -> thm list -> thm list) -> T -> T val trim_context: T -> T val transfer: theory -> T -> T val transform: morphism -> T -> T val init_assignable: T -> T val assign: value option -> T -> T val evaluate: ('a -> value) -> (T -> 'a) -> T -> 'a val closure: T -> T val pretty_value: Proof.context -> T -> Pretty.T val name_of_src: src -> string * Position.T val args_of_src: src -> T list val checked_src: src -> bool val check_src: Proof.context -> (Proof.context -> 'a Name_Space.table) -> src -> src * 'a val pretty_src: Proof.context -> src -> Pretty.T val ident_or_symbolic: string -> bool val read_cartouche: Symbol_Pos.T list -> T val tokenize: Keyword.keywords -> {strict: bool} -> Symbol_Pos.T list -> T list val explode: Keyword.keywords -> Position.T -> string -> T list val explode0: Keyword.keywords -> string -> T list val print_name: Keyword.keywords -> string -> string val print_properties: Keyword.keywords -> Properties.T -> string val make: (int * int) * string -> Position.T -> T * Position.T val make_string: string * Position.T -> T val make_string0: string -> T val make_int: int -> T list val make_src: string * Position.T -> T list -> src type 'a parser = T list -> 'a * T list type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list) val syntax_generic: 'a context_parser -> src -> Context.generic -> 'a * Context.generic val syntax: 'a context_parser -> src -> Proof.context -> 'a * Proof.context end; structure Token: TOKEN = struct (** tokens **) (* token kind *) datatype kind = (*immediate source*) Command | Keyword | Ident | Long_Ident | Sym_Ident | Var | Type_Ident | Type_Var | Nat | Float | Space | (*delimited content*) String | Alt_String | Cartouche | Control of Antiquote.control | Comment of Comment.kind option | (*special content*) Error of string | EOF; val control_kind = Control Antiquote.no_control; fun equiv_kind kind kind' = (case (kind, kind') of (Control _, Control _) => true | (Error _, Error _) => true | _ => kind = kind'); val str_of_kind = fn Command => "command" | Keyword => "keyword" | Ident => "identifier" | Long_Ident => "long identifier" | Sym_Ident => "symbolic identifier" | Var => "schematic variable" | Type_Ident => "type variable" | Type_Var => "schematic type variable" | Nat => "natural number" | Float => "floating-point number" | Space => "white space" | String => "quoted string" | Alt_String => "back-quoted string" | Cartouche => "text cartouche" | Control _ => "control cartouche" | Comment NONE => "informal comment" | Comment (SOME _) => "formal comment" | Error _ => "bad input" | EOF => "end-of-input"; val immediate_kinds = Vector.fromList [Command, Keyword, Ident, Long_Ident, Sym_Ident, Var, Type_Ident, Type_Var, Nat, Float, Space]; val delimited_kind = (fn String => true | Alt_String => true | Cartouche => true | Control _ => true | Comment _ => true | _ => false); (* datatype token *) (*The value slot assigns an (optional) internal value to a token, usually as a side-effect of special scanner setup (see also args.ML). Note that an assignable ref designates an intermediate state of internalization -- it is NOT meant to persist.*) type file = {src_path: Path.T, lines: string list, digest: SHA1.digest, pos: Position.T}; type name_value = {name: string, kind: string, print: Proof.context -> Markup.T * xstring}; datatype T = Token of (Symbol_Pos.text * Position.range) * (kind * string) * slot and slot = Slot | Value of value option | Assignable of value option Unsynchronized.ref and value = Source of T list | Literal of bool * Markup.T | Name of name_value * morphism | Typ of typ | Term of term | Fact of string option * thm list | (*optional name for dynamic fact, i.e. fact "variable"*) Attribute of attribute Morphism.entity | Declaration of Morphism.declaration_entity | Files of file Exn.result list | Output of XML.body option; type src = T list; (* position *) fun pos_of (Token ((_, (pos, _)), _, _)) = pos; fun end_pos_of (Token ((_, (_, pos)), _, _)) = pos; fun adjust_offsets adjust (Token ((x, range), y, z)) = Token ((x, apply2 (Position.adjust_offsets adjust) range), y, z); (* stopper *) fun mk_eof pos = Token (("", (pos, Position.none)), (EOF, ""), Slot); val eof = mk_eof Position.none; fun is_eof (Token (_, (EOF, _), _)) = true | is_eof _ = false; val not_eof = not o is_eof; val stopper = Scan.stopper (fn [] => eof | toks => mk_eof (end_pos_of (List.last toks))) is_eof; (* kind of token *) fun kind_of (Token (_, (k, _), _)) = k; fun is_kind k (Token (_, (k', _), _)) = equiv_kind k k'; fun get_control tok = (case kind_of tok of Control control => SOME control | _ => NONE); val is_command = is_kind Command; fun keyword_with pred (Token (_, (Keyword, x), _)) = pred x | keyword_with _ _ = false; val is_command_modifier = keyword_with (fn x => x = "private" orelse x = "qualified"); fun ident_with pred (Token (_, (Ident, x), _)) = pred x | ident_with _ _ = false; fun is_ignored (Token (_, (Space, _), _)) = true | is_ignored (Token (_, (Comment NONE, _), _)) = true | is_ignored _ = false; fun is_proper (Token (_, (Space, _), _)) = false | is_proper (Token (_, (Comment _, _), _)) = false | is_proper _ = true; fun is_comment (Token (_, (Comment _, _), _)) = true | is_comment _ = false; fun is_informal_comment (Token (_, (Comment NONE, _), _)) = true | is_informal_comment _ = false; fun is_formal_comment (Token (_, (Comment (SOME _), _), _)) = true | is_formal_comment _ = false; fun is_document_marker (Token (_, (Comment (SOME Comment.Marker), _), _)) = true | is_document_marker _ = false; fun is_begin_ignore (Token (_, (Comment NONE, "<"), _)) = true | is_begin_ignore _ = false; fun is_end_ignore (Token (_, (Comment NONE, ">"), _)) = true | is_end_ignore _ = false; fun is_error (Token (_, (Error _, _), _)) = true | is_error _ = false; (* blanks and newlines -- space tokens obey lines *) fun is_space (Token (_, (Space, _), _)) = true | is_space _ = false; fun is_blank (Token (_, (Space, x), _)) = not (String.isSuffix "\n" x) | is_blank _ = false; fun is_newline (Token (_, (Space, x), _)) = String.isSuffix "\n" x | is_newline _ = false; (* range of tokens *) fun range_of (toks as tok :: _) = let val pos' = end_pos_of (List.last toks) in Position.range (pos_of tok, pos') end | range_of [] = Position.no_range; val core_range_of = drop_prefix is_ignored #> drop_suffix is_ignored #> range_of; (* token content *) fun content_of (Token (_, (_, x), _)) = x; fun source_of (Token ((source, _), _, _)) = source; fun input_of (Token ((source, range), (kind, _), _)) = Input.source (delimited_kind kind) source range; fun inner_syntax_of tok = let val x = content_of tok in if YXML.detect x then x else Syntax.implode_input (input_of tok) end; (* markup reports *) local val token_kind_markup = fn Var => (Markup.var, "") | Type_Ident => (Markup.tfree, "") | Type_Var => (Markup.tvar, "") | String => (Markup.string, "") | Alt_String => (Markup.alt_string, "") | Cartouche => (Markup.cartouche, "") | Control _ => (Markup.cartouche, "") | Comment _ => (Markup.comment, "") | Error msg => (Markup.bad (), msg) | _ => (Markup.empty, ""); fun keyword_reports tok = map (fn markup => ((pos_of tok, markup), "")); fun command_markups keywords x = if Keyword.is_theory_end keywords x then [Markup.keyword2 |> Markup.keyword_properties] else (if Keyword.is_proof_asm keywords x then [Markup.keyword3] else if Keyword.is_improper keywords x then [Markup.keyword1, Markup.improper] else [Markup.keyword1]) |> map Markup.command_properties; in fun keyword_markup (important, keyword) x = if important orelse Symbol.is_ascii_identifier x then keyword else Markup.delimiter; fun completion_report tok = if is_kind Keyword tok then map (fn m => ((pos_of tok, m), "")) (Completion.suppress_abbrevs (content_of tok)) else []; fun reports keywords tok = if is_command tok then keyword_reports tok (command_markups keywords (content_of tok)) else if is_kind Keyword tok then keyword_reports tok [keyword_markup (false, Markup.keyword2 |> Markup.keyword_properties) (content_of tok)] else let val pos = pos_of tok; val (m, text) = token_kind_markup (kind_of tok); val deleted = Symbol_Pos.explode_deleted (source_of tok, pos); in ((pos, m), text) :: map (fn p => ((p, Markup.delete), "")) deleted end; fun markups keywords = map (#2 o #1) o reports keywords; end; (* unparse *) fun unparse (Token (_, (kind, x), _)) = (case kind of String => Symbol_Pos.quote_string_qq x | Alt_String => Symbol_Pos.quote_string_bq x | Cartouche => cartouche x | Control control => Symbol_Pos.content (Antiquote.control_symbols control) | Comment NONE => enclose "(*" "*)" x | EOF => "" | _ => x); fun print tok = Markup.markups (markups Keyword.empty_keywords tok) (unparse tok); fun text_of tok = let val k = str_of_kind (kind_of tok); val ms = markups Keyword.empty_keywords tok; val s = unparse tok; in if s = "" then (k, "") else if size s < 40 andalso not (exists_string (fn c => c = "\n") s) then (k ^ " " ^ Markup.markups ms s, "") else (k, Markup.markups ms s) end; (** associated values **) (* inlined file content *) fun file_source (file: file) = let val text = cat_lines (#lines file); val end_pos = Position.symbol_explode text (#pos file); in Input.source true text (Position.range (#pos file, end_pos)) end; fun get_files (Token (_, _, Value (SOME (Files files)))) = files | get_files _ = []; fun put_files [] tok = tok | put_files files (Token (x, y, Slot)) = Token (x, y, Value (SOME (Files files))) | put_files _ tok = raise Fail ("Cannot put inlined files here" ^ Position.here (pos_of tok)); (* document output *) fun get_output (Token (_, _, Value (SOME (Output output)))) = output | get_output _ = NONE; fun put_output output (Token (x, y, Slot)) = Token (x, y, Value (SOME (Output (SOME output)))) | put_output _ tok = raise Fail ("Cannot put document output here" ^ Position.here (pos_of tok)); (* access values *) fun get_value (Token (_, _, Value v)) = v | get_value _ = NONE; fun map_value f (Token (x, y, Value (SOME v))) = Token (x, y, Value (SOME (f v))) | map_value _ tok = tok; (* reports of value *) fun get_assignable_value (Token (_, _, Assignable r)) = ! r | get_assignable_value (Token (_, _, Value v)) = v | get_assignable_value _ = NONE; fun reports_of_value tok = (case get_assignable_value tok of SOME (Literal markup) => let val pos = pos_of tok; val x = content_of tok; in if Position.is_reported pos then map (pair pos) (keyword_markup markup x :: Completion.suppress_abbrevs x) else [] end | _ => []); (* name value *) fun name_value a = Name (a, Morphism.identity); fun get_name tok = (case get_assignable_value tok of SOME (Name (a, _)) => SOME a | _ => NONE); (* maxidx *) fun declare_maxidx tok = (case get_value tok of SOME (Source src) => fold declare_maxidx src | SOME (Typ T) => Variable.declare_maxidx (Term.maxidx_of_typ T) | SOME (Term t) => Variable.declare_maxidx (Term.maxidx_of_term t) | SOME (Fact (_, ths)) => fold (Variable.declare_maxidx o Thm.maxidx_of) ths | SOME (Attribute _) => I (* FIXME !? *) | SOME (Declaration decl) => (fn ctxt => let val ctxt' = Context.proof_map (Morphism.form decl) ctxt in Variable.declare_maxidx (Variable.maxidx_of ctxt') ctxt end) | _ => I); (* fact values *) fun map_facts f = map_value (fn v => (case v of Source src => Source (map (map_facts f) src) | Fact (a, ths) => Fact (a, f a ths) | _ => v)); (* implicit context *) local fun context thm_context morphism_context attribute_context declaration_context = let fun token_context tok = map_value (fn Source src => Source (map token_context src) | Fact (a, ths) => Fact (a, map thm_context ths) | Name (a, phi) => Name (a, morphism_context phi) | Attribute a => Attribute (attribute_context a) | Declaration a => Declaration (declaration_context a) | v => v) tok; in token_context end; in val trim_context = context Thm.trim_context Morphism.reset_context Morphism.entity_reset_context Morphism.entity_reset_context; fun transfer thy = context (Thm.transfer thy) (Morphism.set_context thy) (Morphism.entity_set_context thy) (Morphism.entity_set_context thy); end; (* transform *) fun transform phi = map_value (fn v => (case v of Source src => Source (map (transform phi) src) | Literal _ => v | Name (a, psi) => Name (a, psi $> phi) | Typ T => Typ (Morphism.typ phi T) | Term t => Term (Morphism.term phi t) | Fact (a, ths) => Fact (a, Morphism.fact phi ths) | Attribute att => Attribute (Morphism.transform phi att) | Declaration decl => Declaration (Morphism.transform phi decl) | Files _ => v | Output _ => v)); (* static binding *) (*1st stage: initialize assignable slots*) fun init_assignable tok = (case tok of Token (x, y, Slot) => Token (x, y, Assignable (Unsynchronized.ref NONE)) | Token (_, _, Value _) => tok | Token (_, _, Assignable r) => (r := NONE; tok)); (*2nd stage: assign values as side-effect of scanning*) fun assign v tok = (case tok of Token (x, y, Slot) => Token (x, y, Value v) | Token (_, _, Value _) => tok | Token (_, _, Assignable r) => (r := v; tok)); fun evaluate mk eval arg = let val x = eval arg in (assign (SOME (mk x)) arg; x) end; (*3rd stage: static closure of final values*) fun closure (Token (x, y, Assignable (Unsynchronized.ref v))) = Token (x, y, Value v) | closure tok = tok; (* pretty *) +fun pretty_keyword3 tok = + let val props = Position.properties_of (pos_of tok) + in Pretty.mark_str (Markup.properties props Markup.keyword3, unparse tok) end; + fun pretty_value ctxt tok = (case get_value tok of SOME (Literal markup) => let val x = content_of tok in Pretty.mark_str (keyword_markup markup x, x) end | SOME (Name ({print, ...}, _)) => Pretty.quote (Pretty.mark_str (print ctxt)) | SOME (Typ T) => Syntax.pretty_typ ctxt T | SOME (Term t) => Syntax.pretty_term ctxt t | SOME (Fact (_, ths)) => Pretty.enclose "(" ")" (Pretty.breaks (map (Pretty.cartouche o Thm.pretty_thm ctxt) ths)) + | SOME (Attribute _) => pretty_keyword3 tok + | SOME (Declaration _) => pretty_keyword3 tok | _ => Pretty.marks_str (markups Keyword.empty_keywords tok, unparse tok)); (* src *) fun dest_src ([]: src) = raise Fail "Empty token source" | dest_src (head :: args) = (head, args); fun name_of_src src = let val head = #1 (dest_src src); val name = (case get_name head of SOME {name, ...} => name | NONE => content_of head); in (name, pos_of head) end; val args_of_src = #2 o dest_src; fun pretty_src ctxt src = let val (head, args) = dest_src src; val prt_name = (case get_name head of SOME {print, ...} => Pretty.mark_str (print ctxt) | NONE => Pretty.str (content_of head)); in Pretty.block (Pretty.breaks (Pretty.quote prt_name :: map (pretty_value ctxt) args)) end; fun checked_src (head :: _) = is_some (get_name head) | checked_src [] = true; fun check_src ctxt get_table src = let val (head, args) = dest_src src; val table = get_table ctxt; in (case get_name head of SOME {name, ...} => (src, Name_Space.get table name) | NONE => let val pos = pos_of head; val (name, x) = Name_Space.check (Context.Proof ctxt) table (content_of head, pos); val _ = Context_Position.report ctxt pos Markup.operator; val kind = Name_Space.kind_of (Name_Space.space_of_table table); fun print ctxt' = Name_Space.markup_extern ctxt' (Name_Space.space_of_table (get_table ctxt')) name; val value = name_value {name = name, kind = kind, print = print}; val head' = closure (assign (SOME value) head); in (head' :: args, x) end) end; (** scanners **) open Basic_Symbol_Pos; val err_prefix = "Outer lexical error: "; fun !!! msg = Symbol_Pos.!!! (fn () => err_prefix ^ msg); (* scan symbolic idents *) val scan_symid = Scan.many1 (Symbol.is_symbolic_char o Symbol_Pos.symbol) || Scan.one (Symbol.is_symbolic o Symbol_Pos.symbol) >> single; fun is_symid str = (case try Symbol.explode str of SOME [s] => Symbol.is_symbolic s orelse Symbol.is_symbolic_char s | SOME ss => forall Symbol.is_symbolic_char ss | _ => false); fun ident_or_symbolic "begin" = false | ident_or_symbolic ":" = true | ident_or_symbolic "::" = true | ident_or_symbolic s = Symbol_Pos.is_identifier s orelse is_symid s; (* scan cartouche *) val scan_cartouche = Symbol_Pos.scan_pos -- ((Symbol_Pos.scan_cartouche err_prefix >> Symbol_Pos.cartouche_content) -- Symbol_Pos.scan_pos); (* scan space *) fun space_symbol (s, _) = Symbol.is_blank s andalso s <> "\n"; val scan_space = Scan.many1 space_symbol @@@ Scan.optional ($$$ "\n") [] || Scan.many space_symbol @@@ $$$ "\n"; (* scan comment *) val scan_comment = Symbol_Pos.scan_pos -- (Symbol_Pos.scan_comment_body err_prefix -- Symbol_Pos.scan_pos); (** token sources **) local fun token_leq ((_, syms1), (_, syms2)) = length syms1 <= length syms2; fun token k ss = Token ((Symbol_Pos.implode ss, Symbol_Pos.range ss), (k, Symbol_Pos.content ss), Slot); fun token_range k (pos1, (ss, pos2)) = Token (Symbol_Pos.implode_range (pos1, pos2) ss, (k, Symbol_Pos.content ss), Slot); fun scan_token keywords = !!! "bad input" (Symbol_Pos.scan_string_qq err_prefix >> token_range String || Symbol_Pos.scan_string_bq err_prefix >> token_range Alt_String || scan_comment >> token_range (Comment NONE) || Comment.scan_outer >> (fn (k, ss) => token (Comment (SOME k)) ss) || scan_cartouche >> token_range Cartouche || Antiquote.scan_control err_prefix >> (fn control => token (Control control) (Antiquote.control_symbols control)) || scan_space >> token Space || (Scan.max token_leq (Scan.max token_leq (Scan.literal (Keyword.major_keywords keywords) >> pair Command) (Scan.literal (Keyword.minor_keywords keywords) >> pair Keyword)) (Lexicon.scan_longid >> pair Long_Ident || Lexicon.scan_id >> pair Ident || Lexicon.scan_var >> pair Var || Lexicon.scan_tid >> pair Type_Ident || Lexicon.scan_tvar >> pair Type_Var || Symbol_Pos.scan_float >> pair Float || Symbol_Pos.scan_nat >> pair Nat || scan_symid >> pair Sym_Ident) >> uncurry token)); fun recover msg = (Symbol_Pos.recover_string_qq || Symbol_Pos.recover_string_bq || Symbol_Pos.recover_cartouche || Symbol_Pos.recover_comment || Scan.one (Symbol.not_eof o Symbol_Pos.symbol) >> single) >> (single o token (Error msg)); in fun make_source keywords {strict} = let val scan_strict = Scan.bulk (scan_token keywords); val scan = if strict then scan_strict else Scan.recover scan_strict recover; in Source.source Symbol_Pos.stopper scan end; fun read_cartouche syms = (case Scan.read Symbol_Pos.stopper (scan_cartouche >> token_range Cartouche) syms of SOME tok => tok | NONE => error ("Single cartouche expected" ^ Position.here (#1 (Symbol_Pos.range syms)))); end; (* explode *) fun tokenize keywords strict syms = Source.of_list syms |> make_source keywords strict |> Source.exhaust; fun explode keywords pos text = Symbol_Pos.explode (text, pos) |> tokenize keywords {strict = false}; fun explode0 keywords = explode keywords Position.none; (* print names in parsable form *) fun print_name keywords name = ((case explode keywords Position.none name of [tok] => not (member (op =) [Ident, Long_Ident, Sym_Ident, Nat] (kind_of tok)) | _ => true) ? Symbol_Pos.quote_string_qq) name; fun print_properties keywords = map (apply2 (print_name keywords) #> (fn (a, b) => a ^ " = " ^ b)) #> commas #> enclose "[" "]"; (* make *) fun make ((k, n), s) pos = let val pos' = Position.shift_offsets {remove_id = false} n pos; val range = Position.range (pos, pos'); val tok = if 0 <= k andalso k < Vector.length immediate_kinds then Token ((s, range), (Vector.nth immediate_kinds k, s), Slot) else (case explode Keyword.empty_keywords pos s of [tok] => tok | _ => Token ((s, range), (Error (err_prefix ^ "exactly one token expected"), s), Slot)) in (tok, pos') end; fun make_string (s, pos) = let val Token ((x, _), y, z) = #1 (make ((~1, 0), Symbol_Pos.quote_string_qq s) Position.none); val pos' = Position.no_range_position pos; in Token ((x, (pos', pos')), y, z) end; fun make_string0 s = make_string (s, Position.none); val make_int = explode Keyword.empty_keywords Position.none o signed_string_of_int; fun make_src a args = make_string a :: args; (** parsers **) type 'a parser = T list -> 'a * T list; type 'a context_parser = Context.generic * T list -> 'a * (Context.generic * T list); (* wrapped syntax *) fun syntax_generic scan src0 context = let val src = map (transfer (Context.theory_of context)) src0; val (name, pos) = name_of_src src; val old_reports = maps reports_of_value src; val args1 = map init_assignable (args_of_src src); fun reported_text () = if Context_Position.reports_enabled_generic context then let val new_reports = maps (reports_of_value o closure) args1 in if old_reports <> new_reports then map (fn (p, m) => Position.reported_text p m "") new_reports else [] end else []; in (case Scan.error (Scan.finite' stopper (Scan.option scan)) (context, args1) of (SOME x, (context', [])) => let val _ = Output.report (reported_text ()) in (x, context') end | (_, (context', args2)) => let val print_name = (case get_name (hd src) of NONE => quote name | SOME {kind, print, ...} => let val ctxt' = Context.proof_of context'; val (markup, xname) = print ctxt'; in plain_words kind ^ " " ^ quote (Markup.markup markup xname) end); val print_args = if null args2 then "" else ":\n " ^ space_implode " " (map print args2); in error ("Bad arguments for " ^ print_name ^ Position.here pos ^ print_args ^ Markup.markup_report (implode (reported_text ()))) end) end; fun syntax scan src = apsnd Context.the_proof o syntax_generic scan src o Context.Proof; end; type 'a parser = 'a Token.parser; type 'a context_parser = 'a Token.context_parser; diff --git a/src/Pure/ex/Def.thy b/src/Pure/ex/Def.thy --- a/src/Pure/ex/Def.thy +++ b/src/Pure/ex/Def.thy @@ -1,106 +1,106 @@ (* Title: Pure/ex/Def.thy Author: Makarius Primitive constant definition, without fact definition; automatic expansion via Simplifier (simproc). *) theory Def imports Pure keywords "def" :: thy_defn begin ML \ signature DEF = sig val get_def: Proof.context -> cterm -> thm option val def: (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term -> local_theory -> term * local_theory val def_cmd: (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string -> local_theory -> term * local_theory end; structure Def: DEF = struct (* context data *) type def = {lhs: term, eq: thm}; val eq_def : def * def -> bool = op aconv o apply2 #lhs; fun transform_def phi ({lhs, eq}: def) = {lhs = Morphism.term phi lhs, eq = Morphism.thm phi eq}; structure Data = Generic_Data ( type T = def Item_Net.T; val empty : T = Item_Net.init eq_def (single o #lhs); val merge = Item_Net.merge; ); fun declare_def lhs eq lthy = let val def0: def = {lhs = lhs, eq = Thm.trim_context eq} in - lthy |> Local_Theory.declaration {syntax = false, pervasive = true} + lthy |> Local_Theory.declaration {syntax = false, pervasive = true, pos = \<^here>} (fn phi => fn context => let val psi = Morphism.set_trim_context'' context phi in (Data.map o Item_Net.update) (transform_def psi def0) context end) end; fun get_def ctxt ct = let val thy = Proof_Context.theory_of ctxt; val data = Data.get (Context.Proof ctxt); val t = Thm.term_of ct; fun match_def {lhs, eq} = if Pattern.matches thy (lhs, t) then let val inst = Thm.match (Thm.cterm_of ctxt lhs, ct) in SOME (Thm.instantiate inst (Thm.transfer thy eq)) end else NONE; in Item_Net.retrieve_matching data t |> get_first match_def end; (* simproc setup *) val _ = (Theory.setup o Named_Target.theory_map) (Simplifier.define_simproc \<^binding>\expand_def\ {lhss = [Free ("x", TFree ("'a", []))], proc = K get_def}); (* Isar command *) fun gen_def prep_spec raw_var raw_params raw_spec lthy = let val ((vars, xs, get_pos, spec), _) = lthy |> prep_spec (the_list raw_var) raw_params [] raw_spec; val (((x, _), rhs), prove) = Local_Defs.derived_def lthy get_pos {conditional = false} spec; val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val ((lhs, (_, eq)), lthy') = lthy |> Local_Theory.define_internal ((b, mx), (Binding.empty_atts, rhs)); (*sanity check for original specification*) val _: thm = prove lthy' eq; in (lhs, declare_def lhs eq lthy') end; val def = gen_def Specification.check_spec_open; val def_cmd = gen_def Specification.read_spec_open; val _ = Outer_Syntax.local_theory \<^command_keyword>\def\ "primitive constant definition, without fact definition" (Scan.option Parse_Spec.constdecl -- Parse.prop -- Parse.for_fixes >> (fn ((decl, spec), params) => #2 o def_cmd decl params spec)); end; \ end diff --git a/src/Pure/simplifier.ML b/src/Pure/simplifier.ML --- a/src/Pure/simplifier.ML +++ b/src/Pure/simplifier.ML @@ -1,433 +1,434 @@ (* Title: Pure/simplifier.ML Author: Tobias Nipkow and Markus Wenzel, TU Muenchen Generic simplifier, suitable for most logics (see also raw_simplifier.ML for the actual meta-level rewriting engine). *) signature BASIC_SIMPLIFIER = sig include BASIC_RAW_SIMPLIFIER val simp_tac: Proof.context -> int -> tactic val asm_simp_tac: Proof.context -> int -> tactic val full_simp_tac: Proof.context -> int -> tactic val asm_lr_simp_tac: Proof.context -> int -> tactic val asm_full_simp_tac: Proof.context -> int -> tactic val safe_simp_tac: Proof.context -> int -> tactic val safe_asm_simp_tac: Proof.context -> int -> tactic val safe_full_simp_tac: Proof.context -> int -> tactic val safe_asm_lr_simp_tac: Proof.context -> int -> tactic val safe_asm_full_simp_tac: Proof.context -> int -> tactic val simplify: Proof.context -> thm -> thm val asm_simplify: Proof.context -> thm -> thm val full_simplify: Proof.context -> thm -> thm val asm_lr_simplify: Proof.context -> thm -> thm val asm_full_simplify: Proof.context -> thm -> thm end; signature SIMPLIFIER = sig include BASIC_SIMPLIFIER val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val attrib: (thm -> Proof.context -> Proof.context) -> attribute val simp_add: attribute val simp_del: attribute val simp_flip: attribute val cong_add: attribute val cong_del: attribute val check_simproc: Proof.context -> xstring * Position.T -> string val the_simproc: Proof.context -> string -> simproc type 'a simproc_spec = {lhss: 'a list, proc: morphism -> Proof.context -> cterm -> thm option} val make_simproc: Proof.context -> string -> term simproc_spec -> simproc val define_simproc: binding -> term simproc_spec -> local_theory -> local_theory val define_simproc_cmd: binding -> string simproc_spec -> local_theory -> local_theory val pretty_simpset: bool -> Proof.context -> Pretty.T val default_mk_sym: Proof.context -> thm -> thm option val prems_of: Proof.context -> thm list val add_simp: thm -> Proof.context -> Proof.context val del_simp: thm -> Proof.context -> Proof.context val init_simpset: thm list -> Proof.context -> Proof.context val add_eqcong: thm -> Proof.context -> Proof.context val del_eqcong: thm -> Proof.context -> Proof.context val add_cong: thm -> Proof.context -> Proof.context val del_cong: thm -> Proof.context -> Proof.context val add_prems: thm list -> Proof.context -> Proof.context val mksimps: Proof.context -> thm -> thm list val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context val set_term_ord: term ord -> Proof.context -> Proof.context val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context type trace_ops val set_trace_ops: trace_ops -> theory -> theory val rewrite: Proof.context -> conv val asm_rewrite: Proof.context -> conv val full_rewrite: Proof.context -> conv val asm_lr_rewrite: Proof.context -> conv val asm_full_rewrite: Proof.context -> conv val cong_modifiers: Method.modifier parser list val simp_modifiers': Method.modifier parser list val simp_modifiers: Method.modifier parser list val method_setup: Method.modifier parser list -> theory -> theory val unsafe_solver_tac: Proof.context -> int -> tactic val unsafe_solver: solver val safe_solver_tac: Proof.context -> int -> tactic val safe_solver: solver end; structure Simplifier: SIMPLIFIER = struct open Raw_Simplifier; (** declarations **) (* attributes *) fun attrib f = Thm.declaration_attribute (map_ss o f); val simp_add = attrib add_simp; val simp_del = attrib del_simp; val simp_flip = attrib flip_simp; val cong_add = attrib add_cong; val cong_del = attrib del_cong; (** named simprocs **) structure Simprocs = Generic_Data ( type T = simproc Name_Space.table; val empty : T = Name_Space.empty_table "simproc"; fun merge data : T = Name_Space.merge_tables data; ); (* get simprocs *) val get_simprocs = Simprocs.get o Context.Proof; fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1; val the_simproc = Name_Space.get o get_simprocs; val _ = Theory.setup (ML_Antiquotation.value_embedded \<^binding>\simproc\ (Args.context -- Scan.lift Parse.embedded_position >> (fn (ctxt, name) => "Simplifier.the_simproc ML_context " ^ ML_Syntax.print_string (check_simproc ctxt name)))); (* define simprocs *) type 'a simproc_spec = {lhss: 'a list, proc: morphism -> Proof.context -> cterm -> thm option}; fun make_simproc ctxt name {lhss, proc} = let val ctxt' = fold Proof_Context.augment lhss ctxt; val lhss' = Variable.export_terms ctxt' ctxt lhss; in cert_simproc (Proof_Context.theory_of ctxt) name {lhss = lhss', proc = Morphism.entity proc} end; local fun def_simproc prep b {lhss, proc} lthy = let val simproc = make_simproc lthy (Local_Theory.full_name lthy b) {lhss = prep lthy lhss, proc = proc}; in - lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => - let - val b' = Morphism.binding phi b; - val simproc' = transform_simproc phi simproc; - in - context - |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) - |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) - end) + lthy |> Local_Theory.declaration {syntax = false, pervasive = false, pos = Binding.pos_of b} + (fn phi => fn context => + let + val b' = Morphism.binding phi b; + val simproc' = transform_simproc phi simproc; + in + context + |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) + |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) + end) end; in val define_simproc = def_simproc Syntax.check_terms; val define_simproc_cmd = def_simproc Syntax.read_terms; end; (** congruence rule to protect foundational terms of local definitions **) local fun add_foundation_cong (binding, (const, target_params)) gthy = if null target_params then gthy else let val thy = Context.theory_of gthy; val cong = list_comb (const, target_params) |> Logic.varify_global |> Thm.global_cterm_of thy |> Thm.reflexive |> Thm.close_derivation \<^here>; val cong_binding = Binding.qualify_name true binding "cong"; in gthy |> Attrib.generic_notes Thm.theoremK [((cong_binding, []), [([cong], [])])] |> #2 end; val _ = Theory.setup (Generic_Target.add_foundation_interpretation add_foundation_cong); in end; (** pretty_simpset **) fun pretty_simpset verbose ctxt = let val pretty_term = Syntax.pretty_term ctxt; val pretty_thm = Thm.pretty_thm ctxt; val pretty_thm_item = Thm.pretty_thm_item ctxt; fun pretty_simproc (name, lhss) = Pretty.block (Pretty.mark_str name :: Pretty.str ":" :: Pretty.fbrk :: Pretty.fbreaks (map (Pretty.item o single o pretty_term) lhss)); fun pretty_cong_name (const, name) = pretty_term ((if const then Const else Free) (name, dummyT)); fun pretty_cong (name, thm) = Pretty.block [pretty_cong_name name, Pretty.str ":", Pretty.brk 1, pretty_thm thm]; val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss (simpset_of ctxt); val simprocs = Name_Space.markup_entries verbose ctxt (Name_Space.space_of_table (get_simprocs ctxt)) procs; in [Pretty.big_list "simplification rules:" (map (pretty_thm_item o #2) simps), Pretty.big_list "simplification procedures:" (map pretty_simproc simprocs), Pretty.big_list "congruences:" (map pretty_cong congs), Pretty.strs ("loopers:" :: map quote loopers), Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers), Pretty.strs ("safe solvers:" :: map quote safe_solvers)] |> Pretty.chunks end; (** simplification tactics and rules **) fun solve_all_tac solvers ctxt = let val subgoal_tac = Raw_Simplifier.subgoal_tac (Raw_Simplifier.set_solvers solvers ctxt); val solve_tac = subgoal_tac THEN_ALL_NEW (K no_tac); in DEPTH_SOLVE (solve_tac 1) end; (*NOTE: may instantiate unknowns that appear also in other subgoals*) fun generic_simp_tac safe mode ctxt = let val loop_tac = Raw_Simplifier.loop_tac ctxt; val (unsafe_solvers, solvers) = Raw_Simplifier.solvers ctxt; val solve_tac = FIRST' (map (Raw_Simplifier.solver ctxt) (rev (if safe then solvers else unsafe_solvers))); fun simp_loop_tac i = Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ctxt i THEN (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); in PREFER_GOAL (simp_loop_tac 1) end; local fun simp rew mode ctxt thm = let val (unsafe_solvers, _) = Raw_Simplifier.solvers ctxt; val tacf = solve_all_tac (rev unsafe_solvers); fun prover s th = Option.map #1 (Seq.pull (tacf s th)); in rew mode prover ctxt thm end; in val simp_thm = simp Raw_Simplifier.rewrite_thm; val simp_cterm = simp Raw_Simplifier.rewrite_cterm; end; (* tactics *) val simp_tac = generic_simp_tac false (false, false, false); val asm_simp_tac = generic_simp_tac false (false, true, false); val full_simp_tac = generic_simp_tac false (true, false, false); val asm_lr_simp_tac = generic_simp_tac false (true, true, false); val asm_full_simp_tac = generic_simp_tac false (true, true, true); (*not totally safe: may instantiate unknowns that appear also in other subgoals*) val safe_simp_tac = generic_simp_tac true (false, false, false); val safe_asm_simp_tac = generic_simp_tac true (false, true, false); val safe_full_simp_tac = generic_simp_tac true (true, false, false); val safe_asm_lr_simp_tac = generic_simp_tac true (true, true, false); val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); (* conversions *) val simplify = simp_thm (false, false, false); val asm_simplify = simp_thm (false, true, false); val full_simplify = simp_thm (true, false, false); val asm_lr_simplify = simp_thm (true, true, false); val asm_full_simplify = simp_thm (true, true, true); val rewrite = simp_cterm (false, false, false); val asm_rewrite = simp_cterm (false, true, false); val full_rewrite = simp_cterm (true, false, false); val asm_lr_rewrite = simp_cterm (true, true, false); val asm_full_rewrite = simp_cterm (true, true, true); (** concrete syntax of attributes **) (* add / del *) val simpN = "simp"; val flipN = "flip" val congN = "cong"; val onlyN = "only"; val no_asmN = "no_asm"; val no_asm_useN = "no_asm_use"; val no_asm_simpN = "no_asm_simp"; val asm_lrN = "asm_lr"; (* simprocs *) local val add_del = (Args.del -- Args.colon >> K (op delsimprocs) || Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) >> (fn f => fn simproc => Morphism.entity (fn phi => Thm.declaration_attribute (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc])))))); in val simproc_att = (Args.context -- Scan.lift add_del) :|-- (fn (ctxt, decl) => Scan.repeat1 (Scan.lift (Args.named_attribute (decl o the_simproc ctxt o check_simproc ctxt)))) >> (fn atts => Thm.declaration_attribute (fn th => fold (fn att => Thm.attribute_declaration (Morphism.form att) th) atts)); end; (* conversions *) local fun conv_mode x = ((Args.parens (Args.$$$ no_asmN) >> K simplify || Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || Scan.succeed asm_full_simplify) |> Scan.lift) x; in val simplified = conv_mode -- Attrib.thms >> (fn (f, ths) => Thm.rule_attribute ths (fn context => f ((if null ths then I else Raw_Simplifier.clear_simpset) (Context.proof_of context) addsimps ths))); end; (* setup attributes *) val _ = Theory.setup (Attrib.setup \<^binding>\simp\ (Attrib.add_del simp_add simp_del) "declaration of Simplifier rewrite rule" #> Attrib.setup \<^binding>\cong\ (Attrib.add_del cong_add cong_del) "declaration of Simplifier congruence rule" #> Attrib.setup \<^binding>\simproc\ simproc_att "declaration of simplification procedures" #> Attrib.setup \<^binding>\simplified\ simplified "simplified rule"); (** method syntax **) val cong_modifiers = [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>)]; val simp_modifiers = [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), Args.$$$ simpN -- Args.$$$ flipN -- Args.colon >> K (Method.modifier simp_flip \<^here>), Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}] @ cong_modifiers; val simp_modifiers' = [Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), Args.$$$ flipN -- Args.colon >> K (Method.modifier simp_flip \<^here>), Args.$$$ onlyN -- Args.colon >> K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = \<^here>}] @ cong_modifiers; val simp_options = (Args.parens (Args.$$$ no_asmN) >> K simp_tac || Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || Scan.succeed asm_full_simp_tac); fun simp_method more_mods meth = Scan.lift simp_options --| Method.sections (more_mods @ simp_modifiers') >> (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts)); (** setup **) fun method_setup more_mods = Method.setup \<^binding>\simp\ (simp_method more_mods (fn ctxt => fn tac => fn facts => HEADGOAL (Method.insert_tac ctxt facts THEN' (CHANGED_PROP oo tac) ctxt))) "simplification" #> Method.setup \<^binding>\simp_all\ (simp_method more_mods (fn ctxt => fn tac => fn facts => ALLGOALS (Method.insert_tac ctxt facts) THEN (CHANGED_PROP o PARALLEL_ALLGOALS o tac) ctxt)) "simplification (all goals)"; fun unsafe_solver_tac ctxt = FIRST' [resolve_tac ctxt (Drule.reflexive_thm :: Raw_Simplifier.prems_of ctxt), assume_tac ctxt]; val unsafe_solver = mk_solver "Pure unsafe" unsafe_solver_tac; (*no premature instantiation of variables during simplification*) fun safe_solver_tac ctxt = FIRST' [match_tac ctxt (Drule.reflexive_thm :: Raw_Simplifier.prems_of ctxt), eq_assume_tac]; val safe_solver = mk_solver "Pure safe" safe_solver_tac; val _ = Theory.setup (method_setup [] #> Context.theory_map (map_ss (fn ctxt => empty_simpset ctxt setSSolver safe_solver setSolver unsafe_solver |> set_subgoaler asm_simp_tac))); end; structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier; open Basic_Simplifier; diff --git a/src/ZF/OrdQuant.thy b/src/ZF/OrdQuant.thy --- a/src/ZF/OrdQuant.thy +++ b/src/ZF/OrdQuant.thy @@ -1,362 +1,360 @@ (* Title: ZF/OrdQuant.thy Authors: Krzysztof Grabczewski and L C Paulson *) section \Special quantifiers\ theory OrdQuant imports Ordinal begin subsection \Quantifiers and union operator for ordinals\ definition (* Ordinal Quantifiers *) oall :: "[i, i \ o] \ o" where "oall(A, P) \ \x. x P(x)" definition oex :: "[i, i \ o] \ o" where "oex(A, P) \ \x. x P(x)" definition (* Ordinal Union *) OUnion :: "[i, i \ i] \ i" where "OUnion(i,B) \ {z: \x\i. B(x). Ord(i)}" syntax "_oall" :: "[idt, i, o] \ o" (\(3\_<_./ _)\ 10) "_oex" :: "[idt, i, o] \ o" (\(3\_<_./ _)\ 10) "_OUNION" :: "[idt, i, i] \ i" (\(3\_<_./ _)\ 10) translations "\x "CONST oall(a, \x. P)" "\x "CONST oex(a, \x. P)" "\x "CONST OUnion(a, \x. B)" subsubsection \simplification of the new quantifiers\ (*MOST IMPORTANT that this is added to the simpset BEFORE Ord_atomize is proved. Ord_atomize would convert this rule to x < 0 \ P(x) \ True, which causes dire effects!*) lemma [simp]: "(\x<0. P(x))" by (simp add: oall_def) lemma [simp]: "\(\x<0. P(x))" by (simp add: oex_def) lemma [simp]: "(\x (Ord(i) \ P(i) \ (\xx (Ord(i) \ (P(i) | (\xUnion over ordinals\ lemma Ord_OUN [intro,simp]: "\\x. x Ord(B(x))\ \ Ord(\xax \ i < (\xab(a); Ord(\x \ i \ (\x (\xi\nat.i)=nat"}! *) lemma OUN_least: "(\x. x B(x) \ C) \ (\x C" by (simp add: OUnion_def UN_least ltI) lemma OUN_least_le: "\Ord(i); \x. x b(x) \ i\ \ (\x i" by (simp add: OUnion_def UN_least_le ltI Ord_0_le) lemma le_implies_OUN_le_OUN: "\\x. x c(x) \ d(x)\ \ (\x (\xx. x \ A \ Ord(B(x))) \ (\z < (\x\A. B(x)). C(z)) = (\x\A. \z < B(x). C(z))" by (simp add: OUnion_def) lemma OUN_Union_eq: "(\x. x \ X \ Ord(x)) \ (\z < \(X). C(z)) = (\x\X. \z < x. C(z))" by (simp add: OUnion_def) (*So that rule_format will get rid of this quantifier...*) lemma atomize_oall [symmetric, rulify]: "(\x. x P(x)) \ Trueprop (\xuniversal quantifier for ordinals\ lemma oallI [intro!]: "\\x. x P(x)\ \ \x\x \ P(x)" by (simp add: oall_def) lemma oallE: "\\x Q; \x Q\ \ Q" by (simp add: oall_def, blast) lemma rev_oallE [elim]: "\\xx Q; P(x) \ Q\ \ Q" by (simp add: oall_def, blast) (*Trival rewrite rule. @{term"(\xP"} holds only if a is not 0!*) lemma oall_simp [simp]: "(\x True" by blast (*Congruence rule for rewriting*) lemma oall_cong [cong]: "\a=a'; \x. x P(x) <-> P'(x)\ \ oall(a, \x. P(x)) <-> oall(a', \x. P'(x))" by (simp add: oall_def) subsubsection \existential quantifier for ordinals\ lemma oexI [intro]: "\P(x); x \ \x\xP(x) \ P(a); a \ \x\xx. \x \ Q\ \ Q" apply (simp add: oex_def, blast) done lemma oex_cong [cong]: "\a=a'; \x. x P(x) <-> P'(x)\ \ oex(a, \x. P(x)) <-> oex(a', \x. P'(x))" apply (simp add: oex_def cong add: conj_cong) done subsubsection \Rules for Ordinal-Indexed Unions\ lemma OUN_I [intro]: "\a B(a)\ \ b: (\zb \ (\za.\b \ B(a); a \ R\ \ R" apply (unfold OUnion_def lt_def, blast) done lemma OUN_iff: "b \ (\x (\x B(x))" by (unfold OUnion_def oex_def lt_def, blast) lemma OUN_cong [cong]: "\i=j; \x. x C(x)=D(x)\ \ (\xxix.\xy \ P(x)\ \ P(i)" apply (simp add: lt_def oall_def) apply (erule conjE) apply (erule Ord_induct, assumption, blast) done subsection \Quantification over a class\ definition "rall" :: "[i\o, i\o] \ o" where "rall(M, P) \ \x. M(x) \ P(x)" definition "rex" :: "[i\o, i\o] \ o" where "rex(M, P) \ \x. M(x) \ P(x)" syntax "_rall" :: "[pttrn, i\o, o] \ o" (\(3\_[_]./ _)\ 10) "_rex" :: "[pttrn, i\o, o] \ o" (\(3\_[_]./ _)\ 10) translations "\x[M]. P" \ "CONST rall(M, \x. P)" "\x[M]. P" \ "CONST rex(M, \x. P)" subsubsection\Relativized universal quantifier\ lemma rallI [intro!]: "\\x. M(x) \ P(x)\ \ \x[M]. P(x)" by (simp add: rall_def) lemma rspec: "\\x[M]. P(x); M(x)\ \ P(x)" by (simp add: rall_def) (*Instantiates x first: better for automatic theorem proving?*) lemma rev_rallE [elim]: "\\x[M]. P(x); \ M(x) \ Q; P(x) \ Q\ \ Q" by (simp add: rall_def, blast) lemma rallE: "\\x[M]. P(x); P(x) \ Q; \ M(x) \ Q\ \ Q" by blast (*Trival rewrite rule; (\x[M].P)<->P holds only if A is nonempty!*) lemma rall_triv [simp]: "(\x[M]. P) \ ((\x. M(x)) \ P)" by (simp add: rall_def) (*Congruence rule for rewriting*) lemma rall_cong [cong]: "(\x. M(x) \ P(x) <-> P'(x)) \ (\x[M]. P(x)) <-> (\x[M]. P'(x))" by (simp add: rall_def) subsubsection\Relativized existential quantifier\ lemma rexI [intro]: "\P(x); M(x)\ \ \x[M]. P(x)" by (simp add: rex_def, blast) (*The best argument order when there is only one M(x)*) lemma rev_rexI: "\M(x); P(x)\ \ \x[M]. P(x)" by blast (*Not of the general form for such rules... *) lemma rexCI: "\\x[M]. \P(x) \ P(a); M(a)\ \ \x[M]. P(x)" by blast lemma rexE [elim!]: "\\x[M]. P(x); \x. \M(x); P(x)\ \ Q\ \ Q" by (simp add: rex_def, blast) (*We do not even have (\x[M]. True) <-> True unless A is nonempty\*) lemma rex_triv [simp]: "(\x[M]. P) \ ((\x. M(x)) \ P)" by (simp add: rex_def) lemma rex_cong [cong]: "(\x. M(x) \ P(x) <-> P'(x)) \ (\x[M]. P(x)) <-> (\x[M]. P'(x))" by (simp add: rex_def cong: conj_cong) lemma rall_is_ball [simp]: "(\x[\z. z\A]. P(x)) <-> (\x\A. P(x))" by blast lemma rex_is_bex [simp]: "(\x[\z. z\A]. P(x)) <-> (\x\A. P(x))" by blast lemma atomize_rall: "(\x. M(x) \ P(x)) \ Trueprop (\x[M]. P(x))" by (simp add: rall_def atomize_all atomize_imp) declare atomize_rall [symmetric, rulify] lemma rall_simps1: "(\x[M]. P(x) \ Q) <-> (\x[M]. P(x)) \ ((\x[M]. False) | Q)" "(\x[M]. P(x) | Q) <-> ((\x[M]. P(x)) | Q)" "(\x[M]. P(x) \ Q) <-> ((\x[M]. P(x)) \ Q)" "(\(\x[M]. P(x))) <-> (\x[M]. \P(x))" by blast+ lemma rall_simps2: "(\x[M]. P \ Q(x)) <-> ((\x[M]. False) | P) \ (\x[M]. Q(x))" "(\x[M]. P | Q(x)) <-> (P | (\x[M]. Q(x)))" "(\x[M]. P \ Q(x)) <-> (P \ (\x[M]. Q(x)))" by blast+ lemmas rall_simps [simp] = rall_simps1 rall_simps2 lemma rall_conj_distrib: "(\x[M]. P(x) \ Q(x)) <-> ((\x[M]. P(x)) \ (\x[M]. Q(x)))" by blast lemma rex_simps1: "(\x[M]. P(x) \ Q) <-> ((\x[M]. P(x)) \ Q)" "(\x[M]. P(x) | Q) <-> (\x[M]. P(x)) | ((\x[M]. True) \ Q)" "(\x[M]. P(x) \ Q) <-> ((\x[M]. P(x)) \ ((\x[M]. True) \ Q))" "(\(\x[M]. P(x))) <-> (\x[M]. \P(x))" by blast+ lemma rex_simps2: "(\x[M]. P \ Q(x)) <-> (P \ (\x[M]. Q(x)))" "(\x[M]. P | Q(x)) <-> ((\x[M]. True) \ P) | (\x[M]. Q(x))" "(\x[M]. P \ Q(x)) <-> (((\x[M]. False) | P) \ (\x[M]. Q(x)))" by blast+ lemmas rex_simps [simp] = rex_simps1 rex_simps2 lemma rex_disj_distrib: "(\x[M]. P(x) | Q(x)) <-> ((\x[M]. P(x)) | (\x[M]. Q(x)))" by blast subsubsection\One-point rule for bounded quantifiers\ lemma rex_triv_one_point1 [simp]: "(\x[M]. x=a) <-> ( M(a))" by blast lemma rex_triv_one_point2 [simp]: "(\x[M]. a=x) <-> ( M(a))" by blast lemma rex_one_point1 [simp]: "(\x[M]. x=a \ P(x)) <-> ( M(a) \ P(a))" by blast lemma rex_one_point2 [simp]: "(\x[M]. a=x \ P(x)) <-> ( M(a) \ P(a))" by blast lemma rall_one_point1 [simp]: "(\x[M]. x=a \ P(x)) <-> ( M(a) \ P(a))" by blast lemma rall_one_point2 [simp]: "(\x[M]. a=x \ P(x)) <-> ( M(a) \ P(a))" by blast subsubsection\Sets as Classes\ definition setclass :: "[i,i] \ o" (\##_\ [40] 40) where "setclass(A) \ \x. x \ A" lemma setclass_iff [simp]: "setclass(A,x) <-> x \ A" by (simp add: setclass_def) lemma rall_setclass_is_ball [simp]: "(\x[##A]. P(x)) <-> (\x\A. P(x))" by auto lemma rex_setclass_is_bex [simp]: "(\x[##A]. P(x)) <-> (\x\A. P(x))" by auto ML \ val Ord_atomize = atomize ([(\<^const_name>\oall\, @{thms ospec}), (\<^const_name>\rall\, @{thms rspec})] @ ZF_conn_pairs, ZF_mem_pairs); \ declaration \fn _ => Simplifier.map_ss (Simplifier.set_mksimps (fn ctxt => map mk_eq o Ord_atomize o Variable.gen_all ctxt)) \ text \Setting up the one-point-rule simproc\ simproc_setup defined_rex ("\x[M]. P(x) \ Q(x)") = \ - fn _ => Quantifier1.rearrange_Bex - (fn ctxt => unfold_tac ctxt @{thms rex_def}) + K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms rex_def})) \ simproc_setup defined_rall ("\x[M]. P(x) \ Q(x)") = \ - fn _ => Quantifier1.rearrange_Ball - (fn ctxt => unfold_tac ctxt @{thms rall_def}) + K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms rall_def})) \ end diff --git a/src/ZF/pair.thy b/src/ZF/pair.thy --- a/src/ZF/pair.thy +++ b/src/ZF/pair.thy @@ -1,185 +1,183 @@ (* Title: ZF/pair.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) section\Ordered Pairs\ theory pair imports upair begin ML_file \simpdata.ML\ setup \ map_theory_simpset (Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt) #> Simplifier.add_cong @{thm if_weak_cong}) \ ML \val ZF_ss = simpset_of \<^context>\ simproc_setup defined_Bex ("\x\A. P(x) \ Q(x)") = \ - fn _ => Quantifier1.rearrange_Bex - (fn ctxt => unfold_tac ctxt @{thms Bex_def}) + K (Quantifier1.rearrange_Bex (fn ctxt => unfold_tac ctxt @{thms Bex_def})) \ simproc_setup defined_Ball ("\x\A. P(x) \ Q(x)") = \ - fn _ => Quantifier1.rearrange_Ball - (fn ctxt => unfold_tac ctxt @{thms Ball_def}) + K (Quantifier1.rearrange_Ball (fn ctxt => unfold_tac ctxt @{thms Ball_def})) \ (** Lemmas for showing that \a,b\ uniquely determines a and b **) lemma singleton_eq_iff [iff]: "{a} = {b} \ a=b" by (rule extension [THEN iff_trans], blast) lemma doubleton_eq_iff: "{a,b} = {c,d} \ (a=c \ b=d) | (a=d \ b=c)" by (rule extension [THEN iff_trans], blast) lemma Pair_iff [simp]: "\a,b\ = \c,d\ \ a=c \ b=d" by (simp add: Pair_def doubleton_eq_iff, blast) lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!] lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1] lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2] lemma Pair_not_0: "\a,b\ \ 0" unfolding Pair_def apply (blast elim: equalityE) done lemmas Pair_neq_0 = Pair_not_0 [THEN notE, elim!] declare sym [THEN Pair_neq_0, elim!] lemma Pair_neq_fst: "\a,b\=a \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = a" have "{a, a} \ {{a, a}, {a, b}}" by (rule consI1) hence "{a, a} \ a" by (simp add: eq) moreover have "a \ {a, a}" by (rule consI1) ultimately show "P" by (rule mem_asym) qed lemma Pair_neq_snd: "\a,b\=b \ P" proof (unfold Pair_def) assume eq: "{{a, a}, {a, b}} = b" have "{a, b} \ {{a, a}, {a, b}}" by blast hence "{a, b} \ b" by (simp add: eq) moreover have "b \ {a, b}" by blast ultimately show "P" by (rule mem_asym) qed subsection\Sigma: Disjoint Union of a Family of Sets\ text\Generalizes Cartesian product\ lemma Sigma_iff [simp]: "\a,b\: Sigma(A,B) \ a \ A \ b \ B(a)" by (simp add: Sigma_def) lemma SigmaI [TC,intro!]: "\a \ A; b \ B(a)\ \ \a,b\ \ Sigma(A,B)" by simp lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1] lemmas SigmaD2 = Sigma_iff [THEN iffD1, THEN conjunct2] (*The general elimination rule*) lemma SigmaE [elim!]: "\c \ Sigma(A,B); \x y.\x \ A; y \ B(x); c=\x,y\\ \ P \ \ P" by (unfold Sigma_def, blast) lemma SigmaE2 [elim!]: "\\a,b\ \ Sigma(A,B); \a \ A; b \ B(a)\ \ P \ \ P" by (unfold Sigma_def, blast) lemma Sigma_cong: "\A=A'; \x. x \ A' \ B(x)=B'(x)\ \ Sigma(A,B) = Sigma(A',B')" by (simp add: Sigma_def) (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause flex-flex pairs and the "Check your prover" error. Most Sigmas and Pis are abbreviated as * or -> *) lemma Sigma_empty1 [simp]: "Sigma(0,B) = 0" by blast lemma Sigma_empty2 [simp]: "A*0 = 0" by blast lemma Sigma_empty_iff: "A*B=0 \ A=0 | B=0" by blast subsection\Projections \<^term>\fst\ and \<^term>\snd\\ lemma fst_conv [simp]: "fst(\a,b\) = a" by (simp add: fst_def) lemma snd_conv [simp]: "snd(\a,b\) = b" by (simp add: snd_def) lemma fst_type [TC]: "p \ Sigma(A,B) \ fst(p) \ A" by auto lemma snd_type [TC]: "p \ Sigma(A,B) \ snd(p) \ B(fst(p))" by auto lemma Pair_fst_snd_eq: "a \ Sigma(A,B) \ = a" by auto subsection\The Eliminator, \<^term>\split\\ (*A META-equality, so that it applies to higher types as well...*) lemma split [simp]: "split(\x y. c(x,y), \a,b\) \ c(a,b)" by (simp add: split_def) lemma split_type [TC]: "\p \ Sigma(A,B); \x y.\x \ A; y \ B(x)\ \ c(x,y):C(\x,y\) \ \ split(\x y. c(x,y), p) \ C(p)" by (erule SigmaE, auto) lemma expand_split: "u \ A*B \ R(split(c,u)) \ (\x\A. \y\B. u = \x,y\ \ R(c(x,y)))" by (auto simp add: split_def) subsection\A version of \<^term>\split\ for Formulae: Result Type \<^typ>\o\\ lemma splitI: "R(a,b) \ split(R, \a,b\)" by (simp add: split_def) lemma splitE: "\split(R,z); z \ Sigma(A,B); \x y. \z = \x,y\; R(x,y)\ \ P \ \ P" by (auto simp add: split_def) lemma splitD: "split(R,\a,b\) \ R(a,b)" by (simp add: split_def) text \ \bigskip Complex rules for Sigma. \ lemma split_paired_Bex_Sigma [simp]: "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P(\x,y\))" by blast lemma split_paired_Ball_Sigma [simp]: "(\z \ Sigma(A,B). P(z)) \ (\x \ A. \y \ B(x). P(\x,y\))" by blast end