diff --git a/src/Doc/Isar_Ref/Spec.thy b/src/Doc/Isar_Ref/Spec.thy --- a/src/Doc/Isar_Ref/Spec.thy +++ b/src/Doc/Isar_Ref/Spec.thy @@ -1,1548 +1,1548 @@ (*:maxLineLen=78:*) theory Spec imports Main Base begin chapter \Specifications\ text \ The Isabelle/Isar theory format integrates specifications and proofs, with support for interactive development by continuous document editing. There is a separate document preparation system (see \chref{ch:document-prep}), for typesetting formal developments together with informal text. The resulting hyper-linked PDF documents can be used both for WWW presentation and printed copies. The Isar proof language (see \chref{ch:proofs}) is embedded into the theory language as a proper sub-language. Proof mode is entered by stating some \<^theory_text>\theorem\ or \<^theory_text>\lemma\ at the theory level, and left again with the final conclusion (e.g.\ via \<^theory_text>\qed\). \ section \Defining theories \label{sec:begin-thy}\ text \ \begin{matharray}{rcl} @{command_def "theory"} & : & \toplevel \ theory\ \\ @{command_def (global) "end"} & : & \theory \ toplevel\ \\ @{command_def "thy_deps"}\\<^sup>*\ & : & \theory \\ \\ \end{matharray} Isabelle/Isar theories are defined via theory files, which consist of an outermost sequence of definition--statement--proof elements. Some definitions are self-sufficient (e.g.\ \<^theory_text>\fun\ in Isabelle/HOL), with foundational proofs performed internally. Other definitions require an explicit proof as justification (e.g.\ \<^theory_text>\function\ and \<^theory_text>\termination\ in Isabelle/HOL). Plain statements like \<^theory_text>\theorem\ or \<^theory_text>\lemma\ are merely a special case of that, defining a theorem from a given proposition and its proof. The theory body may be sub-structured by means of \<^emph>\local theory targets\, such as \<^theory_text>\locale\ and \<^theory_text>\class\. It is also possible to use \<^theory_text>\context begin \ end\ blocks to delimited a local theory context: a \<^emph>\named context\ to augment a locale or class specification, or an \<^emph>\unnamed context\ to refer to local parameters and assumptions that are discharged later. See \secref{sec:target} for more details. \<^medskip> A theory is commenced by the \<^theory_text>\theory\ command, which indicates imports of previous theories, according to an acyclic foundational order. Before the initial \<^theory_text>\theory\ command, there may be optional document header material (like \<^theory_text>\section\ or \<^theory_text>\text\, see \secref{sec:markup}). The document header is outside of the formal theory context, though. A theory is concluded by a final @{command (global) "end"} command, one that does not belong to a local theory target. No further commands may follow such a global @{command (global) "end"}. \<^rail>\ @@{command theory} @{syntax system_name} @'imports' (@{syntax system_name} +) \ keywords? abbrevs? @'begin' ; keywords: @'keywords' (keyword_decls + @'and') ; keyword_decls: (@{syntax string} +) ('::' @{syntax name} @{syntax tags})? ; abbrevs: @'abbrevs' (((text+) '=' (text+)) + @'and') ; @@{command thy_deps} (thy_bounds thy_bounds?)? ; thy_bounds: @{syntax name} | '(' (@{syntax name} + @'|') ')' \ \<^descr> \<^theory_text>\theory A imports B\<^sub>1 \ B\<^sub>n begin\ starts a new theory \A\ based on the merge of existing theories \B\<^sub>1 \ B\<^sub>n\. Due to the possibility to import more than one ancestor, the resulting theory structure of an Isabelle session forms a directed acyclic graph (DAG). Isabelle takes care that sources contributing to the development graph are always up-to-date: changed files are automatically rechecked whenever a theory header specification is processed. Empty imports are only allowed in the bootstrap process of the special theory \<^theory>\Pure\, which is the start of any other formal development based on Isabelle. Regular user theories usually refer to some more complex entry point, such as theory \<^theory>\Main\ in Isabelle/HOL. The @{keyword_def "keywords"} specification declares outer syntax (\chref{ch:outer-syntax}) that is introduced in this theory later on (rare in end-user applications). Both minor keywords and major keywords of the Isar command language need to be specified, in order to make parsing of proof documents work properly. Command keywords need to be classified according to their structural role in the formal text. Examples may be seen in Isabelle/HOL sources itself, such as @{keyword "keywords"}~\<^verbatim>\"typedef"\ \:: thy_goal_defn\ or @{keyword "keywords"}~\<^verbatim>\"datatype"\ \:: thy_defn\ for theory-level definitions with and without proof, respectively. Additional @{syntax tags} provide defaults for document preparation (\secref{sec:document-markers}). The @{keyword_def "abbrevs"} specification declares additional abbreviations for syntactic completion. The default for a new keyword is just its name, but completion may be avoided by defining @{keyword "abbrevs"} with empty text. \<^descr> @{command (global) "end"} concludes the current theory definition. Note that some other commands, e.g.\ local theory targets \<^theory_text>\locale\ or \<^theory_text>\class\ may involve a \<^theory_text>\begin\ that needs to be matched by @{command (local) "end"}, according to the usual rules for nested blocks. \<^descr> \<^theory_text>\thy_deps\ visualizes the theory hierarchy as a directed acyclic graph. By default, all imported theories are shown. This may be restricted by specifying bounds wrt. the theory inclusion relation. \ section \Local theory targets \label{sec:target}\ text \ \begin{matharray}{rcll} @{command_def "context"} & : & \theory \ local_theory\ \\ @{command_def (local) "end"} & : & \local_theory \ theory\ \\ @{keyword_def "private"} \\ @{keyword_def "qualified"} \\ \end{matharray} A local theory target is a specification context that is managed separately within the enclosing theory. Contexts may introduce parameters (fixed variables) and assumptions (hypotheses). Definitions and theorems depending on the context may be added incrementally later on. \<^emph>\Named contexts\ refer to locales (cf.\ \secref{sec:locale}) or type classes (cf.\ \secref{sec:class}); the name ``\-\'' signifies the global theory context. \<^emph>\Unnamed contexts\ may introduce additional parameters and assumptions, and results produced in the context are generalized accordingly. Such auxiliary contexts may be nested within other targets, like \<^theory_text>\locale\, \<^theory_text>\class\, \<^theory_text>\instantiation\, \<^theory_text>\overloading\. \<^rail>\ @@{command context} @{syntax name} @{syntax_ref "opening"}? @'begin' ; @@{command context} @{syntax_ref "includes"}? (@{syntax context_elem} * ) @'begin' ; @{syntax_def target}: '(' @'in' @{syntax name} ')' \ \<^descr> \<^theory_text>\context c bundles begin\ opens a named context, by recommencing an existing locale or class \c\. Note that locale and class definitions allow to include the \<^theory_text>\begin\ keyword as well, in order to continue the local theory immediately after the initial specification. Optionally given \bundles\ only take effect in the surface context within the \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\context bundles elements begin\ opens an unnamed context, by extending the enclosing global or local theory target by the given declaration bundles (\secref{sec:bundle}) and context elements (\<^theory_text>\fixes\, \<^theory_text>\assumes\ etc.). This means any results stemming from definitions and proofs in the extended context will be exported into the enclosing target by lifting over extra parameters and premises. \<^descr> @{command (local) "end"} concludes the current local theory, according to the nesting of contexts. Note that a global @{command (global) "end"} has a different meaning: it concludes the theory itself (\secref{sec:begin-thy}). \<^descr> \<^theory_text>\private\ or \<^theory_text>\qualified\ may be given as modifiers before any local theory command. This restricts name space accesses to the local scope, as determined by the enclosing \<^theory_text>\context begin \ end\ block. Outside its scope, a \<^theory_text>\private\ name is inaccessible, and a \<^theory_text>\qualified\ name is only accessible with some qualification. Neither a global \<^theory_text>\theory\ nor a \<^theory_text>\locale\ target provides a local scope by itself: an extra unnamed context is required to use \<^theory_text>\private\ or \<^theory_text>\qualified\ here. \<^descr> \(\@{keyword_def "in"}~\c)\ given after any local theory command specifies an immediate target, e.g.\ ``\<^theory_text>\definition (in c)\'' or ``\<^theory_text>\theorem (in c)\''. This works both in a local or global theory context; the current target context will be suspended for this command only. Note that ``\<^theory_text>\(in -)\'' will always produce a global result independently of the current target context. Any specification element that operates on \local_theory\ according to this manual implicitly allows the above target syntax \<^theory_text>\(in c)\, but individual syntax diagrams omit that aspect for clarity. \<^medskip> The exact meaning of results produced within a local theory context depends on the underlying target infrastructure (locale, type class etc.). The general idea is as follows, considering a context named \c\ with parameter \x\ and assumption \A[x]\. Definitions are exported by introducing a global version with additional arguments; a syntactic abbreviation links the long form with the abstract version of the target context. For example, \a \ t[x]\ becomes \c.a ?x \ t[?x]\ at the theory level (for arbitrary \?x\), together with a local abbreviation \a \ c.a x\ in the target context (for the fixed parameter \x\). Theorems are exported by discharging the assumptions and generalizing the parameters of the context. For example, \a: B[x]\ becomes \c.a: A[?x] \ B[?x]\, again for arbitrary \?x\. \ section \Bundled declarations \label{sec:bundle}\ text \ \begin{matharray}{rcl} @{command_def "bundle"} & : & \local_theory \ local_theory\ \\ @{command "bundle"} & : & \theory \ local_theory\ \\ @{command_def "print_bundles"}\\<^sup>*\ & : & \context \\ \\ @{command_def "include"} & : & \proof(state) \ proof(state)\ \\ @{command_def "including"} & : & \proof(prove) \ proof(prove)\ \\ @{keyword_def "includes"} & : & syntax \\ \end{matharray} The outer syntax of fact expressions (\secref{sec:syn-att}) involves theorems and attributes, which are evaluated in the context and applied to it. Attributes may declare theorems to the context, as in \this_rule [intro] that_rule [elim]\ for example. Configuration options (\secref{sec:config}) are special declaration attributes that operate on the context without a theorem, as in \[[show_types = false]]\ for example. Expressions of this form may be defined as \<^emph>\bundled declarations\ in the context, and included in other situations later on. Including declaration bundles augments a local context casually without logical dependencies, which is in contrast to locales and locale interpretation (\secref{sec:locale}). \<^rail>\ @@{command bundle} @{syntax name} ( '=' @{syntax thms} @{syntax for_fixes} | @'begin') ; @@{command print_bundles} ('!'?) ; (@@{command include} | @@{command including}) (@{syntax name}+) ; @{syntax_def "includes"}: @'includes' (@{syntax name}+) ; @{syntax_def "opening"}: @'opening' (@{syntax name}+) ; @@{command unbundle} (@{syntax name}+) \ \<^descr> \<^theory_text>\bundle b = decls\ defines a bundle of declarations in the current context. The RHS is similar to the one of the \<^theory_text>\declare\ command. Bundles defined in local theory targets are subject to transformations via morphisms, when moved into different application contexts; this works analogously to any other local theory specification. \<^descr> \<^theory_text>\bundle b begin body end\ defines a bundle of declarations from the \body\ of local theory specifications. It may consist of commands that are technically equivalent to \<^theory_text>\declare\ or \<^theory_text>\declaration\, which also includes \<^theory_text>\notation\, for example. Named fact declarations like ``\<^theory_text>\lemmas a [simp] = b\'' or ``\<^theory_text>\lemma a [simp]: B \\'' are also admitted, but the name bindings are not recorded in the bundle. \<^descr> \<^theory_text>\print_bundles\ prints the named bundles that are available in the current context; the ``\!\'' option indicates extra verbosity. \<^descr> \<^theory_text>\include b\<^sub>1 \ b\<^sub>n\ activates the declarations from the given bundles in a proof body (forward mode). This is analogous to \<^theory_text>\note\ (\secref{sec:proof-facts}) with the expanded bundles. \<^descr> \<^theory_text>\including b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\include\, but works in proof refinement (backward mode). This is analogous to \<^theory_text>\using\ (\secref{sec:proof-facts}) with the expanded bundles. \<^descr> \<^theory_text>\includes b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\include\, but applies to a confined specification context: unnamed \<^theory_text>\context\s and long statements of \<^theory_text>\theorem\. \<^descr> \<^theory_text>\opening b\<^sub>1 \ b\<^sub>n\ is similar to \<^theory_text>\includes\, but applies to a named specification context: \<^theory_text>\locale\s, \<^theory_text>\class\es and named \<^theory_text>\context\s. The effect is confined to the surface context within the specification block itself and the corresponding \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\unbundle b\<^sub>1 \ b\<^sub>n\ activates the declarations from the given bundles in the current local theory context. This is analogous to \<^theory_text>\lemmas\ (\secref{sec:theorems}) with the expanded bundles. Here is an artificial example of bundling various configuration options: \ (*<*)experiment begin(*>*) bundle trace = [[simp_trace, linarith_trace, metis_trace, smt_trace]] lemma "x = x" including trace by metis (*<*)end(*>*) section \Term definitions \label{sec:term-definitions}\ text \ \begin{matharray}{rcll} @{command_def "definition"} & : & \local_theory \ local_theory\ \\ @{attribute_def "defn"} & : & \attribute\ \\ @{command_def "print_defn_rules"}\\<^sup>*\ & : & \context \\ \\ @{command_def "abbreviation"} & : & \local_theory \ local_theory\ \\ @{command_def "print_abbrevs"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} Term definitions may either happen within the logic (as equational axioms of a certain form (see also \secref{sec:overloading}), or outside of it as rewrite system on abstract syntax. The second form is called ``abbreviation''. \<^rail>\ @@{command definition} decl? definition ; @@{command abbreviation} @{syntax mode}? decl? abbreviation ; @@{command print_abbrevs} ('!'?) ; decl: @{syntax name} ('::' @{syntax type})? @{syntax mixfix}? @'where' ; definition: @{syntax thmdecl}? @{syntax prop} @{syntax spec_prems} @{syntax for_fixes} ; abbreviation: @{syntax prop} @{syntax for_fixes} \ \<^descr> \<^theory_text>\definition c where eq\ produces an internal definition \c \ t\ according to the specification given as \eq\, which is then turned into a proven fact. The given proposition may deviate from internal meta-level equality according to the rewrite rules declared as @{attribute defn} by the object-logic. This usually covers object-level equality \x = y\ and equivalence \A \ B\. End-users normally need not change the @{attribute defn} setup. Definitions may be presented with explicit arguments on the LHS, as well as additional conditions, e.g.\ \f x y = t\ instead of \f \ \x y. t\ and \y \ 0 \ g x y = u\ instead of an unrestricted \g \ \x y. u\. \<^descr> \<^theory_text>\print_defn_rules\ prints the definitional rewrite rules declared via @{attribute defn} in the current context. \<^descr> \<^theory_text>\abbreviation c where eq\ introduces a syntactic constant which is associated with a certain term according to the meta-level equality \eq\. Abbreviations participate in the usual type-inference process, but are expanded before the logic ever sees them. Pretty printing of terms involves higher-order rewriting with rules stemming from reverted abbreviations. This needs some care to avoid overlapping or looping syntactic replacements! The optional \mode\ specification restricts output to a particular print mode; using ``\input\'' here achieves the effect of one-way abbreviations. The mode may also include an ``\<^theory_text>\output\'' qualifier that affects the concrete syntax declared for abbreviations, cf.\ \<^theory_text>\syntax\ in \secref{sec:syn-trans}. \<^descr> \<^theory_text>\print_abbrevs\ prints all constant abbreviations of the current context; the ``\!\'' option indicates extra verbosity. \ section \Axiomatizations \label{sec:axiomatizations}\ text \ \begin{matharray}{rcll} @{command_def "axiomatization"} & : & \theory \ theory\ & (axiomatic!) \\ \end{matharray} \<^rail>\ @@{command axiomatization} @{syntax vars}? (@'where' axiomatization)? ; axiomatization: (@{syntax thmdecl} @{syntax prop} + @'and') @{syntax spec_prems} @{syntax for_fixes} \ \<^descr> \<^theory_text>\axiomatization c\<^sub>1 \ c\<^sub>m where \\<^sub>1 \ \\<^sub>n\ introduces several constants simultaneously and states axiomatic properties for these. The constants are marked as being specified once and for all, which prevents additional specifications for the same constants later on, but it is always possible to emit axiomatizations without referring to particular constants. Note that lack of precise dependency tracking of axiomatizations may disrupt the well-formedness of an otherwise definitional theory. Axiomatization is restricted to a global theory context: support for local theory targets \secref{sec:target} would introduce an extra dimension of uncertainty what the written specifications really are, and make it infeasible to argue why they are correct. Axiomatic specifications are required when declaring a new logical system within Isabelle/Pure, but in an application environment like Isabelle/HOL the user normally stays within definitional mechanisms provided by the logic and its libraries. \ section \Generic declarations\ text \ \begin{matharray}{rcl} @{command_def "declaration"} & : & \local_theory \ local_theory\ \\ @{command_def "syntax_declaration"} & : & \local_theory \ local_theory\ \\ @{command_def "declare"} & : & \local_theory \ local_theory\ \\ \end{matharray} Arbitrary operations on the background context may be wrapped-up as generic declaration elements. Since the underlying concept of local theories may be subject to later re-interpretation, there is an additional dependency on a morphism that tells the difference of the original declaration context wrt.\ the application context encountered later on. A fact declaration is an important special case: it consists of a theorem which is applied to the context by means of an attribute. \<^rail>\ (@@{command declaration} | @@{command syntax_declaration}) ('(' @'pervasive' ')')? \ @{syntax text} ; @@{command declare} (@{syntax thms} + @'and') \ - \<^descr> \<^theory_text>\declaration d\ adds the declaration function \d\ of ML type \<^ML_type>\declaration\, to the current local theory under construction. In later + \<^descr> \<^theory_text>\declaration d\ adds the declaration function \d\ of ML type \<^ML_type>\Morphism.declaration_fn\, to the current local theory under construction. In later application contexts, the function is transformed according to the morphisms being involved in the interpretation hierarchy. If the \<^theory_text>\(pervasive)\ option is given, the corresponding declaration is applied to all possible contexts involved, including the global background theory. \<^descr> \<^theory_text>\syntax_declaration\ is similar to \<^theory_text>\declaration\, but is meant to affect only ``syntactic'' tools by convention (such as notation and type-checking information). \<^descr> \<^theory_text>\declare thms\ declares theorems to the current local theory context. No theorem binding is involved here, unlike \<^theory_text>\lemmas\ (cf.\ \secref{sec:theorems}), so \<^theory_text>\declare\ only has the effect of applying attributes as included in the theorem specification. \ section \Locales \label{sec:locale}\ text \ A locale is a functor that maps parameters (including implicit type parameters) and a specification to a list of declarations. The syntax of locales is modeled after the Isar proof context commands (cf.\ \secref{sec:proof-context}). Locale hierarchies are supported by maintaining a graph of dependencies between locale instances in the global theory. Dependencies may be introduced through import (where a locale is defined as sublocale of the imported instances) or by proving that an existing locale is a sublocale of one or several locale instances. A locale may be opened with the purpose of appending to its list of declarations (cf.\ \secref{sec:target}). When opening a locale declarations from all dependencies are collected and are presented as a local theory. In this process, which is called \<^emph>\roundup\, redundant locale instances are omitted. A locale instance is redundant if it is subsumed by an instance encountered earlier. A more detailed description of this process is available elsewhere \<^cite>\Ballarin2014\. \ subsection \Locale expressions \label{sec:locale-expr}\ text \ A \<^emph>\locale expression\ denotes a context composed of instances of existing locales. The context consists of the declaration elements from the locale instances. Redundant locale instances are omitted according to roundup. \<^rail>\ @{syntax_def locale_expr}: (instance + '+') @{syntax for_fixes} ; instance: (qualifier ':')? @{syntax name} (pos_insts | named_insts) \ rewrites? ; qualifier: @{syntax name} ('?')? ; pos_insts: ('_' | @{syntax term})* ; named_insts: @'where' (@{syntax name} '=' @{syntax term} + @'and') ; rewrites: @'rewrites' (@{syntax thmdecl}? @{syntax prop} + @'and') \ A locale instance consists of a reference to a locale and either positional or named parameter instantiations optionally followed by rewrites clauses. Identical instantiations (that is, those that instantiate a parameter by itself) may be omitted. The notation ``\_\'' enables to omit the instantiation for a parameter inside a positional instantiation. Terms in instantiations are from the context the locale expressions is declared in. Local names may be added to this context with the optional \<^theory_text>\for\ clause. This is useful for shadowing names bound in outer contexts, and for declaring syntax. In addition, syntax declarations from one instance are effective when parsing subsequent instances of the same expression. Instances have an optional qualifier which applies to names in declarations. Names include local definitions and theorem names. If present, the qualifier itself is either mandatory (default) or non-mandatory (when followed by ``\<^verbatim>\?\''). Non-mandatory means that the qualifier may be omitted on input. Qualifiers only affect name spaces; they play no role in determining whether one locale instance subsumes another. Rewrite clauses amend instances with equations that act as rewrite rules. This is particularly useful for changing concepts introduced through definitions. Rewrite clauses are available only in interpretation commands (see \secref{sec:locale-interpretation} below) and must be proved the user. \ subsection \Locale declarations\ text \ \begin{tabular}{rcl} @{command_def "locale"} & : & \theory \ local_theory\ \\ @{command_def "experiment"} & : & \theory \ local_theory\ \\ @{command_def "print_locale"}\\<^sup>*\ & : & \context \\ \\ @{command_def "print_locales"}\\<^sup>*\ & : & \context \\ \\ @{command_def "locale_deps"}\\<^sup>*\ & : & \context \\ \\ \end{tabular} @{index_ref \\<^theory_text>\fixes\ (element)\} @{index_ref \\<^theory_text>\constrains\ (element)\} @{index_ref \\<^theory_text>\assumes\ (element)\} @{index_ref \\<^theory_text>\defines\ (element)\} @{index_ref \\<^theory_text>\notes\ (element)\} \<^rail>\ @@{command locale} @{syntax name} ('=' @{syntax locale})? @'begin'? ; @@{command experiment} (@{syntax context_elem}*) @'begin' ; @@{command print_locale} '!'? @{syntax name} ; @@{command print_locales} ('!'?) ; @{syntax_def locale}: @{syntax context_elem}+ | @{syntax_ref "opening"} ('+' (@{syntax context_elem}+))? | @{syntax locale_expr} @{syntax_ref "opening"}? ('+' (@{syntax context_elem}+))? ; @{syntax_def context_elem}: @'fixes' @{syntax vars} | @'constrains' (@{syntax name} '::' @{syntax type} + @'and') | @'assumes' (@{syntax props} + @'and') | @'defines' (@{syntax thmdecl}? @{syntax prop} @{syntax prop_pat}? + @'and') | @'notes' (@{syntax thmdef}? @{syntax thms} + @'and') \ \<^descr> \<^theory_text>\locale loc = import opening bundles + body\ defines a new locale \loc\ as a context consisting of a certain view of existing locales (\import\) plus some additional elements (\body\) with declaration \bundles\ enriching the context of the command itself. All \import\, \bundles\ and \body\ are optional; the degenerate form \<^theory_text>\locale loc\ defines an empty locale, which may still be useful to collect declarations of facts later on. Type-inference on locale expressions automatically takes care of the most general typing that the combined context elements may acquire. The \import\ consists of a locale expression; see \secref{sec:locale-expr} above. Its \<^theory_text>\for\ clause defines the parameters of \import\. These are parameters of the defined locale. Locale parameters whose instantiation is omitted automatically extend the (possibly empty) \<^theory_text>\for\ clause: they are inserted at its beginning. This means that these parameters may be referred to from within the expression and also in the subsequent context elements and provides a notational convenience for the inheritance of parameters in locale declarations. Declarations from \bundles\, see \secref{sec:bundle}, are effective in the entire command including a subsequent \<^theory_text>\begin\ / \<^theory_text>\end\ block, but they do not contribute to the declarations stored in the locale. The \body\ consists of context elements: \<^descr> @{element "fixes"}~\x :: \ (mx)\ declares a local parameter of type \\\ and mixfix annotation \mx\ (both are optional). The special syntax declaration ``\(\@{keyword_ref "structure"}\)\'' means that \x\ may be referenced implicitly in this context. \<^descr> @{element "constrains"}~\x :: \\ introduces a type constraint \\\ on the local parameter \x\. This element is deprecated. The type constraint should be introduced in the \<^theory_text>\for\ clause or the relevant @{element "fixes"} element. \<^descr> @{element "assumes"}~\a: \\<^sub>1 \ \\<^sub>n\ introduces local premises, similar to \<^theory_text>\assume\ within a proof (cf.\ \secref{sec:proof-context}). \<^descr> @{element "defines"}~\a: x \ t\ defines a previously declared parameter. This is similar to \<^theory_text>\define\ within a proof (cf.\ \secref{sec:proof-context}), but @{element "defines"} is restricted to Pure equalities and the defined variable needs to be declared beforehand via @{element "fixes"}. The left-hand side of the equation may have additional arguments, e.g.\ ``@{element "defines"}~\f x\<^sub>1 \ x\<^sub>n \ t\'', which need to be free in the context. \<^descr> @{element "notes"}~\a = b\<^sub>1 \ b\<^sub>n\ reconsiders facts within a local context. Most notably, this may include arbitrary declarations in any attribute specifications included here, e.g.\ a local @{attribute simp} rule. Both @{element "assumes"} and @{element "defines"} elements contribute to the locale specification. When defining an operation derived from the parameters, \<^theory_text>\definition\ (\secref{sec:term-definitions}) is usually more appropriate. Note that ``\<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\'' patterns given in the syntax of @{element "assumes"} and @{element "defines"} above are illegal in locale definitions. In the long goal format of \secref{sec:goals}, term bindings may be included as expected, though. \<^medskip> Locale specifications are ``closed up'' by turning the given text into a predicate definition \loc_axioms\ and deriving the original assumptions as local lemmas (modulo local definitions). The predicate statement covers only the newly specified assumptions, omitting the content of included locale expressions. The full cumulative view is only provided on export, involving another predicate \loc\ that refers to the complete specification text. In any case, the predicate arguments are those locale parameters that actually occur in the respective piece of text. Also these predicates operate at the meta-level in theory, but the locale packages attempts to internalize statements according to the object-logic setup (e.g.\ replacing \\\ by \\\, and \\\ by \\\ in HOL; see also \secref{sec:object-logic}). Separate introduction rules \loc_axioms.intro\ and \loc.intro\ are provided as well. \<^descr> \<^theory_text>\experiment body begin\ opens an anonymous locale context with private naming policy. Specifications in its body are inaccessible from outside. This is useful to perform experiments, without polluting the name space. \<^descr> \<^theory_text>\print_locale "locale"\ prints the contents of the named locale. The command omits @{element "notes"} elements by default. Use \<^theory_text>\print_locale!\ to have them included. \<^descr> \<^theory_text>\print_locales\ prints the names of all locales of the current theory; the ``\!\'' option indicates extra verbosity. \<^descr> \<^theory_text>\locale_deps\ visualizes all locales and their relations as a Hasse diagram. This includes locales defined as type classes (\secref{sec:class}). \ subsection \Locale interpretation \label{sec:locale-interpretation}\ text \ \begin{matharray}{rcl} @{command "interpretation"} & : & \local_theory \ proof(prove)\ \\ @{command_def "interpret"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "global_interpretation"} & : & \theory | local_theory \ proof(prove)\ \\ @{command_def "sublocale"} & : & \theory | local_theory \ proof(prove)\ \\ @{command_def "print_interps"}\\<^sup>*\ & : & \context \\ \\ @{method_def intro_locales} & : & \method\ \\ @{method_def unfold_locales} & : & \method\ \\ @{attribute_def trace_locales} & : & \mbox{\attribute\ \quad default \false\} \\ \end{matharray} Locales may be instantiated, and the resulting instantiated declarations added to the current context. This requires proof (of the instantiated specification) and is called \<^emph>\locale interpretation\. Interpretation is possible within arbitrary local theories (\<^theory_text>\interpretation\), within proof bodies (\<^theory_text>\interpret\), into global theories (\<^theory_text>\global_interpretation\) and into locales (\<^theory_text>\sublocale\). \<^rail>\ @@{command interpretation} @{syntax locale_expr} ; @@{command interpret} @{syntax locale_expr} ; @@{command global_interpretation} @{syntax locale_expr} definitions? ; @@{command sublocale} (@{syntax name} ('<' | '\'))? @{syntax locale_expr} \ definitions? ; @@{command print_interps} @{syntax name} ; definitions: @'defines' (@{syntax thmdecl}? @{syntax name} \ @{syntax mixfix}? '=' @{syntax term} + @'and'); \ The core of each interpretation command is a locale expression \expr\; the command generates proof obligations for the instantiated specifications. Once these are discharged by the user, instantiated declarations (in particular, facts) are added to the context in a post-processing phase, in a manner specific to each command. Interpretation commands are aware of interpretations that are already active: post-processing is achieved through a variant of roundup that takes interpretations of the current global or local theory into account. In order to simplify the proof obligations according to existing interpretations use methods @{method intro_locales} or @{method unfold_locales}. Rewrites clauses \<^theory_text>\rewrites eqns\ occur within expressions. They amend the morphism through which a locale instance is interpreted with rewrite rules, also called rewrite morphisms. This is particularly useful for interpreting concepts introduced through definitions. The equations must be proved the user. To enable syntax of the instantiated locale within the equation, while reading a locale expression, equations of a locale instance are read in a temporary context where the instance is already activated. If activation fails, typically due to duplicate constant declarations, processing falls back to reading the equation first. Given definitions \defs\ produce corresponding definitions in the local theory's underlying target \<^emph>\and\ amend the morphism with rewrite rules stemming from the symmetric of those definitions. Hence these need not be proved explicitly the user. Such rewrite definitions are a even more useful device for interpreting concepts introduced through definitions, but they are only supported for interpretation commands operating in a local theory whose implementing target actually supports this. Note that despite the suggestive \<^theory_text>\and\ connective, \defs\ are processed sequentially without mutual recursion. \<^descr> \<^theory_text>\interpretation expr\ interprets \expr\ into a local theory such that its lifetime is limited to the current context block (e.g. a locale or unnamed context). At the closing @{command end} of the block the interpretation and its declarations disappear. Hence facts based on interpretation can be established without creating permanent links to the interpreted locale instances, as would be the case with @{command sublocale}. When used on the level of a global theory, there is no end of a current context block, hence \<^theory_text>\interpretation\ behaves identically to \<^theory_text>\global_interpretation\ then. \<^descr> \<^theory_text>\interpret expr\ interprets \expr\ into a proof context: the interpretation and its declarations disappear when closing the current proof block. Note that for \<^theory_text>\interpret\ the \eqns\ should be explicitly universally quantified. \<^descr> \<^theory_text>\global_interpretation expr defines defs\ interprets \expr\ into a global theory. When adding declarations to locales, interpreted versions of these declarations are added to the global theory for all interpretations in the global theory as well. That is, interpretations into global theories dynamically participate in any declarations added to locales. Free variables in the interpreted expression are allowed. They are turned into schematic variables in the generated declarations. In order to use a free variable whose name is already bound in the context --- for example, because a constant of that name exists --- add it to the \<^theory_text>\for\ clause. When used in a nested target, resulting declarations are propagated through the whole target stack. \<^descr> \<^theory_text>\sublocale name \ expr defines defs\ interprets \expr\ into the locale \name\. A proof that the specification of \name\ implies the specification of \expr\ is required. As in the localized version of the theorem command, the proof is in the context of \name\. After the proof obligation has been discharged, the locale hierarchy is changed as if \name\ imported \expr\ (hence the name \<^theory_text>\sublocale\). When the context of \name\ is subsequently entered, traversing the locale hierarchy will involve the locale instances of \expr\, and their declarations will be added to the context. This makes \<^theory_text>\sublocale\ dynamic: extensions of a locale that is instantiated in \expr\ may take place after the \<^theory_text>\sublocale\ declaration and still become available in the context. Circular \<^theory_text>\sublocale\ declarations are allowed as long as they do not lead to infinite chains. If interpretations of \name\ exist in the current global theory, the command adds interpretations for \expr\ as well, with the same qualifier, although only for fragments of \expr\ that are not interpreted in the theory already. Rewrites clauses in the expression or rewrite definitions \defs\ can help break infinite chains induced by circular \<^theory_text>\sublocale\ declarations. In a named context block the \<^theory_text>\sublocale\ command may also be used, but the locale argument must be omitted. The command then refers to the locale (or class) target of the context block. \<^descr> \<^theory_text>\print_interps name\ lists all interpretations of locale \name\ in the current theory or proof context, including those due to a combination of an \<^theory_text>\interpretation\ or \<^theory_text>\interpret\ and one or several \<^theory_text>\sublocale\ declarations. \<^descr> @{method intro_locales} and @{method unfold_locales} repeatedly expand all introduction rules of locale predicates of the theory. While @{method intro_locales} only applies the \loc.intro\ introduction rules and therefore does not descend to assumptions, @{method unfold_locales} is more aggressive and applies \loc_axioms.intro\ as well. Both methods are aware of locale specifications entailed by the context, both from target statements, and from interpretations (see below). New goals that are entailed by the current context are discharged automatically. While @{method unfold_locales} is part of the default method for \<^theory_text>\proof\ and often invoked ``behind the scenes'', @{method intro_locales} helps understand which proof obligations originated from which locale instances. The latter method is useful while developing proofs but rare in finished developments. \<^descr> @{attribute trace_locales}, when set to \true\, prints the locale instances activated during roundup. Use this when locale commands yield obscure errors or for understanding local theories created by complex locale hierarchies. \begin{warn} If a global theory inherits declarations (body elements) for a locale from one parent and an interpretation of that locale from another parent, the interpretation will not be applied to the declarations. \end{warn} \begin{warn} Since attributes are applied to interpreted theorems, interpretation may modify the context of common proof tools, e.g.\ the Simplifier or Classical Reasoner. As the behaviour of such tools is \<^emph>\not\ stable under interpretation morphisms, manual declarations might have to be added to the target context of the interpretation to revert such declarations. \end{warn} \begin{warn} An interpretation in a local theory or proof context may subsume previous interpretations. This happens if the same specification fragment is interpreted twice and the instantiation of the second interpretation is more general than the interpretation of the first. The locale package does not attempt to remove subsumed interpretations. \end{warn} \begin{warn} While \<^theory_text>\interpretation (in c) \\ is admissible, it is not useful since its result is discarded immediately. \end{warn} \ section \Classes \label{sec:class}\ text \ \begin{matharray}{rcl} @{command_def "class"} & : & \theory \ local_theory\ \\ @{command_def "instantiation"} & : & \theory \ local_theory\ \\ @{command_def "instance"} & : & \local_theory \ local_theory\ \\ @{command "instance"} & : & \theory \ proof(prove)\ \\ @{command_def "subclass"} & : & \local_theory \ local_theory\ \\ @{command_def "print_classes"}\\<^sup>*\ & : & \context \\ \\ @{command_def "class_deps"}\\<^sup>*\ & : & \context \\ \\ @{method_def intro_classes} & : & \method\ \\ \end{matharray} A class is a particular locale with \<^emph>\exactly one\ type variable \\\. Beyond the underlying locale, a corresponding type class is established which is interpreted logically as axiomatic type class \<^cite>\"Wenzel:1997:TPHOL"\ whose logical content are the assumptions of the locale. Thus, classes provide the full generality of locales combined with the commodity of type classes (notably type-inference). See \<^cite>\"isabelle-classes"\ for a short tutorial. \<^rail>\ @@{command class} class_spec @'begin'? ; class_spec: @{syntax name} '=' ((@{syntax name} @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) | @{syntax name} @{syntax_ref "opening"}? | @{syntax_ref "opening"}? '+' (@{syntax context_elem}+)) ; @@{command instantiation} (@{syntax name} + @'and') '::' @{syntax arity} @'begin' ; @@{command instance} (() | (@{syntax name} + @'and') '::' @{syntax arity} | @{syntax name} ('<' | '\') @{syntax name} ) ; @@{command subclass} @{syntax name} ; @@{command class_deps} (class_bounds class_bounds?)? ; class_bounds: @{syntax sort} | '(' (@{syntax sort} + @'|') ')' \ \<^descr> \<^theory_text>\class c = superclasses bundles + body\ defines a new class \c\, inheriting from \superclasses\. This introduces a locale \c\ with import of all locales \superclasses\. Any @{element "fixes"} in \body\ are lifted to the global theory level (\<^emph>\class operations\ \f\<^sub>1, \, f\<^sub>n\ of class \c\), mapping the local type parameter \\\ to a schematic type variable \?\ :: c\. Likewise, @{element "assumes"} in \body\ are also lifted, mapping each local parameter \f :: \[\]\ to its corresponding global constant \f :: \[?\ :: c]\. The corresponding introduction rule is provided as \c_class_axioms.intro\. This rule should be rarely needed directly --- the @{method intro_classes} method takes care of the details of class membership proofs. Optionally given \bundles\ take effect in the surface context within the \body\ and the potentially following \<^theory_text>\begin\ / \<^theory_text>\end\ block. \<^descr> \<^theory_text>\instantiation t :: (s\<^sub>1, \, s\<^sub>n)s begin\ opens a target (cf.\ \secref{sec:target}) which allows to specify class operations \f\<^sub>1, \, f\<^sub>n\ corresponding to sort \s\ at the particular type instance \(\\<^sub>1 :: s\<^sub>1, \, \\<^sub>n :: s\<^sub>n) t\. A plain \<^theory_text>\instance\ command in the target body poses a goal stating these type arities. The target is concluded by an @{command_ref (local) "end"} command. Note that a list of simultaneous type constructors may be given; this corresponds nicely to mutually recursive type definitions, e.g.\ in Isabelle/HOL. \<^descr> \<^theory_text>\instance\ in an instantiation target body sets up a goal stating the type arities claimed at the opening \<^theory_text>\instantiation\. The proof would usually proceed by @{method intro_classes}, and then establish the characteristic theorems of the type classes involved. After finishing the proof, the background theory will be augmented by the proven type arities. On the theory level, \<^theory_text>\instance t :: (s\<^sub>1, \, s\<^sub>n)s\ provides a convenient way to instantiate a type class with no need to specify operations: one can continue with the instantiation proof immediately. \<^descr> \<^theory_text>\subclass c\ in a class context for class \d\ sets up a goal stating that class \c\ is logically contained in class \d\. After finishing the proof, class \d\ is proven to be subclass \c\ and the locale \c\ is interpreted into \d\ simultaneously. A weakened form of this is available through a further variant of @{command instance}: \<^theory_text>\instance c\<^sub>1 \ c\<^sub>2\ opens a proof that class \c\<^sub>2\ implies \c\<^sub>1\ without reference to the underlying locales; this is useful if the properties to prove the logical connection are not sufficient on the locale level but on the theory level. \<^descr> \<^theory_text>\print_classes\ prints all classes in the current theory. \<^descr> \<^theory_text>\class_deps\ visualizes classes and their subclass relations as a directed acyclic graph. By default, all classes from the current theory context are show. This may be restricted by optional bounds as follows: \<^theory_text>\class_deps upper\ or \<^theory_text>\class_deps upper lower\. A class is visualized, iff it is a subclass of some sort from \upper\ and a superclass of some sort from \lower\. \<^descr> @{method intro_classes} repeatedly expands all class introduction rules of this theory. Note that this method usually needs not be named explicitly, as it is already included in the default proof step (e.g.\ of \<^theory_text>\proof\). In particular, instantiation of trivial (syntactic) classes may be performed by a single ``\<^theory_text>\..\'' proof step. \ subsection \The class target\ text \ %FIXME check A named context may refer to a locale (cf.\ \secref{sec:target}). If this locale is also a class \c\, apart from the common locale target behaviour the following happens. \<^item> Local constant declarations \g[\]\ referring to the local type parameter \\\ and local parameters \f[\]\ are accompanied by theory-level constants \g[?\ :: c]\ referring to theory-level class operations \f[?\ :: c]\. \<^item> Local theorem bindings are lifted as are assumptions. \<^item> Local syntax refers to local operations \g[\]\ and global operations \g[?\ :: c]\ uniformly. Type inference resolves ambiguities. In rare cases, manual type annotations are needed. \ subsection \Co-regularity of type classes and arities\ text \ The class relation together with the collection of type-constructor arities must obey the principle of \<^emph>\co-regularity\ as defined below. \<^medskip> For the subsequent formulation of co-regularity we assume that the class relation is closed by transitivity and reflexivity. Moreover the collection of arities \t :: (\<^vec>s)c\ is completed such that \t :: (\<^vec>s)c\ and \c \ c'\ implies \t :: (\<^vec>s)c'\ for all such declarations. Treating sorts as finite sets of classes (meaning the intersection), the class relation \c\<^sub>1 \ c\<^sub>2\ is extended to sorts as follows: \[ \s\<^sub>1 \ s\<^sub>2 \ \c\<^sub>2 \ s\<^sub>2. \c\<^sub>1 \ s\<^sub>1. c\<^sub>1 \ c\<^sub>2\ \] This relation on sorts is further extended to tuples of sorts (of the same length) in the component-wise way. \<^medskip> Co-regularity of the class relation together with the arities relation means: \[ \t :: (\<^vec>s\<^sub>1)c\<^sub>1 \ t :: (\<^vec>s\<^sub>2)c\<^sub>2 \ c\<^sub>1 \ c\<^sub>2 \ \<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ \] for all such arities. In other words, whenever the result classes of some type-constructor arities are related, then the argument sorts need to be related in the same way. \<^medskip> Co-regularity is a very fundamental property of the order-sorted algebra of types. For example, it entails principal types and most general unifiers, e.g.\ see \<^cite>\"nipkow-prehofer"\. \ section \Overloaded constant definitions \label{sec:overloading}\ text \ Definitions essentially express abbreviations within the logic. The simplest form of a definition is \c :: \ \ t\, where \c\ is a new constant and \t\ is a closed term that does not mention \c\. Moreover, so-called \<^emph>\hidden polymorphism\ is excluded: all type variables in \t\ need to occur in its type \\\. \<^emph>\Overloading\ means that a constant being declared as \c :: \ decl\ may be defined separately on type instances \c :: (\\<^sub>1, \, \\<^sub>n)\ decl\ for each type constructor \\\. At most occasions overloading will be used in a Haskell-like fashion together with type classes by means of \<^theory_text>\instantiation\ (see \secref{sec:class}). Sometimes low-level overloading is desirable; this is supported by \<^theory_text>\consts\ and \<^theory_text>\overloading\ explained below. The right-hand side of overloaded definitions may mention overloaded constants recursively at type instances corresponding to the immediate argument types \\\<^sub>1, \, \\<^sub>n\. Incomplete specification patterns impose global constraints on all occurrences. E.g.\ \d :: \ \ \\ on the left-hand side means that all corresponding occurrences on some right-hand side need to be an instance of this, and general \d :: \ \ \\ will be disallowed. Full details are given by Kun\v{c}ar \<^cite>\"Kuncar:2015"\. \<^medskip> The \<^theory_text>\consts\ command and the \<^theory_text>\overloading\ target provide a convenient interface for end-users. Regular specification elements such as @{command definition}, @{command inductive}, @{command function} may be used in the body. It is also possible to use \<^theory_text>\consts c :: \\ with later \<^theory_text>\overloading c \ c :: \\ to keep the declaration and definition of a constant separate. \begin{matharray}{rcl} @{command_def "consts"} & : & \theory \ theory\ \\ @{command_def "overloading"} & : & \theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command consts} ((@{syntax name} '::' @{syntax type} @{syntax mixfix}?) +) ; @@{command overloading} ( spec + ) @'begin' ; spec: @{syntax name} ( '\' | '==' ) @{syntax term} ( '(' @'unchecked' ')' )? \ \<^descr> \<^theory_text>\consts c :: \\ declares constant \c\ to have any instance of type scheme \\\. The optional mixfix annotations may attach concrete syntax to the constants declared. \<^descr> \<^theory_text>\overloading x\<^sub>1 \ c\<^sub>1 :: \\<^sub>1 \ x\<^sub>n \ c\<^sub>n :: \\<^sub>n begin \ end\ defines a theory target (cf.\ \secref{sec:target}) which allows to specify already declared constants via definitions in the body. These are identified by an explicitly given mapping from variable names \x\<^sub>i\ to constants \c\<^sub>i\ at particular type instances. The definitions themselves are established using common specification tools, using the names \x\<^sub>i\ as reference to the corresponding constants. Option \<^theory_text>\(unchecked)\ disables global dependency checks for the corresponding definition, which is occasionally useful for exotic overloading; this is a form of axiomatic specification. It is at the discretion of the user to avoid malformed theory specifications! \ subsubsection \Example\ consts Length :: "'a \ nat" overloading Length\<^sub>0 \ "Length :: unit \ nat" Length\<^sub>1 \ "Length :: 'a \ unit \ nat" Length\<^sub>2 \ "Length :: 'a \ 'b \ unit \ nat" Length\<^sub>3 \ "Length :: 'a \ 'b \ 'c \ unit \ nat" begin fun Length\<^sub>0 :: "unit \ nat" where "Length\<^sub>0 () = 0" fun Length\<^sub>1 :: "'a \ unit \ nat" where "Length\<^sub>1 (a, ()) = 1" fun Length\<^sub>2 :: "'a \ 'b \ unit \ nat" where "Length\<^sub>2 (a, b, ()) = 2" fun Length\<^sub>3 :: "'a \ 'b \ 'c \ unit \ nat" where "Length\<^sub>3 (a, b, c, ()) = 3" end lemma "Length (a, b, c, ()) = 3" by simp lemma "Length ((a, b), (c, d), ()) = 2" by simp lemma "Length ((a, b, c, d, e), ()) = 1" by simp section \Incorporating ML code \label{sec:ML}\ text \ \begin{matharray}{rcl} @{command_def "SML_file"} & : & \local_theory \ local_theory\ \\ @{command_def "SML_file_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "SML_file_no_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_file_no_debug"} & : & \local_theory \ local_theory\ \\ @{command_def "ML"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_export"} & : & \local_theory \ local_theory\ \\ @{command_def "ML_prf"} & : & \proof \ proof\ \\ @{command_def "ML_val"} & : & \any \\ \\ @{command_def "ML_command"} & : & \any \\ \\ @{command_def "setup"} & : & \theory \ theory\ \\ @{command_def "local_setup"} & : & \local_theory \ local_theory\ \\ @{command_def "attribute_setup"} & : & \local_theory \ local_theory\ \\ \end{matharray} \begin{tabular}{rcll} @{attribute_def ML_print_depth} & : & \attribute\ & default 10 \\ @{attribute_def ML_source_trace} & : & \attribute\ & default \false\ \\ @{attribute_def ML_debugger} & : & \attribute\ & default \false\ \\ @{attribute_def ML_exception_trace} & : & \attribute\ & default \false\ \\ @{attribute_def ML_exception_debugger} & : & \attribute\ & default \false\ \\ @{attribute_def ML_environment} & : & \attribute\ & default \Isabelle\ \\ \end{tabular} \<^rail>\ (@@{command SML_file} | @@{command SML_file_debug} | @@{command SML_file_no_debug} | @@{command ML_file} | @@{command ML_file_debug} | @@{command ML_file_no_debug}) @{syntax name} ';'? ; (@@{command ML} | @@{command ML_export} | @@{command ML_prf} | @@{command ML_val} | @@{command ML_command} | @@{command setup} | @@{command local_setup}) @{syntax text} ; @@{command attribute_setup} @{syntax name} '=' @{syntax text} @{syntax text}? \ \<^descr> \<^theory_text>\SML_file name\ reads and evaluates the given Standard ML file. Top-level SML bindings are stored within the (global or local) theory context; the initial environment is restricted to the Standard ML implementation of Poly/ML, without the many add-ons of Isabelle/ML. Multiple \<^theory_text>\SML_file\ commands may be used to build larger Standard ML projects, independently of the regular Isabelle/ML environment. \<^descr> \<^theory_text>\ML_file name\ reads and evaluates the given ML file. The current theory context is passed down to the ML toplevel and may be modified, using \<^ML>\Context.>>\ or derived ML commands. Top-level ML bindings are stored within the (global or local) theory context. \<^descr> \<^theory_text>\SML_file_debug\, \<^theory_text>\SML_file_no_debug\, \<^theory_text>\ML_file_debug\, and \<^theory_text>\ML_file_no_debug\ change the @{attribute ML_debugger} option locally while the given file is compiled. \<^descr> \<^theory_text>\ML\ is similar to \<^theory_text>\ML_file\, but evaluates directly the given \text\. Top-level ML bindings are stored within the (global or local) theory context. \<^descr> \<^theory_text>\ML_export\ is similar to \<^theory_text>\ML\, but the resulting toplevel bindings are exported to the global bootstrap environment of the ML process --- it has a lasting effect that cannot be retracted. This allows ML evaluation without a formal theory context, e.g. for command-line tools via @{tool process} \<^cite>\"isabelle-system"\. \<^descr> \<^theory_text>\ML_prf\ is analogous to \<^theory_text>\ML\ but works within a proof context. Top-level ML bindings are stored within the proof context in a purely sequential fashion, disregarding the nested proof structure. ML bindings introduced by \<^theory_text>\ML_prf\ are discarded at the end of the proof. \<^descr> \<^theory_text>\ML_val\ and \<^theory_text>\ML_command\ are diagnostic versions of \<^theory_text>\ML\, which means that the context may not be updated. \<^theory_text>\ML_val\ echos the bindings produced at the ML toplevel, but \<^theory_text>\ML_command\ is silent. \<^descr> \<^theory_text>\setup "text"\ changes the current theory context by applying \text\, which refers to an ML expression of type \<^ML_type>\theory -> theory\. This enables to initialize any object-logic specific tools and packages written in ML, for example. \<^descr> \<^theory_text>\local_setup\ is similar to \<^theory_text>\setup\ for a local theory context, and an ML expression of type \<^ML_type>\local_theory -> local_theory\. This allows to invoke local theory specification packages without going through concrete outer syntax, for example. \<^descr> \<^theory_text>\attribute_setup name = "text" description\ defines an attribute in the current context. The given \text\ has to be an ML expression of type \<^ML_type>\attribute context_parser\, cf.\ basic parsers defined in structure \<^ML_structure>\Args\ and \<^ML_structure>\Attrib\. In principle, attributes can operate both on a given theorem and the implicit context, although in practice only one is modified and the other serves as parameter. Here are examples for these two cases: \ (*<*)experiment begin(*>*) attribute_setup my_rule = \Attrib.thms >> (fn ths => Thm.rule_attribute ths (fn context: Context.generic => fn th: thm => let val th' = th OF ths in th' end))\ attribute_setup my_declaration = \Attrib.thms >> (fn ths => Thm.declaration_attribute (fn th: thm => fn context: Context.generic => let val context' = context in context' end))\ (*<*)end(*>*) text \ \<^descr> @{attribute ML_print_depth} controls the printing depth of the ML toplevel pretty printer. Typically the limit should be less than 10. Bigger values such as 100--1000 are occasionally useful for debugging. \<^descr> @{attribute ML_source_trace} indicates whether the source text that is given to the ML compiler should be output: it shows the raw Standard ML after expansion of Isabelle/ML antiquotations. \<^descr> @{attribute ML_debugger} controls compilation of sources with or without debugging information. The global system option @{system_option_ref ML_debugger} does the same when building a session image. It is also possible use commands like \<^theory_text>\ML_file_debug\ etc. The ML debugger is explained further in \<^cite>\"isabelle-jedit"\. \<^descr> @{attribute ML_exception_trace} indicates whether the ML run-time system should print a detailed stack trace on exceptions. The result is dependent on various ML compiler optimizations. The boundary for the exception trace is the current Isar command transactions: it is occasionally better to insert the combinator \<^ML>\Runtime.exn_trace\ into ML code for debugging \<^cite>\"isabelle-implementation"\, closer to the point where it actually happens. \<^descr> @{attribute ML_exception_debugger} controls detailed exception trace via the Poly/ML debugger, at the cost of extra compile-time and run-time overhead. Relevant ML modules need to be compiled beforehand with debugging enabled, see @{attribute ML_debugger} above. \<^descr> @{attribute ML_environment} determines the named ML environment for toplevel declarations, e.g.\ in command \<^theory_text>\ML\ or \<^theory_text>\ML_file\. The following ML environments are predefined in Isabelle/Pure: \<^item> \Isabelle\ for Isabelle/ML. It contains all modules of Isabelle/Pure and further add-ons, e.g. material from Isabelle/HOL. \<^item> \SML\ for official Standard ML. It contains only the initial basis according to \<^url>\http://sml-family.org/Basis/overview.html\. The Isabelle/ML function \<^ML>\ML_Env.setup\ defines a new ML environment. This is useful to incorporate big SML projects in an isolated name space, possibly with variations on ML syntax; the existing setup of \<^ML>\ML_Env.SML_operations\ follows the official standard. It is also possible to move toplevel bindings between ML environments, using a notation with ``\>\'' as separator. For example: \ (*<*)experiment begin(*>*) declare [[ML_environment = "Isabelle>SML"]] ML \val println = writeln\ declare [[ML_environment = "SML"]] ML \println "test"\ declare [[ML_environment = "Isabelle"]] ML \ML \println\ (*bad*) handle ERROR msg => warning msg\ (*<*)end(*>*) section \Generated files and exported files\ text \ Write access to the physical file-system is incompatible with the stateless model of processing Isabelle documents. To avoid bad effects, the following concepts for abstract file-management are provided by Isabelle: \<^descr>[Generated files] are stored within the theory context in Isabelle/ML. This allows to operate on the content in Isabelle/ML, e.g. via the command @{command compile_generated_files}. \<^descr>[Exported files] are stored within the session database in Isabelle/Scala. This allows to deliver artefacts to external tools, see also \<^cite>\"isabelle-system"\ for session \<^verbatim>\ROOT\ declaration \<^theory_text>\export_files\, and @{tool build} option \<^verbatim>\-e\. A notable example is the command @{command_ref export_code} (\chref{ch:export-code}): it uses both concepts simultaneously. File names are hierarchically structured, using a slash as separator. The (long) theory name is used as a prefix: the resulting name needs to be globally unique. \begin{matharray}{rcll} @{command_def "generate_file"} & : & \local_theory \ local_theory\ \\ @{command_def "export_generated_files"} & : & \context \\ \\ @{command_def "compile_generated_files"} & : & \context \\ \\ @{command_def "external_file"} & : & \any \ any\ \\ \end{matharray} \<^rail>\ @@{command generate_file} path '=' content ; path: @{syntax embedded} ; content: @{syntax embedded} ; @@{command export_generated_files} (files_in_theory + @'and') ; files_in_theory: (@'_' | (path+)) (('(' @'in' @{syntax name} ')')?) ; @@{command compile_generated_files} (files_in_theory + @'and') \ (@'external_files' (external_files + @'and'))? \ (@'export_files' (export_files + @'and'))? \ (@'export_prefix' path)? ; external_files: (path+) (('(' @'in' path ')')?) ; export_files: (path+) (executable?) ; executable: '(' ('exe' | 'executable') ')' ; @@{command external_file} @{syntax name} ';'? \ \<^descr> \<^theory_text>\generate_file path = content\ augments the table of generated files within the current theory by a new entry: duplicates are not allowed. The name extension determines a pre-existent file-type; the \content\ is a string that is preprocessed according to rules of this file-type. For example, Isabelle/Pure supports \<^verbatim>\.hs\ as file-type for Haskell: embedded cartouches are evaluated as Isabelle/ML expressions of type \<^ML_type>\string\, the result is inlined in Haskell string syntax. \<^descr> \<^theory_text>\export_generated_files paths (in thy)\ retrieves named generated files from the given theory (that needs be reachable via imports of the current one). By default, the current theory node is used. Using ``\<^verbatim>\_\'' (underscore) instead of explicit path names refers to \emph{all} files of a theory node. The overall list of files is prefixed with the respective (long) theory name and exported to the session database. In Isabelle/jEdit the result can be browsed via the virtual file-system with prefix ``\<^verbatim>\isabelle-export:\'' (using the regular file-browser). \<^descr> \<^theory_text>\scala_build_generated_files paths (in thy)\ retrieves named generated files as for \<^theory_text>\export_generated_files\ and writes them into a temporary directory, which is taken as starting point for build process of Isabelle/Scala/Java modules (see \<^cite>\"isabelle-system"\). The corresponding @{path \build.props\} file is expected directly in the toplevel directory, instead of @{path \etc/build.props\} for Isabelle system components. These properties need to specify sources, resources, services etc. as usual. The resulting JAR module becomes an export artefact of the session database, with a name of the form ``\theory\\<^verbatim>\:classpath/\\module\\<^verbatim>\.jar\''. \<^descr> \<^theory_text>\compile_generated_files paths (in thy) where compile_body\ retrieves named generated files as for \<^theory_text>\export_generated_files\ and writes them into a temporary directory, such that the \compile_body\ may operate on them as an ML function of type \<^ML_type>\Path.T -> unit\. This may create further files, e.g.\ executables produced by a compiler that is invoked as external process (e.g.\ via \<^ML>\Isabelle_System.bash\), or any other files. The option ``\<^theory_text>\external_files paths (in base_dir)\'' copies files from the physical file-system into the temporary directory, \emph{before} invoking \compile_body\. The \base_dir\ prefix is removed from each of the \paths\, but the remaining sub-directory structure is reconstructed in the target directory. The option ``\<^theory_text>\export_files paths\'' exports the specified files from the temporary directory to the session database, \emph{after} invoking \compile_body\. Entries may be decorated with ``\<^theory_text>\(exe)\'' to say that it is a platform-specific executable program: the executable file-attribute will be set, and on Windows the \<^verbatim>\.exe\ file-extension will be included; ``\<^theory_text>\(executable)\'' only refers to the file-attribute, without special treatment of the \<^verbatim>\.exe\ extension. The option ``\<^theory_text>\export_prefix path\'' specifies an extra path prefix for all exports of \<^theory_text>\export_files\ above. \<^descr> \<^theory_text>\external_file name\ declares the formal dependency on the given file name, such that the Isabelle build process knows about it (see also \<^cite>\"isabelle-system"\). This is required for any files mentioned in \<^theory_text>\compile_generated_files / external_files\ above, in order to document source dependencies properly. It is also possible to use \<^theory_text>\external_file\ alone, e.g.\ when other Isabelle/ML tools use \<^ML>\File.read\, without specific management of content by the Prover IDE. \ section \Primitive specification elements\ subsection \Sorts\ text \ \begin{matharray}{rcll} @{command_def "default_sort"} & : & \local_theory \ local_theory\ \end{matharray} \<^rail>\ @@{command default_sort} @{syntax sort} \ \<^descr> \<^theory_text>\default_sort s\ makes sort \s\ the new default sort for any type variable that is given explicitly in the text, but lacks a sort constraint (wrt.\ the current context). Type variables generated by type inference are not affected. Usually the default sort is only changed when defining a new object-logic. For example, the default sort in Isabelle/HOL is \<^class>\type\, the class of all HOL types. When merging theories, the default sorts of the parents are logically intersected, i.e.\ the representations as lists of classes are joined. \ subsection \Types \label{sec:types-pure}\ text \ \begin{matharray}{rcll} @{command_def "type_synonym"} & : & \local_theory \ local_theory\ \\ @{command_def "typedecl"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command type_synonym} (@{syntax typespec} '=' @{syntax type} @{syntax mixfix}?) ; @@{command typedecl} @{syntax typespec} @{syntax mixfix}? \ \<^descr> \<^theory_text>\type_synonym (\\<^sub>1, \, \\<^sub>n) t = \\ introduces a \<^emph>\type synonym\ \(\\<^sub>1, \, \\<^sub>n) t\ for the existing type \\\. Unlike the semantic type definitions in Isabelle/HOL, type synonyms are merely syntactic abbreviations without any logical significance. Internally, type synonyms are fully expanded. \<^descr> \<^theory_text>\typedecl (\\<^sub>1, \, \\<^sub>n) t\ declares a new type constructor \t\. If the object-logic defines a base sort \s\, then the constructor is declared to operate on that, via the axiomatic type-class instance \t :: (s, \, s)s\. \begin{warn} If you introduce a new type axiomatically, i.e.\ via @{command_ref typedecl} and @{command_ref axiomatization} (\secref{sec:axiomatizations}), the minimum requirement is that it has a non-empty model, to avoid immediate collapse of the logical environment. Moreover, one needs to demonstrate that the interpretation of such free-form axiomatizations can coexist with other axiomatization schemes for types, notably @{command_def typedef} in Isabelle/HOL (\secref{sec:hol-typedef}), or any other extension that people might have introduced elsewhere. \end{warn} \ section \Naming existing theorems \label{sec:theorems}\ text \ \begin{matharray}{rcll} @{command_def "lemmas"} & : & \local_theory \ local_theory\ \\ @{command_def "named_theorems"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command lemmas} (@{syntax thmdef}? @{syntax thms} + @'and') @{syntax for_fixes} ; @@{command named_theorems} (@{syntax name} @{syntax text}? + @'and') \ \<^descr> \<^theory_text>\lemmas a = b\<^sub>1 \ b\<^sub>n\~@{keyword_def "for"}~\x\<^sub>1 \ x\<^sub>m\ evaluates given facts (with attributes) in the current context, which may be augmented by local variables. Results are standardized before being stored, i.e.\ schematic variables are renamed to enforce index \0\ uniformly. \<^descr> \<^theory_text>\named_theorems name description\ declares a dynamic fact within the context. The same \name\ is used to define an attribute with the usual \add\/\del\ syntax (e.g.\ see \secref{sec:simp-rules}) to maintain the content incrementally, in canonical declaration order of the text structure. \ section \Oracles \label{sec:oracles}\ text \ \begin{matharray}{rcll} @{command_def "oracle"} & : & \theory \ theory\ & (axiomatic!) \\ @{command_def "thm_oracles"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} Oracles allow Isabelle to take advantage of external reasoners such as arithmetic decision procedures, model checkers, fast tautology checkers or computer algebra systems. Invoked as an oracle, an external reasoner can create arbitrary Isabelle theorems. It is the responsibility of the user to ensure that the external reasoner is as trustworthy as the application requires. Another typical source of errors is the linkup between Isabelle and the external tool, not just its concrete implementation, but also the required translation between two different logical environments. Isabelle merely guarantees well-formedness of the propositions being asserted, and records within the internal derivation object how presumed theorems depend on unproven suppositions. This also includes implicit type-class reasoning via the order-sorted algebra of class relations and type arities (see also @{command_ref instantiation} and @{command_ref instance}). \<^rail>\ @@{command oracle} @{syntax name} '=' @{syntax text} ; @@{command thm_oracles} @{syntax thms} \ \<^descr> \<^theory_text>\oracle name = "text"\ turns the given ML expression \text\ of type \<^ML_text>\'a -> cterm\ into an ML function of type \<^ML_text>\'a -> thm\, which is bound to the global identifier \<^ML_text>\name\. This acts like an infinitary specification of axioms! Invoking the oracle only works within the scope of the resulting theory. See \<^file>\~~/src/HOL/Examples/Iff_Oracle.thy\ for a worked example of defining a new primitive rule as oracle, and turning it into a proof method. \<^descr> \<^theory_text>\thm_oracles thms\ displays all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. \ section \Name spaces\ text \ \begin{matharray}{rcl} @{command_def "alias"} & : & \local_theory \ local_theory\ \\ @{command_def "type_alias"} & : & \local_theory \ local_theory\ \\ @{command_def "hide_class"} & : & \theory \ theory\ \\ @{command_def "hide_type"} & : & \theory \ theory\ \\ @{command_def "hide_const"} & : & \theory \ theory\ \\ @{command_def "hide_fact"} & : & \theory \ theory\ \\ \end{matharray} \<^rail>\ (@{command alias} | @{command type_alias}) @{syntax name} '=' @{syntax name} ; (@{command hide_class} | @{command hide_type} | @{command hide_const} | @{command hide_fact}) ('(' @'open' ')')? (@{syntax name} + ) \ Isabelle organizes any kind of name declarations (of types, constants, theorems etc.) by separate hierarchically structured name spaces. Normally the user does not have to control the behaviour of name spaces by hand, yet the following commands provide some way to do so. \<^descr> \<^theory_text>\alias\ and \<^theory_text>\type_alias\ introduce aliases for constants and type constructors, respectively. This allows adhoc changes to name-space accesses. \<^descr> \<^theory_text>\type_alias b = c\ introduces an alias for an existing type constructor. \<^descr> \<^theory_text>\hide_class names\ fully removes class declarations from a given name space; with the \(open)\ option, only the unqualified base name is hidden. Note that hiding name space accesses has no impact on logical declarations --- they remain valid internally. Entities that are no longer accessible to the user are printed with the special qualifier ``\??\'' prefixed to the full internal name. \<^descr> \<^theory_text>\hide_type\, \<^theory_text>\hide_const\, and \<^theory_text>\hide_fact\ are similar to \<^theory_text>\hide_class\, but hide types, constants, and facts, respectively. \ end diff --git a/src/HOL/Decision_Procs/ferrante_rackoff_data.ML b/src/HOL/Decision_Procs/ferrante_rackoff_data.ML --- a/src/HOL/Decision_Procs/ferrante_rackoff_data.ML +++ b/src/HOL/Decision_Procs/ferrante_rackoff_data.ML @@ -1,140 +1,140 @@ (* Title: HOL/Decision_Procs/ferrante_rackoff_data.ML Author: Amine Chaieb, TU Muenchen Context data for Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders. *) signature FERRANTE_RACKOF_DATA = sig datatype ord = Lt | Le | Gt | Ge | Eq | NEq | Nox type entry val get: Proof.context -> (thm * entry) list val del: attribute val add: entry -> attribute val funs: thm -> {isolate_conv: morphism -> Proof.context -> cterm list -> cterm -> thm, whatis: morphism -> cterm -> cterm -> ord, - simpset: morphism -> Proof.context -> simpset} -> declaration + simpset: morphism -> Proof.context -> simpset} -> Morphism.declaration_fn val match: Proof.context -> cterm -> entry option end; structure Ferrante_Rackoff_Data: FERRANTE_RACKOF_DATA = struct (* data *) datatype ord = Lt | Le | Gt | Ge | Eq | NEq | Nox type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list, ld: thm list, qe: thm, atoms : cterm list} * {isolate_conv: Proof.context -> cterm list -> cterm -> thm, whatis : cterm -> cterm -> ord, simpset : simpset}; val eq_key = Thm.eq_thm; fun eq_data arg = eq_fst eq_key arg; structure Data = Generic_Data ( type T = (thm * entry) list; val empty = []; fun merge data : T = AList.merge eq_key (K true) data; ); val get = Data.get o Context.Proof; fun del_data key = remove eq_data (key, []); val del = Thm.declaration_attribute (Data.map o del_data); fun add entry = Thm.declaration_attribute (fn key => fn context => context |> Data.map (del_data key #> cons (key, entry))); (* extra-logical functions *) fun funs raw_key {isolate_conv = icv, whatis = wi, simpset = ss} phi context = context |> Data.map (fn data => let val key = Morphism.thm phi raw_key; val _ = AList.defined eq_key data key orelse raise THM ("No data entry for structure key", 0, [key]); val fns = {isolate_conv = icv phi, whatis = wi phi, simpset = ss phi (Context.proof_of context)}; in AList.map_entry eq_key key (apsnd (K fns)) data end); fun match ctxt tm = let fun match_inst ({minf, pinf, nmi, npi, ld, qe, atoms}, fns) pat = let fun h instT = let val substT = Thm.instantiate (instT, Vars.empty); val substT_cterm = Drule.cterm_rule substT; val minf' = map substT minf val pinf' = map substT pinf val nmi' = map substT nmi val npi' = map substT npi val ld' = map substT ld val qe' = substT qe val atoms' = map substT_cterm atoms val result = ({minf = minf', pinf = pinf', nmi = nmi', npi = npi', ld = ld', qe = qe', atoms = atoms'}, 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 ({atoms = atoms, ...}, _): entry) = get_first (match_inst entry) atoms; in get_first match_struct (get ctxt) end; (* concrete syntax *) local val minfN = "minf"; val pinfN = "pinf"; val nmiN = "nmi"; val npiN = "npi"; val lin_denseN = "lindense"; val qeN = "qe" val atomsN = "atoms" val simpsN = "simps" fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); val any_keyword = keyword minfN || keyword pinfN || keyword nmiN || keyword npiN || keyword lin_denseN || keyword qeN || keyword atomsN || keyword simpsN; val thms = Scan.repeats (Scan.unless any_keyword Attrib.multi_thm); val terms = thms >> map Drule.dest_term; in val _ = Theory.setup (Attrib.setup \<^binding>\ferrack\ ((keyword minfN |-- thms) -- (keyword pinfN |-- thms) -- (keyword nmiN |-- thms) -- (keyword npiN |-- thms) -- (keyword lin_denseN |-- thms) -- (keyword qeN |-- thms) -- (keyword atomsN |-- terms) >> (fn ((((((minf,pinf),nmi),npi),lin_dense),qe), atoms)=> if length qe = 1 then add ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = lin_dense, qe = hd qe, atoms = atoms}, {isolate_conv = undefined, whatis = undefined, simpset = HOL_ss}) else error "only one theorem for qe!")) "Ferrante Rackoff data"); end; end; diff --git a/src/HOL/Tools/Function/partial_function.ML b/src/HOL/Tools/Function/partial_function.ML --- a/src/HOL/Tools/Function/partial_function.ML +++ b/src/HOL/Tools/Function/partial_function.ML @@ -1,322 +1,322 @@ (* Title: HOL/Tools/Function/partial_function.ML Author: Alexander Krauss, TU Muenchen Partial function definitions based on least fixed points in ccpos. *) signature PARTIAL_FUNCTION = sig - val init: string -> term -> term -> thm -> thm -> thm option -> declaration + val init: string -> term -> term -> thm -> thm -> thm option -> Morphism.declaration_fn val mono_tac: Proof.context -> int -> tactic val add_partial_function: string -> (binding * typ option * mixfix) list -> Attrib.binding * term -> local_theory -> (term * thm) * local_theory val add_partial_function_cmd: string -> (binding * string option * mixfix) list -> Attrib.binding * string -> local_theory -> (term * thm) * local_theory val transform_result: morphism -> term * thm -> term * thm end; structure Partial_Function: PARTIAL_FUNCTION = struct open Function_Lib (*** Context Data ***) datatype setup_data = Setup_Data of {fixp: term, mono: term, fixp_eq: thm, fixp_induct: thm, fixp_induct_user: thm option}; fun transform_setup_data phi (Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user}) = let val term = Morphism.term phi; val thm = Morphism.thm phi; in Setup_Data {fixp = term fixp, mono = term mono, fixp_eq = thm fixp_eq, fixp_induct = thm fixp_induct, fixp_induct_user = Option.map thm fixp_induct_user} end; structure Modes = Generic_Data ( type T = setup_data Symtab.table; val empty = Symtab.empty; fun merge data = Symtab.merge (K true) data; ) fun init mode fixp mono fixp_eq fixp_induct fixp_induct_user phi = let val data' = Setup_Data {fixp = fixp, mono = mono, fixp_eq = fixp_eq, fixp_induct = fixp_induct, fixp_induct_user = fixp_induct_user} |> transform_setup_data (phi $> Morphism.trim_context_morphism); in Modes.map (Symtab.update (mode, data')) end; val known_modes = Symtab.keys o Modes.get o Context.Proof; fun lookup_mode ctxt = Symtab.lookup (Modes.get (Context.Proof ctxt)) #> Option.map (transform_setup_data (Morphism.transfer_morphism' ctxt)); (*** Automated monotonicity proofs ***) (*rewrite conclusion with k-th assumtion*) fun rewrite_with_asm_tac ctxt k = Subgoal.FOCUS (fn {context = ctxt', prems, ...} => Local_Defs.unfold0_tac ctxt' [nth prems k]) ctxt; fun dest_case ctxt t = case strip_comb t of (Const (case_comb, _), args) => (case Ctr_Sugar.ctr_sugar_of_case ctxt case_comb of NONE => NONE | SOME {case_thms, ...} => let val lhs = Thm.prop_of (hd case_thms) |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst; val arity = length (snd (strip_comb lhs)); val conv = funpow (length args - arity) Conv.fun_conv (Conv.rewrs_conv (map mk_meta_eq case_thms)); in SOME (nth args (arity - 1), conv) end) | _ => NONE; (*split on case expressions*) val split_cases_tac = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} => SUBGOAL (fn (t, i) => case t of _ $ (_ $ Abs (_, _, body)) => (case dest_case ctxt body of NONE => no_tac | SOME (arg, conv) => let open Conv in if Term.is_open arg then no_tac else ((DETERM o Induct.cases_tac ctxt false [[SOME arg]] NONE []) THEN_ALL_NEW (rewrite_with_asm_tac ctxt 0) THEN_ALL_NEW eresolve_tac ctxt @{thms thin_rl} THEN_ALL_NEW (CONVERSION (params_conv ~1 (fn ctxt' => arg_conv (arg_conv (abs_conv (K conv) ctxt'))) ctxt))) i end) | _ => no_tac) 1); (*monotonicity proof: apply rules + split case expressions*) fun mono_tac ctxt = K (Local_Defs.unfold0_tac ctxt [@{thm curry_def}]) THEN' (TRY o REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\partial_function_mono\)) ORELSE' split_cases_tac ctxt)); (*** Auxiliary functions ***) (*Returns t $ u, but instantiates the type of t to make the application type correct*) fun apply_inst ctxt t u = let val thy = Proof_Context.theory_of ctxt; val T = domain_type (fastype_of t); val T' = fastype_of u; val subst = Sign.typ_match thy (T, T') Vartab.empty handle Type.TYPE_MATCH => raise TYPE ("apply_inst", [T, T'], [t, u]) in map_types (Envir.norm_type subst) t $ u end; fun head_conv cv ct = if can Thm.dest_comb ct then Conv.fun_conv (head_conv cv) ct else cv ct; (*** currying transformation ***) fun curry_const (A, B, C) = Const (\<^const_name>\Product_Type.curry\, [HOLogic.mk_prodT (A, B) --> C, A, B] ---> C); fun mk_curry f = case fastype_of f of Type ("fun", [Type (_, [S, T]), U]) => curry_const (S, T, U) $ f | T => raise TYPE ("mk_curry", [T], [f]); (* iterated versions. Nonstandard left-nested tuples arise naturally from "split o split o split"*) fun curry_n arity = funpow (arity - 1) mk_curry; fun uncurry_n arity = funpow (arity - 1) HOLogic.mk_case_prod; val curry_uncurry_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm Product_Type.curry_case_prod}, @{thm Product_Type.case_prod_curry}]) val split_conv_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm Product_Type.split_conv}]); val curry_K_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm Product_Type.curry_K}]); (* instantiate generic fixpoint induction and eliminate the canonical assumptions; curry induction predicate *) fun specialize_fixp_induct ctxt fT fT_uc curry uncurry mono_thm f_def rule = let val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt val P_inst = Abs ("f", fT_uc, Free (P, fT --> HOLogic.boolT) $ (curry $ Bound 0)) in (* FIXME ctxt vs. ctxt' (!?) *) rule |> infer_instantiate' ctxt ((map o Option.map) (Thm.cterm_of ctxt) [SOME uncurry, NONE, SOME curry, NONE, SOME P_inst]) |> Tactic.rule_by_tactic ctxt (Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 (* discharge U (C f) = f *) THEN Simplifier.simp_tac (put_simpset curry_K_ss ctxt) 4 (* simplify bot case *) THEN Simplifier.full_simp_tac (put_simpset curry_uncurry_ss ctxt) 5) (* simplify induction step *) |> (fn thm => thm OF [mono_thm, f_def]) |> Conv.fconv_rule (Conv.concl_conv ~1 (* simplify conclusion *) (Raw_Simplifier.rewrite ctxt false [mk_meta_eq @{thm Product_Type.curry_case_prod}])) |> singleton (Variable.export ctxt' ctxt) end fun mk_curried_induct args ctxt inst_rule = let val cert = Thm.cterm_of ctxt (* FIXME ctxt vs. ctxt' (!?) *) val ([P], ctxt') = Variable.variant_fixes ["P"] ctxt val split_paired_all_conv = Conv.every_conv (replicate (length args - 1) (Conv.rewr_conv @{thm split_paired_all})) val split_params_conv = Conv.params_conv ~1 (fn _ => Conv.implies_conv split_paired_all_conv Conv.all_conv) val (P_var, x_var) = Thm.prop_of inst_rule |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> strip_comb |> apsnd hd |> apply2 dest_Var val P_rangeT = range_type (snd P_var) val PT = map (snd o dest_Free) args ---> P_rangeT val x_inst = cert (foldl1 HOLogic.mk_prod args) val P_inst = cert (uncurry_n (length args) (Free (P, PT))) val inst_rule' = inst_rule |> Tactic.rule_by_tactic ctxt (Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 4 THEN Simplifier.simp_tac (put_simpset curry_uncurry_ss ctxt) 3 THEN CONVERSION (split_params_conv ctxt then_conv (Conv.forall_conv (K split_paired_all_conv) ctxt)) 3) |> Thm.instantiate (TVars.empty, Vars.make2 (P_var, P_inst) (x_var, x_inst)) |> Simplifier.full_simplify (put_simpset split_conv_ss ctxt) |> singleton (Variable.export ctxt' ctxt) in inst_rule' end; (*** partial_function definition ***) fun transform_result phi (t, thm) = (Morphism.term phi t, Morphism.thm phi thm); fun gen_add_partial_function prep mode fixes_raw eqn_raw lthy = let val setup_data = the (lookup_mode lthy mode) handle Option.Option => error (cat_lines ["Unknown mode " ^ quote mode ^ ".", "Known modes are " ^ commas_quote (known_modes lthy) ^ "."]); val Setup_Data {fixp, mono, fixp_eq, fixp_induct, fixp_induct_user} = setup_data; val ((fixes, [(eq_abinding, eqn)]), _) = prep fixes_raw [(eqn_raw, [], [])] lthy; val ((_, plain_eqn), args_ctxt) = Variable.focus NONE eqn lthy; val ((f_binding, fT), mixfix) = the_single fixes; val f_bname = Binding.name_of f_binding; fun note_qualified (name, thms) = Local_Theory.note ((derived_name f_binding name, []), thms) #> snd val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop plain_eqn); val (head, args) = strip_comb lhs; val argnames = map (fst o dest_Free) args; val F = fold_rev lambda (head :: args) rhs; val arity = length args; val (aTs, bTs) = chop arity (binder_types fT); val tupleT = foldl1 HOLogic.mk_prodT aTs; val fT_uc = tupleT :: bTs ---> body_type fT; val f_uc = Var ((f_bname, 0), fT_uc); val x_uc = Var (("x", 1), tupleT); val uncurry = lambda head (uncurry_n arity head); val curry = lambda f_uc (curry_n arity f_uc); val F_uc = lambda f_uc (uncurry_n arity (F $ curry_n arity f_uc)); val mono_goal = apply_inst lthy mono (lambda f_uc (F_uc $ f_uc $ x_uc)) |> HOLogic.mk_Trueprop |> Logic.all x_uc; val mono_thm = Goal.prove_internal lthy [] (Thm.cterm_of lthy mono_goal) (K (mono_tac lthy 1)) val inst_mono_thm = Thm.forall_elim (Thm.cterm_of lthy x_uc) mono_thm val f_def_rhs = curry_n arity (apply_inst lthy fixp F_uc); val f_def_binding = Thm.make_def_binding (Config.get lthy Function_Lib.function_internals) f_binding val ((f, (_, f_def)), lthy') = Local_Theory.define ((f_binding, mixfix), ((f_def_binding, []), f_def_rhs)) lthy; val eqn = HOLogic.mk_eq (list_comb (f, args), Term.betapplys (F, f :: args)) |> HOLogic.mk_Trueprop; val unfold = (infer_instantiate' lthy' (map (SOME o Thm.cterm_of lthy') [uncurry, F, curry]) fixp_eq OF [inst_mono_thm, f_def]) |> Tactic.rule_by_tactic lthy' (Simplifier.simp_tac (put_simpset curry_uncurry_ss lthy') 1); val specialized_fixp_induct = specialize_fixp_induct lthy' fT fT_uc curry uncurry inst_mono_thm f_def fixp_induct |> Drule.rename_bvars' (map SOME (f_bname :: f_bname :: argnames)); val mk_raw_induct = infer_instantiate' args_ctxt ((map o Option.map) (Thm.cterm_of args_ctxt) [SOME uncurry, NONE, SOME curry]) #> mk_curried_induct args args_ctxt #> singleton (Variable.export args_ctxt lthy') #> (fn thm => infer_instantiate' lthy' [SOME (Thm.cterm_of lthy' F)] thm OF [inst_mono_thm, f_def]) #> Drule.rename_bvars' (map SOME (f_bname :: argnames @ argnames)) val raw_induct = Option.map mk_raw_induct fixp_induct_user val rec_rule = let open Conv in Goal.prove lthy' (map (fst o dest_Free) args) [] eqn (fn _ => CONVERSION ((arg_conv o arg1_conv o head_conv o rewr_conv) (mk_meta_eq unfold)) 1 THEN resolve_tac lthy' @{thms refl} 1) end; val ((_, [rec_rule']), lthy'') = lthy' |> Local_Theory.note (eq_abinding, [rec_rule]) in lthy'' |> Spec_Rules.add Binding.empty Spec_Rules.equational_recdef [f] [rec_rule'] |> note_qualified ("simps", [rec_rule']) |> note_qualified ("mono", [mono_thm]) |> (case raw_induct of NONE => I | SOME thm => note_qualified ("raw_induct", [thm])) |> note_qualified ("fixp_induct", [specialized_fixp_induct]) |> pair (f, rec_rule') end; val add_partial_function = gen_add_partial_function Specification.check_multi_specs; val add_partial_function_cmd = gen_add_partial_function Specification.read_multi_specs; val mode = \<^keyword>\(\ |-- Parse.name --| \<^keyword>\)\; val _ = Outer_Syntax.local_theory \<^command_keyword>\partial_function\ "define partial function" ((mode -- (Parse.vars -- (Parse.where_ |-- Parse_Spec.simple_spec))) >> (fn (mode, (vars, spec)) => add_partial_function_cmd mode vars spec #> #2)); end; diff --git a/src/HOL/Tools/groebner.ML b/src/HOL/Tools/groebner.ML --- a/src/HOL/Tools/groebner.ML +++ b/src/HOL/Tools/groebner.ML @@ -1,986 +1,986 @@ (* Title: HOL/Tools/groebner.ML Author: Amine Chaieb, TU Muenchen *) signature GROEBNER = sig val ring_and_ideal_conv: {idom: thm list, ring: cterm list * thm list, field: cterm list * thm list, vars: cterm list, semiring: cterm list * thm list, ideal : thm list} -> (cterm -> Rat.rat) -> (Rat.rat -> cterm) -> conv -> conv -> {ring_conv: Proof.context -> conv, simple_ideal: cterm list -> cterm -> cterm ord -> cterm list, multi_ideal: cterm list -> cterm list -> cterm list -> (cterm * cterm) list, poly_eq_ss: simpset, unwind_conv: Proof.context -> conv} val ring_tac: thm list -> thm list -> Proof.context -> int -> tactic val ideal_tac: thm list -> thm list -> Proof.context -> int -> tactic val algebra_tac: thm list -> thm list -> Proof.context -> int -> tactic end structure Groebner : GROEBNER = struct 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_binary ct ct' = if is_binop ct ct' then Thm.dest_binop ct' else raise CTERM ("dest_binary: bad binop", [ct, ct']) val denominator_rat = Rat.dest #> snd #> Rat.of_int; fun int_of_rat a = case Rat.dest a of (i,1) => i | _ => error "int_of_rat: not an int"; val lcm_rat = fn x => fn y => Rat.of_int (Integer.lcm (int_of_rat x) (int_of_rat y)); val (eqF_intr, eqF_elim) = let val [th1,th2] = @{thms PFalse} in (fn th => th COMP th2, fn th => th COMP th1) end; val (PFalse, PFalse') = let val PFalse_eq = nth @{thms simp_thms} 13 in (PFalse_eq RS iffD1, PFalse_eq RS iffD2) end; (* Type for recording history, i.e. how a polynomial was obtained. *) datatype history = Start of int | Mmul of (Rat.rat * int list) * history | Add of history * history; (* Monomial ordering. *) fun morder_lt m1 m2= let fun lexorder l1 l2 = case (l1,l2) of ([],[]) => false | (x1::o1,x2::o2) => x1 > x2 orelse x1 = x2 andalso lexorder o1 o2 | _ => error "morder: inconsistent monomial lengths" val n1 = Integer.sum m1 val n2 = Integer.sum m2 in n1 < n2 orelse n1 = n2 andalso lexorder m1 m2 end; (* Arithmetic on canonical polynomials. *) fun grob_neg l = map (fn (c,m) => (Rat.neg c,m)) l; fun grob_add l1 l2 = case (l1,l2) of ([],l2) => l2 | (l1,[]) => l1 | ((c1,m1)::o1,(c2,m2)::o2) => if m1 = m2 then let val c = c1 + c2 val rest = grob_add o1 o2 in if c = @0 then rest else (c,m1)::rest end else if morder_lt m2 m1 then (c1,m1)::(grob_add o1 l2) else (c2,m2)::(grob_add l1 o2); fun grob_sub l1 l2 = grob_add l1 (grob_neg l2); fun grob_mmul (c1,m1) (c2,m2) = (c1 * c2, ListPair.map (op +) (m1, m2)); fun grob_cmul cm pol = map (grob_mmul cm) pol; fun grob_mul l1 l2 = case l1 of [] => [] | (h1::t1) => grob_add (grob_cmul h1 l2) (grob_mul t1 l2); fun grob_inv l = case l of [(c,vs)] => if (forall (fn x => x = 0) vs) then if c = @0 then error "grob_inv: division by zero" else [(@1 / c,vs)] else error "grob_inv: non-constant divisor polynomial" | _ => error "grob_inv: non-constant divisor polynomial"; fun grob_div l1 l2 = case l2 of [(c,l)] => if (forall (fn x => x = 0) l) then if c = @0 then error "grob_div: division by zero" else grob_cmul (@1 / c,l) l1 else error "grob_div: non-constant divisor polynomial" | _ => error "grob_div: non-constant divisor polynomial"; fun grob_pow vars l n = if n < 0 then error "grob_pow: negative power" else if n = 0 then [(@1,map (K 0) vars)] else grob_mul l (grob_pow vars l (n - 1)); (* Monomial division operation. *) fun mdiv (c1,m1) (c2,m2) = (c1 / c2, map2 (fn n1 => fn n2 => if n1 < n2 then error "mdiv" else n1 - n2) m1 m2); (* Lowest common multiple of two monomials. *) fun mlcm (_,m1) (_,m2) = (@1, ListPair.map Int.max (m1, m2)); (* Reduce monomial cm by polynomial pol, returning replacement for cm. *) fun reduce1 cm (pol,hpol) = case pol of [] => error "reduce1" | cm1::cms => ((let val (c,m) = mdiv cm cm1 in (grob_cmul (~ c, m) cms, Mmul ((~ c,m),hpol)) end) handle ERROR _ => error "reduce1"); (* Try this for all polynomials in a basis. *) fun tryfind f l = case l of [] => error "tryfind" | (h::t) => ((f h) handle ERROR _ => tryfind f t); fun reduceb cm basis = tryfind (fn p => reduce1 cm p) basis; (* Reduction of a polynomial (always picking largest monomial possible). *) fun reduce basis (pol,hist) = case pol of [] => (pol,hist) | cm::ptl => ((let val (q,hnew) = reduceb cm basis in reduce basis (grob_add q ptl,Add(hnew,hist)) end) handle (ERROR _) => (let val (q,hist') = reduce basis (ptl,hist) in (cm::q,hist') end)); (* Check for orthogonality w.r.t. LCM. *) fun orthogonal l p1 p2 = snd l = snd(grob_mmul (hd p1) (hd p2)); (* Compute S-polynomial of two polynomials. *) fun spoly cm ph1 ph2 = case (ph1,ph2) of (([],h),_) => ([],h) | (_,([],h)) => ([],h) | ((cm1::ptl1,his1),(cm2::ptl2,his2)) => (grob_sub (grob_cmul (mdiv cm cm1) ptl1) (grob_cmul (mdiv cm cm2) ptl2), Add(Mmul(mdiv cm cm1,his1), Mmul(mdiv (~ (fst cm),snd cm) cm2,his2))); (* Make a polynomial monic. *) fun monic (pol,hist) = if null pol then (pol,hist) else let val (c',m') = hd pol in (map (fn (c,m) => (c / c',m)) pol, Mmul((@1 / c',map (K 0) m'),hist)) end; (* The most popular heuristic is to order critical pairs by LCM monomial. *) fun forder ((_,m1),_) ((_,m2),_) = morder_lt m1 m2; fun poly_lt p q = case (p,q) of (_,[]) => false | ([],_) => true | ((c1,m1)::o1,(c2,m2)::o2) => c1 < c2 orelse c1 = c2 andalso ((morder_lt m1 m2) orelse m1 = m2 andalso poly_lt o1 o2); fun align ((p,hp),(q,hq)) = if poly_lt p q then ((p,hp),(q,hq)) else ((q,hq),(p,hp)); fun poly_eq p1 p2 = eq_list (fn ((c1, m1), (c2, m2)) => c1 = c2 andalso (m1: int list) = m2) (p1, p2); fun memx ((p1,_),(p2,_)) ppairs = not (exists (fn ((q1,_),(q2,_)) => poly_eq p1 q1 andalso poly_eq p2 q2) ppairs); (* Buchberger's second criterion. *) fun criterion2 basis (lcm,((p1,h1),(p2,h2))) opairs = exists (fn g => not(poly_eq (fst g) p1) andalso not(poly_eq (fst g) p2) andalso can (mdiv lcm) (hd(fst g)) andalso not(memx (align (g,(p1,h1))) (map snd opairs)) andalso not(memx (align (g,(p2,h2))) (map snd opairs))) basis; (* Test for hitting constant polynomial. *) fun constant_poly p = length p = 1 andalso forall (fn x => x = 0) (snd(hd p)); (* Grobner basis algorithm. *) (* FIXME: try to get rid of mergesort? *) fun merge ord l1 l2 = case l1 of [] => l2 | h1::t1 => case l2 of [] => l1 | h2::t2 => if ord h1 h2 then h1::(merge ord t1 l2) else h2::(merge ord l1 t2); fun mergesort ord l = let fun mergepairs l1 l2 = case (l1,l2) of ([s],[]) => s | (l,[]) => mergepairs [] l | (l,[s1]) => mergepairs (s1::l) [] | (l,(s1::s2::ss)) => mergepairs ((merge ord s1 s2)::l) ss in if null l then [] else mergepairs [] (map (fn x => [x]) l) end; fun grobner_basis basis pairs = case pairs of [] => basis | (l,(p1,p2))::opairs => let val (sph as (sp,_)) = monic (reduce basis (spoly l p1 p2)) in if null sp orelse criterion2 basis (l,(p1,p2)) opairs then grobner_basis basis opairs else if constant_poly sp then grobner_basis (sph::basis) [] else let val rawcps = map (fn p => (mlcm (hd(fst p)) (hd sp),align(p,sph))) basis val newcps = filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) rawcps in grobner_basis (sph::basis) (merge forder opairs (mergesort forder newcps)) end end; (* Interreduce initial polynomials. *) fun grobner_interreduce rpols ipols = case ipols of [] => map monic (rev rpols) | p::ps => let val p' = reduce (rpols @ ps) p in if null (fst p') then grobner_interreduce rpols ps else grobner_interreduce (p'::rpols) ps end; (* Overall function. *) fun grobner pols = let val npols = map_index (fn (n, p) => (p, Start n)) pols val phists = filter (fn (p,_) => not (null p)) npols val bas = grobner_interreduce [] (map monic phists) val prs0 = map_product pair bas bas val prs1 = filter (fn ((x,_),(y,_)) => poly_lt x y) prs0 val prs2 = map (fn (p,q) => (mlcm (hd(fst p)) (hd(fst q)),(p,q))) prs1 val prs3 = filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) prs2 in grobner_basis bas (mergesort forder prs3) end; (* Get proof of contradiction from Grobner basis. *) fun find p l = case l of [] => error "find" | (h::t) => if p(h) then h else find p t; fun grobner_refute pols = let val gb = grobner pols in snd(find (fn (p,_) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb) end; (* Turn proof into a certificate as sum of multipliers. *) (* In principle this is very inefficient: in a heavily shared proof it may *) (* make the same calculation many times. Could put in a cache or something. *) fun resolve_proof vars prf = case prf of Start(~1) => [] | Start m => [(m,[(@1,map (K 0) vars)])] | Mmul(pol,lin) => let val lis = resolve_proof vars lin in map (fn (n,p) => (n,grob_cmul pol p)) lis end | Add(lin1,lin2) => let val lis1 = resolve_proof vars lin1 val lis2 = resolve_proof vars lin2 val dom = distinct (op =) (union (op =) (map fst lis1) (map fst lis2)) in map (fn n => let val a = these (AList.lookup (op =) lis1 n) val b = these (AList.lookup (op =) lis2 n) in (n,grob_add a b) end) dom end; (* Run the procedure and produce Weak Nullstellensatz certificate. *) fun grobner_weak vars pols = let val cert = resolve_proof vars (grobner_refute pols) val l = fold_rev (fold_rev (lcm_rat o denominator_rat o fst) o snd) cert @1 in (l,map (fn (i,p) => (i,map (fn (d,m) => (l * d,m)) p)) cert) end; (* Prove a polynomial is in ideal generated by others, using Grobner basis. *) fun grobner_ideal vars pols pol = let val (pol',h) = reduce (grobner pols) (grob_neg pol,Start(~1)) in if not (null pol') then error "grobner_ideal: not in the ideal" else resolve_proof vars h end; (* Produce Strong Nullstellensatz certificate for a power of pol. *) fun grobner_strong vars pols pol = let val vars' = \<^cterm>\True\::vars val grob_z = [(@1, 1::(map (K 0) vars))] val grob_1 = [(@1, (map (K 0) vars'))] fun augment p= map (fn (c,m) => (c,0::m)) p val pols' = map augment pols val pol' = augment pol val allpols = (grob_sub (grob_mul grob_z pol') grob_1)::pols' val (l,cert) = grobner_weak vars' allpols val d = fold (fold (Integer.max o hd o snd) o snd) cert 0 fun transform_monomial (c,m) = grob_cmul (c,tl m) (grob_pow vars pol (d - hd m)) fun transform_polynomial q = fold_rev (grob_add o transform_monomial) q [] val cert' = map (fn (c,q) => (c-1,transform_polynomial q)) (filter (fn (k,_) => k <> 0) cert) in (d,l,cert') end; (* Overall parametrized universal procedure for (semi)rings. *) (* We return an ideal_conv and the actual ring prover. *) fun refute_disj rfn tm = case Thm.term_of tm of Const(\<^const_name>\HOL.disj\,_)$_$_ => Drule.compose (refute_disj rfn (Thm.dest_arg tm), 2, Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE)) | _ => rfn tm ; val notnotD = @{thm notnotD}; fun mk_binop ct x y = Thm.apply (Thm.apply ct x) y fun is_neg t = case Thm.term_of t of (Const(\<^const_name>\Not\,_)$_) => true | _ => false; fun is_eq t = case Thm.term_of t of (Const(\<^const_name>\HOL.eq\,_)$_$_) => true | _ => false; fun end_itlist f l = case l of [] => error "end_itlist" | [x] => x | (h::t) => f h (end_itlist f t); val list_mk_binop = fn b => end_itlist (mk_binop b); val list_dest_binop = fn b => let fun h acc t = ((let val (l,r) = dest_binary b t in h (h acc r) l end) handle CTERM _ => (t::acc)) (* Why had I handle _ => ? *) in h [] end; val strip_exists = let fun h (acc, t) = case Thm.term_of t of \<^Const_>\Ex _ for \Abs _\\ => h (Thm.dest_abs_global (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) end; fun is_forall t = case Thm.term_of t of (Const(\<^const_name>\All\,_)$Abs(_,_,_)) => true | _ => false; val nnf_simps = @{thms nnf_simps}; fun weak_dnf_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps}); val initial_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps nnf_simps addsimps [not_all, not_ex] addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps})); fun initial_conv ctxt = Simplifier.rewrite (put_simpset initial_ss ctxt); val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); val cTrp = \<^cterm>\Trueprop\; val cConj = \<^cterm>\HOL.conj\; val (cNot,false_tm) = (\<^cterm>\Not\, \<^cterm>\False\); val assume_Trueprop = Thm.apply cTrp #> Thm.assume; val list_mk_conj = list_mk_binop cConj; val conjs = list_dest_binop cConj; val mk_neg = Thm.apply cNot; fun striplist dest = let fun h acc x = case try dest x of SOME (a,b) => h (h acc b) a | NONE => x::acc in h [] end; fun list_mk_binop b = foldr1 (fn (s,t) => Thm.apply (Thm.apply b s) t); val eq_commute = mk_meta_eq @{thm eq_commute}; fun sym_conv eq = let val (l,r) = Thm.dest_binop eq in Thm.instantiate' [SOME (Thm.ctyp_of_cterm l)] [SOME l, SOME r] eq_commute end; (* FIXME : copied from cqe.ML -- complex QE*) fun conjuncts ct = case Thm.term_of ct of \<^term>\HOL.conj\$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct)) | _ => [ct]; fun fold1 f = foldr1 (uncurry f); fun mk_conj_tab th = let fun h acc th = case Thm.prop_of th of \<^term>\Trueprop\$(\<^term>\HOL.conj\$_$_) => h (h acc (th RS conjunct2)) (th RS conjunct1) | \<^term>\Trueprop\$p => (p,th)::acc in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; fun is_conj (\<^term>\HOL.conj\$_$_) = true | is_conj _ = false; fun prove_conj tab cjs = case cjs of [c] => if is_conj (Thm.term_of c) then prove_conj tab (conjuncts c) else tab c | c::cs => conjI OF [prove_conj tab [c], prove_conj tab cs]; fun conj_ac_rule eq = let val (l,r) = Thm.dest_equals eq val ctabl = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ l)) val ctabr = mk_conj_tab (Thm.assume (Thm.apply \<^cterm>\Trueprop\ r)) fun tabl c = the (Termtab.lookup ctabl (Thm.term_of c)) fun tabr c = the (Termtab.lookup ctabr (Thm.term_of c)) val thl = prove_conj tabl (conjuncts r) |> implies_intr_hyps val thr = prove_conj tabr (conjuncts l) |> implies_intr_hyps val eqI = Thm.instantiate' [] [SOME l, SOME r] @{thm iffI} in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; (* END FIXME.*) (* Conversion for the equivalence of existential statements where EX quantifiers are rearranged differently *) fun ext ctxt T = Thm.cterm_of ctxt (Const (\<^const_name>\Ex\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) fun mk_ex ctxt v t = Thm.apply (ext ctxt (Thm.typ_of_cterm v)) (Thm.lambda v t) fun choose v th th' = case Thm.concl_of th of \<^term>\Trueprop\ $ (Const(\<^const_name>\Ex\,_)$_) => let val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th val T = Thm.dest_ctyp0 (Thm.ctyp_of_cterm p) val th0 = Conv.fconv_rule (Thm.beta_conversion true) (Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) val pv = (Thm.rhs_of o Thm.beta_conversion true) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply p v)) val th1 = Thm.forall_intr v (Thm.implies_intr pv th') in Thm.implies_elim (Thm.implies_elim th0 th) th1 end | _ => error "" (* FIXME ? *) fun simple_choose ctxt v th = choose v (Thm.assume ((Thm.apply \<^cterm>\Trueprop\ o mk_ex ctxt v) (Thm.dest_arg (hd (Thm.chyps_of th))))) th fun mkexi v th = let val p = Thm.lambda v (Thm.dest_arg (Thm.cprop_of th)) in Thm.implies_elim (Conv.fconv_rule (Thm.beta_conversion true) (Thm.instantiate' [SOME (Thm.ctyp_of_cterm v)] [SOME p, SOME v] @{thm exI})) th end fun ex_eq_conv ctxt t = let val (p0,q0) = Thm.dest_binop t val (vs',P) = strip_exists p0 val (vs,_) = strip_exists q0 val th = Thm.assume (Thm.apply \<^cterm>\Trueprop\ P) val th1 = implies_intr_hyps (fold (simple_choose ctxt) vs' (fold mkexi vs th)) val th2 = implies_intr_hyps (fold (simple_choose ctxt) vs (fold mkexi vs' th)) val p = (Thm.dest_arg o Thm.dest_arg1 o Thm.cprop_of) th1 val q = (Thm.dest_arg o Thm.dest_arg o Thm.cprop_of) th1 in Thm.implies_elim (Thm.implies_elim (Thm.instantiate' [] [SOME p, SOME q] iffI) th1) th2 |> mk_meta_eq end; fun getname v = case Thm.term_of v of Free(s,_) => s | Var ((s,_),_) => s | _ => "x" fun mk_eq s t = Thm.apply (Thm.apply \<^cterm>\(\) :: bool \ _\ s) t fun mk_exists ctxt v th = Drule.arg_cong_rule (ext ctxt (Thm.typ_of_cterm v)) (Thm.abstract_rule (getname v) v th) fun simp_ex_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)}) fun free_in v t = Cterms.defined (Cterms.build (Drule.add_frees_cterm t)) v; val vsubst = let fun vsubst (t,v) tm = (Thm.rhs_of o Thm.beta_conversion false) (Thm.apply (Thm.lambda v tm) t) in fold vsubst end; (** main **) fun ring_and_ideal_conv {vars = _, semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), idom, ideal} dest_const mk_const ring_eq_conv ring_normalize_conv = let val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; val [ring_add_tm, ring_mul_tm, ring_pow_tm] = map Thm.dest_fun2 [add_pat, mul_pat, pow_pat]; val (ring_sub_tm, ring_neg_tm) = (case r_ops of [sub_pat, neg_pat] => (Thm.dest_fun2 sub_pat, Thm.dest_fun neg_pat) |_ => (\<^cterm>\True\, \<^cterm>\True\)); val (field_div_tm, field_inv_tm) = (case f_ops of [div_pat, inv_pat] => (Thm.dest_fun2 div_pat, Thm.dest_fun inv_pat) | _ => (\<^cterm>\True\, \<^cterm>\True\)); val [idom_thm, neq_thm] = idom; val [idl_sub, idl_add0] = if length ideal = 2 then ideal else [eq_commute, eq_commute] fun ring_dest_neg t = let val (l,r) = Thm.dest_comb t in if Term.could_unify(Thm.term_of l, Thm.term_of ring_neg_tm) then r else raise CTERM ("ring_dest_neg", [t]) end fun field_dest_inv t = let val (l,r) = Thm.dest_comb t in if Term.could_unify (Thm.term_of l, Thm.term_of field_inv_tm) then r else raise CTERM ("field_dest_inv", [t]) end val ring_dest_add = dest_binary ring_add_tm; val ring_mk_add = mk_binop ring_add_tm; val ring_dest_sub = dest_binary ring_sub_tm; val ring_dest_mul = dest_binary ring_mul_tm; val ring_mk_mul = mk_binop ring_mul_tm; val field_dest_div = dest_binary field_div_tm; val ring_dest_pow = dest_binary ring_pow_tm; val ring_mk_pow = mk_binop ring_pow_tm ; fun grobvars tm acc = if can dest_const tm then acc else if can ring_dest_neg tm then grobvars (Thm.dest_arg tm) acc else if can ring_dest_pow tm then grobvars (Thm.dest_arg1 tm) acc else if can ring_dest_add tm orelse can ring_dest_sub tm orelse can ring_dest_mul tm then grobvars (Thm.dest_arg1 tm) (grobvars (Thm.dest_arg tm) acc) else if can field_dest_inv tm then let val gvs = grobvars (Thm.dest_arg tm) [] in if null gvs then acc else tm::acc end else if can field_dest_div tm then let val lvs = grobvars (Thm.dest_arg1 tm) acc val gvs = grobvars (Thm.dest_arg tm) [] in if null gvs then lvs else tm::acc end else tm::acc ; fun grobify_term vars tm = ((if not (member (op aconvc) vars tm) then raise CTERM ("Not a variable", [tm]) else [(@1, map (fn i => if i aconvc tm then 1 else 0) vars)]) handle CTERM _ => ((let val x = dest_const tm in if x = @0 then [] else [(x,map (K 0) vars)] end) handle ERROR _ => ((grob_neg(grobify_term vars (ring_dest_neg tm))) handle CTERM _ => ( (grob_inv(grobify_term vars (field_dest_inv tm))) handle CTERM _ => ((let val (l,r) = ring_dest_add tm in grob_add (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_sub tm in grob_sub (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_mul tm in grob_mul (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ( (let val (l,r) = field_dest_div tm in grob_div (grobify_term vars l) (grobify_term vars r) end) handle CTERM _ => ((let val (l,r) = ring_dest_pow tm in grob_pow vars (grobify_term vars l) ((Thm.term_of #> HOLogic.dest_number #> snd) r) end) handle CTERM _ => error "grobify_term: unknown or invalid term"))))))))); val eq_tm = idom_thm |> concl |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_fun2; val dest_eq = dest_binary eq_tm; fun grobify_equation vars tm = let val (l,r) = dest_binary eq_tm tm in grob_sub (grobify_term vars l) (grobify_term vars r) end; fun grobify_equations tm = let val cjs = conjs tm val rawvars = fold_rev (fn eq => fn a => grobvars (Thm.dest_arg1 eq) (grobvars (Thm.dest_arg eq) a)) cjs [] val vars = sort Thm.term_ord (distinct (op aconvc) rawvars) in (vars,map (grobify_equation vars) cjs) end; val holify_polynomial = let fun holify_varpow (v,n) = if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber \<^ctyp>\nat\ n) (* FIXME *) fun holify_monomial vars (c,m) = let val xps = map holify_varpow (filter (fn (_,n) => n <> 0) (vars ~~ m)) in end_itlist ring_mk_mul (mk_const c :: xps) end fun holify_polynomial vars p = if null p then mk_const @0 else end_itlist ring_mk_add (map (holify_monomial vars) p) in holify_polynomial end ; fun idom_rule ctxt = simplify (put_simpset HOL_basic_ss ctxt addsimps [idom_thm]); fun prove_nz n = eqF_elim (ring_eq_conv(mk_binop eq_tm (mk_const n) (mk_const @0))); val neq_01 = prove_nz @1; fun neq_rule n th = [prove_nz n, th] MRS neq_thm; fun mk_add th1 = Thm.combination (Drule.arg_cong_rule ring_add_tm th1); fun refute ctxt tm = if tm aconvc false_tm then assume_Trueprop tm else ((let val (nths0,eths0) = List.partition (is_neg o concl) (HOLogic.conj_elims ctxt (assume_Trueprop tm)) val nths = filter (is_eq o Thm.dest_arg o concl) nths0 val eths = filter (is_eq o concl) eths0 in if null eths then let val th1 = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths val th2 = Conv.fconv_rule ((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1 val conc = th2 |> concl |> Thm.dest_arg val (l,_) = conc |> dest_eq in Thm.implies_intr (Thm.apply cTrp tm) (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th2)) (HOLogic.mk_obj_eq (Thm.reflexive l))) end else let val (vars,l,cert,noteqth) =( if null nths then let val (vars,pols) = grobify_equations(list_mk_conj(map concl eths)) val (l,cert) = grobner_weak vars pols in (vars,l,cert,neq_01) end else let val nth = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths val (vars,pol::pols) = grobify_equations(list_mk_conj(Thm.dest_arg(concl nth)::map concl eths)) val (deg,l,cert) = grobner_strong vars pols pol val th1 = Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) nth val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr ctxt th1) neq_01 in (vars,l,cert,th2) end) val cert_pos = map (fn (i,p) => (i,filter (fn (c,_) => c > @0) p)) cert val cert_neg = map (fn (i,p) => (i,map (fn (c,m) => (~ c,m)) (filter (fn (c,_) => c < @0) p))) cert val herts_pos = map (fn (i,p) => (i,holify_polynomial vars p)) cert_pos val herts_neg = map (fn (i,p) => (i,holify_polynomial vars p)) cert_neg fun thm_fn pols = if null pols then Thm.reflexive(mk_const @0) else end_itlist mk_add (map (fn (i,p) => Drule.arg_cong_rule (Thm.apply ring_mul_tm p) (nth eths i |> mk_meta_eq)) pols) val th1 = thm_fn herts_pos val th2 = thm_fn herts_neg val th3 = HOLogic.conj_intr ctxt (HOLogic.mk_obj_eq (mk_add (Thm.symmetric th1) th2)) noteqth val th4 = Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv o Conv.binop_conv) ring_normalize_conv) (neq_rule l th3) val (l, _) = dest_eq(Thm.dest_arg(concl th4)) in Thm.implies_intr (Thm.apply cTrp tm) (Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th4)) (HOLogic.mk_obj_eq (Thm.reflexive l))) end end) handle ERROR _ => raise CTERM ("Groebner-refute: unable to refute",[tm])) fun ring ctxt tm = let fun mk_forall x p = let val T = Thm.typ_of_cterm x; val all = Thm.cterm_of ctxt (Const (\<^const_name>\All\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) in Thm.apply all (Thm.lambda x p) end val avs = Cterms.build (Drule.add_frees_cterm tm) val P' = fold mk_forall (Cterms.list_set_rev avs) tm val th1 = initial_conv ctxt (mk_neg P') val (evs,bod) = strip_exists(concl th1) in if is_forall bod then raise CTERM("ring: non-universal formula",[tm]) else let val th1a = weak_dnf_conv ctxt bod val boda = concl th1a val th2a = refute_disj (refute ctxt) boda val th2b = [HOLogic.mk_obj_eq th1a, (th2a COMP notI) COMP PFalse'] MRS trans val th2 = fold (fn v => fn th => (Thm.forall_intr v th) COMP allI) evs (th2b RS PFalse) val th3 = Thm.equal_elim (Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [not_ex RS sym]) (th2 |> Thm.cprop_of)) th2 in specl (Cterms.list_set_rev avs) ([[[HOLogic.mk_obj_eq th1, th3 RS PFalse'] MRS trans] MRS PFalse] MRS notnotD) end end fun ideal tms tm ord = let val rawvars = fold_rev grobvars (tm::tms) [] val vars = sort ord (distinct (fn (x,y) => (Thm.term_of x) aconv (Thm.term_of y)) rawvars) val pols = map (grobify_term vars) tms val pol = grobify_term vars tm val cert = grobner_ideal vars pols pol in map_range (fn n => these (AList.lookup (op =) cert n) |> holify_polynomial vars) (length pols) end fun poly_eq_conv t = let val (a,b) = Thm.dest_binop t in Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv ring_normalize_conv)) (Thm.instantiate' [] [SOME a, SOME b] idl_sub) end val poly_eq_simproc = let fun proc ct = let val th = poly_eq_conv ct in if Thm.is_reflexive th then NONE else SOME th end in Simplifier.cert_simproc (Thm.theory_of_thm idl_sub) "poly_eq_simproc" {lhss = [Thm.term_of (Thm.lhs_of idl_sub)], - proc = fn _ => fn _ => proc} + proc = Morphism.entity (fn _ => fn _ => proc)} end; val poly_eq_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms} addsimprocs [poly_eq_simproc]) local fun is_defined v t = let val mons = striplist(dest_binary ring_add_tm) t in member (op aconvc) mons v andalso forall (fn m => v aconvc m orelse not(Cterms.defined (Cterms.build (Drule.add_frees_cterm m)) v)) mons end fun isolate_variable vars tm = let val th = poly_eq_conv tm val th' = (sym_conv then_conv poly_eq_conv) tm val (v,th1) = case find_first(fn v=> is_defined v (Thm.dest_arg1 (Thm.rhs_of th))) vars of SOME v => (v,th') | NONE => (the (find_first (fn v => is_defined v (Thm.dest_arg1 (Thm.rhs_of th'))) vars) ,th) val th2 = Thm.transitive th1 (Thm.instantiate' [] [(SOME o Thm.dest_arg1 o Thm.rhs_of) th1, SOME v] idl_add0) in Conv.fconv_rule(funpow 2 Conv.arg_conv ring_normalize_conv) th2 end in fun unwind_polys_conv ctxt tm = let val (vars,bod) = strip_exists tm val cjs = striplist (dest_binary \<^cterm>\HOL.conj\) bod val th1 = (the (get_first (try (isolate_variable vars)) cjs) handle Option.Option => raise CTERM ("unwind_polys_conv",[tm])) val eq = Thm.lhs_of th1 val bod' = list_mk_binop \<^cterm>\HOL.conj\ (eq::(remove (op aconvc) eq cjs)) val th2 = conj_ac_rule (mk_eq bod bod') val th3 = Thm.transitive th2 (Drule.binop_cong_rule \<^cterm>\HOL.conj\ th1 (Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2)))) val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3)) val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists ctxt v th3) val th5 = ex_eq_conv ctxt (mk_eq tm (fold (mk_ex ctxt) (remove op aconvc v vars) (Thm.lhs_of th4))) in Thm.transitive th5 (fold (mk_exists ctxt) (remove op aconvc v vars) th4) end; end local fun scrub_var v m = let val ps = striplist ring_dest_mul m val ps' = remove op aconvc v ps in if null ps' then one_tm else fold1 ring_mk_mul ps' end fun find_multipliers v mons = let val mons1 = filter (fn m => free_in v m) mons val mons2 = map (scrub_var v) mons1 in if null mons2 then zero_tm else fold1 ring_mk_add mons2 end fun isolate_monomials vars tm = let val (vmons, cmons) = List.partition (fn m => let val frees = Cterms.build (Drule.add_frees_cterm m) in exists (Cterms.defined frees) vars end) (striplist ring_dest_add tm) val cofactors = map (fn v => find_multipliers v vmons) vars val cnc = if null cmons then zero_tm else Thm.apply ring_neg_tm (list_mk_binop ring_add_tm cmons) in (cofactors,cnc) end; fun isolate_variables evs ps eq = let val vars = filter (fn v => free_in v eq) evs val (qs,p) = isolate_monomials vars eq val rs = ideal (qs @ ps) p Thm.term_ord in (eq, take (length qs) rs ~~ vars) end; fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p)); in fun solve_idealism evs ps eqs = if null evs then [] else let val (eq,cfs) = get_first (try (isolate_variables evs ps)) eqs |> the val evs' = subtract op aconvc evs (map snd cfs) val eqs' = map (subst_in_poly cfs) (remove op aconvc eq eqs) in cfs @ solve_idealism evs' ps eqs' end; end; in {ring_conv = ring, simple_ideal = ideal, multi_ideal = solve_idealism, poly_eq_ss = poly_eq_ss, unwind_conv = unwind_polys_conv} end; fun find_term tm ctxt = (case Thm.term_of tm of Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args tm ctxt else (Thm.dest_arg tm, ctxt) | Const (\<^const_name>\Not\, _) $ _ => find_term (Thm.dest_arg tm) ctxt | Const (\<^const_name>\All\, _) $ _ => find_body (Thm.dest_arg tm) ctxt | Const (\<^const_name>\Ex\, _) $ _ => find_body (Thm.dest_arg tm) ctxt | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args tm ctxt | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args tm ctxt | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args tm ctxt | \<^term>\Pure.imp\ $_$_ => find_args tm ctxt | Const("(==)",_)$_$_ => find_args tm ctxt (* FIXME proper const name *) | \<^term>\Trueprop\$_ => find_term (Thm.dest_arg tm) ctxt | _ => raise TERM ("find_term", [])) and find_args tm ctxt = let val (t, u) = Thm.dest_binop tm in (find_term t ctxt handle TERM _ => find_term u ctxt) end and find_body b ctxt = let val ((_, b'), ctxt') = Variable.dest_abs_cterm b ctxt in find_term b' ctxt' end; fun get_ring_ideal_convs ctxt form = case \<^try>\find_term form ctxt\ of NONE => NONE | SOME (tm, ctxt') => (case Semiring_Normalizer.match ctxt' tm of NONE => NONE | SOME (res as (theory, {is_const = _, dest_const, mk_const, conv = ring_eq_conv})) => SOME (ring_and_ideal_conv theory dest_const (mk_const (Thm.ctyp_of_cterm tm)) (ring_eq_conv ctxt') (Semiring_Normalizer.semiring_normalize_wrapper ctxt' res))) fun presimplify ctxt add_thms del_thms = asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\algebra\) delsimps del_thms addsimps add_thms); fun ring_tac add_ths del_ths ctxt = Object_Logic.full_atomize_tac ctxt THEN' presimplify ctxt add_ths del_ths THEN' CSUBGOAL (fn (p, i) => resolve_tac ctxt [let val form = Object_Logic.dest_judgment ctxt p in case get_ring_ideal_convs ctxt form of NONE => Thm.reflexive form | SOME thy => #ring_conv thy ctxt form end] i handle TERM _ => no_tac | CTERM _ => no_tac | THM _ => no_tac); local fun lhs t = case Thm.term_of t of Const(\<^const_name>\HOL.eq\,_)$_$_ => Thm.dest_arg1 t | _=> raise CTERM ("ideal_tac - lhs",[t]) fun exitac _ NONE = no_tac | exitac ctxt (SOME y) = resolve_tac ctxt [Thm.instantiate' [SOME (Thm.ctyp_of_cterm y)] [NONE,SOME y] exI] 1 val claset = claset_of \<^context> in fun ideal_tac add_ths del_ths ctxt = presimplify ctxt add_ths del_ths THEN' CSUBGOAL (fn (p, i) => case get_ring_ideal_convs ctxt p of NONE => no_tac | SOME thy => let fun poly_exists_tac {asms = asms, concl = concl, prems = prems, params = _, context = ctxt, schematics = _} = let val (evs,bod) = strip_exists (Thm.dest_arg concl) val ps = map_filter (try (lhs o Thm.dest_arg)) asms val cfs = (map swap o #multi_ideal thy evs ps) (map Thm.dest_arg1 (conjuncts bod)) val ws = map (exitac ctxt o AList.lookup op aconvc cfs) evs in EVERY (rev ws) THEN Method.insert_tac ctxt prems 1 THEN ring_tac add_ths del_ths ctxt 1 end in clarify_tac (put_claset claset ctxt) i THEN Object_Logic.full_atomize_tac ctxt i THEN asm_full_simp_tac (put_simpset (#poly_eq_ss thy) ctxt) i THEN clarify_tac (put_claset claset ctxt) i THEN (REPEAT (CONVERSION (#unwind_conv thy ctxt) i)) THEN SUBPROOF poly_exists_tac ctxt i end handle TERM _ => no_tac | CTERM _ => no_tac | THM _ => no_tac); end; fun algebra_tac add_ths del_ths ctxt i = ring_tac add_ths del_ths ctxt i ORELSE ideal_tac add_ths del_ths ctxt i end; diff --git a/src/Pure/Isar/args.ML b/src/Pure/Isar/args.ML --- a/src/Pure/Isar/args.ML +++ b/src/Pure/Isar/args.ML @@ -1,189 +1,189 @@ (* Title: Pure/Isar/args.ML Author: Markus Wenzel, TU Muenchen Quasi-inner syntax based on outer tokens: concrete argument syntax of attributes, methods etc. *) signature ARGS = sig val context: Proof.context context_parser val theory: theory context_parser val symbolic: Token.T parser val $$$ : string -> string parser val add: string parser val del: string parser val colon: string parser val query: string parser val bang: string parser val query_colon: string parser val bang_colon: string parser val parens: 'a parser -> 'a parser val bracks: 'a parser -> 'a parser val mode: string -> bool parser val maybe: 'a parser -> 'a option parser val name_token: Token.T parser val name: string parser val name_position: (string * Position.T) parser val cartouche_inner_syntax: string parser val cartouche_input: Input.source parser val binding: binding parser val alt_name: string parser val liberal_name: string parser val var: indexname parser val internal_source: Token.src parser val internal_name: Token.name_value parser val internal_typ: typ parser val internal_term: term parser val internal_fact: thm list parser - val internal_attribute: (morphism -> attribute) parser - val internal_declaration: declaration parser + val internal_attribute: attribute Morphism.entity parser + val internal_declaration: Morphism.declaration parser val named_source: (Token.T -> Token.src) -> Token.src parser val named_typ: (string -> typ) -> typ parser val named_term: (string -> term) -> term parser val named_fact: (string -> string option * thm list) -> thm list parser - val named_attribute: (string * Position.T -> morphism -> attribute) -> - (morphism -> attribute) parser - val embedded_declaration: (Input.source -> declaration) -> declaration parser + val named_attribute: (string * Position.T -> attribute Morphism.entity) -> + attribute Morphism.entity parser + val embedded_declaration: (Input.source -> Morphism.declaration) -> Morphism.declaration parser val typ_abbrev: typ context_parser val typ: typ context_parser val term: term context_parser val term_pattern: term context_parser val term_abbrev: term context_parser val prop: term context_parser val type_name: {proper: bool, strict: bool} -> string context_parser val const: {proper: bool, strict: bool} -> string context_parser val goal_spec: ((int -> tactic) -> tactic) context_parser end; structure Args: ARGS = struct (** argument scanners **) (* context *) fun context x = (Scan.state >> Context.proof_of) x; fun theory x = (Scan.state >> Context.theory_of) x; (* basic *) val ident = Parse.token (Parse.short_ident || Parse.long_ident || Parse.sym_ident || Parse.term_var || Parse.type_ident || Parse.type_var || Parse.number); val string = Parse.token Parse.string; val alt_string = Parse.token (Parse.alt_string || Parse.cartouche); val symbolic = Parse.token (Parse.keyword_with Token.ident_or_symbolic); fun $$$ x = (ident || Parse.token Parse.keyword) :|-- (fn tok => let val y = Token.content_of tok in if x = y then (Token.assign (SOME (Token.Literal (false, Markup.quasi_keyword))) tok; Scan.succeed x) else Scan.fail end); val add = $$$ "add"; val del = $$$ "del"; val colon = $$$ ":"; val query = $$$ "?"; val bang = $$$ "!"; val query_colon = $$$ "?" ^^ $$$ ":"; val bang_colon = $$$ "!" ^^ $$$ ":"; fun parens scan = $$$ "(" |-- scan --| $$$ ")"; fun bracks scan = $$$ "[" |-- scan --| $$$ "]"; fun mode s = Scan.optional (parens ($$$ s) >> K true) false; fun maybe scan = $$$ "_" >> K NONE || scan >> SOME; val name_token = ident || string; val name = name_token >> Token.content_of; val name_position = name_token >> (Input.source_content o Token.input_of); val cartouche = Parse.token Parse.cartouche; val cartouche_inner_syntax = cartouche >> Token.inner_syntax_of; val cartouche_input = cartouche >> Token.input_of; val binding = Parse.input name >> (Binding.make o Input.source_content); val alt_name = alt_string >> Token.content_of; val liberal_name = (symbolic >> Token.content_of) || name; val var = (ident >> Token.content_of) :|-- (fn x => (case Lexicon.read_variable x of SOME v => Scan.succeed v | NONE => Scan.fail)); (* values *) fun value dest = Scan.some (fn arg => (case Token.get_value arg of SOME v => (SOME (dest v) handle Match => NONE) | NONE => NONE)); val internal_source = value (fn Token.Source src => src); val internal_name = value (fn Token.Name (a, _) => a); val internal_typ = value (fn Token.Typ T => T); val internal_term = value (fn Token.Term t => t); val internal_fact = value (fn Token.Fact (_, ths) => ths); val internal_attribute = value (fn Token.Attribute att => att); val internal_declaration = value (fn Token.Declaration decl => decl); fun named_source read = internal_source || name_token >> Token.evaluate Token.Source read; fun named_typ read = internal_typ || Parse.token Parse.embedded >> Token.evaluate Token.Typ (read o Token.inner_syntax_of); fun named_term read = internal_term || Parse.token Parse.embedded >> Token.evaluate Token.Term (read o Token.inner_syntax_of); fun named_fact get = internal_fact || name_token >> Token.evaluate Token.Fact (get o Token.content_of) >> #2 || alt_string >> Token.evaluate Token.Fact (get o Token.inner_syntax_of) >> #2; fun named_attribute att = internal_attribute || name_token >> Token.evaluate Token.Attribute (fn tok => att (Token.content_of tok, Token.pos_of tok)); fun embedded_declaration read = internal_declaration || Parse.token Parse.embedded >> Token.evaluate Token.Declaration (read o Token.input_of); (* terms and types *) val typ_abbrev = Scan.peek (named_typ o Proof_Context.read_typ_abbrev o Context.proof_of); val typ = Scan.peek (named_typ o Syntax.read_typ o Context.proof_of); val term = Scan.peek (named_term o Syntax.read_term o Context.proof_of); val term_pattern = Scan.peek (named_term o Proof_Context.read_term_pattern o Context.proof_of); val term_abbrev = Scan.peek (named_term o Proof_Context.read_term_abbrev o Context.proof_of); val prop = Scan.peek (named_term o Syntax.read_prop o Context.proof_of); (* type and constant names *) fun type_name flags = Scan.peek (named_typ o Proof_Context.read_type_name flags o Context.proof_of) >> (fn Type (c, _) => c | TFree (a, _) => a | _ => ""); fun const flags = Scan.peek (named_term o Proof_Context.read_const flags o Context.proof_of) >> (fn Const (c, _) => c | Free (x, _) => x | _ => ""); (* improper method arguments *) val from_to = Parse.nat -- ($$$ "-" |-- Parse.nat) >> (fn (i, j) => fn tac => Seq.INTERVAL tac i j) || Parse.nat --| $$$ "-" >> (fn i => fn tac => fn st => Seq.INTERVAL tac i (Thm.nprems_of st) st) || Parse.nat >> (fn i => fn tac => tac i) || $$$ "!" >> K ALLGOALS; val goal = Parse.keyword_improper "[" |-- Parse.!!! (from_to --| Parse.keyword_improper "]"); fun goal_spec x = Scan.lift (Scan.optional goal (fn tac => tac 1)) x; end; 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,662 +1,661 @@ (* 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: declaration -> thms + val internal_declaration: Morphism.declaration -> 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 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 internal_name ctxt name = Token.make_src (name, Position.none) [] |> check_src ctxt |> hd; local val _ = Theory.setup (setup (Binding.make ("attribute", \<^here>)) (Scan.lift Args.internal_attribute >> Morphism.form || Scan.lift Args.internal_declaration >> (Thm.declaration_attribute o K o Morphism.form)) "internal attribute"); val internal_attribute_name = internal_name (Context.the_local_context ()) "attribute"; fun internal_source value = internal_attribute_name :: [Token.make_string ("", Position.none) |> Token.assign (SOME value)]; in -fun internal att = internal_source (Token.Attribute att); - -fun internal_declaration decl = - [([Drule.dummy_thm], [internal_source (Token.Declaration decl)])]; +fun internal arg = internal_source (Token.Attribute (Morphism.entity arg)); +fun internal_declaration arg = [([Drule.dummy_thm], [internal_source (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) "configuration option" + |> 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 internal = internal_name (Context.the_local_context ()); val consumes_name = internal "consumes"; val constraints_name = internal "constraints"; val cases_open_name = internal "cases_open"; val case_names_name = internal "case_names"; val case_conclusion_name = internal "case_conclusion"; fun make_string s = Token.make_string (s, Position.none); 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 make_string names; fun case_conclusion (name, names) = case_conclusion_name :: map make_string (name :: names); end; end; \ No newline at end of file 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,479 +1,480 @@ (* 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: declaration -> local_theory -> local_theory + val background_declaration: Morphism.declaration -> 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 -> Context.generic -> Context.generic) -> local_theory -> local_theory + val standard_declaration: (int * int -> bool) -> Morphism.declaration -> + 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: declaration -> local_theory -> local_theory + val theory_declaration: Morphism.declaration -> 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 -> declaration -> local_theory -> local_theory + val locale_target_declaration: string -> bool -> Morphism.declaration -> + 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} -> declaration -> + val locale_declaration: string -> {syntax: bool, pervasive: bool} -> Morphism.declaration -> 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) phi context = +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 (Proof_Context.generic_notation true prmode [(rhs', mx)]) + |> 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 (Proof_Context.generic_notation true prmode [(const, mx)])))) + Morphism.form_entity (Proof_Context.generic_notation true prmode [(const, mx)])))) end - else context; + 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 |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_declaration locale syntax (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 operations **) fun locale_declaration locale {syntax, pervasive} decl = pervasive ? background_declaration decl #> locale_target_declaration locale syntax 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/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,437 +1,439 @@ (* 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 -> (morphism -> 'a) -> 'a + 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} -> declaration -> local_theory -> local_theory + val declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration_fn -> + 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} -> declaration -> local_theory -> local_theory, + declaration: {syntax: bool, pervasive: bool} -> Morphism.declaration -> + 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.norm_export_morphism lthy ctxt $> 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; -val declaration = operation2 #declaration; +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 => 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} (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} (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} (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} (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 => 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,800 +1,800 @@ (* 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 -> - declaration list -> + Morphism.declaration 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 -> declaration -> Proof.context -> Proof.context + val add_declaration: string -> bool -> Morphism.declaration -> 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: (declaration * serial) list, + syntax_decls: (Morphism.declaration * 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 => (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; in context - |> fold_rev (fn (decl, _) => decl morph) syntax_decls + |> fold_rev (Morphism.form o Morphism.transform morph o #1) syntax_decls 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 context => let val elem' = Element.transform_ctxt (Morphism.set_context'' context morph) elem in Element.init elem' context end); val apply_registrations = Context.theory_map (fn context => fold_rev (apply_notes o #2) (registrations_of context loc) context); 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 = syntax ? Proof_Context.background_theory ((change_locale loc o apfst o apfst) (cons (decl, serial ()))) #> add_facts loc "" [(Binding.empty_atts, Attrib.internal_declaration decl)]; (*** 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,829 +1,831 @@ (* 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 - fn phi => set_tactic (fn _ => Context.setmp_generic_context (SOME context) (tac phi)) - end)) >> (fn decl => Morphism.form (the_tactic (Morphism.form decl context)))); + 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)]), bs)); - fun decl phi = - Context.mapping I init #> - Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd; + val decl = + Morphism.entity (fn phi => + Context.mapping I init #> + Attrib.generic_notes "" (Attrib.transform_facts phi facts) #> snd); 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/token.ML b/src/Pure/Isar/token.ML --- a/src/Pure/Isar/token.ML +++ b/src/Pure/Isar/token.ML @@ -1,822 +1,822 @@ (* 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 morphism -> attribute | - Declaration of declaration | + Attribute of attribute Morphism.entity | + Declaration of Morphism.declaration | 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_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 morphism -> attribute | - Declaration of declaration | + Attribute of attribute Morphism.entity | + Declaration of Morphism.declaration | 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 = 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) | v => v) tok; in token_context end; in val trim_context = context Thm.trim_context Morphism.reset_context; fun transfer thy = context (Thm.transfer thy) (Morphism.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_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)) | _ => 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; 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 src context = let 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/morphism.ML b/src/Pure/morphism.ML --- a/src/Pure/morphism.ML +++ b/src/Pure/morphism.ML @@ -1,231 +1,243 @@ (* Title: Pure/morphism.ML Author: Makarius Abstract morphisms on formal entities. *) infix 1 $> signature BASIC_MORPHISM = sig type morphism - type declaration = morphism -> Context.generic -> Context.generic val $> : morphism * morphism -> morphism end signature MORPHISM = sig include BASIC_MORPHISM exception MORPHISM of string * exn val the_theory: theory option -> theory val set_context: theory -> morphism -> morphism val set_context': Proof.context -> morphism -> morphism val set_context'': Context.generic -> morphism -> morphism val reset_context: morphism -> morphism val morphism: string -> {binding: (theory option -> binding -> binding) list, typ: (theory option -> typ -> typ) list, term: (theory option -> term -> term) list, fact: (theory option -> thm list -> thm list) list} -> morphism val is_identity: morphism -> bool val is_empty: morphism -> bool val pretty: morphism -> Pretty.T val binding: morphism -> binding -> binding val binding_prefix: morphism -> (string * bool) list val typ: morphism -> typ -> typ val term: morphism -> term -> term val fact: morphism -> thm list -> thm list val thm: morphism -> thm -> thm val cterm: morphism -> cterm -> cterm val identity: morphism val default: morphism option -> morphism val compose: morphism -> morphism -> morphism - val transform: morphism -> (morphism -> 'a) -> morphism -> 'a - val form: (morphism -> 'a) -> 'a + type 'a entity + val entity: (morphism -> 'a) -> 'a entity + val transform: morphism -> 'a entity -> 'a entity + val form: 'a entity -> 'a + val form_entity: (morphism -> 'a) -> 'a + type declaration = (Context.generic -> Context.generic) entity + type declaration_fn = morphism -> Context.generic -> Context.generic val binding_morphism: string -> (binding -> binding) -> morphism val typ_morphism': string -> (theory -> typ -> typ) -> morphism val typ_morphism: string -> (typ -> typ) -> morphism val term_morphism': string -> (theory -> term -> term) -> morphism val term_morphism: string -> (term -> term) -> morphism val fact_morphism: string -> (thm list -> thm list) -> morphism val thm_morphism: string -> (thm -> thm) -> morphism val transfer_morphism: theory -> morphism val transfer_morphism': Proof.context -> morphism val transfer_morphism'': Context.generic -> morphism val trim_context_morphism: morphism val set_trim_context: theory -> morphism -> morphism val set_trim_context': Proof.context -> morphism -> morphism val set_trim_context'': Context.generic -> morphism -> morphism val instantiate_frees_morphism: ctyp TFrees.table * cterm Frees.table -> morphism val instantiate_morphism: ctyp TVars.table * cterm Vars.table -> morphism end; structure Morphism: MORPHISM = struct (* named functions *) type 'a funs = (string * (theory option -> 'a -> 'a)) list; exception MORPHISM of string * exn; fun app context (name, f) x = f context x handle exn => if Exn.is_interrupt exn then Exn.reraise exn else raise MORPHISM (name, exn); (* optional context *) fun the_theory (SOME thy) = thy | the_theory NONE = raise Fail "Morphism lacks theory context"; fun join_transfer (SOME thy) = Thm.join_transfer thy | join_transfer NONE = I; val join_context = join_options Context.join_certificate_theory; (* type morphism *) datatype morphism = Morphism of {context: theory option, names: string list, binding: binding funs, typ: typ funs, term: term funs, fact: thm list funs}; -type declaration = morphism -> Context.generic -> Context.generic; - fun rep (Morphism args) = args; fun apply which phi = let val args = rep phi in fold_rev (app (#context args)) (which args) end; fun put_context context (Morphism {context = _, names, binding, typ, term, fact}) = Morphism {context = context, names = names, binding = binding, typ = typ, term = term, fact = fact}; val set_context = put_context o SOME; val set_context' = set_context o Proof_Context.theory_of; val set_context'' = set_context o Context.theory_of; val reset_context = put_context NONE; fun morphism a {binding, typ, term, fact} = Morphism { context = NONE, names = if a = "" then [] else [a], binding = map (pair a) binding, typ = map (pair a) typ, term = map (pair a) term, fact = map (pair a) fact}; (*syntactic test only!*) fun is_identity (Morphism {context = _, names, binding, typ, term, fact}) = null names andalso null binding andalso null typ andalso null term andalso null fact; fun is_empty phi = is_none (#context (rep phi)) andalso is_identity phi; fun pretty phi = Pretty.enum ";" "{" "}" (map Pretty.str (rev (#names (rep phi)))); val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty); val binding = apply #binding; fun binding_prefix morph = Binding.name "x" |> binding morph |> Binding.prefix_of; val typ = apply #typ; val term = apply #term; fun fact phi = map (join_transfer (#context (rep phi))) #> apply #fact phi; val thm = singleton o fact; val cterm = Drule.cterm_rule o thm; (* morphism combinators *) val identity = morphism "" {binding = [], typ = [], term = [], fact = []}; val default = the_default identity; fun compose phi1 phi2 = if is_empty phi1 then phi2 else if is_empty phi2 then phi1 else let val {context = context1, names = names1, binding = binding1, typ = typ1, term = term1, fact = fact1} = rep phi1; val {context = context2, names = names2, binding = binding2, typ = typ2, term = term2, fact = fact2} = rep phi2; in Morphism { context = join_context (context1, context2), names = names1 @ names2, binding = binding1 @ binding2, typ = typ1 @ typ2, term = term1 @ term2, fact = fact1 @ fact2} end; fun phi1 $> phi2 = compose phi2 phi1; -fun transform phi f = fn psi => f (phi $> psi); -fun form f = f identity; + +(* abstract entities *) + +datatype 'a entity = Entity of morphism -> 'a +fun entity f = Entity f; + +fun transform phi (Entity f) = Entity (fn psi => f (phi $> psi)); +fun form (Entity f) = f identity; +fun form_entity f = f identity; + +type declaration = (Context.generic -> Context.generic) entity; +type declaration_fn = morphism -> Context.generic -> Context.generic; (* concrete morphisms *) fun binding_morphism a binding = morphism a {binding = [K binding], typ = [], term = [], fact = []}; fun typ_morphism' a typ = morphism a {binding = [], typ = [typ o the_theory], term = [], fact = []}; fun typ_morphism a typ = morphism a {binding = [], typ = [K typ], term = [], fact = []}; fun term_morphism' a term = morphism a {binding = [], typ = [], term = [term o the_theory], fact = []}; fun term_morphism a term = morphism a {binding = [], typ = [], term = [K term], fact = []}; fun fact_morphism a fact = morphism a {binding = [], typ = [], term = [], fact = [K fact]}; fun thm_morphism a thm = morphism a {binding = [], typ = [], term = [], fact = [K (map thm)]}; fun transfer_morphism thy = fact_morphism "transfer" I |> set_context thy; val transfer_morphism' = transfer_morphism o Proof_Context.theory_of; val transfer_morphism'' = transfer_morphism o Context.theory_of; val trim_context_morphism = thm_morphism "trim_context" Thm.trim_context; fun set_trim_context thy phi = set_context thy phi $> trim_context_morphism; val set_trim_context' = set_trim_context o Proof_Context.theory_of; val set_trim_context'' = set_trim_context o Context.theory_of; (* instantiate *) fun instantiate_frees_morphism (cinstT, cinst) = if TFrees.is_empty cinstT andalso Frees.is_empty cinst then identity else let val instT = TFrees.map (K Thm.typ_of) cinstT; val inst = Frees.map (K Thm.term_of) cinst; in morphism "instantiate_frees" {binding = [], typ = if TFrees.is_empty instT then [] else [K (Term_Subst.instantiateT_frees instT)], term = [K (Term_Subst.instantiate_frees (instT, inst))], fact = [K (map (Thm.instantiate_frees (cinstT, cinst)))]} end; fun instantiate_morphism (cinstT, cinst) = if TVars.is_empty cinstT andalso Vars.is_empty cinst then identity else let val instT = TVars.map (K Thm.typ_of) cinstT; val inst = Vars.map (K Thm.term_of) cinst; in morphism "instantiate" {binding = [], typ = if TVars.is_empty instT then [] else [K (Term_Subst.instantiateT instT)], term = [K (Term_Subst.instantiate (instT, inst))], fact = [K (map (Thm.instantiate (cinstT, cinst)))]} end; end; structure Basic_Morphism: BASIC_MORPHISM = Morphism; open Basic_Morphism; diff --git a/src/Pure/raw_simplifier.ML b/src/Pure/raw_simplifier.ML --- a/src/Pure/raw_simplifier.ML +++ b/src/Pure/raw_simplifier.ML @@ -1,1460 +1,1460 @@ (* Title: Pure/raw_simplifier.ML Author: Tobias Nipkow and Stefan Berghofer, TU Muenchen Higher-order Simplification. *) infix 4 addsimps delsimps addsimprocs delsimprocs setloop addloop delloop setSSolver addSSolver setSolver addSolver; signature BASIC_RAW_SIMPLIFIER = sig val simp_depth_limit: int Config.T val simp_trace_depth_limit: int Config.T val simp_debug: bool Config.T val simp_trace: bool Config.T type cong_name = bool * string type rrule val mk_rrules: Proof.context -> thm list -> rrule list val eq_rrule: rrule * rrule -> bool type proc type solver val mk_solver: string -> (Proof.context -> int -> tactic) -> solver type simpset val empty_ss: simpset val merge_ss: simpset * simpset -> simpset val dest_ss: simpset -> {simps: (string * thm) list, procs: (string * term list) list, congs: (cong_name * thm) list, weak_congs: cong_name list, loopers: string list, unsafe_solvers: string list, safe_solvers: string list} type simproc val eq_simproc: simproc * simproc -> bool val cert_simproc: theory -> string -> - {lhss: term list, proc: morphism -> Proof.context -> cterm -> thm option} -> simproc + {lhss: term list, proc: (Proof.context -> cterm -> thm option) Morphism.entity} -> simproc val transform_simproc: morphism -> simproc -> simproc val simpset_of: Proof.context -> simpset val put_simpset: simpset -> Proof.context -> Proof.context val simpset_map: Proof.context -> (Proof.context -> Proof.context) -> simpset -> simpset val map_theory_simpset: (Proof.context -> Proof.context) -> theory -> theory val empty_simpset: Proof.context -> Proof.context val clear_simpset: Proof.context -> Proof.context val addsimps: Proof.context * thm list -> Proof.context val delsimps: Proof.context * thm list -> Proof.context val addsimprocs: Proof.context * simproc list -> Proof.context val delsimprocs: Proof.context * simproc list -> Proof.context val setloop: Proof.context * (Proof.context -> int -> tactic) -> Proof.context val addloop: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context val delloop: Proof.context * string -> Proof.context val setSSolver: Proof.context * solver -> Proof.context val addSSolver: Proof.context * solver -> Proof.context val setSolver: Proof.context * solver -> Proof.context val addSolver: Proof.context * solver -> Proof.context val rewrite_rule: Proof.context -> thm list -> thm -> thm val rewrite_goals_rule: Proof.context -> thm list -> thm -> thm val rewrite_goals_tac: Proof.context -> thm list -> tactic val rewrite_goal_tac: Proof.context -> thm list -> int -> tactic val prune_params_tac: Proof.context -> tactic val fold_rule: Proof.context -> thm list -> thm -> thm val fold_goals_tac: Proof.context -> thm list -> tactic val norm_hhf: Proof.context -> thm -> thm val norm_hhf_protect: Proof.context -> thm -> thm val norm_hhf_protect_without_context: thm -> thm end; signature RAW_SIMPLIFIER = sig include BASIC_RAW_SIMPLIFIER exception SIMPLIFIER of string * thm list type trace_ops val set_trace_ops: trace_ops -> theory -> theory val subgoal_tac: Proof.context -> int -> tactic val loop_tac: Proof.context -> int -> tactic val solvers: Proof.context -> solver list * solver list val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val prems_of: Proof.context -> thm list val add_simp: thm -> Proof.context -> Proof.context val del_simp: thm -> Proof.context -> Proof.context val flip_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 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 val solver: Proof.context -> solver -> int -> tactic val default_mk_sym: Proof.context -> thm -> thm option val add_prems: thm list -> Proof.context -> Proof.context val set_reorient: (Proof.context -> term list -> term -> term -> bool) -> Proof.context -> Proof.context val set_solvers: solver list -> Proof.context -> Proof.context val rewrite_cterm: bool * bool * bool -> (Proof.context -> thm -> thm option) -> Proof.context -> conv val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term val rewrite_thm: bool * bool * bool -> (Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm val generic_rewrite_goal_tac: bool * bool * bool -> (Proof.context -> tactic) -> Proof.context -> int -> tactic val rewrite: Proof.context -> bool -> thm list -> conv end; structure Raw_Simplifier: RAW_SIMPLIFIER = struct (** datatype simpset **) (* congruence rules *) type cong_name = bool * string; fun cong_name (Const (a, _)) = SOME (true, a) | cong_name (Free (a, _)) = SOME (false, a) | cong_name _ = NONE; structure Congtab = Table(type key = cong_name val ord = prod_ord bool_ord fast_string_ord); (* rewrite rules *) type rrule = {thm: thm, (*the rewrite rule*) name: string, (*name of theorem from which rewrite rule was extracted*) lhs: term, (*the left-hand side*) elhs: cterm, (*the eta-contracted lhs*) extra: bool, (*extra variables outside of elhs*) fo: bool, (*use first-order matching*) perm: bool}; (*the rewrite rule is permutative*) fun trim_context_rrule ({thm, name, lhs, elhs, extra, fo, perm}: rrule) = {thm = Thm.trim_context thm, name = name, lhs = lhs, elhs = Thm.trim_context_cterm elhs, extra = extra, fo = fo, perm = perm}; (* Remarks: - elhs is used for matching, lhs only for preservation of bound variable names; - fo is set iff either elhs is first-order (no Var is applied), in which case fo-matching is complete, or elhs is not a pattern, in which case there is nothing better to do; *) fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) = Thm.eq_thm_prop (thm1, thm2); (* FIXME: it seems that the conditions on extra variables are too liberal if prems are nonempty: does solving the prems really guarantee instantiation of all its Vars? Better: a dynamic check each time a rule is applied. *) fun rewrite_rule_extra_vars prems elhs erhs = let val elhss = elhs :: prems; val tvars = TVars.build (fold TVars.add_tvars elhss); val vars = Vars.build (fold Vars.add_vars elhss); in erhs |> Term.exists_type (Term.exists_subtype (fn TVar v => not (TVars.defined tvars v) | _ => false)) orelse erhs |> Term.exists_subterm (fn Var v => not (Vars.defined vars v) | _ => false) end; fun rrule_extra_vars elhs thm = rewrite_rule_extra_vars [] (Thm.term_of elhs) (Thm.full_prop_of thm); fun mk_rrule2 {thm, name, lhs, elhs, perm} = let val t = Thm.term_of elhs; val fo = Pattern.first_order t orelse not (Pattern.pattern t); val extra = rrule_extra_vars elhs thm; in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end; (*simple test for looping rewrite rules and stupid orientations*) fun default_reorient ctxt prems lhs rhs = rewrite_rule_extra_vars prems lhs rhs orelse is_Var (head_of lhs) orelse (* turns t = x around, which causes a headache if x is a local variable - usually it is very useful :-( is_Free rhs andalso not(is_Free lhs) andalso not(Logic.occs(rhs,lhs)) andalso not(exists_subterm is_Var lhs) orelse *) exists (fn t => Logic.occs (lhs, t)) (rhs :: prems) orelse null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs) (*the condition "null prems" is necessary because conditional rewrites with extra variables in the conditions may terminate although the rhs is an instance of the lhs; example: ?m < ?n \ f ?n \ f ?m *) orelse is_Const lhs andalso not (is_Const rhs); (* simplification procedures *) datatype proc = Proc of {name: string, lhs: term, proc: Proof.context -> cterm -> thm option, stamp: stamp}; fun eq_proc (Proc {stamp = stamp1, ...}, Proc {stamp = stamp2, ...}) = stamp1 = stamp2; (* solvers *) datatype solver = Solver of {name: string, solver: Proof.context -> int -> tactic, id: stamp}; fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()}; fun solver_name (Solver {name, ...}) = name; fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt; fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2); (* simplification sets *) (*A simpset contains data required during conversion: rules: discrimination net of rewrite rules; prems: current premises; depth: simp_depth and exceeded flag; congs: association list of congruence rules and a list of `weak' congruence constants. A congruence is `weak' if it avoids normalization of some argument. procs: discrimination net of simplification procedures (functions that prove rewrite rules on the fly); mk_rews: mk: turn simplification thms into rewrite rules; mk_cong: prepare congruence rules; mk_sym: turn \ around; mk_eq_True: turn P into P \ True; term_ord: for ordered rewriting;*) datatype simpset = Simpset of {rules: rrule Net.net, prems: thm list, depth: int * bool Unsynchronized.ref} * {congs: thm Congtab.table * cong_name list, procs: proc Net.net, mk_rews: {mk: Proof.context -> thm -> thm list, mk_cong: Proof.context -> thm -> thm, mk_sym: Proof.context -> thm -> thm option, mk_eq_True: Proof.context -> thm -> thm option, reorient: Proof.context -> term list -> term -> term -> bool}, term_ord: term ord, subgoal_tac: Proof.context -> int -> tactic, loop_tacs: (string * (Proof.context -> int -> tactic)) list, solvers: solver list * solver list}; fun internal_ss (Simpset (_, ss2)) = ss2; fun make_ss1 (rules, prems, depth) = {rules = rules, prems = prems, depth = depth}; fun map_ss1 f {rules, prems, depth} = make_ss1 (f (rules, prems, depth)); fun make_ss2 (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) = {congs = congs, procs = procs, mk_rews = mk_rews, term_ord = term_ord, subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers}; fun map_ss2 f {congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers} = make_ss2 (f (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2); fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) = {simps = Net.entries rules |> map (fn {name, thm, ...} => (name, thm)), procs = Net.entries procs |> map (fn Proc {name, lhs, stamp, ...} => ((name, lhs), stamp)) |> partition_eq (eq_snd op =) |> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)), congs = congs |> fst |> Congtab.dest, weak_congs = congs |> snd, loopers = map fst loop_tacs, unsafe_solvers = map solver_name (#1 solvers), safe_solvers = map solver_name (#2 solvers)}; (* empty *) fun init_ss depth mk_rews term_ord subgoal_tac solvers = make_simpset ((Net.empty, [], depth), ((Congtab.empty, []), Net.empty, mk_rews, term_ord, subgoal_tac, [], solvers)); fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm); val empty_ss = init_ss (0, Unsynchronized.ref false) {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [], mk_cong = K I, mk_sym = default_mk_sym, mk_eq_True = K (K NONE), reorient = default_reorient} Term_Ord.term_ord (K (K no_tac)) ([], []); (* merge *) (*NOTE: ignores some fields of 2nd simpset*) fun merge_ss (ss1, ss2) = if pointer_eq (ss1, ss2) then ss1 else let val Simpset ({rules = rules1, prems = prems1, depth = depth1}, {congs = (congs1, weak1), procs = procs1, mk_rews, term_ord, subgoal_tac, loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1; val Simpset ({rules = rules2, prems = prems2, depth = depth2}, {congs = (congs2, weak2), procs = procs2, mk_rews = _, term_ord = _, subgoal_tac = _, loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2; val rules' = Net.merge eq_rrule (rules1, rules2); val prems' = Thm.merge_thms (prems1, prems2); val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1; val congs' = Congtab.merge (K true) (congs1, congs2); val weak' = merge (op =) (weak1, weak2); val procs' = Net.merge eq_proc (procs1, procs2); val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2); val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2); val solvers' = merge eq_solver (solvers1, solvers2); in make_simpset ((rules', prems', depth'), ((congs', weak'), procs', mk_rews, term_ord, subgoal_tac, loop_tacs', (unsafe_solvers', solvers'))) end; (** context data **) structure Simpset = Generic_Data ( type T = simpset; val empty = empty_ss; val merge = merge_ss; ); val simpset_of = Simpset.get o Context.Proof; fun map_simpset f = Context.proof_map (Simpset.map f); fun map_simpset1 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (map_ss1 f ss1, ss2)); fun map_simpset2 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (ss1, map_ss2 f ss2)); fun put_simpset ss = map_simpset (K ss); fun simpset_map ctxt f ss = ctxt |> put_simpset ss |> f |> simpset_of; val empty_simpset = put_simpset empty_ss; fun map_theory_simpset f thy = let val ctxt' = f (Proof_Context.init_global thy); val thy' = Proof_Context.theory_of ctxt'; in Context.theory_map (Simpset.map (K (simpset_of ctxt'))) thy' end; fun map_ss f = Context.mapping (map_theory_simpset (f o Context_Position.not_really)) f; val clear_simpset = map_simpset (fn Simpset ({depth, ...}, {mk_rews, term_ord, subgoal_tac, solvers, ...}) => init_ss depth mk_rews term_ord subgoal_tac solvers); (* accessors for tactis *) fun subgoal_tac ctxt = (#subgoal_tac o internal_ss o simpset_of) ctxt ctxt; fun loop_tac ctxt = FIRST' (map (fn (_, tac) => tac ctxt) (rev ((#loop_tacs o internal_ss o simpset_of) ctxt))); val solvers = #solvers o internal_ss o simpset_of (* simp depth *) (* The simp_depth_limit is meant to abort infinite recursion of the simplifier early but should not terminate "normal" executions. As of 2017, 25 would suffice; 40 builds in a safety margin. *) val simp_depth_limit = Config.declare_int ("simp_depth_limit", \<^here>) (K 40); val simp_trace_depth_limit = Config.declare_int ("simp_trace_depth_limit", \<^here>) (K 1); fun inc_simp_depth ctxt = ctxt |> map_simpset1 (fn (rules, prems, (depth, exceeded)) => (rules, prems, (depth + 1, if depth = Config.get ctxt simp_trace_depth_limit then Unsynchronized.ref false else exceeded))); fun simp_depth ctxt = let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt in depth end; (* diagnostics *) exception SIMPLIFIER of string * thm list; val simp_debug = Config.declare_bool ("simp_debug", \<^here>) (K false); val simp_trace = Config.declare_bool ("simp_trace", \<^here>) (K false); fun cond_warning ctxt msg = if Context_Position.is_really_visible ctxt then warning (msg ()) else (); fun cond_tracing' ctxt flag msg = if Config.get ctxt flag then let val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt; val depth_limit = Config.get ctxt simp_trace_depth_limit; in if depth > depth_limit then if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true) else (tracing (enclose "[" "]" (string_of_int depth) ^ msg ()); exceeded := false) end else (); fun cond_tracing ctxt = cond_tracing' ctxt simp_trace; fun print_term ctxt s t = s ^ "\n" ^ Syntax.string_of_term ctxt t; fun print_thm ctxt s (name, th) = print_term ctxt (if name = "" then s else s ^ " " ^ quote name ^ ":") (Thm.full_prop_of th); (** simpset operations **) (* prems *) fun prems_of ctxt = let val Simpset ({prems, ...}, _) = simpset_of ctxt in prems end; fun add_prems ths = map_simpset1 (fn (rules, prems, depth) => (rules, ths @ prems, depth)); (* maintain simp rules *) fun del_rrule loud (rrule as {thm, elhs, ...}) ctxt = ctxt |> map_simpset1 (fn (rules, prems, depth) => (Net.delete_term eq_rrule (Thm.term_of elhs, rrule) rules, prems, depth)) handle Net.DELETE => (if not loud then () else cond_warning ctxt (fn () => print_thm ctxt "Rewrite rule not in simpset:" ("", thm)); ctxt); fun insert_rrule (rrule as {thm, name, ...}) ctxt = (cond_tracing ctxt (fn () => print_thm ctxt "Adding rewrite rule" (name, thm)); ctxt |> map_simpset1 (fn (rules, prems, depth) => let val rrule2 as {elhs, ...} = mk_rrule2 rrule; val rules' = Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule2) rules; in (rules', prems, depth) end) handle Net.INSERT => (cond_warning ctxt (fn () => print_thm ctxt "Ignoring duplicate rewrite rule:" ("", thm)); ctxt)); val vars_set = Vars.build o Vars.add_vars; local fun vperm (Var _, Var _) = true | vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t) | vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2) | vperm (t, u) = (t = u); fun var_perm (t, u) = vperm (t, u) andalso Vars.eq_set (apply2 vars_set (t, u)); in fun decomp_simp thm = let val prop = Thm.prop_of thm; val prems = Logic.strip_imp_prems prop; val concl = Drule.strip_imp_concl (Thm.cprop_of thm); val (lhs, rhs) = Thm.dest_equals concl handle TERM _ => raise SIMPLIFIER ("Rewrite rule not a meta-equality", [thm]); val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs)); val erhs = Envir.eta_contract (Thm.term_of rhs); val perm = var_perm (Thm.term_of elhs, erhs) andalso not (Thm.term_of elhs aconv erhs) andalso not (is_Var (Thm.term_of elhs)); in (prems, Thm.term_of lhs, elhs, Thm.term_of rhs, perm) end; end; fun decomp_simp' thm = let val (_, lhs, _, rhs, _) = decomp_simp thm in if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", [thm]) else (lhs, rhs) end; fun mk_eq_True ctxt (thm, name) = let val Simpset (_, {mk_rews = {mk_eq_True, ...}, ...}) = simpset_of ctxt in (case mk_eq_True ctxt thm of NONE => [] | SOME eq_True => let val (_, lhs, elhs, _, _) = decomp_simp eq_True; in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end) end; (*create the rewrite rule and possibly also the eq_True variant, in case there are extra vars on the rhs*) fun rrule_eq_True ctxt thm name lhs elhs rhs thm2 = let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in if rewrite_rule_extra_vars [] lhs rhs then mk_eq_True ctxt (thm2, name) @ [rrule] else [rrule] end; fun mk_rrule ctxt (thm, name) = let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] else (*weak test for loops*) if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (Thm.term_of elhs) then mk_eq_True ctxt (thm, name) else rrule_eq_True ctxt thm name lhs elhs rhs thm end |> map (fn {thm, name, lhs, elhs, perm} => {thm = Thm.trim_context thm, name = name, lhs = lhs, elhs = Thm.trim_context_cterm elhs, perm = perm}); fun orient_rrule ctxt (thm, name) = let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm; val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = simpset_of ctxt; in if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}] else if reorient ctxt prems lhs rhs then if reorient ctxt prems rhs lhs then mk_eq_True ctxt (thm, name) else (case mk_sym ctxt thm of NONE => [] | SOME thm' => let val (_, lhs', elhs', rhs', _) = decomp_simp thm' in rrule_eq_True ctxt thm' name lhs' elhs' rhs' thm end) else rrule_eq_True ctxt thm name lhs elhs rhs thm end; fun extract_rews ctxt sym thms = let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt; val mk = if sym then fn ctxt => fn th => (mk ctxt th) RL [Drule.symmetric_thm] else mk in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end; fun extract_safe_rrules ctxt thm = maps (orient_rrule ctxt) (extract_rews ctxt false [thm]); fun mk_rrules ctxt thms = let val rews = extract_rews ctxt false thms val raw_rrules = flat (map (mk_rrule ctxt) rews) in map mk_rrule2 raw_rrules end (* add/del rules explicitly *) local fun comb_simps ctxt comb mk_rrule sym thms = let val rews = extract_rews ctxt sym (map (Thm.transfer' ctxt) thms); in fold (fold comb o mk_rrule) rews ctxt end; (* This code checks if the symetric version of a rule is already in the simpset. However, the variable names in the two versions of the rule may differ. Thus the current test modulo eq_rrule is too weak to be useful and needs to be refined. fun present ctxt rules (rrule as {thm, elhs, ...}) = (Net.insert_term eq_rrule (Thm.term_of elhs, trim_context_rrule rrule) rules; false) handle Net.INSERT => (cond_warning ctxt (fn () => print_thm ctxt "Symmetric rewrite rule already in simpset:" ("", thm)); true); fun sym_present ctxt thms = let val rews = extract_rews ctxt true (map (Thm.transfer' ctxt) thms); val rrules = map mk_rrule2 (flat(map (mk_rrule ctxt) rews)) val Simpset({rules, ...},_) = simpset_of ctxt in exists (present ctxt rules) rrules end *) in fun ctxt addsimps thms = comb_simps ctxt insert_rrule (mk_rrule ctxt) false thms; fun addsymsimps ctxt thms = comb_simps ctxt insert_rrule (mk_rrule ctxt) true thms; fun ctxt delsimps thms = comb_simps ctxt (del_rrule true) (map mk_rrule2 o mk_rrule ctxt) false thms; fun delsimps_quiet ctxt thms = comb_simps ctxt (del_rrule false) (map mk_rrule2 o mk_rrule ctxt) false thms; fun add_simp thm ctxt = ctxt addsimps [thm]; (* with check for presence of symmetric version: if sym_present ctxt [thm] then (cond_warning ctxt (fn () => print_thm ctxt "Ignoring rewrite rule:" ("", thm)); ctxt) else ctxt addsimps [thm]; *) fun del_simp thm ctxt = ctxt delsimps [thm]; fun flip_simp thm ctxt = addsymsimps (delsimps_quiet ctxt [thm]) [thm]; end; fun init_simpset thms ctxt = ctxt |> Context_Position.set_visible false |> empty_simpset |> fold add_simp thms |> Context_Position.restore_visible ctxt; (* congs *) local fun is_full_cong_prems [] [] = true | is_full_cong_prems [] _ = false | is_full_cong_prems (p :: prems) varpairs = (case Logic.strip_assums_concl p of Const ("Pure.eq", _) $ lhs $ rhs => let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in is_Var x andalso forall is_Bound xs andalso not (has_duplicates (op =) xs) andalso xs = ys andalso member (op =) varpairs (x, y) andalso is_full_cong_prems prems (remove (op =) (x, y) varpairs) end | _ => false); fun is_full_cong thm = let val prems = Thm.prems_of thm and concl = Thm.concl_of thm; val (lhs, rhs) = Logic.dest_equals concl; val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs; in f = g andalso not (has_duplicates (op =) (xs @ ys)) andalso length xs = length ys andalso is_full_cong_prems prems (xs ~~ ys) end; fun mk_cong ctxt = let val Simpset (_, {mk_rews = {mk_cong = f, ...}, ...}) = simpset_of ctxt in f ctxt end; in fun add_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]); (*val lhs = Envir.eta_contract lhs;*) val a = the (cong_name (head_of lhs)) handle Option.Option => raise SIMPLIFIER ("Congruence must start with a constant or free variable", [thm]); val (xs, weak) = congs; val xs' = Congtab.update (a, Thm.trim_context thm) xs; val weak' = if is_full_cong thm then weak else a :: weak; in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end); fun del_eqcong thm ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val (lhs, _) = Logic.dest_equals (Thm.concl_of thm) handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]); (*val lhs = Envir.eta_contract lhs;*) val a = the (cong_name (head_of lhs)) handle Option.Option => raise SIMPLIFIER ("Congruence must start with a constant", [thm]); val (xs, _) = congs; val xs' = Congtab.delete_safe a xs; val weak' = Congtab.fold (fn (a, th) => if is_full_cong th then I else insert (op =) a) xs' []; in ((xs', weak'), procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) end); fun add_cong thm ctxt = add_eqcong (mk_cong ctxt thm) ctxt; fun del_cong thm ctxt = del_eqcong (mk_cong ctxt thm) ctxt; end; (* simprocs *) datatype simproc = Simproc of {name: string, lhss: term list, - proc: morphism -> Proof.context -> cterm -> thm option, + proc: (Proof.context -> cterm -> thm option) Morphism.entity, stamp: stamp}; fun eq_simproc (Simproc {stamp = stamp1, ...}, Simproc {stamp = stamp2, ...}) = stamp1 = stamp2; fun cert_simproc thy name {lhss, proc} = Simproc {name = name, lhss = map (Sign.cert_term thy) lhss, proc = proc, stamp = stamp ()}; fun transform_simproc phi (Simproc {name, lhss, proc, stamp}) = Simproc {name = name, lhss = map (Morphism.term phi) lhss, proc = Morphism.transform phi proc, stamp = stamp}; local fun add_proc (proc as Proc {name, lhs, ...}) ctxt = (cond_tracing ctxt (fn () => print_term ctxt ("Adding simplification procedure " ^ quote name ^ " for") lhs); ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, Net.insert_term eq_proc (lhs, proc) procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)) handle Net.INSERT => (cond_warning ctxt (fn () => "Ignoring duplicate simplification procedure " ^ quote name); ctxt)); fun del_proc (proc as Proc {name, lhs, ...}) ctxt = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, Net.delete_term eq_proc (lhs, proc) procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)) handle Net.DELETE => (cond_warning ctxt (fn () => "Simplification procedure " ^ quote name ^ " not in simpset"); ctxt); fun prep_procs (Simproc {name, lhss, proc, stamp}) = lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, stamp = stamp}); in fun ctxt addsimprocs ps = fold (fold add_proc o prep_procs) ps ctxt; fun ctxt delsimprocs ps = fold (fold del_proc o prep_procs) ps ctxt; end; (* mk_rews *) local fun map_mk_rews f = map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => let val {mk, mk_cong, mk_sym, mk_eq_True, reorient} = mk_rews; val (mk', mk_cong', mk_sym', mk_eq_True', reorient') = f (mk, mk_cong, mk_sym, mk_eq_True, reorient); val mk_rews' = {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True', reorient = reorient'}; in (congs, procs, mk_rews', term_ord, subgoal_tac, loop_tacs, solvers) end); in fun mksimps ctxt = let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt in mk ctxt end; fun set_mksimps mk = map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mkcong mk_cong = map_mk_rews (fn (mk, _, mk_sym, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mksym mk_sym = map_mk_rews (fn (mk, mk_cong, _, mk_eq_True, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_mkeqTrue mk_eq_True = map_mk_rews (fn (mk, mk_cong, mk_sym, _, reorient) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); fun set_reorient reorient = map_mk_rews (fn (mk, mk_cong, mk_sym, mk_eq_True, _) => (mk, mk_cong, mk_sym, mk_eq_True, reorient)); end; (* term_ord *) fun set_term_ord term_ord = map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); (* tactics *) fun set_subgoaler subgoal_tac = map_simpset2 (fn (congs, procs, mk_rews, term_ord, _, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers)); fun ctxt setloop tac = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, _, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, [("", tac)], solvers)); fun ctxt addloop (name, tac) = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, AList.update (op =) (name, tac) loop_tacs, solvers)); fun ctxt delloop name = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, solvers) => (congs, procs, mk_rews, term_ord, subgoal_tac, (if AList.defined (op =) loop_tacs name then () else cond_warning ctxt (fn () => "No such looper in simpset: " ^ quote name); AList.delete (op =) name loop_tacs), solvers)); fun ctxt setSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, _)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, [solver]))); fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers))); fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, ([solver], solvers))); fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers))); fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, term_ord, subgoal_tac, loop_tacs, (solvers, solvers))); (* trace operations *) type trace_ops = {trace_invoke: {depth: int, term: term} -> Proof.context -> Proof.context, trace_apply: {unconditional: bool, term: term, thm: thm, rrule: rrule} -> Proof.context -> (Proof.context -> (thm * term) option) -> (thm * term) option}; structure Trace_Ops = Theory_Data ( type T = trace_ops; val empty: T = {trace_invoke = fn _ => fn ctxt => ctxt, trace_apply = fn _ => fn ctxt => fn cont => cont ctxt}; fun merge (trace_ops, _) = trace_ops; ); val set_trace_ops = Trace_Ops.put; val trace_ops = Trace_Ops.get o Proof_Context.theory_of; fun trace_invoke args ctxt = #trace_invoke (trace_ops ctxt) args ctxt; fun trace_apply args ctxt = #trace_apply (trace_ops ctxt) args ctxt; (** rewriting **) (* Uses conversions, see: L C Paulson, A higher-order implementation of rewriting, Science of Computer Programming 3 (1983), pages 119-149. *) fun check_conv ctxt msg thm thm' = let val thm'' = Thm.transitive thm thm' handle THM _ => let val nthm' = Thm.transitive (Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm' in Thm.transitive thm nthm' handle THM _ => let val nthm = Thm.transitive thm (Drule.beta_eta_conversion (Thm.rhs_of thm)) in Thm.transitive nthm nthm' end end val _ = if msg then cond_tracing ctxt (fn () => print_thm ctxt "SUCCEEDED" ("", thm')) else (); in SOME thm'' end handle THM _ => let val _ $ _ $ prop0 = Thm.prop_of thm; val _ = cond_tracing ctxt (fn () => print_thm ctxt "Proved wrong theorem (bad subgoaler?)" ("", thm') ^ "\n" ^ print_term ctxt "Should have proved:" prop0); in NONE end; (* mk_procrule *) fun mk_procrule ctxt thm = let val (prems, lhs, elhs, rhs, _) = decomp_simp thm val thm' = Thm.close_derivation \<^here> thm; in if rewrite_rule_extra_vars prems lhs rhs then (cond_warning ctxt (fn () => print_thm ctxt "Extra vars on rhs:" ("", thm)); []) else [mk_rrule2 {thm = thm', name = "", lhs = lhs, elhs = elhs, perm = false}] end; (* rewritec: conversion to apply the meta simpset to a term *) (*Since the rewriting strategy is bottom-up, we avoid re-normalizing already normalized terms by carrying around the rhs of the rewrite rule just applied. This is called the `skeleton'. It is decomposed in parallel with the term. Once a Var is encountered, the corresponding term is already in normal form. skel0 is a dummy skeleton that is to enforce complete normalization.*) val skel0 = Bound 0; (*Use rhs as skeleton only if the lhs does not contain unnormalized bits. The latter may happen iff there are weak congruence rules for constants in the lhs.*) fun uncond_skel ((_, weak), (lhs, rhs)) = if null weak then rhs (*optimization*) else if exists_subterm (fn Const (a, _) => member (op =) weak (true, a) | Free (a, _) => member (op =) weak (false, a) | _ => false) lhs then skel0 else rhs; (*Behaves like unconditional rule if rhs does not contain vars not in the lhs. Otherwise those vars may become instantiated with unnormalized terms while the premises are solved.*) fun cond_skel (args as (_, (lhs, rhs))) = if Vars.subset (vars_set rhs, vars_set lhs) then uncond_skel args else skel0; (* Rewriting -- we try in order: (1) beta reduction (2) unconditional rewrite rules (3) conditional rewrite rules (4) simplification procedures IMPORTANT: rewrite rules must not introduce new Vars or TVars! *) fun rewritec (prover, maxt) ctxt t = let val thy = Proof_Context.theory_of ctxt; val Simpset ({rules, ...}, {congs, procs, term_ord, ...}) = simpset_of ctxt; val eta_thm = Thm.eta_conversion t; val eta_t' = Thm.rhs_of eta_thm; val eta_t = Thm.term_of eta_t'; fun rew rrule = let val {thm = thm0, name, lhs, elhs = elhs0, extra, fo, perm} = rrule; val thm = Thm.transfer thy thm0; val elhs = Thm.transfer_cterm thy elhs0; val prop = Thm.prop_of thm; val (rthm, elhs') = if maxt = ~1 orelse not extra then (thm, elhs) else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs); val insts = if fo then Thm.first_order_match (elhs', eta_t') else Thm.match (elhs', eta_t'); val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm); val prop' = Thm.prop_of thm'; val unconditional = Logic.no_prems prop'; val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop'); val trace_args = {unconditional = unconditional, term = eta_t, thm = thm', rrule = rrule}; in if perm andalso is_greater_equal (term_ord (rhs', lhs')) then (cond_tracing ctxt (fn () => print_thm ctxt "Cannot apply permutative rewrite rule" (name, thm) ^ "\n" ^ print_thm ctxt "Term does not become smaller:" ("", thm')); NONE) else (cond_tracing ctxt (fn () => print_thm ctxt "Applying instance of rewrite rule" (name, thm)); if unconditional then (cond_tracing ctxt (fn () => print_thm ctxt "Rewriting:" ("", thm')); trace_apply trace_args ctxt (fn ctxt' => let val lr = Logic.dest_equals prop; val SOME thm'' = check_conv ctxt' false eta_thm thm'; in SOME (thm'', uncond_skel (congs, lr)) end)) else (cond_tracing ctxt (fn () => print_thm ctxt "Trying to rewrite:" ("", thm')); if simp_depth ctxt > Config.get ctxt simp_depth_limit then (cond_tracing ctxt (fn () => "simp_depth_limit exceeded - giving up"); NONE) else trace_apply trace_args ctxt (fn ctxt' => (case prover ctxt' thm' of NONE => (cond_tracing ctxt' (fn () => print_thm ctxt' "FAILED" ("", thm')); NONE) | SOME thm2 => (case check_conv ctxt' true eta_thm thm2 of NONE => NONE | SOME thm2' => let val concl = Logic.strip_imp_concl prop; val lr = Logic.dest_equals concl; in SOME (thm2', cond_skel (congs, lr)) end))))) end; fun rews [] = NONE | rews (rrule :: rrules) = let val opt = rew rrule handle Pattern.MATCH => NONE in (case opt of NONE => rews rrules | some => some) end; fun sort_rrules rrs = let fun is_simple ({thm, ...}: rrule) = (case Thm.prop_of thm of Const ("Pure.eq", _) $ _ $ _ => true | _ => false); fun sort [] (re1, re2) = re1 @ re2 | sort (rr :: rrs) (re1, re2) = if is_simple rr then sort rrs (rr :: re1, re2) else sort rrs (re1, rr :: re2); in sort rrs ([], []) end; fun proc_rews [] = NONE | proc_rews (Proc {name, proc, lhs, ...} :: ps) = if Pattern.matches (Proof_Context.theory_of ctxt) (lhs, Thm.term_of t) then (cond_tracing' ctxt simp_debug (fn () => print_term ctxt ("Trying procedure " ^ quote name ^ " on:") eta_t); (case proc ctxt eta_t' of NONE => (cond_tracing' ctxt simp_debug (fn () => "FAILED"); proc_rews ps) | SOME raw_thm => (cond_tracing ctxt (fn () => print_thm ctxt ("Procedure " ^ quote name ^ " produced rewrite rule:") ("", raw_thm)); (case rews (mk_procrule ctxt raw_thm) of NONE => (cond_tracing ctxt (fn () => print_term ctxt ("IGNORED result of simproc " ^ quote name ^ " -- does not match") (Thm.term_of t)); proc_rews ps) | some => some)))) else proc_rews ps; in (case eta_t of Abs _ $ _ => SOME (Thm.transitive eta_thm (Thm.beta_conversion false eta_t'), skel0) | _ => (case rews (sort_rrules (Net.match_term rules eta_t)) of NONE => proc_rews (Net.match_term procs eta_t) | some => some)) end; (* conversion to apply a congruence rule to a term *) fun congc prover ctxt maxt cong t = let val rthm = Thm.incr_indexes (maxt + 1) cong; val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (Thm.cprop_of rthm))); val insts = Thm.match (rlhs, t) (* Thm.match can raise Pattern.MATCH; is handled when congc is called *) val thm' = Thm.instantiate insts (Thm.rename_boundvars (Thm.term_of rlhs) (Thm.term_of t) rthm); val _ = cond_tracing ctxt (fn () => print_thm ctxt "Applying congruence rule:" ("", thm')); fun err (msg, thm) = (cond_tracing ctxt (fn () => print_thm ctxt msg ("", thm)); NONE); in (case prover thm' of NONE => err ("Congruence proof failed. Could not prove", thm') | SOME thm2 => (case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of NONE => err ("Congruence proof failed. Should not have proved", thm2) | SOME thm2' => if op aconv (apply2 Thm.term_of (Thm.dest_equals (Thm.cprop_of thm2'))) then NONE else SOME thm2')) end; val vA = (("A", 0), propT); val vB = (("B", 0), propT); val vC = (("C", 0), propT); fun transitive1 NONE NONE = NONE | transitive1 (SOME thm1) NONE = SOME thm1 | transitive1 NONE (SOME thm2) = SOME thm2 | transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2); fun transitive2 thm = transitive1 (SOME thm); fun transitive3 thm = transitive1 thm o SOME; fun bottomc ((simprem, useprem, mutsimp), prover, maxidx) = let fun botc skel ctxt t = if is_Var skel then NONE else (case subc skel ctxt t of some as SOME thm1 => (case rewritec (prover, maxidx) ctxt (Thm.rhs_of thm1) of SOME (thm2, skel2) => transitive2 (Thm.transitive thm1 thm2) (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => some) | NONE => (case rewritec (prover, maxidx) ctxt t of SOME (thm2, skel2) => transitive2 thm2 (botc skel2 ctxt (Thm.rhs_of thm2)) | NONE => NONE)) and try_botc ctxt t = (case botc skel0 ctxt t of SOME trec1 => trec1 | NONE => Thm.reflexive t) and subc skel ctxt t0 = let val Simpset (_, {congs, ...}) = simpset_of ctxt in (case Thm.term_of t0 of Abs (a, _, _) => let val ((v, t'), ctxt') = Variable.dest_abs_cterm t0 ctxt; val skel' = (case skel of Abs (_, _, sk) => sk | _ => skel0); in (case botc skel' ctxt' t' of SOME thm => SOME (Thm.abstract_rule a v thm) | NONE => NONE) end | t $ _ => (case t of Const ("Pure.imp", _) $ _ => impc t0 ctxt | Abs _ => let val thm = Thm.beta_conversion false t0 in (case subc skel0 ctxt (Thm.rhs_of thm) of NONE => SOME thm | SOME thm' => SOME (Thm.transitive thm thm')) end | _ => let fun appc () = let val (tskel, uskel) = (case skel of tskel $ uskel => (tskel, uskel) | _ => (skel0, skel0)); val (ct, cu) = Thm.dest_comb t0; in (case botc tskel ctxt ct of SOME thm1 => (case botc uskel ctxt cu of SOME thm2 => SOME (Thm.combination thm1 thm2) | NONE => SOME (Thm.combination thm1 (Thm.reflexive cu))) | NONE => (case botc uskel ctxt cu of SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1) | NONE => NONE)) end; val (h, ts) = strip_comb t; in (case cong_name h of SOME a => (case Congtab.lookup (fst congs) a of NONE => appc () | SOME cong => (*post processing: some partial applications h t1 ... tj, j <= length ts, may be a redex. Example: map (\x. x) = (\xs. xs) wrt map_cong*) (let val thm = congc (prover ctxt) ctxt maxidx cong t0; val t = the_default t0 (Option.map Thm.rhs_of thm); val (cl, cr) = Thm.dest_comb t val dVar = Var(("", 0), dummyT) val skel = list_comb (h, replicate (length ts) dVar) in (case botc skel ctxt cl of NONE => thm | SOME thm' => transitive3 thm (Thm.combination thm' (Thm.reflexive cr))) end handle Pattern.MATCH => appc ())) | _ => appc ()) end) | _ => NONE) end and impc ct ctxt = if mutsimp then mut_impc0 [] ct [] [] ctxt else nonmut_impc ct ctxt and rules_of_prem prem ctxt = if maxidx_of_term (Thm.term_of prem) <> ~1 then (cond_tracing ctxt (fn () => print_term ctxt "Cannot add premise as rewrite rule because it contains (type) unknowns:" (Thm.term_of prem)); (([], NONE), ctxt)) else let val (asm, ctxt') = Thm.assume_hyps prem ctxt in ((extract_safe_rrules ctxt' asm, SOME asm), ctxt') end and add_rrules (rrss, asms) ctxt = (fold o fold) insert_rrule rrss ctxt |> add_prems (map_filter I asms) and disch r prem eq = let val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq); val eq' = Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, lhs) (vC, rhs)) Drule.imp_cong) (Thm.implies_intr prem eq); in if not r then eq' else let val (prem', concl) = Thm.dest_implies lhs; val (prem'', _) = Thm.dest_implies rhs; in Thm.transitive (Thm.transitive (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem') (vB, prem) (vC, concl)) Drule.swap_prems_eq) eq') (Thm.instantiate (TVars.empty, Vars.make3 (vA, prem) (vB, prem'') (vC, concl)) Drule.swap_prems_eq) end end and rebuild [] _ _ _ _ eq = eq | rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq = let val ctxt' = add_rrules (rev rrss, rev asms) ctxt; val concl' = Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq)); val dprem = Option.map (disch false prem); in (case rewritec (prover, maxidx) ctxt' concl' of NONE => rebuild prems concl' rrss asms ctxt (dprem eq) | SOME (eq', _) => transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq'))) (mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ctxt)) end and mut_impc0 prems concl rrss asms ctxt = let val prems' = strip_imp_prems concl; val ((rrss', asms'), ctxt') = fold_map rules_of_prem prems' ctxt |>> split_list; in mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss') (asms @ asms') [] [] [] [] ctxt' ~1 ~1 end and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k = transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1 (Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE) (if changed > 0 then mut_impc (rev prems') concl (rev rrss') (rev asms') [] [] [] [] ctxt ~1 changed else rebuild prems' concl rrss' asms' ctxt (botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl)) | mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms) prems' rrss' asms' eqns ctxt changed k = (case (if k = 0 then NONE else botc skel0 (add_rrules (rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of NONE => mut_impc prems concl rrss asms (prem :: prems') (rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed (if k = 0 then 0 else k - 1) | SOME eqn => let val prem' = Thm.rhs_of eqn; val tprems = map Thm.term_of prems; val i = 1 + fold Integer.max (map (fn p => find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1; val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt; in mut_impc prems concl rrss asms (prem' :: prems') (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true) (take i prems) (Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies (drop i prems, concl))))) :: eqns) ctxt' (length prems') ~1 end) (*legacy code -- only for backwards compatibility*) and nonmut_impc ct ctxt = let val (prem, conc) = Thm.dest_implies ct; val thm1 = if simprem then botc skel0 ctxt prem else NONE; val prem1 = the_default prem (Option.map Thm.rhs_of thm1); val ctxt1 = if not useprem then ctxt else let val ((rrs, asm), ctxt') = rules_of_prem prem1 ctxt in add_rrules ([rrs], [asm]) ctxt' end; in (case botc skel0 ctxt1 conc of NONE => (case thm1 of NONE => NONE | SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc))) | SOME thm2 => let val thm2' = disch false prem1 thm2 in (case thm1 of NONE => SOME thm2' | SOME thm1' => SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2')) end) end; in try_botc end; (* Meta-rewriting: rewrites t to u and returns the theorem t \ u *) (* Parameters: mode = (simplify A, use A in simplifying B, use prems of B (if B is again a meta-impl.) to simplify A) when simplifying A \ B prover: how to solve premises in conditional rewrites and congruences *) fun rewrite_cterm mode prover raw_ctxt raw_ct = let val thy = Proof_Context.theory_of raw_ctxt; val ct = raw_ct |> Thm.transfer_cterm thy |> Thm.adjust_maxidx_cterm ~1; val maxidx = Thm.maxidx_of_cterm ct; val ctxt = raw_ctxt |> Variable.set_body true |> Context_Position.set_visible false |> inc_simp_depth |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = Thm.term_of ct} ctxt); val _ = cond_tracing ctxt (fn () => print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (Thm.term_of ct)); in ct |> bottomc (mode, Option.map (Drule.flexflex_unique (SOME ctxt)) oo prover, maxidx) ctxt |> Thm.solve_constraints end; val simple_prover = SINGLE o (fn ctxt => ALLGOALS (resolve_tac ctxt (prems_of ctxt))); fun rewrite _ _ [] = Thm.reflexive | rewrite ctxt full thms = rewrite_cterm (full, false, false) simple_prover (init_simpset thms ctxt); fun rewrite_rule ctxt = Conv.fconv_rule o rewrite ctxt true; (*simple term rewriting -- no proof*) fun rewrite_term thy rules procs = Pattern.rewrite_term thy (map decomp_simp' rules) procs; fun rewrite_thm mode prover ctxt = Conv.fconv_rule (rewrite_cterm mode prover ctxt); (*Rewrite the subgoals of a proof state (represented by a theorem)*) fun rewrite_goals_rule ctxt thms th = Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover (init_simpset thms ctxt))) th; (** meta-rewriting tactics **) (*Rewrite all subgoals*) fun rewrite_goals_tac ctxt defs = PRIMITIVE (rewrite_goals_rule ctxt defs); (*Rewrite one subgoal*) fun generic_rewrite_goal_tac mode prover_tac ctxt i thm = if 0 < i andalso i <= Thm.nprems_of thm then Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ctxt) i thm) else Seq.empty; fun rewrite_goal_tac ctxt thms = generic_rewrite_goal_tac (true, false, false) (K no_tac) (init_simpset thms ctxt); (*Prunes all redundant parameters from the proof state by rewriting.*) fun prune_params_tac ctxt = rewrite_goals_tac ctxt [Drule.triv_forall_equality]; (* for folding definitions, handling critical pairs *) (*The depth of nesting in a term*) fun term_depth (Abs (_, _, t)) = 1 + term_depth t | term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t) | term_depth _ = 0; val lhs_of_thm = #1 o Logic.dest_equals o Thm.prop_of; (*folding should handle critical pairs! E.g. K \ Inl 0, S \ Inr (Inl 0) Returns longest lhs first to avoid folding its subexpressions.*) fun sort_lhs_depths defs = let val keylist = AList.make (term_depth o lhs_of_thm) defs val keys = sort_distinct (rev_order o int_ord) (map #2 keylist) in map (AList.find (op =) keylist) keys end; val rev_defs = sort_lhs_depths o map Thm.symmetric; fun fold_rule ctxt defs = fold (rewrite_rule ctxt) (rev_defs defs); fun fold_goals_tac ctxt defs = EVERY (map (rewrite_goals_tac ctxt) (rev_defs defs)); (* HHF normal form: \ before \, outermost \ generalized *) local fun gen_norm_hhf protect ss ctxt0 th0 = let val (ctxt, th) = Thm.join_transfer_context (ctxt0, th0); val th' = if Drule.is_norm_hhf protect (Thm.prop_of th) then th else Conv.fconv_rule (rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset ss ctxt)) th; in th' |> Thm.adjust_maxidx_thm ~1 |> Variable.gen_all ctxt end; val hhf_ss = Context.the_local_context () |> init_simpset Drule.norm_hhf_eqs |> simpset_of; val hhf_protect_ss = Context.the_local_context () |> init_simpset Drule.norm_hhf_eqs |> add_eqcong Drule.protect_cong |> simpset_of; in val norm_hhf = gen_norm_hhf {protect = false} hhf_ss; val norm_hhf_protect = gen_norm_hhf {protect = true} hhf_protect_ss; fun norm_hhf_protect_without_context th = norm_hhf_protect (Proof_Context.init_global (Thm.theory_of_thm th)) th; end; end; structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier; open Basic_Meta_Simplifier; 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 = proc} + 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 psi = Morphism.set_trim_context'' context phi; val b' = Morphism.binding psi b; val simproc' = transform_simproc psi 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 => fn phi => Thm.declaration_attribute - (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc]))))); + >> (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;