diff --git a/src/Doc/Implementation/Logic.thy b/src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy +++ b/src/Doc/Implementation/Logic.thy @@ -1,1259 +1,1258 @@ (*:maxLineLen=78:*) theory Logic imports Base begin chapter \Primitive logic \label{ch:logic}\ text \ The logical foundations of Isabelle/Isar are that of the Pure logic, which has been introduced as a Natural Deduction framework in @{cite paulson700}. This is essentially the same logic as ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, although there are some key differences in the specific treatment of simple types in Isabelle/Pure. Following type-theoretic parlance, the Pure logic consists of three levels of \\\-calculus with corresponding arrows, \\\ for syntactic function space (terms depending on terms), \\\ for universal quantification (proofs depending on terms), and \\\ for implication (proofs depending on proofs). Derivations are relative to a logical theory, which declares type constructors, constants, and axioms. Theory declarations support schematic polymorphism, which is strictly speaking outside the logic.\<^footnote>\This is the deeper logical reason, why the theory context \\\ is separate from the proof context \\\ of the core calculus: type constructors, term constants, and facts (proof constants) may involve arbitrary type schemes, but the type of a locally fixed term parameter is also fixed!\ \ section \Types \label{sec:types}\ text \ The language of types is an uninterpreted order-sorted first-order algebra; types are qualified by ordered type classes. \<^medskip> A \<^emph>\type class\ is an abstract syntactic entity declared in the theory context. The \<^emph>\subclass relation\ \c\<^sub>1 \ c\<^sub>2\ is specified by stating an acyclic generating relation; the transitive closure is maintained internally. The resulting relation is an ordering: reflexive, transitive, and antisymmetric. A \<^emph>\sort\ is a list of type classes written as \s = {c\<^sub>1, \, c\<^sub>m}\, it represents symbolic intersection. Notationally, the curly braces are omitted for singleton intersections, i.e.\ any class \c\ may be read as a sort \{c}\. The ordering on type classes is extended to sorts according to the meaning of intersections: \{c\<^sub>1, \ c\<^sub>m} \ {d\<^sub>1, \, d\<^sub>n}\ iff \\j. \i. c\<^sub>i \ d\<^sub>j\. The empty intersection \{}\ refers to the universal sort, which is the largest element wrt.\ the sort order. Thus \{}\ represents the ``full sort'', not the empty one! The intersection of all (finitely many) classes declared in the current theory is the least element wrt.\ the sort ordering. \<^medskip> A \<^emph>\fixed type variable\ is a pair of a basic name (starting with a \'\ character) and a sort constraint, e.g.\ \('a, s)\ which is usually printed as \\\<^sub>s\. A \<^emph>\schematic type variable\ is a pair of an indexname and a sort constraint, e.g.\ \(('a, 0), s)\ which is usually printed as \?\\<^sub>s\. Note that \<^emph>\all\ syntactic components contribute to the identity of type variables: basic name, index, and sort constraint. The core logic handles type variables with the same name but different sorts as different, although the type-inference layer (which is outside the core) rejects anything like that. A \<^emph>\type constructor\ \\\ is a \k\-ary operator on types declared in the theory. Type constructor application is written postfix as \(\\<^sub>1, \, \\<^sub>k)\\. For \k = 0\ the argument tuple is omitted, e.g.\ \prop\ instead of \()prop\. For \k = 1\ the parentheses are omitted, e.g.\ \\ list\ instead of \(\)list\. Further notation is provided for specific constructors, notably the right-associative infix \\ \ \\ instead of \(\, \)fun\. The logical category \<^emph>\type\ is defined inductively over type variables and type constructors as follows: \\ = \\<^sub>s | ?\\<^sub>s | (\\<^sub>1, \, \\<^sub>k)\\. A \<^emph>\type abbreviation\ is a syntactic definition \(\<^vec>\)\ = \\ of an arbitrary type expression \\\ over variables \\<^vec>\\. Type abbreviations appear as type constructors in the syntax, but are expanded before entering the logical core. A \<^emph>\type arity\ declares the image behavior of a type constructor wrt.\ the algebra of sorts: \\ :: (s\<^sub>1, \, s\<^sub>k)s\ means that \(\\<^sub>1, \, \\<^sub>k)\\ is of sort \s\ if every argument type \\\<^sub>i\ is of sort \s\<^sub>i\. Arity declarations are implicitly completed, i.e.\ \\ :: (\<^vec>s)c\ entails \\ :: (\<^vec>s)c'\ for any \c' \ c\. \<^medskip> The sort algebra is always maintained as \<^emph>\coregular\, which means that type arities are consistent with the subclass relation: for any type constructor \\\, and classes \c\<^sub>1 \ c\<^sub>2\, and arities \\ :: (\<^vec>s\<^sub>1)c\<^sub>1\ and \\ :: (\<^vec>s\<^sub>2)c\<^sub>2\ holds \\<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ component-wise. The key property of a coregular order-sorted algebra is that sort constraints can be solved in a most general fashion: for each type constructor \\\ and sort \s\ there is a most general vector of argument sorts \(s\<^sub>1, \, s\<^sub>k)\ such that a type scheme \(\\<^bsub>s\<^sub>1\<^esub>, \, \\<^bsub>s\<^sub>k\<^esub>)\\ is of sort \s\. Consequently, type unification has most general solutions (modulo equivalence of sorts), so type-inference produces primary types as expected @{cite "nipkow-prehofer"}. \ text %mlref \ \begin{mldecls} @{define_ML_type class = string} \\ @{define_ML_type sort = "class list"} \\ @{define_ML_type arity = "string * sort list * sort"} \\ @{define_ML_type typ} \\ @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\ @{define_ML Sign.add_type_abbrev: "Proof.context -> binding * string list * typ -> theory -> theory"} \\ @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\ @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\class\ represents type classes. \<^descr> Type \<^ML_type>\sort\ represents sorts, i.e.\ finite intersections of classes. The empty list \<^ML>\[]: sort\ refers to the empty class intersection, i.e.\ the ``full sort''. \<^descr> Type \<^ML_type>\arity\ represents type arities. A triple \(\, \<^vec>s, s) : arity\ represents \\ :: (\<^vec>s)s\ as described above. \<^descr> Type \<^ML_type>\typ\ represents types; this is a datatype with constructors \<^ML>\TFree\, \<^ML>\TVar\, \<^ML>\Type\. \<^descr> \<^ML>\Term.map_atyps\~\f \\ applies the mapping \f\ to all atomic types (\<^ML>\TFree\, \<^ML>\TVar\) occurring in \\\. \<^descr> \<^ML>\Term.fold_atyps\~\f \\ iterates the operation \f\ over all occurrences of atomic types (\<^ML>\TFree\, \<^ML>\TVar\) in \\\; the type structure is traversed from left to right. \<^descr> \<^ML>\Sign.subsort\~\thy (s\<^sub>1, s\<^sub>2)\ tests the subsort relation \s\<^sub>1 \ s\<^sub>2\. \<^descr> \<^ML>\Sign.of_sort\~\thy (\, s)\ tests whether type \\\ is of sort \s\. \<^descr> \<^ML>\Sign.add_type\~\ctxt (\, k, mx)\ declares a new type constructors \\\ with \k\ arguments and optional mixfix syntax. \<^descr> \<^ML>\Sign.add_type_abbrev\~\ctxt (\, \<^vec>\, \)\ defines a new type abbreviation \(\<^vec>\)\ = \\. \<^descr> \<^ML>\Sign.primitive_class\~\(c, [c\<^sub>1, \, c\<^sub>n])\ declares a new class \c\, together with class relations \c \ c\<^sub>i\, for \i = 1, \, n\. \<^descr> \<^ML>\Sign.primitive_classrel\~\(c\<^sub>1, c\<^sub>2)\ declares the class relation \c\<^sub>1 \ c\<^sub>2\. \<^descr> \<^ML>\Sign.primitive_arity\~\(\, \<^vec>s, s)\ declares the arity \\ :: (\<^vec>s)s\. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "class"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "sort"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "nonterminal"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "typ"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation class} embedded ; @@{ML_antiquotation sort} sort ; (@@{ML_antiquotation type_name} | @@{ML_antiquotation type_abbrev} | @@{ML_antiquotation nonterminal}) embedded ; @@{ML_antiquotation typ} type \ \<^descr> \@{class c}\ inlines the internalized class \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{sort s}\ inlines the internalized sort \s\ --- as \<^ML_type>\string list\ literal. \<^descr> \@{type_name c}\ inlines the internalized type constructor \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{type_abbrev c}\ inlines the internalized type abbreviation \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{nonterminal c}\ inlines the internalized syntactic type~/ grammar nonterminal \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{typ \}\ inlines the internalized type \\\ --- as constructor term for datatype \<^ML_type>\typ\. \ section \Terms \label{sec:terms}\ text \ The language of terms is that of simply-typed \\\-calculus with de-Bruijn indices for bound variables (cf.\ @{cite debruijn72} or @{cite "paulson-ml2"}), with the types being determined by the corresponding binders. In contrast, free variables and constants have an explicit name and type in each occurrence. \<^medskip> A \<^emph>\bound variable\ is a natural number \b\, which accounts for the number of intermediate binders between the variable occurrence in the body and its binding position. For example, the de-Bruijn term \\\<^bsub>bool\<^esub>. \\<^bsub>bool\<^esub>. 1 \ 0\ would correspond to \\x\<^bsub>bool\<^esub>. \y\<^bsub>bool\<^esub>. x \ y\ in a named representation. Note that a bound variable may be represented by different de-Bruijn indices at different occurrences, depending on the nesting of abstractions. A \<^emph>\loose variable\ is a bound variable that is outside the scope of local binders. The types (and names) for loose variables can be managed as a separate context, that is maintained as a stack of hypothetical binders. The core logic operates on closed terms, without any loose variables. A \<^emph>\fixed variable\ is a pair of a basic name and a type, e.g.\ \(x, \)\ which is usually printed \x\<^sub>\\ here. A \<^emph>\schematic variable\ is a pair of an indexname and a type, e.g.\ \((x, 0), \)\ which is likewise printed as \?x\<^sub>\\. \<^medskip> A \<^emph>\constant\ is a pair of a basic name and a type, e.g.\ \(c, \)\ which is usually printed as \c\<^sub>\\ here. Constants are declared in the context as polymorphic families \c :: \\, meaning that all substitution instances \c\<^sub>\\ for \\ = \\\ are valid. The vector of \<^emph>\type arguments\ of constant \c\<^sub>\\ wrt.\ the declaration \c :: \\ is defined as the codomain of the matcher \\ = {?\\<^sub>1 \ \\<^sub>1, \, ?\\<^sub>n \ \\<^sub>n}\ presented in canonical order \(\\<^sub>1, \, \\<^sub>n)\, corresponding to the left-to-right occurrences of the \\\<^sub>i\ in \\\. Within a given theory context, there is a one-to-one correspondence between any constant \c\<^sub>\\ and the application \c(\\<^sub>1, \, \\<^sub>n)\ of its type arguments. For example, with \plus :: \ \ \ \ \\, the instance \plus\<^bsub>nat \ nat \ nat\<^esub>\ corresponds to \plus(nat)\. Constant declarations \c :: \\ may contain sort constraints for type variables in \\\. These are observed by type-inference as expected, but \<^emph>\ignored\ by the core logic. This means the primitive logic is able to reason with instances of polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. \<^medskip> An \<^emph>\atomic term\ is either a variable or constant. The logical category \<^emph>\term\ is defined inductively over atomic terms, with abstraction and application as follows: \t = b | x\<^sub>\ | ?x\<^sub>\ | c\<^sub>\ | \\<^sub>\. t | t\<^sub>1 t\<^sub>2\. Parsing and printing takes care of converting between an external representation with named bound variables. Subsequently, we shall use the latter notation instead of internal de-Bruijn representation. The inductive relation \t :: \\ assigns a (unique) type to a term according to the structure of atomic terms, abstractions, and applications: \[ \infer{\a\<^sub>\ :: \\}{} \qquad \infer{\(\x\<^sub>\. t) :: \ \ \\}{\t :: \\} \qquad \infer{\t u :: \\}{\t :: \ \ \\ & \u :: \\} \] A \<^emph>\well-typed term\ is a term that can be typed according to these rules. Typing information can be omitted: type-inference is able to reconstruct the most general type of a raw term, while assigning most general types to all of its variables and constants. Type-inference depends on a context of type constraints for fixed variables, and declarations for polymorphic constants. The identity of atomic terms consists both of the name and the type component. This means that different variables \x\<^bsub>\\<^sub>1\<^esub>\ and \x\<^bsub>\\<^sub>2\<^esub>\ may become the same after type instantiation. Type-inference rejects variables of the same name, but different types. In contrast, mixed instances of polymorphic constants occur routinely. \<^medskip> The \<^emph>\hidden polymorphism\ of a term \t :: \\ is the set of type variables occurring in \t\, but not in its type \\\. This means that the term implicitly depends on type arguments that are not accounted in the result type, i.e.\ there are different type instances \t\ :: \\ and \t\' :: \\ with the same type. This slightly pathological situation notoriously demands additional care. \<^medskip> A \<^emph>\term abbreviation\ is a syntactic definition \c\<^sub>\ \ t\ of a closed term \t\ of type \\\, without any hidden polymorphism. A term abbreviation looks like a constant in the syntax, but is expanded before entering the logical core. Abbreviations are usually reverted when printing terms, using \t \ c\<^sub>\\ as rules for higher-order rewriting. \<^medskip> Canonical operations on \\\-terms include \\\\\-conversion: \\\-conversion refers to capture-free renaming of bound variables; \\\-conversion contracts an abstraction applied to an argument term, substituting the argument in the body: \(\x. b)a\ becomes \b[a/x]\; \\\-conversion contracts vacuous application-abstraction: \\x. f x\ becomes \f\, provided that the bound variable does not occur in \f\. Terms are normally treated modulo \\\-conversion, which is implicit in the de-Bruijn representation. Names for bound variables in abstractions are maintained separately as (meaningless) comments, mostly for parsing and printing. Full \\\\\-conversion is commonplace in various standard operations (\secref{sec:obj-rules}) that are based on higher-order unification and matching. \ text %mlref \ \begin{mldecls} @{define_ML_type term} \\ @{define_ML_infix "aconv": "term * term -> bool"} \\ @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML fastype_of: "term -> typ"} \\ @{define_ML lambda: "term -> term -> term"} \\ @{define_ML betapply: "term * term -> term"} \\ @{define_ML incr_boundvars: "int -> term -> term"} \\ @{define_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{define_ML Sign.add_abbrev: "string -> binding * term -> theory -> (term * term) * theory"} \\ @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\term\ represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}. \<^descr> \t\~\<^ML_text>\aconv\~\u\ checks \\\-equivalence of two terms. This is the basic equality relation on type \<^ML_type>\term\; raw datatype equality should only be used for operations related to parsing or printing! \<^descr> \<^ML>\Term.map_types\~\f t\ applies the mapping \f\ to all types occurring in \t\. \<^descr> \<^ML>\Term.fold_types\~\f t\ iterates the operation \f\ over all occurrences of types in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\Term.map_aterms\~\f t\ applies the mapping \f\ to all atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) occurring in \t\. \<^descr> \<^ML>\Term.fold_aterms\~\f t\ iterates the operation \f\ over all occurrences of atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\fastype_of\~\t\ determines the type of a well-typed term. This operation is relatively slow, despite the omission of any sanity checks. \<^descr> \<^ML>\lambda\~\a b\ produces an abstraction \\a. b\, where occurrences of the atomic term \a\ in the body \b\ are replaced by bound variables. \<^descr> \<^ML>\betapply\~\(t, u)\ produces an application \t u\, with topmost \\\-conversion if \t\ is an abstraction. \<^descr> \<^ML>\incr_boundvars\~\j\ increments a term's dangling bound variables by the offset \j\. This is required when moving a subterm into a context where it is enclosed by a different number of abstractions. Bound variables with a matching abstraction are unaffected. \<^descr> \<^ML>\Sign.declare_const\~\ctxt ((c, \), mx)\ declares a new constant \c :: \\ with optional mixfix syntax. \<^descr> \<^ML>\Sign.add_abbrev\~\print_mode (c, t)\ introduces a new term abbreviation \c \ t\. \<^descr> \<^ML>\Sign.const_typargs\~\thy (c, \)\ and \<^ML>\Sign.const_instance\~\thy (c, [\\<^sub>1, \, \\<^sub>n])\ convert between two representations of polymorphic constants: full type instance vs.\ compact type arguments form. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "const_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "term"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "prop"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ (@@{ML_antiquotation const_name} | @@{ML_antiquotation const_abbrev}) embedded ; @@{ML_antiquotation const} ('(' (type + ',') ')')? ; @@{ML_antiquotation term} term ; @@{ML_antiquotation prop} prop \ \<^descr> \@{const_name c}\ inlines the internalized logical constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const_abbrev c}\ inlines the internalized abbreviated constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const c(\<^vec>\)}\ inlines the internalized constant \c\ with precise type instantiation in the sense of \<^ML>\Sign.const_instance\ --- as \<^ML>\Const\ constructor term for datatype \<^ML_type>\term\. \<^descr> \@{term t}\ inlines the internalized term \t\ --- as constructor term for datatype \<^ML_type>\term\. \<^descr> \@{prop \}\ inlines the internalized proposition \\\ --- as constructor term for datatype \<^ML_type>\term\. \ section \Theorems \label{sec:thms}\ text \ A \<^emph>\proposition\ is a well-typed term of type \prop\, a \<^emph>\theorem\ is a proven proposition (depending on a context of hypotheses and the background theory). Primitive inferences include plain Natural Deduction rules for the primary connectives \\\ and \\\ of the framework. There is also a builtin notion of equality/equivalence \\\. \ subsection \Primitive connectives and rules \label{sec:prim-rules}\ text \ The theory \Pure\ contains constant declarations for the primitive connectives \\\, \\\, and \\\ of the logical framework, see \figref{fig:pure-connectives}. The derivability judgment \A\<^sub>1, \, A\<^sub>n \ B\ is defined inductively by the primitive inferences given in \figref{fig:prim-rules}, with the global restriction that the hypotheses must \<^emph>\not\ contain any schematic variables. The builtin equality is conceptually axiomatized as shown in \figref{fig:pure-equality}, although the implementation works directly with derived inferences. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \all :: (\ \ prop) \ prop\ & universal quantification (binder \\\) \\ \\ :: prop \ prop \ prop\ & implication (right associative infix) \\ \\ :: \ \ \ \ prop\ & equality relation (infix) \\ \end{tabular} \caption{Primitive connectives of Pure}\label{fig:pure-connectives} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \[ \infer[\(axiom)\]{\\ A\}{\A \ \\} \qquad \infer[\(assume)\]{\A \ A\}{} \] \[ \infer[\(\\intro)\]{\\ \ \x. B[x]\}{\\ \ B[x]\ & \x \ \\} \qquad \infer[\(\\elim)\]{\\ \ B[a]\}{\\ \ \x. B[x]\} \] \[ \infer[\(\\intro)\]{\\ - A \ A \ B\}{\\ \ B\} \qquad \infer[\(\\elim)\]{\\\<^sub>1 \ \\<^sub>2 \ B\}{\\\<^sub>1 \ A \ B\ & \\\<^sub>2 \ A\} \] \caption{Primitive inferences of Pure}\label{fig:prim-rules} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \\ (\x. b[x]) a \ b[a]\ & \\\-conversion \\ \\ x \ x\ & reflexivity \\ \\ x \ y \ P x \ P y\ & substitution \\ \\ (\x. f x \ g x) \ f \ g\ & extensionality \\ \\ (A \ B) \ (B \ A) \ A \ B\ & logical equivalence \\ \end{tabular} \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality} \end{center} \end{figure} The introduction and elimination rules for \\\ and \\\ are analogous to formation of dependently typed \\\-terms representing the underlying proof objects. Proof terms are irrelevant in the Pure logic, though; they cannot occur within propositions. The system provides a runtime option to record explicit proof terms for primitive inferences, see also \secref{sec:proof-terms}. Thus all three levels of \\\-calculus become explicit: \\\ for terms, and \\/\\ for proofs (cf.\ @{cite "Berghofer-Nipkow:2000:TPHOL"}). Observe that locally fixed parameters (as in \\\intro\) need not be recorded in the hypotheses, because the simple syntactic types of Pure are always inhabitable. ``Assumptions'' \x :: \\ for type-membership are only present as long as some \x\<^sub>\\ occurs in the statement body.\<^footnote>\This is the key difference to ``\\HOL\'' in the PTS framework @{cite "Barendregt-Geuvers:2001"}, where hypotheses \x : A\ are treated uniformly for propositions and types.\ \<^medskip> The axiomatization of a theory is implicitly closed by forming all instances of type and term variables: \\ A\\ holds for any substitution instance of an axiom \\ A\. By pushing substitutions through derivations inductively, we also get admissible \generalize\ and \instantiate\ rules as shown in \figref{fig:subst-rules}. \begin{figure}[htb] \begin{center} \[ \infer{\\ \ B[?\]\}{\\ \ B[\]\ & \\ \ \\} \quad \infer[\quad\(generalize)\]{\\ \ B[?x]\}{\\ \ B[x]\ & \x \ \\} \] \[ \infer{\\ \ B[\]\}{\\ \ B[?\]\} \quad \infer[\quad\(instantiate)\]{\\ \ B[t]\}{\\ \ B[?x]\} \] \caption{Admissible substitution rules}\label{fig:subst-rules} \end{center} \end{figure} Note that \instantiate\ does not require an explicit side-condition, because \\\ may never contain schematic variables. In principle, variables could be substituted in hypotheses as well, but this would disrupt the monotonicity of reasoning: deriving \\\ \ B\\ from \\ \ B\ is correct, but \\\ \ \\ does not necessarily hold: the result belongs to a different proof context. \<^medskip> An \<^emph>\oracle\ is a function that produces axioms on the fly. Logically, this is an instance of the \axiom\ rule (\figref{fig:prim-rules}), but there is an operational difference. The inference kernel records oracle invocations within derivations of theorems by a unique tag. 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}). Axiomatizations should be limited to the bare minimum, typically as part of the initial logical basis of an object-logic formalization. Later on, theories are usually developed in a strictly definitional fashion, by stating only certain equalities over new constants. A \<^emph>\simple definition\ consists of a constant declaration \c :: \\ together with an axiom \\ c \ t\, where \t :: \\ is a closed term without any hidden polymorphism. The RHS may depend on further defined constants, but not \c\ itself. Definitions of functions may be presented as \c \<^vec>x \ t\ instead of the puristic \c \ \\<^vec>x. t\. An \<^emph>\overloaded definition\ consists of a collection of axioms for the same constant, with zero or one equations \c((\<^vec>\)\) \ t\ for each type constructor \\\ (for distinct variables \\<^vec>\\). The RHS may mention previously defined constants as above, or arbitrary constants \d(\\<^sub>i)\ for some \\\<^sub>i\ projected from \\<^vec>\\. Thus overloaded definitions essentially work by primitive recursion over the syntactic structure of a single type argument. See also @{cite \\S4.3\ "Haftmann-Wenzel:2006:classes"}. \ text %mlref \ \begin{mldecls} @{define_ML Logic.all: "term -> term -> term"} \\ @{define_ML Logic.mk_implies: "term * term -> term"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type ctyp} \\ @{define_ML_type cterm} \\ @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\ @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type thm} \\ @{define_ML Thm.transfer: "theory -> thm -> thm"} \\ @{define_ML Thm.assume: "cterm -> thm"} \\ @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\ @{define_ML Thm.generalize: "Names.set * Names.set -> int -> thm -> thm"} \\ - @{define_ML Thm.instantiate: "((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list - -> thm -> thm"} \\ + @{define_ML Thm.instantiate: "ctyp TVars.table * cterm Vars.table -> thm -> thm"} \\ @{define_ML Thm.add_axiom: "Proof.context -> binding * term -> theory -> (string * thm) * theory"} \\ @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory"} \\ @{define_ML Thm.add_def: "Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} @{define_ML Theory.add_deps: "Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory"} \\ @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\ \end{mldecls} \<^descr> \<^ML>\Logic.all\~\a B\ produces a Pure quantification \\a. B\, where occurrences of the atomic term \a\ in the body proposition \B\ are replaced by bound variables. (See also \<^ML>\lambda\ on terms.) \<^descr> \<^ML>\Logic.mk_implies\~\(A, B)\ produces a Pure implication \A \ B\. \<^descr> Types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc.\ in the background theory. The abstract types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are part of the same inference kernel that is mainly responsible for \<^ML_type>\thm\. Thus syntactic operations on \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are located in the \<^ML_structure>\Thm\ module, even though theorems are not yet involved at that stage. \<^descr> \<^ML>\Thm.ctyp_of\~\ctxt \\ and \<^ML>\Thm.cterm_of\~\ctxt t\ explicitly check types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the underlying theory context. Full re-certification is relatively slow and should be avoided in tight reasoning loops. \<^descr> \<^ML>\Thm.apply\, \<^ML>\Thm.lambda\, \<^ML>\Thm.all\, \<^ML>\Drule.mk_implies\ etc.\ compose certified terms (or propositions) incrementally. This is equivalent to \<^ML>\Thm.cterm_of\ after unchecked \<^ML_infix>\$\, \<^ML>\lambda\, \<^ML>\Logic.all\, \<^ML>\Logic.mk_implies\ etc., but there can be a big difference in performance when large existing entities are composed by a few extra constructions on top. There are separate operations to decompose certified terms and theorems to produce certified terms again. \<^descr> Type \<^ML_type>\thm\ represents proven propositions. This is an abstract datatype that guarantees that its values have been constructed by basic principles of the \<^ML_structure>\Thm\ module. Every \<^ML_type>\thm\ value refers its background theory, cf.\ \secref{sec:context-theory}. \<^descr> \<^ML>\Thm.transfer\~\thy thm\ transfers the given theorem to a \<^emph>\larger\ theory, see also \secref{sec:context}. This formal adjustment of the background context has no logical significance, but is occasionally required for formal reasons, e.g.\ when theorems that are imported from more basic theories are used in the current situation. \<^descr> \<^ML>\Thm.assume\, \<^ML>\Thm.forall_intr\, \<^ML>\Thm.forall_elim\, \<^ML>\Thm.implies_intr\, and \<^ML>\Thm.implies_elim\ correspond to the primitive inferences of \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.generalize\~\(\<^vec>\, \<^vec>x)\ corresponds to the \generalize\ rules of \figref{fig:subst-rules}. Here collections of type and term variables are generalized simultaneously, specified by the given sets of basic names. \<^descr> \<^ML>\Thm.instantiate\~\(\<^vec>\\<^sub>s, \<^vec>x\<^sub>\)\ corresponds to the \instantiate\ rules of \figref{fig:subst-rules}. Type variables are substituted before term variables. Note that the types in \\<^vec>x\<^sub>\\ refer to the instantiated versions. \<^descr> \<^ML>\Thm.add_axiom\~\ctxt (name, A)\ declares an arbitrary proposition as axiom, and retrieves it as a theorem from the resulting theory, cf.\ \axiom\ in \figref{fig:prim-rules}. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Thm.add_oracle\~\(binding, oracle)\ produces a named oracle rule, essentially generating arbitrary axioms on the fly, cf.\ \axiom\ in \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.add_def\~\ctxt unchecked overloaded (name, c \<^vec>x \ t)\ states a definitional axiom for an existing constant \c\. Dependencies are recorded via \<^ML>\Theory.add_deps\, unless the \unchecked\ option is set. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Theory.add_deps\~\ctxt name c\<^sub>\ \<^vec>d\<^sub>\\ declares dependencies of a named specification for constant \c\<^sub>\\, relative to existing specifications for constants \\<^vec>d\<^sub>\\. This also works for type constructors. \<^descr> \<^ML>\Thm_Deps.all_oracles\~\thms\ returns all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. The result contains an authentic oracle name; if @{ML Proofterm.proofs} was at least @{ML 1} during the oracle inference, it also contains the position of the oracle invocation and its proposition. See also the command @{command_ref "thm_oracles"}. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "ctyp"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cterm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cprop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thms"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "lemma"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "oracle_name"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation ctyp} typ ; @@{ML_antiquotation cterm} term ; @@{ML_antiquotation cprop} prop ; @@{ML_antiquotation thm} thm ; @@{ML_antiquotation thms} thms ; @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \ @'by' method method? ; @@{ML_antiquotation oracle_name} embedded \ \<^descr> \@{ctyp \}\ produces a certified type wrt.\ the current background theory --- as abstract value of type \<^ML_type>\ctyp\. \<^descr> \@{cterm t}\ and \@{cprop \}\ produce a certified term wrt.\ the current background theory --- as abstract value of type \<^ML_type>\cterm\. \<^descr> \@{thm a}\ produces a singleton fact --- as abstract value of type \<^ML_type>\thm\. \<^descr> \@{thms a}\ produces a general fact --- as abstract value of type \<^ML_type>\thm list\. \<^descr> \@{lemma \ by meth}\ produces a fact that is proven on the spot according to the minimal proof, which imitates a terminal Isar proof. The result is an abstract value of type \<^ML_type>\thm\ or \<^ML_type>\thm list\, depending on the number of propositions given here. The internal derivation object lacks a proper theorem name, but it is formally closed, unless the \(open)\ option is specified (this may impact performance of applications with proof terms). Since ML antiquotations are always evaluated at compile-time, there is no run-time overhead even for non-trivial proofs. Nonetheless, the justification is syntactically limited to a single @{command "by"} step. More complex Isar proofs should be done in regular theory source, before compiling the corresponding ML text that uses the result. \<^descr> \@{oracle_name a}\ inlines the internalized oracle name \a\ --- as \<^ML_type>\string\ literal. \ subsection \Auxiliary connectives \label{sec:logic-aux}\ text \ Theory \Pure\ provides a few auxiliary connectives that are defined on top of the primitive ones, see \figref{fig:pure-aux}. These special constants are useful in certain internal encodings, and are normally not directly exposed to the user. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \conjunction :: prop \ prop \ prop\ & (infix \&&&\) \\ \\ A &&& B \ (\C. (A \ B \ C) \ C)\ \\[1ex] \prop :: prop \ prop\ & (prefix \#\, suppressed) \\ \#A \ A\ \\[1ex] \term :: \ \ prop\ & (prefix \TERM\) \\ \term x \ (\A. A \ A)\ \\[1ex] \type :: \ itself\ & (prefix \TYPE\) \\ \(unspecified)\ \\ \end{tabular} \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} \end{center} \end{figure} The introduction \A \ B \ A &&& B\, and eliminations (projections) \A &&& B \ A\ and \A &&& B \ B\ are available as derived rules. Conjunction allows to treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \A \ B \ C &&& D\. In particular, the goal mechanism represents multiple claims as explicit conjunction internally, but this is refined (via backwards introduction) into separate sub-goals before the user commences the proof; the final result is projected into a list of theorems using eliminations (cf.\ \secref{sec:tactical-goals}). The \prop\ marker (\#\) makes arbitrarily complex propositions appear as atomic, without changing the meaning: \\ \ A\ and \\ \ #A\ are interchangeable. See \secref{sec:tactical-goals} for specific operations. The \term\ marker turns any well-typed term into a derivable proposition: \\ TERM t\ holds unconditionally. Although this is logically vacuous, it allows to treat terms and proofs uniformly, similar to a type-theoretic framework. The \TYPE\ constructor is the canonical representative of the unspecified type \\ itself\; it essentially injects the language of types into that of terms. There is specific notation \TYPE(\)\ for \TYPE\<^bsub>\ itself\<^esub>\. Although being devoid of any particular meaning, the term \TYPE(\)\ accounts for the type \\\ within the term language. In particular, \TYPE(\)\ may be used as formal argument in primitive definitions, in order to circumvent hidden polymorphism (cf.\ \secref{sec:terms}). For example, \c TYPE(\) \ A[\]\ defines \c :: \ itself \ prop\ in terms of a proposition \A\ that depends on an additional type argument, which is essentially a predicate on types. \ text %mlref \ \begin{mldecls} @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\ @{define_ML Conjunction.elim: "thm -> thm * thm"} \\ @{define_ML Drule.mk_term: "cterm -> thm"} \\ @{define_ML Drule.dest_term: "thm -> cterm"} \\ @{define_ML Logic.mk_type: "typ -> term"} \\ @{define_ML Logic.dest_type: "term -> typ"} \\ \end{mldecls} \<^descr> \<^ML>\Conjunction.intr\ derives \A &&& B\ from \A\ and \B\. \<^descr> \<^ML>\Conjunction.elim\ derives \A\ and \B\ from \A &&& B\. \<^descr> \<^ML>\Drule.mk_term\ derives \TERM t\. \<^descr> \<^ML>\Drule.dest_term\ recovers term \t\ from \TERM t\. \<^descr> \<^ML>\Logic.mk_type\~\\\ produces the term \TYPE(\)\. \<^descr> \<^ML>\Logic.dest_type\~\TYPE(\)\ recovers the type \\\. \ subsection \Sort hypotheses\ text \ Type variables are decorated with sorts, as explained in \secref{sec:types}. This constrains type instantiation to certain ranges of types: variable \\\<^sub>s\ may only be assigned to types \\\ that belong to sort \s\. Within the logic, sort constraints act like implicit preconditions on the result \\\\<^sub>1 : s\<^sub>1\, \, \\\<^sub>n : s\<^sub>n\, \ \ \\ where the type variables \\\<^sub>1, \, \\<^sub>n\ cover the propositions \\\, \\\, as well as the proof of \\ \ \\. These \<^emph>\sort hypothesis\ of a theorem are passed monotonically through further derivations. They are redundant, as long as the statement of a theorem still contains the type variables that are accounted here. The logical significance of sort hypotheses is limited to the boundary case where type variables disappear from the proposition, e.g.\ \\\\<^sub>s : s\ \ \\. Since such dangling type variables can be renamed arbitrarily without changing the proposition \\\, the inference kernel maintains sort hypotheses in anonymous form \s \ \\. In most practical situations, such extra sort hypotheses may be stripped in a final bookkeeping step, e.g.\ at the end of a proof: they are typically left over from intermediate reasoning with type classes that can be satisfied by some concrete type \\\ of sort \s\ to replace the hypothetical type variable \\\<^sub>s\. \ text %mlref \ \begin{mldecls} @{define_ML Thm.extra_shyps: "thm -> sort list"} \\ @{define_ML Thm.strip_shyps: "thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Thm.extra_shyps\~\thm\ determines the extraneous sort hypotheses of the given theorem, i.e.\ the sorts that are not present within type variables of the statement. \<^descr> \<^ML>\Thm.strip_shyps\~\thm\ removes any extraneous sort hypotheses that can be witnessed from the type signature. \ text %mlex \ The following artificial example demonstrates the derivation of \<^prop>\False\ with a pending sort hypothesis involving a logically empty sort. \ class empty = assumes bad: "\(x::'a) y. x \ y" theorem (in empty) false: False using bad by blast ML_val \\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\empty\])\ text \ Thanks to the inference kernel managing sort hypothesis according to their logical significance, this example is merely an instance of \<^emph>\ex falso quodlibet consequitur\ --- not a collapse of the logical framework! \ section \Object-level rules \label{sec:obj-rules}\ text \ The primitive inferences covered so far mostly serve foundational purposes. User-level reasoning usually works via object-level rules that are represented as theorems of Pure. Composition of rules involves \<^emph>\backchaining\, \<^emph>\higher-order unification\ modulo \\\\\-conversion of \\\-terms, and so-called \<^emph>\lifting\ of rules into a context of \\\ and \\\ connectives. Thus the full power of higher-order Natural Deduction in Isabelle/Pure becomes readily available. \ subsection \Hereditary Harrop Formulae\ text \ The idea of object-level rules is to model Natural Deduction inferences in the style of Gentzen @{cite "Gentzen:1935"}, but we allow arbitrary nesting similar to @{cite "Schroeder-Heister:1984"}. The most basic rule format is that of a \<^emph>\Horn Clause\: \[ \infer{\A\}{\A\<^sub>1\ & \\\ & \A\<^sub>n\} \] where \A, A\<^sub>1, \, A\<^sub>n\ are atomic propositions of the framework, usually of the form \Trueprop B\, where \B\ is a (compound) object-level statement. This object-level inference corresponds to an iterated implication in Pure like this: \[ \A\<^sub>1 \ \ A\<^sub>n \ A\ \] As an example consider conjunction introduction: \A \ B \ A \ B\. Any parameters occurring in such rule statements are conceptionally treated as arbitrary: \[ \\x\<^sub>1 \ x\<^sub>m. A\<^sub>1 x\<^sub>1 \ x\<^sub>m \ \ A\<^sub>n x\<^sub>1 \ x\<^sub>m \ A x\<^sub>1 \ x\<^sub>m\ \] Nesting of rules means that the positions of \A\<^sub>i\ may again hold compound rules, not just atomic propositions. Propositions of this format are called \<^emph>\Hereditary Harrop Formulae\ in the literature @{cite "Miller:1991"}. Here we give an inductive characterization as follows: \<^medskip> \begin{tabular}{ll} \\<^bold>x\ & set of variables \\ \\<^bold>A\ & set of atomic propositions \\ \\<^bold>H = \\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \ \<^bold>A\ & set of Hereditary Harrop Formulas \\ \end{tabular} \<^medskip> Thus we essentially impose nesting levels on propositions formed from \\\ and \\\. At each level there is a prefix of parameters and compound premises, concluding an atomic proposition. Typical examples are \\\-introduction \(A \ B) \ A \ B\ or mathematical induction \P 0 \ (\n. P n \ P (Suc n)) \ P n\. Even deeper nesting occurs in well-founded induction \(\x. (\y. y \ x \ P y) \ P x) \ P x\, but this already marks the limit of rule complexity that is usually seen in practice. \<^medskip> Regular user-level inferences in Isabelle/Pure always maintain the following canonical form of results: \<^item> Normalization by \(A \ (\x. B x)) \ (\x. A \ B x)\, which is a theorem of Pure, means that quantifiers are pushed in front of implication at each level of nesting. The normal form is a Hereditary Harrop Formula. \<^item> The outermost prefix of parameters is represented via schematic variables: instead of \\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x\ we have \\<^vec>H ?\<^vec>x \ A ?\<^vec>x\. Note that this representation looses information about the order of parameters, and vacuous quantifiers vanish automatically. \ text %mlref \ \begin{mldecls} @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Simplifier.norm_hhf\~\ctxt thm\ normalizes the given theorem according to the canonical form specified above. This is occasionally helpful to repair some low-level tools that do not handle Hereditary Harrop Formulae properly. \ subsection \Rule composition\ text \ The rule calculus of Isabelle/Pure provides two main inferences: @{inference resolution} (i.e.\ back-chaining of rules) and @{inference assumption} (i.e.\ closing a branch), both modulo higher-order unification. There are also combined variants, notably @{inference elim_resolution} and @{inference dest_resolution}. To understand the all-important @{inference resolution} principle, we first consider raw @{inference_def composition} (modulo higher-order unification with substitution \\\): \[ \infer[(@{inference_def composition})]{\\<^vec>A\ \ C\\} {\\<^vec>A \ B\ & \B' \ C\ & \B\ = B'\\} \] Here the conclusion of the first rule is unified with the premise of the second; the resulting rule instance inherits the premises of the first and conclusion of the second. Note that \C\ can again consist of iterated implications. We can also permute the premises of the second rule back-and-forth in order to compose with \B'\ in any position (subsequently we shall always refer to position 1 w.l.o.g.). In @{inference composition} the internal structure of the common part \B\ and \B'\ is not taken into account. For proper @{inference resolution} we require \B\ to be atomic, and explicitly observe the structure \\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x\ of the premise of the second rule. The idea is to adapt the first rule by ``lifting'' it into this context, by means of iterated application of the following inferences: \[ \infer[(@{inference_def imp_lift})]{\(\<^vec>H \ \<^vec>A) \ (\<^vec>H \ B)\}{\\<^vec>A \ B\} \] \[ \infer[(@{inference_def all_lift})]{\(\\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \ (\\<^vec>x. B (?\<^vec>a \<^vec>x))\}{\\<^vec>A ?\<^vec>a \ B ?\<^vec>a\} \] By combining raw composition with lifting, we get full @{inference resolution} as follows: \[ \infer[(@{inference_def resolution})] {\(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (?\<^vec>a \<^vec>x))\ \ C\\} {\begin{tabular}{l} \\<^vec>A ?\<^vec>a \ B ?\<^vec>a\ \\ \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \(\\<^vec>x. B (?\<^vec>a \<^vec>x))\ = B'\\ \\ \end{tabular}} \] Continued resolution of rules allows to back-chain a problem towards more and sub-problems. Branches are closed either by resolving with a rule of 0 premises, or by producing a ``short-circuit'' within a solved situation (again modulo unification): \[ \infer[(@{inference_def assumption})]{\C\\} {\(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C\ & \A\ = H\<^sub>i\\~~\mbox{(for some~\i\)}} \] %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution} \ text %mlref \ \begin{mldecls} @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\ @{define_ML_infix "RS": "thm * thm -> thm"} \\ @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\ @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\ @{define_ML_infix "MRS": "thm list * thm -> thm"} \\ @{define_ML_infix "OF": "thm * thm list -> thm"} \\ \end{mldecls} \<^descr> \rule\<^sub>1 RSN (i, rule\<^sub>2)\ resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, according to the @{inference resolution} principle explained above. Unless there is precisely one resolvent it raises exception \<^ML>\THM\. This corresponds to the rule attribute @{attribute THEN} in Isar source language. \<^descr> \rule\<^sub>1 RS rule\<^sub>2\ abbreviates \rule\<^sub>1 RSN (1, rule\<^sub>2)\. \<^descr> \rules\<^sub>1 RLN (i, rules\<^sub>2)\ joins lists of rules. For every \rule\<^sub>1\ in \rules\<^sub>1\ and \rule\<^sub>2\ in \rules\<^sub>2\, it resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, accumulating multiple results in one big list. Note that such strict enumerations of higher-order unifications can be inefficient compared to the lazy variant seen in elementary tactics like \<^ML>\resolve_tac\. \<^descr> \rules\<^sub>1 RL rules\<^sub>2\ abbreviates \rules\<^sub>1 RLN (1, rules\<^sub>2)\. \<^descr> \[rule\<^sub>1, \, rule\<^sub>n] MRS rule\ resolves \rule\<^sub>i\ against premise \i\ of \rule\, for \i = n, \, 1\. By working from right to left, newly emerging premises are concatenated in the result, without interfering. \<^descr> \rule OF rules\ is an alternative notation for \rules MRS rule\, which makes rule composition look more like function application. Note that the argument \rules\ need not be atomic. This corresponds to the rule attribute @{attribute OF} in Isar source language. \ section \Proof terms \label{sec:proof-terms}\ text \ The Isabelle/Pure inference kernel can record the proof of each theorem as a proof term that contains all logical inferences in detail. Rule composition by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken down to primitive rules of the logical framework. The proof term can be inspected by a separate proof-checker, for example. According to the well-known \<^emph>\Curry-Howard isomorphism\, a proof can be viewed as a \\\-term. Following this idea, proofs in Isabelle are internally represented by a datatype similar to the one for terms described in \secref{sec:terms}. On top of these syntactic terms, two more layers of \\\-calculus are added, which correspond to \\x :: \. B x\ and \A \ B\ according to the propositions-as-types principle. The resulting 3-level \\\-calculus resembles ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, if some fine points like schematic polymorphism and type classes are ignored. \<^medskip> \<^emph>\Proof abstractions\ of the form \\<^bold>\x :: \. prf\ or \\<^bold>\p : A. prf\ correspond to introduction of \\\/\\\, and \<^emph>\proof applications\ of the form \p \ t\ or \p \ q\ correspond to elimination of \\\/\\\. Actual types \\\, propositions \A\, and terms \t\ might be suppressed and reconstructed from the overall proof term. \<^medskip> Various atomic proofs indicate special situations within the proof construction as follows. A \<^emph>\bound proof variable\ is a natural number \b\ that acts as de-Bruijn index for proof term abstractions. A \<^emph>\minimal proof\ ``\?\'' is a dummy proof term. This indicates some unrecorded part of the proof. \Hyp A\ refers to some pending hypothesis by giving its proposition. This indicates an open context of implicit hypotheses, similar to loose bound variables or free variables within a term (\secref{sec:terms}). An \<^emph>\axiom\ or \<^emph>\oracle\ \a : A[\<^vec>\]\ refers some postulated \proof constant\, which is subject to schematic polymorphism of theory content, and the particular type instantiation may be given explicitly. The vector of types \\<^vec>\\ refers to the schematic type variables in the generic proposition \A\ in canonical order. A \<^emph>\proof promise\ \a : A[\<^vec>\]\ is a placeholder for some proof of polymorphic proposition \A\, with explicit type instantiation as given by the vector \\<^vec>\\, as above. Unlike axioms or oracles, proof promises may be \<^emph>\fulfilled\ eventually, by substituting \a\ by some particular proof \q\ at the corresponding type instance. This acts like Hindley-Milner \let\-polymorphism: a generic local proof definition may get used at different type instances, and is replaced by the concrete instance eventually. A \<^emph>\named theorem\ wraps up some concrete proof as a closed formal entity, in the manner of constant definitions for proof terms. The \<^emph>\proof body\ of such boxed theorems involves some digest about oracles and promises occurring in the original proof. This allows the inference kernel to manage this critical information without the full overhead of explicit proof terms. \ subsection \Reconstructing and checking proof terms\ text \ Fully explicit proof terms can be large, but most of this information is redundant and can be reconstructed from the context. Therefore, the Isabelle/Pure inference kernel records only \<^emph>\implicit\ proof terms, by omitting all typing information in terms, all term and type labels of proof abstractions, and some argument terms of applications \p \ t\ (if possible). There are separate operations to reconstruct the full proof term later on, using \<^emph>\higher-order pattern unification\ @{cite "nipkow-patterns" and "Berghofer-Nipkow:2000:TPHOL"}. The \<^emph>\proof checker\ expects a fully reconstructed proof term, and can turn it into a theorem by replaying its primitive inferences within the kernel. \ subsection \Concrete syntax of proof terms\ text \ The concrete syntax of proof terms is a slight extension of the regular inner syntax of Isabelle/Pure @{cite "isabelle-isar-ref"}. Its main syntactic category @{syntax (inner) proof} is defined as follows: \begin{center} \begin{supertabular}{rclr} @{syntax_def (inner) proof} & = & \\<^bold>\\ \params\ \<^verbatim>\.\ \proof\ \\ & \|\ & \proof\ \\\ \any\ \\ & \|\ & \proof\ \\\ \proof\ \\ & \|\ & \id | longid\ \\ \\ \param\ & = & \idt\ \\ & \|\ & \idt\ \<^verbatim>\:\ \prop\ \\ & \|\ & \<^verbatim>\(\ \param\ \<^verbatim>\)\ \\ \\ \params\ & = & \param\ \\ & \|\ & \param\ \params\ \\ \end{supertabular} \end{center} Implicit term arguments in partial proofs are indicated by ``\_\''. Type arguments for theorems and axioms may be specified using \p \ TYPE(type)\ (they must appear before any other term argument of a theorem or axiom, but may be omitted altogether). \<^medskip> There are separate read and print operations for proof terms, in order to avoid conflicts with the regular term language. \ text %mlref \ \begin{mldecls} @{define_ML_type proof} \\ @{define_ML_type proof_body} \\ @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{define_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{define_ML Proofterm.expand_proof: "theory -> (Proofterm.thm_header -> string option) -> proof -> proof"} \\ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\proof\ represents proof terms; this is a datatype with constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"}, @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained above. %FIXME PClass (!?) \<^descr> Type \<^ML_type>\proof_body\ represents the nested proof information of a named theorem, consisting of a digest of oracles and named theorem over some proof term. The digest only covers the directly visible part of the proof: in order to get the full information, the implicit graph of nested theorems needs to be traversed (e.g.\ using \<^ML>\Proofterm.fold_body_thms\). \<^descr> \<^ML>\Thm.proof_of\~\thm\ and \<^ML>\Thm.proof_body_of\~\thm\ produce the proof term or proof body (with digest of oracles and theorems) from a given theorem. Note that this involves a full join of internal futures that fulfill pending proof promises, and thus disrupts the natural bottom-up construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. \<^descr> \<^ML>\Proofterm.proofs\ specifies the detail of proof recording within \<^ML_type>\thm\ values produced by the inference kernel: \<^ML>\0\ records only the names of oracles, \<^ML>\1\ records oracle names and propositions, \<^ML>\2\ additionally records full proof terms. Officially named theorems that contribute to a result are recorded in any case. \<^descr> \<^ML>\Proofterm.reconstruct_proof\~\thy prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. Reconstruction may fail if \prf\ is not a proof of \prop\, or if it does not contain sufficient information for reconstruction. Failure may only happen for proofs that are constructed manually, but not for those produced automatically by the inference kernel. \<^descr> \<^ML>\Proofterm.expand_proof\~\thy expand prf\ reconstructs and expands the proofs of nested theorems according to the given \expand\ function: a result of @{ML \SOME ""\} means full expansion, @{ML \SOME\}~\name\ means to keep the theorem node but replace its internal name, @{ML NONE} means no change. \<^descr> \<^ML>\Proof_Checker.thm_of_proof\~\thy prf\ turns the given (full) proof into a theorem, by replaying it using only primitive rules of the inference kernel. \<^descr> \<^ML>\Proof_Syntax.read_proof\~\thy b\<^sub>1 b\<^sub>2 s\ reads in a proof term. The Boolean flags indicate the use of sort and type information. Usually, typing information is left implicit and is inferred during proof reconstruction. %FIXME eliminate flags!? \<^descr> \<^ML>\Proof_Syntax.pretty_proof\~\ctxt prf\ pretty-prints the given proof term. \ text %mlex \ \<^item> \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\ provides basic examples involving proof terms. \<^item> \<^file>\~~/src/HOL/Proofs/ex/XML_Data.thy\ demonstrates export and import of proof terms via XML/ML data representation. \ end diff --git a/src/HOL/Analysis/normarith.ML b/src/HOL/Analysis/normarith.ML --- a/src/HOL/Analysis/normarith.ML +++ b/src/HOL/Analysis/normarith.ML @@ -1,413 +1,414 @@ (* Title: HOL/Analysis/normarith.ML Author: Amine Chaieb, University of Cambridge Simple decision procedure for linear problems in Euclidean space. *) signature NORM_ARITH = sig val norm_arith : Proof.context -> conv val norm_arith_tac : Proof.context -> int -> tactic end structure NormArith : NORM_ARITH = struct open Conv; val bool_eq = op = : bool *bool -> bool fun dest_ratconst t = case Thm.term_of t of Const(\<^const_name>\divide\, _)$a$b => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) | Const(\<^const_name>\inverse\, _)$a => Rat.make(1, HOLogic.dest_number a |> snd) | _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd) fun is_ratconst t = can dest_ratconst t fun augment_norm b t acc = case Thm.term_of t of Const(\<^const_name>\norm\, _) $ _ => insert (eq_pair bool_eq (op aconvc)) (b,Thm.dest_arg t) acc | _ => acc fun find_normedterms t acc = case Thm.term_of t of \<^term>\(+) :: real => _\$_$_ => find_normedterms (Thm.dest_arg1 t) (find_normedterms (Thm.dest_arg t) acc) | \<^term>\(*) :: real => _\$_$_ => if not (is_ratconst (Thm.dest_arg1 t)) then acc else augment_norm (dest_ratconst (Thm.dest_arg1 t) >= @0) (Thm.dest_arg t) acc | _ => augment_norm true t acc val cterm_lincomb_neg = FuncUtil.Ctermfunc.map (K ~) fun cterm_lincomb_cmul c t = if c = @0 then FuncUtil.Ctermfunc.empty else FuncUtil.Ctermfunc.map (fn _ => fn x => x * c) t fun cterm_lincomb_add l r = FuncUtil.Ctermfunc.combine (curry op +) (fn x => x = @0) l r fun cterm_lincomb_sub l r = cterm_lincomb_add l (cterm_lincomb_neg r) fun cterm_lincomb_eq l r = FuncUtil.Ctermfunc.is_empty (cterm_lincomb_sub l r) (* val int_lincomb_neg = FuncUtil.Intfunc.map (K ~) *) fun int_lincomb_cmul c t = if c = @0 then FuncUtil.Intfunc.empty else FuncUtil.Intfunc.map (fn _ => fn x => x * c) t fun int_lincomb_add l r = FuncUtil.Intfunc.combine (curry op +) (fn x => x = @0) l r (* fun int_lincomb_sub l r = int_lincomb_add l (int_lincomb_neg r) fun int_lincomb_eq l r = FuncUtil.Intfunc.is_empty (int_lincomb_sub l r) *) fun vector_lincomb t = case Thm.term_of t of Const(\<^const_name>\plus\, _) $ _ $ _ => cterm_lincomb_add (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | Const(\<^const_name>\minus\, _) $ _ $ _ => cterm_lincomb_sub (vector_lincomb (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | Const(\<^const_name>\scaleR\, _)$_$_ => cterm_lincomb_cmul (dest_ratconst (Thm.dest_arg1 t)) (vector_lincomb (Thm.dest_arg t)) | Const(\<^const_name>\uminus\, _)$_ => cterm_lincomb_neg (vector_lincomb (Thm.dest_arg t)) (* FIXME: how should we handle numerals? | Const(@ {const_name vec},_)$_ => let val b = ((snd o HOLogic.dest_number o term_of o Thm.dest_arg) t = 0 handle TERM _=> false) in if b then FuncUtil.Ctermfunc.onefunc (t,@1) else FuncUtil.Ctermfunc.empty end *) | _ => FuncUtil.Ctermfunc.onefunc (t,@1) fun vector_lincombs ts = fold_rev (fn t => fn fns => case AList.lookup (op aconvc) fns t of NONE => let val f = vector_lincomb t in case find_first (fn (_,f') => cterm_lincomb_eq f f') fns of SOME (_,f') => (t,f') :: fns | NONE => (t,f) :: fns end | SOME _ => fns) ts [] fun replacenegnorms cv t = case Thm.term_of t of \<^term>\(+) :: real => _\$_$_ => binop_conv (replacenegnorms cv) t | \<^term>\(*) :: real => _\$_$_ => if dest_ratconst (Thm.dest_arg1 t) < @0 then arg_conv cv t else Thm.reflexive t | _ => Thm.reflexive t (* fun flip v eq = if FuncUtil.Ctermfunc.defined eq v then FuncUtil.Ctermfunc.update (v, ~ (FuncUtil.Ctermfunc.apply eq v)) eq else eq *) fun allsubsets s = case s of [] => [[]] |(a::t) => let val res = allsubsets t in map (cons a) res @ res end fun evaluate env lin = FuncUtil.Intfunc.fold (fn (x,c) => fn s => s + c * (FuncUtil.Intfunc.apply env x)) lin @0 fun solve (vs,eqs) = case (vs,eqs) of ([],[]) => SOME (FuncUtil.Intfunc.onefunc (0,@1)) |(_,eq::oeqs) => (case filter (member (op =) vs) (FuncUtil.Intfunc.dom eq) of (*FIXME use find_first here*) [] => NONE | v::_ => if FuncUtil.Intfunc.defined eq v then let val c = FuncUtil.Intfunc.apply eq v val vdef = int_lincomb_cmul (~ (Rat.inv c)) eq fun eliminate eqn = if not (FuncUtil.Intfunc.defined eqn v) then eqn else int_lincomb_add (int_lincomb_cmul (FuncUtil.Intfunc.apply eqn v) vdef) eqn in (case solve (remove (op =) v vs, map eliminate oeqs) of NONE => NONE | SOME soln => SOME (FuncUtil.Intfunc.update (v, evaluate soln (FuncUtil.Intfunc.delete_safe v vdef)) soln)) end else NONE) fun combinations k l = if k = 0 then [[]] else case l of [] => [] | h::t => map (cons h) (combinations (k - 1) t) @ combinations k t fun vertices vs eqs = let fun vertex cmb = case solve(vs,cmb) of NONE => NONE | SOME soln => SOME (map (fn v => FuncUtil.Intfunc.tryapplyd soln v @0) vs) val rawvs = map_filter vertex (combinations (length vs) eqs) val unset = filter (forall (fn c => c >= @0)) rawvs in fold_rev (insert (eq_list op =)) unset [] end val subsumes = eq_list (fn (x, y) => Rat.abs x <= Rat.abs y) fun subsume todo dun = case todo of [] => dun |v::ovs => let val dun' = if exists (fn w => subsumes (w, v)) dun then dun else v:: filter (fn w => not (subsumes (v, w))) dun in subsume ovs dun' end; fun match_mp PQ P = P RS PQ; fun cterm_of_rat x = let val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber \<^ctyp>\real\ a else Thm.apply (Thm.apply \<^cterm>\(/) :: real => _\ (Numeral.mk_cnumber \<^ctyp>\real\ a)) (Numeral.mk_cnumber \<^ctyp>\real\ b) end; fun norm_cmul_rule c th = Thm.instantiate' [] [SOME (cterm_of_rat c)] (th RS @{thm norm_cmul_rule_thm}); fun norm_add_rule th1 th2 = [th1, th2] MRS @{thm norm_add_rule_thm}; (* I think here the static context should be sufficient!! *) fun inequality_canon_rule ctxt = let (* FIXME : Should be computed statically!! *) val real_poly_conv = Semiring_Normalizer.semiring_normalize_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) in fconv_rule (arg_conv ((rewr_conv @{thm ge_iff_diff_ge_0}) then_conv arg_conv (Numeral_Simprocs.field_comp_conv ctxt then_conv real_poly_conv))) end; val apply_pth1 = rewr_conv @{thm pth_1}; val apply_pth2 = rewr_conv @{thm pth_2}; val apply_pth3 = rewr_conv @{thm pth_3}; val apply_pth4 = rewrs_conv @{thms pth_4}; val apply_pth5 = rewr_conv @{thm pth_5}; val apply_pth6 = rewr_conv @{thm pth_6}; val apply_pth7 = rewrs_conv @{thms pth_7}; fun apply_pth8 ctxt = rewr_conv @{thm pth_8} then_conv arg1_conv (Numeral_Simprocs.field_comp_conv ctxt) then_conv (try_conv (rewr_conv (mk_meta_eq @{thm scaleR_zero_left}))); fun apply_pth9 ctxt = rewrs_conv @{thms pth_9} then_conv arg1_conv (arg1_conv (Numeral_Simprocs.field_comp_conv ctxt)); val apply_ptha = rewr_conv @{thm pth_a}; val apply_pthb = rewrs_conv @{thms pth_b}; val apply_pthc = rewrs_conv @{thms pth_c}; val apply_pthd = try_conv (rewr_conv @{thm pth_d}); fun headvector t = case t of Const(\<^const_name>\plus\, _)$ (Const(\<^const_name>\scaleR\, _)$_$v)$_ => v | Const(\<^const_name>\scaleR\, _)$_$v => v | _ => error "headvector: non-canonical term" fun vector_cmul_conv ctxt ct = ((apply_pth5 then_conv arg1_conv (Numeral_Simprocs.field_comp_conv ctxt)) else_conv (apply_pth6 then_conv binop_conv (vector_cmul_conv ctxt))) ct fun vector_add_conv ctxt ct = apply_pth7 ct handle CTERM _ => (apply_pth8 ctxt ct handle CTERM _ => (case Thm.term_of ct of Const(\<^const_name>\plus\,_)$lt$rt => let val l = headvector lt val r = headvector rt in (case Term_Ord.fast_term_ord (l,r) of LESS => (apply_pthb then_conv arg_conv (vector_add_conv ctxt) then_conv apply_pthd) ct | GREATER => (apply_pthc then_conv arg_conv (vector_add_conv ctxt) then_conv apply_pthd) ct | EQUAL => (apply_pth9 ctxt then_conv ((apply_ptha then_conv (vector_add_conv ctxt)) else_conv arg_conv (vector_add_conv ctxt) then_conv apply_pthd)) ct) end | _ => Thm.reflexive ct)) fun vector_canon_conv ctxt ct = case Thm.term_of ct of Const(\<^const_name>\plus\,_)$_$_ => let val ((p,l),r) = Thm.dest_comb ct |>> Thm.dest_comb val lth = vector_canon_conv ctxt l val rth = vector_canon_conv ctxt r val th = Drule.binop_cong_rule p lth rth in fconv_rule (arg_conv (vector_add_conv ctxt)) th end | Const(\<^const_name>\scaleR\, _)$_$_ => let val (p,r) = Thm.dest_comb ct val rth = Drule.arg_cong_rule p (vector_canon_conv ctxt r) in fconv_rule (arg_conv (apply_pth4 else_conv (vector_cmul_conv ctxt))) rth end | Const(\<^const_name>\minus\,_)$_$_ => (apply_pth2 then_conv (vector_canon_conv ctxt)) ct | Const(\<^const_name>\uminus\,_)$_ => (apply_pth3 then_conv (vector_canon_conv ctxt)) ct (* FIXME | Const(@{const_name vec},_)$n => let val n = Thm.dest_arg ct in if is_ratconst n andalso not (dest_ratconst n =/ @0) then Thm.reflexive ct else apply_pth1 ct end *) | _ => apply_pth1 ct fun norm_canon_conv ctxt ct = case Thm.term_of ct of Const(\<^const_name>\norm\,_)$_ => arg_conv (vector_canon_conv ctxt) ct | _ => raise CTERM ("norm_canon_conv", [ct]) fun int_flip v eq = if FuncUtil.Intfunc.defined eq v then FuncUtil.Intfunc.update (v, ~ (FuncUtil.Intfunc.apply eq v)) eq else eq; local val pth_zero = @{thm norm_zero} val tv_n = (dest_TVar o Thm.typ_of_cterm o Thm.dest_arg o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of) pth_zero val concl = Thm.dest_arg o Thm.cprop_of fun real_vector_combo_prover ctxt translator (nubs,ges,gts) = let (* FIXME: Should be computed statically!!*) val real_poly_conv = Semiring_Normalizer.semiring_normalize_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) val sources = map (Thm.dest_arg o Thm.dest_arg1 o concl) nubs val rawdests = fold_rev (find_normedterms o Thm.dest_arg o concl) (ges @ gts) [] val _ = if not (forall fst rawdests) then error "real_vector_combo_prover: Sanity check" else () val dests = distinct (op aconvc) (map snd rawdests) val srcfuns = map vector_lincomb sources val destfuns = map vector_lincomb dests val vvs = fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) (srcfuns @ destfuns) [] val n = length srcfuns val nvs = 1 upto n val srccombs = srcfuns ~~ nvs fun consider d = let fun coefficients x = let val inp = if FuncUtil.Ctermfunc.defined d x then FuncUtil.Intfunc.onefunc (0, ~ (FuncUtil.Ctermfunc.apply d x)) else FuncUtil.Intfunc.empty in fold_rev (fn (f,v) => fn g => if FuncUtil.Ctermfunc.defined f x then FuncUtil.Intfunc.update (v, FuncUtil.Ctermfunc.apply f x) g else g) srccombs inp end val equations = map coefficients vvs val inequalities = map (fn n => FuncUtil.Intfunc.onefunc (n,@1)) nvs fun plausiblevertices f = let val flippedequations = map (fold_rev int_flip f) equations val constraints = flippedequations @ inequalities val rawverts = vertices nvs constraints fun check_solution v = let val f = fold_rev FuncUtil.Intfunc.update (nvs ~~ v) (FuncUtil.Intfunc.onefunc (0, @1)) in forall (fn e => evaluate f e = @0) flippedequations end val goodverts = filter check_solution rawverts val signfixups = map (fn n => if member (op =) f n then ~1 else 1) nvs in map (map2 (fn s => fn c => Rat.of_int s * c) signfixups) goodverts end val allverts = fold_rev append (map plausiblevertices (allsubsets nvs)) [] in subsume allverts [] end fun compute_ineq v = let val ths = map_filter (fn (v,t) => if v = @0 then NONE else SOME(norm_cmul_rule v t)) (v ~~ nubs) fun end_itlist f xs = split_last xs |> uncurry (fold_rev f) in inequality_canon_rule ctxt (end_itlist norm_add_rule ths) end val ges' = map_filter (try compute_ineq) (fold_rev (append o consider) destfuns []) @ map (inequality_canon_rule ctxt) nubs @ ges val zerodests = filter (fn t => null (FuncUtil.Ctermfunc.dom (vector_lincomb t))) (map snd rawdests) in fst (RealArith.real_linear_prover translator - (map (fn t => Drule.instantiate_normalize ([(tv_n, Thm.ctyp_of_cterm t)],[]) pth_zero) + (map (fn t => + Drule.instantiate_normalize (TVars.make [(tv_n, Thm.ctyp_of_cterm t)], Vars.empty) pth_zero) zerodests, map (fconv_rule (try_conv (Conv.top_sweep_conv (K (norm_canon_conv ctxt)) ctxt) then_conv arg_conv (arg_conv real_poly_conv))) ges', map (fconv_rule (try_conv (Conv.top_sweep_conv (K (norm_canon_conv ctxt)) ctxt) then_conv arg_conv (arg_conv real_poly_conv))) gts)) end in val real_vector_combo_prover = real_vector_combo_prover end; local val pth = @{thm norm_imp_pos_and_ge} val norm_mp = match_mp pth val concl = Thm.dest_arg o Thm.cprop_of fun conjunct1 th = th RS @{thm conjunct1} fun conjunct2 th = th RS @{thm conjunct2} fun real_vector_ineq_prover ctxt translator (ges,gts) = let (* val _ = error "real_vector_ineq_prover: pause" *) val ntms = fold_rev find_normedterms (map (Thm.dest_arg o concl) (ges @ gts)) [] val lctab = vector_lincombs (map snd (filter (not o fst) ntms)) val (fxns, ctxt') = Variable.variant_fixes (replicate (length lctab) "x") ctxt fun instantiate_cterm' ty tms = Drule.cterm_rule (Thm.instantiate' ty tms) fun mk_norm t = let val T = Thm.typ_of_cterm t in Thm.apply (Thm.cterm_of ctxt' (Const (\<^const_name>\norm\, T --> \<^typ>\real\))) t end fun mk_equals l r = let val T = Thm.typ_of_cterm l val eq = Thm.cterm_of ctxt (Const (\<^const_name>\Pure.eq\, T --> T --> propT)) in Thm.apply (Thm.apply eq l) r end val asl = map2 (fn (t,_) => fn n => Thm.assume (mk_equals (mk_norm t) (Thm.cterm_of ctxt' (Free(n,\<^typ>\real\))))) lctab fxns val replace_conv = try_conv (rewrs_conv asl) val replace_rule = fconv_rule (funpow 2 arg_conv (replacenegnorms replace_conv)) val ges' = fold_rev (fn th => fn ths => conjunct1(norm_mp th)::ths) asl (map replace_rule ges) val gts' = map replace_rule gts val nubs = map (conjunct2 o norm_mp) asl val th1 = real_vector_combo_prover ctxt' translator (nubs,ges',gts') val shs = filter (member (fn (t,th) => t aconvc Thm.cprop_of th) asl) (Thm.chyps_of th1) val th11 = hd (Variable.export ctxt' ctxt [fold Thm.implies_intr shs th1]) val cps = map (swap o Thm.dest_equals) (cprems_of th11) - val th12 = Drule.instantiate_normalize ([], map (apfst (dest_Var o Thm.term_of)) cps) th11 + val th12 = Drule.instantiate_normalize (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) cps)) th11 val th13 = fold Thm.elim_implies (map (Thm.reflexive o snd) cps) th12; in hd (Variable.export ctxt' ctxt [th13]) end in val real_vector_ineq_prover = real_vector_ineq_prover end; local val rawrule = fconv_rule (arg_conv (rewr_conv @{thm real_eq_0_iff_le_ge_0})) fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2}) (* FIXME: Lookup in the context every time!!! Fix this !!!*) fun splitequation ctxt th acc = let val real_poly_neg_conv = #neg (Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) Thm.term_ord) val (th1,th2) = conj_pair(rawrule th) in th1::fconv_rule (arg_conv (arg_conv (real_poly_neg_conv ctxt))) th2::acc end in fun real_vector_prover ctxt _ translator (eqs,ges,gts) = (real_vector_ineq_prover ctxt translator (fold_rev (splitequation ctxt) eqs ges,gts), RealArith.Trivial) end; fun init_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps ([(*@{thm vec_0}, @{thm vec_1},*) @{thm dist_norm}, @{thm right_minus}, @{thm diff_self}, @{thm norm_zero}] @ @{thms arithmetic_simps} @ @{thms norm_pths})) then_conv Numeral_Simprocs.field_comp_conv ctxt then_conv nnf_conv ctxt fun pure ctxt = fst o RealArith.gen_prover_real_arith ctxt (real_vector_prover ctxt); fun norm_arith ctxt ct = let val ctxt' = Variable.declare_term (Thm.term_of ct) ctxt val th = init_conv ctxt' ct in Thm.equal_elim (Drule.arg_cong_rule \<^cterm>\Trueprop\ (Thm.symmetric th)) (pure ctxt' (Thm.rhs_of th)) end fun norm_arith_tac ctxt = clarify_tac (put_claset HOL_cs ctxt) THEN' Object_Logic.full_atomize_tac ctxt THEN' CSUBGOAL ( fn (p,i) => resolve_tac ctxt [norm_arith ctxt (Thm.dest_arg p )] i); end; diff --git a/src/HOL/Decision_Procs/approximation.ML b/src/HOL/Decision_Procs/approximation.ML --- a/src/HOL/Decision_Procs/approximation.ML +++ b/src/HOL/Decision_Procs/approximation.ML @@ -1,294 +1,296 @@ (* Title: HOL/Decision_Procs/approximation.ML Author: Johannes Hoelzl, TU Muenchen *) signature APPROXIMATION = sig val reify_form: Proof.context -> term -> term val approx: int -> Proof.context -> term -> term val approximate: Proof.context -> term -> term val approximation_tac : int -> (string * int) list -> int option -> Proof.context -> int -> tactic end structure Approximation = struct fun reorder_bounds_tac ctxt prems i = let fun variable_of_bound (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Set.member\, _) $ Free (name, _) $ _)) = name | variable_of_bound (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ Free (name, _) $ _)) = name | variable_of_bound t = raise TERM ("variable_of_bound", [t]) val variable_bounds = map (`(variable_of_bound o Thm.prop_of)) prems fun add_deps (name, bnds) = Graph.add_deps_acyclic (name, remove (op =) name (Term.add_free_names (Thm.prop_of bnds) [])) val order = Graph.empty |> fold Graph.new_node variable_bounds |> fold add_deps variable_bounds |> Graph.strong_conn |> map the_single |> rev |> map_filter (AList.lookup (op =) variable_bounds) fun prepend_prem th tac = tac THEN resolve_tac ctxt [th RSN (2, @{thm mp})] i in fold prepend_prem order all_tac end fun approximate ctxt t = case fastype_of t of \<^typ>\bool\ => Approximation_Computation.approx_bool ctxt t | \<^typ>\(float interval) option\ => Approximation_Computation.approx_arith ctxt t | \<^typ>\(float interval) option list\ => Approximation_Computation.approx_form_eval ctxt t | _ => error ("Bad term: " ^ Syntax.string_of_term ctxt t); fun rewrite_interpret_form_tac ctxt prec splitting taylor i st = let fun lookup_splitting (Free (name, _)) = (case AList.lookup (op =) splitting name of SOME s => HOLogic.mk_number \<^typ>\nat\ s | NONE => \<^term>\0 :: nat\) | lookup_splitting t = raise TERM ("lookup_splitting", [t]) val vs = nth (Thm.prems_of st) (i - 1) |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> Term.strip_comb |> snd |> List.last |> HOLogic.dest_list val p = prec |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt in case taylor of NONE => let val n = vs |> length |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> HOLogic.mk_list \<^typ>\nat\ |> Thm.cterm_of ctxt in - (resolve_tac ctxt [Thm.instantiate ([], [((("n", 0), \<^typ>\nat\), n), - ((("prec", 0), \<^typ>\nat\), p), - ((("ss", 0), \<^typ>\nat list\), s)]) + (resolve_tac ctxt [Thm.instantiate (TVars.empty, + Vars.make [((("n", 0), \<^typ>\nat\), n), + ((("prec", 0), \<^typ>\nat\), p), + ((("ss", 0), \<^typ>\nat list\), s)]) @{thm approx_form}] i THEN simp_tac (put_simpset (simpset_of \<^context>) ctxt) i) st end | SOME t => if length vs <> 1 then raise (TERM ("More than one variable used for taylor series expansion", [Thm.prop_of st])) else let val t = t |> HOLogic.mk_number \<^typ>\nat\ |> Thm.cterm_of ctxt val s = vs |> map lookup_splitting |> hd |> Thm.cterm_of ctxt in - resolve_tac ctxt [Thm.instantiate ([], [((("s", 0), \<^typ>\nat\), s), - ((("t", 0), \<^typ>\nat\), t), - ((("prec", 0), \<^typ>\nat\), p)]) + resolve_tac ctxt [Thm.instantiate (TVars.empty, + Vars.make [((("s", 0), \<^typ>\nat\), s), + ((("t", 0), \<^typ>\nat\), t), + ((("prec", 0), \<^typ>\nat\), p)]) @{thm approx_tse_form}] i st end end fun calculated_subterms (\<^const>\Trueprop\ $ t) = calculated_subterms t | calculated_subterms (\<^const>\HOL.implies\ $ _ $ t) = calculated_subterms t | calculated_subterms (\<^term>\(\) :: real \ real \ bool\ $ t1 $ t2) = [t1, t2] | calculated_subterms (\<^term>\(<) :: real \ real \ bool\ $ t1 $ t2) = [t1, t2] | calculated_subterms (\<^term>\(\) :: real \ real set \ bool\ $ t1 $ (\<^term>\atLeastAtMost :: real \ real \ real set\ $ t2 $ t3)) = [t1, t2, t3] | calculated_subterms t = raise TERM ("calculated_subterms", [t]) fun dest_interpret_form (\<^const>\interpret_form\ $ b $ xs) = (b, xs) | dest_interpret_form t = raise TERM ("dest_interpret_form", [t]) fun dest_interpret (\<^const>\interpret_floatarith\ $ b $ xs) = (b, xs) | dest_interpret t = raise TERM ("dest_interpret", [t]) fun dest_interpret_env (\<^const>\interpret_form\ $ _ $ xs) = xs | dest_interpret_env (\<^const>\interpret_floatarith\ $ _ $ xs) = xs | dest_interpret_env t = raise TERM ("dest_interpret_env", [t]) fun dest_float (\<^const>\Float\ $ m $ e) = (snd (HOLogic.dest_number m), snd (HOLogic.dest_number e)) | dest_float t = raise TERM ("dest_float", [t]) fun dest_ivl (Const (\<^const_name>\Some\, _) $ (Const (\<^const_name>\Interval\, _) $ ((Const (\<^const_name>\Pair\, _) $ u $ l)))) = SOME (dest_float u, dest_float l) | dest_ivl (Const (\<^const_name>\None\, _)) = NONE | dest_ivl t = raise TERM ("dest_result", [t]) fun mk_approx' prec t = (\<^const>\approx'\ $ HOLogic.mk_number \<^typ>\nat\ prec $ t $ \<^term>\[] :: (float interval) option list\) fun mk_approx_form_eval prec t xs = (\<^const>\approx_form_eval\ $ HOLogic.mk_number \<^typ>\nat\ prec $ t $ xs) fun float2_float10 prec round_down (m, e) = ( let val (m, e) = (if e < 0 then (m,e) else (m * Integer.pow e 2, 0)) fun frac _ _ 0 digits cnt = (digits, cnt, 0) | frac _ 0 r digits cnt = (digits, cnt, r) | frac c p r digits cnt = (let val (d, r) = Integer.div_mod (r * 10) (Integer.pow (~e) 2) in frac (c orelse d <> 0) (if d <> 0 orelse c then p - 1 else p) r (digits * 10 + d) (cnt + 1) end) val sgn = Int.sign m val m = abs m val round_down = (sgn = 1 andalso round_down) orelse (sgn = ~1 andalso not round_down) val (x, r) = Integer.div_mod m (Integer.pow (~e) 2) val p = ((if x = 0 then prec else prec - (Integer.log2 x + 1)) * 3) div 10 + 1 val (digits, e10, r) = if p > 0 then frac (x <> 0) p r 0 0 else (0,0,0) val digits = if round_down orelse r = 0 then digits else digits + 1 in (sgn * (digits + x * (Integer.pow e10 10)), ~e10) end) fun mk_result prec (SOME (l, u)) = (let fun mk_float10 rnd x = (let val (m, e) = float2_float10 prec rnd x in if e = 0 then HOLogic.mk_number \<^typ>\real\ m else if e = 1 then \<^term>\divide :: real \ real \ real\ $ HOLogic.mk_number \<^typ>\real\ m $ \<^term>\10\ else \<^term>\divide :: real \ real \ real\ $ HOLogic.mk_number \<^typ>\real\ m $ (\<^term>\power 10 :: nat \ real\ $ HOLogic.mk_number \<^typ>\nat\ (~e)) end) in \<^term>\atLeastAtMost :: real \ real \ real set\ $ mk_float10 true l $ mk_float10 false u end) | mk_result _ NONE = \<^term>\UNIV :: real set\ fun realify t = let val t = Logic.varify_global t val m = map (fn (name, _) => (name, \<^typ>\real\)) (Term.add_tvars t []) val t = Term.subst_TVars m t in t end fun apply_tactic ctxt term tactic = Thm.cterm_of ctxt term |> Goal.init |> SINGLE tactic |> the |> Thm.prems_of |> hd fun preproc_form_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps (Named_Theorems.get ctxt \<^named_theorems>\approximation_preproc\)) fun reify_form_conv ctxt ct = let val thm = Reification.conv ctxt @{thms interpret_form.simps interpret_floatarith.simps} ct handle ERROR msg => cat_error ("Reification failed: " ^ msg) ("Approximation does not support " ^ quote (Syntax.string_of_term ctxt (Thm.term_of ct))) fun check_env (Free _) = () | check_env (Var _) = () | check_env t = cat_error "Term not supported by approximation:" (Syntax.string_of_term ctxt t) val _ = Thm.rhs_of thm |> Thm.term_of |> dest_interpret_env |> HOLogic.dest_list |> map check_env in thm end fun reify_form_tac ctxt i = CONVERSION (Conv.arg_conv (reify_form_conv ctxt)) i fun prepare_form_tac ctxt i = REPEAT (FIRST' [eresolve_tac ctxt @{thms intervalE}, eresolve_tac ctxt @{thms meta_eqE}, resolve_tac ctxt @{thms impI}] i) THEN Subgoal.FOCUS (fn {prems, context = ctxt', ...} => reorder_bounds_tac ctxt' prems i) ctxt i THEN DETERM (TRY (filter_prems_tac ctxt (K false) i)) THEN CONVERSION (Conv.arg_conv (preproc_form_conv ctxt)) i fun prepare_form ctxt term = apply_tactic ctxt term (prepare_form_tac ctxt 1) fun apply_reify_form ctxt t = apply_tactic ctxt t (reify_form_tac ctxt 1) fun reify_form ctxt t = HOLogic.mk_Trueprop t |> prepare_form ctxt |> apply_reify_form ctxt |> HOLogic.dest_Trueprop fun approx_form prec ctxt t = realify t |> prepare_form ctxt |> (fn arith_term => apply_reify_form ctxt arith_term |> HOLogic.dest_Trueprop |> dest_interpret_form |> (fn (data, xs) => mk_approx_form_eval prec data (HOLogic.mk_list \<^typ>\(float interval) option\ (map (fn _ => \<^term>\None :: (float interval) option\) (HOLogic.dest_list xs))) |> approximate ctxt |> HOLogic.dest_list |> curry ListPair.zip (HOLogic.dest_list xs @ calculated_subterms arith_term) |> map (fn (elem, s) => \<^term>\(\) :: real \ real set \ bool\ $ elem $ mk_result prec (dest_ivl s)) |> foldr1 HOLogic.mk_conj)) fun approx_arith prec ctxt t = realify t |> Thm.cterm_of ctxt |> (preproc_form_conv ctxt then_conv reify_form_conv ctxt) |> Thm.prop_of |> Logic.dest_equals |> snd |> dest_interpret |> fst |> mk_approx' prec |> approximate ctxt |> dest_ivl |> mk_result prec fun approx prec ctxt t = if type_of t = \<^typ>\prop\ then approx_form prec ctxt t else if type_of t = \<^typ>\bool\ then approx_form prec ctxt (\<^const>\Trueprop\ $ t) else approx_arith prec ctxt t fun approximate_cmd modes raw_t state = let val ctxt = Toplevel.context_of state; val t = Syntax.read_term ctxt raw_t; val t' = approx 30 ctxt t; val ty' = Term.type_of t'; val ctxt' = Proof_Context.augment t' ctxt; in Print_Mode.with_modes modes (fn () => Pretty.block [Pretty.quote (Syntax.pretty_term ctxt' t'), Pretty.fbrk, Pretty.str "::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ ctxt' ty')]) () end |> Pretty.writeln; val opt_modes = Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) []; val _ = Outer_Syntax.command \<^command_keyword>\approximate\ "print approximation of term" (opt_modes -- Parse.term >> (fn (modes, t) => Toplevel.keep (approximate_cmd modes t))); fun approximation_tac prec splitting taylor ctxt = prepare_form_tac ctxt THEN' reify_form_tac ctxt THEN' rewrite_interpret_form_tac ctxt prec splitting taylor THEN' CONVERSION (Approximation_Computation.approx_conv ctxt) THEN' resolve_tac ctxt [TrueI] end; diff --git a/src/HOL/Decision_Procs/ferrante_rackoff.ML b/src/HOL/Decision_Procs/ferrante_rackoff.ML --- a/src/HOL/Decision_Procs/ferrante_rackoff.ML +++ b/src/HOL/Decision_Procs/ferrante_rackoff.ML @@ -1,234 +1,234 @@ (* Title: HOL/Decision_Procs/ferrante_rackoff.ML Author: Amine Chaieb, TU Muenchen Ferrante and Rackoff's algorithm for quantifier elimination in dense linear orders. Proof-synthesis and tactic. *) signature FERRANTE_RACKOFF = sig val dlo_conv: Proof.context -> conv val dlo_tac: Proof.context -> int -> tactic end; structure FerranteRackoff: FERRANTE_RACKOFF = struct open Ferrante_Rackoff_Data; open Conv; type entry = {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list, ld: thm list, qe: thm, atoms : cterm list} * {isolate_conv: cterm list -> cterm -> thm, whatis : cterm -> cterm -> ord, simpset : simpset}; fun get_p1 th = funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) (funpow 2 Thm.dest_arg (Thm.cprop_of th)) |> Thm.dest_arg fun ferrack_conv ctxt (entr as ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = ld, qe = qe, atoms = atoms}, {isolate_conv = icv, whatis = wi, simpset = simpset}):entry) = let fun uset (vars as (x::vs)) p = case Thm.term_of p of Const(\<^const_name>\HOL.conj\, _)$ _ $ _ => let val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb val (lS,lth) = uset vars l val (rS, rth) = uset vars r in (lS@rS, Drule.binop_cong_rule b lth rth) end | Const(\<^const_name>\HOL.disj\, _)$ _ $ _ => let val ((b,l),r) = Thm.dest_comb p |>> Thm.dest_comb val (lS,lth) = uset vars l val (rS, rth) = uset vars r in (lS@rS, Drule.binop_cong_rule b lth rth) end | _ => let val th = icv vars p val p' = Thm.rhs_of th val c = wi x p' val S = (if member (op =) [Lt, Le, Eq] c then single o Thm.dest_arg else if member (op =) [Gt, Ge] c then single o Thm.dest_arg1 else if c = NEq then single o Thm.dest_arg o Thm.dest_arg else K []) p' in (S,th) end val ((p1_v,p2_v),(mp1_v,mp2_v)) = funpow 2 (Thm.dest_arg o snd o Thm.dest_abs NONE) (funpow 4 Thm.dest_arg (Thm.cprop_of (hd minf))) |> Thm.dest_binop |> apply2 Thm.dest_binop |> apfst (apply2 Thm.dest_fun) |> apply2 (apply2 (dest_Var o Thm.term_of)) fun myfwd (th1, th2, th3, th4, th5) p1 p2 [(th_1,th_2,th_3,th_4,th_5), (th_1',th_2',th_3',th_4',th_5')] = let val (mp1, mp2) = (get_p1 th_1, get_p1 th_1') val (pp1, pp2) = (get_p1 th_2, get_p1 th_2') fun fw mi th th' th'' = let val th0 = if mi then - Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th - else Drule.instantiate_normalize ([],[(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th + Drule.instantiate_normalize (TVars.empty, Vars.make [(p1_v, p1),(p2_v, p2),(mp1_v, mp1), (mp2_v, mp2)]) th + else Drule.instantiate_normalize (TVars.empty, Vars.make [(p1_v, p1),(p2_v, p2),(mp1_v, pp1), (mp2_v, pp2)]) th in Thm.implies_elim (Thm.implies_elim th0 th') th'' end in (fw true th1 th_1 th_1', fw false th2 th_2 th_2', fw true th3 th_3 th_3', fw false th4 th_4 th_4', fw true th5 th_5 th_5') end val U_v = dest_Var (Thm.term_of (Thm.dest_arg (Thm.dest_arg (Thm.dest_arg1 (Thm.cprop_of qe))))) fun main vs p = let val ((xn,ce),(x,fm)) = (case Thm.term_of p of Const(\<^const_name>\Ex\,_)$Abs(xn,xT,_) => Thm.dest_comb p ||> Thm.dest_abs (SOME xn) |>> pair xn | _ => raise CTERM ("main QE only treats existential quantifiers!", [p])) val cT = Thm.ctyp_of_cterm x val (u,nth) = uset (x::vs) fm |>> distinct (op aconvc) val nthx = Thm.abstract_rule xn x nth val q = Thm.rhs_of nth val qx = Thm.rhs_of nthx val enth = Drule.arg_cong_rule ce nthx val [th0,th1] = map (Thm.instantiate' [SOME cT] []) @{thms "finite.intros"} fun ins x th = Thm.implies_elim (Thm.instantiate' [] [(SOME o Thm.dest_arg o Thm.dest_arg) (Thm.cprop_of th), SOME x] th1) th val fU = fold ins u th0 val cU = funpow 2 Thm.dest_arg (Thm.cprop_of fU) local val insI1 = Thm.instantiate' [SOME cT] [] @{thm "insertI1"} val insI2 = Thm.instantiate' [SOME cT] [] @{thm "insertI2"} in fun provein x S = case Thm.term_of S of Const(\<^const_name>\Orderings.bot\, _) => raise CTERM ("provein : not a member!", [S]) | Const(\<^const_name>\insert\, _) $ y $_ => let val (cy,S') = Thm.dest_binop S in if Thm.term_of x aconv y then Thm.instantiate' [] [SOME x, SOME S'] insI1 else Thm.implies_elim (Thm.instantiate' [] [SOME x, SOME S', SOME cy] insI2) (provein x S') end end val tabU = fold (fn t => fn tab => Termtab.update (Thm.term_of t, provein t cU) tab) u Termtab.empty val U = the o Termtab.lookup tabU o Thm.term_of val [minf_conj, minf_disj, minf_eq, minf_neq, minf_lt, minf_le, minf_gt, minf_ge, minf_P] = minf val [pinf_conj, pinf_disj, pinf_eq, pinf_neq, pinf_lt, pinf_le, pinf_gt, pinf_ge, pinf_P] = pinf val [nmi_conj, nmi_disj, nmi_eq, nmi_neq, nmi_lt, - nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) nmi + nmi_le, nmi_gt, nmi_ge, nmi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) nmi val [npi_conj, npi_disj, npi_eq, npi_neq, npi_lt, - npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) npi + npi_le, npi_gt, npi_ge, npi_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) npi val [ld_conj, ld_disj, ld_eq, ld_neq, ld_lt, - ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize ([],[(U_v,cU)])) ld + ld_le, ld_gt, ld_ge, ld_P] = map (Drule.instantiate_normalize (TVars.empty, Vars.make [(U_v,cU)])) ld fun decomp_mpinf fm = case Thm.term_of fm of Const(\<^const_name>\HOL.conj\,_)$_$_ => let val (p,q) = Thm.dest_binop fm in ([p,q], myfwd (minf_conj,pinf_conj, nmi_conj, npi_conj,ld_conj) (Thm.lambda x p) (Thm.lambda x q)) end | Const(\<^const_name>\HOL.disj\,_)$_$_ => let val (p,q) = Thm.dest_binop fm in ([p,q],myfwd (minf_disj, pinf_disj, nmi_disj, npi_disj,ld_disj) (Thm.lambda x p) (Thm.lambda x q)) end | _ => (let val c = wi x fm val t = (if c=Nox then I else if member (op =) [Lt, Le, Eq] c then Thm.dest_arg else if member (op =) [Gt, Ge] c then Thm.dest_arg1 else if c = NEq then (Thm.dest_arg o Thm.dest_arg) else raise Fail "decomp_mpinf: Impossible case!!") fm val [mi_th, pi_th, nmi_th, npi_th, ld_th] = if c = Nox then map (Thm.instantiate' [] [SOME fm]) [minf_P, pinf_P, nmi_P, npi_P, ld_P] else let val [mi_th,pi_th,nmi_th,npi_th,ld_th] = map (Thm.instantiate' [] [SOME t]) (case c of Lt => [minf_lt, pinf_lt, nmi_lt, npi_lt, ld_lt] | Le => [minf_le, pinf_le, nmi_le, npi_le, ld_le] | Gt => [minf_gt, pinf_gt, nmi_gt, npi_gt, ld_gt] | Ge => [minf_ge, pinf_ge, nmi_ge, npi_ge, ld_ge] | Eq => [minf_eq, pinf_eq, nmi_eq, npi_eq, ld_eq] | NEq => [minf_neq, pinf_neq, nmi_neq, npi_neq, ld_neq]) val tU = U t fun Ufw th = Thm.implies_elim th tU in [mi_th, pi_th, Ufw nmi_th, Ufw npi_th, Ufw ld_th] end in ([], K (mi_th, pi_th, nmi_th, npi_th, ld_th)) end) val (minf_th, pinf_th, nmi_th, npi_th, ld_th) = divide_and_conquer decomp_mpinf q val qe_th = Drule.implies_elim_list ((fconv_rule (Thm.beta_conversion true)) (Thm.instantiate' [] (map SOME [cU, qx, get_p1 minf_th, get_p1 pinf_th]) qe)) [fU, ld_th, nmi_th, npi_th, minf_th, pinf_th] val bex_conv = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms bex_simps(1-5)}) val result_th = fconv_rule (arg_conv bex_conv) (Thm.transitive enth qe_th) in result_th end in main end; val grab_atom_bop = let fun h bounds tm = (case Thm.term_of tm of Const (\<^const_name>\HOL.eq\, T) $ _ $ _ => if domain_type T = HOLogic.boolT then find_args bounds tm else Thm.dest_fun2 tm | Const (\<^const_name>\Not\, _) $ _ => h bounds (Thm.dest_arg tm) | Const (\<^const_name>\All\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Ex\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.disj\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.imp\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => find_args bounds tm | Const (\<^const_name>\Pure.all\, _) $ _ => find_body bounds (Thm.dest_arg tm) | Const (\<^const_name>\Trueprop\, _) $ _ => h bounds (Thm.dest_arg tm) | _ => Thm.dest_fun2 tm) and find_args bounds tm = (h bounds (Thm.dest_arg tm) handle CTERM _ => Thm.dest_arg1 tm) and find_body bounds b = let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b in h (bounds + 1) b' end; in h end; fun raw_ferrack_qe_conv ctxt (thy, {isolate_conv, whatis, simpset = ss}) tm = let val ss' = merge_ss (simpset_of (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps all_simps not_all all_not_ex ex_disj_distrib}), ss); val pcv = Simplifier.rewrite (put_simpset ss' ctxt); val postcv = Simplifier.rewrite (put_simpset ss ctxt); val nnf = K (nnf_conv ctxt then_conv postcv) val env = Cterms.list_set_rev (Cterms.build (Drule.add_frees_cterm tm)) val qe_conv = Qelim.gen_qelim_conv ctxt pcv postcv pcv cons env (isolate_conv ctxt) nnf (fn vs => ferrack_conv ctxt (thy,{isolate_conv = isolate_conv ctxt, whatis = whatis, simpset = ss}) vs then_conv postcv) in (Simplifier.rewrite (put_simpset ss ctxt) then_conv qe_conv) tm end; fun dlo_instance ctxt tm = Ferrante_Rackoff_Data.match ctxt (grab_atom_bop 0 tm); fun dlo_conv ctxt tm = (case dlo_instance ctxt tm of NONE => raise CTERM ("ferrackqe_conv: no corresponding instance in context!", [tm]) | SOME instance => raw_ferrack_qe_conv ctxt instance tm); fun dlo_tac ctxt = CSUBGOAL (fn (p, i) => (case dlo_instance ctxt p of NONE => no_tac | SOME instance => Object_Logic.full_atomize_tac ctxt i THEN simp_tac (put_simpset (#simpset (snd instance)) ctxt) i THEN (* FIXME already part of raw_ferrack_qe_conv? *) CONVERSION (Object_Logic.judgment_conv ctxt (raw_ferrack_qe_conv ctxt instance)) i THEN simp_tac ctxt i)); (* FIXME really? *) 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,141 +1,141 @@ (* 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 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 = []; val extend = I; 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, []); + 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/Decision_Procs/langford_data.ML b/src/HOL/Decision_Procs/langford_data.ML --- a/src/HOL/Decision_Procs/langford_data.ML +++ b/src/HOL/Decision_Procs/langford_data.ML @@ -1,113 +1,113 @@ signature LANGFORD_DATA = sig type entry val get: Proof.context -> simpset * (thm * entry) list val add: entry -> attribute val del: attribute val match: Proof.context -> cterm -> entry option end; structure Langford_Data: LANGFORD_DATA = struct (* data *) type entry = {qe_bnds: thm, qe_nolb : thm , qe_noub: thm, gs : thm list, gst : thm, atoms: cterm list}; val eq_key = Thm.eq_thm; fun eq_data arg = eq_fst eq_key arg; structure Data = Generic_Data ( type T = simpset * (thm * entry) list; val empty = (HOL_basic_ss, []); val extend = I; fun merge ((ss1, es1), (ss2, es2)) : T = (merge_ss (ss1, ss2), AList.merge eq_key (K true) (es1, es2)); ); val get = Data.get o Context.Proof; fun del_data key = apsnd (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 #> apsnd (cons (key, entry)))); val add_simp = Thm.declaration_attribute (fn th => fn context => (Data.map o apfst) (simpset_map (Context.proof_of context) (Simplifier.add_simp th)) context); val del_simp = Thm.declaration_attribute (fn th => fn context => (Data.map o apfst) (simpset_map (Context.proof_of context) (Simplifier.del_simp th)) context); fun match ctxt tm = let fun match_inst {qe_bnds, qe_nolb, qe_noub, gs, gst, atoms} pat = let fun h instT = let - val substT = Thm.instantiate (instT, []); + val substT = Thm.instantiate (instT, Vars.empty); val substT_cterm = Drule.cterm_rule substT; val qe_bnds' = substT qe_bnds; val qe_nolb' = substT qe_nolb; val qe_noub' = substT qe_noub; val gs' = map substT gs; val gst' = substT gst; val atoms' = map substT_cterm atoms; val result = {qe_bnds = qe_bnds', qe_nolb = qe_nolb', qe_noub = qe_noub', gs = gs', gst = gst', atoms = atoms'}; 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 (snd (get ctxt)) end; (* concrete syntax *) local val qeN = "qe"; val gatherN = "gather"; val atomsN = "atoms" fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); val any_keyword = keyword qeN || keyword gatherN || keyword atomsN; 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>\langford\ ((keyword qeN |-- thms) -- (keyword gatherN |-- thms) -- (keyword atomsN |-- terms) >> (fn ((qes, gs), atoms) => if length qes = 3 andalso length gs > 1 then let val [q1, q2, q3] = qes; val gst :: gss = gs; val entry = {qe_bnds = q1, qe_nolb = q2, qe_noub = q3, gs = gss, gst = gst, atoms = atoms}; in add entry end else error "Need 3 theorems for qe and at least one for gs")) "Langford data"); end; val _ = Theory.setup (Attrib.setup \<^binding>\langfordsimp\ (Attrib.add_del add_simp del_simp) "Langford simpset"); end; diff --git a/src/HOL/Eisbach/eisbach_rule_insts.ML b/src/HOL/Eisbach/eisbach_rule_insts.ML --- a/src/HOL/Eisbach/eisbach_rule_insts.ML +++ b/src/HOL/Eisbach/eisbach_rule_insts.ML @@ -1,233 +1,233 @@ (* Title: HOL/Eisbach/eisbach_rule_insts.ML Author: Daniel Matichuk, NICTA/UNSW Eisbach-aware variants of the "where" and "of" attributes. Alternate syntax for rule_insts.ML participates in token closures by examining the behaviour of Rule_Insts.where_rule and instantiating token values accordingly. Instantiations in re-interpretation are done with infer_instantiate. *) structure Eisbach_Rule_Insts: sig end = struct fun restore_tags thm = Thm.map_tags (K (Thm.get_tags thm)); val mk_term_type_internal = Logic.protect o Logic.mk_term o Logic.mk_type; fun term_type_cases f g t = (case \<^try>\Logic.dest_type (Logic.dest_term (Logic.unprotect t))\ of SOME T => f T | NONE => (case \<^try>\Logic.dest_term t\ of SOME t => g t | NONE => raise Fail "Lost encoded instantiation")) fun add_thm_insts ctxt thm = let val tyvars = Thm.fold_terms {hyps = false} Term.add_tvars thm []; val tyvars' = tyvars |> map (mk_term_type_internal o TVar); val tvars = Thm.fold_terms {hyps = false} Term.add_vars thm []; val tvars' = tvars |> map (Logic.mk_term o Var); val conj = Logic.mk_conjunction_list (tyvars' @ tvars') |> Thm.cterm_of ctxt |> Drule.mk_term; in ((tyvars, tvars), Conjunction.intr thm conj) end; fun get_thm_insts thm = let val (thm', insts) = Conjunction.elim thm; val insts' = insts |> Drule.dest_term |> Thm.term_of |> Logic.dest_conjunction_list |> (fn f => fold (fn t => fn (tys, ts) => term_type_cases (fn T => (T :: tys, ts)) (fn t => (tys, t :: ts)) t) f ([], [])) ||> rev |>> rev; in (thm', insts') end; fun instantiate_xis ctxt insts thm = let val tyvars = Thm.fold_terms {hyps = false} Term.add_tvars thm []; val tvars = Thm.fold_terms {hyps = false} Term.add_vars thm []; fun add_inst (xi, t) (Ts, ts) = (case AList.lookup (op =) tyvars xi of SOME S => (((xi, S), Thm.ctyp_of ctxt (Logic.dest_type t)) :: Ts, ts) | NONE => (case AList.lookup (op =) tvars xi of SOME _ => (Ts, (xi, Thm.cterm_of ctxt t) :: ts) | NONE => error "indexname not found in thm")); val (instT, inst) = fold add_inst insts ([], []); in - (Thm.instantiate (instT, []) thm + (Thm.instantiate (TVars.make instT, Vars.empty) thm |> infer_instantiate ctxt inst COMP_INCR asm_rl) |> Thm.adjust_maxidx_thm ~1 |> restore_tags thm end; datatype rule_inst = Named_Insts of ((indexname * string) * (term -> unit)) list * (binding * string option * mixfix) list | Term_Insts of (indexname * term) list | Unchecked_Term_Insts of term option list * term option list; fun mk_pair (t, t') = Logic.mk_conjunction (Logic.mk_term t, Logic.mk_term t'); fun dest_pair t = apply2 Logic.dest_term (Logic.dest_conjunction t); fun embed_indexname ((xi, s), f) = let fun wrap_xi xi t = mk_pair (Var (xi, fastype_of t), t); in ((xi, s), f o wrap_xi xi) end; fun unembed_indexname t = dest_pair t |> apfst (Term.dest_Var #> fst); fun read_where_insts (insts, fixes) = let val insts' = if forall (fn (_, v) => Parse_Tools.is_real_val v) insts then Term_Insts (map (unembed_indexname o Parse_Tools.the_real_val o snd) insts) else Named_Insts (map (fn (xi, p) => embed_indexname ((xi, Parse_Tools.the_parse_val p), Parse_Tools.the_parse_fun p)) insts, fixes); in insts' end; fun of_rule thm (args, concl_args) = let fun zip_vars _ [] = [] | zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest | zip_vars ((x, _) :: xs) (SOME t :: rest) = (x, t) :: zip_vars xs rest | zip_vars [] _ = error "More instantiations than variables in theorem"; val insts = zip_vars (rev (Term.add_vars (Thm.full_prop_of thm) [])) args @ zip_vars (rev (Term.add_vars (Thm.concl_of thm) [])) concl_args; in insts end; val inst = Args.maybe Parse_Tools.name_term; val concl = Args.$$$ "concl" -- Args.colon; fun close_unchecked_insts context ((insts, concl_inst), fixes) = let val ctxt = Context.proof_of context; val ctxt1 = ctxt |> Proof_Context.add_fixes_cmd fixes |> #2; val insts' = insts @ concl_inst; val term_insts = map (the_list o (Option.map Parse_Tools.the_parse_val)) insts' |> burrow (Syntax.read_terms ctxt1 #> Variable.export_terms ctxt1 ctxt) |> map (try the_single); val _ = (insts', term_insts) |> ListPair.app (fn (SOME p, SOME t) => Parse_Tools.the_parse_fun p t | _ => ()); val (insts'', concl_insts'') = chop (length insts) term_insts; in Unchecked_Term_Insts (insts'', concl_insts'') end; fun read_of_insts checked context ((insts, concl_insts), fixes) = if forall (fn SOME t => Parse_Tools.is_real_val t | NONE => true) (insts @ concl_insts) then if checked then (fn _ => Term_Insts (map (unembed_indexname o Parse_Tools.the_real_val) (map_filter I (insts @ concl_insts)))) else (fn _ => Unchecked_Term_Insts (map (Option.map Parse_Tools.the_real_val) insts, map (Option.map Parse_Tools.the_real_val) concl_insts)) else if checked then (fn thm => Named_Insts (apply2 (map (Option.map (fn p => (Parse_Tools.the_parse_val p, Parse_Tools.the_parse_fun p)))) (insts, concl_insts) |> of_rule thm |> map ((fn (xi, (nm, f)) => embed_indexname ((xi, nm), f))), fixes)) else let val result = close_unchecked_insts context ((insts, concl_insts), fixes); in fn _ => result end; fun read_instantiate_closed ctxt (Named_Insts (insts, fixes)) thm = let val insts' = map (fn ((v, t), _) => ((v, Position.none), t)) insts; val (thm_insts, thm') = add_thm_insts ctxt thm; val (thm'', thm_insts') = Rule_Insts.where_rule ctxt insts' fixes thm' |> get_thm_insts; val tyinst = ListPair.zip (fst thm_insts, fst thm_insts') |> map (fn ((xi, _), typ) => (xi, typ)); val tinst = ListPair.zip (snd thm_insts, snd thm_insts') |> map (fn ((xi, _), t) => (xi, t)); val _ = map (fn ((xi, _), f) => (case AList.lookup (op =) tyinst xi of SOME typ => f (Logic.mk_type typ) | NONE => (case AList.lookup (op =) tinst xi of SOME t => f t | NONE => error "Lost indexname in instantiated theorem"))) insts; in (thm'' |> restore_tags thm) end | read_instantiate_closed ctxt (Unchecked_Term_Insts insts) thm = let val (xis, ts) = ListPair.unzip (of_rule thm insts); val ctxt' = Variable.declare_maxidx (Thm.maxidx_of thm) ctxt; val (ts', ctxt'') = Variable.import_terms false ts ctxt'; val ts'' = Variable.export_terms ctxt'' ctxt ts'; val insts' = ListPair.zip (xis, ts''); in instantiate_xis ctxt insts' thm end | read_instantiate_closed ctxt (Term_Insts insts) thm = instantiate_xis ctxt insts thm; val _ = Theory.setup (Attrib.setup \<^binding>\where\ (Scan.lift (Parse.and_list1 (Args.var -- (Args.$$$ "=" |-- Parse_Tools.name_term)) -- Parse.for_fixes) >> (fn args => let val args' = read_where_insts args in fn (context, thm) => (NONE, SOME (read_instantiate_closed (Context.proof_of context) args' thm)) end)) "named instantiation of theorem"); val _ = Theory.setup (Attrib.setup \<^binding>\of\ (Scan.lift (Args.mode "unchecked" -- (Scan.repeat (Scan.unless concl inst) -- Scan.optional (concl |-- Scan.repeat inst) [] -- Parse.for_fixes)) -- Scan.state >> (fn ((unchecked, args), context) => let val read_insts = read_of_insts (not unchecked) context args; in fn (context, thm) => let val thm' = if Thm.is_free_dummy thm andalso unchecked then Drule.free_dummy_thm else read_instantiate_closed (Context.proof_of context) (read_insts thm) thm in (NONE, SOME thm') end end)) "positional instantiation of theorem"); end; diff --git a/src/HOL/Eisbach/match_method.ML b/src/HOL/Eisbach/match_method.ML --- a/src/HOL/Eisbach/match_method.ML +++ b/src/HOL/Eisbach/match_method.ML @@ -1,750 +1,750 @@ (* Title: HOL/Eisbach/match_method.ML Author: Daniel Matichuk, NICTA/UNSW Setup for "match" proof method. It provides basic fact/term matching in addition to premise/conclusion matching through Subgoal.focus, and binds fact names from matches as well as term patterns within matches. *) signature MATCH_METHOD = sig val focus_schematics: Proof.context -> Envir.tenv val focus_params: Proof.context -> term list (* FIXME proper ML interface for the main thing *) end structure Match_Method : MATCH_METHOD = struct (* Filter premises by predicate, with premise order; recovers premise order afterwards.*) fun filter_prems_tac' ctxt prem = let fun Then NONE tac = SOME tac | Then (SOME tac) tac' = SOME (tac THEN' tac'); fun thins idxH (tac, n, i) = if prem idxH then (tac, n + 1 , i) else (Then tac (rotate_tac n THEN' eresolve_tac ctxt [thin_rl]), 0, i + n); in SUBGOAL (fn (goal, i) => let val idxHs = tag_list 0 (Logic.strip_assums_hyp goal) in (case fold thins idxHs (NONE, 0, 0) of (NONE, _, _) => no_tac | (SOME tac, _, n) => tac i THEN rotate_tac (~ n) i) end) end; datatype match_kind = Match_Term of term Item_Net.T | Match_Fact of thm Item_Net.T | Match_Concl | Match_Prems of bool; val aconv_net = Item_Net.init (op aconv) single; val parse_match_kind = Scan.lift \<^keyword>\conclusion\ >> K Match_Concl || Scan.lift (\<^keyword>\premises\ |-- Args.mode "local") >> Match_Prems || Scan.lift (\<^keyword>\(\) |-- Args.term --| Scan.lift (\<^keyword>\)\) >> (fn t => Match_Term (Item_Net.update t aconv_net)) || Attrib.thms >> (fn thms => Match_Fact (fold Item_Net.update thms Thm.item_net)); fun nameable_match m = (case m of Match_Fact _ => true | Match_Prems _ => true | _ => false); fun prop_match m = (case m of Match_Term _ => false | _ => true); val bound_thm : (thm, binding) Parse_Tools.parse_val parser = Parse_Tools.parse_thm_val Parse.binding; val bound_term : (term, binding) Parse_Tools.parse_val parser = Parse_Tools.parse_term_val Parse.binding; val fixes = Parse.and_list1 (Scan.repeat1 (Parse.position bound_term) -- Scan.option (\<^keyword>\::\ |-- Parse.!!! Parse.typ) >> (fn (xs, T) => map (fn (x, pos) => ((x, T), pos)) xs)) >> flat; val for_fixes = Scan.optional (\<^keyword>\for\ |-- fixes) []; fun pos_of dyn = Parse_Tools.the_parse_val dyn |> Binding.pos_of; (*FIXME: Dynamic facts modify the background theory, so we have to resort to token replacement for matched facts. *) val dynamic_fact = Scan.lift bound_thm -- Attrib.opt_attribs; type match_args = {multi : bool, cut : int}; val parse_match_args = Scan.optional (Args.parens (Parse.enum1 "," (Args.$$$ "multi" >> (fn _ => fn {cut, ...} => {multi = true, cut = cut}) || Args.$$$ "cut" |-- Scan.optional Parse.nat 1 >> (fn n => fn {multi, ...} => {multi = multi, cut = n})))) [] >> (fn fs => fold I fs {multi = false, cut = ~1}); fun parse_named_pats match_kind = Args.context -- Parse.and_list1' (Scan.option (dynamic_fact --| Scan.lift Args.colon) :-- (fn opt_dyn => if is_none opt_dyn orelse nameable_match match_kind then Scan.lift (Parse_Tools.name_term -- parse_match_args) else let val b = #1 (the opt_dyn) in error ("Cannot bind fact name in term match" ^ Position.here (pos_of b)) end)) -- Scan.lift (for_fixes -- (\<^keyword>\\\ |-- Parse.token Parse.text)) >> (fn ((ctxt, ts), (fixes, body)) => (case Token.get_value body of SOME (Token.Source src) => let val text = Method.read ctxt src; val ts' = map (fn (b, (Parse_Tools.Real_Val v, match_args)) => ((Option.map (fn (b, att) => (Parse_Tools.the_real_val b, att)) b, match_args), v) | _ => raise Fail "Expected closed term") ts val fixes' = map (fn ((p, _), _) => Parse_Tools.the_real_val p) fixes in (ts', fixes', text) end | _ => let val (fix_names, ctxt3) = ctxt |> Proof_Context.add_fixes_cmd (map (fn ((pb, otyp), _) => (Parse_Tools.the_parse_val pb, otyp, NoSyn)) fixes) ||> Proof_Context.set_mode Proof_Context.mode_schematic; fun parse_term term = if prop_match match_kind then Syntax.parse_prop ctxt3 term else Syntax.parse_term ctxt3 term; fun drop_judgment_dummy t = (case t of Const (judgment, _) $ (Const (\<^syntax_const>\_type_constraint_\, T) $ Const (\<^const_name>\Pure.dummy_pattern\, _)) => if judgment = Object_Logic.judgment_name ctxt then Const (\<^syntax_const>\_type_constraint_\, T) $ Const (\<^const_name>\Pure.dummy_pattern\, propT) else t | t1 $ t2 => drop_judgment_dummy t1 $ drop_judgment_dummy t2 | Abs (a, T, b) => Abs (a, T, drop_judgment_dummy b) | _ => t); val pats = map (fn (_, (term, _)) => parse_term (Parse_Tools.the_parse_val term)) ts |> map drop_judgment_dummy |> (fn ts => fold_map Term.replace_dummy_patterns ts (Variable.maxidx_of ctxt3 + 1)) |> fst |> Syntax.check_terms ctxt3; val pat_fixes = fold (Term.add_frees) pats [] |> map fst; val _ = map2 (fn nm => fn (_, pos) => member (op =) pat_fixes nm orelse error ("For-fixed variable must be bound in some pattern" ^ Position.here pos)) fix_names fixes; val _ = map (Term.map_types Type.no_tvars) pats; val ctxt4 = fold Variable.declare_term pats ctxt3; val (real_fixes, ctxt5) = ctxt4 |> fold_map Proof_Context.inferred_param fix_names |>> map Free; fun reject_extra_free (Free (x, _)) () = if Variable.is_fixed ctxt5 x then () else error ("Illegal use of free (unfixed) variable " ^ quote x) | reject_extra_free _ () = (); val _ = (fold o fold_aterms) reject_extra_free pats (); val binds = map (fn (b, _) => Option.map (fn (b, att) => (Parse_Tools.the_parse_val b, att)) b) ts; fun upd_ctxt (SOME (b, att)) pat (tms, ctxt) = let val ([nm], ctxt') = Variable.variant_fixes [Name.internal (Binding.name_of b)] ctxt; val abs_nms = Term.strip_all_vars pat; val param_thm = map (Drule.mk_term o Thm.cterm_of ctxt' o Free) abs_nms |> Conjunction.intr_balanced |> Drule.generalize (Names.empty, Names.make_set (map fst abs_nms)) |> Thm.tag_free_dummy; val atts = map (Attrib.attribute ctxt') att; val (param_thm', ctxt'') = Thm.proof_attributes atts param_thm ctxt'; fun label_thm thm = Thm.cterm_of ctxt'' (Free (nm, propT)) |> Drule.mk_term |> not (null abs_nms) ? Conjunction.intr thm val [head_thm, body_thm] = Drule.zero_var_indexes_list (map label_thm [param_thm, param_thm']) |> map Thm.tag_free_dummy; val ctxt''' = Attrib.local_notes "" [((b, []), [([body_thm], [])])] ctxt'' |> snd |> Variable.declare_maxidx (Thm.maxidx_of head_thm); in (SOME (head_thm, att) :: tms, ctxt''') end | upd_ctxt NONE _ (tms, ctxt) = (NONE :: tms, ctxt); val (binds, ctxt6) = ctxt5 |> (fn ctxt => fold2 upd_ctxt binds pats ([], ctxt) |> apfst rev) ||> Proof_Context.restore_mode ctxt; val (text, src) = Method.read_closure_input ctxt6 (Token.input_of body); val morphism = Variable.export_morphism ctxt6 (ctxt |> fold Token.declare_maxidx src |> Variable.declare_maxidx (Variable.maxidx_of ctxt6)); val pats' = map (Term.map_types Type_Infer.paramify_vars #> Morphism.term morphism) pats; val _ = ListPair.app (fn ((_, (v, _)), t) => Parse_Tools.the_parse_fun v t) (ts, pats'); fun close_src src = let val src' = src |> map (Token.closure #> Token.transform morphism); val _ = (Token.args_of_src src ~~ Token.args_of_src src') |> List.app (fn (tok, tok') => (case Token.get_value tok' of SOME value => ignore (Token.assign (SOME value) tok) | NONE => ())); in src' end; val binds' = map (Option.map (fn (t, atts) => (Morphism.thm morphism t, map close_src atts))) binds; val _ = ListPair.app (fn ((SOME ((v, _)), _), SOME (t, _)) => Parse_Tools.the_parse_fun v t | ((NONE, _), NONE) => () | _ => error "Mismatch between real and parsed bound variables") (ts, binds'); val real_fixes' = map (Morphism.term morphism) real_fixes; val _ = ListPair.app (fn (((v, _) , _), t) => Parse_Tools.the_parse_fun v t) (fixes, real_fixes'); val match_args = map (fn (_, (_, match_args)) => match_args) ts; val binds'' = (binds' ~~ match_args) ~~ pats'; val src' = map (Token.transform morphism) src; val _ = Token.assign (SOME (Token.Source src')) body; in (binds'', real_fixes', text) end)); fun dest_internal_term t = (case try Logic.dest_conjunction t of SOME (params, head) => (params |> Logic.dest_conjunctions |> map Logic.dest_term, head |> Logic.dest_term) | NONE => ([], t |> Logic.dest_term)); fun dest_internal_fact thm = dest_internal_term (Thm.prop_of thm); fun inst_thm ctxt env ts params thm = let val ts' = map (Envir.norm_term env) ts; val insts = map (#1 o dest_Var) ts' ~~ map (Thm.cterm_of ctxt) params; in infer_instantiate ctxt insts thm end; fun do_inst fact_insts' env text ctxt = let val fact_insts = map_filter (fn ((((SOME ((_, head), att), _), _), _), thms) => SOME (head, (thms, att)) | _ => NONE) fact_insts'; fun try_dest_term thm = \<^try>\#2 (dest_internal_fact thm)\; fun expand_fact fact_insts thm = the_default [thm] (case try_dest_term thm of SOME t_ident => AList.lookup (op aconv) fact_insts t_ident | NONE => NONE); fun fact_morphism fact_insts = Morphism.term_morphism "do_inst.term" (Envir.norm_term env) $> Morphism.typ_morphism "do_inst.type" (Envir.norm_type (Envir.type_env env)) $> Morphism.fact_morphism "do_inst.fact" (maps (expand_fact fact_insts)); fun apply_attribute (head, (fact, atts)) (fact_insts, ctxt) = let val morphism = fact_morphism fact_insts; val atts' = map (Attrib.attribute ctxt o map (Token.transform morphism)) atts; val (fact'', ctxt') = fold_map (Thm.proof_attributes atts') fact ctxt; in ((head, fact'') :: fact_insts, ctxt') end; (*TODO: What to do about attributes that raise errors?*) val (fact_insts', ctxt') = fold_rev (apply_attribute) fact_insts ([], ctxt); val text' = (Method.map_source o map) (Token.transform (fact_morphism fact_insts')) text; in (text', ctxt') end; fun prep_fact_pat ((x, args), pat) ctxt = let val ((params, pat'), ctxt') = Variable.focus NONE pat ctxt; val params' = map (Free o snd) params; val morphism = Variable.export_morphism ctxt' (ctxt |> Variable.declare_maxidx (Variable.maxidx_of ctxt')); val pat'' :: params'' = map (Morphism.term morphism) (pat' :: params'); fun prep_head (t, att) = (dest_internal_fact t, att); in ((((Option.map prep_head x, args), params''), pat''), ctxt') end; fun morphism_env morphism env = let val tenv = Envir.term_env env |> Vartab.map (K (fn (T, t) => (Morphism.typ morphism T, Morphism.term morphism t))); val tyenv = Envir.type_env env |> Vartab.map (K (fn (S, T) => (S, Morphism.typ morphism T))); in Envir.Envir {maxidx = Envir.maxidx_of env, tenv = tenv, tyenv = tyenv} end; fun export_with_params ctxt morphism (SOME ts, params) thm env = let val outer_env = morphism_env morphism env; val thm' = Morphism.thm morphism thm; in inst_thm ctxt outer_env params ts thm' end | export_with_params _ morphism (NONE, _) thm _ = Morphism.thm morphism thm; fun match_filter_env is_newly_fixed pat_vars fixes params env = let val param_vars = map Term.dest_Var params; val tenv = Envir.term_env env; val params' = map (fn (xi, _) => Vartab.lookup tenv xi) param_vars; val fixes_vars = map Term.dest_Var fixes; val all_vars = Vartab.keys tenv; val extra_vars = subtract (fn ((xi, _), xi') => xi = xi') fixes_vars all_vars; val tenv' = tenv |> fold (Vartab.delete_safe) extra_vars; val env' = Envir.Envir {maxidx = Envir.maxidx_of env, tenv = tenv', tyenv = Envir.type_env env}; val all_params_bound = forall (fn SOME (_, Free (x, _)) => is_newly_fixed x | _ => false) params'; val all_params_distinct = not (has_duplicates (eq_option (eq_pair (op =) (op aconv))) params'); val pat_fixes = inter (eq_fst (op =)) fixes_vars pat_vars; val all_pat_fixes_bound = forall (fn (xi, _) => is_some (Vartab.lookup tenv' xi)) pat_fixes; in if all_params_bound andalso all_pat_fixes_bound andalso all_params_distinct then SOME env' else NONE end; (* Slightly hacky way of uniquely identifying focus premises *) val prem_idN = "premise_id"; fun prem_id_eq ((id, _ : thm), (id', _ : thm)) = id = id'; val prem_rules : (int * thm) Item_Net.T = Item_Net.init prem_id_eq (single o Thm.full_prop_of o snd); fun raw_thm_to_id thm = (case Properties.get (Thm.get_tags thm) prem_idN of NONE => NONE | SOME id => Int.fromString id) |> the_default ~1; structure Focus_Data = Proof_Data ( type T = (int * (int * thm) Item_Net.T) * (*prems*) Envir.tenv * (*schematics*) term list (*params*) fun init _ : T = ((0, prem_rules), Vartab.empty, []) ); (* focus prems *) val focus_prems = #1 o Focus_Data.get; fun transfer_focus_prems from_ctxt = Focus_Data.map (@{apply 3(1)} (K (focus_prems from_ctxt))) fun add_focus_prem prem = `(Focus_Data.get #> #1 #> #1) ##> (Focus_Data.map o @{apply 3(1)}) (fn (next, net) => (next + 1, Item_Net.update (next, Thm.tag_rule (prem_idN, string_of_int next) prem) net)); fun remove_focus_prem' (ident, thm) = (Focus_Data.map o @{apply 3(1)} o apsnd) (Item_Net.remove (ident, thm)); fun remove_focus_prem thm = remove_focus_prem' (raw_thm_to_id thm, thm); (*TODO: Preliminary analysis to see if we're trying to clear in a non-focus match?*) val _ = Theory.setup (Attrib.setup \<^binding>\thin\ (Scan.succeed (Thm.declaration_attribute (fn th => Context.mapping I (remove_focus_prem th)))) "clear premise inside match method"); (* focus schematics *) val focus_schematics = #2 o Focus_Data.get; fun add_focus_schematics schematics = (Focus_Data.map o @{apply 3(2)}) (fold (fn ((xi, T), ct) => Vartab.update_new (xi, (T, Thm.term_of ct))) schematics); (* focus params *) val focus_params = #3 o Focus_Data.get; fun add_focus_params params = (Focus_Data.map o @{apply 3(3)}) (append (map (fn (_, ct) => Thm.term_of ct) params)); (* Add focus elements as proof data *) fun augment_focus (focus: Subgoal.focus) : (int list * Subgoal.focus) = let val {context, params, prems, asms, concl, schematics} = focus; val (prem_ids, ctxt') = context |> add_focus_params params - |> add_focus_schematics (snd schematics) + |> add_focus_schematics (Vars.dest (snd schematics)) |> fold_map add_focus_prem (rev prems) in (prem_ids, {context = ctxt', params = params, prems = prems, concl = concl, schematics = schematics, asms = asms}) end; (* Fix schematics in the goal *) fun focus_concl ctxt i bindings goal = let val ({context = ctxt', concl, params, prems, asms, schematics}, goal') = Subgoal.focus_params ctxt i bindings goal; val ((_, inst), ctxt'') = Variable.import_inst true [Thm.term_of concl] ctxt'; val schematic_terms = Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt'' t)) inst []; - val goal'' = Thm.instantiate ([], schematic_terms) goal'; - val concl' = Thm.instantiate_cterm ([], schematic_terms) concl; + val goal'' = Thm.instantiate (TVars.empty, Vars.make schematic_terms) goal'; + val concl' = Thm.instantiate_cterm (TVars.empty, Vars.make schematic_terms) concl; val (schematic_types, schematic_terms') = schematics; - val schematics' = (schematic_types, schematic_terms @ schematic_terms'); + val schematics' = (schematic_types, fold Vars.add schematic_terms schematic_terms'); in ({context = ctxt'', concl = concl', params = params, prems = prems, schematics = schematics', asms = asms} : Subgoal.focus, goal'') end; fun deduplicate eq prev seq = Seq.make (fn () => (case Seq.pull seq of SOME (x, seq') => if member eq prev x then Seq.pull (deduplicate eq prev seq') else SOME (x, deduplicate eq (x :: prev) seq') | NONE => NONE)); fun consistent_env env = let val tenv = Envir.term_env env; val tyenv = Envir.type_env env; in forall (fn (_, (T, t)) => Envir.norm_type tyenv T = fastype_of t) (Vartab.dest tenv) end; fun term_eq_wrt (env1, env2) (t1, t2) = Envir.eta_contract (Envir.norm_term env1 t1) aconv Envir.eta_contract (Envir.norm_term env2 t2); fun type_eq_wrt (env1, env2) (T1, T2) = Envir.norm_type (Envir.type_env env1) T1 = Envir.norm_type (Envir.type_env env2) T2; fun eq_env (env1, env2) = Envir.maxidx_of env1 = Envir.maxidx_of env1 andalso ListPair.allEq (fn ((var, (_, t)), (var', (_, t'))) => (var = var' andalso term_eq_wrt (env1, env2) (t, t'))) (apply2 Vartab.dest (Envir.term_env env1, Envir.term_env env2)) andalso ListPair.allEq (fn ((var, (_, T)), (var', (_, T'))) => var = var' andalso type_eq_wrt (env1, env2) (T, T')) (apply2 Vartab.dest (Envir.type_env env1, Envir.type_env env2)); fun merge_env (env1, env2) = let val tenv = Vartab.merge (eq_snd (term_eq_wrt (env1, env2))) (Envir.term_env env1, Envir.term_env env2); val tyenv = Vartab.merge (eq_snd (type_eq_wrt (env1, env2)) andf eq_fst (op =)) (Envir.type_env env1, Envir.type_env env2); val maxidx = Int.max (Envir.maxidx_of env1, Envir.maxidx_of env2); in Envir.Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv} end; fun import_with_tags thms ctxt = let val ((_, thms'), ctxt') = Variable.import false thms ctxt; val thms'' = map2 (fn thm => Thm.map_tags (K (Thm.get_tags thm))) thms thms'; in (thms'', ctxt') end; fun try_merge (env, env') = SOME (merge_env (env, env')) handle Vartab.DUP _ => NONE fun Seq_retrieve seq f = let fun retrieve' (list, seq) f = (case Seq.pull seq of SOME (x, seq') => if f x then (SOME x, (list, seq')) else retrieve' (list @ [x], seq') f | NONE => (NONE, (list, seq))); val (result, (list, seq)) = retrieve' ([], seq) f; in (result, Seq.append (Seq.of_list list) seq) end; fun match_facts ctxt fixes prop_pats get = let fun is_multi (((_, x : match_args), _), _) = #multi x; fun get_cut (((_, x : match_args), _), _) = #cut x; fun do_cut n = if n = ~1 then I else Seq.take n; val raw_thmss = map (get o snd) prop_pats; val (thmss, ctxt') = fold_burrow import_with_tags raw_thmss ctxt; val newly_fixed = Variable.is_newly_fixed ctxt' ctxt; val morphism = Variable.export_morphism ctxt' ctxt; fun match_thm (((x, params), pat), thm) = let val pat_vars = Term.add_vars pat []; val ts = Option.map (fst o fst) (fst x); val item' = Thm.prop_of thm; val matches = (Unify.matchers (Context.Proof ctxt) [(pat, item')]) |> Seq.filter consistent_env |> Seq.map_filter (fn env' => (case match_filter_env newly_fixed pat_vars fixes params env' of SOME env'' => SOME (export_with_params ctxt morphism (ts, params) thm env', env'') | NONE => NONE)) |> Seq.map (apfst (Thm.map_tags (K (Thm.get_tags thm)))) |> deduplicate (eq_pair Thm.eq_thm_prop eq_env) [] in matches end; val all_matches = map2 pair prop_pats thmss |> map (fn (pat, matches) => (pat, map (fn thm => match_thm (pat, thm)) matches)); fun proc_multi_match (pat, thmenvs) (pats, env) = do_cut (get_cut pat) (if is_multi pat then let fun maximal_set tail seq envthms = Seq.make (fn () => (case Seq.pull seq of SOME ((thm, env'), seq') => let val (result, envthms') = Seq_retrieve envthms (fn (env, _) => eq_env (env, env')); in (case result of SOME (_, thms) => SOME ((env', thm :: thms), maximal_set tail seq' envthms') | NONE => Seq.pull (maximal_set (tail @ [(env', [thm])]) seq' envthms')) end | NONE => Seq.pull (Seq.append envthms (Seq.of_list tail)))); val maximal_sets = fold (maximal_set []) thmenvs Seq.empty; in maximal_sets |> Seq.map swap |> Seq.filter (fn (thms, _) => not (null thms)) |> Seq.map_filter (fn (thms, env') => (case try_merge (env, env') of SOME env'' => SOME ((pat, thms) :: pats, env'') | NONE => NONE)) end else let fun just_one (thm, env') = (case try_merge (env, env') of SOME env'' => SOME ((pat, [thm]) :: pats, env'') | NONE => NONE); in fold (fn seq => Seq.append (Seq.map_filter just_one seq)) thmenvs Seq.empty end); val all_matches = Seq.EVERY (map proc_multi_match all_matches) ([], Envir.init); in all_matches |> Seq.map (apsnd (morphism_env morphism)) end; fun real_match using outer_ctxt fixes m text pats st = let val goal_ctxt = fold Variable.declare_term fixes outer_ctxt (*FIXME Is this a good idea? We really only care about the maxidx*) |> fold (fn (_, t) => Variable.declare_term t) pats; fun make_fact_matches ctxt get = let val (pats', ctxt') = fold_map prep_fact_pat pats ctxt; in match_facts ctxt' fixes pats' get |> Seq.map (fn (fact_insts, env) => do_inst fact_insts env text ctxt') end; fun make_term_matches ctxt get = let val pats' = map (fn ((SOME _, _), _) => error "Cannot name term match" | ((_, x), t) => (((NONE, x), []), Logic.mk_term t)) pats; val thm_of = Drule.mk_term o Thm.cterm_of ctxt; fun get' t = get (Logic.dest_term t) |> map thm_of; in match_facts ctxt fixes pats' get' |> Seq.map (fn (fact_insts, env) => do_inst fact_insts env text ctxt) end; in (case m of Match_Fact net => make_fact_matches goal_ctxt (Item_Net.retrieve net) |> Seq.map (fn (text, ctxt') => Method.evaluate_runtime text ctxt' using (ctxt', st) |> Seq.filter_results |> Seq.map (fn (_, thm) => (outer_ctxt, thm))) | Match_Term net => make_term_matches goal_ctxt (Item_Net.retrieve net) |> Seq.map (fn (text, ctxt') => Method.evaluate_runtime text ctxt' using (ctxt', st) |> Seq.filter_results |> Seq.map (fn (_, thm) => (outer_ctxt, thm))) | match_kind => if Thm.no_prems st then Seq.empty else let fun focus_cases f g = (case match_kind of Match_Prems b => f b | Match_Concl => g | _ => raise Fail "Match kind fell through"); val ((local_premids, {context = focus_ctxt, params, asms, concl, prems, ...}), focused_goal) = focus_cases (K Subgoal.focus_prems) focus_concl goal_ctxt 1 NONE st |>> augment_focus; val texts = focus_cases (fn is_local => fn _ => make_fact_matches focus_ctxt (Item_Net.retrieve (focus_prems focus_ctxt |> snd) #> is_local ? filter (fn (p, _) => exists (fn id' => id' = p) local_premids) #> order_list)) (fn _ => make_term_matches focus_ctxt (fn _ => [Logic.strip_imp_concl (Thm.term_of concl)])) (); (*TODO: How to handle cases? *) fun do_retrofit (inner_ctxt, st1) = let val (_, before_prems) = focus_prems focus_ctxt; val (_, after_prems) = focus_prems inner_ctxt; val removed_prems = Item_Net.filter (null o Item_Net.lookup after_prems) before_prems val removed_local_prems = Item_Net.content removed_prems |> filter (fn (id, _) => member (op =) local_premids id) |> map (fn (_, prem) => Thm.prop_of prem) fun filter_removed_prems prems = Item_Net.filter (null o Item_Net.lookup removed_prems) prems; val outer_ctxt' = outer_ctxt |> Focus_Data.map (@{apply 3(1)} (apsnd filter_removed_prems)); val n_subgoals = Thm.nprems_of st1; val removed_prem_idxs = prems |> tag_list 0 |> filter (member (op aconv) removed_local_prems o Thm.prop_of o snd) |> map fst fun filter_prem (i, _) = not (member (op =) removed_prem_idxs i); in Subgoal.retrofit inner_ctxt goal_ctxt params asms 1 st1 st |> focus_cases (fn _ => (n_subgoals > 0 andalso length removed_local_prems > 0) ? (Seq.map (Goal.restrict 1 n_subgoals) #> Seq.maps (ALLGOALS (fn i => DETERM (filter_prems_tac' goal_ctxt filter_prem i))) #> Seq.map (Goal.unrestrict 1))) I |> Seq.map (pair outer_ctxt') end; fun apply_text (text, ctxt') = Method.evaluate_runtime text ctxt' using (ctxt', focused_goal) |> Seq.filter_results |> Seq.maps (Seq.DETERM do_retrofit) in Seq.map apply_text texts end) end; val _ = Theory.setup (Method.setup \<^binding>\match\ (parse_match_kind :-- (fn kind => Scan.lift \<^keyword>\in\ |-- Parse.enum1' "\" (parse_named_pats kind)) >> (fn (matches, bodies) => fn ctxt => CONTEXT_METHOD (fn using => Method.RUNTIME (fn (goal_ctxt, st) => let val ctxt' = transfer_focus_prems goal_ctxt ctxt; fun exec (pats, fixes, text) st' = real_match using ctxt' fixes matches text pats st'; in Seq.flat (Seq.FIRST (map exec bodies) st) |> Seq.map (apfst (fn ctxt' => transfer_focus_prems ctxt' goal_ctxt)) |> Seq.make_results end)))) "structural analysis/matching on goals"); end; diff --git a/src/HOL/Import/import_rule.ML b/src/HOL/Import/import_rule.ML --- a/src/HOL/Import/import_rule.ML +++ b/src/HOL/Import/import_rule.ML @@ -1,451 +1,451 @@ (* Title: HOL/Import/import_rule.ML Author: Cezary Kaliszyk, University of Innsbruck Author: Alexander Krauss, QAware GmbH Importer proof rules and processing of lines and files. Based on earlier code by Steven Obua and Sebastian Skalberg. *) signature IMPORT_RULE = sig val beta : cterm -> thm val eq_mp : thm -> thm -> thm val comb : thm -> thm -> thm val trans : thm -> thm -> thm val deduct : thm -> thm -> thm val conj1 : thm -> thm val conj2 : thm -> thm val refl : cterm -> thm val abs : cterm -> thm -> thm val mdef : string -> theory -> thm val def : string -> cterm -> theory -> thm * theory val mtydef : string -> theory -> thm val tydef : string -> string -> string -> cterm -> cterm -> thm -> theory -> thm * theory val inst_type : (ctyp * ctyp) list -> thm -> theory -> thm val inst : (cterm * cterm) list -> thm -> thm type state val init_state : state val process_line : string -> (theory * state) -> (theory * state) val process_file : Path.T -> theory -> theory end structure Import_Rule: IMPORT_RULE = struct val init_state = ((Inttab.empty, 0), (Inttab.empty, 0), (Inttab.empty, 0)) type state = (ctyp Inttab.table * int) * (cterm Inttab.table * int) * (thm Inttab.table * int) fun implies_elim_all th = implies_elim_list th (map Thm.assume (cprems_of th)) fun meta_mp th1 th2 = let val th1a = implies_elim_all th1 val th1b = Thm.implies_intr (strip_imp_concl (Thm.cprop_of th2)) th1a val th2a = implies_elim_all th2 val th3 = Thm.implies_elim th1b th2a in implies_intr_hyps th3 end fun meta_eq_to_obj_eq th = let val (tml, tmr) = Thm.dest_binop (strip_imp_concl (Thm.cprop_of th)) val cty = Thm.ctyp_of_cterm tml val i = Thm.instantiate' [SOME cty] [SOME tml, SOME tmr] @{thm meta_eq_to_obj_eq} in Thm.implies_elim i th end fun beta ct = meta_eq_to_obj_eq (Thm.beta_conversion false ct) fun eq_mp th1 th2 = let val (tm1l, tm1r) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1))) val i1 = Thm.instantiate' [] [SOME tm1l, SOME tm1r] @{thm iffD1} val i2 = meta_mp i1 th1 in meta_mp i2 th2 end fun comb th1 th2 = let val t1c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1)) val t2c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th2)) val (cf, cg) = Thm.dest_binop t1c val (cx, cy) = Thm.dest_binop t2c val [fd, fr] = Thm.dest_ctyp (Thm.ctyp_of_cterm cf) val i1 = Thm.instantiate' [SOME fd, SOME fr] [SOME cf, SOME cg, SOME cx, SOME cy] @{thm cong} val i2 = meta_mp i1 th1 in meta_mp i2 th2 end fun trans th1 th2 = let val t1c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1)) val t2c = Thm.dest_arg (strip_imp_concl (Thm.cprop_of th2)) val (r, s) = Thm.dest_binop t1c val (_, t) = Thm.dest_binop t2c val ty = Thm.ctyp_of_cterm r val i1 = Thm.instantiate' [SOME ty] [SOME r, SOME s, SOME t] @{thm trans} val i2 = meta_mp i1 th1 in meta_mp i2 th2 end fun deduct th1 th2 = let val th1c = strip_imp_concl (Thm.cprop_of th1) val th2c = strip_imp_concl (Thm.cprop_of th2) val th1a = implies_elim_all th1 val th2a = implies_elim_all th2 val th1b = Thm.implies_intr th2c th1a val th2b = Thm.implies_intr th1c th2a val i = Thm.instantiate' [] [SOME (Thm.dest_arg th1c), SOME (Thm.dest_arg th2c)] @{thm iffI} val i1 = Thm.implies_elim i (Thm.assume (Thm.cprop_of th2b)) val i2 = Thm.implies_elim i1 th1b val i3 = Thm.implies_intr (Thm.cprop_of th2b) i2 val i4 = Thm.implies_elim i3 th2b in implies_intr_hyps i4 end fun conj1 th = let val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th))) val i = Thm.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct1} in meta_mp i th end fun conj2 th = let val (tml, tmr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th))) val i = Thm.instantiate' [] [SOME tml, SOME tmr] @{thm conjunct2} in meta_mp i th end fun refl ctm = let val cty = Thm.ctyp_of_cterm ctm in Thm.instantiate' [SOME cty] [SOME ctm] @{thm refl} end fun abs cv th = let val th1 = implies_elim_all th val (tl, tr) = Thm.dest_binop (Thm.dest_arg (strip_imp_concl (Thm.cprop_of th1))) val (ll, lr) = (Thm.lambda cv tl, Thm.lambda cv tr) val (al, ar) = (Thm.apply ll cv, Thm.apply lr cv) val bl = beta al val br = meta_eq_to_obj_eq (Thm.symmetric (Thm.beta_conversion false ar)) val th2 = trans (trans bl th1) br val th3 = implies_elim_all th2 val th4 = Thm.forall_intr cv th3 val i = Thm.instantiate' [SOME (Thm.ctyp_of_cterm cv), SOME (Thm.ctyp_of_cterm tl)] [SOME ll, SOME lr] @{thm ext2} in meta_mp i th4 end fun freezeT thy thm = let val tvars = Term.add_tvars (Thm.prop_of thm) [] val tfrees = map (fn ((t, _), s) => TFree (t, s)) tvars in - Thm.instantiate ((tvars ~~ map (Thm.global_ctyp_of thy) tfrees), []) thm + Thm.instantiate (TVars.make (tvars ~~ map (Thm.global_ctyp_of thy) tfrees), Vars.empty) thm end fun def' constname rhs thy = let val rhs = Thm.term_of rhs val typ = type_of rhs val constbinding = Binding.name constname val thy1 = Sign.add_consts [(constbinding, typ, NoSyn)] thy val eq = Logic.mk_equals (Const (Sign.full_name thy1 constbinding, typ), rhs) val (thms, thy2) = Global_Theory.add_defs false [((Binding.suffix_name "_hldef" constbinding, eq), [])] thy1 val def_thm = freezeT thy1 (hd thms) in (meta_eq_to_obj_eq def_thm, thy2) end fun mdef name thy = case Import_Data.get_const_def name thy of SOME th => th | NONE => error ("constant mapped but no definition: " ^ name) fun def constname rhs thy = case Import_Data.get_const_def constname thy of SOME _ => let val () = warning ("Const mapped but def provided: " ^ constname) in (mdef constname thy, thy) end | NONE => def' constname rhs thy fun typedef_hollight th thy = let val (th_s, cn) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)) val (th_s, abst) = Thm.dest_comb th_s val rept = Thm.dest_arg th_s val P = Thm.dest_arg cn val [nty, oty] = Thm.dest_ctyp (Thm.ctyp_of_cterm rept) in Thm.instantiate' [SOME nty, SOME oty] [SOME rept, SOME abst, SOME P, SOME (Thm.global_cterm_of thy (Free ("a", Thm.typ_of nty))), SOME (Thm.global_cterm_of thy (Free ("r", Thm.typ_of oty)))] @{thm typedef_hol2hollight} end fun tydef' tycname abs_name rep_name cP ct td_th thy = let val ctT = Thm.ctyp_of_cterm ct val nonempty = Thm.instantiate' [SOME ctT] [SOME cP, SOME ct] @{thm light_ex_imp_nonempty} val th2 = meta_mp nonempty td_th val c = case Thm.concl_of th2 of _ $ (Const(\<^const_name>\Ex\,_) $ Abs(_,_,Const(\<^const_name>\Set.member\,_) $ _ $ c)) => c | _ => error "type_introduction: bad type definition theorem" val tfrees = Term.add_tfrees c [] val tnames = sort_strings (map fst tfrees) val typedef_bindings = {Rep_name = Binding.name rep_name, Abs_name = Binding.name abs_name, type_definition_name = Binding.name ("type_definition_" ^ tycname)} val ((_, typedef_info), thy') = Named_Target.theory_map_result (apsnd o Typedef.transform_info) (Typedef.add_typedef {overloaded = false} (Binding.name tycname, map (rpair dummyS) tnames, NoSyn) c (SOME typedef_bindings) (fn ctxt => resolve_tac ctxt [th2] 1)) thy val aty = #abs_type (#1 typedef_info) val th = freezeT thy' (#type_definition (#2 typedef_info)) val (th_s, _) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)) val (th_s, abst) = Thm.dest_comb th_s val rept = Thm.dest_arg th_s val [nty, oty] = Thm.dest_ctyp (Thm.ctyp_of_cterm rept) val typedef_th = Thm.instantiate' [SOME nty, SOME oty] [SOME rept, SOME abst, SOME cP, SOME (Thm.global_cterm_of thy' (Free ("a", aty))), SOME (Thm.global_cterm_of thy' (Free ("r", Thm.typ_of ctT)))] @{thm typedef_hol2hollight} val th4 = typedef_th OF [#type_definition (#2 typedef_info)] in (th4, thy') end fun mtydef name thy = case Import_Data.get_typ_def name thy of SOME thn => meta_mp (typedef_hollight thn thy) thn | NONE => error ("type mapped but no tydef thm registered: " ^ name) fun tydef tycname abs_name rep_name P t td_th thy = case Import_Data.get_typ_def tycname thy of SOME _ => let val () = warning ("Type mapped but proofs provided: " ^ tycname) in (mtydef tycname thy, thy) end | NONE => tydef' tycname abs_name rep_name P t td_th thy fun inst_type lambda th thy = let fun assoc _ [] = error "assoc" | assoc x ((x',y)::rest) = if x = x' then y else assoc x rest val lambda = map (fn (a, b) => (Thm.typ_of a, b)) lambda val tys_before = Term.add_tfrees (Thm.prop_of th) [] val th1 = Thm.varifyT_global th val tys_after = Term.add_tvars (Thm.prop_of th1) [] val tyinst = map2 (fn bef => fn iS => (case try (assoc (TFree bef)) lambda of SOME cty => (iS, cty) | NONE => (iS, Thm.global_ctyp_of thy (TFree bef)))) tys_before tys_after in - Thm.instantiate (tyinst,[]) th1 + Thm.instantiate (TVars.make tyinst, Vars.empty) th1 end fun inst sigma th = let val (dom, rng) = ListPair.unzip (rev sigma) in th |> forall_intr_list dom |> forall_elim_list rng end fun transl_dotc #"." = "dot" | transl_dotc c = Char.toString c val transl_dot = String.translate transl_dotc fun transl_qmc #"?" = "t" | transl_qmc c = Char.toString c val transl_qm = String.translate transl_qmc fun getconstname s thy = case Import_Data.get_const_map s thy of SOME s => s | NONE => Sign.full_name thy (Binding.name (transl_dot s)) fun gettyname s thy = case Import_Data.get_typ_map s thy of SOME s => s | NONE => Sign.full_name thy (Binding.name s) fun get (map, no) s = case Int.fromString s of NONE => error "Import_Rule.get: not a number" | SOME i => (case Inttab.lookup map (Int.abs i) of NONE => error "Import_Rule.get: lookup failed" | SOME res => (res, (if i < 0 then Inttab.delete (Int.abs i) map else map, no))) fun getty i (thy, (tyi, tmi, thi)) = let val (i, tyi) = (get tyi i) in (i, (thy, (tyi, tmi, thi))) end fun gettm i (thy, (tyi, tmi, thi)) = let val (i, tmi) = (get tmi i) in (i, (thy, (tyi, tmi, thi))) end fun getth i (thy, (tyi, tmi, thi)) = let val (i, thi) = (get thi i) in (i, (thy, (tyi, tmi, thi))) end fun set (map, no) v = (Inttab.update_new (no + 1, v) map, no + 1) fun setty v (thy, (tyi, tmi, thi)) = (thy, (set tyi v, tmi, thi)) fun settm v (thy, (tyi, tmi, thi)) = (thy, (tyi, set tmi v, thi)) fun setth v (thy, (tyi, tmi, thi)) = (thy, (tyi, tmi, set thi v)) fun last_thm (_, _, (map, no)) = case Inttab.lookup map no of NONE => error "Import_Rule.last_thm: lookup failed" | SOME thm => thm fun listLast (h1 :: (h2 :: t)) = apfst (fn t => h1 :: h2 :: t) (listLast t) | listLast [p] = ([], p) | listLast [] = error "listLast: empty" fun pairList (h1 :: (h2 :: t)) = ((h1, h2) :: pairList t) | pairList [] = [] | pairList _ = error "pairList: odd list length" fun store_thm binding thm thy = let val ctxt = Proof_Context.init_global thy val thm = Drule.export_without_context_open thm val tvs = Term.add_tvars (Thm.prop_of thm) [] val tns = map (fn (_, _) => "'") tvs val nms = fst (fold_map Name.variant tns (Variable.names_of ctxt)) val vs = map TVar ((nms ~~ (map (snd o fst) tvs)) ~~ (map snd tvs)) - val thm' = Thm.instantiate ((tvs ~~ map (Thm.ctyp_of ctxt) vs), []) thm + val thm' = Thm.instantiate (TVars.make (tvs ~~ map (Thm.ctyp_of ctxt) vs), Vars.empty) thm in snd (Global_Theory.add_thm ((binding, thm'), []) thy) end fun log_timestamp () = let val time = Time.now () val millis = nth (space_explode "." (Time.fmt 3 time)) 1 in Date.fmt "%d.%m.%Y %H:%M:%S." (Date.fromTimeLocal time) ^ millis end fun process_line str tstate = let fun process tstate (#"R", [t]) = gettm t tstate |>> refl |-> setth | process tstate (#"B", [t]) = gettm t tstate |>> beta |-> setth | process tstate (#"1", [th]) = getth th tstate |>> conj1 |-> setth | process tstate (#"2", [th]) = getth th tstate |>> conj2 |-> setth | process tstate (#"H", [t]) = gettm t tstate |>> Thm.apply \<^cterm>\Trueprop\ |>> Thm.trivial |-> setth | process tstate (#"A", [_, t]) = gettm t tstate |>> Thm.apply \<^cterm>\Trueprop\ |>> Skip_Proof.make_thm_cterm |-> setth | process tstate (#"C", [th1, th2]) = getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => comb t1 t2) |-> setth | process tstate (#"T", [th1, th2]) = getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => trans t1 t2) |-> setth | process tstate (#"E", [th1, th2]) = getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => eq_mp t1 t2) |-> setth | process tstate (#"D", [th1, th2]) = getth th1 tstate ||>> getth th2 |>> (fn (t1, t2) => deduct t1 t2) |-> setth | process tstate (#"L", [t, th]) = gettm t tstate ||>> (fn ti => getth th ti) |>> (fn (tm, th) => abs tm th) |-> setth | process (thy, state) (#"M", [s]) = let val ctxt = Proof_Context.init_global thy val thm = freezeT thy (Global_Theory.get_thm thy s) val ((_, [th']), _) = Variable.import true [thm] ctxt in setth th' (thy, state) end | process (thy, state) (#"Q", l) = let val (tys, th) = listLast l val (th, tstate) = getth th (thy, state) val (tys, tstate) = fold_map getty tys tstate in setth (inst_type (pairList tys) th thy) tstate end | process tstate (#"S", l) = let val (tms, th) = listLast l val (th, tstate) = getth th tstate val (tms, tstate) = fold_map gettm tms tstate in setth (inst (pairList tms) th) tstate end | process tstate (#"F", [name, t]) = let val (tm, (thy, state)) = gettm t tstate val (th, thy) = def (transl_dot name) tm thy in setth th (thy, state) end | process (thy, state) (#"F", [name]) = setth (mdef name thy) (thy, state) | process tstate (#"Y", [name, absname, repname, t1, t2, th]) = let val (th, tstate) = getth th tstate val (t1, tstate) = gettm t1 tstate val (t2, (thy, state)) = gettm t2 tstate val (th, thy) = tydef name absname repname t1 t2 th thy in setth th (thy, state) end | process (thy, state) (#"Y", [name, _, _]) = setth (mtydef name thy) (thy, state) | process (thy, state) (#"t", [n]) = setty (Thm.global_ctyp_of thy (TFree ("'" ^ (transl_qm n), \<^sort>\type\))) (thy, state) | process (thy, state) (#"a", n :: l) = fold_map getty l (thy, state) |>> (fn tys => Thm.global_ctyp_of thy (Type (gettyname n thy, map Thm.typ_of tys))) |-> setty | process (thy, state) (#"v", [n, ty]) = getty ty (thy, state) |>> (fn ty => Thm.global_cterm_of thy (Free (transl_dot n, Thm.typ_of ty))) |-> settm | process (thy, state) (#"c", [n, ty]) = getty ty (thy, state) |>> (fn ty => Thm.global_cterm_of thy (Const (getconstname n thy, Thm.typ_of ty))) |-> settm | process tstate (#"f", [t1, t2]) = gettm t1 tstate ||>> gettm t2 |>> (fn (t1, t2) => Thm.apply t1 t2) |-> settm | process tstate (#"l", [t1, t2]) = gettm t1 tstate ||>> gettm t2 |>> (fn (t1, t2) => Thm.lambda t1 t2) |-> settm | process (thy, state) (#"+", [s]) = (store_thm (Binding.name (transl_dot s)) (last_thm state) thy, state) | process _ (c, _) = error ("process: unknown command: " ^ String.implode [c]) fun parse_line s = case String.tokens (fn x => (x = #"\n" orelse x = #" ")) s of [] => error "parse_line: empty" | h :: t => (case String.explode h of [] => error "parse_line: empty command" | sh :: st => (sh, (String.implode st) :: t)) in process tstate (parse_line str) end fun process_file path thy = (thy, init_state) |> File.fold_lines process_line path |> fst val _ = Outer_Syntax.command \<^command_keyword>\import_file\ "import a recorded proof file" (Parse.path >> (fn name => Toplevel.theory (fn thy => process_file (Path.explode name) thy))) end diff --git a/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML b/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML --- a/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML +++ b/src/HOL/Library/Sum_of_Squares/positivstellensatz.ML @@ -1,779 +1,779 @@ (* Title: HOL/Library/positivstellensatz.ML Author: Amine Chaieb, University of Cambridge A generic arithmetic prover based on Positivstellensatz certificates --- also implements Fourier-Motzkin elimination as a special case Fourier-Motzkin elimination. *) (* A functor for finite mappings based on Tables *) signature FUNC = sig include TABLE val apply : 'a table -> key -> 'a val applyd :'a table -> (key -> 'a) -> key -> 'a val combine : ('a -> 'a -> 'a) -> ('a -> bool) -> 'a table -> 'a table -> 'a table val dom : 'a table -> key list val tryapplyd : 'a table -> key -> 'a -> 'a val updatep : (key * 'a -> bool) -> key * 'a -> 'a table -> 'a table val choose : 'a table -> key * 'a val onefunc : key * 'a -> 'a table end; functor FuncFun(Key: KEY) : FUNC = struct structure Tab = Table(Key); open Tab; fun dom a = sort Key.ord (Tab.keys a); fun applyd f d x = case Tab.lookup f x of SOME y => y | NONE => d x; fun apply f x = applyd f (fn _ => raise Tab.UNDEF x) x; fun tryapplyd f a d = applyd f (K d) a; fun updatep p (k,v) t = if p (k, v) then t else update (k,v) t fun combine f z a b = let fun h (k,v) t = case Tab.lookup t k of NONE => Tab.update (k,v) t | SOME v' => let val w = f v v' in if z w then Tab.delete k t else Tab.update (k,w) t end; in Tab.fold h a b end; fun choose f = (case Tab.min f of SOME entry => entry | NONE => error "FuncFun.choose : Completely empty function") fun onefunc kv = update kv empty end; (* Some standard functors and utility functions for them *) structure FuncUtil = struct structure Intfunc = FuncFun(type key = int val ord = int_ord); structure Ratfunc = FuncFun(type key = Rat.rat val ord = Rat.ord); structure Intpairfunc = FuncFun(type key = int*int val ord = prod_ord int_ord int_ord); structure Symfunc = FuncFun(type key = string val ord = fast_string_ord); structure Termfunc = FuncFun(type key = term val ord = Term_Ord.fast_term_ord); structure Ctermfunc = FuncFun(type key = cterm val ord = Thm.fast_term_ord); type monomial = int Ctermfunc.table; val monomial_ord = list_ord (prod_ord Thm.fast_term_ord int_ord) o apply2 Ctermfunc.dest structure Monomialfunc = FuncFun(type key = monomial val ord = monomial_ord) type poly = Rat.rat Monomialfunc.table; (* The ordering so we can create canonical HOL polynomials. *) fun dest_monomial mon = sort (Thm.fast_term_ord o apply2 fst) (Ctermfunc.dest mon); fun monomial_order (m1,m2) = if Ctermfunc.is_empty m2 then LESS else if Ctermfunc.is_empty m1 then GREATER else let val mon1 = dest_monomial m1 val mon2 = dest_monomial m2 val deg1 = fold (Integer.add o snd) mon1 0 val deg2 = fold (Integer.add o snd) mon2 0 in if deg1 < deg2 then GREATER else if deg1 > deg2 then LESS else list_ord (prod_ord Thm.fast_term_ord int_ord) (mon1,mon2) end; end (* positivstellensatz datatype and prover generation *) signature REAL_ARITH = sig datatype positivstellensatz = Axiom_eq of int | Axiom_le of int | Axiom_lt of int | Rational_eq of Rat.rat | Rational_le of Rat.rat | Rational_lt of Rat.rat | Square of FuncUtil.poly | Eqmul of FuncUtil.poly * positivstellensatz | Sum of positivstellensatz * positivstellensatz | Product of positivstellensatz * positivstellensatz; datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree datatype tree_choice = Left | Right type prover = tree_choice list -> (thm list * thm list * thm list -> positivstellensatz -> thm) -> thm list * thm list * thm list -> thm * pss_tree type cert_conv = cterm -> thm * pss_tree val gen_gen_real_arith : Proof.context -> (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv val real_linear_prover : (thm list * thm list * thm list -> positivstellensatz -> thm) -> thm list * thm list * thm list -> thm * pss_tree val gen_real_arith : Proof.context -> (Rat.rat -> cterm) * conv * conv * conv * conv * conv * conv * conv * prover -> cert_conv val gen_prover_real_arith : Proof.context -> prover -> cert_conv val is_ratconst : cterm -> bool val dest_ratconst : cterm -> Rat.rat val cterm_of_rat : Rat.rat -> cterm end structure RealArith : REAL_ARITH = struct open Conv (* ------------------------------------------------------------------------- *) (* Data structure for Positivstellensatz refutations. *) (* ------------------------------------------------------------------------- *) datatype positivstellensatz = Axiom_eq of int | Axiom_le of int | Axiom_lt of int | Rational_eq of Rat.rat | Rational_le of Rat.rat | Rational_lt of Rat.rat | Square of FuncUtil.poly | Eqmul of FuncUtil.poly * positivstellensatz | Sum of positivstellensatz * positivstellensatz | Product of positivstellensatz * positivstellensatz; (* Theorems used in the procedure *) datatype pss_tree = Trivial | Cert of positivstellensatz | Branch of pss_tree * pss_tree datatype tree_choice = Left | Right type prover = tree_choice list -> (thm list * thm list * thm list -> positivstellensatz -> thm) -> thm list * thm list * thm list -> thm * pss_tree type cert_conv = cterm -> thm * pss_tree (* Some useful derived rules *) fun deduct_antisym_rule tha thb = Thm.equal_intr (Thm.implies_intr (Thm.cprop_of thb) tha) (Thm.implies_intr (Thm.cprop_of tha) thb); fun prove_hyp tha thb = if exists (curry op aconv (Thm.concl_of tha)) (Thm.hyps_of thb) (* FIXME !? *) then Thm.equal_elim (Thm.symmetric (deduct_antisym_rule tha thb)) tha else thb; val pth = @{lemma "(((x::real) < y) \ (y - x > 0))" and "((x \ y) \ (y - x \ 0))" and "((x = y) \ (x - y = 0))" and "((\(x < y)) \ (x - y \ 0))" and "((\(x \ y)) \ (x - y > 0))" and "((\(x = y)) \ (x - y > 0 \ -(x - y) > 0))" by (atomize (full), auto simp add: less_diff_eq le_diff_eq not_less)}; val pth_final = @{lemma "(\p \ False) \ p" by blast} val pth_add = @{lemma "(x = (0::real) \ y = 0 \ x + y = 0 )" and "( x = 0 \ y \ 0 \ x + y \ 0)" and "(x = 0 \ y > 0 \ x + y > 0)" and "(x \ 0 \ y = 0 \ x + y \ 0)" and "(x \ 0 \ y \ 0 \ x + y \ 0)" and "(x \ 0 \ y > 0 \ x + y > 0)" and "(x > 0 \ y = 0 \ x + y > 0)" and "(x > 0 \ y \ 0 \ x + y > 0)" and "(x > 0 \ y > 0 \ x + y > 0)" by simp_all}; val pth_mul = @{lemma "(x = (0::real) \ y = 0 \ x * y = 0)" and "(x = 0 \ y \ 0 \ x * y = 0)" and "(x = 0 \ y > 0 \ x * y = 0)" and "(x \ 0 \ y = 0 \ x * y = 0)" and "(x \ 0 \ y \ 0 \ x * y \ 0)" and "(x \ 0 \ y > 0 \ x * y \ 0)" and "(x > 0 \ y = 0 \ x * y = 0)" and "(x > 0 \ y \ 0 \ x * y \ 0)" and "(x > 0 \ y > 0 \ x * y > 0)" by (auto intro: mult_mono[where a="0::real" and b="x" and d="y" and c="0", simplified] mult_strict_mono[where b="x" and d="y" and a="0" and c="0", simplified])}; val pth_emul = @{lemma "y = (0::real) \ x * y = 0" by simp}; val pth_square = @{lemma "x * x \ (0::real)" by simp}; val weak_dnf_simps = List.take (@{thms simp_thms}, 34) @ @{lemma "((P \ (Q \ R)) = ((P\Q) \ (P\R)))" and "((Q \ R) \ P) = ((Q\P) \ (R\P))" and "(P \ Q) = (Q \ P)" and "((P \ Q) = (Q \ P))" by blast+}; (* val nnfD_simps = @{lemma "((~(P & Q)) = (~P | ~Q))" and "((~(P | Q)) = (~P & ~Q) )" and "((P --> Q) = (~P | Q) )" and "((P = Q) = ((P & Q) | (~P & ~ Q)))" and "((~(P = Q)) = ((P & ~ Q) | (~P & Q)) )" and "((~ ~(P)) = P)" by blast+}; *) val choice_iff = @{lemma "(\x. \y. P x y) = (\f. \x. P x (f x))" by metis}; val prenex_simps = map (fn th => th RS sym) ([@{thm "all_conj_distrib"}, @{thm "ex_disj_distrib"}] @ @{thms "HOL.all_simps"(1-4)} @ @{thms "ex_simps"(1-4)}); val real_abs_thms1 = @{lemma "((-1 * \x::real\ \ r) = (-1 * x \ r \ 1 * x \ r))" and "((-1 * \x\ + a \ r) = (a + -1 * x \ r \ a + 1 * x \ r))" and "((a + -1 * \x\ \ r) = (a + -1 * x \ r \ a + 1 * x \ r))" and "((a + -1 * \x\ + b \ r) = (a + -1 * x + b \ r \ a + 1 * x + b \ r))" and "((a + b + -1 * \x\ \ r) = (a + b + -1 * x \ r \ a + b + 1 * x \ r))" and "((a + b + -1 * \x\ + c \ r) = (a + b + -1 * x + c \ r \ a + b + 1 * x + c \ r))" and "((-1 * max x y \ r) = (-1 * x \ r \ -1 * y \ r))" and "((-1 * max x y + a \ r) = (a + -1 * x \ r \ a + -1 * y \ r))" and "((a + -1 * max x y \ r) = (a + -1 * x \ r \ a + -1 * y \ r))" and "((a + -1 * max x y + b \ r) = (a + -1 * x + b \ r \ a + -1 * y + b \ r))" and "((a + b + -1 * max x y \ r) = (a + b + -1 * x \ r \ a + b + -1 * y \ r))" and "((a + b + -1 * max x y + c \ r) = (a + b + -1 * x + c \ r \ a + b + -1 * y + c \ r))" and "((1 * min x y \ r) = (1 * x \ r \ 1 * y \ r))" and "((1 * min x y + a \ r) = (a + 1 * x \ r \ a + 1 * y \ r))" and "((a + 1 * min x y \ r) = (a + 1 * x \ r \ a + 1 * y \ r))" and "((a + 1 * min x y + b \ r) = (a + 1 * x + b \ r \ a + 1 * y + b \ r))" and "((a + b + 1 * min x y \ r) = (a + b + 1 * x \ r \ a + b + 1 * y \ r))" and "((a + b + 1 * min x y + c \ r) = (a + b + 1 * x + c \ r \ a + b + 1 * y + c \ r))" and "((min x y \ r) = (x \ r \ y \ r))" and "((min x y + a \ r) = (a + x \ r \ a + y \ r))" and "((a + min x y \ r) = (a + x \ r \ a + y \ r))" and "((a + min x y + b \ r) = (a + x + b \ r \ a + y + b \ r))" and "((a + b + min x y \ r) = (a + b + x \ r \ a + b + y \ r))" and "((a + b + min x y + c \ r) = (a + b + x + c \ r \ a + b + y + c \ r))" and "((-1 * \x\ > r) = (-1 * x > r \ 1 * x > r))" and "((-1 * \x\ + a > r) = (a + -1 * x > r \ a + 1 * x > r))" and "((a + -1 * \x\ > r) = (a + -1 * x > r \ a + 1 * x > r))" and "((a + -1 * \x\ + b > r) = (a + -1 * x + b > r \ a + 1 * x + b > r))" and "((a + b + -1 * \x\ > r) = (a + b + -1 * x > r \ a + b + 1 * x > r))" and "((a + b + -1 * \x\ + c > r) = (a + b + -1 * x + c > r \ a + b + 1 * x + c > r))" and "((-1 * max x y > r) = ((-1 * x > r) \ -1 * y > r))" and "((-1 * max x y + a > r) = (a + -1 * x > r \ a + -1 * y > r))" and "((a + -1 * max x y > r) = (a + -1 * x > r \ a + -1 * y > r))" and "((a + -1 * max x y + b > r) = (a + -1 * x + b > r \ a + -1 * y + b > r))" and "((a + b + -1 * max x y > r) = (a + b + -1 * x > r \ a + b + -1 * y > r))" and "((a + b + -1 * max x y + c > r) = (a + b + -1 * x + c > r \ a + b + -1 * y + c > r))" and "((min x y > r) = (x > r \ y > r))" and "((min x y + a > r) = (a + x > r \ a + y > r))" and "((a + min x y > r) = (a + x > r \ a + y > r))" and "((a + min x y + b > r) = (a + x + b > r \ a + y + b > r))" and "((a + b + min x y > r) = (a + b + x > r \ a + b + y > r))" and "((a + b + min x y + c > r) = (a + b + x + c > r \ a + b + y + c > r))" by auto}; val abs_split' = @{lemma "P \x::'a::linordered_idom\ == (x \ 0 \ P x \ x < 0 \ P (-x))" by (atomize (full)) (auto split: abs_split)}; val max_split = @{lemma "P (max x y) \ ((x::'a::linorder) \ y \ P y \ x > y \ P x)" by (atomize (full)) (cases "x \ y", auto simp add: max_def)}; val min_split = @{lemma "P (min x y) \ ((x::'a::linorder) \ y \ P x \ x > y \ P y)" by (atomize (full)) (cases "x \ y", auto simp add: min_def)}; (* Miscellaneous *) fun literals_conv bops uops cv = let fun h t = (case Thm.term_of t of b$_$_ => if member (op aconv) bops b then binop_conv h t else cv t | u$_ => if member (op aconv) uops u then arg_conv h t else cv t | _ => cv t) in h end; fun cterm_of_rat x = let val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber \<^ctyp>\real\ a else Thm.apply (Thm.apply \<^cterm>\(/) :: real \ _\ (Numeral.mk_cnumber \<^ctyp>\real\ a)) (Numeral.mk_cnumber \<^ctyp>\real\ b) end; fun dest_ratconst t = case Thm.term_of t of Const(\<^const_name>\divide\, _)$a$b => Rat.make(HOLogic.dest_number a |> snd, HOLogic.dest_number b |> snd) | _ => Rat.of_int (HOLogic.dest_number (Thm.term_of t) |> snd) fun is_ratconst t = can dest_ratconst t (* fun find_term p t = if p t then t else case t of a$b => (find_term p a handle TERM _ => find_term p b) | Abs (_,_,t') => find_term p t' | _ => raise TERM ("find_term",[t]); *) fun find_cterm p t = if p t then t else case Thm.term_of t of _$_ => (find_cterm p (Thm.dest_fun t) handle CTERM _ => find_cterm p (Thm.dest_arg t)) | Abs (_,_,_) => find_cterm p (Thm.dest_abs NONE t |> snd) | _ => raise CTERM ("find_cterm",[t]); fun is_comb t = (case Thm.term_of t of _ $ _ => true | _ => false); fun is_binop ct ct' = ct aconvc (Thm.dest_fun (Thm.dest_fun ct')) handle CTERM _ => false; (* Map back polynomials to HOL. *) fun cterm_of_varpow x k = if k = 1 then x else Thm.apply (Thm.apply \<^cterm>\(^) :: real \ _\ x) (Numeral.mk_cnumber \<^ctyp>\nat\ k) fun cterm_of_monomial m = if FuncUtil.Ctermfunc.is_empty m then \<^cterm>\1::real\ else let val m' = FuncUtil.dest_monomial m val vps = fold_rev (fn (x,k) => cons (cterm_of_varpow x k)) m' [] in foldr1 (fn (s, t) => Thm.apply (Thm.apply \<^cterm>\(*) :: real \ _\ s) t) vps end fun cterm_of_cmonomial (m,c) = if FuncUtil.Ctermfunc.is_empty m then cterm_of_rat c else if c = @1 then cterm_of_monomial m else Thm.apply (Thm.apply \<^cterm>\(*)::real \ _\ (cterm_of_rat c)) (cterm_of_monomial m); fun cterm_of_poly p = if FuncUtil.Monomialfunc.is_empty p then \<^cterm>\0::real\ else let val cms = map cterm_of_cmonomial (sort (prod_ord FuncUtil.monomial_order (K EQUAL)) (FuncUtil.Monomialfunc.dest p)) in foldr1 (fn (t1, t2) => Thm.apply(Thm.apply \<^cterm>\(+) :: real \ _\ t1) t2) cms end; (* A general real arithmetic prover *) fun gen_gen_real_arith ctxt (mk_numeric, numeric_eq_conv,numeric_ge_conv,numeric_gt_conv, poly_conv,poly_neg_conv,poly_add_conv,poly_mul_conv, absconv1,absconv2,prover) = let val pre_ss = put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms ex_simps all_simps not_all not_ex ex_disj_distrib all_conj_distrib if_bool_eq_disj} val prenex_ss = put_simpset HOL_basic_ss ctxt addsimps prenex_simps val skolemize_ss = put_simpset HOL_basic_ss ctxt addsimps [choice_iff] val presimp_conv = Simplifier.rewrite pre_ss val prenex_conv = Simplifier.rewrite prenex_ss val skolemize_conv = Simplifier.rewrite skolemize_ss val weak_dnf_ss = put_simpset HOL_basic_ss ctxt addsimps weak_dnf_simps val weak_dnf_conv = Simplifier.rewrite weak_dnf_ss fun eqT_elim th = Thm.equal_elim (Thm.symmetric th) @{thm TrueI} fun oprconv cv ct = let val g = Thm.dest_fun2 ct in if g aconvc \<^cterm>\(\) :: real \ _\ orelse g aconvc \<^cterm>\(<) :: real \ _\ then arg_conv cv ct else arg1_conv cv ct end fun real_ineq_conv th ct = let val th' = (Thm.instantiate (Thm.match (Thm.lhs_of th, ct)) th handle Pattern.MATCH => raise CTERM ("real_ineq_conv", [ct])) in Thm.transitive th' (oprconv poly_conv (Thm.rhs_of th')) end val [real_lt_conv, real_le_conv, real_eq_conv, real_not_lt_conv, real_not_le_conv, _] = map real_ineq_conv pth fun match_mp_rule ths ths' = let fun f ths ths' = case ths of [] => raise THM("match_mp_rule",0,ths) | th::ths => (ths' MRS th handle THM _ => f ths ths') in f ths ths' end fun mul_rule th th' = fconv_rule (arg_conv (oprconv poly_mul_conv)) (match_mp_rule pth_mul [th, th']) fun add_rule th th' = fconv_rule (arg_conv (oprconv poly_add_conv)) (match_mp_rule pth_add [th, th']) fun emul_rule ct th = fconv_rule (arg_conv (oprconv poly_mul_conv)) (Thm.instantiate' [] [SOME ct] (th RS pth_emul)) fun square_rule t = fconv_rule (arg_conv (oprconv poly_conv)) (Thm.instantiate' [] [SOME t] pth_square) fun hol_of_positivstellensatz(eqs,les,lts) proof = let fun translate prf = case prf of Axiom_eq n => nth eqs n | Axiom_le n => nth les n | Axiom_lt n => nth lts n | Rational_eq x => eqT_elim(numeric_eq_conv(Thm.apply \<^cterm>\Trueprop\ (Thm.apply (Thm.apply \<^cterm>\(=)::real \ _\ (mk_numeric x)) \<^cterm>\0::real\))) | Rational_le x => eqT_elim(numeric_ge_conv(Thm.apply \<^cterm>\Trueprop\ (Thm.apply (Thm.apply \<^cterm>\(\)::real \ _\ \<^cterm>\0::real\) (mk_numeric x)))) | Rational_lt x => eqT_elim(numeric_gt_conv(Thm.apply \<^cterm>\Trueprop\ (Thm.apply (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) (mk_numeric x)))) | Square pt => square_rule (cterm_of_poly pt) | Eqmul(pt,p) => emul_rule (cterm_of_poly pt) (translate p) | Sum(p1,p2) => add_rule (translate p1) (translate p2) | Product(p1,p2) => mul_rule (translate p1) (translate p2) in fconv_rule (first_conv [numeric_ge_conv, numeric_gt_conv, numeric_eq_conv, all_conv]) (translate proof) end val init_conv = presimp_conv then_conv nnf_conv ctxt then_conv skolemize_conv then_conv prenex_conv then_conv weak_dnf_conv val concl = Thm.dest_arg o Thm.cprop_of fun is_binop opr ct = (Thm.dest_fun2 ct aconvc opr handle CTERM _ => false) val is_req = is_binop \<^cterm>\(=):: real \ _\ val is_ge = is_binop \<^cterm>\(\):: real \ _\ val is_gt = is_binop \<^cterm>\(<):: real \ _\ val is_conj = is_binop \<^cterm>\HOL.conj\ val is_disj = is_binop \<^cterm>\HOL.disj\ fun conj_pair th = (th RS @{thm conjunct1}, th RS @{thm conjunct2}) fun disj_cases th th1 th2 = let val (p,q) = Thm.dest_binop (concl th) val c = concl th1 val _ = if c aconvc (concl th2) then () else error "disj_cases : conclusions not alpha convertible" in Thm.implies_elim (Thm.implies_elim (Thm.implies_elim (Thm.instantiate' [] (map SOME [p,q,c]) @{thm disjE}) th) (Thm.implies_intr (Thm.apply \<^cterm>\Trueprop\ p) th1)) (Thm.implies_intr (Thm.apply \<^cterm>\Trueprop\ q) th2) end fun overall cert_choice dun ths = case ths of [] => let val (eq,ne) = List.partition (is_req o concl) dun val (le,nl) = List.partition (is_ge o concl) ne val lt = filter (is_gt o concl) nl in prover (rev cert_choice) hol_of_positivstellensatz (eq,le,lt) end | th::oths => let val ct = concl th in if is_conj ct then let val (th1,th2) = conj_pair th in overall cert_choice dun (th1::th2::oths) end else if is_disj ct then let val (th1, cert1) = overall (Left::cert_choice) dun (Thm.assume (Thm.apply \<^cterm>\Trueprop\ (Thm.dest_arg1 ct))::oths) val (th2, cert2) = overall (Right::cert_choice) dun (Thm.assume (Thm.apply \<^cterm>\Trueprop\ (Thm.dest_arg ct))::oths) in (disj_cases th th1 th2, Branch (cert1, cert2)) end else overall cert_choice (th::dun) oths end fun dest_binary b ct = if is_binop b ct then Thm.dest_binop ct else raise CTERM ("dest_binary",[b,ct]) val dest_eq = dest_binary \<^cterm>\(=) :: real \ _\ val neq_th = nth pth 5 fun real_not_eq_conv ct = let val (l,r) = dest_eq (Thm.dest_arg ct) - val th = Thm.instantiate ([],[((("x", 0), \<^typ>\real\),l),((("y", 0), \<^typ>\real\),r)]) neq_th + val th = Thm.instantiate (TVars.empty, Vars.make [((("x", 0), \<^typ>\real\),l),((("y", 0), \<^typ>\real\),r)]) neq_th val th_p = poly_conv(Thm.dest_arg(Thm.dest_arg1(Thm.rhs_of th))) val th_x = Drule.arg_cong_rule \<^cterm>\uminus :: real \ _\ th_p val th_n = fconv_rule (arg_conv poly_neg_conv) th_x val th' = Drule.binop_cong_rule \<^cterm>\HOL.disj\ (Drule.arg_cong_rule (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) th_p) (Drule.arg_cong_rule (Thm.apply \<^cterm>\(<)::real \ _\ \<^cterm>\0::real\) th_n) in Thm.transitive th th' end fun equal_implies_1_rule PQ = let val P = Thm.lhs_of PQ in Thm.implies_intr P (Thm.equal_elim PQ (Thm.assume P)) end (*FIXME!!! Copied from groebner.ml*) val strip_exists = let fun h (acc, t) = case Thm.term_of t of Const(\<^const_name>\Ex\,_)$Abs(_,_,_) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) end fun name_of x = case Thm.term_of x of Free(s,_) => s | Var ((s,_),_) => s | _ => "x" fun mk_forall x th = let val T = Thm.typ_of_cterm x val all = Thm.cterm_of ctxt (Const (\<^const_name>\All\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) in Drule.arg_cong_rule all (Thm.abstract_rule (name_of x) x th) end val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); fun ext T = Thm.cterm_of ctxt (Const (\<^const_name>\Ex\, (T --> \<^typ>\bool\) --> \<^typ>\bool\)) fun mk_ex v t = Thm.apply (ext (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 = 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 | _ => raise THM ("choose",0,[th, th']) fun simple_choose v th = choose v (Thm.assume ((Thm.apply \<^cterm>\Trueprop\ o mk_ex v) (Thm.dest_arg (hd (Thm.chyps_of th))))) th val strip_forall = let fun h (acc, t) = case Thm.term_of t of Const(\<^const_name>\All\,_)$Abs(_,_,_) => h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) | _ => (acc,t) in fn t => h ([],t) end fun f ct = let val nnf_norm_conv' = nnf_conv ctxt then_conv literals_conv [\<^term>\HOL.conj\, \<^term>\HOL.disj\] [] (Conv.cache_conv (first_conv [real_lt_conv, real_le_conv, real_eq_conv, real_not_lt_conv, real_not_le_conv, real_not_eq_conv, all_conv])) fun absremover ct = (literals_conv [\<^term>\HOL.conj\, \<^term>\HOL.disj\] [] (try_conv (absconv1 then_conv binop_conv (arg_conv poly_conv))) then_conv try_conv (absconv2 then_conv nnf_norm_conv' then_conv binop_conv absremover)) ct val nct = Thm.apply \<^cterm>\Trueprop\ (Thm.apply \<^cterm>\Not\ ct) val th0 = (init_conv then_conv arg_conv nnf_norm_conv') nct val tm0 = Thm.dest_arg (Thm.rhs_of th0) val (th, certificates) = if tm0 aconvc \<^cterm>\False\ then (equal_implies_1_rule th0, Trivial) else let val (evs,bod) = strip_exists tm0 val (avs,ibod) = strip_forall bod val th1 = Drule.arg_cong_rule \<^cterm>\Trueprop\ (fold mk_forall avs (absremover ibod)) val (th2, certs) = overall [] [] [specl avs (Thm.assume (Thm.rhs_of th1))] val th3 = fold simple_choose evs (prove_hyp (Thm.equal_elim th1 (Thm.assume (Thm.apply \<^cterm>\Trueprop\ bod))) th2) in (Drule.implies_intr_hyps (prove_hyp (Thm.equal_elim th0 (Thm.assume nct)) th3), certs) end in (Thm.implies_elim (Thm.instantiate' [] [SOME ct] pth_final) th, certificates) end in f end; (* A linear arithmetic prover *) local val linear_add = FuncUtil.Ctermfunc.combine (curry op +) (fn z => z = @0) fun linear_cmul c = FuncUtil.Ctermfunc.map (fn _ => fn x => c * x) val one_tm = \<^cterm>\1::real\ fun contradictory p (e,_) = ((FuncUtil.Ctermfunc.is_empty e) andalso not(p @0)) orelse ((eq_set (op aconvc) (FuncUtil.Ctermfunc.dom e, [one_tm])) andalso not(p(FuncUtil.Ctermfunc.apply e one_tm))) fun linear_ineqs vars (les,lts) = case find_first (contradictory (fn x => x > @0)) lts of SOME r => r | NONE => (case find_first (contradictory (fn x => x > @0)) les of SOME r => r | NONE => if null vars then error "linear_ineqs: no contradiction" else let val ineqs = les @ lts fun blowup v = length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) ineqs) + length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) ineqs) * length(filter (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 < @0) ineqs) val v = fst(hd(sort (fn ((_,i),(_,j)) => int_ord (i,j)) (map (fn v => (v,blowup v)) vars))) fun addup (e1,p1) (e2,p2) acc = let val c1 = FuncUtil.Ctermfunc.tryapplyd e1 v @0 val c2 = FuncUtil.Ctermfunc.tryapplyd e2 v @0 in if c1 * c2 >= @0 then acc else let val e1' = linear_cmul (abs c2) e1 val e2' = linear_cmul (abs c1) e2 val p1' = Product(Rational_lt (abs c2), p1) val p2' = Product(Rational_lt (abs c1), p2) in (linear_add e1' e2',Sum(p1',p2'))::acc end end val (les0,les1) = List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) les val (lts0,lts1) = List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 = @0) lts val (lesp,lesn) = List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) les1 val (ltsp,ltsn) = List.partition (fn (e,_) => FuncUtil.Ctermfunc.tryapplyd e v @0 > @0) lts1 val les' = fold_rev (fn ep1 => fold_rev (addup ep1) lesp) lesn les0 val lts' = fold_rev (fn ep1 => fold_rev (addup ep1) (lesp@ltsp)) ltsn (fold_rev (fn ep1 => fold_rev (addup ep1) (lesn@ltsn)) ltsp lts0) in linear_ineqs (remove (op aconvc) v vars) (les',lts') end) fun linear_eqs(eqs,les,lts) = case find_first (contradictory (fn x => x = @0)) eqs of SOME r => r | NONE => (case eqs of [] => let val vars = remove (op aconvc) one_tm (fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom o fst) (les@lts) []) in linear_ineqs vars (les,lts) end | (e,p)::es => if FuncUtil.Ctermfunc.is_empty e then linear_eqs (es,les,lts) else let val (x,c) = FuncUtil.Ctermfunc.choose (FuncUtil.Ctermfunc.delete_safe one_tm e) fun xform (inp as (t,q)) = let val d = FuncUtil.Ctermfunc.tryapplyd t x @0 in if d = @0 then inp else let val k = ~ d * abs c / c val e' = linear_cmul k e val t' = linear_cmul (abs c) t val p' = Eqmul(FuncUtil.Monomialfunc.onefunc (FuncUtil.Ctermfunc.empty, k),p) val q' = Product(Rational_lt (abs c), q) in (linear_add e' t',Sum(p',q')) end end in linear_eqs(map xform es,map xform les,map xform lts) end) fun linear_prover (eq,le,lt) = let val eqs = map_index (fn (n, p) => (p,Axiom_eq n)) eq val les = map_index (fn (n, p) => (p,Axiom_le n)) le val lts = map_index (fn (n, p) => (p,Axiom_lt n)) lt in linear_eqs(eqs,les,lts) end fun lin_of_hol ct = if ct aconvc \<^cterm>\0::real\ then FuncUtil.Ctermfunc.empty else if not (is_comb ct) then FuncUtil.Ctermfunc.onefunc (ct, @1) else if is_ratconst ct then FuncUtil.Ctermfunc.onefunc (one_tm, dest_ratconst ct) else let val (lop,r) = Thm.dest_comb ct in if not (is_comb lop) then FuncUtil.Ctermfunc.onefunc (ct, @1) else let val (opr,l) = Thm.dest_comb lop in if opr aconvc \<^cterm>\(+) :: real \ _\ then linear_add (lin_of_hol l) (lin_of_hol r) else if opr aconvc \<^cterm>\(*) :: real \ _\ andalso is_ratconst l then FuncUtil.Ctermfunc.onefunc (r, dest_ratconst l) else FuncUtil.Ctermfunc.onefunc (ct, @1) end end fun is_alien ct = case Thm.term_of ct of Const(\<^const_name>\of_nat\, _)$ n => not (can HOLogic.dest_number n) | Const(\<^const_name>\of_int\, _)$ n => not (can HOLogic.dest_number n) | _ => false in fun real_linear_prover translator (eq,le,lt) = let val lhs = lin_of_hol o Thm.dest_arg1 o Thm.dest_arg o Thm.cprop_of val rhs = lin_of_hol o Thm.dest_arg o Thm.dest_arg o Thm.cprop_of val eq_pols = map lhs eq val le_pols = map rhs le val lt_pols = map rhs lt val aliens = filter is_alien (fold_rev (union (op aconvc) o FuncUtil.Ctermfunc.dom) (eq_pols @ le_pols @ lt_pols) []) val le_pols' = le_pols @ map (fn v => FuncUtil.Ctermfunc.onefunc (v,@1)) aliens val (_,proof) = linear_prover (eq_pols,le_pols',lt_pols) val le' = le @ map (fn a => Thm.instantiate' [] [SOME (Thm.dest_arg a)] @{thm of_nat_0_le_iff}) aliens in ((translator (eq,le',lt) proof), Trivial) end end; (* A less general generic arithmetic prover dealing with abs,max and min*) local val absmaxmin_elim_ss1 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps real_abs_thms1) fun absmaxmin_elim_conv1 ctxt = Simplifier.rewrite (put_simpset absmaxmin_elim_ss1 ctxt) val absmaxmin_elim_conv2 = let val pth_abs = Thm.instantiate' [SOME \<^ctyp>\real\] [] abs_split' val pth_max = Thm.instantiate' [SOME \<^ctyp>\real\] [] max_split val pth_min = Thm.instantiate' [SOME \<^ctyp>\real\] [] min_split val abs_tm = \<^cterm>\abs :: real \ _\ val p_v = (("P", 0), \<^typ>\real \ bool\) val x_v = (("x", 0), \<^typ>\real\) val y_v = (("y", 0), \<^typ>\real\) val is_max = is_binop \<^cterm>\max :: real \ _\ val is_min = is_binop \<^cterm>\min :: real \ _\ fun is_abs t = is_comb t andalso Thm.dest_fun t aconvc abs_tm fun eliminate_construct p c tm = let val t = find_cterm p tm val th0 = (Thm.symmetric o Thm.beta_conversion false) (Thm.apply (Thm.lambda t tm) t) val (p,ax) = (Thm.dest_comb o Thm.rhs_of) th0 in fconv_rule(arg_conv(binop_conv (arg_conv (Thm.beta_conversion false)))) (Thm.transitive th0 (c p ax)) end val elim_abs = eliminate_construct is_abs (fn p => fn ax => - Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax)]) pth_abs) + Thm.instantiate (TVars.empty, Vars.make [(p_v,p), (x_v, Thm.dest_arg ax)]) pth_abs) val elim_max = eliminate_construct is_max (fn p => fn ax => let val (ax,y) = Thm.dest_comb ax - in Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) + in Thm.instantiate (TVars.empty, Vars.make [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) pth_max end) val elim_min = eliminate_construct is_min (fn p => fn ax => let val (ax,y) = Thm.dest_comb ax - in Thm.instantiate ([], [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) + in Thm.instantiate (TVars.empty, Vars.make [(p_v,p), (x_v, Thm.dest_arg ax), (y_v,y)]) pth_min end) in first_conv [elim_abs, elim_max, elim_min, all_conv] end; in fun gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul,prover) = gen_gen_real_arith ctxt (mkconst,eq,ge,gt,norm,neg,add,mul, absmaxmin_elim_conv1 ctxt,absmaxmin_elim_conv2,prover) end; (* An instance for reals*) fun gen_prover_real_arith ctxt prover = let val {add, mul, neg, pow = _, sub = _, main} = Semiring_Normalizer.semiring_normalizers_ord_wrapper ctxt (the (Semiring_Normalizer.match ctxt \<^cterm>\(0::real) + 1\)) Thm.term_ord in gen_real_arith ctxt (cterm_of_rat, Numeral_Simprocs.field_comp_conv ctxt, Numeral_Simprocs.field_comp_conv ctxt, Numeral_Simprocs.field_comp_conv ctxt, main ctxt, neg ctxt, add ctxt, mul ctxt, prover) end; end diff --git a/src/HOL/Library/cconv.ML b/src/HOL/Library/cconv.ML --- a/src/HOL/Library/cconv.ML +++ b/src/HOL/Library/cconv.ML @@ -1,256 +1,259 @@ (* Title: HOL/Library/cconv.ML Author: Christoph Traut, Lars Noschinski, TU Muenchen FIXME!? *) infix 1 then_cconv infix 0 else_cconv type cconv = conv signature BASIC_CCONV = sig val then_cconv: cconv * cconv -> cconv val else_cconv: cconv * cconv -> cconv val CCONVERSION: cconv -> int -> tactic end signature CCONV = sig include BASIC_CCONV val no_cconv: cconv val all_cconv: cconv val first_cconv: cconv list -> cconv val abs_cconv: (cterm * Proof.context -> cconv) -> Proof.context -> cconv val combination_cconv: cconv -> cconv -> cconv val comb_cconv: cconv -> cconv val arg_cconv: cconv -> cconv val fun_cconv: cconv -> cconv val arg1_cconv: cconv -> cconv val fun2_cconv: cconv -> cconv val rewr_cconv: thm -> cconv val rewrs_cconv: thm list -> cconv val params_cconv: int -> (Proof.context -> cconv) -> Proof.context -> cconv val prems_cconv: int -> cconv -> cconv val with_prems_cconv: int -> cconv -> cconv val concl_cconv: int -> cconv -> cconv val fconv_rule: cconv -> thm -> thm val gconv_rule: cconv -> int -> thm -> thm end structure CConv : CCONV = struct val concl_lhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_lhs val concl_rhs_of = Thm.cprop_of #> Drule.strip_imp_concl #> Thm.dest_equals_rhs fun transitive th1 th2 = Drule.transitive_thm OF [th1, th2] val combination_thm = let val fg = \<^cprop>\f :: 'a :: {} \ 'b :: {} \ g\ val st = \<^cprop>\s :: 'a :: {} \ t\ val thm = Thm.combination (Thm.assume fg) (Thm.assume st) |> Thm.implies_intr st |> Thm.implies_intr fg in Drule.export_without_context thm end fun abstract_rule_thm n = let val eq = \<^cprop>\\x :: 'a :: {}. (s :: 'a \ 'b :: {}) x \ t x\ val x = \<^cterm>\x :: 'a :: {}\ val thm = eq |> Thm.assume |> Thm.forall_elim x |> Thm.abstract_rule n x |> Thm.implies_intr eq in Drule.export_without_context thm end val no_cconv = Conv.no_conv val all_cconv = Conv.all_conv fun (cv1 else_cconv cv2) ct = (cv1 ct handle THM _ => cv2 ct | CTERM _ => cv2 ct | TERM _ => cv2 ct | TYPE _ => cv2 ct) fun (cv1 then_cconv cv2) ct = let val eq1 = cv1 ct val eq2 = cv2 (concl_rhs_of eq1) in if Thm.is_reflexive eq1 then eq2 else if Thm.is_reflexive eq2 then eq1 else transitive eq1 eq2 end fun first_cconv cvs = fold_rev (curry op else_cconv) cvs no_cconv fun rewr_cconv rule ct = let val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule val lhs = concl_lhs_of rule1 val rule2 = Thm.rename_boundvars (Thm.term_of lhs) (Thm.term_of ct) rule1 val rule3 = Thm.instantiate (Thm.match (lhs, ct)) rule2 handle Pattern.MATCH => raise CTERM ("rewr_cconv", [lhs, ct]) val concl = rule3 |> Thm.cprop_of |> Drule.strip_imp_concl val rule4 = if Thm.dest_equals_lhs concl aconvc ct then rule3 else let val ceq = Thm.dest_fun2 concl in rule3 RS Thm.trivial (Thm.mk_binop ceq ct (Thm.dest_equals_rhs concl)) end in transitive rule4 (Thm.beta_conversion true (concl_rhs_of rule4)) end fun rewrs_cconv rules = first_cconv (map rewr_cconv rules) fun combination_cconv cv1 cv2 cterm = let val (l, r) = Thm.dest_comb cterm in combination_thm OF [cv1 l, cv2 r] end fun comb_cconv cv = combination_cconv cv cv fun fun_cconv conversion = combination_cconv conversion all_cconv fun arg_cconv conversion = combination_cconv all_cconv conversion fun abs_cconv cv ctxt ct = (case Thm.term_of ct of Abs (x, _, _) => let (* Instantiate the rule properly and apply it to the eq theorem. *) fun abstract_rule u v eq = let (* Take a variable v and an equality theorem of form: P1 \ ... \ Pn \ L v \ R v And build a term of form: \v. (\x. L x) v \ (\x. R x) v *) fun mk_concl var eq = let val certify = Thm.cterm_of ctxt fun abs term = (Term.lambda var term) $ var fun equals_cong f t = Logic.dest_equals t |> (fn (a, b) => (f a, f b)) |> Logic.mk_equals in Thm.concl_of eq |> equals_cong abs |> Logic.all var |> certify end val rule = abstract_rule_thm x val inst = Thm.match (Drule.cprems_of rule |> hd, mk_concl (Thm.term_of v) eq) in (Drule.instantiate_normalize inst rule OF [Drule.generalize (Names.empty, Names.make_set [u]) eq]) |> Drule.zero_var_indexes end (* Destruct the abstraction and apply the conversion. *) val (u, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt val (v, ct') = Thm.dest_abs (SOME u) ct val eq = cv (v, ctxt') ct' in if Thm.is_reflexive eq then all_cconv ct else abstract_rule u v eq end | _ => raise CTERM ("abs_cconv", [ct])) val arg1_cconv = fun_cconv o arg_cconv val fun2_cconv = fun_cconv o fun_cconv (* conversions on HHF rules *) (*rewrite B in \x1 ... xn. B*) fun params_cconv n cv ctxt ct = if n <> 0 andalso Logic.is_all (Thm.term_of ct) then arg_cconv (abs_cconv (params_cconv (n - 1) cv o #2) ctxt) ct else cv ctxt ct (* TODO: This code behaves not exactly like Conv.prems_cconv does. Fix this! *) (*rewrite the A's in A1 \ ... \ An \ B*) fun prems_cconv 0 cv ct = cv ct | prems_cconv n cv ct = (case ct |> Thm.term_of of (Const (\<^const_name>\Pure.imp\, _) $ _) $ _ => ((if n = 1 then fun_cconv else I) o arg_cconv) (prems_cconv (n-1) cv) ct | _ => cv ct) -fun inst_imp_cong ct = Thm.instantiate ([], [((("A", 0), propT), ct)]) Drule.imp_cong +fun inst_imp_cong ct = + Thm.instantiate (TVars.empty, Vars.make [((("A", 0), propT), ct)]) Drule.imp_cong (*rewrite B in A1 \ ... \ An \ B*) fun concl_cconv 0 cv ct = cv ct | concl_cconv n cv ct = (case try Thm.dest_implies ct of NONE => cv ct | SOME (A,B) => (concl_cconv (n-1) cv B) RS inst_imp_cong A) (* Rewrites A in A \ A1 \ An \ B. The premises of the resulting theorem assume A1, ..., An *) fun with_prems_cconv n cv ct = let fun strip_prems 0 As B = (As, B) | strip_prems i As B = case try Thm.dest_implies B of NONE => (As, B) | SOME (A,B) => strip_prems (i - 1) (A::As) B val (prem, (prems, concl)) = ct |> Thm.dest_implies ||> strip_prems n [] - val rewr_imp_concl = Thm.instantiate ([], [((("C", 0), propT), concl)]) @{thm rewr_imp} + val rewr_imp_concl = + Thm.instantiate (TVars.empty, Vars.make [((("C", 0), propT), concl)]) @{thm rewr_imp} val th1 = cv prem RS rewr_imp_concl val nprems = Thm.nprems_of th1 - fun inst_cut_rl ct = Thm.instantiate ([], [((("psi", 0), propT), ct)]) cut_rl + fun inst_cut_rl ct = + Thm.instantiate (TVars.empty, Vars.make [((("psi", 0), propT), ct)]) cut_rl fun f p th = (th RS inst_cut_rl p) |> Conv.fconv_rule (Conv.concl_conv nprems (Conv.rewr_conv @{thm imp_cong_eq})) in fold f prems th1 end (*forward conversion, cf. FCONV_RULE in LCF*) fun fconv_rule cv th = let val eq = cv (Thm.cprop_of th) in if Thm.is_reflexive eq then th else th COMP (Thm.permute_prems 0 (Thm.nprems_of eq) (eq RS Drule.equal_elim_rule1)) end (*goal conversion*) fun gconv_rule cv i th = (case try (Thm.cprem_of th) i of SOME ct => let val eq = cv ct (* Drule.with_subgoal assumes that there are no new premises generated and thus rotates the premises wrongly. *) fun with_subgoal i f thm = let val num_prems = Thm.nprems_of thm val rotate_to_front = rotate_prems (i - 1) fun rotate_back thm = rotate_prems (1 - i + num_prems - Thm.nprems_of thm) thm in thm |> rotate_to_front |> f |> rotate_back end in if Thm.is_reflexive eq then th else with_subgoal i (fconv_rule (arg1_cconv (K eq))) th end | NONE => raise THM ("gconv_rule", i, [th])) (* Conditional conversions as tactics. *) fun CCONVERSION cv i st = Seq.single (gconv_rule cv i st) handle THM _ => Seq.empty | CTERM _ => Seq.empty | TERM _ => Seq.empty | TYPE _ => Seq.empty end structure Basic_CConv: BASIC_CCONV = CConv open Basic_CConv diff --git a/src/HOL/Library/old_recdef.ML b/src/HOL/Library/old_recdef.ML --- a/src/HOL/Library/old_recdef.ML +++ b/src/HOL/Library/old_recdef.ML @@ -1,2891 +1,2892 @@ (* Title: HOL/Library/old_recdef.ML Author: Konrad Slind, Cambridge University Computer Laboratory Author: Lucas Dixon, University of Edinburgh Old TFL/recdef package. *) signature CASE_SPLIT = sig (* try to recursively split conjectured thm to given list of thms *) val splitto : Proof.context -> thm list -> thm -> thm end; signature UTILS = sig exception ERR of {module: string, func: string, mesg: string} val end_itlist: ('a -> 'a -> 'a) -> 'a list -> 'a val itlist2: ('a -> 'b -> 'c -> 'c) -> 'a list -> 'b list -> 'c -> 'c val pluck: ('a -> bool) -> 'a list -> 'a * 'a list val zip3: 'a list -> 'b list -> 'c list -> ('a*'b*'c) list val take: ('a -> 'b) -> int * 'a list -> 'b list end; signature USYNTAX = sig datatype lambda = VAR of {Name : string, Ty : typ} | CONST of {Name : string, Ty : typ} | COMB of {Rator: term, Rand : term} | LAMB of {Bvar : term, Body : term} val alpha : typ (* Types *) val type_vars : typ -> typ list val type_varsl : typ list -> typ list val mk_vartype : string -> typ val is_vartype : typ -> bool val strip_prod_type : typ -> typ list (* Terms *) val free_vars_lr : term -> term list val type_vars_in_term : term -> typ list val dest_term : term -> lambda (* Prelogic *) val inst : (typ*typ) list -> term -> term (* Construction routines *) val mk_abs :{Bvar : term, Body : term} -> term val mk_imp :{ant : term, conseq : term} -> term val mk_select :{Bvar : term, Body : term} -> term val mk_forall :{Bvar : term, Body : term} -> term val mk_exists :{Bvar : term, Body : term} -> term val mk_conj :{conj1 : term, conj2 : term} -> term val mk_disj :{disj1 : term, disj2 : term} -> term val mk_pabs :{varstruct : term, body : term} -> term (* Destruction routines *) val dest_const: term -> {Name : string, Ty : typ} val dest_comb : term -> {Rator : term, Rand : term} val dest_abs : string list -> term -> {Bvar : term, Body : term} * string list val dest_eq : term -> {lhs : term, rhs : term} val dest_imp : term -> {ant : term, conseq : term} val dest_forall : term -> {Bvar : term, Body : term} val dest_exists : term -> {Bvar : term, Body : term} val dest_neg : term -> term val dest_conj : term -> {conj1 : term, conj2 : term} val dest_disj : term -> {disj1 : term, disj2 : term} val dest_pair : term -> {fst : term, snd : term} val dest_pabs : string list -> term -> {varstruct : term, body : term, used : string list} val lhs : term -> term val rhs : term -> term val rand : term -> term (* Query routines *) val is_imp : term -> bool val is_forall : term -> bool val is_exists : term -> bool val is_neg : term -> bool val is_conj : term -> bool val is_disj : term -> bool val is_pair : term -> bool val is_pabs : term -> bool (* Construction of a term from a list of Preterms *) val list_mk_abs : (term list * term) -> term val list_mk_imp : (term list * term) -> term val list_mk_forall : (term list * term) -> term val list_mk_conj : term list -> term (* Destructing a term to a list of Preterms *) val strip_comb : term -> (term * term list) val strip_abs : term -> (term list * term) val strip_imp : term -> (term list * term) val strip_forall : term -> (term list * term) val strip_exists : term -> (term list * term) val strip_disj : term -> term list (* Miscellaneous *) val mk_vstruct : typ -> term list -> term val gen_all : term -> term val find_term : (term -> bool) -> term -> term option val dest_relation : term -> term * term * term val is_WFR : term -> bool val ARB : typ -> term end; signature DCTERM = sig val dest_comb: cterm -> cterm * cterm val dest_abs: string option -> cterm -> cterm * cterm val capply: cterm -> cterm -> cterm val cabs: cterm -> cterm -> cterm val mk_conj: cterm * cterm -> cterm val mk_disj: cterm * cterm -> cterm val mk_exists: cterm * cterm -> cterm val dest_conj: cterm -> cterm * cterm val dest_const: cterm -> {Name: string, Ty: typ} val dest_disj: cterm -> cterm * cterm val dest_eq: cterm -> cterm * cterm val dest_exists: cterm -> cterm * cterm val dest_forall: cterm -> cterm * cterm val dest_imp: cterm -> cterm * cterm val dest_neg: cterm -> cterm val dest_pair: cterm -> cterm * cterm val dest_var: cterm -> {Name:string, Ty:typ} val is_conj: cterm -> bool val is_disj: cterm -> bool val is_eq: cterm -> bool val is_exists: cterm -> bool val is_forall: cterm -> bool val is_imp: cterm -> bool val is_neg: cterm -> bool val is_pair: cterm -> bool val list_mk_disj: cterm list -> cterm val strip_abs: cterm -> cterm list * cterm val strip_comb: cterm -> cterm * cterm list val strip_disj: cterm -> cterm list val strip_exists: cterm -> cterm list * cterm val strip_forall: cterm -> cterm list * cterm val strip_imp: cterm -> cterm list * cterm val drop_prop: cterm -> cterm val mk_prop: cterm -> cterm end; signature RULES = sig val dest_thm: thm -> term list * term (* Inference rules *) val REFL: cterm -> thm val ASSUME: cterm -> thm val MP: thm -> thm -> thm val MATCH_MP: thm -> thm -> thm val CONJUNCT1: thm -> thm val CONJUNCT2: thm -> thm val CONJUNCTS: thm -> thm list val DISCH: cterm -> thm -> thm val UNDISCH: thm -> thm val SPEC: cterm -> thm -> thm val ISPEC: cterm -> thm -> thm val ISPECL: cterm list -> thm -> thm val GEN: Proof.context -> cterm -> thm -> thm val GENL: Proof.context -> cterm list -> thm -> thm val LIST_CONJ: thm list -> thm val SYM: thm -> thm val DISCH_ALL: thm -> thm val FILTER_DISCH_ALL: (term -> bool) -> thm -> thm val SPEC_ALL: thm -> thm val GEN_ALL: Proof.context -> thm -> thm val IMP_TRANS: thm -> thm -> thm val PROVE_HYP: thm -> thm -> thm val CHOOSE: Proof.context -> cterm * thm -> thm -> thm val EXISTS: Proof.context -> cterm * cterm -> thm -> thm val IT_EXISTS: Proof.context -> (cterm * cterm) list -> thm -> thm val EVEN_ORS: thm list -> thm list val DISJ_CASESL: thm -> thm list -> thm val list_beta_conv: cterm -> cterm list -> thm val SUBS: Proof.context -> thm list -> thm -> thm val simpl_conv: Proof.context -> thm list -> cterm -> thm val rbeta: thm -> thm val tracing: bool Unsynchronized.ref val CONTEXT_REWRITE_RULE: Proof.context -> term * term list * thm * thm list -> thm -> thm * term list val RIGHT_ASSOC: Proof.context -> thm -> thm val prove: Proof.context -> bool -> term * tactic -> thm end; signature THRY = sig val match_term: theory -> term -> term -> (term * term) list * (typ * typ) list val match_type: theory -> typ -> typ -> (typ * typ) list val typecheck: theory -> term -> cterm (*datatype facts of various flavours*) val match_info: theory -> string -> {constructors: term list, case_const: term} option val induct_info: theory -> string -> {constructors: term list, nchotomy: thm} option val extract_info: theory -> {case_congs: thm list, case_rewrites: thm list} end; signature PRIM = sig val trace: bool Unsynchronized.ref val trace_thms: Proof.context -> string -> thm list -> unit val trace_cterm: Proof.context -> string -> cterm -> unit type pattern val mk_functional: theory -> term list -> {functional: term, pats: pattern list} val wfrec_definition0: string -> term -> term -> theory -> thm * theory val post_definition: Proof.context -> thm list -> thm * pattern list -> {rules: thm, rows: int list, TCs: term list list, full_pats_TCs: (term * term list) list} val mk_induction: Proof.context -> {fconst: term, R: term, SV: term list, pat_TCs_list: (term * term list) list} -> thm val postprocess: Proof.context -> bool -> {wf_tac: tactic, terminator: tactic, simplifier: cterm -> thm} -> {rules: thm, induction: thm, TCs: term list list} -> {rules: thm, induction: thm, nested_tcs: thm list} end; signature TFL = sig val define_i: bool -> thm list -> thm list -> xstring -> term -> term list -> Proof.context -> {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context val define: bool -> thm list -> thm list -> xstring -> string -> string list -> Proof.context -> {lhs: term, rules: (thm * int) list, induct: thm, tcs: term list} * Proof.context end; signature OLD_RECDEF = sig val get_recdef: theory -> string -> {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} option val get_hints: Proof.context -> {simps: thm list, congs: (string * thm) list, wfs: thm list} val simp_add: attribute val simp_del: attribute val cong_add: attribute val cong_del: attribute val wf_add: attribute val wf_del: attribute val add_recdef: bool -> xstring -> string -> ((binding * string) * Token.src list) list -> Token.src option -> theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} val add_recdef_i: bool -> xstring -> term -> ((binding * term) * attribute list) list -> theory -> theory * {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list} end; structure Old_Recdef: OLD_RECDEF = struct (*** extra case splitting for TFL ***) structure CaseSplit: CASE_SPLIT = struct (* make a casethm from an induction thm *) fun cases_thm_of_induct_thm ctxt = Seq.hd o (ALLGOALS (fn i => REPEAT (eresolve_tac ctxt [Drule.thin_rl] i))); (* get the case_thm (my version) from a type *) fun case_thm_of_ty ctxt ty = let val thy = Proof_Context.theory_of ctxt val ty_str = case ty of Type(ty_str, _) => ty_str | TFree(s,_) => error ("Free type: " ^ s) | TVar((s,_),_) => error ("Free variable: " ^ s) val {induct, ...} = BNF_LFP_Compat.the_info thy [BNF_LFP_Compat.Keep_Nesting] ty_str in cases_thm_of_induct_thm ctxt induct end; (* for use when there are no prems to the subgoal *) (* does a case split on the given variable *) fun mk_casesplit_goal_thm ctxt (vstr,ty) gt = let val thy = Proof_Context.theory_of ctxt; val x = Free(vstr,ty); val abst = Abs(vstr, ty, Term.abstract_over (x, gt)); val case_thm = case_thm_of_ty ctxt ty; val abs_ct = Thm.cterm_of ctxt abst; val free_ct = Thm.cterm_of ctxt x; val (Pv, Dv, type_insts) = case (Thm.concl_of case_thm) of (_ $ (Pv $ (Dv as Var(_, Dty)))) => (Pv, Dv, Sign.typ_match thy (Dty, ty) Vartab.empty) | _ => error "not a valid case thm"; val type_cinsts = map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T)) (Vartab.dest type_insts); val Pv = dest_Var (Envir.subst_term_types type_insts Pv); val Dv = dest_Var (Envir.subst_term_types type_insts Dv); in Conv.fconv_rule Drule.beta_eta_conversion (case_thm - |> Thm.instantiate (type_cinsts, []) - |> Thm.instantiate ([], [(Pv, abs_ct), (Dv, free_ct)])) + |> Thm.instantiate (TVars.make type_cinsts, Vars.empty) + |> Thm.instantiate (TVars.empty, Vars.make [(Pv, abs_ct), (Dv, free_ct)])) end; (* the find_XXX_split functions are simply doing a lightwieght (I think) term matching equivalent to find where to do the next split *) (* assuming two twems are identical except for a free in one at a subterm, or constant in another, ie assume that one term is a plit of another, then gives back the free variable that has been split. *) exception find_split_exp of string fun find_term_split (Free v, _ $ _) = SOME v | find_term_split (Free v, Const _) = SOME v | find_term_split (Free v, Abs _) = SOME v (* do we really want this case? *) | find_term_split (Free _, Var _) = NONE (* keep searching *) | find_term_split (a $ b, a2 $ b2) = (case find_term_split (a, a2) of NONE => find_term_split (b,b2) | vopt => vopt) | find_term_split (Abs(_,_,t1), Abs(_,_,t2)) = find_term_split (t1, t2) | find_term_split (Const (x,_), Const(x2,_)) = if x = x2 then NONE else (* keep searching *) raise find_split_exp (* stop now *) "Terms are not identical upto a free varaible! (Consts)" | find_term_split (Bound i, Bound j) = if i = j then NONE else (* keep searching *) raise find_split_exp (* stop now *) "Terms are not identical upto a free varaible! (Bound)" | find_term_split _ = raise find_split_exp (* stop now *) "Terms are not identical upto a free varaible! (Other)"; (* assume that "splitth" is a case split form of subgoal i of "genth", then look for a free variable to split, breaking the subgoal closer to splitth. *) fun find_thm_split splitth i genth = find_term_split (Logic.get_goal (Thm.prop_of genth) i, Thm.concl_of splitth) handle find_split_exp _ => NONE; (* as above but searches "splitths" for a theorem that suggest a case split *) fun find_thms_split splitths i genth = Library.get_first (fn sth => find_thm_split sth i genth) splitths; (* split the subgoal i of "genth" until we get to a member of splitths. Assumes that genth will be a general form of splitths, that can be case-split, as needed. Otherwise fails. Note: We assume that all of "splitths" are split to the same level, and thus it doesn't matter which one we choose to look for the next split. Simply add search on splitthms and split variable, to change this. *) (* Note: possible efficiency measure: when a case theorem is no longer useful, drop it? *) (* Note: This should not be a separate tactic but integrated into the case split done during recdef's case analysis, this would avoid us having to (re)search for variables to split. *) fun splitto ctxt splitths genth = let val _ = not (null splitths) orelse error "splitto: no given splitths"; (* check if we are a member of splitths - FIXME: quicker and more flexible with discrim net. *) fun solve_by_splitth th split = Thm.biresolution (SOME ctxt) false [(false,split)] 1 th; fun split th = (case find_thms_split splitths 1 th of NONE => (writeln (cat_lines (["th:", Thm.string_of_thm ctxt th, "split ths:"] @ map (Thm.string_of_thm ctxt) splitths @ ["\n--"])); error "splitto: cannot find variable to split on") | SOME v => let val gt = HOLogic.dest_Trueprop (#1 (Logic.dest_implies (Thm.prop_of th))); val split_thm = mk_casesplit_goal_thm ctxt v gt; val (subthms, expf) = IsaND.fixed_subgoal_thms ctxt split_thm; in expf (map recsplitf subthms) end) and recsplitf th = (* note: multiple unifiers! we only take the first element, probably fine -- there is probably only one anyway. *) (case get_first (Seq.pull o solve_by_splitth th) splitths of NONE => split th | SOME (solved_th, _) => solved_th); in recsplitf genth end; end; (*** basic utilities ***) structure Utils: UTILS = struct (*standard exception for TFL*) exception ERR of {module: string, func: string, mesg: string}; fun UTILS_ERR func mesg = ERR {module = "Utils", func = func, mesg = mesg}; fun end_itlist _ [] = raise (UTILS_ERR "end_itlist" "list too short") | end_itlist _ [x] = x | end_itlist f (x :: xs) = f x (end_itlist f xs); fun itlist2 f L1 L2 base_value = let fun it ([],[]) = base_value | it ((a::rst1),(b::rst2)) = f a b (it (rst1,rst2)) | it _ = raise UTILS_ERR "itlist2" "different length lists" in it (L1,L2) end; fun pluck p = let fun remv ([],_) = raise UTILS_ERR "pluck" "item not found" | remv (h::t, A) = if p h then (h, rev A @ t) else remv (t,h::A) in fn L => remv(L,[]) end; fun take f = let fun grab(0, _) = [] | grab(n, x::rst) = f x::grab(n-1,rst) in grab end; fun zip3 [][][] = [] | zip3 (x::l1) (y::l2) (z::l3) = (x,y,z)::zip3 l1 l2 l3 | zip3 _ _ _ = raise UTILS_ERR "zip3" "different lengths"; end; (*** emulation of HOL's abstract syntax functions ***) structure USyntax: USYNTAX = struct infix 4 ##; fun USYN_ERR func mesg = Utils.ERR {module = "USyntax", func = func, mesg = mesg}; (*--------------------------------------------------------------------------- * * Types * *---------------------------------------------------------------------------*) val mk_prim_vartype = TVar; fun mk_vartype s = mk_prim_vartype ((s, 0), \<^sort>\type\); (* But internally, it's useful *) fun dest_vtype (TVar x) = x | dest_vtype _ = raise USYN_ERR "dest_vtype" "not a flexible type variable"; val is_vartype = can dest_vtype; val type_vars = map mk_prim_vartype o Misc_Legacy.typ_tvars fun type_varsl L = distinct (op =) (fold (curry op @ o type_vars) L []); val alpha = mk_vartype "'a" val strip_prod_type = HOLogic.flatten_tupleT; (*--------------------------------------------------------------------------- * * Terms * *---------------------------------------------------------------------------*) (* Free variables, in order of occurrence, from left to right in the * syntax tree. *) fun free_vars_lr tm = let fun memb x = let fun m[] = false | m(y::rst) = (x=y)orelse m rst in m end fun add (t, frees) = case t of Free _ => if (memb t frees) then frees else t::frees | Abs (_,_,body) => add(body,frees) | f$t => add(t, add(f, frees)) | _ => frees in rev(add(tm,[])) end; val type_vars_in_term = map mk_prim_vartype o Misc_Legacy.term_tvars; (* Prelogic *) fun dest_tybinding (v,ty) = (#1(dest_vtype v),ty) fun inst theta = subst_vars (map dest_tybinding theta,[]) (* Construction routines *) fun mk_abs{Bvar as Var((s,_),ty),Body} = Abs(s,ty,abstract_over(Bvar,Body)) | mk_abs{Bvar as Free(s,ty),Body} = Abs(s,ty,abstract_over(Bvar,Body)) | mk_abs _ = raise USYN_ERR "mk_abs" "Bvar is not a variable"; fun mk_imp{ant,conseq} = let val c = Const(\<^const_name>\HOL.implies\,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) in list_comb(c,[ant,conseq]) end; fun mk_select (r as {Bvar,Body}) = let val ty = type_of Bvar val c = Const(\<^const_name>\Eps\,(ty --> HOLogic.boolT) --> ty) in list_comb(c,[mk_abs r]) end; fun mk_forall (r as {Bvar,Body}) = let val ty = type_of Bvar val c = Const(\<^const_name>\All\,(ty --> HOLogic.boolT) --> HOLogic.boolT) in list_comb(c,[mk_abs r]) end; fun mk_exists (r as {Bvar,Body}) = let val ty = type_of Bvar val c = Const(\<^const_name>\Ex\,(ty --> HOLogic.boolT) --> HOLogic.boolT) in list_comb(c,[mk_abs r]) end; fun mk_conj{conj1,conj2} = let val c = Const(\<^const_name>\HOL.conj\,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) in list_comb(c,[conj1,conj2]) end; fun mk_disj{disj1,disj2} = let val c = Const(\<^const_name>\HOL.disj\,HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) in list_comb(c,[disj1,disj2]) end; fun prod_ty ty1 ty2 = HOLogic.mk_prodT (ty1,ty2); local fun mk_uncurry (xt, yt, zt) = Const(\<^const_name>\case_prod\, (xt --> yt --> zt) --> prod_ty xt yt --> zt) fun dest_pair(Const(\<^const_name>\Pair\,_) $ M $ N) = {fst=M, snd=N} | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair" fun is_var (Var _) = true | is_var (Free _) = true | is_var _ = false in fun mk_pabs{varstruct,body} = let fun mpa (varstruct, body) = if is_var varstruct then mk_abs {Bvar = varstruct, Body = body} else let val {fst, snd} = dest_pair varstruct in mk_uncurry (type_of fst, type_of snd, type_of body) $ mpa (fst, mpa (snd, body)) end in mpa (varstruct, body) end handle TYPE _ => raise USYN_ERR "mk_pabs" ""; end; (* Destruction routines *) datatype lambda = VAR of {Name : string, Ty : typ} | CONST of {Name : string, Ty : typ} | COMB of {Rator: term, Rand : term} | LAMB of {Bvar : term, Body : term}; fun dest_term(Var((s,_),ty)) = VAR{Name = s, Ty = ty} | dest_term(Free(s,ty)) = VAR{Name = s, Ty = ty} | dest_term(Const(s,ty)) = CONST{Name = s, Ty = ty} | dest_term(M$N) = COMB{Rator=M,Rand=N} | dest_term(Abs(s,ty,M)) = let val v = Free(s,ty) in LAMB{Bvar = v, Body = Term.betapply (M,v)} end | dest_term(Bound _) = raise USYN_ERR "dest_term" "Bound"; fun dest_const(Const(s,ty)) = {Name = s, Ty = ty} | dest_const _ = raise USYN_ERR "dest_const" "not a constant"; fun dest_comb(t1 $ t2) = {Rator = t1, Rand = t2} | dest_comb _ = raise USYN_ERR "dest_comb" "not a comb"; fun dest_abs used (a as Abs(s, ty, _)) = let val s' = singleton (Name.variant_list used) s; val v = Free(s', ty); in ({Bvar = v, Body = Term.betapply (a,v)}, s'::used) end | dest_abs _ _ = raise USYN_ERR "dest_abs" "not an abstraction"; fun dest_eq(Const(\<^const_name>\HOL.eq\,_) $ M $ N) = {lhs=M, rhs=N} | dest_eq _ = raise USYN_ERR "dest_eq" "not an equality"; fun dest_imp(Const(\<^const_name>\HOL.implies\,_) $ M $ N) = {ant=M, conseq=N} | dest_imp _ = raise USYN_ERR "dest_imp" "not an implication"; fun dest_forall(Const(\<^const_name>\All\,_) $ (a as Abs _)) = fst (dest_abs [] a) | dest_forall _ = raise USYN_ERR "dest_forall" "not a forall"; fun dest_exists(Const(\<^const_name>\Ex\,_) $ (a as Abs _)) = fst (dest_abs [] a) | dest_exists _ = raise USYN_ERR "dest_exists" "not an existential"; fun dest_neg(Const(\<^const_name>\Not\,_) $ M) = M | dest_neg _ = raise USYN_ERR "dest_neg" "not a negation"; fun dest_conj(Const(\<^const_name>\HOL.conj\,_) $ M $ N) = {conj1=M, conj2=N} | dest_conj _ = raise USYN_ERR "dest_conj" "not a conjunction"; fun dest_disj(Const(\<^const_name>\HOL.disj\,_) $ M $ N) = {disj1=M, disj2=N} | dest_disj _ = raise USYN_ERR "dest_disj" "not a disjunction"; fun mk_pair{fst,snd} = let val ty1 = type_of fst val ty2 = type_of snd val c = Const(\<^const_name>\Pair\,ty1 --> ty2 --> prod_ty ty1 ty2) in list_comb(c,[fst,snd]) end; fun dest_pair(Const(\<^const_name>\Pair\,_) $ M $ N) = {fst=M, snd=N} | dest_pair _ = raise USYN_ERR "dest_pair" "not a pair"; local fun ucheck t = (if #Name (dest_const t) = \<^const_name>\case_prod\ then t else raise Match) in fun dest_pabs used tm = let val ({Bvar,Body}, used') = dest_abs used tm in {varstruct = Bvar, body = Body, used = used'} end handle Utils.ERR _ => let val {Rator,Rand} = dest_comb tm val _ = ucheck Rator val {varstruct = lv, body, used = used'} = dest_pabs used Rand val {varstruct = rv, body, used = used''} = dest_pabs used' body in {varstruct = mk_pair {fst = lv, snd = rv}, body = body, used = used''} end end; val lhs = #lhs o dest_eq val rhs = #rhs o dest_eq val rand = #Rand o dest_comb (* Query routines *) val is_imp = can dest_imp val is_forall = can dest_forall val is_exists = can dest_exists val is_neg = can dest_neg val is_conj = can dest_conj val is_disj = can dest_disj val is_pair = can dest_pair val is_pabs = can (dest_pabs []) (* Construction of a cterm from a list of Terms *) fun list_mk_abs(L,tm) = fold_rev (fn v => fn M => mk_abs{Bvar=v, Body=M}) L tm; (* These others are almost never used *) fun list_mk_imp(A,c) = fold_rev (fn a => fn tm => mk_imp{ant=a,conseq=tm}) A c; fun list_mk_forall(V,t) = fold_rev (fn v => fn b => mk_forall{Bvar=v, Body=b})V t; val list_mk_conj = Utils.end_itlist(fn c1 => fn tm => mk_conj{conj1=c1, conj2=tm}) (* Need to reverse? *) fun gen_all tm = list_mk_forall(Misc_Legacy.term_frees tm, tm); (* Destructing a cterm to a list of Terms *) fun strip_comb tm = let fun dest(M$N, A) = dest(M, N::A) | dest x = x in dest(tm,[]) end; fun strip_abs(tm as Abs _) = let val ({Bvar,Body}, _) = dest_abs [] tm val (bvs, core) = strip_abs Body in (Bvar::bvs, core) end | strip_abs M = ([],M); fun strip_imp fm = if (is_imp fm) then let val {ant,conseq} = dest_imp fm val (was,wb) = strip_imp conseq in ((ant::was), wb) end else ([],fm); fun strip_forall fm = if (is_forall fm) then let val {Bvar,Body} = dest_forall fm val (bvs,core) = strip_forall Body in ((Bvar::bvs), core) end else ([],fm); fun strip_exists fm = if (is_exists fm) then let val {Bvar, Body} = dest_exists fm val (bvs,core) = strip_exists Body in (Bvar::bvs, core) end else ([],fm); fun strip_disj w = if (is_disj w) then let val {disj1,disj2} = dest_disj w in (strip_disj disj1@strip_disj disj2) end else [w]; (* Miscellaneous *) fun mk_vstruct ty V = let fun follow_prod_type (Type(\<^type_name>\Product_Type.prod\,[ty1,ty2])) vs = let val (ltm,vs1) = follow_prod_type ty1 vs val (rtm,vs2) = follow_prod_type ty2 vs1 in (mk_pair{fst=ltm, snd=rtm}, vs2) end | follow_prod_type _ (v::vs) = (v,vs) in #1 (follow_prod_type ty V) end; (* Search a term for a sub-term satisfying the predicate p. *) fun find_term p = let fun find tm = if (p tm) then SOME tm else case tm of Abs(_,_,body) => find body | (t$u) => (case find t of NONE => find u | some => some) | _ => NONE in find end; fun dest_relation tm = if (type_of tm = HOLogic.boolT) then let val (Const(\<^const_name>\Set.member\,_) $ (Const(\<^const_name>\Pair\,_)$y$x) $ R) = tm in (R,y,x) end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure" else raise USYN_ERR "dest_relation" "not a boolean term"; fun is_WFR (Const(\<^const_name>\Wellfounded.wf\,_)$_) = true | is_WFR _ = false; fun ARB ty = mk_select{Bvar=Free("v",ty), Body=Const(\<^const_name>\True\,HOLogic.boolT)}; end; (*** derived cterm destructors ***) structure Dcterm: DCTERM = struct fun ERR func mesg = Utils.ERR {module = "Dcterm", func = func, mesg = mesg}; fun dest_comb t = Thm.dest_comb t handle CTERM (msg, _) => raise ERR "dest_comb" msg; fun dest_abs a t = Thm.dest_abs a t handle CTERM (msg, _) => raise ERR "dest_abs" msg; fun capply t u = Thm.apply t u handle CTERM (msg, _) => raise ERR "capply" msg; fun cabs a t = Thm.lambda a t handle CTERM (msg, _) => raise ERR "cabs" msg; (*--------------------------------------------------------------------------- * Some simple constructor functions. *---------------------------------------------------------------------------*) val mk_hol_const = Thm.cterm_of \<^theory_context>\HOL\ o Const; fun mk_exists (r as (Bvar, Body)) = let val ty = Thm.typ_of_cterm Bvar val c = mk_hol_const(\<^const_name>\Ex\, (ty --> HOLogic.boolT) --> HOLogic.boolT) in capply c (uncurry cabs r) end; local val c = mk_hol_const(\<^const_name>\HOL.conj\, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) in fun mk_conj(conj1,conj2) = capply (capply c conj1) conj2 end; local val c = mk_hol_const(\<^const_name>\HOL.disj\, HOLogic.boolT --> HOLogic.boolT --> HOLogic.boolT) in fun mk_disj(disj1,disj2) = capply (capply c disj1) disj2 end; (*--------------------------------------------------------------------------- * The primitives. *---------------------------------------------------------------------------*) fun dest_const ctm = (case Thm.term_of ctm of Const(s,ty) => {Name = s, Ty = ty} | _ => raise ERR "dest_const" "not a constant"); fun dest_var ctm = (case Thm.term_of ctm of Var((s,_),ty) => {Name=s, Ty=ty} | Free(s,ty) => {Name=s, Ty=ty} | _ => raise ERR "dest_var" "not a variable"); (*--------------------------------------------------------------------------- * Derived destructor operations. *---------------------------------------------------------------------------*) fun dest_monop expected tm = let fun err () = raise ERR "dest_monop" ("Not a(n) " ^ quote expected); val (c, N) = dest_comb tm handle Utils.ERR _ => err (); val name = #Name (dest_const c handle Utils.ERR _ => err ()); in if name = expected then N else err () end; fun dest_binop expected tm = let fun err () = raise ERR "dest_binop" ("Not a(n) " ^ quote expected); val (M, N) = dest_comb tm handle Utils.ERR _ => err () in (dest_monop expected M, N) handle Utils.ERR _ => err () end; fun dest_binder expected tm = dest_abs NONE (dest_monop expected tm) handle Utils.ERR _ => raise ERR "dest_binder" ("Not a(n) " ^ quote expected); val dest_neg = dest_monop \<^const_name>\Not\ val dest_pair = dest_binop \<^const_name>\Pair\ val dest_eq = dest_binop \<^const_name>\HOL.eq\ val dest_imp = dest_binop \<^const_name>\HOL.implies\ val dest_conj = dest_binop \<^const_name>\HOL.conj\ val dest_disj = dest_binop \<^const_name>\HOL.disj\ val dest_exists = dest_binder \<^const_name>\Ex\ val dest_forall = dest_binder \<^const_name>\All\ (* Query routines *) val is_eq = can dest_eq val is_imp = can dest_imp val is_forall = can dest_forall val is_exists = can dest_exists val is_neg = can dest_neg val is_conj = can dest_conj val is_disj = can dest_disj val is_pair = can dest_pair (*--------------------------------------------------------------------------- * Iterated creation. *---------------------------------------------------------------------------*) val list_mk_disj = Utils.end_itlist (fn d1 => fn tm => mk_disj (d1, tm)); (*--------------------------------------------------------------------------- * Iterated destruction. (To the "right" in a term.) *---------------------------------------------------------------------------*) fun strip break tm = let fun dest (p as (ctm,accum)) = let val (M,N) = break ctm in dest (N, M::accum) end handle Utils.ERR _ => p in dest (tm,[]) end; fun rev2swap (x,l) = (rev l, x); val strip_comb = strip (Library.swap o dest_comb) (* Goes to the "left" *) val strip_imp = rev2swap o strip dest_imp val strip_abs = rev2swap o strip (dest_abs NONE) val strip_forall = rev2swap o strip dest_forall val strip_exists = rev2swap o strip dest_exists val strip_disj = rev o (op::) o strip dest_disj (*--------------------------------------------------------------------------- * Going into and out of prop *---------------------------------------------------------------------------*) fun is_Trueprop ct = (case Thm.term_of ct of Const (\<^const_name>\Trueprop\, _) $ _ => true | _ => false); fun mk_prop ct = if is_Trueprop ct then ct else Thm.apply \<^cterm>\Trueprop\ ct; fun drop_prop ct = if is_Trueprop ct then Thm.dest_arg ct else ct; end; (*** emulation of HOL inference rules for TFL ***) structure Rules: RULES = struct fun RULES_ERR func mesg = Utils.ERR {module = "Rules", func = func, mesg = mesg}; fun cconcl thm = Dcterm.drop_prop (Thm.cprop_of thm); fun chyps thm = map Dcterm.drop_prop (Thm.chyps_of thm); fun dest_thm thm = (map HOLogic.dest_Trueprop (Thm.hyps_of thm), HOLogic.dest_Trueprop (Thm.prop_of thm)) handle TERM _ => raise RULES_ERR "dest_thm" "missing Trueprop"; (* Inference rules *) (*--------------------------------------------------------------------------- * Equality (one step) *---------------------------------------------------------------------------*) fun REFL tm = HOLogic.mk_obj_eq (Thm.reflexive tm) handle THM (msg, _, _) => raise RULES_ERR "REFL" msg; fun SYM thm = thm RS sym handle THM (msg, _, _) => raise RULES_ERR "SYM" msg; fun ALPHA thm ctm1 = let val ctm2 = Thm.cprop_of thm; val ctm2_eq = Thm.reflexive ctm2; val ctm1_eq = Thm.reflexive ctm1; in Thm.equal_elim (Thm.transitive ctm2_eq ctm1_eq) thm end handle THM (msg, _, _) => raise RULES_ERR "ALPHA" msg; fun rbeta th = (case Dcterm.strip_comb (cconcl th) of (_, [_, r]) => Thm.transitive th (Thm.beta_conversion false r) | _ => raise RULES_ERR "rbeta" ""); (*---------------------------------------------------------------------------- * Implication and the assumption list * * Assumptions get stuck on the meta-language assumption list. Implications * are in the object language, so discharging an assumption "A" from theorem * "B" results in something that looks like "A --> B". *---------------------------------------------------------------------------*) fun ASSUME ctm = Thm.assume (Dcterm.mk_prop ctm); (*--------------------------------------------------------------------------- * Implication in TFL is -->. Meta-language implication (==>) is only used * in the implementation of some of the inference rules below. *---------------------------------------------------------------------------*) fun MP th1 th2 = th2 RS (th1 RS mp) handle THM (msg, _, _) => raise RULES_ERR "MP" msg; (*forces the first argument to be a proposition if necessary*) fun DISCH tm thm = Thm.implies_intr (Dcterm.mk_prop tm) thm COMP impI handle THM (msg, _, _) => raise RULES_ERR "DISCH" msg; fun DISCH_ALL thm = fold_rev DISCH (Thm.chyps_of thm) thm; fun FILTER_DISCH_ALL P thm = let fun check tm = P (Thm.term_of tm) in fold_rev (fn tm => fn th => if check tm then DISCH tm th else th) (chyps thm) thm end; fun UNDISCH thm = let val tm = Dcterm.mk_prop (#1 (Dcterm.dest_imp (cconcl thm))) in Thm.implies_elim (thm RS mp) (ASSUME tm) end handle Utils.ERR _ => raise RULES_ERR "UNDISCH" "" | THM _ => raise RULES_ERR "UNDISCH" ""; fun PROVE_HYP ath bth = MP (DISCH (cconcl ath) bth) ath; fun IMP_TRANS th1 th2 = th2 RS (th1 RS @{thm tfl_imp_trans}) handle THM (msg, _, _) => raise RULES_ERR "IMP_TRANS" msg; (*---------------------------------------------------------------------------- * Conjunction *---------------------------------------------------------------------------*) fun CONJUNCT1 thm = thm RS conjunct1 handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT1" msg; fun CONJUNCT2 thm = thm RS conjunct2 handle THM (msg, _, _) => raise RULES_ERR "CONJUNCT2" msg; fun CONJUNCTS th = CONJUNCTS (CONJUNCT1 th) @ CONJUNCTS (CONJUNCT2 th) handle Utils.ERR _ => [th]; fun LIST_CONJ [] = raise RULES_ERR "LIST_CONJ" "empty list" | LIST_CONJ [th] = th | LIST_CONJ (th :: rst) = MP (MP (conjI COMP (impI RS impI)) th) (LIST_CONJ rst) handle THM (msg, _, _) => raise RULES_ERR "LIST_CONJ" msg; (*---------------------------------------------------------------------------- * Disjunction *---------------------------------------------------------------------------*) local val prop = Thm.prop_of disjI1 val [_,Q] = Misc_Legacy.term_vars prop val disj1 = Thm.forall_intr (Thm.cterm_of \<^context> Q) disjI1 in fun DISJ1 thm tm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj1) handle THM (msg, _, _) => raise RULES_ERR "DISJ1" msg; end; local val prop = Thm.prop_of disjI2 val [P,_] = Misc_Legacy.term_vars prop val disj2 = Thm.forall_intr (Thm.cterm_of \<^context> P) disjI2 in fun DISJ2 tm thm = thm RS (Thm.forall_elim (Dcterm.drop_prop tm) disj2) handle THM (msg, _, _) => raise RULES_ERR "DISJ2" msg; end; (*---------------------------------------------------------------------------- * * A1 |- M1, ..., An |- Mn * --------------------------------------------------- * [A1 |- M1 \/ ... \/ Mn, ..., An |- M1 \/ ... \/ Mn] * *---------------------------------------------------------------------------*) fun EVEN_ORS thms = let fun blue ldisjs [] _ = [] | blue ldisjs (th::rst) rdisjs = let val tail = tl rdisjs val rdisj_tl = Dcterm.list_mk_disj tail in fold_rev DISJ2 ldisjs (DISJ1 th rdisj_tl) :: blue (ldisjs @ [cconcl th]) rst tail end handle Utils.ERR _ => [fold_rev DISJ2 ldisjs th] in blue [] thms (map cconcl thms) end; (*---------------------------------------------------------------------------- * * A |- P \/ Q B,P |- R C,Q |- R * --------------------------------------------------- * A U B U C |- R * *---------------------------------------------------------------------------*) fun DISJ_CASES th1 th2 th3 = let val c = Dcterm.drop_prop (cconcl th1); val (disj1, disj2) = Dcterm.dest_disj c; val th2' = DISCH disj1 th2; val th3' = DISCH disj2 th3; in th3' RS (th2' RS (th1 RS @{thm tfl_disjE})) handle THM (msg, _, _) => raise RULES_ERR "DISJ_CASES" msg end; (*----------------------------------------------------------------------------- * * |- A1 \/ ... \/ An [A1 |- M, ..., An |- M] * --------------------------------------------------- * |- M * * Note. The list of theorems may be all jumbled up, so we have to * first organize it to align with the first argument (the disjunctive * theorem). *---------------------------------------------------------------------------*) fun organize eq = (* a bit slow - analogous to insertion sort *) let fun extract a alist = let fun ex (_,[]) = raise RULES_ERR "organize" "not a permutation.1" | ex(left,h::t) = if (eq h a) then (h,rev left@t) else ex(h::left,t) in ex ([],alist) end fun place [] [] = [] | place (a::rst) alist = let val (item,next) = extract a alist in item::place rst next end | place _ _ = raise RULES_ERR "organize" "not a permutation.2" in place end; fun DISJ_CASESL disjth thl = let val c = cconcl disjth fun eq th atm = exists (fn t => HOLogic.dest_Trueprop t aconv Thm.term_of atm) (Thm.hyps_of th) val tml = Dcterm.strip_disj c fun DL _ [] = raise RULES_ERR "DISJ_CASESL" "no cases" | DL th [th1] = PROVE_HYP th th1 | DL th [th1,th2] = DISJ_CASES th th1 th2 | DL th (th1::rst) = let val tm = #2 (Dcterm.dest_disj (Dcterm.drop_prop(cconcl th))) in DISJ_CASES th th1 (DL (ASSUME tm) rst) end in DL disjth (organize eq tml thl) end; (*---------------------------------------------------------------------------- * Universals *---------------------------------------------------------------------------*) local (* this is fragile *) val prop = Thm.prop_of spec val x = hd (tl (Misc_Legacy.term_vars prop)) val TV = dest_TVar (type_of x) val gspec = Thm.forall_intr (Thm.cterm_of \<^context> x) spec in fun SPEC tm thm = - let val gspec' = Drule.instantiate_normalize ([(TV, Thm.ctyp_of_cterm tm)], []) gspec + let val gspec' = + Drule.instantiate_normalize (TVars.make [(TV, Thm.ctyp_of_cterm tm)], Vars.empty) gspec in thm RS (Thm.forall_elim tm gspec') end end; fun SPEC_ALL thm = fold SPEC (#1 (Dcterm.strip_forall(cconcl thm))) thm; val ISPEC = SPEC val ISPECL = fold ISPEC; (* Not optimized! Too complicated. *) local val prop = Thm.prop_of allI val [P] = Misc_Legacy.add_term_vars (prop, []) fun cty_theta ctxt = map (fn (i, (S, ty)) => ((i, S), Thm.ctyp_of ctxt ty)) fun ctm_theta ctxt = map (fn (i, (_, tm2)) => let val ctm2 = Thm.cterm_of ctxt tm2 in ((i, Thm.typ_of_cterm ctm2), ctm2) end) fun certify ctxt (ty_theta,tm_theta) = - (cty_theta ctxt (Vartab.dest ty_theta), - ctm_theta ctxt (Vartab.dest tm_theta)) + (TVars.make (cty_theta ctxt (Vartab.dest ty_theta)), + Vars.make (ctm_theta ctxt (Vartab.dest tm_theta))) in fun GEN ctxt v th = let val gth = Thm.forall_intr v th val thy = Proof_Context.theory_of ctxt val Const(\<^const_name>\Pure.all\,_)$Abs(x,ty,rst) = Thm.prop_of gth val P' = Abs(x,ty, HOLogic.dest_Trueprop rst) (* get rid of trueprop *) val theta = Pattern.match thy (P,P') (Vartab.empty, Vartab.empty); val allI2 = Drule.instantiate_normalize (certify ctxt theta) allI val thm = Thm.implies_elim allI2 gth val tp $ (A $ Abs(_,_,M)) = Thm.prop_of thm val prop' = tp $ (A $ Abs(x,ty,M)) in ALPHA thm (Thm.cterm_of ctxt prop') end end; fun GENL ctxt = fold_rev (GEN ctxt); fun GEN_ALL ctxt thm = let val prop = Thm.prop_of thm val vlist = map (Thm.cterm_of ctxt) (Misc_Legacy.add_term_vars (prop, [])) in GENL ctxt vlist thm end; fun MATCH_MP th1 th2 = if (Dcterm.is_forall (Dcterm.drop_prop(cconcl th1))) then MATCH_MP (th1 RS spec) th2 else MP th1 th2; (*---------------------------------------------------------------------------- * Existentials *---------------------------------------------------------------------------*) (*--------------------------------------------------------------------------- * Existential elimination * * A1 |- ?x.t[x] , A2, "t[v]" |- t' * ------------------------------------ (variable v occurs nowhere) * A1 u A2 |- t' * *---------------------------------------------------------------------------*) fun CHOOSE ctxt (fvar, exth) fact = let val lam = #2 (Dcterm.dest_comb (Dcterm.drop_prop (cconcl exth))) val redex = Dcterm.capply lam fvar val t$u = Thm.term_of redex val residue = Thm.cterm_of ctxt (Term.betapply (t, u)) in GEN ctxt fvar (DISCH residue fact) RS (exth RS @{thm tfl_exE}) handle THM (msg, _, _) => raise RULES_ERR "CHOOSE" msg end; fun EXISTS ctxt (template,witness) thm = let val abstr = #2 (Dcterm.dest_comb template) in thm RS (infer_instantiate ctxt [(("P", 0), abstr), (("x", 0), witness)] exI) handle THM (msg, _, _) => raise RULES_ERR "EXISTS" msg end; (*---------------------------------------------------------------------------- * * A |- M[x_1,...,x_n] * ---------------------------- [(x |-> y)_1,...,(x |-> y)_n] * A |- ?y_1...y_n. M * *---------------------------------------------------------------------------*) (* Could be improved, but needs "subst_free" for certified terms *) fun IT_EXISTS ctxt blist th = let val blist' = map (apply2 Thm.term_of) blist fun ex v M = Thm.cterm_of ctxt (USyntax.mk_exists{Bvar=v,Body = M}) in fold_rev (fn (b as (r1,r2)) => fn thm => EXISTS ctxt (ex r2 (subst_free [b] (HOLogic.dest_Trueprop(Thm.prop_of thm))), Thm.cterm_of ctxt r1) thm) blist' th end; (*---------------------------------------------------------------------------- * Rewriting *---------------------------------------------------------------------------*) fun SUBS ctxt thl = rewrite_rule ctxt (map (fn th => th RS eq_reflection handle THM _ => th) thl); val rew_conv = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)); fun simpl_conv ctxt thl ctm = HOLogic.mk_obj_eq (rew_conv (ctxt addsimps thl) ctm); fun RIGHT_ASSOC ctxt = rewrite_rule ctxt @{thms tfl_disj_assoc}; (*--------------------------------------------------------------------------- * TERMINATION CONDITION EXTRACTION *---------------------------------------------------------------------------*) (* Object language quantifier, i.e., "!" *) fun Forall v M = USyntax.mk_forall{Bvar=v, Body=M}; (* Fragile: it's a cong if it is not "R y x ==> cut f R x y = f y" *) fun is_cong thm = case (Thm.prop_of thm) of (Const(\<^const_name>\Pure.imp\,_)$(Const(\<^const_name>\Trueprop\,_)$ _) $ (Const(\<^const_name>\Pure.eq\,_) $ (Const (\<^const_name>\Wfrec.cut\,_) $ _ $ _ $ _ $ _) $ _)) => false | _ => true; fun dest_equal(Const (\<^const_name>\Pure.eq\,_) $ (Const (\<^const_name>\Trueprop\,_) $ lhs) $ (Const (\<^const_name>\Trueprop\,_) $ rhs)) = {lhs=lhs, rhs=rhs} | dest_equal(Const (\<^const_name>\Pure.eq\,_) $ lhs $ rhs) = {lhs=lhs, rhs=rhs} | dest_equal tm = USyntax.dest_eq tm; fun get_lhs tm = #lhs(dest_equal (HOLogic.dest_Trueprop tm)); fun dest_all used (Const(\<^const_name>\Pure.all\,_) $ (a as Abs _)) = USyntax.dest_abs used a | dest_all _ _ = raise RULES_ERR "dest_all" "not a !!"; val is_all = can (dest_all []); fun strip_all used fm = if (is_all fm) then let val ({Bvar, Body}, used') = dest_all used fm val (bvs, core, used'') = strip_all used' Body in ((Bvar::bvs), core, used'') end else ([], fm, used); fun list_break_all(Const(\<^const_name>\Pure.all\,_) $ Abs (s,ty,body)) = let val (L,core) = list_break_all body in ((s,ty)::L, core) end | list_break_all tm = ([],tm); (*--------------------------------------------------------------------------- * Rename a term of the form * * !!x1 ...xn. x1=M1 ==> ... ==> xn=Mn * ==> ((%v1...vn. Q) x1 ... xn = g x1 ... xn. * to one of * * !!v1 ... vn. v1=M1 ==> ... ==> vn=Mn * ==> ((%v1...vn. Q) v1 ... vn = g v1 ... vn. * * This prevents name problems in extraction, and helps the result to read * better. There is a problem with varstructs, since they can introduce more * than n variables, and some extra reasoning needs to be done. *---------------------------------------------------------------------------*) fun get ([],_,L) = rev L | get (ant::rst,n,L) = case (list_break_all ant) of ([],_) => get (rst, n+1,L) | (_,body) => let val eq = Logic.strip_imp_concl body val (f,_) = USyntax.strip_comb (get_lhs eq) val (vstrl,_) = USyntax.strip_abs f val names = Name.variant_list (Misc_Legacy.add_term_names(body, [])) (map (#1 o dest_Free) vstrl) in get (rst, n+1, (names,n)::L) end handle TERM _ => get (rst, n+1, L) | Utils.ERR _ => get (rst, n+1, L); (* Note: Thm.rename_params_rule counts from 1, not 0 *) fun rename thm = let val ants = Logic.strip_imp_prems (Thm.prop_of thm) val news = get (ants,1,[]) in fold Thm.rename_params_rule news thm end; (*--------------------------------------------------------------------------- * Beta-conversion to the rhs of an equation (taken from hol90/drule.sml) *---------------------------------------------------------------------------*) fun list_beta_conv tm = let fun rbeta th = Thm.transitive th (Thm.beta_conversion false (#2(Dcterm.dest_eq(cconcl th)))) fun iter [] = Thm.reflexive tm | iter (v::rst) = rbeta (Thm.combination(iter rst) (Thm.reflexive v)) in iter end; (*--------------------------------------------------------------------------- * Trace information for the rewriter *---------------------------------------------------------------------------*) val tracing = Unsynchronized.ref false; fun say s = if !tracing then writeln s else (); fun print_thms ctxt s L = say (cat_lines (s :: map (Thm.string_of_thm ctxt) L)); fun print_term ctxt s t = say (cat_lines [s, Syntax.string_of_term ctxt t]); (*--------------------------------------------------------------------------- * General abstraction handlers, should probably go in USyntax. *---------------------------------------------------------------------------*) fun mk_aabs (vstr, body) = USyntax.mk_abs {Bvar = vstr, Body = body} handle Utils.ERR _ => USyntax.mk_pabs {varstruct = vstr, body = body}; fun list_mk_aabs (vstrl,tm) = fold_rev (fn vstr => fn tm => mk_aabs(vstr,tm)) vstrl tm; fun dest_aabs used tm = let val ({Bvar,Body}, used') = USyntax.dest_abs used tm in (Bvar, Body, used') end handle Utils.ERR _ => let val {varstruct, body, used} = USyntax.dest_pabs used tm in (varstruct, body, used) end; fun strip_aabs used tm = let val (vstr, body, used') = dest_aabs used tm val (bvs, core, used'') = strip_aabs used' body in (vstr::bvs, core, used'') end handle Utils.ERR _ => ([], tm, used); fun dest_combn tm 0 = (tm,[]) | dest_combn tm n = let val {Rator,Rand} = USyntax.dest_comb tm val (f,rands) = dest_combn Rator (n-1) in (f,Rand::rands) end; local fun dest_pair M = let val {fst,snd} = USyntax.dest_pair M in (fst,snd) end fun mk_fst tm = let val ty as Type(\<^type_name>\Product_Type.prod\, [fty,sty]) = type_of tm in Const (\<^const_name>\Product_Type.fst\, ty --> fty) $ tm end fun mk_snd tm = let val ty as Type(\<^type_name>\Product_Type.prod\, [fty,sty]) = type_of tm in Const (\<^const_name>\Product_Type.snd\, ty --> sty) $ tm end in fun XFILL tych x vstruct = let fun traverse p xocc L = if (is_Free p) then tych xocc::L else let val (p1,p2) = dest_pair p in traverse p1 (mk_fst xocc) (traverse p2 (mk_snd xocc) L) end in traverse vstruct x [] end end; (*--------------------------------------------------------------------------- * Replace a free tuple (vstr) by a universally quantified variable (a). * Note that the notion of "freeness" for a tuple is different than for a * variable: if variables in the tuple also occur in any other place than * an occurrences of the tuple, they aren't "free" (which is thus probably * the wrong word to use). *---------------------------------------------------------------------------*) fun VSTRUCT_ELIM ctxt tych a vstr th = let val L = USyntax.free_vars_lr vstr val bind1 = tych (HOLogic.mk_Trueprop (HOLogic.mk_eq(a,vstr))) val thm1 = Thm.implies_intr bind1 (SUBS ctxt [SYM(Thm.assume bind1)] th) val thm2 = forall_intr_list (map tych L) thm1 val thm3 = forall_elim_list (XFILL tych a vstr) thm2 in refl RS rewrite_rule ctxt [Thm.symmetric (@{thm surjective_pairing} RS eq_reflection)] thm3 end; fun PGEN ctxt tych a vstr th = let val a1 = tych a in Thm.forall_intr a1 (VSTRUCT_ELIM ctxt tych a vstr th) end; (*--------------------------------------------------------------------------- * Takes apart a paired beta-redex, looking like "(\(x,y).N) vstr", into * * (([x,y],N),vstr) *---------------------------------------------------------------------------*) fun dest_pbeta_redex used M n = let val (f,args) = dest_combn M n val _ = dest_aabs used f in (strip_aabs used f,args) end; fun pbeta_redex M n = can (fn t => dest_pbeta_redex [] t n) M; fun dest_impl tm = let val ants = Logic.strip_imp_prems tm val eq = Logic.strip_imp_concl tm in (ants,get_lhs eq) end; fun restricted t = is_some (USyntax.find_term (fn (Const(\<^const_name>\Wfrec.cut\,_)) =>true | _ => false) t) fun CONTEXT_REWRITE_RULE main_ctxt (func, G, cut_lemma, congs) th = let val globals = func::G val ctxt0 = empty_simpset main_ctxt val pbeta_reduce = simpl_conv ctxt0 [@{thm split_conv} RS eq_reflection]; val tc_list = Unsynchronized.ref []: term list Unsynchronized.ref val cut_lemma' = cut_lemma RS eq_reflection fun prover used ctxt thm = let fun cong_prover ctxt thm = let val _ = say "cong_prover:" val cntxt = Simplifier.prems_of ctxt val _ = print_thms ctxt "cntxt:" cntxt val _ = say "cong rule:" val _ = say (Thm.string_of_thm ctxt thm) (* Unquantified eliminate *) fun uq_eliminate (thm,imp) = let val tych = Thm.cterm_of ctxt val _ = print_term ctxt "To eliminate:" imp val ants = map tych (Logic.strip_imp_prems imp) val eq = Logic.strip_imp_concl imp val lhs = tych(get_lhs eq) val ctxt' = Simplifier.add_prems (map ASSUME ants) ctxt val lhs_eq_lhs1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used) ctxt' lhs handle Utils.ERR _ => Thm.reflexive lhs val _ = print_thms ctxt' "proven:" [lhs_eq_lhs1] val lhs_eq_lhs2 = implies_intr_list ants lhs_eq_lhs1 val lhs_eeq_lhs2 = HOLogic.mk_obj_eq lhs_eq_lhs2 in lhs_eeq_lhs2 COMP thm end fun pq_eliminate (thm, vlist, imp_body, lhs_eq) = let val ((vstrl, _, used'), args) = dest_pbeta_redex used lhs_eq (length vlist) val _ = forall (op aconv) (ListPair.zip (vlist, args)) orelse error "assertion failed in CONTEXT_REWRITE_RULE" val imp_body1 = subst_free (ListPair.zip (args, vstrl)) imp_body val tych = Thm.cterm_of ctxt val ants1 = map tych (Logic.strip_imp_prems imp_body1) val eq1 = Logic.strip_imp_concl imp_body1 val Q = get_lhs eq1 val QeqQ1 = pbeta_reduce (tych Q) val Q1 = #2(Dcterm.dest_eq(cconcl QeqQ1)) val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt val Q1eeqQ2 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' Q1 handle Utils.ERR _ => Thm.reflexive Q1 val Q2 = #2 (Logic.dest_equals (Thm.prop_of Q1eeqQ2)) val Q3 = tych(list_comb(list_mk_aabs(vstrl,Q2),vstrl)) val Q2eeqQ3 = Thm.symmetric(pbeta_reduce Q3 RS eq_reflection) val thA = Thm.transitive(QeqQ1 RS eq_reflection) Q1eeqQ2 val QeeqQ3 = Thm.transitive thA Q2eeqQ3 handle THM _ => (HOLogic.mk_obj_eq Q2eeqQ3 RS (HOLogic.mk_obj_eq thA RS trans)) RS eq_reflection val impth = implies_intr_list ants1 QeeqQ3 val impth1 = HOLogic.mk_obj_eq impth (* Need to abstract *) val ant_th = Utils.itlist2 (PGEN ctxt' tych) args vstrl impth1 in ant_th COMP thm end fun q_eliminate (thm, imp) = let val (vlist, imp_body, used') = strip_all used imp val (ants,Q) = dest_impl imp_body in if (pbeta_redex Q) (length vlist) then pq_eliminate (thm, vlist, imp_body, Q) else let val tych = Thm.cterm_of ctxt val ants1 = map tych ants val ctxt' = Simplifier.add_prems (map ASSUME ants1) ctxt val Q_eeq_Q1 = Raw_Simplifier.rewrite_cterm (false,true,false) (prover used') ctxt' (tych Q) handle Utils.ERR _ => Thm.reflexive (tych Q) val lhs_eeq_lhs2 = implies_intr_list ants1 Q_eeq_Q1 val lhs_eq_lhs2 = HOLogic.mk_obj_eq lhs_eeq_lhs2 val ant_th = forall_intr_list(map tych vlist)lhs_eq_lhs2 in ant_th COMP thm end end fun eliminate thm = case Thm.prop_of thm of Const(\<^const_name>\Pure.imp\,_) $ imp $ _ => eliminate (if not(is_all imp) then uq_eliminate (thm, imp) else q_eliminate (thm, imp)) (* Assume that the leading constant is ==, *) | _ => thm (* if it is not a ==> *) in SOME(eliminate (rename thm)) end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) fun restrict_prover ctxt thm = let val _ = say "restrict_prover:" val cntxt = rev (Simplifier.prems_of ctxt) val _ = print_thms ctxt "cntxt:" cntxt val Const(\<^const_name>\Pure.imp\,_) $ (Const(\<^const_name>\Trueprop\,_) $ A) $ _ = Thm.prop_of thm fun genl tm = let val vlist = subtract (op aconv) globals (Misc_Legacy.add_term_frees(tm,[])) in fold_rev Forall vlist tm end (*-------------------------------------------------------------- * This actually isn't quite right, since it will think that * not-fully applied occs. of "f" in the context mean that the * current call is nested. The real solution is to pass in a * term "f v1..vn" which is a pattern that any full application * of "f" will match. *-------------------------------------------------------------*) val func_name = #1(dest_Const func) fun is_func (Const (name,_)) = (name = func_name) | is_func _ = false val rcontext = rev cntxt val cncl = HOLogic.dest_Trueprop o Thm.prop_of val antl = case rcontext of [] => [] | _ => [USyntax.list_mk_conj(map cncl rcontext)] val TC = genl(USyntax.list_mk_imp(antl, A)) val _ = print_term ctxt "func:" func val _ = print_term ctxt "TC:" (HOLogic.mk_Trueprop TC) val _ = tc_list := (TC :: !tc_list) val nestedp = is_some (USyntax.find_term is_func TC) val _ = if nestedp then say "nested" else say "not_nested" val th' = if nestedp then raise RULES_ERR "solver" "nested function" else let val cTC = Thm.cterm_of ctxt (HOLogic.mk_Trueprop TC) in case rcontext of [] => SPEC_ALL(ASSUME cTC) | _ => MP (SPEC_ALL (ASSUME cTC)) (LIST_CONJ rcontext) end val th'' = th' RS thm in SOME (th'') end handle Utils.ERR _ => NONE (* FIXME handle THM as well?? *) in (if (is_cong thm) then cong_prover else restrict_prover) ctxt thm end val ctm = Thm.cprop_of th val names = Misc_Legacy.add_term_names (Thm.term_of ctm, []) val th1 = Raw_Simplifier.rewrite_cterm (false, true, false) (prover names) (ctxt0 addsimps [cut_lemma'] |> fold Simplifier.add_eqcong congs) ctm val th2 = Thm.equal_elim th1 th in (th2, filter_out restricted (!tc_list)) end; fun prove ctxt strict (t, tac) = let val ctxt' = Proof_Context.augment t ctxt; in if strict then Goal.prove ctxt' [] [] t (K tac) else Goal.prove ctxt' [] [] t (K tac) handle ERROR msg => (warning msg; raise RULES_ERR "prove" msg) end; end; (*** theory operations ***) structure Thry: THRY = struct fun THRY_ERR func mesg = Utils.ERR {module = "Thry", func = func, mesg = mesg}; (*--------------------------------------------------------------------------- * Matching *---------------------------------------------------------------------------*) local fun tybind (ixn, (S, T)) = (TVar (ixn, S), T); in fun match_term thry pat ob = let val (ty_theta, tm_theta) = Pattern.match thry (pat,ob) (Vartab.empty, Vartab.empty); fun tmbind (ixn, (T, t)) = (Var (ixn, Envir.subst_type ty_theta T), t) in (map tmbind (Vartab.dest tm_theta), map tybind (Vartab.dest ty_theta)) end; fun match_type thry pat ob = map tybind (Vartab.dest (Sign.typ_match thry (pat, ob) Vartab.empty)); end; (*--------------------------------------------------------------------------- * Typing *---------------------------------------------------------------------------*) fun typecheck thy t = Thm.global_cterm_of thy t handle TYPE (msg, _, _) => raise THRY_ERR "typecheck" msg | TERM (msg, _) => raise THRY_ERR "typecheck" msg; (*--------------------------------------------------------------------------- * Get information about datatypes *---------------------------------------------------------------------------*) fun match_info thy dtco = case (BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco, BNF_LFP_Compat.get_constrs thy dtco) of (SOME {case_name, ... }, SOME constructors) => SOME {case_const = Const (case_name, Sign.the_const_type thy case_name), constructors = map Const constructors} | _ => NONE; fun induct_info thy dtco = case BNF_LFP_Compat.get_info thy [BNF_LFP_Compat.Keep_Nesting] dtco of NONE => NONE | SOME {nchotomy, ...} => SOME {nchotomy = nchotomy, constructors = (map Const o the o BNF_LFP_Compat.get_constrs thy) dtco}; fun extract_info thy = let val infos = map snd (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting])) in {case_congs = map (mk_meta_eq o #case_cong) infos, case_rewrites = maps (map mk_meta_eq o #case_rewrites) infos} end; end; (*** first part of main module ***) structure Prim: PRIM = struct val trace = Unsynchronized.ref false; fun TFL_ERR func mesg = Utils.ERR {module = "Tfl", func = func, mesg = mesg}; val concl = #2 o Rules.dest_thm; val list_mk_type = Utils.end_itlist (curry (op -->)); fun front_last [] = raise TFL_ERR "front_last" "empty list" | front_last [x] = ([],x) | front_last (h::t) = let val (pref,x) = front_last t in (h::pref,x) end; (*--------------------------------------------------------------------------- * The next function is common to pattern-match translation and * proof of completeness of cases for the induction theorem. * * The curried function "gvvariant" returns a function to generate distinct * variables that are guaranteed not to be in names. The names of * the variables go u, v, ..., z, aa, ..., az, ... The returned * function contains embedded refs! *---------------------------------------------------------------------------*) fun gvvariant names = let val slist = Unsynchronized.ref names val vname = Unsynchronized.ref "u" fun new() = if member (op =) (!slist) (!vname) then (vname := Symbol.bump_string (!vname); new()) else (slist := !vname :: !slist; !vname) in fn ty => Free(new(), ty) end; (*--------------------------------------------------------------------------- * Used in induction theorem production. This is the simple case of * partitioning up pattern rows by the leading constructor. *---------------------------------------------------------------------------*) fun ipartition gv (constructors,rows) = let fun pfail s = raise TFL_ERR "partition.part" s fun part {constrs = [], rows = [], A} = rev A | part {constrs = [], rows = _::_, A} = pfail"extra cases in defn" | part {constrs = _::_, rows = [], A} = pfail"cases missing in defn" | part {constrs = c::crst, rows, A} = let val (c, T) = dest_Const c val L = binder_types T val (in_group, not_in_group) = fold_rev (fn (row as (p::rst, rhs)) => fn (in_group,not_in_group) => let val (pc,args) = USyntax.strip_comb p in if (#1(dest_Const pc) = c) then ((args@rst, rhs)::in_group, not_in_group) else (in_group, row::not_in_group) end) rows ([],[]) val col_types = Utils.take type_of (length L, #1(hd in_group)) in part{constrs = crst, rows = not_in_group, A = {constructor = c, new_formals = map gv col_types, group = in_group}::A} end in part{constrs = constructors, rows = rows, A = []} end; (*--------------------------------------------------------------------------- * Each pattern carries with it a tag (i,b) where * i is the clause it came from and * b=true indicates that clause was given by the user * (or is an instantiation of a user supplied pattern) * b=false --> i = ~1 *---------------------------------------------------------------------------*) type pattern = term * (int * bool) fun pattern_map f (tm,x) = (f tm, x); fun pattern_subst theta = pattern_map (subst_free theta); val pat_of = fst; fun row_of_pat x = fst (snd x); fun given x = snd (snd x); (*--------------------------------------------------------------------------- * Produce an instance of a constructor, plus genvars for its arguments. *---------------------------------------------------------------------------*) fun fresh_constr ty_match colty gv c = let val (_,Ty) = dest_Const c val L = binder_types Ty and ty = body_type Ty val ty_theta = ty_match ty colty val c' = USyntax.inst ty_theta c val gvars = map (USyntax.inst ty_theta o gv) L in (c', gvars) end; (*--------------------------------------------------------------------------- * Goes through a list of rows and picks out the ones beginning with a * pattern with constructor = name. *---------------------------------------------------------------------------*) fun mk_group name rows = fold_rev (fn (row as ((prfx, p::rst), rhs)) => fn (in_group,not_in_group) => let val (pc,args) = USyntax.strip_comb p in if ((#1 (Term.dest_Const pc) = name) handle TERM _ => false) then (((prfx,args@rst), rhs)::in_group, not_in_group) else (in_group, row::not_in_group) end) rows ([],[]); (*--------------------------------------------------------------------------- * Partition the rows. Not efficient: we should use hashing. *---------------------------------------------------------------------------*) fun partition _ _ (_,_,_,[]) = raise TFL_ERR "partition" "no rows" | partition gv ty_match (constructors, colty, res_ty, rows as (((prfx,_),_)::_)) = let val fresh = fresh_constr ty_match colty gv fun part {constrs = [], rows, A} = rev A | part {constrs = c::crst, rows, A} = let val (c',gvars) = fresh c val (in_group, not_in_group) = mk_group (#1 (dest_Const c')) rows val in_group' = if (null in_group) (* Constructor not given *) then [((prfx, #2(fresh c)), (USyntax.ARB res_ty, (~1,false)))] else in_group in part{constrs = crst, rows = not_in_group, A = {constructor = c', new_formals = gvars, group = in_group'}::A} end in part{constrs=constructors, rows=rows, A=[]} end; (*--------------------------------------------------------------------------- * Misc. routines used in mk_case *---------------------------------------------------------------------------*) fun mk_pat (c,l) = let val L = length (binder_types (type_of c)) fun build (prfx,tag,plist) = let val (args, plist') = chop L plist in (prfx,tag,list_comb(c,args)::plist') end in map build l end; fun v_to_prfx (prfx, v::pats) = (v::prfx,pats) | v_to_prfx _ = raise TFL_ERR "mk_case" "v_to_prfx"; fun v_to_pats (v::prfx,tag, pats) = (prfx, tag, v::pats) | v_to_pats _ = raise TFL_ERR "mk_case" "v_to_pats"; (*---------------------------------------------------------------------------- * Translation of pattern terms into nested case expressions. * * This performs the translation and also builds the full set of patterns. * Thus it supports the construction of induction theorems even when an * incomplete set of patterns is given. *---------------------------------------------------------------------------*) fun mk_case ty_info ty_match usednames range_ty = let fun mk_case_fail s = raise TFL_ERR "mk_case" s val fresh_var = gvvariant usednames val divide = partition fresh_var ty_match fun expand _ ty ((_,[]), _) = mk_case_fail"expand_var_row" | expand constructors ty (row as ((prfx, p::rst), rhs)) = if (is_Free p) then let val fresh = fresh_constr ty_match ty fresh_var fun expnd (c,gvs) = let val capp = list_comb(c,gvs) in ((prfx, capp::rst), pattern_subst[(p,capp)] rhs) end in map expnd (map fresh constructors) end else [row] fun mk{rows=[],...} = mk_case_fail"no rows" | mk{path=[], rows = ((prfx, []), (tm,tag))::_} = (* Done *) ([(prfx,tag,[])], tm) | mk{path=[], rows = _::_} = mk_case_fail"blunder" | mk{path as u::rstp, rows as ((prfx, []), rhs)::rst} = mk{path = path, rows = ((prfx, [fresh_var(type_of u)]), rhs)::rst} | mk{path = u::rstp, rows as ((_, p::_), _)::_} = let val (pat_rectangle,rights) = ListPair.unzip rows val col0 = map(hd o #2) pat_rectangle in if (forall is_Free col0) then let val rights' = map (fn(v,e) => pattern_subst[(v,u)] e) (ListPair.zip (col0, rights)) val pat_rectangle' = map v_to_prfx pat_rectangle val (pref_patl,tm) = mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')} in (map v_to_pats pref_patl, tm) end else let val pty as Type (ty_name,_) = type_of p in case (ty_info ty_name) of NONE => mk_case_fail("Not a known datatype: "^ty_name) | SOME{case_const,constructors} => let val case_const_name = #1(dest_Const case_const) val nrows = maps (expand constructors pty) rows val subproblems = divide(constructors, pty, range_ty, nrows) val groups = map #group subproblems and new_formals = map #new_formals subproblems and constructors' = map #constructor subproblems val news = map (fn (nf,rows) => {path = nf@rstp, rows=rows}) (ListPair.zip (new_formals, groups)) val rec_calls = map mk news val (pat_rect,dtrees) = ListPair.unzip rec_calls val case_functions = map USyntax.list_mk_abs (ListPair.zip (new_formals, dtrees)) val types = map type_of (case_functions@[u]) @ [range_ty] val case_const' = Const(case_const_name, list_mk_type types) val tree = list_comb(case_const', case_functions@[u]) val pat_rect1 = flat (ListPair.map mk_pat (constructors', pat_rect)) in (pat_rect1,tree) end end end in mk end; (* Repeated variable occurrences in a pattern are not allowed. *) fun FV_multiset tm = case (USyntax.dest_term tm) of USyntax.VAR{Name = c, Ty = T} => [Free(c, T)] | USyntax.CONST _ => [] | USyntax.COMB{Rator, Rand} => FV_multiset Rator @ FV_multiset Rand | USyntax.LAMB _ => raise TFL_ERR "FV_multiset" "lambda"; fun no_repeat_vars thy pat = let fun check [] = true | check (v::rst) = if member (op aconv) rst v then raise TFL_ERR "no_repeat_vars" (quote (#1 (dest_Free v)) ^ " occurs repeatedly in the pattern " ^ quote (Syntax.string_of_term_global thy pat)) else check rst in check (FV_multiset pat) end; fun dest_atom (Free p) = p | dest_atom (Const p) = p | dest_atom _ = raise TFL_ERR "dest_atom" "function name not an identifier"; fun same_name (p,q) = #1(dest_atom p) = #1(dest_atom q); local fun mk_functional_err s = raise TFL_ERR "mk_functional" s fun single [_$_] = mk_functional_err "recdef does not allow currying" | single [f] = f | single fs = (*multiple function names?*) if length (distinct same_name fs) < length fs then mk_functional_err "The function being declared appears with multiple types" else mk_functional_err (string_of_int (length fs) ^ " distinct function names being declared") in fun mk_functional thy clauses = let val (L,R) = ListPair.unzip (map HOLogic.dest_eq clauses handle TERM _ => raise TFL_ERR "mk_functional" "recursion equations must use the = relation") val (funcs,pats) = ListPair.unzip (map (fn (t$u) =>(t,u)) L) val atom = single (distinct (op aconv) funcs) val (fname,ftype) = dest_atom atom val _ = map (no_repeat_vars thy) pats val rows = ListPair.zip (map (fn x => ([]:term list,[x])) pats, map_index (fn (i, t) => (t,(i,true))) R) val names = List.foldr Misc_Legacy.add_term_names [] R val atype = type_of(hd pats) and aname = singleton (Name.variant_list names) "a" val a = Free(aname,atype) val ty_info = Thry.match_info thy val ty_match = Thry.match_type thy val range_ty = type_of (hd R) val (patts, case_tm) = mk_case ty_info ty_match (aname::names) range_ty {path=[a], rows=rows} val patts1 = map (fn (_,tag,[pat]) => (pat,tag)) patts handle Match => mk_functional_err "error in pattern-match translation" val patts2 = Library.sort (Library.int_ord o apply2 row_of_pat) patts1 val finals = map row_of_pat patts2 val originals = map (row_of_pat o #2) rows val _ = case (subtract (op =) finals originals) of [] => () | L => mk_functional_err ("The following clauses are redundant (covered by preceding clauses): " ^ commas (map (fn i => string_of_int (i + 1)) L)) in {functional = Abs(Long_Name.base_name fname, ftype, abstract_over (atom, absfree (aname,atype) case_tm)), pats = patts2} end end; (*---------------------------------------------------------------------------- * * PRINCIPLES OF DEFINITION * *---------------------------------------------------------------------------*) (*For Isabelle, the lhs of a definition must be a constant.*) fun const_def sign (c, Ty, rhs) = singleton (Syntax.check_terms (Proof_Context.init_global sign)) (Const(\<^const_name>\Pure.eq\,dummyT) $ Const(c,Ty) $ rhs); (*Make all TVars available for instantiation by adding a ? to the front*) fun poly_tvars (Type(a,Ts)) = Type(a, map (poly_tvars) Ts) | poly_tvars (TFree (a,sort)) = TVar (("?" ^ a, 0), sort) | poly_tvars (TVar ((a,i),sort)) = TVar (("?" ^ a, i+1), sort); local val f_eq_wfrec_R_M = #ant(USyntax.dest_imp(#2(USyntax.strip_forall (concl @{thm tfl_wfrec})))) val {lhs=f, rhs} = USyntax.dest_eq f_eq_wfrec_R_M val _ = dest_Free f val (wfrec,_) = USyntax.strip_comb rhs in fun wfrec_definition0 fid R (functional as Abs(x, Ty, _)) thy = let val def_name = Thm.def_name (Long_Name.base_name fid) val wfrec_R_M = map_types poly_tvars (wfrec $ map_types poly_tvars R) $ functional val def_term = const_def thy (fid, Ty, wfrec_R_M) val ([def], thy') = Global_Theory.add_defs false [Thm.no_attributes (Binding.name def_name, def_term)] thy in (def, thy') end; end; (*--------------------------------------------------------------------------- * This structure keeps track of congruence rules that aren't derived * from a datatype definition. *---------------------------------------------------------------------------*) fun extraction_thms thy = let val {case_rewrites,case_congs} = Thry.extract_info thy in (case_rewrites, case_congs) end; (*--------------------------------------------------------------------------- * Pair patterns with termination conditions. The full list of patterns for * a definition is merged with the TCs arising from the user-given clauses. * There can be fewer clauses than the full list, if the user omitted some * cases. This routine is used to prepare input for mk_induction. *---------------------------------------------------------------------------*) fun merge full_pats TCs = let fun insert (p,TCs) = let fun insrt ((x as (h,[]))::rst) = if (p aconv h) then (p,TCs)::rst else x::insrt rst | insrt (x::rst) = x::insrt rst | insrt[] = raise TFL_ERR "merge.insert" "pattern not found" in insrt end fun pass ([],ptcl_final) = ptcl_final | pass (ptcs::tcl, ptcl) = pass(tcl, insert ptcs ptcl) in pass (TCs, map (fn p => (p,[])) full_pats) end; fun givens pats = map pat_of (filter given pats); fun post_definition ctxt meta_tflCongs (def, pats) = let val thy = Proof_Context.theory_of ctxt val tych = Thry.typecheck thy val f = #lhs(USyntax.dest_eq(concl def)) val corollary = Rules.MATCH_MP @{thm tfl_wfrec} def val pats' = filter given pats val given_pats = map pat_of pats' val rows = map row_of_pat pats' val WFR = #ant(USyntax.dest_imp(concl corollary)) val R = #Rand(USyntax.dest_comb WFR) val corollary' = Rules.UNDISCH corollary (* put WF R on assums *) val corollaries = map (fn pat => Rules.SPEC (tych pat) corollary') given_pats val (case_rewrites,context_congs) = extraction_thms thy (*case_ss causes minimal simplification: bodies of case expressions are not simplified. Otherwise large examples (Red-Black trees) are too slow.*) val case_simpset = put_simpset HOL_basic_ss ctxt addsimps case_rewrites |> fold (Simplifier.add_cong o #case_cong_weak o snd) (Symtab.dest (BNF_LFP_Compat.get_all thy [BNF_LFP_Compat.Keep_Nesting])) val corollaries' = map (Simplifier.simplify case_simpset) corollaries val extract = Rules.CONTEXT_REWRITE_RULE ctxt (f, [R], @{thm cut_apply}, meta_tflCongs @ context_congs) val (rules, TCs) = ListPair.unzip (map extract corollaries') val rules0 = map (rewrite_rule ctxt @{thms tfl_cut_def}) rules val mk_cond_rule = Rules.FILTER_DISCH_ALL(not o curry (op aconv) WFR) val rules1 = Rules.LIST_CONJ(map mk_cond_rule rules0) in {rules = rules1, rows = rows, full_pats_TCs = merge (map pat_of pats) (ListPair.zip (given_pats, TCs)), TCs = TCs} end; (*---------------------------------------------------------------------------- * * INDUCTION THEOREM * *---------------------------------------------------------------------------*) (*------------------------ Miscellaneous function -------------------------- * * [x_1,...,x_n] ?v_1...v_n. M[v_1,...,v_n] * ----------------------------------------------------------- * ( M[x_1,...,x_n], [(x_i,?v_1...v_n. M[v_1,...,v_n]), * ... * (x_j,?v_n. M[x_1,...,x_(n-1),v_n])] ) * * This function is totally ad hoc. Used in the production of the induction * theorem. The nchotomy theorem can have clauses that look like * * ?v1..vn. z = C vn..v1 * * in which the order of quantification is not the order of occurrence of the * quantified variables as arguments to C. Since we have no control over this * aspect of the nchotomy theorem, we make the correspondence explicit by * pairing the incoming new variable with the term it gets beta-reduced into. *---------------------------------------------------------------------------*) fun alpha_ex_unroll (xlist, tm) = let val (qvars,body) = USyntax.strip_exists tm val vlist = #2 (USyntax.strip_comb (USyntax.rhs body)) val plist = ListPair.zip (vlist, xlist) val args = map (the o AList.lookup (op aconv) plist) qvars handle Option.Option => raise Fail "TFL.alpha_ex_unroll: no correspondence" fun build ex [] = [] | build (_$rex) (v::rst) = let val ex1 = Term.betapply(rex, v) in ex1 :: build ex1 rst end val (nex::exl) = rev (tm::build tm args) in (nex, ListPair.zip (args, rev exl)) end; (*---------------------------------------------------------------------------- * * PROVING COMPLETENESS OF PATTERNS * *---------------------------------------------------------------------------*) fun mk_case ctxt ty_info usednames = let val thy = Proof_Context.theory_of ctxt val divide = ipartition (gvvariant usednames) val tych = Thry.typecheck thy fun tych_binding(x,y) = (tych x, tych y) fun fail s = raise TFL_ERR "mk_case" s fun mk{rows=[],...} = fail"no rows" | mk{path=[], rows = [([], (thm, bindings))]} = Rules.IT_EXISTS ctxt (map tych_binding bindings) thm | mk{path = u::rstp, rows as (p::_, _)::_} = let val (pat_rectangle,rights) = ListPair.unzip rows val col0 = map hd pat_rectangle val pat_rectangle' = map tl pat_rectangle in if (forall is_Free col0) (* column 0 is all variables *) then let val rights' = map (fn ((thm,theta),v) => (thm,theta@[(u,v)])) (ListPair.zip (rights, col0)) in mk{path = rstp, rows = ListPair.zip (pat_rectangle', rights')} end else (* column 0 is all constructors *) let val Type (ty_name,_) = type_of p in case (ty_info ty_name) of NONE => fail("Not a known datatype: "^ty_name) | SOME{constructors,nchotomy} => let val thm' = Rules.ISPEC (tych u) nchotomy val disjuncts = USyntax.strip_disj (concl thm') val subproblems = divide(constructors, rows) val groups = map #group subproblems and new_formals = map #new_formals subproblems val existentials = ListPair.map alpha_ex_unroll (new_formals, disjuncts) val constraints = map #1 existentials val vexl = map #2 existentials fun expnd tm (pats,(th,b)) = (pats, (Rules.SUBS ctxt [Rules.ASSUME (tych tm)] th, b)) val news = map (fn (nf,rows,c) => {path = nf@rstp, rows = map (expnd c) rows}) (Utils.zip3 new_formals groups constraints) val recursive_thms = map mk news val build_exists = Library.foldr (fn((x,t), th) => Rules.CHOOSE ctxt (tych x, Rules.ASSUME (tych t)) th) val thms' = ListPair.map build_exists (vexl, recursive_thms) val same_concls = Rules.EVEN_ORS thms' in Rules.DISJ_CASESL thm' same_concls end end end in mk end; fun complete_cases ctxt = let val thy = Proof_Context.theory_of ctxt val tych = Thry.typecheck thy val ty_info = Thry.induct_info thy in fn pats => let val names = List.foldr Misc_Legacy.add_term_names [] pats val T = type_of (hd pats) val aname = singleton (Name.variant_list names) "a" val vname = singleton (Name.variant_list (aname::names)) "v" val a = Free (aname, T) val v = Free (vname, T) val a_eq_v = HOLogic.mk_eq(a,v) val ex_th0 = Rules.EXISTS ctxt (tych (USyntax.mk_exists{Bvar=v,Body=a_eq_v}), tych a) (Rules.REFL (tych a)) val th0 = Rules.ASSUME (tych a_eq_v) val rows = map (fn x => ([x], (th0,[]))) pats in Rules.GEN ctxt (tych a) (Rules.RIGHT_ASSOC ctxt (Rules.CHOOSE ctxt (tych v, ex_th0) (mk_case ctxt ty_info (vname::aname::names) {path=[v], rows=rows}))) end end; (*--------------------------------------------------------------------------- * Constructing induction hypotheses: one for each recursive call. * * Note. R will never occur as a variable in the ind_clause, because * to do so, it would have to be from a nested definition, and we don't * allow nested defns to have R variable. * * Note. When the context is empty, there can be no local variables. *---------------------------------------------------------------------------*) local infix 5 ==> fun (tm1 ==> tm2) = USyntax.mk_imp{ant = tm1, conseq = tm2} in fun build_ih f (P,SV) (pat,TCs) = let val pat_vars = USyntax.free_vars_lr pat val globals = pat_vars@SV fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm) fun dest_TC tm = let val (cntxt,R_y_pat) = USyntax.strip_imp(#2(USyntax.strip_forall tm)) val (R,y,_) = USyntax.dest_relation R_y_pat val P_y = if (nested tm) then R_y_pat ==> P$y else P$y in case cntxt of [] => (P_y, (tm,[])) | _ => let val imp = USyntax.list_mk_conj cntxt ==> P_y val lvs = subtract (op aconv) globals (USyntax.free_vars_lr imp) val locals = #2(Utils.pluck (curry (op aconv) P) lvs) handle Utils.ERR _ => lvs in (USyntax.list_mk_forall(locals,imp), (tm,locals)) end end in case TCs of [] => (USyntax.list_mk_forall(pat_vars, P$pat), []) | _ => let val (ihs, TCs_locals) = ListPair.unzip(map dest_TC TCs) val ind_clause = USyntax.list_mk_conj ihs ==> P$pat in (USyntax.list_mk_forall(pat_vars,ind_clause), TCs_locals) end end end; (*--------------------------------------------------------------------------- * This function makes good on the promise made in "build_ih". * * Input is tm = "(!y. R y pat ==> P y) ==> P pat", * TCs = TC_1[pat] ... TC_n[pat] * thm = ih1 /\ ... /\ ih_n |- ih[pat] *---------------------------------------------------------------------------*) fun prove_case ctxt f (tm,TCs_locals,thm) = let val tych = Thry.typecheck (Proof_Context.theory_of ctxt) val antc = tych(#ant(USyntax.dest_imp tm)) val thm' = Rules.SPEC_ALL thm fun nested tm = is_some (USyntax.find_term (curry (op aconv) f) tm) fun get_cntxt TC = tych(#ant(USyntax.dest_imp(#2(USyntax.strip_forall(concl TC))))) fun mk_ih ((TC,locals),th2,nested) = Rules.GENL ctxt (map tych locals) (if nested then Rules.DISCH (get_cntxt TC) th2 handle Utils.ERR _ => th2 else if USyntax.is_imp (concl TC) then Rules.IMP_TRANS TC th2 else Rules.MP th2 TC) in Rules.DISCH antc (if USyntax.is_imp(concl thm') (* recursive calls in this clause *) then let val th1 = Rules.ASSUME antc val TCs = map #1 TCs_locals val ylist = map (#2 o USyntax.dest_relation o #2 o USyntax.strip_imp o #2 o USyntax.strip_forall) TCs val TClist = map (fn(TC,lvs) => (Rules.SPEC_ALL(Rules.ASSUME(tych TC)),lvs)) TCs_locals val th2list = map (fn t => Rules.SPEC (tych t) th1) ylist val nlist = map nested TCs val triples = Utils.zip3 TClist th2list nlist val Pylist = map mk_ih triples in Rules.MP thm' (Rules.LIST_CONJ Pylist) end else thm') end; (*--------------------------------------------------------------------------- * * x = (v1,...,vn) |- M[x] * --------------------------------------------- * ?v1 ... vn. x = (v1,...,vn) |- M[x] * *---------------------------------------------------------------------------*) fun LEFT_ABS_VSTRUCT ctxt tych thm = let fun CHOOSER v (tm,thm) = let val ex_tm = USyntax.mk_exists{Bvar=v,Body=tm} in (ex_tm, Rules.CHOOSE ctxt (tych v, Rules.ASSUME (tych ex_tm)) thm) end val [veq] = filter (can USyntax.dest_eq) (#1 (Rules.dest_thm thm)) val {lhs,rhs} = USyntax.dest_eq veq val L = USyntax.free_vars_lr rhs in #2 (fold_rev CHOOSER L (veq,thm)) end; (*---------------------------------------------------------------------------- * Input : f, R, and [(pat1,TCs1),..., (patn,TCsn)] * * Instantiates tfl_wf_induct, getting Sinduct and then tries to prove * recursion induction (Rinduct) by proving the antecedent of Sinduct from * the antecedent of Rinduct. *---------------------------------------------------------------------------*) fun mk_induction ctxt {fconst, R, SV, pat_TCs_list} = let val thy = Proof_Context.theory_of ctxt val tych = Thry.typecheck thy val Sinduction = Rules.UNDISCH (Rules.ISPEC (tych R) @{thm tfl_wf_induct}) val (pats,TCsl) = ListPair.unzip pat_TCs_list val case_thm = complete_cases ctxt pats val domain = (type_of o hd) pats val Pname = singleton (Name.variant_list (List.foldr (Library.foldr Misc_Legacy.add_term_names) [] (pats::TCsl))) "P" val P = Free(Pname, domain --> HOLogic.boolT) val Sinduct = Rules.SPEC (tych P) Sinduction val Sinduct_assumf = USyntax.rand ((#ant o USyntax.dest_imp o concl) Sinduct) val Rassums_TCl' = map (build_ih fconst (P,SV)) pat_TCs_list val (Rassums,TCl') = ListPair.unzip Rassums_TCl' val Rinduct_assum = Rules.ASSUME (tych (USyntax.list_mk_conj Rassums)) val cases = map (fn pat => Term.betapply (Sinduct_assumf, pat)) pats val tasks = Utils.zip3 cases TCl' (Rules.CONJUNCTS Rinduct_assum) val proved_cases = map (prove_case ctxt fconst) tasks val v = Free (singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] (map concl proved_cases))) "v", domain) val vtyped = tych v val substs = map (Rules.SYM o Rules.ASSUME o tych o (curry HOLogic.mk_eq v)) pats val proved_cases1 = ListPair.map (fn (th,th') => Rules.SUBS ctxt [th]th') (substs, proved_cases) val abs_cases = map (LEFT_ABS_VSTRUCT ctxt tych) proved_cases1 val dant = Rules.GEN ctxt vtyped (Rules.DISJ_CASESL (Rules.ISPEC vtyped case_thm) abs_cases) val dc = Rules.MP Sinduct dant val Parg_ty = type_of(#Bvar(USyntax.dest_forall(concl dc))) val vars = map (gvvariant[Pname]) (USyntax.strip_prod_type Parg_ty) val dc' = fold_rev (Rules.GEN ctxt o tych) vars (Rules.SPEC (tych(USyntax.mk_vstruct Parg_ty vars)) dc) in Rules.GEN ctxt (tych P) (Rules.DISCH (tych(concl Rinduct_assum)) dc') end handle Utils.ERR _ => raise TFL_ERR "mk_induction" "failed derivation"; (*--------------------------------------------------------------------------- * * POST PROCESSING * *---------------------------------------------------------------------------*) fun simplify_induction thy hth ind = let val tych = Thry.typecheck thy val (asl,_) = Rules.dest_thm ind val (_,tc_eq_tc') = Rules.dest_thm hth val tc = USyntax.lhs tc_eq_tc' fun loop [] = ind | loop (asm::rst) = if (can (Thry.match_term thy asm) tc) then Rules.UNDISCH (Rules.MATCH_MP (Rules.MATCH_MP @{thm tfl_simp_thm} (Rules.DISCH (tych asm) ind)) hth) else loop rst in loop asl end; (*--------------------------------------------------------------------------- * The termination condition is an antecedent to the rule, and an * assumption to the theorem. *---------------------------------------------------------------------------*) fun elim_tc tcthm (rule,induction) = (Rules.MP rule tcthm, Rules.PROVE_HYP tcthm induction) fun trace_thms ctxt s L = if !trace then writeln (cat_lines (s :: map (Thm.string_of_thm ctxt) L)) else (); fun trace_cterm ctxt s ct = if !trace then writeln (cat_lines [s, Syntax.string_of_term ctxt (Thm.term_of ct)]) else (); fun postprocess ctxt strict {wf_tac, terminator, simplifier} {rules,induction,TCs} = let val thy = Proof_Context.theory_of ctxt; val tych = Thry.typecheck thy; (*--------------------------------------------------------------------- * Attempt to eliminate WF condition. It's the only assumption of rules *---------------------------------------------------------------------*) val (rules1,induction1) = let val thm = Rules.prove ctxt strict (HOLogic.mk_Trueprop (hd(#1(Rules.dest_thm rules))), wf_tac) in (Rules.PROVE_HYP thm rules, Rules.PROVE_HYP thm induction) end handle Utils.ERR _ => (rules,induction); (*---------------------------------------------------------------------- * The termination condition (tc) is simplified to |- tc = tc' (there * might not be a change!) and then 3 attempts are made: * * 1. if |- tc = T, then eliminate it with tfl_eq_True; otherwise, * 2. apply the terminator to tc'. If |- tc' = T then eliminate; else * 3. replace tc by tc' in both the rules and the induction theorem. *---------------------------------------------------------------------*) fun simplify_tc tc (r,ind) = let val tc1 = tych tc val _ = trace_cterm ctxt "TC before simplification: " tc1 val tc_eq = simplifier tc1 val _ = trace_thms ctxt "result: " [tc_eq] in elim_tc (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq) (r,ind) handle Utils.ERR _ => (elim_tc (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq) (Rules.prove ctxt strict (HOLogic.mk_Trueprop(USyntax.rhs(concl tc_eq)), terminator))) (r,ind) handle Utils.ERR _ => (Rules.UNDISCH(Rules.MATCH_MP (Rules.MATCH_MP @{thm tfl_simp_thm} r) tc_eq), simplify_induction thy tc_eq ind)) end (*---------------------------------------------------------------------- * Nested termination conditions are harder to get at, since they are * left embedded in the body of the function (and in induction * theorem hypotheses). Our "solution" is to simplify them, and try to * prove termination, but leave the application of the resulting theorem * to a higher level. So things go much as in "simplify_tc": the * termination condition (tc) is simplified to |- tc = tc' (there might * not be a change) and then 2 attempts are made: * * 1. if |- tc = T, then return |- tc; otherwise, * 2. apply the terminator to tc'. If |- tc' = T then return |- tc; else * 3. return |- tc = tc' *---------------------------------------------------------------------*) fun simplify_nested_tc tc = let val tc_eq = simplifier (tych (#2 (USyntax.strip_forall tc))) in Rules.GEN_ALL ctxt (Rules.MATCH_MP @{thm tfl_eq_True} tc_eq handle Utils.ERR _ => (Rules.MATCH_MP(Rules.MATCH_MP @{thm tfl_rev_eq_mp} tc_eq) (Rules.prove ctxt strict (HOLogic.mk_Trueprop (USyntax.rhs(concl tc_eq)), terminator)) handle Utils.ERR _ => tc_eq)) end (*------------------------------------------------------------------- * Attempt to simplify the termination conditions in each rule and * in the induction theorem. *-------------------------------------------------------------------*) fun strip_imp tm = if USyntax.is_neg tm then ([],tm) else USyntax.strip_imp tm fun loop ([],extras,R,ind) = (rev R, ind, extras) | loop ((r,ftcs)::rst, nthms, R, ind) = let val tcs = #1(strip_imp (concl r)) val extra_tcs = subtract (op aconv) tcs ftcs val extra_tc_thms = map simplify_nested_tc extra_tcs val (r1,ind1) = fold simplify_tc tcs (r,ind) val r2 = Rules.FILTER_DISCH_ALL(not o USyntax.is_WFR) r1 in loop(rst, nthms@extra_tc_thms, r2::R, ind1) end val rules_tcs = ListPair.zip (Rules.CONJUNCTS rules1, TCs) val (rules2,ind2,extras) = loop(rules_tcs,[],[],induction1) in {induction = ind2, rules = Rules.LIST_CONJ rules2, nested_tcs = extras} end; end; (*** second part of main module (postprocessing of TFL definitions) ***) structure Tfl: TFL = struct (* misc *) (*--------------------------------------------------------------------------- * Extract termination goals so that they can be put it into a goalstack, or * have a tactic directly applied to them. *--------------------------------------------------------------------------*) fun termination_goals rules = map (Type.legacy_freeze o HOLogic.dest_Trueprop) (fold_rev (union (op aconv) o Thm.prems_of) rules []); (*--------------------------------------------------------------------------- * Three postprocessors are applied to the definition. It * attempts to prove wellfoundedness of the given relation, simplifies the * non-proved termination conditions, and finally attempts to prove the * simplified termination conditions. *--------------------------------------------------------------------------*) fun std_postprocessor ctxt strict wfs = Prim.postprocess ctxt strict {wf_tac = REPEAT (ares_tac ctxt wfs 1), terminator = asm_simp_tac ctxt 1 THEN TRY (Arith_Data.arith_tac ctxt 1 ORELSE fast_force_tac (ctxt addSDs @{thms not0_implies_Suc}) 1), simplifier = Rules.simpl_conv ctxt []}; val concl = #2 o Rules.dest_thm; (*--------------------------------------------------------------------------- * Postprocess a definition made by "define". This is a separate stage of * processing from the definition stage. *---------------------------------------------------------------------------*) local (* The rest of these local definitions are for the tricky nested case *) val solved = not o can USyntax.dest_eq o #2 o USyntax.strip_forall o concl fun id_thm th = let val {lhs,rhs} = USyntax.dest_eq (#2 (USyntax.strip_forall (#2 (Rules.dest_thm th)))); in lhs aconv rhs end handle Utils.ERR _ => false; val P_imp_P_eq_True = @{thm eqTrueI} RS eq_reflection; fun mk_meta_eq r = (case Thm.concl_of r of Const(\<^const_name>\Pure.eq\,_)$_$_ => r | _ $(Const(\<^const_name>\HOL.eq\,_)$_$_) => r RS eq_reflection | _ => r RS P_imp_P_eq_True) (*Is this the best way to invoke the simplifier??*) fun rewrite ctxt L = rewrite_rule ctxt (map mk_meta_eq (filter_out id_thm L)) fun join_assums ctxt th = let val tych = Thm.cterm_of ctxt val {lhs,rhs} = USyntax.dest_eq(#2 (USyntax.strip_forall (concl th))) val cntxtl = (#1 o USyntax.strip_imp) lhs (* cntxtl should = cntxtr *) val cntxtr = (#1 o USyntax.strip_imp) rhs (* but union is solider *) val cntxt = union (op aconv) cntxtl cntxtr in Rules.GEN_ALL ctxt (Rules.DISCH_ALL (rewrite ctxt (map (Rules.ASSUME o tych) cntxt) (Rules.SPEC_ALL th))) end val gen_all = USyntax.gen_all in fun proof_stage ctxt strict wfs {f, R, rules, full_pats_TCs, TCs} = let val _ = writeln "Proving induction theorem ..." val ind = Prim.mk_induction ctxt {fconst=f, R=R, SV=[], pat_TCs_list=full_pats_TCs} val _ = writeln "Postprocessing ..."; val {rules, induction, nested_tcs} = std_postprocessor ctxt strict wfs {rules=rules, induction=ind, TCs=TCs} in case nested_tcs of [] => {induction=induction, rules=rules,tcs=[]} | L => let val _ = writeln "Simplifying nested TCs ..." val (solved,simplified,stubborn) = fold_rev (fn th => fn (So,Si,St) => if (id_thm th) then (So, Si, th::St) else if (solved th) then (th::So, Si, St) else (So, th::Si, St)) nested_tcs ([],[],[]) val simplified' = map (join_assums ctxt) simplified val dummy = (Prim.trace_thms ctxt "solved =" solved; Prim.trace_thms ctxt "simplified' =" simplified') val rewr = full_simplify (ctxt addsimps (solved @ simplified')); val dummy = Prim.trace_thms ctxt "Simplifying the induction rule..." [induction] val induction' = rewr induction val dummy = Prim.trace_thms ctxt "Simplifying the recursion rules..." [rules] val rules' = rewr rules val _ = writeln "... Postprocessing finished"; in {induction = induction', rules = rules', tcs = map (gen_all o USyntax.rhs o #2 o USyntax.strip_forall o concl) (simplified@stubborn)} end end; (*lcp: curry the predicate of the induction rule*) fun curry_rule ctxt rl = Split_Rule.split_rule_var ctxt (Term.head_of (HOLogic.dest_Trueprop (Thm.concl_of rl))) rl; (*lcp: put a theorem into Isabelle form, using meta-level connectives*) fun meta_outer ctxt = curry_rule ctxt o Drule.export_without_context o rule_by_tactic ctxt (REPEAT (FIRSTGOAL (resolve_tac ctxt [allI, impI, conjI] ORELSE' eresolve_tac ctxt [conjE]))); (*Strip off the outer !P*) val spec'= Rule_Insts.read_instantiate \<^context> [((("x", 0), Position.none), "P::'b=>bool")] [] spec; fun simplify_defn ctxt strict congs wfs pats def0 = let val thy = Proof_Context.theory_of ctxt; val def = HOLogic.mk_obj_eq (Thm.unvarify_global thy def0) val {rules, rows, TCs, full_pats_TCs} = Prim.post_definition ctxt congs (def, pats) val {lhs=f,rhs} = USyntax.dest_eq (concl def) val (_,[R,_]) = USyntax.strip_comb rhs val _ = Prim.trace_thms ctxt "congs =" congs (*the next step has caused simplifier looping in some cases*) val {induction, rules, tcs} = proof_stage ctxt strict wfs {f = f, R = R, rules = rules, full_pats_TCs = full_pats_TCs, TCs = TCs} val rules' = map (Drule.export_without_context o Object_Logic.rulify_no_asm ctxt) (Rules.CONJUNCTS rules) in {induct = meta_outer ctxt (Object_Logic.rulify_no_asm ctxt (induction RS spec')), rules = ListPair.zip(rules', rows), tcs = (termination_goals rules') @ tcs} end handle Utils.ERR {mesg,func,module} => error (mesg ^ "\n (In TFL function " ^ module ^ "." ^ func ^ ")"); (* Derive the initial equations from the case-split rules to meet the users specification of the recursive function. *) local fun get_related_thms i = map_filter ((fn (r,x) => if x = i then SOME r else NONE)); fun solve_eq _ (_, [], _) = error "derive_init_eqs: missing rules" | solve_eq _ (_, [a], i) = [(a, i)] | solve_eq ctxt (th, splitths, i) = (writeln "Proving unsplit equation..."; [((Drule.export_without_context o Object_Logic.rulify_no_asm ctxt) (CaseSplit.splitto ctxt splitths th), i)]) handle ERROR s => (warning ("recdef (solve_eq): " ^ s); map (fn x => (x,i)) splitths); in fun derive_init_eqs ctxt rules eqs = map (Thm.trivial o Thm.cterm_of ctxt o HOLogic.mk_Trueprop) eqs |> map_index (fn (i, e) => solve_eq ctxt (e, (get_related_thms i rules), i)) |> flat; end; (*--------------------------------------------------------------------------- * Defining a function with an associated termination relation. *---------------------------------------------------------------------------*) fun define_i strict congs wfs fid R eqs ctxt = let val thy = Proof_Context.theory_of ctxt val {functional, pats} = Prim.mk_functional thy eqs val (def, thy') = Prim.wfrec_definition0 fid R functional thy val ctxt' = Proof_Context.transfer thy' ctxt val (lhs, _) = Logic.dest_equals (Thm.prop_of def) val {induct, rules, tcs} = simplify_defn ctxt' strict congs wfs pats def val rules' = if strict then derive_init_eqs ctxt' rules eqs else rules in ({lhs = lhs, rules = rules', induct = induct, tcs = tcs}, ctxt') end; fun define strict congs wfs fid R seqs ctxt = define_i strict congs wfs fid (Syntax.read_term ctxt R) (map (Syntax.read_term ctxt) seqs) ctxt handle Utils.ERR {mesg,...} => error mesg; end; end; (*** wrappers for Isar ***) (** recdef hints **) (* type hints *) type hints = {simps: thm list, congs: (string * thm) list, wfs: thm list}; fun mk_hints (simps, congs, wfs) = {simps = simps, congs = congs, wfs = wfs}: hints; fun map_hints f ({simps, congs, wfs}: hints) = mk_hints (f (simps, congs, wfs)); fun map_simps f = map_hints (fn (simps, congs, wfs) => (f simps, congs, wfs)); fun map_congs f = map_hints (fn (simps, congs, wfs) => (simps, f congs, wfs)); fun map_wfs f = map_hints (fn (simps, congs, wfs) => (simps, congs, f wfs)); (* congruence rules *) local val cong_head = fst o Term.dest_Const o Term.head_of o fst o Logic.dest_equals o Thm.concl_of; fun prep_cong raw_thm = let val thm = safe_mk_meta_eq raw_thm in (cong_head thm, thm) end; in fun add_cong raw_thm congs = let val (c, thm) = prep_cong raw_thm; val _ = if AList.defined (op =) congs c then warning ("Overwriting recdef congruence rule for " ^ quote c) else (); in AList.update (op =) (c, thm) congs end; fun del_cong raw_thm congs = let val (c, _) = prep_cong raw_thm; val _ = if AList.defined (op =) congs c then () else warning ("No recdef congruence rule for " ^ quote c); in AList.delete (op =) c congs end; end; (** global and local recdef data **) (* theory data *) type recdef_info = {lhs: term, simps: thm list, rules: thm list list, induct: thm, tcs: term list}; structure Data = Generic_Data ( type T = recdef_info Symtab.table * hints; val empty = (Symtab.empty, mk_hints ([], [], [])): T; val extend = I; fun merge ((tab1, {simps = simps1, congs = congs1, wfs = wfs1}), (tab2, {simps = simps2, congs = congs2, wfs = wfs2})) : T = (Symtab.merge (K true) (tab1, tab2), mk_hints (Thm.merge_thms (simps1, simps2), AList.merge (op =) (K true) (congs1, congs2), Thm.merge_thms (wfs1, wfs2))); ); val get_recdef = Symtab.lookup o #1 o Data.get o Context.Theory; fun put_recdef name info = (Context.theory_map o Data.map o apfst) (fn tab => Symtab.update_new (name, info) tab handle Symtab.DUP _ => error ("Duplicate recursive function definition " ^ quote name)); val get_hints = #2 o Data.get o Context.Proof; val map_hints = Data.map o apsnd; (* attributes *) fun attrib f = Thm.declaration_attribute (map_hints o f); val simp_add = attrib (map_simps o Thm.add_thm); val simp_del = attrib (map_simps o Thm.del_thm); val cong_add = attrib (map_congs o add_cong); val cong_del = attrib (map_congs o del_cong); val wf_add = attrib (map_wfs o Thm.add_thm); val wf_del = attrib (map_wfs o Thm.del_thm); (* modifiers *) val recdef_simpN = "recdef_simp"; val recdef_congN = "recdef_cong"; val recdef_wfN = "recdef_wf"; val recdef_modifiers = [Args.$$$ recdef_simpN -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ recdef_simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add \<^here>), Args.$$$ recdef_simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del \<^here>), Args.$$$ recdef_congN -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ recdef_congN -- Args.add -- Args.colon >> K (Method.modifier cong_add \<^here>), Args.$$$ recdef_congN -- Args.del -- Args.colon >> K (Method.modifier cong_del \<^here>), Args.$$$ recdef_wfN -- Args.colon >> K (Method.modifier wf_add \<^here>), Args.$$$ recdef_wfN -- Args.add -- Args.colon >> K (Method.modifier wf_add \<^here>), Args.$$$ recdef_wfN -- Args.del -- Args.colon >> K (Method.modifier wf_del \<^here>)] @ Clasimp.clasimp_modifiers; (** prepare hints **) fun prepare_hints opt_src ctxt = let val ctxt' = (case opt_src of NONE => ctxt | SOME src => #2 (Token.syntax (Method.sections recdef_modifiers) src ctxt)); val {simps, congs, wfs} = get_hints ctxt'; val ctxt'' = ctxt' addsimps simps |> Simplifier.del_cong @{thm imp_cong}; in ((rev (map snd congs), wfs), ctxt'') end; fun prepare_hints_i () ctxt = let val {simps, congs, wfs} = get_hints ctxt; val ctxt' = ctxt addsimps simps |> Simplifier.del_cong @{thm imp_cong}; in ((rev (map snd congs), wfs), ctxt') end; (** add_recdef(_i) **) fun gen_add_recdef tfl_fn prep_att prep_hints not_permissive raw_name R eq_srcs hints thy = let val _ = legacy_feature "Old 'recdef' command -- use 'fun' or 'function' instead"; val name = Sign.intern_const thy raw_name; val bname = Long_Name.base_name name; val _ = writeln ("Defining recursive function " ^ quote name ^ " ..."); val ((eq_names, eqs), raw_eq_atts) = apfst split_list (split_list eq_srcs); val eq_atts = map (map (prep_att thy)) raw_eq_atts; val ((congs, wfs), ctxt) = prep_hints hints (Proof_Context.init_global thy); (*We must remove imp_cong to prevent looping when the induction rule is simplified. Many induction rules have nested implications that would give rise to looping conditional rewriting.*) val ({lhs, rules = rules_idx, induct, tcs}, ctxt1) = tfl_fn not_permissive congs wfs name R eqs ctxt; val rules = (map o map) fst (partition_eq (eq_snd (op = : int * int -> bool)) rules_idx); val simp_att = if null tcs then [Simplifier.simp_add, Named_Theorems.add \<^named_theorems>\nitpick_simp\] else []; val ((simps' :: rules', [induct']), thy2) = Proof_Context.theory_of ctxt1 |> Sign.add_path bname |> Global_Theory.add_thmss (((Binding.name "simps", flat rules), simp_att) :: ((eq_names ~~ rules) ~~ eq_atts)) ||>> Global_Theory.add_thms [((Binding.name "induct", induct), [])] ||> Spec_Rules.add_global (Binding.name bname) Spec_Rules.equational_recdef [lhs] (flat rules) ||> null tcs ? Code.declare_default_eqns_global (map (rpair true) (flat rules)); val result = {lhs = lhs, simps = simps', rules = rules', induct = induct', tcs = tcs}; val thy3 = thy2 |> put_recdef name result |> Sign.parent_path; in (thy3, result) end; val add_recdef = gen_add_recdef Tfl.define Attrib.attribute_cmd_global prepare_hints; fun add_recdef_i x y z w = gen_add_recdef Tfl.define_i (K I) prepare_hints_i x y z w (); (** package setup **) (* setup theory *) val _ = Theory.setup (Attrib.setup \<^binding>\recdef_simp\ (Attrib.add_del simp_add simp_del) "declaration of recdef simp rule" #> Attrib.setup \<^binding>\recdef_cong\ (Attrib.add_del cong_add cong_del) "declaration of recdef cong rule" #> Attrib.setup \<^binding>\recdef_wf\ (Attrib.add_del wf_add wf_del) "declaration of recdef wf rule"); (* outer syntax *) val hints = \<^keyword>\(\ |-- Parse.!!! ((Parse.token \<^keyword>\hints\ ::: Parse.args) --| \<^keyword>\)\); val recdef_decl = Scan.optional (\<^keyword>\(\ -- Parse.!!! (\<^keyword>\permissive\ -- \<^keyword>\)\) >> K false) true -- Parse.name -- Parse.term -- Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- Scan.option hints >> (fn ((((p, f), R), eqs), src) => #1 o add_recdef p f R (map (fn ((x, y), z) => ((x, z), y)) eqs) src); val _ = Outer_Syntax.command \<^command_keyword>\recdef\ "define general recursive functions (obsolete TFL)" (recdef_decl >> Toplevel.theory); end; diff --git a/src/HOL/Library/rewrite.ML b/src/HOL/Library/rewrite.ML --- a/src/HOL/Library/rewrite.ML +++ b/src/HOL/Library/rewrite.ML @@ -1,453 +1,453 @@ (* Title: HOL/Library/rewrite.ML Author: Christoph Traut, Lars Noschinski, TU Muenchen This is a rewrite method that supports subterm-selection based on patterns. The patterns accepted by rewrite are of the following form: ::= | "concl" | "asm" | "for" "(" ")" ::= (in | at ) [] ::= [] ("to" ) This syntax was clearly inspired by Gonthier's and Tassi's language of patterns but has diverged significantly during its development. We also allow introduction of identifiers for bound variables, which can then be used to match arbitrary subterms inside abstractions. *) infix 1 then_pconv; infix 0 else_pconv; signature REWRITE = sig type patconv = Proof.context -> Type.tyenv * (string * term) list -> cconv val then_pconv: patconv * patconv -> patconv val else_pconv: patconv * patconv -> patconv val abs_pconv: patconv -> string option * typ -> patconv (*XXX*) val fun_pconv: patconv -> patconv val arg_pconv: patconv -> patconv val imp_pconv: patconv -> patconv val params_pconv: patconv -> patconv val forall_pconv: patconv -> string option * typ option -> patconv val all_pconv: patconv val for_pconv: patconv -> (string option * typ option) list -> patconv val concl_pconv: patconv -> patconv val asm_pconv: patconv -> patconv val asms_pconv: patconv -> patconv val judgment_pconv: patconv -> patconv val in_pconv: patconv -> patconv val match_pconv: patconv -> term * (string option * typ) list -> patconv val rewrs_pconv: term option -> thm list -> patconv datatype ('a, 'b) pattern = At | In | Term of 'a | Concl | Asm | For of 'b list val mk_hole: int -> typ -> term val rewrite_conv: Proof.context -> (term * (string * typ) list, string * typ option) pattern list * term option -> thm list -> conv end structure Rewrite : REWRITE = struct datatype ('a, 'b) pattern = At | In | Term of 'a | Concl | Asm | For of 'b list exception NO_TO_MATCH val holeN = Name.internal "_hole" fun prep_meta_eq ctxt = Simplifier.mksimps ctxt #> map Drule.zero_var_indexes (* holes *) fun mk_hole i T = Var ((holeN, i), T) fun is_hole (Var ((name, _), _)) = (name = holeN) | is_hole _ = false fun is_hole_const (Const (\<^const_name>\rewrite_HOLE\, _)) = true | is_hole_const _ = false val hole_syntax = let (* Modified variant of Term.replace_hole *) fun replace_hole Ts (Const (\<^const_name>\rewrite_HOLE\, T)) i = (list_comb (mk_hole i (Ts ---> T), map_range Bound (length Ts)), i + 1) | replace_hole Ts (Abs (x, T, t)) i = let val (t', i') = replace_hole (T :: Ts) t i in (Abs (x, T, t'), i') end | replace_hole Ts (t $ u) i = let val (t', i') = replace_hole Ts t i val (u', i'') = replace_hole Ts u i' in (t' $ u', i'') end | replace_hole _ a i = (a, i) fun prep_holes ts = #1 (fold_map (replace_hole []) ts 1) in Context.proof_map (Syntax_Phases.term_check 101 "hole_expansion" (K prep_holes)) #> Proof_Context.set_mode Proof_Context.mode_pattern end (* pattern conversions *) type patconv = Proof.context -> Type.tyenv * (string * term) list -> cterm -> thm fun (cv1 then_pconv cv2) ctxt tytenv ct = (cv1 ctxt tytenv then_conv cv2 ctxt tytenv) ct fun (cv1 else_pconv cv2) ctxt tytenv ct = (cv1 ctxt tytenv else_conv cv2 ctxt tytenv) ct fun raw_abs_pconv cv ctxt tytenv ct = case Thm.term_of ct of Abs _ => CConv.abs_cconv (fn (x, ctxt') => cv x ctxt' tytenv) ctxt ct | t => raise TERM ("raw_abs_pconv", [t]) fun raw_fun_pconv cv ctxt tytenv ct = case Thm.term_of ct of _ $ _ => CConv.fun_cconv (cv ctxt tytenv) ct | t => raise TERM ("raw_fun_pconv", [t]) fun raw_arg_pconv cv ctxt tytenv ct = case Thm.term_of ct of _ $ _ => CConv.arg_cconv (cv ctxt tytenv) ct | t => raise TERM ("raw_arg_pconv", [t]) fun abs_pconv cv (s,T) ctxt (tyenv, ts) ct = let val u = Thm.term_of ct in case try (fastype_of #> dest_funT) u of NONE => raise TERM ("abs_pconv: no function type", [u]) | SOME (U, _) => let val tyenv' = if T = dummyT then tyenv else Sign.typ_match (Proof_Context.theory_of ctxt) (T, U) tyenv val eta_expand_cconv = case u of Abs _=> Thm.reflexive | _ => CConv.rewr_cconv @{thm eta_expand} fun add_ident NONE _ l = l | add_ident (SOME name) ct l = (name, Thm.term_of ct) :: l val abs_cv = CConv.abs_cconv (fn (ct, ctxt) => cv ctxt (tyenv', add_ident s ct ts)) ctxt in (eta_expand_cconv then_conv abs_cv) ct end handle Pattern.MATCH => raise TYPE ("abs_pconv: types don't match", [T,U], [u]) end fun fun_pconv cv ctxt tytenv ct = case Thm.term_of ct of _ $ _ => CConv.fun_cconv (cv ctxt tytenv) ct | Abs (_, T, _ $ Bound 0) => abs_pconv (fun_pconv cv) (NONE, T) ctxt tytenv ct | t => raise TERM ("fun_pconv", [t]) local fun arg_pconv_gen cv0 cv ctxt tytenv ct = case Thm.term_of ct of _ $ _ => cv0 (cv ctxt tytenv) ct | Abs (_, T, _ $ Bound 0) => abs_pconv (arg_pconv_gen cv0 cv) (NONE, T) ctxt tytenv ct | t => raise TERM ("arg_pconv_gen", [t]) in fun arg_pconv ctxt = arg_pconv_gen CConv.arg_cconv ctxt fun imp_pconv ctxt = arg_pconv_gen (CConv.concl_cconv 1) ctxt end (* Move to B in !!x_1 ... x_n. B. Do not eta-expand *) fun params_pconv cv ctxt tytenv ct = let val pconv = case Thm.term_of ct of Const (\<^const_name>\Pure.all\, _) $ Abs _ => (raw_arg_pconv o raw_abs_pconv) (fn _ => params_pconv cv) | Const (\<^const_name>\Pure.all\, _) => raw_arg_pconv (params_pconv cv) | _ => cv in pconv ctxt tytenv ct end fun forall_pconv cv ident ctxt tytenv ct = case Thm.term_of ct of Const (\<^const_name>\Pure.all\, T) $ _ => let val def_U = T |> dest_funT |> fst |> dest_funT |> fst val ident' = apsnd (the_default (def_U)) ident in arg_pconv (abs_pconv cv ident') ctxt tytenv ct end | t => raise TERM ("forall_pconv", [t]) fun all_pconv _ _ = Thm.reflexive fun for_pconv cv idents ctxt tytenv ct = let fun f rev_idents (Const (\<^const_name>\Pure.all\, _) $ t) = let val (rev_idents', cv') = f rev_idents (case t of Abs (_,_,u) => u | _ => t) in case rev_idents' of [] => ([], forall_pconv cv' (NONE, NONE)) | (x :: xs) => (xs, forall_pconv cv' x) end | f rev_idents _ = (rev_idents, cv) in case f (rev idents) (Thm.term_of ct) of ([], cv') => cv' ctxt tytenv ct | _ => raise CTERM ("for_pconv", [ct]) end fun concl_pconv cv ctxt tytenv ct = case Thm.term_of ct of (Const (\<^const_name>\Pure.imp\, _) $ _) $ _ => imp_pconv (concl_pconv cv) ctxt tytenv ct | _ => cv ctxt tytenv ct fun asm_pconv cv ctxt tytenv ct = case Thm.term_of ct of (Const (\<^const_name>\Pure.imp\, _) $ _) $ _ => CConv.with_prems_cconv ~1 (cv ctxt tytenv) ct | t => raise TERM ("asm_pconv", [t]) fun asms_pconv cv ctxt tytenv ct = case Thm.term_of ct of (Const (\<^const_name>\Pure.imp\, _) $ _) $ _ => ((CConv.with_prems_cconv ~1 oo cv) else_pconv imp_pconv (asms_pconv cv)) ctxt tytenv ct | t => raise TERM ("asms_pconv", [t]) fun judgment_pconv cv ctxt tytenv ct = if Object_Logic.is_judgment ctxt (Thm.term_of ct) then arg_pconv cv ctxt tytenv ct else cv ctxt tytenv ct fun in_pconv cv ctxt tytenv ct = (cv else_pconv raw_fun_pconv (in_pconv cv) else_pconv raw_arg_pconv (in_pconv cv) else_pconv raw_abs_pconv (fn _ => in_pconv cv)) ctxt tytenv ct fun replace_idents idents t = let fun subst ((n1, s)::ss) (t as Free (n2, _)) = if n1 = n2 then s else subst ss t | subst _ t = t in Term.map_aterms (subst idents) t end fun match_pconv cv (t,fixes) ctxt (tyenv, env_ts) ct = let val t' = replace_idents env_ts t val thy = Proof_Context.theory_of ctxt val u = Thm.term_of ct fun descend_hole fixes (Abs (_, _, t)) = (case descend_hole fixes t of NONE => NONE | SOME (fix :: fixes', pos) => SOME (fixes', abs_pconv pos fix) | SOME ([], _) => raise Match (* less fixes than abstractions on path to hole *)) | descend_hole fixes (t as l $ r) = let val (f, _) = strip_comb t in if is_hole f then SOME (fixes, cv) else (case descend_hole fixes l of SOME (fixes', pos) => SOME (fixes', fun_pconv pos) | NONE => (case descend_hole fixes r of SOME (fixes', pos) => SOME (fixes', arg_pconv pos) | NONE => NONE)) end | descend_hole fixes t = if is_hole t then SOME (fixes, cv) else NONE val to_hole = descend_hole (rev fixes) #> the_default ([], cv) #> snd in case try (Pattern.match thy (apply2 Logic.mk_term (t',u))) (tyenv, Vartab.empty) of NONE => raise TERM ("match_pconv: Does not match pattern", [t, t',u]) | SOME (tyenv', _) => to_hole t ctxt (tyenv', env_ts) ct end fun rewrs_pconv to thms ctxt (tyenv, env_ts) = let fun instantiate_normalize_env ctxt env thm = let val prop = Thm.prop_of thm val norm_type = Envir.norm_type o Envir.type_env val insts = Term.add_vars prop [] |> map (fn x as (s, T) => ((s, norm_type env T), Thm.cterm_of ctxt (Envir.norm_term env (Var x)))) val tyinsts = Term.add_tvars prop [] |> map (fn x => (x, Thm.ctyp_of ctxt (norm_type env (TVar x)))) - in Drule.instantiate_normalize (tyinsts, insts) thm end + in Drule.instantiate_normalize (TVars.make tyinsts, Vars.make insts) thm end fun unify_with_rhs context to env thm = let val (_, rhs) = thm |> Thm.concl_of |> Logic.dest_equals val env' = Pattern.unify context (Logic.mk_term to, Logic.mk_term rhs) env handle Pattern.Unif => raise NO_TO_MATCH in env' end fun inst_thm_to _ (NONE, _) thm = thm | inst_thm_to (ctxt : Proof.context) (SOME to, env) thm = instantiate_normalize_env ctxt (unify_with_rhs (Context.Proof ctxt) to env thm) thm fun inst_thm ctxt idents (to, tyenv) thm = let (* Replace any identifiers with their corresponding bound variables. *) val maxidx = Term.maxidx_typs (map (snd o snd) (Vartab.dest tyenv)) 0 val env = Envir.Envir {maxidx = maxidx, tenv = Vartab.empty, tyenv = tyenv} val maxidx = Envir.maxidx_of env |> fold Term.maxidx_term (the_list to) val thm' = Thm.incr_indexes (maxidx + 1) thm in SOME (inst_thm_to ctxt (Option.map (replace_idents idents) to, env) thm') end handle NO_TO_MATCH => NONE in CConv.rewrs_cconv (map_filter (inst_thm ctxt env_ts (to, tyenv)) thms) end fun rewrite_conv ctxt (pattern, to) thms ct = let fun apply_pat At = judgment_pconv | apply_pat In = in_pconv | apply_pat Asm = params_pconv o asms_pconv | apply_pat Concl = params_pconv o concl_pconv | apply_pat (For idents) = (fn cv => for_pconv cv (map (apfst SOME) idents)) | apply_pat (Term x) = (fn cv => match_pconv cv (apsnd (map (apfst SOME)) x)) val cv = fold_rev apply_pat pattern fun distinct_prems th = case Seq.pull (distinct_subgoals_tac th) of NONE => th | SOME (th', _) => th' val rewrite = rewrs_pconv to (maps (prep_meta_eq ctxt) thms) in cv rewrite ctxt (Vartab.empty, []) ct |> distinct_prems end fun rewrite_export_tac ctxt (pat, pat_ctxt) thms = let val export = case pat_ctxt of NONE => I | SOME inner => singleton (Proof_Context.export inner ctxt) in CCONVERSION (export o rewrite_conv ctxt pat thms) end val _ = Theory.setup let fun mk_fix s = (Binding.name s, NONE, NoSyn) val raw_pattern : (string, binding * string option * mixfix) pattern list parser = let val sep = (Args.$$$ "at" >> K At) || (Args.$$$ "in" >> K In) val atom = (Args.$$$ "asm" >> K Asm) || (Args.$$$ "concl" >> K Concl) || (Args.$$$ "for" |-- Args.parens (Scan.optional Parse.vars []) >> For) || (Parse.term >> Term) val sep_atom = sep -- atom >> (fn (s,a) => [s,a]) fun append_default [] = [Concl, In] | append_default (ps as Term _ :: _) = Concl :: In :: ps | append_default [For x, In] = [For x, Concl, In] | append_default (For x :: (ps as In :: Term _:: _)) = For x :: Concl :: ps | append_default ps = ps in Scan.repeats sep_atom >> (rev #> append_default) end fun context_lift (scan : 'a parser) f = fn (context : Context.generic, toks) => let val (r, toks') = scan toks val (r', context') = Context.map_proof_result (fn ctxt => f ctxt r) context in (r', (context', toks' : Token.T list)) end fun read_fixes fixes ctxt = let fun read_typ (b, rawT, mx) = (b, Option.map (Syntax.read_typ ctxt) rawT, mx) in Proof_Context.add_fixes (map read_typ fixes) ctxt end fun prep_pats ctxt (ps : (string, binding * string option * mixfix) pattern list) = let fun add_constrs ctxt n (Abs (x, T, t)) = let val (x', ctxt') = yield_singleton Proof_Context.add_fixes (mk_fix x) ctxt in (case add_constrs ctxt' (n+1) t of NONE => NONE | SOME ((ctxt'', n', xs), t') => let val U = Type_Infer.mk_param n [] val u = Type.constraint (U --> dummyT) (Abs (x, T, t')) in SOME ((ctxt'', n', (x', U) :: xs), u) end) end | add_constrs ctxt n (l $ r) = (case add_constrs ctxt n l of SOME (c, l') => SOME (c, l' $ r) | NONE => (case add_constrs ctxt n r of SOME (c, r') => SOME (c, l $ r') | NONE => NONE)) | add_constrs ctxt n t = if is_hole_const t then SOME ((ctxt, n, []), t) else NONE fun prep (Term s) (n, ctxt) = let val t = Syntax.parse_term ctxt s val ((ctxt', n', bs), t') = the_default ((ctxt, n, []), t) (add_constrs ctxt (n+1) t) in (Term (t', bs), (n', ctxt')) end | prep (For ss) (n, ctxt) = let val (ns, ctxt') = read_fixes ss ctxt in (For ns, (n, ctxt')) end | prep At (n,ctxt) = (At, (n, ctxt)) | prep In (n,ctxt) = (In, (n, ctxt)) | prep Concl (n,ctxt) = (Concl, (n, ctxt)) | prep Asm (n,ctxt) = (Asm, (n, ctxt)) val (xs, (_, ctxt')) = fold_map prep ps (0, ctxt) in (xs, ctxt') end fun prep_args ctxt (((raw_pats, raw_to), raw_ths)) = let fun check_terms ctxt ps to = let fun safe_chop (0: int) xs = ([], xs) | safe_chop n (x :: xs) = chop (n - 1) xs |>> cons x | safe_chop _ _ = raise Match fun reinsert_pat _ (Term (_, cs)) (t :: ts) = let val (cs', ts') = safe_chop (length cs) ts in (Term (t, map dest_Free cs'), ts') end | reinsert_pat _ (Term _) [] = raise Match | reinsert_pat ctxt (For ss) ts = let val fixes = map (fn s => (s, Variable.default_type ctxt s)) ss in (For fixes, ts) end | reinsert_pat _ At ts = (At, ts) | reinsert_pat _ In ts = (In, ts) | reinsert_pat _ Concl ts = (Concl, ts) | reinsert_pat _ Asm ts = (Asm, ts) fun free_constr (s,T) = Type.constraint T (Free (s, dummyT)) fun mk_free_constrs (Term (t, cs)) = t :: map free_constr cs | mk_free_constrs _ = [] val ts = maps mk_free_constrs ps @ the_list to |> Syntax.check_terms (hole_syntax ctxt) val ctxt' = fold Variable.declare_term ts ctxt val (ps', (to', ts')) = fold_map (reinsert_pat ctxt') ps ts ||> (fn xs => case to of NONE => (NONE, xs) | SOME _ => (SOME (hd xs), tl xs)) val _ = case ts' of (_ :: _) => raise Match | [] => () in ((ps', to'), ctxt') end val (pats, ctxt') = prep_pats ctxt raw_pats val ths = Attrib.eval_thms ctxt' raw_ths val to = Option.map (Syntax.parse_term ctxt') raw_to val ((pats', to'), ctxt'') = check_terms ctxt' pats to in ((pats', ths, (to', ctxt)), ctxt'') end val to_parser = Scan.option ((Args.$$$ "to") |-- Parse.term) val subst_parser = let val scan = raw_pattern -- to_parser -- Parse.thms1 in context_lift scan prep_args end in Method.setup \<^binding>\rewrite\ (subst_parser >> (fn (pattern, inthms, (to, pat_ctxt)) => fn orig_ctxt => SIMPLE_METHOD' (rewrite_export_tac orig_ctxt ((pattern, to), SOME pat_ctxt) inthms))) "single-step rewriting, allowing subterm selection via patterns." end end diff --git a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML --- a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML +++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML @@ -1,467 +1,467 @@ (* Title: HOL/Matrix_LP/Compute_Oracle/linker.ML Author: Steven Obua This module solves the problem that the computing oracle does not instantiate polymorphic rules. By going through the PCompute interface, all possible instantiations are resolved by compiling new programs, if necessary. The obvious disadvantage of this approach is that in the worst case for each new term to be rewritten, a new program may be compiled. *) (* Given constants/frees c_1::t_1, c_2::t_2, ...., c_n::t_n, and constants/frees d_1::d_1, d_2::s_2, ..., d_m::s_m Find all substitutions S such that a) the domain of S is tvars (t_1, ..., t_n) b) there are indices i_1, ..., i_k, and j_1, ..., j_k with 1. S (c_i_1::t_i_1) = d_j_1::s_j_1, ..., S (c_i_k::t_i_k) = d_j_k::s_j_k 2. tvars (t_i_1, ..., t_i_k) = tvars (t_1, ..., t_n) *) signature LINKER = sig exception Link of string datatype constant = Constant of bool * string * typ val constant_of : term -> constant type instances type subst = Type.tyenv val empty : constant list -> instances val typ_of_constant : constant -> typ val add_instances : theory -> instances -> constant list -> subst list * instances val substs_of : instances -> subst list val is_polymorphic : constant -> bool val distinct_constants : constant list -> constant list val collect_consts : term list -> constant list end structure Linker : LINKER = struct exception Link of string; type subst = Type.tyenv datatype constant = Constant of bool * string * typ fun constant_of (Const (name, ty)) = Constant (false, name, ty) | constant_of (Free (name, ty)) = Constant (true, name, ty) | constant_of _ = raise Link "constant_of" fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL) fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3)) fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2)) structure Consttab = Table(type key = constant val ord = constant_ord); structure ConsttabModTy = Table(type key = constant val ord = constant_modty_ord); fun typ_of_constant (Constant (_, _, ty)) = ty val empty_subst = (Vartab.empty : Type.tyenv) fun merge_subst (A:Type.tyenv) (B:Type.tyenv) = SOME (Vartab.fold (fn (v, t) => fn tab => (case Vartab.lookup tab v of NONE => Vartab.update (v, t) tab | SOME t' => if t = t' then tab else raise Type.TYPE_MATCH)) A B) handle Type.TYPE_MATCH => NONE fun subst_ord (A:Type.tyenv, B:Type.tyenv) = (list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B) structure Substtab = Table(type key = Type.tyenv val ord = subst_ord); fun substtab_union c = Substtab.fold Substtab.update c fun substtab_unions [] = Substtab.empty | substtab_unions [c] = c | substtab_unions (c::cs) = substtab_union c (substtab_unions cs) datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty [])) fun distinct_constants cs = Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty) fun empty cs = let val cs = distinct_constants (filter is_polymorphic cs) val old_cs = cs (* fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (Misc_Legacy.typ_tvars ty) tab val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty)) fun tvars_of ty = collect_tvars ty Typtab.empty val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs fun tyunion A B = Typtab.fold (fn (v,()) => fn tab => Typtab.update (v, case Typtab.lookup tab v of NONE => 1 | SOME n => n+1) tab) A B fun is_essential A B = Typtab.fold (fn (v, ()) => fn essential => essential orelse (case Typtab.lookup B v of NONE => raise Link "is_essential" | SOME n => n=1)) A false fun add_minimal (c', tvs') (tvs, cs) = let val tvs = tyunion tvs' tvs val cs = (c', tvs')::cs in if forall (fn (c',tvs') => is_essential tvs' tvs) cs then SOME (tvs, cs) else NONE end fun is_spanning (tvs, _) = (length (Typtab.keys tvs) = tvars_count) fun generate_minimal_subsets subsets [] = subsets | generate_minimal_subsets subsets (c::cs) = let val subsets' = map_filter (add_minimal c) subsets in generate_minimal_subsets (subsets@subsets') cs end*) val minimal_subsets = [old_cs] (*map (fn (tvs, cs) => map fst cs) (filter is_spanning (generate_minimal_subsets [(Typtab.empty, [])] cs))*) val constants = Consttab.keys (fold (fold (fn c => Consttab.update (c, ()))) minimal_subsets Consttab.empty) in Instances ( fold (fn c => fn tab => ConsttabModTy.update (c, ()) tab) constants ConsttabModTy.empty, Consttab.make (map (fn c => (c, Consttab.empty : Type.tyenv Consttab.table)) constants), minimal_subsets, Substtab.empty) end local fun calc ctab substtab [] = substtab | calc ctab substtab (c::cs) = let val csubsts = map snd (Consttab.dest (the (Consttab.lookup ctab c))) fun merge_substs substtab subst = Substtab.fold (fn (s,_) => fn tab => (case merge_subst subst s of NONE => tab | SOME s => Substtab.update (s, ()) tab)) substtab Substtab.empty val substtab = substtab_unions (map (merge_substs substtab) csubsts) in calc ctab substtab cs end in fun calc_substs ctab (cs:constant list) = calc ctab (Substtab.update (empty_subst, ()) Substtab.empty) cs end fun add_instances thy (Instances (cfilter, ctab,minsets,substs)) cs = let (* val _ = writeln (makestring ("add_instances: ", length_cs, length cs, length (Consttab.keys ctab)))*) fun calc_instantiations (constant as Constant (free, name, ty)) instantiations = Consttab.fold (fn (constant' as Constant (free', name', ty'), insttab) => fn instantiations => if free <> free' orelse name <> name' then instantiations else case Consttab.lookup insttab constant of SOME _ => instantiations | NONE => ((constant', (constant, Sign.typ_match thy (ty', ty) empty_subst))::instantiations handle Type.TYPE_MATCH => instantiations)) ctab instantiations val instantiations = fold calc_instantiations cs [] (*val _ = writeln ("instantiations = "^(makestring (length instantiations)))*) fun update_ctab (constant', entry) ctab = (case Consttab.lookup ctab constant' of NONE => raise Link "internal error: update_ctab" | SOME tab => Consttab.update (constant', Consttab.update entry tab) ctab) val ctab = fold update_ctab instantiations ctab val new_substs = fold (fn minset => fn substs => substtab_union (calc_substs ctab minset) substs) minsets Substtab.empty val (added_substs, substs) = Substtab.fold (fn (ns, _) => fn (added, substtab) => (case Substtab.lookup substs ns of NONE => (ns::added, Substtab.update (ns, ()) substtab) | SOME () => (added, substtab))) new_substs ([], substs) in (added_substs, Instances (cfilter, ctab, minsets, substs)) end fun substs_of (Instances (_,_,_,substs)) = Substtab.keys substs local fun collect (Var _) tab = tab | collect (Bound _) tab = tab | collect (a $ b) tab = collect b (collect a tab) | collect (Abs (_, _, body)) tab = collect body tab | collect t tab = Consttab.update (constant_of t, ()) tab in fun collect_consts tms = Consttab.keys (fold collect tms Consttab.empty) end end signature PCOMPUTE = sig type pcomputer val make : Compute.machine -> theory -> thm list -> Linker.constant list -> pcomputer val make_with_cache : Compute.machine -> theory -> term list -> thm list -> Linker.constant list -> pcomputer val add_instances : pcomputer -> Linker.constant list -> bool val add_instances' : pcomputer -> term list -> bool val rewrite : pcomputer -> cterm list -> thm list val simplify : pcomputer -> Compute.theorem -> thm val make_theorem : pcomputer -> thm -> string list -> Compute.theorem val instantiate : pcomputer -> (string * cterm) list -> Compute.theorem -> Compute.theorem val evaluate_prem : pcomputer -> int -> Compute.theorem -> Compute.theorem val modus_ponens : pcomputer -> int -> thm -> Compute.theorem -> Compute.theorem end structure PCompute : PCOMPUTE = struct exception PCompute of string datatype theorem = MonoThm of thm | PolyThm of thm * Linker.instances * thm list datatype pattern = MonoPattern of term | PolyPattern of term * Linker.instances * term list datatype pcomputer = PComputer of theory * Compute.computer * theorem list Unsynchronized.ref * pattern list Unsynchronized.ref (*fun collect_consts (Var x) = [] | collect_consts (Bound _) = [] | collect_consts (a $ b) = (collect_consts a)@(collect_consts b) | collect_consts (Abs (_, _, body)) = collect_consts body | collect_consts t = [Linker.constant_of t]*) fun computer_of (PComputer (_,computer,_,_)) = computer fun collect_consts_of_thm th = let val th = Thm.prop_of th val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th) val (left, right) = Logic.dest_equals th in (Linker.collect_consts [left], Linker.collect_consts (right::prems)) end fun create_theorem th = let val (left, right) = collect_consts_of_thm th val polycs = filter Linker.is_polymorphic left val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (Misc_Legacy.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty fun check_const (c::cs) cs' = let val tvars = Misc_Legacy.typ_tvars (Linker.typ_of_constant c) val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false in if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side" else if null (tvars) then check_const cs (c::cs') else check_const cs cs' end | check_const [] cs' = cs' val monocs = check_const right [] in if null (polycs) then (monocs, MonoThm th) else (monocs, PolyThm (th, Linker.empty polycs, [])) end fun create_pattern pat = let val cs = Linker.collect_consts [pat] val polycs = filter Linker.is_polymorphic cs in if null (polycs) then MonoPattern pat else PolyPattern (pat, Linker.empty polycs, []) end fun create_computer machine thy pats ths = let fun add (MonoThm th) ths = th::ths | add (PolyThm (_, _, ths')) ths = ths'@ths fun addpat (MonoPattern p) pats = p::pats | addpat (PolyPattern (_, _, ps)) pats = ps@pats val ths = fold_rev add ths [] val pats = fold_rev addpat pats [] in Compute.make_with_cache machine thy pats ths end fun update_computer computer pats ths = let fun add (MonoThm th) ths = th::ths | add (PolyThm (_, _, ths')) ths = ths'@ths fun addpat (MonoPattern p) pats = p::pats | addpat (PolyPattern (_, _, ps)) pats = ps@pats val ths = fold_rev add ths [] val pats = fold_rev addpat pats [] in Compute.update_with_cache computer pats ths end fun conv_subst thy (subst : Type.tyenv) = map (fn (iname, (sort, ty)) => ((iname, sort), Thm.global_ctyp_of thy ty)) (Vartab.dest subst) fun add_monos thy monocs pats ths = let val changed = Unsynchronized.ref false fun add monocs (th as (MonoThm _)) = ([], th) | add monocs (PolyThm (th, instances, instanceths)) = let val (newsubsts, instances) = Linker.add_instances thy instances monocs val _ = if not (null newsubsts) then changed := true else () - val newths = map (fn subst => Thm.instantiate (conv_subst thy subst, []) th) newsubsts + val newths = map (fn subst => Thm.instantiate (TVars.make (conv_subst thy subst), Vars.empty) th) newsubsts (* val _ = if not (null newths) then (print ("added new theorems: ", newths); ()) else ()*) val newmonos = fold (fn th => fn monos => (snd (collect_consts_of_thm th))@monos) newths [] in (newmonos, PolyThm (th, instances, instanceths@newths)) end fun addpats monocs (pat as (MonoPattern _)) = pat | addpats monocs (PolyPattern (p, instances, instancepats)) = let val (newsubsts, instances) = Linker.add_instances thy instances monocs val _ = if not (null newsubsts) then changed := true else () val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts in PolyPattern (p, instances, instancepats@newpats) end fun step monocs ths = fold_rev (fn th => fn (newmonos, ths) => let val (newmonos', th') = add monocs th in (newmonos'@newmonos, th'::ths) end) ths ([], []) fun loop monocs pats ths = let val (monocs', ths') = step monocs ths val pats' = map (addpats monocs) pats in if null (monocs') then (pats', ths') else loop monocs' pats' ths' end val result = loop monocs pats ths in (!changed, result) end datatype cthm = ComputeThm of term list * sort list * term fun thm2cthm th = ComputeThm (Thm.hyps_of th, Thm.shyps_of th, Thm.prop_of th) val cthm_ord' = prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2)) structure CThmtab = Table(type key = cthm val ord = cthm_ord) fun remove_duplicates ths = let val counter = Unsynchronized.ref 0 val tab = Unsynchronized.ref (CThmtab.empty : unit CThmtab.table) val thstab = Unsynchronized.ref (Inttab.empty : thm Inttab.table) fun update th = let val key = thm2cthm th in case CThmtab.lookup (!tab) key of NONE => ((tab := CThmtab.update_new (key, ()) (!tab)); thstab := Inttab.update_new (!counter, th) (!thstab); counter := !counter + 1) | _ => () end val _ = map update ths in map snd (Inttab.dest (!thstab)) end fun make_with_cache machine thy pats ths cs = let val ths = remove_duplicates ths val (monocs, ths) = fold_rev (fn th => fn (monocs, ths) => let val (m, t) = create_theorem th in (m@monocs, t::ths) end) ths (cs, []) val pats = map create_pattern pats val (_, (pats, ths)) = add_monos thy monocs pats ths val computer = create_computer machine thy pats ths in PComputer (thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats) end fun make machine thy ths cs = make_with_cache machine thy [] ths cs fun add_instances (PComputer (thy, computer, rths, rpats)) cs = let val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths) in if changed then (update_computer computer pats ths; rths := ths; rpats := pats; true) else false end fun add_instances' pc ts = add_instances pc (Linker.collect_consts ts) fun rewrite pc cts = let val _ = add_instances' pc (map Thm.term_of cts) val computer = (computer_of pc) in map (fn ct => Compute.rewrite computer ct) cts end fun simplify pc th = Compute.simplify (computer_of pc) th fun make_theorem pc th vars = let val _ = add_instances' pc [Thm.prop_of th] in Compute.make_theorem (computer_of pc) th vars end fun instantiate pc insts th = let val _ = add_instances' pc (map (Thm.term_of o snd) insts) in Compute.instantiate (computer_of pc) insts th end fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th fun modus_ponens pc prem_no th' th = let val _ = add_instances' pc [Thm.prop_of th'] in Compute.modus_ponens (computer_of pc) prem_no th' th end end diff --git a/src/HOL/Nominal/nominal_datatype.ML b/src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML +++ b/src/HOL/Nominal/nominal_datatype.ML @@ -1,2109 +1,2109 @@ (* Title: HOL/Nominal/nominal_datatype.ML Author: Stefan Berghofer and Christian Urban, TU Muenchen Nominal datatype package for Isabelle/HOL. *) signature NOMINAL_DATATYPE = sig val nominal_datatype : Old_Datatype_Aux.config -> Old_Datatype.spec list -> theory -> theory val nominal_datatype_cmd : Old_Datatype_Aux.config -> Old_Datatype.spec_cmd list -> theory -> theory type descr type nominal_datatype_info val get_nominal_datatypes : theory -> nominal_datatype_info Symtab.table val get_nominal_datatype : theory -> string -> nominal_datatype_info option val mk_perm: typ list -> term -> term -> term val perm_of_pair: term * term -> term val mk_not_sym: thm list -> thm list val perm_simproc: simproc val fresh_const: typ -> typ -> term val fresh_star_const: typ -> typ -> term end structure NominalDatatype : NOMINAL_DATATYPE = struct val finite_emptyI = @{thm finite.emptyI}; val finite_Diff = @{thm finite_Diff}; val finite_Un = @{thm finite_Un}; val Un_iff = @{thm Un_iff}; val Un_assoc = @{thm Un_assoc}; val Collect_disj_eq = @{thm Collect_disj_eq}; val Collect_False_empty = @{thm empty_def [THEN sym, THEN eq_reflection]}; val empty_iff = @{thm empty_iff}; open NominalAtoms; (* theory data *) type descr = (int * (string * Old_Datatype_Aux.dtyp list * (string * (Old_Datatype_Aux.dtyp list * Old_Datatype_Aux.dtyp) list) list)) list; type nominal_datatype_info = {index : int, descr : descr, rec_names : string list, rec_rewrites : thm list, induction : thm, distinct : thm list, inject : thm list}; structure NominalDatatypesData = Theory_Data ( type T = nominal_datatype_info Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data = Symtab.merge (K true) data; ); val get_nominal_datatypes = NominalDatatypesData.get; val put_nominal_datatypes = NominalDatatypesData.put; val map_nominal_datatypes = NominalDatatypesData.map; val get_nominal_datatype = Symtab.lookup o get_nominal_datatypes; (**** make datatype info ****) fun make_dt_info descr induct reccomb_names rec_thms (i, (((_, (tname, _, _)), distinct), inject)) = (tname, {index = i, descr = descr, rec_names = reccomb_names, rec_rewrites = rec_thms, induction = induct, distinct = distinct, inject = inject}); (*******************************) val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (Thm.prems_of Old_Datatype.distinct_lemma); (** simplification procedure for sorting permutations **) val dj_cp = @{thm dj_cp}; fun dest_permT (Type (\<^type_name>\fun\, [Type (\<^type_name>\list\, [Type (\<^type_name>\Product_Type.prod\, [T, _])]), Type (\<^type_name>\fun\, [_, U])])) = (T, U); fun permTs_of (Const (\<^const_name>\Nominal.perm\, T) $ t $ u) = fst (dest_permT T) :: permTs_of u | permTs_of _ = []; fun perm_simproc' ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, T) $ t $ (u as Const (\<^const_name>\Nominal.perm\, U) $ r $ s) => let val thy = Proof_Context.theory_of ctxt; val (aT as Type (a, []), S) = dest_permT T; val (bT as Type (b, []), _) = dest_permT U; in if member (op =) (permTs_of u) aT andalso aT <> bT then let val cp = cp_inst_of thy a b; val dj = dj_thm_of thy b a; val dj_cp' = [cp, dj] MRS dj_cp; val cert = SOME o Thm.cterm_of ctxt in SOME (mk_meta_eq (Thm.instantiate' [SOME (Thm.ctyp_of ctxt S)] [cert t, cert r, cert s] dj_cp')) end else NONE end | _ => NONE); val perm_simproc = Simplifier.make_simproc \<^context> "perm_simp" {lhss = [\<^term>\pi1 \ (pi2 \ x)\], proc = K perm_simproc'}; fun projections ctxt rule = Project_Rule.projections ctxt rule |> map (Drule.export_without_context #> Rule_Cases.save rule); val supp_prod = @{thm supp_prod}; val fresh_prod = @{thm fresh_prod}; val supports_fresh = @{thm supports_fresh}; val supports_def = Simpdata.mk_eq @{thm Nominal.supports_def}; val fresh_def = Simpdata.mk_eq @{thm fresh_def}; val supp_def = Simpdata.mk_eq @{thm supp_def}; val rev_simps = @{thms rev.simps}; val app_simps = @{thms append.simps}; val at_fin_set_supp = @{thm at_fin_set_supp}; val at_fin_set_fresh = @{thm at_fin_set_fresh}; val abs_fun_eq1 = @{thm abs_fun_eq1}; fun collect_simp ctxt = rewrite_rule ctxt [mk_meta_eq mem_Collect_eq]; fun mk_perm Ts t u = let val T = fastype_of1 (Ts, t); val U = fastype_of1 (Ts, u) in Const (\<^const_name>\Nominal.perm\, T --> U --> U) $ t $ u end; fun perm_of_pair (x, y) = let val T = fastype_of x; val pT = mk_permT T in Const (\<^const_name>\Cons\, HOLogic.mk_prodT (T, T) --> pT --> pT) $ HOLogic.mk_prod (x, y) $ Const (\<^const_name>\Nil\, pT) end; fun mk_not_sym ths = maps (fn th => (case Thm.prop_of th of _ $ (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) => [th, th RS not_sym] | _ => [th])) ths; fun fresh_const T U = Const (\<^const_name>\Nominal.fresh\, T --> U --> HOLogic.boolT); fun fresh_star_const T U = Const (\<^const_name>\Nominal.fresh_star\, HOLogic.mk_setT T --> U --> HOLogic.boolT); fun gen_nominal_datatype prep_specs (config: Old_Datatype_Aux.config) dts thy = let val new_type_names = map (fn ((tname, _, _), _) => Binding.name_of tname) dts; val (dts', _) = prep_specs dts thy; val atoms = atoms_of thy; val tyvars = map (fn ((_, tvs, _), _) => tvs) dts'; val sorts = flat tyvars; fun inter_sort thy S S' = Sign.inter_sort thy (S, S'); fun augment_sort_typ thy S = let val S = Sign.minimize_sort thy (Sign.certify_sort thy S) in map_type_tfree (fn (s, S') => TFree (s, if member (op = o apsnd fst) sorts s then inter_sort thy S S' else S')) end; fun augment_sort thy S = map_types (augment_sort_typ thy S); val types_syntax = map (fn ((tname, tvs, mx), constrs) => (tname, mx)) dts'; val constr_syntax = map (fn (_, constrs) => map (fn (cname, cargs, mx) => (cname, mx)) constrs) dts'; val ps = map (fn ((n, _, _), _) => (Sign.full_name thy n, Sign.full_name thy (Binding.suffix_name "_Rep" n))) dts; val rps = map Library.swap ps; fun replace_types (Type (\<^type_name>\ABS\, [T, U])) = Type (\<^type_name>\fun\, [T, Type (\<^type_name>\noption\, [replace_types U])]) | replace_types (Type (s, Ts)) = Type (the_default s (AList.lookup op = ps s), map replace_types Ts) | replace_types T = T; val dts'' = map (fn ((tname, tvs, mx), constrs) => ((Binding.suffix_name "_Rep" tname, tvs, NoSyn), map (fn (cname, cargs, mx) => (Binding.suffix_name "_Rep" cname, map replace_types cargs, NoSyn)) constrs)) dts'; val new_type_names' = map (fn n => n ^ "_Rep") new_type_names; val (full_new_type_names',thy1) = BNF_LFP_Compat.add_datatype [BNF_LFP_Compat.Kill_Type_Args] dts'' thy; val {descr, induct, ...} = BNF_LFP_Compat.the_info thy1 [] (hd full_new_type_names'); fun nth_dtyp i = Old_Datatype_Aux.typ_of_dtyp descr (Old_Datatype_Aux.DtRec i); val big_name = space_implode "_" new_type_names; (**** define permutation functions ****) val permT = mk_permT (TFree ("'x", \<^sort>\type\)); val pi = Free ("pi", permT); val perm_types = map (fn (i, _) => let val T = nth_dtyp i in permT --> T --> T end) descr; val perm_names' = Old_Datatype_Prop.indexify_names (map (fn (i, _) => "perm_" ^ Old_Datatype_Aux.name_of_typ (nth_dtyp i)) descr); val perm_names = replicate (length new_type_names) \<^const_name>\Nominal.perm\ @ map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names)); val perm_names_types = perm_names ~~ perm_types; val perm_names_types' = perm_names' ~~ perm_types; val perm_eqs = maps (fn (i, (_, _, constrs)) => let val T = nth_dtyp i in map (fn (cname, dts) => let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts; val names = Name.variant_list ["pi"] (Old_Datatype_Prop.make_tnames Ts); val args = map Free (names ~~ Ts); val c = Const (cname, Ts ---> T); fun perm_arg (dt, x) = let val T = type_of x in if Old_Datatype_Aux.is_rec_type dt then let val Us = binder_types T in fold_rev (Term.abs o pair "x") Us (Free (nth perm_names_types' (Old_Datatype_Aux.body_index dt)) $ pi $ list_comb (x, map (fn (i, U) => Const (\<^const_name>\Nominal.perm\, permT --> U --> U) $ (Const (\<^const_name>\rev\, permT --> permT) $ pi) $ Bound i) ((length Us - 1 downto 0) ~~ Us))) end else Const (\<^const_name>\Nominal.perm\, permT --> T --> T) $ pi $ x end; in ((Binding.empty_atts, HOLogic.mk_Trueprop (HOLogic.mk_eq (Free (nth perm_names_types' i) $ Free ("pi", mk_permT (TFree ("'x", \<^sort>\type\))) $ list_comb (c, args), list_comb (c, map perm_arg (dts ~~ args))))), [], []) end) constrs end) descr; val (perm_simps, thy2) = BNF_LFP_Compat.primrec_overloaded (map (fn (s, sT) => (s, sT, false)) (List.take (perm_names' ~~ perm_names_types, length new_type_names))) (map (fn s => (Binding.name s, NONE, NoSyn)) perm_names') perm_eqs thy1; (**** prove that permutation functions introduced by unfolding are ****) (**** equivalent to already existing permutation functions ****) val _ = warning ("length descr: " ^ string_of_int (length descr)); val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); val perm_indnames = Old_Datatype_Prop.make_tnames (map body_type perm_types); val perm_fun_def = Simpdata.mk_eq @{thm perm_fun_def}; val unfolded_perm_eq_thms = if length descr = length new_type_names then [] else map Drule.export_without_context (List.drop (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy2 [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (c as (s, T), x) => let val [T1, T2] = binder_types T in HOLogic.mk_eq (Const c $ pi $ Free (x, T2), Const (\<^const_name>\Nominal.perm\, T) $ pi $ Free (x, T2)) end) (perm_names_types ~~ perm_indnames)))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1, ALLGOALS (asm_full_simp_tac (ctxt addsimps [perm_fun_def]))])), length new_type_names)); (**** prove [] \ t = t ****) val _ = warning "perm_empty_thms"; val perm_empty_thms = maps (fn a => let val permT = mk_permT (Type (a, [])) in map Drule.export_without_context (List.take (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy2 [] [] (augment_sort thy2 [pt_class_of thy2 a] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn ((s, T), x) => HOLogic.mk_eq (Const (s, permT --> T --> T) $ Const (\<^const_name>\Nil\, permT) $ Free (x, T), Free (x, T))) (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1, ALLGOALS (asm_full_simp_tac ctxt)])), length new_type_names)) end) atoms; (**** prove (pi1 @ pi2) \ t = pi1 \ (pi2 \ t) ****) val _ = warning "perm_append_thms"; (*FIXME: these should be looked up statically*) val at_pt_inst = Global_Theory.get_thm thy2 "at_pt_inst"; val pt2 = Global_Theory.get_thm thy2 "pt2"; val perm_append_thms = maps (fn a => let val permT = mk_permT (Type (a, [])); val pi1 = Free ("pi1", permT); val pi2 = Free ("pi2", permT); val pt_inst = pt_inst_of thy2 a; val pt2' = pt_inst RS pt2; val pt2_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "2") a); in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy2 [] [] (augment_sort thy2 [pt_class_of thy2 a] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn ((s, T), x) => let val perm = Const (s, permT --> T --> T) in HOLogic.mk_eq (perm $ (Const (\<^const_name>\append\, permT --> permT --> permT) $ pi1 $ pi2) $ Free (x, T), perm $ pi1 $ (perm $ pi2 $ Free (x, T))) end) (perm_names ~~ map body_type perm_types ~~ perm_indnames))))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1, ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt2', pt2_ax]))]))), length new_type_names) end) atoms; (**** prove pi1 ~ pi2 ==> pi1 \ t = pi2 \ t ****) val _ = warning "perm_eq_thms"; val pt3 = Global_Theory.get_thm thy2 "pt3"; val pt3_rev = Global_Theory.get_thm thy2 "pt3_rev"; val perm_eq_thms = maps (fn a => let val permT = mk_permT (Type (a, [])); val pi1 = Free ("pi1", permT); val pi2 = Free ("pi2", permT); val at_inst = at_inst_of thy2 a; val pt_inst = pt_inst_of thy2 a; val pt3' = pt_inst RS pt3; val pt3_rev' = at_inst RS (pt_inst RS pt3_rev); val pt3_ax = Global_Theory.get_thm thy2 (Long_Name.map_base_name (fn s => "pt_" ^ s ^ "3") a); in List.take (map Drule.export_without_context (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy2 [] [] (augment_sort thy2 [pt_class_of thy2 a] (Logic.mk_implies (HOLogic.mk_Trueprop (Const (\<^const_name>\Nominal.prm_eq\, permT --> permT --> HOLogic.boolT) $ pi1 $ pi2), HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn ((s, T), x) => let val perm = Const (s, permT --> T --> T) in HOLogic.mk_eq (perm $ pi1 $ Free (x, T), perm $ pi2 $ Free (x, T)) end) (perm_names ~~ map body_type perm_types ~~ perm_indnames)))))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1, ALLGOALS (asm_full_simp_tac (ctxt addsimps [pt3', pt3_rev', pt3_ax]))]))), length new_type_names) end) atoms; (**** prove pi1 \ (pi2 \ t) = (pi1 \ pi2) \ (pi1 \ t) ****) val cp1 = Global_Theory.get_thm thy2 "cp1"; val dj_cp = Global_Theory.get_thm thy2 "dj_cp"; val pt_perm_compose = Global_Theory.get_thm thy2 "pt_perm_compose"; val pt_perm_compose_rev = Global_Theory.get_thm thy2 "pt_perm_compose_rev"; val dj_perm_perm_forget = Global_Theory.get_thm thy2 "dj_perm_perm_forget"; fun composition_instance name1 name2 thy = let val cp_class = cp_class_of thy name1 name2; val pt_class = if name1 = name2 then [pt_class_of thy name1] else []; val permT1 = mk_permT (Type (name1, [])); val permT2 = mk_permT (Type (name2, [])); val Ts = map body_type perm_types; val cp_inst = cp_inst_of thy name1 name2; fun simps ctxt = ctxt addsimps (perm_fun_def :: (if name1 <> name2 then let val dj = dj_thm_of thy name2 name1 in [dj RS (cp_inst RS dj_cp), dj RS dj_perm_perm_forget] end else let val at_inst = at_inst_of thy name1; val pt_inst = pt_inst_of thy name1; in [cp_inst RS cp1 RS sym, at_inst RS (pt_inst RS pt_perm_compose) RS sym, at_inst RS (pt_inst RS pt_perm_compose_rev) RS sym] end)) val sort = Sign.minimize_sort thy (Sign.certify_sort thy (cp_class :: pt_class)); val thms = Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy [] [] (augment_sort thy sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn ((s, T), x) => let val pi1 = Free ("pi1", permT1); val pi2 = Free ("pi2", permT2); val perm1 = Const (s, permT1 --> T --> T); val perm2 = Const (s, permT2 --> T --> T); val perm3 = Const (\<^const_name>\Nominal.perm\, permT1 --> permT2 --> permT2) in HOLogic.mk_eq (perm1 $ pi1 $ (perm2 $ pi2 $ Free (x, T)), perm2 $ (perm3 $ pi1 $ pi2) $ (perm1 $ pi1 $ Free (x, T))) end) (perm_names ~~ Ts ~~ perm_indnames))))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt induct perm_indnames 1, ALLGOALS (asm_full_simp_tac (simps ctxt))])) in fold (fn (s, tvs) => fn thy => Axclass.prove_arity (s, map (inter_sort thy sort o snd) tvs, [cp_class]) (fn ctxt' => Class.intro_classes_tac ctxt' [] THEN ALLGOALS (resolve_tac ctxt' thms)) thy) (full_new_type_names' ~~ tyvars) thy end; val (perm_thmss,thy3) = thy2 |> fold (fn name1 => fold (composition_instance name1) atoms) atoms |> fold (fn atom => fn thy => let val pt_name = pt_class_of thy atom in fold (fn (s, tvs) => fn thy => Axclass.prove_arity (s, map (inter_sort thy [pt_name] o snd) tvs, [pt_name]) (fn ctxt => EVERY [Class.intro_classes_tac ctxt [], resolve_tac ctxt perm_empty_thms 1, resolve_tac ctxt perm_append_thms 1, resolve_tac ctxt perm_eq_thms 1, assume_tac ctxt 1]) thy) (full_new_type_names' ~~ tyvars) thy end) atoms |> Global_Theory.add_thmss [((Binding.name (space_implode "_" new_type_names ^ "_unfolded_perm_eq"), unfolded_perm_eq_thms), [Simplifier.simp_add]), ((Binding.name (space_implode "_" new_type_names ^ "_perm_empty"), perm_empty_thms), [Simplifier.simp_add]), ((Binding.name (space_implode "_" new_type_names ^ "_perm_append"), perm_append_thms), [Simplifier.simp_add]), ((Binding.name (space_implode "_" new_type_names ^ "_perm_eq"), perm_eq_thms), [Simplifier.simp_add])]; (**** Define representing sets ****) val _ = warning "representing sets"; val rep_set_names = Old_Datatype_Prop.indexify_names (map (fn (i, _) => Old_Datatype_Aux.name_of_typ (nth_dtyp i) ^ "_set") descr); val big_rep_name = space_implode "_" (Old_Datatype_Prop.indexify_names (map_filter (fn (i, (\<^type_name>\noption\, _, _)) => NONE | (i, _) => SOME (Old_Datatype_Aux.name_of_typ (nth_dtyp i))) descr)) ^ "_set"; val _ = warning ("big_rep_name: " ^ big_rep_name); fun strip_option (dtf as Old_Datatype_Aux.DtType ("fun", [dt, Old_Datatype_Aux.DtRec i])) = (case AList.lookup op = descr i of SOME (\<^type_name>\noption\, _, [(_, [dt']), _]) => apfst (cons dt) (strip_option dt') | _ => ([], dtf)) | strip_option (Old_Datatype_Aux.DtType ("fun", [dt, Old_Datatype_Aux.DtType (\<^type_name>\noption\, [dt'])])) = apfst (cons dt) (strip_option dt') | strip_option dt = ([], dt); val dt_atomTs = distinct op = (map (Old_Datatype_Aux.typ_of_dtyp descr) (maps (fn (_, (_, _, cs)) => maps (maps (fst o strip_option) o snd) cs) descr)); val dt_atoms = map (fst o dest_Type) dt_atomTs; fun make_intr s T (cname, cargs) = let fun mk_prem dt (j, j', prems, ts) = let val (dts, dt') = strip_option dt; val (dts', dt'') = Old_Datatype_Aux.strip_dtyp dt'; val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr) dts; val Us = map (Old_Datatype_Aux.typ_of_dtyp descr) dts'; val T = Old_Datatype_Aux.typ_of_dtyp descr dt''; val free = Old_Datatype_Aux.mk_Free "x" (Us ---> T) j; val free' = Old_Datatype_Aux.app_bnds free (length Us); fun mk_abs_fun T (i, t) = let val U = fastype_of t in (i + 1, Const (\<^const_name>\Nominal.abs_fun\, [T, U, T] ---> Type (\<^type_name>\noption\, [U])) $ Old_Datatype_Aux.mk_Free "y" T i $ t) end in (j + 1, j' + length Ts, case dt'' of Old_Datatype_Aux.DtRec k => Logic.list_all (map (pair "x") Us, HOLogic.mk_Trueprop (Free (nth rep_set_names k, T --> HOLogic.boolT) $ free')) :: prems | _ => prems, snd (fold_rev mk_abs_fun Ts (j', free)) :: ts) end; val (_, _, prems, ts) = fold_rev mk_prem cargs (1, 1, [], []); val concl = HOLogic.mk_Trueprop (Free (s, T --> HOLogic.boolT) $ list_comb (Const (cname, map fastype_of ts ---> T), ts)) in Logic.list_implies (prems, concl) end; val (intr_ts, (rep_set_names', recTs')) = apfst flat (apsnd ListPair.unzip (ListPair.unzip (map_filter (fn ((_, (\<^type_name>\noption\, _, _)), _) => NONE | ((i, (_, _, constrs)), rep_set_name) => let val T = nth_dtyp i in SOME (map (make_intr rep_set_name T) constrs, (rep_set_name, T)) end) (descr ~~ rep_set_names)))); val rep_set_names'' = map (Sign.full_bname thy3) rep_set_names'; val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy4) = thy3 |> Sign.concealed |> Named_Target.theory_map_result Inductive.transform_result (Inductive.add_inductive {quiet_mode = false, verbose = false, alt_name = Binding.name big_rep_name, coind = false, no_elim = true, no_ind = false, skip_mono = true} (map (fn (s, T) => ((Binding.name s, T --> HOLogic.boolT), NoSyn)) (rep_set_names' ~~ recTs')) [] (map (fn x => (Binding.empty_atts, x)) intr_ts) []) ||> Sign.restore_naming thy3; (**** Prove that representing set is closed under permutation ****) val _ = warning "proving closure under permutation..."; val abs_perm = Global_Theory.get_thms thy4 "abs_perm"; val perm_indnames' = map_filter (fn (x, (_, (\<^type_name>\noption\, _, _))) => NONE | (x, _) => SOME x) (perm_indnames ~~ descr); fun mk_perm_closed name = map (fn th => Drule.export_without_context (th RS mp)) (List.take (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy4 [] [] (augment_sort thy4 (pt_class_of thy4 name :: map (cp_class_of thy4 name) (remove (op =) name dt_atoms)) (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn ((s, T), x) => let val S = Const (s, T --> HOLogic.boolT); val permT = mk_permT (Type (name, [])) in HOLogic.mk_imp (S $ Free (x, T), S $ (Const (\<^const_name>\Nominal.perm\, permT --> T --> T) $ Free ("pi", permT) $ Free (x, T))) end) (rep_set_names'' ~~ recTs' ~~ perm_indnames'))))) (fn {context = ctxt, ...} => EVERY [Old_Datatype_Aux.ind_tac ctxt rep_induct [] 1, ALLGOALS (simp_tac (ctxt addsimps (Thm.symmetric perm_fun_def :: abs_perm))), ALLGOALS (resolve_tac ctxt rep_intrs THEN_ALL_NEW assume_tac ctxt)])), length new_type_names)); val perm_closed_thmss = map mk_perm_closed atoms; (**** typedef ****) val _ = warning "defining type..."; val (typedefs, thy6) = thy4 |> fold_map (fn (((name, mx), tvs), (cname, U)) => fn thy => Named_Target.theory_map_result (apsnd o Typedef.transform_info) (Typedef.add_typedef {overloaded = false} (name, map (fn (v, _) => (v, dummyS)) tvs, mx) (* FIXME keep constraints!? *) (Const (\<^const_name>\Collect\, (U --> HOLogic.boolT) --> HOLogic.mk_setT U) $ Const (cname, U --> HOLogic.boolT)) NONE (fn ctxt => resolve_tac ctxt [exI] 1 THEN resolve_tac ctxt [CollectI] 1 THEN QUIET_BREADTH_FIRST (has_fewer_prems 1) (resolve_tac ctxt rep_intrs 1))) thy |> (fn ((_, r), thy) => let val permT = mk_permT (TFree (singleton (Name.variant_list (map fst tvs)) "'a", \<^sort>\type\)); val pi = Free ("pi", permT); val T = Type (Sign.full_name thy name, map TFree tvs); in apfst (pair r o hd) (Global_Theory.add_defs_unchecked true [((Binding.map_name (fn n => "prm_" ^ n ^ "_def") name, Logic.mk_equals (Const (\<^const_name>\Nominal.perm\, permT --> T --> T) $ pi $ Free ("x", T), Const (Sign.intern_const thy ("Abs_" ^ Binding.name_of name), U --> T) $ (Const (\<^const_name>\Nominal.perm\, permT --> U --> U) $ pi $ (Const (Sign.intern_const thy ("Rep_" ^ Binding.name_of name), T --> U) $ Free ("x", T))))), [])] thy) end)) (types_syntax ~~ tyvars ~~ List.take (rep_set_names'' ~~ recTs', length new_type_names)); val ctxt6 = Proof_Context.init_global thy6; val perm_defs = map snd typedefs; val Abs_inverse_thms = map (collect_simp ctxt6 o #Abs_inverse o snd o fst) typedefs; val Rep_inverse_thms = map (#Rep_inverse o snd o fst) typedefs; val Rep_thms = map (collect_simp ctxt6 o #Rep o snd o fst) typedefs; (** prove that new types are in class pt_ **) val _ = warning "prove that new types are in class pt_ ..."; fun pt_instance (atom, perm_closed_thms) = fold (fn ((((((Abs_inverse, Rep_inverse), Rep), perm_def), name), tvs), perm_closed) => fn thy => let val pt_class = pt_class_of thy atom; val sort = Sign.minimize_sort thy (Sign.certify_sort thy (pt_class :: map (cp_class_of thy atom) (remove (op =) atom dt_atoms))) in Axclass.prove_arity (Sign.intern_type thy name, map (inter_sort thy sort o snd) tvs, [pt_class]) (fn ctxt => EVERY [Class.intro_classes_tac ctxt [], rewrite_goals_tac ctxt [perm_def], asm_full_simp_tac (ctxt addsimps [Rep_inverse]) 1, asm_full_simp_tac (ctxt addsimps [Rep RS perm_closed RS Abs_inverse]) 1, asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name atom ^ "3")]) 1]) thy end) (Abs_inverse_thms ~~ Rep_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms); (** prove that new types are in class cp__ **) val _ = warning "prove that new types are in class cp__ ..."; fun cp_instance (atom1, perm_closed_thms1) (atom2, perm_closed_thms2) thy = let val cp_class = cp_class_of thy atom1 atom2; val sort = Sign.minimize_sort thy (Sign.certify_sort thy (pt_class_of thy atom1 :: map (cp_class_of thy atom1) (remove (op =) atom1 dt_atoms) @ (if atom1 = atom2 then [cp_class_of thy atom1 atom1] else pt_class_of thy atom2 :: map (cp_class_of thy atom2) (remove (op =) atom2 dt_atoms)))); val cp1' = cp_inst_of thy atom1 atom2 RS cp1 in fold (fn ((((((Abs_inverse, Rep), perm_def), name), tvs), perm_closed1), perm_closed2) => fn thy => Axclass.prove_arity (Sign.intern_type thy name, map (inter_sort thy sort o snd) tvs, [cp_class]) (fn ctxt => EVERY [Class.intro_classes_tac ctxt [], rewrite_goals_tac ctxt [perm_def], asm_full_simp_tac (ctxt addsimps ((Rep RS perm_closed1 RS Abs_inverse) :: (if atom1 = atom2 then [] else [Rep RS perm_closed2 RS Abs_inverse]))) 1, cong_tac ctxt 1, resolve_tac ctxt [refl] 1, resolve_tac ctxt [cp1'] 1]) thy) (Abs_inverse_thms ~~ Rep_thms ~~ perm_defs ~~ new_type_names ~~ tyvars ~~ perm_closed_thms1 ~~ perm_closed_thms2) thy end; val thy7 = fold (fn x => fn thy => thy |> pt_instance x |> fold (cp_instance x) (atoms ~~ perm_closed_thmss)) (atoms ~~ perm_closed_thmss) thy6; (**** constructors ****) fun mk_abs_fun x t = let val T = fastype_of x; val U = fastype_of t in Const (\<^const_name>\Nominal.abs_fun\, T --> U --> T --> Type (\<^type_name>\noption\, [U])) $ x $ t end; val (ty_idxs, _) = List.foldl (fn ((i, (\<^type_name>\noption\, _, _)), p) => p | ((i, _), (ty_idxs, j)) => (ty_idxs @ [(i, j)], j + 1)) ([], 0) descr; fun reindex (Old_Datatype_Aux.DtType (s, dts)) = Old_Datatype_Aux.DtType (s, map reindex dts) | reindex (Old_Datatype_Aux.DtRec i) = Old_Datatype_Aux.DtRec (the (AList.lookup op = ty_idxs i)) | reindex dt = dt; fun strip_suffix i s = implode (List.take (raw_explode s, size s - i)); (* FIXME Symbol.explode (?) *) (** strips the "_Rep" in type names *) fun strip_nth_name i s = let val xs = Long_Name.explode s; in Long_Name.implode (Library.nth_map (length xs - i) (strip_suffix 4) xs) end; val (descr'', ndescr) = ListPair.unzip (map_filter (fn (i, (\<^type_name>\noption\, _, _)) => NONE | (i, (s, dts, constrs)) => let val SOME index = AList.lookup op = ty_idxs i; val (constrs2, constrs1) = map_split (fn (cname, cargs) => apsnd (pair (strip_nth_name 2 (strip_nth_name 1 cname))) (fold_map (fn dt => fn dts => let val (dts', dt') = strip_option dt in ((length dts, length dts'), dts @ dts' @ [reindex dt']) end) cargs [])) constrs in SOME ((index, (strip_nth_name 1 s, map reindex dts, constrs1)), (index, constrs2)) end) descr); val (descr1, descr2) = chop (length new_type_names) descr''; val descr' = [descr1, descr2]; fun partition_cargs idxs xs = map (fn (i, j) => (List.take (List.drop (xs, i), j), nth xs (i + j))) idxs; val pdescr = map (fn ((i, (s, dts, constrs)), (_, idxss)) => (i, (s, dts, map (fn ((cname, cargs), idxs) => (cname, partition_cargs idxs cargs)) (constrs ~~ idxss)))) (descr'' ~~ ndescr); fun nth_dtyp' i = Old_Datatype_Aux.typ_of_dtyp descr'' (Old_Datatype_Aux.DtRec i); val rep_names = map (fn s => Sign.intern_const thy7 ("Rep_" ^ s)) new_type_names; val abs_names = map (fn s => Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; val recTs = Old_Datatype_Aux.get_rec_types descr''; val newTs' = take (length new_type_names) recTs'; val newTs = take (length new_type_names) recTs; val full_new_type_names = map (Sign.full_bname thy) new_type_names; fun make_constr_def tname T T' (((cname_rep, _), (cname, cargs)), (cname', mx)) (thy, defs, eqns) = let fun constr_arg (dts, dt) (j, l_args, r_args) = let val xs = map (fn (dt, i) => Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) i) (dts ~~ (j upto j + length dts - 1)) val x = Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts) in (j + length dts + 1, xs @ x :: l_args, fold_rev mk_abs_fun xs (case dt of Old_Datatype_Aux.DtRec k => if k < length new_type_names then Const (nth rep_names k, Old_Datatype_Aux.typ_of_dtyp descr'' dt --> Old_Datatype_Aux.typ_of_dtyp descr dt) $ x else error "nested recursion not supported" | _ => x) :: r_args) end val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); val abs_name = Sign.intern_const thy ("Abs_" ^ tname); val rep_name = Sign.intern_const thy ("Rep_" ^ tname); val constrT = map fastype_of l_args ---> T; val lhs = list_comb (Const (cname, constrT), l_args); val rhs = list_comb (Const (cname_rep, map fastype_of r_args ---> T'), r_args); val def = Logic.mk_equals (lhs, Const (abs_name, T' --> T) $ rhs); val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (rep_name, T --> T') $ lhs, rhs)); val def_name = (Long_Name.base_name cname) ^ "_def"; val ([def_thm], thy') = thy |> Sign.add_consts [(cname', constrT, mx)] |> (Global_Theory.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]; in (thy', defs @ [def_thm], eqns @ [eqn]) end; fun dt_constr_defs ((((((_, (_, _, constrs)), (_, (_, _, constrs'))), tname), T), T'), constr_syntax) (thy, defs, eqns, dist_lemmas) = let val rep_const = Thm.global_cterm_of thy (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> T')); val dist = Drule.export_without_context (infer_instantiate (Proof_Context.init_global thy) [(#1 (dest_Var distinct_f), rep_const)] Old_Datatype.distinct_lemma); val (thy', defs', eqns') = fold (make_constr_def tname T T') (constrs ~~ constrs' ~~ constr_syntax) (Sign.add_path tname thy, defs, []) in (Sign.parent_path thy', defs', eqns @ [eqns'], dist_lemmas @ [dist]) end; val (thy8, constr_defs, constr_rep_eqns, dist_lemmas) = fold dt_constr_defs (List.take (descr, length new_type_names) ~~ List.take (pdescr, length new_type_names) ~~ new_type_names ~~ newTs ~~ newTs' ~~ constr_syntax) (thy7, [], [], []); val abs_inject_thms = map (collect_simp ctxt6 o #Abs_inject o snd o fst) typedefs val rep_inject_thms = map (#Rep_inject o snd o fst) typedefs (* prove theorem Rep_i (Constr_j ...) = Constr'_j ... *) fun prove_constr_rep_thm eqn = let val inj_thms = map (fn r => r RS iffD1) abs_inject_thms; val rewrites = constr_defs @ map mk_meta_eq Rep_inverse_thms in Goal.prove_global_future thy8 [] [] eqn (fn {context = ctxt, ...} => EVERY [resolve_tac ctxt inj_thms 1, rewrite_goals_tac ctxt rewrites, resolve_tac ctxt [refl] 3, resolve_tac ctxt rep_intrs 2, REPEAT (resolve_tac ctxt Rep_thms 1)]) end; val constr_rep_thmss = map (map prove_constr_rep_thm) constr_rep_eqns; (* prove theorem pi \ Rep_i x = Rep_i (pi \ x) *) fun prove_perm_rep_perm (atom, perm_closed_thms) = map (fn th => let val _ $ (_ $ (Rep $ x)) = Logic.unvarify_global (Thm.prop_of th); val Type ("fun", [T, U]) = fastype_of Rep; val permT = mk_permT (Type (atom, [])); val pi = Free ("pi", permT); in Goal.prove_global_future thy8 [] [] (augment_sort thy8 (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (remove (op =) atom dt_atoms)) (HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (\<^const_name>\Nominal.perm\, permT --> U --> U) $ pi $ (Rep $ x), Rep $ (Const (\<^const_name>\Nominal.perm\, permT --> T --> T) $ pi $ x))))) (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_basic_ss ctxt addsimps (perm_defs @ Abs_inverse_thms @ perm_closed_thms @ Rep_thms)) 1) end) Rep_thms; val perm_rep_perm_thms = maps prove_perm_rep_perm (atoms ~~ perm_closed_thmss); (* prove distinctness theorems *) val distinct_props = Old_Datatype_Prop.make_distincts descr'; val dist_rewrites = map2 (fn rep_thms => fn dist_lemma => dist_lemma :: rep_thms) constr_rep_thmss dist_lemmas; fun prove_distinct_thms _ [] = [] | prove_distinct_thms (p as (rep_thms, dist_lemma)) (t :: ts) = let val dist_thm = Goal.prove_global_future thy8 [] [] t (fn {context = ctxt, ...} => simp_tac (ctxt addsimps (dist_lemma :: rep_thms)) 1) in dist_thm :: Drule.export_without_context (dist_thm RS not_sym) :: prove_distinct_thms p ts end; val distinct_thms = map2 prove_distinct_thms (constr_rep_thmss ~~ dist_lemmas) distinct_props; (** prove equations for permutation functions **) val perm_simps' = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => let val T = nth_dtyp' i in maps (fn (atom, perm_closed_thms) => map (fn ((cname, dts), constr_rep_thm) => let val cname = Sign.intern_const thy8 (Long_Name.append tname (Long_Name.base_name cname)); val permT = mk_permT (Type (atom, [])); val pi = Free ("pi", permT); fun perm t = let val T = fastype_of t in Const (\<^const_name>\Nominal.perm\, permT --> T --> T) $ pi $ t end; fun constr_arg (dts, dt) (j, l_args, r_args) = let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') dts; val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) (Ts ~~ (j upto j + length dts - 1)); val x = Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts); in (j + length dts + 1, xs @ x :: l_args, map perm (xs @ [x]) @ r_args) end val (_, l_args, r_args) = fold_rev constr_arg dts (1, [], []); val c = Const (cname, map fastype_of l_args ---> T) in Goal.prove_global_future thy8 [] [] (augment_sort thy8 (pt_class_of thy8 atom :: map (cp_class_of thy8 atom) (remove (op =) atom dt_atoms)) (HOLogic.mk_Trueprop (HOLogic.mk_eq (perm (list_comb (c, l_args)), list_comb (c, r_args))))) (fn {context = ctxt, ...} => EVERY [simp_tac (ctxt addsimps (constr_rep_thm :: perm_defs)) 1, simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Rep_thms @ Abs_inverse_thms @ constr_defs @ perm_closed_thms)) 1, TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps (Thm.symmetric perm_fun_def :: abs_perm)) 1), TRY (simp_tac (put_simpset HOL_basic_ss ctxt addsimps (perm_fun_def :: perm_defs @ Rep_thms @ Abs_inverse_thms @ perm_closed_thms)) 1)]) end) (constrs ~~ constr_rep_thms)) (atoms ~~ perm_closed_thmss) end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); (** prove injectivity of constructors **) val rep_inject_thms' = map (fn th => th RS sym) rep_inject_thms; val alpha = Global_Theory.get_thms thy8 "alpha"; val abs_fresh = Global_Theory.get_thms thy8 "abs_fresh"; val pt_cp_sort = map (pt_class_of thy8) dt_atoms @ maps (fn s => map (cp_class_of thy8 s) (remove (op =) s dt_atoms)) dt_atoms; val inject_thms = map (fn (((i, (_, _, constrs)), tname), constr_rep_thms) => let val T = nth_dtyp' i in map_filter (fn ((cname, dts), constr_rep_thm) => if null dts then NONE else SOME let val cname = Sign.intern_const thy8 (Long_Name.append tname (Long_Name.base_name cname)); fun make_inj (dts, dt) (j, args1, args2, eqs) = let val Ts_idx = map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1); val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx; val ys = map (fn (T, i) => Old_Datatype_Aux.mk_Free "y" T i) Ts_idx; val x = Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts); val y = Old_Datatype_Aux.mk_Free "y" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts); in (j + length dts + 1, xs @ (x :: args1), ys @ (y :: args2), HOLogic.mk_eq (fold_rev mk_abs_fun xs x, fold_rev mk_abs_fun ys y) :: eqs) end; val (_, args1, args2, eqs) = fold_rev make_inj dts (1, [], [], []); val Ts = map fastype_of args1; val c = Const (cname, Ts ---> T) in Goal.prove_global_future thy8 [] [] (augment_sort thy8 pt_cp_sort (HOLogic.mk_Trueprop (HOLogic.mk_eq (HOLogic.mk_eq (list_comb (c, args1), list_comb (c, args2)), foldr1 HOLogic.mk_conj eqs)))) (fn {context = ctxt, ...} => EVERY [asm_full_simp_tac (ctxt addsimps (constr_rep_thm :: rep_inject_thms')) 1, TRY (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps (fresh_def :: supp_def :: alpha @ abs_perm @ abs_fresh @ rep_inject_thms @ perm_rep_perm_thms)) 1)]) end) (constrs ~~ constr_rep_thms) end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ constr_rep_thmss); (** equations for support and freshness **) val (supp_thms, fresh_thms) = ListPair.unzip (map ListPair.unzip (map (fn ((((i, (_, _, constrs)), tname), inject_thms'), perm_thms') => let val T = nth_dtyp' i in maps (fn (cname, dts) => map (fn atom => let val cname = Sign.intern_const thy8 (Long_Name.append tname (Long_Name.base_name cname)); val atomT = Type (atom, []); fun process_constr (dts, dt) (j, args1, args2) = let val Ts_idx = map (Old_Datatype_Aux.typ_of_dtyp descr'') dts ~~ (j upto j + length dts - 1); val xs = map (fn (T, i) => Old_Datatype_Aux.mk_Free "x" T i) Ts_idx; val x = Old_Datatype_Aux.mk_Free "x" (Old_Datatype_Aux.typ_of_dtyp descr'' dt) (j + length dts); in (j + length dts + 1, xs @ (x :: args1), fold_rev mk_abs_fun xs x :: args2) end; val (_, args1, args2) = fold_rev process_constr dts (1, [], []); val Ts = map fastype_of args1; val c = list_comb (Const (cname, Ts ---> T), args1); fun supp t = Const (\<^const_name>\Nominal.supp\, fastype_of t --> HOLogic.mk_setT atomT) $ t; fun fresh t = fresh_const atomT (fastype_of t) $ Free ("a", atomT) $ t; val supp_thm = Goal.prove_global_future thy8 [] [] (augment_sort thy8 pt_cp_sort (HOLogic.mk_Trueprop (HOLogic.mk_eq (supp c, if null dts then HOLogic.mk_set atomT [] else foldr1 (HOLogic.mk_binop \<^const_abbrev>\union\) (map supp args2))))) (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_basic_ss ctxt addsimps (supp_def :: Un_assoc :: @{thm de_Morgan_conj} :: Collect_disj_eq :: finite_Un :: Collect_False_empty :: finite_emptyI :: @{thms simp_thms} @ abs_perm @ abs_fresh @ inject_thms' @ perm_thms')) 1) in (supp_thm, Goal.prove_global_future thy8 [] [] (augment_sort thy8 pt_cp_sort (HOLogic.mk_Trueprop (HOLogic.mk_eq (fresh c, if null dts then \<^term>\True\ else foldr1 HOLogic.mk_conj (map fresh args2))))) (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_ss ctxt addsimps [Un_iff, empty_iff, fresh_def, supp_thm]) 1)) end) atoms) constrs end) (List.take (pdescr, length new_type_names) ~~ new_type_names ~~ inject_thms ~~ perm_simps'))); (**** weak induction theorem ****) fun mk_indrule_lemma (((i, _), T), U) (prems, concls) = let val Rep_t = Const (nth rep_names i, T --> U) $ Old_Datatype_Aux.mk_Free "x" T i; val Abs_t = Const (nth abs_names i, U --> T); in (prems @ [HOLogic.imp $ (Const (nth rep_set_names'' i, U --> HOLogic.boolT) $ Rep_t) $ (Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], concls @ [Old_Datatype_Aux.mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ Old_Datatype_Aux.mk_Free "x" T i]) end; val (indrule_lemma_prems, indrule_lemma_concls) = fold mk_indrule_lemma (descr'' ~~ recTs ~~ recTs') ([], []); val indrule_lemma = Goal.prove_global_future thy8 [] [] (Logic.mk_implies (HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_prems), HOLogic.mk_Trueprop (Old_Datatype_Aux.mk_conj indrule_lemma_concls))) (fn {context = ctxt, ...} => EVERY [REPEAT (eresolve_tac ctxt [conjE] 1), REPEAT (EVERY [TRY (resolve_tac ctxt [conjI] 1), full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps Rep_inverse_thms) 1, eresolve_tac ctxt [mp] 1, resolve_tac ctxt Rep_thms 1])]); val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of indrule_lemma))); val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else map (Free o apfst fst o dest_Var) Ps; val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; val dt_induct_prop = Old_Datatype_Prop.make_ind descr'; val dt_induct = Goal.prove_global_future thy8 [] (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) (fn {prems, context = ctxt} => let val indrule_lemma' = infer_instantiate ctxt (map (#1 o dest_Var) Ps ~~ map (Thm.cterm_of ctxt) frees) indrule_lemma; in EVERY [resolve_tac ctxt [indrule_lemma'] 1, (Old_Datatype_Aux.ind_tac ctxt rep_induct [] THEN_ALL_NEW Object_Logic.atomize_prems_tac ctxt) 1, EVERY (map (fn (prem, r) => (EVERY [REPEAT (eresolve_tac ctxt Abs_inverse_thms' 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [Thm.symmetric r]) 1, DEPTH_SOLVE_1 (assume_tac ctxt 1 ORELSE resolve_tac ctxt [prem] 1 ORELSE eresolve_tac ctxt [allE] 1)])) (prems ~~ constr_defs))] end); val case_names_induct = Old_Datatype_Data.mk_case_names_induct descr''; (**** prove that new datatypes have finite support ****) val _ = warning "proving finite support for the new datatype"; val indnames = Old_Datatype_Prop.make_tnames recTs; val abs_supp = Global_Theory.get_thms thy8 "abs_supp"; val supp_atm = Global_Theory.get_thms thy8 "supp_atm"; val finite_supp_thms = map (fn atom => let val atomT = Type (atom, []) in map Drule.export_without_context (List.take (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy8 [] [] (augment_sort thy8 (fs_class_of thy8 atom :: pt_cp_sort) (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (s, T) => Const (\<^const_name>\finite\, HOLogic.mk_setT atomT --> HOLogic.boolT) $ (Const (\<^const_name>\Nominal.supp\, T --> HOLogic.mk_setT atomT) $ Free (s, T))) (indnames ~~ recTs))))) (fn {context = ctxt, ...} => Old_Datatype_Aux.ind_tac ctxt dt_induct indnames 1 THEN ALLGOALS (asm_full_simp_tac (ctxt addsimps (abs_supp @ supp_atm @ Global_Theory.get_thms thy8 ("fs_" ^ Long_Name.base_name atom ^ "1") @ flat supp_thms))))), length new_type_names)) end) atoms; val simp_atts = replicate (length new_type_names) [Simplifier.simp_add]; (* Function to add both the simp and eqvt attributes *) (* These two attributes are duplicated on all the types in the mutual nominal datatypes *) val simp_eqvt_atts = replicate (length new_type_names) [Simplifier.simp_add, NominalThmDecls.eqvt_add]; val (_, thy9) = thy8 |> Sign.add_path big_name |> Global_Theory.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>> Global_Theory.add_thmss [((Binding.name "inducts", projections (Proof_Context.init_global thy8) dt_induct), [case_names_induct])] ||> Sign.parent_path ||>> Old_Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> Old_Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> Old_Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> Old_Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>> Old_Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>> Old_Datatype_Aux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> fold (fn (atom, ths) => fn thy => let val class = fs_class_of thy atom; val sort = Sign.minimize_sort thy (Sign.certify_sort thy (class :: pt_cp_sort)); in fold (fn Type (s, Ts) => Axclass.prove_arity (s, map (inter_sort thy sort o snd o dest_TFree) Ts, [class]) (fn ctxt => Class.intro_classes_tac ctxt [] THEN resolve_tac ctxt ths 1)) newTs thy end) (atoms ~~ finite_supp_thms); (**** strong induction theorem ****) val pnames = if length descr'' = 1 then ["P"] else map (fn i => "P" ^ string_of_int i) (1 upto length descr''); val ind_sort = if null dt_atomTs then \<^sort>\type\ else Sign.minimize_sort thy9 (Sign.certify_sort thy9 (map (fs_class_of thy9) dt_atoms)); val fsT = TFree ("'n", ind_sort); val fsT' = TFree ("'n", \<^sort>\type\); val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) (Old_Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); fun make_pred fsT i T = Free (nth pnames i, fsT --> T --> HOLogic.boolT); fun mk_fresh1 xs [] = [] | mk_fresh1 xs ((y as (_, T)) :: ys) = map (fn x => HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (Free y, Free x)))) (filter (fn (_, U) => T = U) (rev xs)) @ mk_fresh1 (y :: xs) ys; fun mk_fresh2 xss [] = [] | mk_fresh2 xss ((p as (ys, _)) :: yss) = maps (fn y as (_, T) => map (fn (_, x as (_, U)) => HOLogic.mk_Trueprop (fresh_const T U $ Free y $ Free x)) (rev xss @ yss)) ys @ mk_fresh2 (p :: xss) yss; fun make_ind_prem fsT f k T ((cname, cargs), idxs) = let val recs = filter Old_Datatype_Aux.is_rec_type cargs; val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs; val recTs' = map (Old_Datatype_Aux.typ_of_dtyp descr'') recs; val tnames = Name.variant_list pnames (Old_Datatype_Prop.make_tnames Ts); val rec_tnames = map fst (filter (Old_Datatype_Aux.is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val frees' = partition_cargs idxs frees; val z = (singleton (Name.variant_list tnames) "z", fsT); fun mk_prem ((dt, s), T) = let val (Us, U) = strip_type T; val l = length Us; in Logic.list_all (z :: map (pair "x") Us, HOLogic.mk_Trueprop (make_pred fsT (Old_Datatype_Aux.body_index dt) U $ Bound l $ Old_Datatype_Aux.app_bnds (Free (s, T)) l)) end; val prems = map mk_prem (recs ~~ rec_tnames ~~ recTs'); val prems' = map (fn p as (_, T) => HOLogic.mk_Trueprop (f T (Free p) (Free z))) (maps fst frees') @ mk_fresh1 [] (maps fst frees') @ mk_fresh2 [] frees' in fold_rev (Logic.all o Free) (frees @ [z]) (Logic.list_implies (prems' @ prems, HOLogic.mk_Trueprop (make_pred fsT k T $ Free z $ list_comb (Const (cname, Ts ---> T), map Free frees)))) end; val ind_prems = maps (fn (((i, (_, _, constrs)), (_, idxss)), T) => map (make_ind_prem fsT (fn T => fn t => fn u => fresh_const T fsT $ t $ u) i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs); val tnames = Old_Datatype_Prop.make_tnames recTs; val zs = Name.variant_list tnames (replicate (length descr'') "z"); val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop \<^const_name>\HOL.conj\) (map (fn ((((i, _), T), tname), z) => make_pred fsT i T $ Free (z, fsT) $ Free (tname, T)) (descr'' ~~ recTs ~~ tnames ~~ zs))); val induct = Logic.list_implies (ind_prems, ind_concl); val ind_prems' = map (fn (_, f as Free (_, T)) => Logic.all (Free ("x", fsT')) (HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, Term.range_type T --> HOLogic.boolT) $ (f $ Free ("x", fsT'))))) fresh_fs @ maps (fn (((i, (_, _, constrs)), (_, idxss)), T) => map (make_ind_prem fsT' (fn T => fn t => fn u => HOLogic.Not $ HOLogic.mk_mem (t, the (AList.lookup op = fresh_fs T) $ u)) i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs); val ind_concl' = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop \<^const_name>\HOL.conj\) (map (fn ((((i, _), T), tname), z) => make_pred fsT' i T $ Free (z, fsT') $ Free (tname, T)) (descr'' ~~ recTs ~~ tnames ~~ zs))); val induct' = Logic.list_implies (ind_prems', ind_concl'); val aux_ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~ map mk_permT dt_atomTs) @ [("z", fsT')]; val aux_ind_Ts = rev (map snd aux_ind_vars); val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop \<^const_name>\HOL.conj\) (map (fn (((i, _), T), tname) => HOLogic.list_all (aux_ind_vars, make_pred fsT' i T $ Bound 0 $ fold_rev (mk_perm aux_ind_Ts) (map Bound (length dt_atomTs downto 1)) (Free (tname, T)))) (descr'' ~~ recTs ~~ tnames))); val fin_set_supp = map (fn s => at_inst_of thy9 s RS at_fin_set_supp) dt_atoms; val fin_set_fresh = map (fn s => at_inst_of thy9 s RS at_fin_set_fresh) dt_atoms; val pt1_atoms = map (fn Type (s, _) => Global_Theory.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "1")) dt_atomTs; val pt2_atoms = map (fn Type (s, _) => Global_Theory.get_thm thy9 ("pt_" ^ Long_Name.base_name s ^ "2") RS sym) dt_atomTs; val exists_fresh' = Global_Theory.get_thms thy9 "exists_fresh'"; val fs_atoms = Global_Theory.get_thms thy9 "fin_supp"; val abs_supp = Global_Theory.get_thms thy9 "abs_supp"; val perm_fresh_fresh = Global_Theory.get_thms thy9 "perm_fresh_fresh"; val calc_atm = Global_Theory.get_thms thy9 "calc_atm"; val fresh_atm = Global_Theory.get_thms thy9 "fresh_atm"; val fresh_left = Global_Theory.get_thms thy9 "fresh_left"; val perm_swap = Global_Theory.get_thms thy9 "perm_swap"; fun obtain_fresh_name' ths ts T (freshs1, freshs2, ctxt) = let val p = foldr1 HOLogic.mk_prod (ts @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn _ => EVERY [resolve_tac ctxt exists_fresh' 1, simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms @ fin_set_supp @ ths)) 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn ctxt' => EVERY [eresolve_tac ctxt' [exE] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, REPEAT (eresolve_tac ctxt' [conjE] 1)]) [ex] ctxt in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; fun fresh_fresh_inst thy a b = let val T = fastype_of a; val SOME th = find_first (fn th => (case Thm.prop_of th of _ $ (_ $ (Const (_, Type (_, [U, _])) $ _ $ _)) $ _ => U = T | _ => false)) perm_fresh_fresh in Thm.instantiate' [] [SOME (Thm.global_cterm_of thy a), NONE, SOME (Thm.global_cterm_of thy b)] th end; val fs_cp_sort = map (fs_class_of thy9) dt_atoms @ maps (fn s => map (cp_class_of thy9 s) (remove (op =) s dt_atoms)) dt_atoms; (********************************************************************** The subgoals occurring in the proof of induct_aux have the following parameters: x_1 ... x_k p_1 ... p_m z where x_i : constructor arguments (introduced by weak induction rule) p_i : permutations (one for each atom type in the data type) z : freshness context ***********************************************************************) val _ = warning "proving strong induction theorem ..."; val induct_aux = Goal.prove_global_future thy9 [] (map (augment_sort thy9 fs_cp_sort) ind_prems') (augment_sort thy9 fs_cp_sort ind_concl') (fn {prems, context} => let val (prems1, prems2) = chop (length dt_atomTs) prems; val ind_ss2 = put_simpset HOL_ss context addsimps finite_Diff :: abs_fresh @ abs_supp @ fs_atoms; val ind_ss1 = ind_ss2 addsimps fresh_left @ calc_atm @ fresh_atm @ rev_simps @ app_simps; val ind_ss3 = put_simpset HOL_ss context addsimps abs_fun_eq1 :: abs_perm @ calc_atm @ perm_swap; val ind_ss4 = put_simpset HOL_basic_ss context addsimps fresh_left @ prems1 @ fin_set_fresh @ calc_atm; val ind_ss5 = put_simpset HOL_basic_ss context addsimps pt1_atoms; val ind_ss6 = put_simpset HOL_basic_ss context addsimps flat perm_simps'; val th = Goal.prove context [] [] (augment_sort thy9 fs_cp_sort aux_ind_concl) (fn {context = context1, ...} => EVERY (Old_Datatype_Aux.ind_tac context1 dt_induct tnames 1 :: maps (fn ((_, (_, _, constrs)), (_, constrs')) => map (fn ((cname, cargs), is) => REPEAT (resolve_tac context1 [allI] 1) THEN SUBPROOF (fn {prems = iprems, params, concl, context = context2, ...} => let val concl' = Thm.term_of concl; val _ $ (_ $ _ $ u) = concl'; val U = fastype_of u; val (xs, params') = chop (length cargs) (map (Thm.term_of o #2) params); val Ts = map fastype_of xs; val cnstr = Const (cname, Ts ---> U); val (pis, z) = split_last params'; val mk_pi = fold_rev (mk_perm []) pis; val xs' = partition_cargs is xs; val xs'' = map (fn (ts, u) => (map mk_pi ts, mk_pi u)) xs'; val ts = maps (fn (ts, u) => ts @ [u]) xs''; val (freshs1, freshs2, context3) = fold (fn t => let val T = fastype_of t in obtain_fresh_name' prems1 (the (AList.lookup op = fresh_fs T) $ z :: ts) T end) (maps fst xs') ([], [], context2); val freshs1' = unflat (map fst xs') freshs1; val freshs2' = map (Simplifier.simplify ind_ss4) (mk_not_sym freshs2); val ind_ss1' = ind_ss1 addsimps freshs2'; val ind_ss3' = ind_ss3 addsimps freshs2'; val rename_eq = if forall (null o fst) xs' then [] else [Goal.prove context3 [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (cnstr, ts), list_comb (cnstr, maps (fn ((bs, t), cs) => cs @ [fold_rev (mk_perm []) (map perm_of_pair (bs ~~ cs)) t]) (xs'' ~~ freshs1'))))) (fn {context = ctxt, ...} => EVERY (simp_tac (put_simpset HOL_ss context3 addsimps flat inject_thms) 1 :: REPEAT (FIRSTGOAL (resolve_tac ctxt [conjI])) :: maps (fn ((bs, t), cs) => if null bs then [] else resolve_tac ctxt [sym] 1 :: maps (fn (b, c) => [resolve_tac ctxt [trans] 1, resolve_tac ctxt [sym] 1, resolve_tac ctxt [fresh_fresh_inst thy9 b c] 1, simp_tac ind_ss1' 1, simp_tac ind_ss2 1, simp_tac ind_ss3' 1]) (bs ~~ cs)) (xs'' ~~ freshs1')))]; val th = Goal.prove context3 [] [] concl' (fn _ => EVERY [simp_tac (ind_ss6 addsimps rename_eq) 1, cut_facts_tac iprems 1, (resolve_tac context2 prems THEN_ALL_NEW SUBGOAL (fn (t, i) => case Logic.strip_assums_concl t of _ $ (Const (\<^const_name>\Nominal.fresh\, _) $ _ $ _) => simp_tac ind_ss1' i | _ $ (Const (\<^const_name>\Not\, _) $ _) => resolve_tac context2 freshs2' i | _ => asm_simp_tac (put_simpset HOL_basic_ss context3 addsimps pt2_atoms addsimprocs [perm_simproc]) i)) 1]) val final = Proof_Context.export context3 context2 [th] in resolve_tac context2 final 1 end) context1 1) (constrs ~~ constrs')) (descr'' ~~ ndescr))) in EVERY [cut_facts_tac [th] 1, REPEAT (eresolve_tac context [conjE, @{thm allE_Nil}] 1), REPEAT (eresolve_tac context [allE] 1), REPEAT (TRY (resolve_tac context [conjI] 1) THEN asm_full_simp_tac ind_ss5 1)] end); - val induct_aux' = Thm.instantiate ([], - map (fn (s, Var (v as (_, T))) => + val induct_aux' = Thm.instantiate (TVars.empty, + Vars.make (map (fn (s, Var (v as (_, T))) => (v, Thm.global_cterm_of thy9 (Free (s, T)))) (pnames ~~ map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (Thm.concl_of induct_aux)))) @ map (fn (_, f) => let val f' = Logic.varify_global f in (dest_Var f', Thm.global_cterm_of thy9 (Const (\<^const_name>\Nominal.supp\, fastype_of f'))) - end) fresh_fs) induct_aux; + end) fresh_fs)) induct_aux; val induct = Goal.prove_global_future thy9 [] (map (augment_sort thy9 fs_cp_sort) ind_prems) (augment_sort thy9 fs_cp_sort ind_concl) (fn {prems, context = ctxt} => EVERY [resolve_tac ctxt [induct_aux'] 1, REPEAT (resolve_tac ctxt fs_atoms 1), REPEAT ((resolve_tac ctxt prems THEN_ALL_NEW (eresolve_tac ctxt @{thms meta_spec} ORELSE' full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [fresh_def]))) 1)]) val (_, thy10) = thy9 |> Sign.add_path big_name |> Global_Theory.add_thms [((Binding.name "strong_induct'", induct_aux), [])] ||>> Global_Theory.add_thms [((Binding.name "strong_induct", induct), [case_names_induct])] ||>> Global_Theory.add_thmss [((Binding.name "strong_inducts", projections (Proof_Context.init_global thy9) induct), [case_names_induct])]; (**** recursion combinator ****) val _ = warning "defining recursion combinator ..."; val used = fold Term.add_tfree_namesT recTs []; val (rec_result_Ts', rec_fn_Ts') = Old_Datatype_Prop.make_primrec_Ts descr' used; val rec_sort = if null dt_atomTs then \<^sort>\type\ else Sign.minimize_sort thy10 (Sign.certify_sort thy10 pt_cp_sort); val rec_result_Ts = map (fn TFree (s, _) => TFree (s, rec_sort)) rec_result_Ts'; val rec_fn_Ts = map (typ_subst_atomic (rec_result_Ts' ~~ rec_result_Ts)) rec_fn_Ts'; val rec_set_Ts = map (fn (T1, T2) => rec_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); val big_rec_name = big_name ^ "_rec_set"; val rec_set_names' = if length descr'' = 1 then [big_rec_name] else map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) (1 upto (length descr'')); val rec_set_names = map (Sign.full_bname thy10) rec_set_names'; val rec_fns = map (uncurry (Old_Datatype_Aux.mk_Free "f")) (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); val rec_sets' = map (fn c => list_comb (Free c, rec_fns)) (rec_set_names' ~~ rec_set_Ts); val rec_sets = map (fn c => list_comb (Const c, rec_fns)) (rec_set_names ~~ rec_set_Ts); (* introduction rules for graph of recursion function *) val rec_preds = map (fn (a, T) => Free (a, T --> HOLogic.boolT)) (pnames ~~ rec_result_Ts); fun mk_fresh3 rs [] = [] | mk_fresh3 rs ((p as (ys, z)) :: yss) = maps (fn y as (_, T) => map_filter (fn ((_, (_, x)), r as (_, U)) => if z = x then NONE else SOME (HOLogic.mk_Trueprop (fresh_const T U $ Free y $ Free r))) rs) ys @ mk_fresh3 rs yss; (* FIXME: avoid collisions with other variable names? *) val rec_ctxt = Free ("z", fsT'); fun make_rec_intr T p rec_set ((cname, cargs), idxs) (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, l) = let val Ts = map (Old_Datatype_Aux.typ_of_dtyp descr'') cargs; val frees = map (fn i => "x" ^ string_of_int i) (1 upto length Ts) ~~ Ts; val frees' = partition_cargs idxs frees; val binders = maps fst frees'; val atomTs = distinct op = (maps (map snd o fst) frees'); val recs = map_filter (fn ((_, Old_Datatype_Aux.DtRec i), p) => SOME (i, p) | _ => NONE) (partition_cargs idxs cargs ~~ frees'); val frees'' = map (fn i => "y" ^ string_of_int i) (1 upto length recs) ~~ map (fn (i, _) => nth rec_result_Ts i) recs; val prems1 = map (fn ((i, (_, x)), y) => HOLogic.mk_Trueprop (nth rec_sets' i $ Free x $ Free y)) (recs ~~ frees''); val prems2 = map (fn f => map (fn p as (_, T) => HOLogic.mk_Trueprop (fresh_const T (fastype_of f) $ Free p $ f)) binders) rec_fns; val prems3 = mk_fresh1 [] binders @ mk_fresh2 [] frees'; val prems4 = map (fn ((i, _), y) => HOLogic.mk_Trueprop (nth rec_preds i $ Free y)) (recs ~~ frees''); val prems5 = mk_fresh3 (recs ~~ frees'') frees'; val prems6 = maps (fn aT => map (fn y as (_, T) => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT aT --> HOLogic.boolT) $ (Const (\<^const_name>\Nominal.supp\, T --> HOLogic.mk_setT aT) $ Free y))) frees'') atomTs; val prems7 = map (fn x as (_, T) => HOLogic.mk_Trueprop (fresh_const T fsT' $ Free x $ rec_ctxt)) binders; val result = list_comb (nth rec_fns l, map Free (frees @ frees'')); val result_freshs = map (fn p as (_, T) => fresh_const T (fastype_of result) $ Free p $ result) binders; val P = HOLogic.mk_Trueprop (p $ result) in (rec_intr_ts @ [Logic.list_implies (flat prems2 @ prems3 @ prems1, HOLogic.mk_Trueprop (rec_set $ list_comb (Const (cname, Ts ---> T), map Free frees) $ result))], rec_prems @ [fold_rev (Logic.all o Free) (frees @ frees'') (Logic.list_implies (prems4, P))], rec_prems' @ map (fn fr => fold_rev (Logic.all o Free) (frees @ frees'') (Logic.list_implies (nth prems2 l @ prems3 @ prems5 @ prems7 @ prems6 @ [P], HOLogic.mk_Trueprop fr))) result_freshs, rec_eq_prems @ [flat prems2 @ prems3], l + 1) end; val (rec_intr_ts, rec_prems, rec_prems', rec_eq_prems, _) = fold (fn ((((d, d'), T), p), rec_set) => fold (make_rec_intr T p rec_set) (#3 (snd d) ~~ snd d')) (descr'' ~~ ndescr ~~ recTs ~~ rec_preds ~~ rec_sets') ([], [], [], [], 0); val ({intrs = rec_intrs, elims = rec_elims, raw_induct = rec_induct, ...}, thy11) = thy10 |> Sign.concealed |> Named_Target.theory_map_result Inductive.transform_result (Inductive.add_inductive {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, coind = false, no_elim = false, no_ind = false, skip_mono = true} (map (fn (s, T) => ((Binding.name s, T), NoSyn)) (rec_set_names' ~~ rec_set_Ts)) (map dest_Free rec_fns) (map (fn x => (Binding.empty_atts, x)) rec_intr_ts) []) ||> Global_Theory.hide_fact true (Long_Name.append (Sign.full_bname thy10 big_rec_name) "induct") ||> Sign.restore_naming thy10; (** equivariance **) val fresh_bij = Global_Theory.get_thms thy11 "fresh_bij"; val perm_bij = Global_Theory.get_thms thy11 "perm_bij"; val (rec_equiv_thms, rec_equiv_thms') = ListPair.unzip (map (fn aT => let val permT = mk_permT aT; val pi = Free ("pi", permT); val rec_fns_pi = map (mk_perm [] pi o uncurry (Old_Datatype_Aux.mk_Free "f")) (rec_fn_Ts ~~ (1 upto (length rec_fn_Ts))); val rec_sets_pi = map (fn c => list_comb (Const c, rec_fns_pi)) (rec_set_names ~~ rec_set_Ts); val ps = map (fn ((((T, U), R), R'), i) => let val x = Free ("x" ^ string_of_int i, T); val y = Free ("y" ^ string_of_int i, U) in (R $ x $ y, R' $ mk_perm [] pi x $ mk_perm [] pi y) end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ rec_sets_pi ~~ (1 upto length recTs)); val ths = map (fn th => Drule.export_without_context (th RS mp)) (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy11 [] [] (augment_sort thy1 pt_cp_sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.mk_imp ps)))) (fn {context = ctxt, ...} => resolve_tac ctxt [rec_induct] 1 THEN REPEAT (simp_tac (put_simpset HOL_basic_ss ctxt addsimps flat perm_simps' addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN (resolve_tac ctxt rec_intrs THEN_ALL_NEW asm_simp_tac (put_simpset HOL_ss ctxt addsimps (fresh_bij @ perm_bij))) 1)))) val ths' = map (fn ((P, Q), th) => Goal.prove_global_future thy11 [] [] (augment_sort thy1 pt_cp_sort (Logic.mk_implies (HOLogic.mk_Trueprop Q, HOLogic.mk_Trueprop P))) (fn {context = ctxt, ...} => - dresolve_tac ctxt [Thm.instantiate ([], - [((("pi", 0), permT), + dresolve_tac ctxt [Thm.instantiate (TVars.empty, + Vars.make [((("pi", 0), permT), Thm.global_cterm_of thy11 (Const (\<^const_name>\rev\, permT --> permT) $ pi))]) th] 1 THEN NominalPermeq.perm_simp_tac (put_simpset HOL_ss ctxt) 1)) (ps ~~ ths) in (ths, ths') end) dt_atomTs); (** finite support **) val rec_fin_supp_thms = map (fn aT => let val name = Long_Name.base_name (fst (dest_Type aT)); val fs_name = Global_Theory.get_thm thy11 ("fs_" ^ name ^ "1"); val aset = HOLogic.mk_setT aT; val finite = Const (\<^const_name>\finite\, aset --> HOLogic.boolT); val fins = map (fn (f, T) => HOLogic.mk_Trueprop (finite $ (Const (\<^const_name>\Nominal.supp\, T --> aset) $ f))) (rec_fns ~~ rec_fn_Ts) in map (fn th => Drule.export_without_context (th RS mp)) (Old_Datatype_Aux.split_conj_thm (Goal.prove_global_future thy11 [] (map (augment_sort thy11 fs_cp_sort) fins) (augment_sort thy11 fs_cp_sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (((T, U), R), i) => let val x = Free ("x" ^ string_of_int i, T); val y = Free ("y" ^ string_of_int i, U) in HOLogic.mk_imp (R $ x $ y, finite $ (Const (\<^const_name>\Nominal.supp\, U --> aset) $ y)) end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ (1 upto length recTs)))))) (fn {prems = fins, context = ctxt} => (resolve_tac ctxt [rec_induct] THEN_ALL_NEW cut_facts_tac fins) 1 THEN REPEAT (NominalPermeq.finite_guess_tac (put_simpset HOL_ss ctxt addsimps [fs_name]) 1)))) end) dt_atomTs; (** freshness **) val finite_premss = map (fn aT => map (fn (f, T) => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT aT --> HOLogic.boolT) $ (Const (\<^const_name>\Nominal.supp\, T --> HOLogic.mk_setT aT) $ f))) (rec_fns ~~ rec_fn_Ts)) dt_atomTs; val rec_fns' = map (augment_sort thy11 fs_cp_sort) rec_fns; val rec_fresh_thms = map (fn ((aT, eqvt_ths), finite_prems) => let val name = Long_Name.base_name (fst (dest_Type aT)); val fs_name = Global_Theory.get_thm thy11 ("fs_" ^ name ^ "1"); val a = Free ("a", aT); val freshs = map (fn (f, fT) => HOLogic.mk_Trueprop (fresh_const aT fT $ a $ f)) (rec_fns ~~ rec_fn_Ts) in map (fn (((T, U), R), eqvt_th) => let val x = Free ("x", augment_sort_typ thy11 fs_cp_sort T); val y = Free ("y", U); val y' = Free ("y'", U) in Drule.export_without_context (Goal.prove (Proof_Context.init_global thy11) [] (map (augment_sort thy11 fs_cp_sort) (finite_prems @ [HOLogic.mk_Trueprop (R $ x $ y), HOLogic.mk_Trueprop (HOLogic.mk_all ("y'", U, HOLogic.mk_imp (R $ x $ y', HOLogic.mk_eq (y', y)))), HOLogic.mk_Trueprop (fresh_const aT T $ a $ x)] @ freshs)) (HOLogic.mk_Trueprop (fresh_const aT U $ a $ y)) (fn {prems, context = ctxt} => let val (finite_prems, rec_prem :: unique_prem :: fresh_prems) = chop (length finite_prems) prems; val unique_prem' = unique_prem RS spec RS mp; val unique = [unique_prem', unique_prem' RS sym] MRS trans; val _ $ (_ $ (_ $ S $ _)) $ _ = Thm.prop_of supports_fresh; val tuple = foldr1 HOLogic.mk_prod (x :: rec_fns') in EVERY [resolve_tac ctxt [infer_instantiate ctxt [(#1 (dest_Var S), Thm.cterm_of ctxt (Const (\<^const_name>\Nominal.supp\, fastype_of tuple --> HOLogic.mk_setT aT) $ tuple))] supports_fresh] 1, simp_tac (put_simpset HOL_basic_ss ctxt addsimps [supports_def, Thm.symmetric fresh_def, fresh_prod]) 1, REPEAT_DETERM (resolve_tac ctxt [allI, impI] 1), REPEAT_DETERM (eresolve_tac ctxt [conjE] 1), resolve_tac ctxt [unique] 1, SUBPROOF (fn {context = ctxt', prems = prems', params = [(_, a), (_, b)], ...} => EVERY [cut_facts_tac [rec_prem] 1, - resolve_tac ctxt' [Thm.instantiate ([], - [((("pi", 0), mk_permT aT), + resolve_tac ctxt' [Thm.instantiate (TVars.empty, + Vars.make [((("pi", 0), mk_permT aT), Thm.cterm_of ctxt' (perm_of_pair (Thm.term_of a, Thm.term_of b)))]) eqvt_th] 1, asm_simp_tac (put_simpset HOL_ss ctxt' addsimps (prems' @ perm_swap @ perm_fresh_fresh)) 1]) ctxt 1, resolve_tac ctxt [rec_prem] 1, simp_tac (put_simpset HOL_ss ctxt addsimps (fs_name :: supp_prod :: finite_Un :: finite_prems)) 1, simp_tac (put_simpset HOL_ss ctxt addsimps (Thm.symmetric fresh_def :: fresh_prod :: fresh_prems)) 1] end)) end) (recTs ~~ rec_result_Ts ~~ rec_sets ~~ eqvt_ths) end) (dt_atomTs ~~ rec_equiv_thms' ~~ finite_premss); (** uniqueness **) val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); val fun_tupleT = fastype_of fun_tuple; val rec_unique_frees = Old_Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs; val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; val rec_unique_frees' = Old_Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; val rec_unique_concls = map (fn ((x, U), R) => Const (\<^const_name>\Ex1\, (U --> HOLogic.boolT) --> HOLogic.boolT) $ Abs ("y", U, R $ Free x $ Bound 0)) (rec_unique_frees ~~ rec_result_Ts ~~ rec_sets); val induct_aux_rec = infer_instantiate (Proof_Context.init_global thy11) (map (apsnd (Thm.global_cterm_of thy11 o augment_sort thy11 fs_cp_sort)) (map (fn (aT, f) => (#1 (dest_Var (Logic.varify_global f)), Abs ("z", HOLogic.unitT, Const (\<^const_name>\Nominal.supp\, fun_tupleT --> HOLogic.mk_setT aT) $ fun_tuple))) fresh_fs @ map (fn (((P, T), (x, U)), Q) => ((P, 0), Abs ("z", HOLogic.unitT, absfree (x, U) Q))) (pnames ~~ recTs ~~ rec_unique_frees ~~ rec_unique_concls) @ map (fn (s, T) => ((s, 0), Free (s, T))) rec_unique_frees)) induct_aux; fun obtain_fresh_name vs ths rec_fin_supp T (freshs1, freshs2, ctxt) = let val p = foldr1 HOLogic.mk_prod (vs @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn _ => EVERY [cut_facts_tac ths 1, REPEAT_DETERM (dresolve_tac ctxt (the (AList.lookup op = rec_fin_supp T)) 1), resolve_tac ctxt exists_fresh' 1, asm_simp_tac (put_simpset HOL_ss ctxt addsimps (supp_prod :: finite_Un :: fs_atoms)) 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn ctxt' => EVERY [eresolve_tac ctxt' [exE] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, REPEAT (eresolve_tac ctxt [conjE] 1)]) [ex] ctxt in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; val finite_ctxt_prems = map (fn aT => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT aT --> HOLogic.boolT) $ (Const (\<^const_name>\Nominal.supp\, fsT' --> HOLogic.mk_setT aT) $ rec_ctxt))) dt_atomTs; val rec_unique_thms = Old_Datatype_Aux.split_conj_thm (Goal.prove (Proof_Context.init_global thy11) (map fst rec_unique_frees) (map (augment_sort thy11 fs_cp_sort) (flat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems')) (augment_sort thy11 fs_cp_sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj rec_unique_concls))) (fn {prems, context} => let val k = length rec_fns; val (finite_thss, ths1) = fold_map (fn T => fn xs => apfst (pair T) (chop k xs)) dt_atomTs prems; val (finite_ctxt_ths, ths2) = chop (length dt_atomTs) ths1; val (P_ind_ths, fcbs) = chop k ths2; val P_ths = map (fn th => th RS mp) (Old_Datatype_Aux.split_conj_thm (Goal.prove context (map fst (rec_unique_frees'' @ rec_unique_frees')) [] (augment_sort thy11 fs_cp_sort (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (((x, y), S), P) => HOLogic.mk_imp (S $ Free x $ Free y, P $ (Free y))) (rec_unique_frees'' ~~ rec_unique_frees' ~~ rec_sets ~~ rec_preds))))) (fn {context = ctxt, ...} => resolve_tac ctxt [rec_induct] 1 THEN REPEAT ((resolve_tac ctxt P_ind_ths THEN_ALL_NEW assume_tac ctxt) 1)))); val rec_fin_supp_thms' = map (fn (ths, (T, fin_ths)) => (T, map (curry op MRS fin_ths) ths)) (rec_fin_supp_thms ~~ finite_thss); in EVERY ([resolve_tac context [induct_aux_rec] 1] @ maps (fn ((_, finite_ths), finite_th) => [cut_facts_tac (finite_th :: finite_ths) 1, asm_simp_tac (put_simpset HOL_ss context addsimps [supp_prod, finite_Un]) 1]) (finite_thss ~~ finite_ctxt_ths) @ maps (fn ((_, idxss), elim) => maps (fn idxs => [full_simp_tac (put_simpset HOL_ss context addsimps [Thm.symmetric fresh_def, supp_prod, Un_iff]) 1, REPEAT_DETERM (eresolve_tac context @{thms conjE ex1E} 1), resolve_tac context @{thms ex1I} 1, (resolve_tac context rec_intrs THEN_ALL_NEW assume_tac context) 1, rotate_tac ~1 1, ((DETERM o eresolve_tac context [elim]) THEN_ALL_NEW full_simp_tac (put_simpset HOL_ss context addsimps flat distinct_thms)) 1] @ (if null idxs then [] else [hyp_subst_tac context 1, SUBPROOF (fn {asms, concl, prems = prems', params, context = context', ...} => let val SOME prem = find_first (can (HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.prop_of)) prems'; val _ $ (_ $ lhs $ rhs) = Thm.prop_of prem; val _ $ (_ $ lhs' $ rhs') = Thm.term_of concl; val rT = fastype_of lhs'; val (c, cargsl) = strip_comb lhs; val cargsl' = partition_cargs idxs cargsl; val boundsl = maps fst cargsl'; val (_, cargsr) = strip_comb rhs; val cargsr' = partition_cargs idxs cargsr; val boundsr = maps fst cargsr'; val (params1, _ :: params2) = chop (length params div 2) (map (Thm.term_of o #2) params); val params' = params1 @ params2; val rec_prems = filter (fn th => (case Thm.prop_of th of _ $ p => (case head_of p of Const (s, _) => member (op =) rec_set_names s | _ => false) | _ => false)) prems'; val fresh_prems = filter (fn th => (case Thm.prop_of th of _ $ (Const (\<^const_name>\Nominal.fresh\, _) $ _ $ _) => true | _ $ (Const (\<^const_name>\Not\, _) $ _) => true | _ => false)) prems'; val Ts = map fastype_of boundsl; val _ = warning "step 1: obtaining fresh names"; val (freshs1, freshs2, context'') = fold (obtain_fresh_name (rec_ctxt :: rec_fns' @ params') (maps snd finite_thss @ finite_ctxt_ths @ rec_prems) rec_fin_supp_thms') Ts ([], [], context'); val pi1 = map perm_of_pair (boundsl ~~ freshs1); val rpi1 = rev pi1; val pi2 = map perm_of_pair (boundsr ~~ freshs1); val rpi2 = rev pi2; val fresh_prems' = mk_not_sym fresh_prems; val freshs2' = mk_not_sym freshs2; (** as, bs, cs # K as ts, K bs us **) val _ = warning "step 2: as, bs, cs # K as ts, K bs us"; val prove_fresh_simpset = put_simpset HOL_ss context'' addsimps (finite_Diff :: flat fresh_thms @ fs_atoms @ abs_fresh @ abs_supp @ fresh_atm); (* FIXME: avoid asm_full_simp_tac ? *) fun prove_fresh ths y x = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const (fastype_of x) (fastype_of y) $ x $ y)) (fn _ => cut_facts_tac ths 1 THEN asm_full_simp_tac prove_fresh_simpset 1); val constr_fresh_thms = map (prove_fresh fresh_prems lhs) boundsl @ map (prove_fresh fresh_prems rhs) boundsr @ map (prove_fresh freshs2 lhs) freshs1 @ map (prove_fresh freshs2 rhs) freshs1; (** pi1 o (K as ts) = pi2 o (K bs us) **) val _ = warning "step 3: pi1 o (K as ts) = pi2 o (K bs us)"; val pi1_pi2_eq = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 rhs))) (fn {context = ctxt, ...} => EVERY [cut_facts_tac constr_fresh_thms 1, asm_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh) 1, resolve_tac ctxt [prem] 1]); (** pi1 o ts = pi2 o us **) val _ = warning "step 4: pi1 o ts = pi2 o us"; val pi1_pi2_eqs = map (fn (t, u) => Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 t, fold_rev (mk_perm []) pi2 u))) (fn _ => EVERY [cut_facts_tac [pi1_pi2_eq] 1, asm_full_simp_tac (put_simpset HOL_ss context'' addsimps (calc_atm @ flat perm_simps' @ fresh_prems' @ freshs2' @ abs_perm @ alpha @ flat inject_thms)) 1])) (map snd cargsl' ~~ map snd cargsr'); (** pi1^-1 o pi2 o us = ts **) val _ = warning "step 5: pi1^-1 o pi2 o us = ts"; val rpi1_pi2_eqs = map (fn ((t, u), eq) => Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) (rpi1 @ pi2) u, t))) (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps ((eq RS sym) :: perm_swap)) 1)) (map snd cargsl' ~~ map snd cargsr' ~~ pi1_pi2_eqs); val (rec_prems1, rec_prems2) = chop (length rec_prems div 2) rec_prems; (** (ts, pi1^-1 o pi2 o vs) in rec_set **) val _ = warning "step 6: (ts, pi1^-1 o pi2 o vs) in rec_set"; val rec_prems' = map (fn th => let val _ $ (S $ x $ y) = Thm.prop_of th; val Const (s, _) = head_of S; val k = find_index (equal s) rec_set_names; val pi = rpi1 @ pi2; fun mk_pi z = fold_rev (mk_perm []) pi z; fun eqvt_tac ctxt p = let val U as Type (_, [Type (_, [T, _])]) = fastype_of p; val l = find_index (equal T) dt_atomTs; val th = nth (nth rec_equiv_thms' l) k; - val th' = Thm.instantiate ([], - [((("pi", 0), U), Thm.global_cterm_of thy11 p)]) th; + val th' = Thm.instantiate (TVars.empty, + Vars.make [((("pi", 0), U), Thm.global_cterm_of thy11 p)]) th; in resolve_tac ctxt [th'] 1 end; val th' = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (S $ mk_pi x $ mk_pi y)) (fn {context = ctxt, ...} => EVERY (map (eqvt_tac ctxt) pi @ [simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @ perm_swap @ perm_fresh_fresh)) 1, resolve_tac ctxt [th] 1])) in Simplifier.simplify (put_simpset HOL_basic_ss context'' addsimps rpi1_pi2_eqs) th' end) rec_prems2; val ihs = filter (fn th => (case Thm.prop_of th of _ $ (Const (\<^const_name>\All\, _) $ _) => true | _ => false)) prems'; (** pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs **) val _ = warning "step 7: pi1 o rs = pi2 o vs , rs = pi1^-1 o pi2 o vs"; val rec_eqns = map (fn (th, ih) => let val th' = th RS (ih RS spec RS mp) RS sym; val _ $ (_ $ lhs $ rhs) = Thm.prop_of th'; fun strip_perm (_ $ _ $ t) = strip_perm t | strip_perm t = t; in Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 lhs, fold_rev (mk_perm []) pi2 (strip_perm rhs)))) (fn _ => simp_tac (put_simpset HOL_basic_ss context'' addsimps (th' :: perm_swap)) 1) end) (rec_prems' ~~ ihs); (** as # rs **) val _ = warning "step 8: as # rs"; val rec_freshs = maps (fn (rec_prem, ih) => let val _ $ (S $ x $ (y as Free (_, T))) = Thm.prop_of rec_prem; val k = find_index (equal S) rec_sets; val atoms = flat (map_filter (fn (bs, z) => if z = x then NONE else SOME bs) cargsl') in map (fn a as Free (_, aT) => let val l = find_index (equal aT) dt_atomTs; in Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT T $ a $ y)) (fn {context = ctxt, ...} => EVERY (resolve_tac ctxt [nth (nth rec_fresh_thms l) k] 1 :: map (fn th => resolve_tac ctxt [th] 1) (snd (nth finite_thss l)) @ [resolve_tac ctxt [rec_prem] 1, resolve_tac ctxt [ih] 1, REPEAT_DETERM (resolve_tac ctxt fresh_prems 1)])) end) atoms end) (rec_prems1 ~~ ihs); (** as # fK as ts rs , bs # fK bs us vs **) val _ = warning "step 9: as # fK as ts rs , bs # fK bs us vs"; fun prove_fresh_result (a as Free (_, aT)) = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ rhs')) (fn _ => EVERY [resolve_tac context'' fcbs 1, REPEAT_DETERM (resolve_tac context'' (fresh_prems @ rec_freshs) 1), REPEAT_DETERM (resolve_tac context'' (maps snd rec_fin_supp_thms') 1 THEN resolve_tac context'' rec_prems 1), resolve_tac context'' P_ind_ths 1, REPEAT_DETERM (resolve_tac context'' (P_ths @ rec_prems) 1)]); val fresh_results'' = map prove_fresh_result boundsl; fun prove_fresh_result'' ((a as Free (_, aT), b), th) = let val th' = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT rT $ fold_rev (mk_perm []) (rpi2 @ pi1) a $ fold_rev (mk_perm []) (rpi2 @ pi1) rhs')) (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_ss context'' addsimps fresh_bij) 1 THEN resolve_tac ctxt [th] 1) in Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT rT $ b $ lhs')) (fn {context = ctxt, ...} => EVERY [cut_facts_tac [th'] 1, full_simp_tac (put_simpset HOL_ss ctxt addsimps rec_eqns @ pi1_pi2_eqs @ perm_swap addsimprocs [NominalPermeq.perm_simproc_app]) 1, full_simp_tac (put_simpset HOL_ss context'' addsimps (calc_atm @ fresh_prems' @ freshs2' @ perm_fresh_fresh)) 1]) end; val fresh_results = fresh_results'' @ map prove_fresh_result'' (boundsl ~~ boundsr ~~ fresh_results''); (** cs # fK as ts rs , cs # fK bs us vs **) val _ = warning "step 10: cs # fK as ts rs , cs # fK bs us vs"; fun prove_fresh_result' recs t (a as Free (_, aT)) = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (fresh_const aT rT $ a $ t)) (fn _ => EVERY [cut_facts_tac recs 1, REPEAT_DETERM (dresolve_tac context'' (the (AList.lookup op = rec_fin_supp_thms' aT)) 1), NominalPermeq.fresh_guess_tac (put_simpset HOL_ss context'' addsimps (freshs2 @ fs_atoms @ fresh_atm @ maps snd finite_thss)) 1]); val fresh_results' = map (prove_fresh_result' rec_prems1 rhs') freshs1 @ map (prove_fresh_result' rec_prems2 lhs') freshs1; (** pi1 o (fK as ts rs) = pi2 o (fK bs us vs) **) val _ = warning "step 11: pi1 o (fK as ts rs) = pi2 o (fK bs us vs)"; val pi1_pi2_result = Goal.prove context'' [] [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (fold_rev (mk_perm []) pi1 rhs', fold_rev (mk_perm []) pi2 lhs'))) (fn _ => simp_tac (put_simpset HOL_ss context'' addsimps pi1_pi2_eqs @ rec_eqns addsimprocs [NominalPermeq.perm_simproc_app]) 1 THEN TRY (simp_tac (put_simpset HOL_ss context'' addsimps (fresh_prems' @ freshs2' @ calc_atm @ perm_fresh_fresh)) 1)); val _ = warning "final result"; val final = Goal.prove context'' [] [] (Thm.term_of concl) (fn _ => cut_facts_tac [pi1_pi2_result RS sym] 1 THEN full_simp_tac (put_simpset HOL_basic_ss context'' addsimps perm_fresh_fresh @ fresh_results @ fresh_results') 1); val final' = Proof_Context.export context'' context' [final]; val _ = warning "finished!" in resolve_tac context' final' 1 end) context 1])) idxss) (ndescr ~~ rec_elims)) end)); val rec_total_thms = map (fn r => r RS @{thm theI'}) rec_unique_thms; (* define primrec combinators *) val big_reccomb_name = space_implode "_" new_type_names ^ "_rec"; val reccomb_names = map (Sign.full_bname thy11) (if length descr'' = 1 then [big_reccomb_name] else (map ((curry (op ^) (big_reccomb_name ^ "_")) o string_of_int) (1 upto (length descr'')))); val reccombs = map (fn ((name, T), T') => list_comb (Const (name, rec_fn_Ts @ [T] ---> T'), rec_fns)) (reccomb_names ~~ recTs ~~ rec_result_Ts); val (reccomb_defs, thy12) = thy11 |> Sign.add_consts (map (fn ((name, T), T') => (Binding.name (Long_Name.base_name name), rec_fn_Ts @ [T] ---> T', NoSyn)) (reccomb_names ~~ recTs ~~ rec_result_Ts)) |> (Global_Theory.add_defs false o map Thm.no_attributes) (map (fn ((((name, comb), set), T), T') => (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (comb, absfree ("x", T) (Const (\<^const_name>\The\, (T' --> HOLogic.boolT) --> T') $ absfree ("y", T') (set $ Free ("x", T) $ Free ("y", T')))))) (reccomb_names ~~ reccombs ~~ rec_sets ~~ recTs ~~ rec_result_Ts)); (* prove characteristic equations for primrec combinators *) val rec_thms = map (fn (prems, concl) => let val _ $ (_ $ (_ $ x) $ _) = concl; val (_, cargs) = strip_comb x; val ps = map (fn (x as Free (_, T), i) => (Free ("x" ^ string_of_int i, T), x)) (cargs ~~ (1 upto length cargs)); val concl' = subst_atomic_types (rec_result_Ts' ~~ rec_result_Ts) concl; val prems' = flat finite_premss @ finite_ctxt_prems @ rec_prems @ rec_prems' @ map (subst_atomic ps) prems; fun solve ctxt rules prems = resolve_tac ctxt rules THEN_ALL_NEW (resolve_tac ctxt prems THEN_ALL_NEW assume_tac ctxt) in Goal.prove_global_future thy12 [] (map (augment_sort thy12 fs_cp_sort) prems') (augment_sort thy12 fs_cp_sort concl') (fn {context = ctxt, prems} => EVERY [rewrite_goals_tac ctxt reccomb_defs, resolve_tac ctxt @{thms the1_equality} 1, solve ctxt rec_unique_thms prems 1, resolve_tac ctxt rec_intrs 1, REPEAT (solve ctxt (prems @ rec_total_thms) prems 1)]) end) (rec_eq_prems ~~ Old_Datatype_Prop.make_primrecs reccomb_names descr' thy12); val dt_infos = map_index (make_dt_info pdescr induct reccomb_names rec_thms) (descr1 ~~ distinct_thms ~~ inject_thms); (* FIXME: theorems are stored in database for testing only *) val (_, thy13) = thy12 |> Global_Theory.add_thmss [((Binding.name "rec_equiv", flat rec_equiv_thms), []), ((Binding.name "rec_equiv'", flat rec_equiv_thms'), []), ((Binding.name "rec_fin_supp", flat rec_fin_supp_thms), []), ((Binding.name "rec_fresh", flat rec_fresh_thms), []), ((Binding.name "rec_unique", map Drule.export_without_context rec_unique_thms), []), ((Binding.name "recs", rec_thms), [])] ||> Sign.parent_path ||> map_nominal_datatypes (fold Symtab.update dt_infos); in thy13 end; val nominal_datatype = gen_nominal_datatype Old_Datatype.check_specs; val nominal_datatype_cmd = gen_nominal_datatype Old_Datatype.read_specs; val spec_cmd = Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- Parse.enum1 "|" (Parse.binding -- Scan.repeat Parse.typ -- Parse.opt_mixfix)) >> (fn (((vs, t), mx), cons) => ((t, vs, mx), map Scan.triple1 cons)); val _ = Outer_Syntax.command \<^command_keyword>\nominal_datatype\ "define nominal datatypes" (Parse.and_list1 spec_cmd >> (Toplevel.theory o nominal_datatype_cmd Old_Datatype_Aux.default_config)); end diff --git a/src/HOL/Nominal/nominal_fresh_fun.ML b/src/HOL/Nominal/nominal_fresh_fun.ML --- a/src/HOL/Nominal/nominal_fresh_fun.ML +++ b/src/HOL/Nominal/nominal_fresh_fun.ML @@ -1,181 +1,181 @@ (* Title: HOL/Nominal/nominal_fresh_fun.ML Authors: Stefan Berghofer and Julien Narboux, TU Muenchen Provides a tactic to generate fresh names and a tactic to analyse instances of the fresh_fun. *) (* FIXME proper ML structure! *) (* FIXME res_inst_tac mostly obsolete, cf. Subgoal.FOCUS *) (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) fun gen_res_inst_tac_term ctxt instf tyinst tinst elim th i st = let val cgoal = nth (cprems_of st) (i - 1); val maxidx = Thm.maxidx_of_cterm cgoal; val j = maxidx + 1; val tyinst' = map (apfst (Logic.incr_tvar j)) tyinst; val ps = Logic.strip_params (Thm.term_of cgoal); val Ts = map snd ps; val tinst' = map (fn (t, u) => (head_of (Logic.incr_indexes ([], Ts, j) t), fold_rev Term.abs ps u)) tinst; val th' = instf (map (apsnd (Thm.ctyp_of ctxt)) tyinst') (map (apsnd (Thm.cterm_of ctxt)) tinst') (Thm.lift_rule cgoal th) in compose_tac ctxt (elim, th', Thm.nprems_of th) i st end handle General.Subscript => Seq.empty; (* FIXME proper SUBGOAL/CSUBGOAL instead of cprems_of etc. *) fun res_inst_tac_term ctxt = gen_res_inst_tac_term ctxt (fn instT => fn inst => Thm.instantiate - (map (apfst dest_TVar) instT, - map (apfst dest_Var) inst)); + (TVars.make (map (apfst dest_TVar) instT), + Vars.make (map (apfst dest_Var) inst))); fun res_inst_tac_term' ctxt = gen_res_inst_tac_term ctxt (fn _ => fn inst => infer_instantiate ctxt (map (apfst (#1 o dest_Var)) inst)) []; fun cut_inst_tac_term' ctxt tinst th = res_inst_tac_term' ctxt tinst false (Rule_Insts.make_elim_preserve ctxt th); fun get_dyn_thm thy name atom_name = Global_Theory.get_thm thy name handle ERROR _ => error ("The atom type "^atom_name^" is not defined."); (* The theorems needed that are known at compile time. *) val at_exists_fresh' = @{thm "at_exists_fresh'"}; val fresh_fun_app' = @{thm "fresh_fun_app'"}; val fresh_prod = @{thm "fresh_prod"}; (* A tactic to generate a name fresh for all the free *) (* variables and parameters of the goal *) fun generate_fresh_tac ctxt atom_name = SUBGOAL (fn (goal, _) => let val thy = Proof_Context.theory_of ctxt; (* the parsing function returns a qualified name, we get back the base name *) val atom_basename = Long_Name.base_name atom_name; val ps = Logic.strip_params goal; val Ts = rev (map snd ps); fun is_of_fs_name T = Sign.of_sort thy (T, [Sign.intern_class thy ("fs_"^atom_basename)]); (* rebuild de bruijn indices *) val bvs = map_index (Bound o fst) ps; (* select variables of the right class *) val vs = filter (fn t => is_of_fs_name (fastype_of1 (Ts, t))) (Misc_Legacy.term_frees goal @ bvs); (* build the tuple *) val s = (Library.foldr1 (fn (v, s) => HOLogic.pair_const (fastype_of1 (Ts, v)) (fastype_of1 (Ts, s)) $ v $ s) vs) handle TERM _ => HOLogic.unit; val fs_name_thm = get_dyn_thm thy ("fs_"^atom_basename^"1") atom_basename; val at_name_inst_thm = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename; val exists_fresh' = at_name_inst_thm RS at_exists_fresh'; (* find the variable we want to instantiate *) val x = hd (Misc_Legacy.term_vars (Thm.prop_of exists_fresh')); in fn st => (cut_inst_tac_term' ctxt [(x,s)] exists_fresh' 1 THEN resolve_tac ctxt [fs_name_thm] 1 THEN eresolve_tac ctxt [exE] 1) st handle List.Empty => all_tac st (* if we collected no variables then we do nothing *) end) 1; fun get_inner_fresh_fun (Bound j) = NONE | get_inner_fresh_fun (v as Free _) = NONE | get_inner_fresh_fun (v as Var _) = NONE | get_inner_fresh_fun (Const _) = NONE | get_inner_fresh_fun (Abs (_, _, t)) = get_inner_fresh_fun t | get_inner_fresh_fun (Const (\<^const_name>\Nominal.fresh_fun\, Type(\<^type_name>\fun\,[Type (\<^type_name>\fun\,[Type (T,_),_]),_])) $ u) = SOME T | get_inner_fresh_fun (t $ u) = let val a = get_inner_fresh_fun u in if a = NONE then get_inner_fresh_fun t else a end; (* This tactic generates a fresh name of the atom type *) (* given by the innermost fresh_fun *) fun generate_fresh_fun_tac ctxt = SUBGOAL (fn (goal, _) => let val atom_name_opt = get_inner_fresh_fun goal; in case atom_name_opt of NONE => all_tac | SOME atom_name => generate_fresh_tac ctxt atom_name end) 1; (* Two substitution tactics which looks for the innermost occurrence in one assumption or in the conclusion *) val search_fun = curry (Seq.flat o uncurry EqSubst.searchf_bt_unify_valid); val search_fun_asm = EqSubst.skip_first_asm_occs_search EqSubst.searchf_bt_unify_valid; fun subst_inner_tac ctxt = EqSubst.eqsubst_tac' ctxt search_fun; fun subst_inner_asm_tac_aux i ctxt = EqSubst.eqsubst_asm_tac' ctxt search_fun_asm i; (* A tactic to substitute in the first assumption which contains an occurrence. *) fun subst_inner_asm_tac ctxt th = curry (curry (FIRST' (map uncurry (map uncurry (map subst_inner_asm_tac_aux (1 upto Thm.nprems_of th)))))) ctxt th; fun fresh_fun_tac ctxt no_asm = SUBGOAL (fn (goal, i) => (* Find the variable we instantiate *) let val thy = Proof_Context.theory_of ctxt; val abs_fresh = Global_Theory.get_thms thy "abs_fresh"; val fresh_perm_app = Global_Theory.get_thms thy "fresh_perm_app"; val simp_ctxt = ctxt addsimps (fresh_prod :: abs_fresh) addsimps fresh_perm_app; val x = hd (tl (Misc_Legacy.term_vars (Thm.prop_of exI))); val atom_name_opt = get_inner_fresh_fun goal; val n = length (Logic.strip_params goal); (* Here we rely on the fact that the variable introduced by generate_fresh_tac *) (* is the last one in the list, the inner one *) in case atom_name_opt of NONE => all_tac | SOME atom_name => let val atom_basename = Long_Name.base_name atom_name; val pt_name_inst = get_dyn_thm thy ("pt_"^atom_basename^"_inst") atom_basename; val at_name_inst = get_dyn_thm thy ("at_"^atom_basename^"_inst") atom_basename; fun inst_fresh vars params i st = let val vars' = Misc_Legacy.term_vars (Thm.prop_of st); in case subtract (op =) vars vars' of [Var v] => Seq.single - (Thm.instantiate ([], - [(v, Thm.cterm_of ctxt (fold_rev Term.abs params (Bound 0)))]) st) + (Thm.instantiate (TVars.empty, + Vars.make [(v, Thm.cterm_of ctxt (fold_rev Term.abs params (Bound 0)))]) st) | _ => error "fresh_fun_simp: Too many variables, please report." end in ((fn st => let val vars = Misc_Legacy.term_vars (Thm.prop_of st); val params = Logic.strip_params (nth (Thm.prems_of st) (i-1)) (* The tactics which solve the subgoals generated by the conditionnal rewrite rule. *) val post_rewrite_tacs = [resolve_tac ctxt [pt_name_inst], resolve_tac ctxt [at_name_inst], TRY o SOLVED' (NominalPermeq.finite_guess_tac simp_ctxt), inst_fresh vars params THEN' (TRY o SOLVED' (NominalPermeq.fresh_guess_tac simp_ctxt)) THEN' (TRY o SOLVED' (asm_full_simp_tac simp_ctxt))] in ((if no_asm then no_tac else (subst_inner_asm_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) ORELSE (subst_inner_tac ctxt fresh_fun_app' i THEN (RANGE post_rewrite_tacs i))) st end)) end end) diff --git a/src/HOL/Nominal/nominal_inductive.ML b/src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML +++ b/src/HOL/Nominal/nominal_inductive.ML @@ -1,694 +1,694 @@ (* Title: HOL/Nominal/nominal_inductive.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. *) signature NOMINAL_INDUCTIVE = sig val prove_strong_ind: string -> (string * string list) list -> local_theory -> Proof.state val prove_eqvt: string -> string list -> local_theory -> local_theory end structure NominalInductive : NOMINAL_INDUCTIVE = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val fresh_prod = @{thm fresh_prod}; val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else insert (op aconv o apply2 fst) (incr_boundvars (~i) u, T)) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn _ => EVERY [cut_facts_tac [th] 1, eresolve_tac ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)), REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; val eta_contract_cterm = Thm.dest_arg o Thm.cprop_of o Thm.eta_conversion; fun first_order_matchs pats objs = Thm.first_order_match (eta_contract_cterm (Conjunction.mk_conjunction_balanced pats), eta_contract_cterm (Conjunction.mk_conjunction_balanced objs)); fun first_order_mrs ths th = ths MRS Thm.instantiate (first_order_matchs (cprems_of th) (map Thm.cprop_of ths)) th; fun prove_strong_ind s avoids ctxt = let val thy = Proof_Context.theory_of ctxt; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global ctxt (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct ctxt raw_induct; val elims = map (atomize_induct ctxt) elims; val monos = Inductive.get_monos ctxt; val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP ctxt (hd names))))); val (raw_induct', ctxt') = ctxt |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = avoids |> forall (fn (a, xs) => null (duplicates (op =) xs) orelse error ("Duplicate variable names for case " ^ quote a)); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); val avoids' = if null induct_cases then replicate (length intrs) ("", []) else map (fn name => (name, the_default [] (AList.lookup op = avoids name))) induct_cases; fun mk_avoids params (name, ps) = let val k = length params - 1 in map (fn x => case find_index (equal x o fst) params of ~1 => error ("No such variable in case " ^ quote name ^ " of inductive definition: " ^ quote x) | i => (Bound (k - i), snd (nth params i))) ps end; val prems = map (fn (prem, avoid) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, fold (add_binders thy 0) (prems @ [concl]) [] @ map (apfst (incr_boundvars 1)) (mk_avoids params avoid), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ avoids'); val atomTs = distinct op = (maps (map snd o #2) prems); val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn T => Sign.intern_class thy ("fs_" ^ Long_Name.base_name (fst (dest_Type T)))) atomTs)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_distinct [] = [] | mk_distinct ((x, T) :: xs) = map_filter (fn (y, U) => if T = U then SOME (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.eq_const T $ x $ y))) else NONE) xs @ mk_distinct xs; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, bvars, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh bvars @ mk_distinct bvars @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term prem params') in (Logic.list_all (params', prem), (rev vs, subst_bounds (vs, prem))) end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val vc_compat = map (fn (params, bvars, prems, (p, ts)) => map (fn q => Logic.list_all (params, incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (mk_distinct bvars @ maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_const U T $ u $ t)) bvars) (ts ~~ binder_types (fastype_of p)))) prems; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn aT => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "2")) atomTs; val eqvt_ss = simpset_of (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]); val fresh_bij = Global_Theory.get_thms thy "fresh_bij"; val perm_bij = Global_Theory.get_thms thy "perm_bij"; val fs_atoms = map (fn aT => Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name (fst (dest_Type aT)) ^ "1")) atomTs; val exists_fresh' = Global_Theory.get_thms thy "exists_fresh'"; val fresh_atm = Global_Theory.get_thms thy "fresh_atm"; val swap_simps = Global_Theory.get_thms thy "swap_simps"; val perm_fresh_fresh = Global_Theory.get_thms thy "perm_fresh_fresh"; fun obtain_fresh_name ts T (freshs1, freshs2, ctxt) = let (** protect terms to avoid that fresh_prod interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts @ freshs1); val ex = Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (HOLogic.exists_const T $ Abs ("x", T, NominalDatatype.fresh_const T (fastype_of p) $ Bound 0 $ p))) (fn _ => EVERY [resolve_tac ctxt exists_fresh' 1, resolve_tac ctxt fs_atoms 1]); val (([(_, cx)], ths), ctxt') = Obtain.result (fn ctxt' => EVERY [eresolve_tac ctxt' [exE] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps (fresh_prod :: fresh_atm)) 1, full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, REPEAT (eresolve_tac ctxt [conjE] 1)]) [ex] ctxt in (freshs1 @ [Thm.term_of cx], freshs2 @ ths, ctxt') end; fun mk_ind_proof ctxt' thss = Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} => let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => resolve_tac context [raw_induct] 1 THEN EVERY (maps (fn ((((_, bvars, oprems, _), vc_compat_ths), ihyp), (vs, ihypt)) => [REPEAT (resolve_tac context [allI] 1), simp_tac (put_simpset eqvt_ss context) 1, SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => let val (params', (pis, z)) = chop (length params - length atomTs - 1) (map (Thm.term_of o #2) params) ||> split_last; val bvars' = map (fn (Bound i, T) => (nth params' (length params' - i), T) | (t, T) => (t, T)) bvars; val pi_bvars = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) bvars'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val (freshs1, freshs2, ctxt'') = fold (obtain_fresh_name (ts @ pi_bvars)) (map snd bvars') ([], [], ctxt'); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis' = map NominalDatatype.perm_of_pair (pi_bvars ~~ freshs1); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold (concat_perm #> map) pis' pis; val env = Pattern.first_order_match thy (ihypt, Thm.prop_of ihyp) (Vartab.empty, Vartab.empty); - val ihyp' = Thm.instantiate ([], - map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t)) + val ihyp' = Thm.instantiate (TVars.empty, + Vars.make (map (fn (v, t) => (dest_Var v, Thm.global_cterm_of thy t)) (map (Envir.subst_term env) vs ~~ map (fold_rev (NominalDatatype.mk_perm []) - (rev pis' @ pis)) params' @ [z])) ihyp; + (rev pis' @ pis)) params' @ [z]))) ihyp; fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (put_simpset eqvt_ss ctxt') (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt') (rev pis' @ pis) th)); val (gprems1, gprems2) = split_list (map (fn (th, t) => if null (preds_of ps t) then (SOME th, mk_pi th) else (map_thm ctxt' (split_conj (K o I) names) (eresolve_tac ctxt' [conjunct1] 1) monos NONE th, mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th)))) (gprems ~~ oprems)) |>> map_filter I; val vc_compat_ths' = map (fn th => let val th' = first_order_mrs gprems1 th; val (bop, lhs, rhs) = (case Thm.concl_of th' of _ $ (fresh $ lhs $ rhs) => (fn t => fn u => fresh $ t $ u, lhs, rhs) | _ $ (_ $ (_ $ lhs $ rhs)) => (curry (HOLogic.mk_not o HOLogic.mk_eq), lhs, rhs)); val th'' = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (bop (fold_rev (NominalDatatype.mk_perm []) pis lhs) (fold_rev (NominalDatatype.mk_perm []) pis rhs))) (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps (fresh_bij @ perm_bij)) 1 THEN resolve_tac ctxt' [th'] 1) in Simplifier.simplify (put_simpset eqvt_ss ctxt'' addsimps fresh_atm) th'' end) vc_compat_ths; val vc_compat_ths'' = NominalDatatype.mk_not_sym vc_compat_ths'; (** Since swap_simps simplifies (pi :: 'a prm) o (x :: 'b) to x **) (** we have to pre-simplify the rewrite rules **) val swap_simps_simpset = put_simpset HOL_ss ctxt'' addsimps swap_simps @ map (Simplifier.simplify (put_simpset HOL_ss ctxt'' addsimps swap_simps)) (vc_compat_ths'' @ freshs2'); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, resolve_tac ctxt'' [ihyp'] 1, REPEAT_DETERM_N (Thm.nprems_of ihyp - length gprems) (simp_tac swap_simps_simpset 1), REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac ctxt'' gprems2 1)])); val final = Goal.prove ctxt'' [] [] (Thm.term_of concl) (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt'' addsimps vc_compat_ths'' @ freshs2' @ perm_fresh_fresh @ fresh_atm) 1); val final' = Proof_Context.export ctxt'' ctxt' [final]; in resolve_tac ctxt' final' 1 end) context 1]) (prems ~~ thss ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac ctxt [conjE] 1) THEN REPEAT (REPEAT (resolve_tac ctxt [conjI, impI] 1) THEN eresolve_tac ctxt [impE] 1 THEN assume_tac ctxt 1 THEN REPEAT (eresolve_tac ctxt @{thms allE_Nil} 1) THEN asm_full_simp_tac ctxt 1) end) |> singleton (Proof_Context.export ctxt' ctxt); (** strong case analysis rule **) val cases_prems = map (fn ((name, avoids), rule) => let val ([rule'], ctxt') = Variable.import_terms false [Thm.prop_of rule] ctxt; val prem :: prems = Logic.strip_imp_prems rule'; val concl = Logic.strip_imp_concl rule' in (prem, List.drop (snd (strip_comb (HOLogic.dest_Trueprop prem)), length ind_params), concl, fold_map (fn (prem, (_, avoid)) => fn ctxt => let val prems = Logic.strip_assums_hyp prem; val params = Logic.strip_params prem; val bnds = fold (add_binders thy 0) prems [] @ mk_avoids params avoid; fun mk_subst (p as (s, T)) (i, j, ctxt, ps, qs, is, ts) = if member (op = o apsnd fst) bnds (Bound i) then let val ([s'], ctxt') = Variable.variant_fixes [s] ctxt; val t = Free (s', T) in (i + 1, j, ctxt', ps, (t, T) :: qs, i :: is, t :: ts) end else (i + 1, j + 1, ctxt, p :: ps, qs, is, Bound j :: ts); val (_, _, ctxt', ps, qs, is, ts) = fold_rev mk_subst params (0, 0, ctxt, [], [], [], []) in ((ps, qs, is, map (curry subst_bounds (rev ts)) prems), ctxt') end) (prems ~~ avoids) ctxt') end) (Inductive.partition_rules' raw_induct (intrs ~~ avoids') ~~ elims); val cases_prems' = map (fn (prem, args, concl, (prems, _)) => let fun mk_prem (ps, [], _, prems) = Logic.list_all (ps, Logic.list_implies (prems, concl)) | mk_prem (ps, qs, _, prems) = Logic.list_all (ps, Logic.mk_implies (Logic.list_implies (mk_distinct qs @ maps (fn (t, T) => map (fn u => HOLogic.mk_Trueprop (NominalDatatype.fresh_const T (fastype_of u) $ t $ u)) args) qs, HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map HOLogic.dest_Trueprop prems))), concl)) in map mk_prem prems end) cases_prems; val cases_eqvt_simpset = put_simpset HOL_ss (Proof_Context.init_global thy) addsimps eqvt_thms @ swap_simps @ perm_pi_simp addsimprocs [NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val simp_fresh_atm = map (Simplifier.simplify (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps fresh_atm)); fun mk_cases_proof ((((name, thss), elim), (prem, args, concl, (prems, ctxt'))), prems') = (name, Goal.prove ctxt' [] (prem :: prems') concl (fn {prems = hyp :: hyps, context = ctxt1} => EVERY (resolve_tac ctxt1 [hyp RS elim] 1 :: map (fn (((_, vc_compat_ths), case_hyp), (_, qs, is, _)) => SUBPROOF (fn {prems = case_hyps, params, context = ctxt2, concl, ...} => if null qs then resolve_tac ctxt2 [first_order_mrs case_hyps case_hyp] 1 else let val params' = map (Thm.term_of o #2 o nth (rev params)) is; val tab = params' ~~ map fst qs; val (hyps1, hyps2) = chop (length args) case_hyps; (* turns a = t and [x1 # t, ..., xn # t] *) (* into [x1 # a, ..., xn # a] *) fun inst_fresh th' ths = let val (ths1, ths2) = chop (length qs) ths in (map (fn th => let val (cf, ct) = Thm.dest_comb (Thm.dest_arg (Thm.cprop_of th)); val arg_cong' = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ct)] [NONE, SOME ct, SOME cf] (arg_cong RS iffD2); val inst = Thm.first_order_match (ct, Thm.dest_arg (Thm.dest_arg (Thm.cprop_of th'))) in [th', th] MRS Thm.instantiate inst arg_cong' end) ths1, ths2) end; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length args * length qs) (map (first_order_mrs hyps2) vc_compat_ths); val vc_compat_ths' = NominalDatatype.mk_not_sym vc_compat_ths1 @ flat (fst (fold_map inst_fresh hyps1 vc_compat_ths2)); val (freshs1, freshs2, ctxt3) = fold (obtain_fresh_name (args @ map fst qs @ params')) (map snd qs) ([], [], ctxt2); val freshs2' = NominalDatatype.mk_not_sym freshs2; val pis = map (NominalDatatype.perm_of_pair) ((freshs1 ~~ map fst qs) @ (params' ~~ freshs1)); val mk_pis = fold_rev (mk_perm_bool ctxt3) (map (Thm.cterm_of ctxt3) pis); val obj = Thm.global_cterm_of thy (foldr1 HOLogic.mk_conj (map (map_aterms (fn x as Free _ => if member (op =) args x then x else (case AList.lookup op = tab x of SOME y => y | NONE => fold_rev (NominalDatatype.mk_perm []) pis x) | x => x) o HOLogic.dest_Trueprop o Thm.prop_of) case_hyps)); val inst = Thm.first_order_match (Thm.dest_arg (Drule.strip_imp_concl (hd (cprems_of case_hyp))), obj); val th = Goal.prove ctxt3 [] [] (Thm.term_of concl) (fn {context = ctxt4, ...} => resolve_tac ctxt4 [Thm.instantiate inst case_hyp] 1 THEN SUBPROOF (fn {context = ctxt5, prems = fresh_hyps, ...} => let val fresh_hyps' = NominalDatatype.mk_not_sym fresh_hyps; val case_simpset = cases_eqvt_simpset addsimps freshs2' @ simp_fresh_atm (vc_compat_ths' @ fresh_hyps'); val fresh_fresh_simpset = case_simpset addsimps perm_fresh_fresh; val hyps1' = map (mk_pis #> Simplifier.simplify fresh_fresh_simpset) hyps1; val hyps2' = map (mk_pis #> Simplifier.simplify case_simpset) hyps2; val case_hyps' = hyps1' @ hyps2' in simp_tac case_simpset 1 THEN REPEAT_DETERM (TRY (resolve_tac ctxt5 [conjI] 1) THEN resolve_tac ctxt5 case_hyps' 1) end) ctxt4 1) val final = Proof_Context.export ctxt3 ctxt2 [th] in resolve_tac ctxt2 final 1 end) ctxt1 1) (thss ~~ hyps ~~ prems))) |> singleton (Proof_Context.export ctxt' ctxt)) in ctxt'' |> Proof.theorem NONE (fn thss => fn ctxt => (* FIXME ctxt/ctxt' should be called lthy/lthy' *) let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr ctxt)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof ctxt thss' |> Inductive.rulify ctxt; val strong_cases = map (mk_cases_proof ##> Inductive.rulify ctxt) (thsss ~~ elims ~~ cases_prems ~~ cases_prems'); val strong_induct_atts = map (Attrib.internal o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note ((rec_qualified (Binding.name "strong_induct"), strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects ctxt (1 upto length names) strong_induct'; in ctxt' |> Local_Theory.notes [((rec_qualified (Binding.name "strong_inducts"), []), strong_inducts |> map (fn th => ([th], [Attrib.internal (K ind_case_names), Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd |> Local_Theory.notes (map (fn ((name, elim), (_, cases)) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "strong_cases"), [Attrib.internal (K (Rule_Cases.case_names (map snd cases))), Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of elim)))]), [([elim], [])])) (strong_cases ~~ induct_cases')) |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; fun prove_eqvt s xatoms ctxt = (* FIXME ctxt should be called lthy *) let val thy = Proof_Context.theory_of ctxt; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global ctxt (Sign.intern_const thy s); val raw_induct = atomize_induct ctxt raw_induct; val elims = map (atomize_induct ctxt) elims; val intrs = map (atomize_intr ctxt) intrs; val monos = Inductive.get_monos ctxt; val intrs' = Inductive.unpartition_rules intrs (map (fn (((s, ths), (_, k)), th) => (s, ths ~~ Inductive.infer_intro_vars thy th k ths)) (Inductive.partition_rules raw_induct intrs ~~ Inductive.arities_of raw_induct ~~ elims)); val k = length (Inductive.params_of raw_induct); val atoms' = NominalAtoms.atoms_of thy; val atoms = if null xatoms then atoms' else let val atoms = map (Sign.intern_type thy) xatoms in (case duplicates op = atoms of [] => () | xs => error ("Duplicate atoms: " ^ commas xs); case subtract (op =) atoms' atoms of [] => () | xs => error ("No such atoms: " ^ commas xs); atoms) end; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val (([t], [pi]), ctxt') = ctxt |> Variable.import_terms false [Thm.concl_of raw_induct] ||>> Variable.variant_fixes ["pi"]; val eqvt_simpset = put_simpset HOL_basic_ss ctxt' addsimps (NominalThmDecls.get_eqvt_thms ctxt' @ perm_pi_simp) addsimprocs [mk_perm_bool_simproc names, NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]; val ps = map (fst o HOLogic.dest_imp) (HOLogic.dest_conj (HOLogic.dest_Trueprop t)); fun eqvt_tac pi (intr, vs) st = let fun eqvt_err s = let val ([t], ctxt'') = Variable.import_terms true [Thm.prop_of intr] ctxt' in error ("Could not prove equivariance for introduction rule\n" ^ Syntax.string_of_term ctxt'' t ^ "\n" ^ s) end; val res = SUBPROOF (fn {context = ctxt'', prems, params, ...} => let val prems' = map (fn th => the_default th (map_thm ctxt'' (split_conj (K I) names) (eresolve_tac ctxt'' [conjunct2] 1) monos NONE th)) prems; val prems'' = map (fn th => Simplifier.simplify eqvt_simpset (mk_perm_bool ctxt'' (Thm.cterm_of ctxt'' pi) th)) prems'; val intr' = infer_instantiate ctxt'' (map (#1 o dest_Var) vs ~~ map (Thm.cterm_of ctxt'' o NominalDatatype.mk_perm [] pi o Thm.term_of o #2) params) intr in (resolve_tac ctxt'' [intr'] THEN_ALL_NEW (TRY o resolve_tac ctxt'' prems'')) 1 end) ctxt' 1 st in case (Seq.pull res handle THM (s, _, _) => eqvt_err s) of NONE => eqvt_err ("Rule does not match goal\n" ^ Syntax.string_of_term ctxt' (hd (Thm.prems_of st))) | SOME (th, _) => Seq.single th end; val thss = map (fn atom => let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) in map (fn th => zero_var_indexes (th RS mp)) (Old_Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => let val (h, ts) = strip_comb p; val (ts1, ts2) = chop k ts in HOLogic.mk_imp (p, list_comb (h, ts1 @ map (NominalDatatype.mk_perm [] pi') ts2)) end) ps))) (fn _ => EVERY (resolve_tac ctxt' [raw_induct] 1 :: map (fn intr_vs => full_simp_tac eqvt_simpset 1 THEN eqvt_tac pi' intr_vs) intrs')) |> singleton (Proof_Context.export ctxt' ctxt))) end) atoms in ctxt |> Local_Theory.notes (map (fn (name, ths) => ((Binding.qualified_name (Long_Name.qualify (Long_Name.base_name name) "eqvt"), [Attrib.internal (K NominalThmDecls.eqvt_add)]), [(ths, [])])) (names ~~ transp thss)) |> snd end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive\ "prove equivariance and strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\avoids\ |-- Parse.and_list1 (Parse.name -- (\<^keyword>\:\ |-- Scan.repeat1 Parse.name))) [] >> (fn (name, avoids) => prove_strong_ind name avoids)); val _ = Outer_Syntax.local_theory \<^command_keyword>\equivariance\ "prove equivariance for inductive predicate involving nominal datatypes" (Parse.name -- Scan.optional (\<^keyword>\[\ |-- Parse.list1 Parse.name --| \<^keyword>\]\) [] >> (fn (name, atoms) => prove_eqvt name atoms)); end diff --git a/src/HOL/Nominal/nominal_inductive2.ML b/src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML +++ b/src/HOL/Nominal/nominal_inductive2.ML @@ -1,498 +1,500 @@ (* Title: HOL/Nominal/nominal_inductive2.ML Author: Stefan Berghofer, TU Muenchen Infrastructure for proving equivariance and strong induction theorems for inductive predicates involving nominal datatypes. Experimental version that allows to avoid lists of atoms. *) signature NOMINAL_INDUCTIVE2 = sig val prove_strong_ind: string -> string option -> (string * string list) list -> local_theory -> Proof.state end structure NominalInductive2 : NOMINAL_INDUCTIVE2 = struct val inductive_forall_def = @{thm HOL.induct_forall_def}; val inductive_atomize = @{thms induct_atomize}; val inductive_rulify = @{thms induct_rulify}; fun rulify_term thy = Raw_Simplifier.rewrite_term thy inductive_rulify []; fun atomize_conv ctxt = Raw_Simplifier.rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset HOL_basic_ss ctxt addsimps inductive_atomize); fun atomize_intr ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (atomize_conv ctxt)); fun atomize_induct ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv ~1 (atomize_conv ctxt))) ctxt)); fun fresh_postprocess ctxt = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fresh_star_set_eq}, @{thm fresh_star_Un_elim}, @{thm fresh_star_insert_elim}, @{thm fresh_star_empty_elim}]); fun preds_of ps t = inter (op = o apsnd dest_Free) ps (Term.add_frees t []); val perm_bool = mk_meta_eq @{thm perm_bool_def}; val perm_boolI = @{thm perm_boolI}; val (_, [perm_boolI_pi, _]) = Drule.strip_comb (snd (Thm.dest_comb (Drule.strip_imp_concl (Thm.cprop_of perm_boolI)))); fun mk_perm_bool ctxt pi th = th RS infer_instantiate ctxt [(#1 (dest_Var (Thm.term_of perm_boolI_pi)), pi)] perm_boolI; fun mk_perm_bool_simproc names = Simplifier.make_simproc \<^context> "perm_bool" {lhss = [\<^term>\perm pi x\], proc = fn _ => fn _ => fn ct => (case Thm.term_of ct of Const (\<^const_name>\Nominal.perm\, _) $ _ $ t => if member (op =) names (the_default "" (try (head_of #> dest_Const #> fst) t)) then SOME perm_bool else NONE | _ => NONE)}; fun transp ([] :: _) = [] | transp xs = map hd xs :: transp (map tl xs); fun add_binders thy i (t as (_ $ _)) bs = (case strip_comb t of (Const (s, T), ts) => (case strip_type T of (Ts, Type (tname, _)) => (case NominalDatatype.get_nominal_datatype thy tname of NONE => fold (add_binders thy i) ts bs | SOME {descr, index, ...} => (case AList.lookup op = (#3 (the (AList.lookup op = descr index))) s of NONE => fold (add_binders thy i) ts bs | SOME cargs => fst (fold (fn (xs, x) => fn (bs', cargs') => let val (cargs1, (u, _) :: cargs2) = chop (length xs) cargs' in (add_binders thy i u (fold (fn (u, T) => if exists (fn j => j < i) (loose_bnos u) then I else AList.map_default op = (T, []) (insert op aconv (incr_boundvars (~i) u))) cargs1 bs'), cargs2) end) cargs (bs, ts ~~ Ts)))) | _ => fold (add_binders thy i) ts bs) | (u, ts) => add_binders thy i u (fold (add_binders thy i) ts bs)) | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; fun split_conj f names (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (f p q) else NONE | _ => NONE) | split_conj _ _ _ _ = NONE; fun strip_all [] t = t | strip_all (_ :: xs) (Const (\<^const_name>\All\, _) $ Abs (s, T, t)) = strip_all xs t; (*********************************************************************) (* maps R ... & (ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t)) *) (* or ALL pi_1 ... pi_n z. P z (pi_1 o ... o pi_n o t) *) (* to R ... & id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* or id (ALL z. P z (pi_1 o ... o pi_n o t)) *) (* *) (* where "id" protects the subformula from simplification *) (*********************************************************************) fun inst_conj_all names ps pis (Const (\<^const_name>\HOL.conj\, _) $ p $ q) _ = (case head_of p of Const (name, _) => if member (op =) names name then SOME (HOLogic.mk_conj (p, Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis q)))) else NONE | _ => NONE) | inst_conj_all names ps pis t u = if member (op aconv) ps (head_of u) then SOME (Const (\<^const_name>\Fun.id\, HOLogic.boolT --> HOLogic.boolT) $ (subst_bounds (pis, strip_all pis t))) else NONE | inst_conj_all _ _ _ _ _ = NONE; fun inst_conj_all_tac ctxt k = EVERY [TRY (EVERY [eresolve_tac ctxt [conjE] 1, resolve_tac ctxt [conjI] 1, assume_tac ctxt 1]), REPEAT_DETERM_N k (eresolve_tac ctxt [allE] 1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm id_apply}]) 1]; fun map_term f t u = (case f t u of NONE => map_term' f t u | x => x) and map_term' f (t $ u) (t' $ u') = (case (map_term f t t', map_term f u u') of (NONE, NONE) => NONE | (SOME t'', NONE) => SOME (t'' $ u) | (NONE, SOME u'') => SOME (t $ u'') | (SOME t'', SOME u'') => SOME (t'' $ u'')) | map_term' f (Abs (s, T, t)) (Abs (s', T', t')) = (case map_term f t t' of NONE => NONE | SOME t'' => SOME (Abs (s, T, t''))) | map_term' _ _ _ = NONE; (*********************************************************************) (* Prove F[f t] from F[t], where F is monotone *) (*********************************************************************) fun map_thm ctxt f tac monos opt th = let val prop = Thm.prop_of th; fun prove t = Goal.prove ctxt [] [] t (fn _ => EVERY [cut_facts_tac [th] 1, eresolve_tac ctxt [rev_mp] 1, REPEAT_DETERM (FIRSTGOAL (resolve_tac ctxt monos)), REPEAT_DETERM (resolve_tac ctxt [impI] 1 THEN (assume_tac ctxt 1 ORELSE tac))]) in Option.map prove (map_term f prop (the_default prop opt)) end; fun abs_params params t = let val vs = map (Var o apfst (rpair 0)) (Term.rename_wrt_term t params) in (Logic.list_all (params, t), (rev vs, subst_bounds (vs, t))) end; fun inst_params thy (vs, p) th cts = let val env = Pattern.first_order_match thy (p, Thm.prop_of th) (Vartab.empty, Vartab.empty) - in Thm.instantiate ([], map (dest_Var o Envir.subst_term env) vs ~~ cts) th end; + in + Thm.instantiate (TVars.empty, Vars.make (map (dest_Var o Envir.subst_term env) vs ~~ cts)) th + end; fun prove_strong_ind s alt_name avoids ctxt = let val thy = Proof_Context.theory_of ctxt; val ({names, ...}, {raw_induct, intrs, elims, ...}) = Inductive.the_inductive_global ctxt (Sign.intern_const thy s); val ind_params = Inductive.params_of raw_induct; val raw_induct = atomize_induct ctxt raw_induct; val elims = map (atomize_induct ctxt) elims; val monos = Inductive.get_monos ctxt; val eqvt_thms = NominalThmDecls.get_eqvt_thms ctxt; val _ = (case subtract (op =) (fold (Term.add_const_names o Thm.prop_of) eqvt_thms []) names of [] => () | xs => error ("Missing equivariance theorem for predicate(s): " ^ commas_quote xs)); val induct_cases = map (fst o fst) (fst (Rule_Cases.get (the (Induct.lookup_inductP ctxt (hd names))))); val induct_cases' = if null induct_cases then replicate (length intrs) "" else induct_cases; val (raw_induct', ctxt') = ctxt |> yield_singleton (Variable.import_terms false) (Thm.prop_of raw_induct); val concls = raw_induct' |> Logic.strip_imp_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_conj |> map (HOLogic.dest_imp ##> strip_comb); val ps = map (fst o snd) concls; val _ = (case duplicates (op = o apply2 fst) avoids of [] => () | xs => error ("Duplicate case names: " ^ commas_quote (map fst xs))); val _ = (case subtract (op =) induct_cases (map fst avoids) of [] => () | xs => error ("No such case(s) in inductive definition: " ^ commas_quote xs)); fun mk_avoids params name sets = let val (_, ctxt') = Proof_Context.add_fixes (map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params) ctxt; fun mk s = let val t = Syntax.read_term ctxt' s; val t' = fold_rev absfree params t |> funpow (length params) (fn Abs (_, _, t) => t) in (t', HOLogic.dest_setT (fastype_of t)) end handle TERM _ => error ("Expression " ^ quote s ^ " to be avoided in case " ^ quote name ^ " is not a set type"); fun add_set p [] = [p] | add_set (t, T) ((u, U) :: ps) = if T = U then let val S = HOLogic.mk_setT T in (Const (\<^const_name>\sup\, S --> S --> S) $ u $ t, T) :: ps end else (u, U) :: add_set (t, T) ps in fold (mk #> add_set) sets [] end; val prems = map (fn (prem, name) => let val prems = map (incr_boundvars 1) (Logic.strip_assums_hyp prem); val concl = incr_boundvars 1 (Logic.strip_assums_concl prem); val params = Logic.strip_params prem in (params, if null avoids then map (fn (T, ts) => (HOLogic.mk_set T ts, T)) (fold (add_binders thy 0) (prems @ [concl]) []) else case AList.lookup op = avoids name of NONE => [] | SOME sets => map (apfst (incr_boundvars 1)) (mk_avoids params name sets), prems, strip_comb (HOLogic.dest_Trueprop concl)) end) (Logic.strip_imp_prems raw_induct' ~~ induct_cases'); val atomTs = distinct op = (maps (map snd o #2) prems); val atoms = map (fst o dest_Type) atomTs; val ind_sort = if null atomTs then \<^sort>\type\ else Sign.minimize_sort thy (Sign.certify_sort thy (map (fn a => Sign.intern_class thy ("fs_" ^ Long_Name.base_name a)) atoms)); val (fs_ctxt_tyname, _) = Name.variant "'n" (Variable.names_of ctxt'); val ([fs_ctxt_name], ctxt'') = Variable.variant_fixes ["z"] ctxt'; val fsT = TFree (fs_ctxt_tyname, ind_sort); val inductive_forall_def' = Thm.instantiate' [SOME (Thm.global_ctyp_of thy fsT)] [] inductive_forall_def; fun lift_pred' t (Free (s, T)) ts = list_comb (Free (s, fsT --> T), t :: ts); val lift_pred = lift_pred' (Bound 0); fun lift_prem (t as (f $ u)) = let val (p, ts) = strip_comb t in if member (op =) ps p then HOLogic.mk_induct_forall fsT $ Abs ("z", fsT, lift_pred p (map (incr_boundvars 1) ts)) else lift_prem f $ lift_prem u end | lift_prem (Abs (s, T, t)) = Abs (s, T, lift_prem t) | lift_prem t = t; fun mk_fresh (x, T) = HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const T fsT $ x $ Bound 0); val (prems', prems'') = split_list (map (fn (params, sets, prems, (p, ts)) => let val params' = params @ [("y", fsT)]; val prem = Logic.list_implies (map mk_fresh sets @ map (fn prem => if null (preds_of ps prem) then prem else lift_prem prem) prems, HOLogic.mk_Trueprop (lift_pred p ts)); in abs_params params' prem end) prems); val ind_vars = (Old_Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); val concl = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, HOLogic.list_all (ind_vars, lift_pred p (map (fold_rev (NominalDatatype.mk_perm ind_Ts) (map Bound (length atomTs downto 1))) ts)))) concls)); val concl' = HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn (prem, (p, ts)) => HOLogic.mk_imp (prem, lift_pred' (Free (fs_ctxt_name, fsT)) p ts)) concls)); val (vc_compat, vc_compat') = map (fn (params, sets, prems, (p, ts)) => map (fn q => abs_params params (incr_boundvars ~1 (Logic.list_implies (map_filter (fn prem => if null (preds_of ps prem) then SOME prem else map_term (split_conj (K o I) names) prem prem) prems, q)))) (maps (fn (t, T) => map (fn (u, U) => HOLogic.mk_Trueprop (NominalDatatype.fresh_star_const U T $ u $ t)) sets) (ts ~~ binder_types (fastype_of p)) @ map (fn (u, U) => HOLogic.mk_Trueprop (Const (\<^const_name>\finite\, HOLogic.mk_setT U --> HOLogic.boolT) $ u)) sets) |> split_list) prems |> split_list; val perm_pi_simp = Global_Theory.get_thms thy "perm_pi_simp"; val pt2_atoms = map (fn a => Global_Theory.get_thm thy ("pt_" ^ Long_Name.base_name a ^ "2")) atoms; val eqvt_ss = simpset_of (put_simpset HOL_basic_ss (Proof_Context.init_global thy) addsimps (eqvt_thms @ perm_pi_simp @ pt2_atoms) addsimprocs [mk_perm_bool_simproc [\<^const_name>\Fun.id\], NominalPermeq.perm_simproc_app, NominalPermeq.perm_simproc_fun]); val fresh_star_bij = Global_Theory.get_thms thy "fresh_star_bij"; val pt_insts = map (NominalAtoms.pt_inst_of thy) atoms; val at_insts = map (NominalAtoms.at_inst_of thy) atoms; val dj_thms = maps (fn a => map (NominalAtoms.dj_thm_of thy a) (remove (op =) a atoms)) atoms; val finite_ineq = map2 (fn th => fn th' => th' RS (th RS @{thm pt_set_finite_ineq})) pt_insts at_insts; val perm_set_forget = map (fn th => th RS @{thm dj_perm_set_forget}) dj_thms; val perm_freshs_freshs = atomTs ~~ map2 (fn th => fn th' => th' RS (th RS @{thm pt_freshs_freshs})) pt_insts at_insts; fun obtain_fresh_name ts sets (T, fin) (freshs, ths1, ths2, ths3, ctxt) = let val thy = Proof_Context.theory_of ctxt; (** protect terms to avoid that fresh_star_prod_set interferes with **) (** pairs used in introduction rules of inductive predicate **) fun protect t = let val T = fastype_of t in Const (\<^const_name>\Fun.id\, T --> T) $ t end; val p = foldr1 HOLogic.mk_prod (map protect ts); val atom = fst (dest_Type T); val {at_inst, ...} = NominalAtoms.the_atom_info thy atom; val fs_atom = Global_Theory.get_thm thy ("fs_" ^ Long_Name.base_name atom ^ "1"); val avoid_th = Thm.instantiate' [SOME (Thm.global_ctyp_of thy (fastype_of p))] [SOME (Thm.global_cterm_of thy p)] ([at_inst, fin, fs_atom] MRS @{thm at_set_avoiding}); val (([(_, cx)], th1 :: th2 :: ths), ctxt') = Obtain.result (fn ctxt' => EVERY [resolve_tac ctxt' [avoid_th] 1, full_simp_tac (put_simpset HOL_ss ctxt' addsimps [@{thm fresh_star_prod_set}]) 1, full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}]) 1, rotate_tac 1 1, REPEAT (eresolve_tac ctxt' [conjE] 1)]) [] ctxt; val (Ts1, _ :: Ts2) = chop_prefix (not o equal T) (map snd sets); val pTs = map NominalAtoms.mk_permT (Ts1 @ Ts2); val (pis1, pis2) = chop (length Ts1) (map Bound (length pTs - 1 downto 0)); val _ $ (f $ (_ $ pi $ l) $ r) = Thm.prop_of th2 val th2' = Goal.prove ctxt' [] [] (Logic.list_all (map (pair "pi") pTs, HOLogic.mk_Trueprop (f $ fold_rev (NominalDatatype.mk_perm (rev pTs)) (pis1 @ pi :: pis2) l $ r))) (fn _ => cut_facts_tac [th2] 1 THEN full_simp_tac (put_simpset HOL_basic_ss ctxt' addsimps perm_set_forget) 1) |> Simplifier.simplify (put_simpset eqvt_ss ctxt') in (freshs @ [Thm.term_of cx], ths1 @ ths, ths2 @ [th1], ths3 @ [th2'], ctxt') end; fun mk_ind_proof ctxt' thss = Goal.prove ctxt' [] prems' concl' (fn {prems = ihyps, context = ctxt} => let val th = Goal.prove ctxt [] [] concl (fn {context, ...} => resolve_tac ctxt [raw_induct] 1 THEN EVERY (maps (fn (((((_, sets, oprems, _), vc_compat_ths), vc_compat_vs), ihyp), vs_ihypt) => [REPEAT (resolve_tac ctxt [allI] 1), simp_tac (put_simpset eqvt_ss context) 1, SUBPROOF (fn {prems = gprems, params, concl, context = ctxt', ...} => let val (cparams', (pis, z)) = chop (length params - length atomTs - 1) (map #2 params) ||> (map Thm.term_of #> split_last); val params' = map Thm.term_of cparams' val sets' = map (apfst (curry subst_bounds (rev params'))) sets; val pi_sets = map (fn (t, _) => fold_rev (NominalDatatype.mk_perm []) pis t) sets'; val (P, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.term_of concl)); val gprems1 = map_filter (fn (th, t) => if null (preds_of ps t) then SOME th else map_thm ctxt' (split_conj (K o I) names) (eresolve_tac ctxt' [conjunct1] 1) monos NONE th) (gprems ~~ oprems); val vc_compat_ths' = map2 (fn th => fn p => let val th' = gprems1 MRS inst_params thy p th cparams'; val (h, ts) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of th')) in Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (list_comb (h, map (fold_rev (NominalDatatype.mk_perm []) pis) ts))) (fn _ => simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (fresh_star_bij @ finite_ineq)) 1 THEN resolve_tac ctxt' [th'] 1) end) vc_compat_ths vc_compat_vs; val (vc_compat_ths1, vc_compat_ths2) = chop (length vc_compat_ths - length sets) vc_compat_ths'; val vc_compat_ths1' = map (Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (Simplifier.rewrite (put_simpset eqvt_ss ctxt'))))) vc_compat_ths1; val (pis', fresh_ths1, fresh_ths2, fresh_ths3, ctxt'') = fold (obtain_fresh_name ts sets) (map snd sets' ~~ vc_compat_ths2) ([], [], [], [], ctxt'); fun concat_perm pi1 pi2 = let val T = fastype_of pi1 in if T = fastype_of pi2 then Const (\<^const_name>\append\, T --> T --> T) $ pi1 $ pi2 else pi2 end; val pis'' = fold_rev (concat_perm #> map) pis' pis; val ihyp' = inst_params thy vs_ihypt ihyp (map (fold_rev (NominalDatatype.mk_perm []) (pis' @ pis) #> Thm.global_cterm_of thy) params' @ [Thm.global_cterm_of thy z]); fun mk_pi th = Simplifier.simplify (put_simpset HOL_basic_ss ctxt' addsimps [@{thm id_apply}] addsimprocs [NominalDatatype.perm_simproc]) (Simplifier.simplify (put_simpset eqvt_ss ctxt') (fold_rev (mk_perm_bool ctxt' o Thm.cterm_of ctxt') (pis' @ pis) th)); val gprems2 = map (fn (th, t) => if null (preds_of ps t) then mk_pi th else mk_pi (the (map_thm ctxt' (inst_conj_all names ps (rev pis'')) (inst_conj_all_tac ctxt' (length pis'')) monos (SOME t) th))) (gprems ~~ oprems); val perm_freshs_freshs' = map (fn (th, (_, T)) => th RS the (AList.lookup op = perm_freshs_freshs T)) (fresh_ths2 ~~ sets); val th = Goal.prove ctxt'' [] [] (HOLogic.mk_Trueprop (list_comb (P $ hd ts, map (fold_rev (NominalDatatype.mk_perm []) pis') (tl ts)))) (fn _ => EVERY ([simp_tac (put_simpset eqvt_ss ctxt'') 1, resolve_tac ctxt'' [ihyp'] 1] @ map (fn th => resolve_tac ctxt'' [th] 1) fresh_ths3 @ [REPEAT_DETERM_N (length gprems) (simp_tac (put_simpset HOL_basic_ss ctxt'' addsimps [inductive_forall_def'] addsimprocs [NominalDatatype.perm_simproc]) 1 THEN resolve_tac ctxt'' gprems2 1)])); val final = Goal.prove ctxt'' [] [] (Thm.term_of concl) (fn _ => cut_facts_tac [th] 1 THEN full_simp_tac (put_simpset HOL_ss ctxt'' addsimps vc_compat_ths1' @ fresh_ths1 @ perm_freshs_freshs') 1); val final' = Proof_Context.export ctxt'' ctxt' [final]; in resolve_tac ctxt' final' 1 end) context 1]) (prems ~~ thss ~~ vc_compat' ~~ ihyps ~~ prems''))) in cut_facts_tac [th] 1 THEN REPEAT (eresolve_tac ctxt' [conjE] 1) THEN REPEAT (REPEAT (resolve_tac ctxt' [conjI, impI] 1) THEN eresolve_tac ctxt' [impE] 1 THEN assume_tac ctxt' 1 THEN REPEAT (eresolve_tac ctxt' @{thms allE_Nil} 1) THEN asm_full_simp_tac ctxt 1) end) |> fresh_postprocess ctxt' |> singleton (Proof_Context.export ctxt' ctxt); in ctxt'' |> Proof.theorem NONE (fn thss => fn ctxt => (* FIXME ctxt/ctxt' should be called lthy/lthy' *) let val rec_name = space_implode "_" (map Long_Name.base_name names); val rec_qualified = Binding.qualify false rec_name; val ind_case_names = Rule_Cases.case_names induct_cases; val induct_cases' = Inductive.partition_rules' raw_induct (intrs ~~ induct_cases); val thss' = map (map (atomize_intr ctxt)) thss; val thsss = Inductive.partition_rules' raw_induct (intrs ~~ thss'); val strong_raw_induct = mk_ind_proof ctxt thss' |> Inductive.rulify ctxt; val strong_induct_atts = map (Attrib.internal o K) [ind_case_names, Rule_Cases.consumes (~ (Thm.nprems_of strong_raw_induct))]; val strong_induct = if length names > 1 then strong_raw_induct else strong_raw_induct RSN (2, rev_mp); val (induct_name, inducts_name) = case alt_name of NONE => (rec_qualified (Binding.name "strong_induct"), rec_qualified (Binding.name "strong_inducts")) | SOME s => (Binding.name s, Binding.name (s ^ "s")); val ((_, [strong_induct']), ctxt') = ctxt |> Local_Theory.note ((induct_name, strong_induct_atts), [strong_induct]); val strong_inducts = Project_Rule.projects ctxt' (1 upto length names) strong_induct' in ctxt' |> Local_Theory.notes [((inducts_name, []), strong_inducts |> map (fn th => ([th], [Attrib.internal (K ind_case_names), Attrib.internal (K (Rule_Cases.consumes (1 - Thm.nprems_of th)))])))] |> snd end) (map (map (rulify_term thy #> rpair [])) vc_compat) end; (* outer syntax *) val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\nominal_inductive2\ "prove strong induction theorem for inductive predicate involving nominal datatypes" (Parse.name -- Scan.option (\<^keyword>\(\ |-- Parse.!!! (Parse.name --| \<^keyword>\)\)) -- (Scan.optional (\<^keyword>\avoids\ |-- Parse.enum1 "|" (Parse.name -- (\<^keyword>\:\ |-- Parse.and_list1 Parse.term))) []) >> (fn ((name, rule_name), avoids) => prove_strong_ind name rule_name avoids)); end diff --git a/src/HOL/Real_Asymp/multiseries_expansion.ML b/src/HOL/Real_Asymp/multiseries_expansion.ML --- a/src/HOL/Real_Asymp/multiseries_expansion.ML +++ b/src/HOL/Real_Asymp/multiseries_expansion.ML @@ -1,2374 +1,2374 @@ signature MULTISERIES_EXPANSION = sig type expansion_thm = thm type trimmed_thm = thm type expr = Exp_Log_Expression.expr type basis = Asymptotic_Basis.basis datatype trim_mode = Simple_Trim | Pos_Trim | Neg_Trim | Sgn_Trim datatype zeroness = IsZero | IsNonZero | IsPos | IsNeg datatype intyness = Nat of thm | Neg_Nat of thm | No_Nat datatype parity = Even of thm | Odd of thm | Unknown_Parity datatype limit = Zero_Limit of bool option | Finite_Limit of term | Infinite_Limit of bool option datatype trim_result = Trimmed of zeroness * trimmed_thm option | Aborted of order val get_intyness : Proof.context -> cterm -> intyness val get_parity : cterm -> parity val get_expansion : thm -> term val get_coeff : term -> term val get_exponent : term -> term val get_expanded_fun : thm -> term val get_eval : term -> term val expands_to_hd : thm -> thm val mk_eval_ctxt : Proof.context -> Lazy_Eval.eval_ctxt val expand : Lazy_Eval.eval_ctxt -> expr -> basis -> expansion_thm * basis val expand_term : Lazy_Eval.eval_ctxt -> term -> basis -> expansion_thm * basis val expand_terms : Lazy_Eval.eval_ctxt -> term list -> basis -> expansion_thm list * basis val limit_of_expansion : bool * bool -> Lazy_Eval.eval_ctxt -> thm * basis -> limit * thm val compute_limit : Lazy_Eval.eval_ctxt -> term -> limit * thm val compare_expansions : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> order * thm * expansion_thm * expansion_thm (* TODO DEBUG *) datatype comparison_result = Cmp_Dominated of order * thm list * zeroness * trimmed_thm * expansion_thm * expansion_thm | Cmp_Asymp_Equiv of thm * thm val compare_expansions' : Lazy_Eval.eval_ctxt -> thm * thm * Asymptotic_Basis.basis -> comparison_result val prove_at_infinity : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_at_top : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_at_bot : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_nhds : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_at_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_at_left_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_at_right_0 : Lazy_Eval.eval_ctxt -> thm * basis -> thm val prove_smallo : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_bigo : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_bigtheta : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_asymp_equiv : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_asymptotic_relation : Lazy_Eval.eval_ctxt -> thm * thm * basis -> order * thm val prove_eventually_less : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_eventually_greater : Lazy_Eval.eval_ctxt -> thm * thm * basis -> thm val prove_eventually_nonzero : Lazy_Eval.eval_ctxt -> thm * basis -> thm val extract_terms : int * bool -> Lazy_Eval.eval_ctxt -> basis -> term -> term * term option (* Internal functions *) val check_expansion : Exp_Log_Expression.expr -> expansion_thm -> expansion_thm val zero_expansion : basis -> expansion_thm val const_expansion : Lazy_Eval.eval_ctxt -> basis -> term -> expansion_thm val ln_expansion : Lazy_Eval.eval_ctxt -> trimmed_thm -> expansion_thm -> basis -> expansion_thm * basis val exp_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> basis -> expansion_thm * basis val powr_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm * basis val powr_const_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm val powr_nat_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm * basis val power_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm val root_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * term * basis -> expansion_thm val sgn_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis -> expansion_thm val min_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm val max_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * expansion_thm * basis -> expansion_thm val arctan_expansion : Lazy_Eval.eval_ctxt -> basis -> expansion_thm -> expansion_thm val ev_zeroness_oracle : Lazy_Eval.eval_ctxt -> term -> thm option val zeroness_oracle : bool -> trim_mode option -> Lazy_Eval.eval_ctxt -> term -> zeroness * thm option val whnf_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> term option * expansion_thm * thm val simplify_expansion : Lazy_Eval.eval_ctxt -> expansion_thm -> expansion_thm val simplify_term : Lazy_Eval.eval_ctxt -> term -> term val trim_expansion_while_greater : bool -> term list option -> bool -> trim_mode option -> Lazy_Eval.eval_ctxt -> thm * Asymptotic_Basis.basis -> thm * trim_result * (zeroness * thm) list val trim_expansion : bool -> trim_mode option -> Lazy_Eval.eval_ctxt -> expansion_thm * basis -> expansion_thm * zeroness * trimmed_thm option val try_drop_leading_term_ex : bool -> Lazy_Eval.eval_ctxt -> expansion_thm -> expansion_thm option val try_prove_real_eq : bool -> Lazy_Eval.eval_ctxt -> term * term -> thm option val try_prove_ev_eq : Lazy_Eval.eval_ctxt -> term * term -> thm option val prove_compare_expansions : order -> thm list -> thm val simplify_trimmed_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * trimmed_thm -> expansion_thm * trimmed_thm val retrim_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis -> expansion_thm * thm val retrim_pos_expansion : Lazy_Eval.eval_ctxt -> expansion_thm * basis * trimmed_thm -> expansion_thm * thm * trimmed_thm val register_sign_oracle : binding * (Proof.context -> int -> tactic) -> Context.generic -> Context.generic val get_sign_oracles : Context.generic -> (string * (Proof.context -> int -> tactic)) list val solve_eval_eq : thm -> thm end structure Multiseries_Expansion : MULTISERIES_EXPANSION = struct open Asymptotic_Basis open Exp_Log_Expression open Lazy_Eval structure Data = Generic_Data ( type T = (Proof.context -> int -> tactic) Name_Space.table; val empty : T = Name_Space.empty_table "sign_oracle_tactic"; val extend = I; fun merge (tactics1, tactics2) : T = Name_Space.merge_tables (tactics1, tactics2); ); fun register_sign_oracle (s, tac) ctxt = Data.map (Name_Space.define ctxt false (s, tac) #> snd) ctxt fun get_sign_oracles ctxt = Name_Space.fold_table cons (Data.get ctxt) [] fun apply_sign_oracles ctxt tac = let val oracles = get_sign_oracles (Context.Proof ctxt) fun tac' {context = ctxt, concl, ...} = if Thm.term_of concl = \<^term>\HOL.Trueprop HOL.False\ then no_tac else FIRST (map (fn tac => HEADGOAL (snd tac ctxt)) oracles) in tac THEN_ALL_NEW (Subgoal.FOCUS_PREMS tac' ctxt) end type expansion_thm = thm type trimmed_thm = thm val dest_fun = dest_comb #> fst val dest_arg = dest_comb #> snd val concl_of' = Thm.concl_of #> HOLogic.dest_Trueprop fun get_expansion thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> Term.dest_comb |> fst |> Term.dest_comb |> snd fun get_expanded_fun thm = thm |> concl_of' |> dest_fun |> dest_fun |> dest_arg (* The following function is useful in order to detect whether a given real constant is an integer, which allows us to use the "f(x) ^ n" operation instead of "f(x) powr n". This usually leads to nicer results. *) datatype intyness = Nat of thm | Neg_Nat of thm | No_Nat fun get_intyness ctxt ct = if Thm.typ_of_cterm ct = \<^typ>\Real.real\ then let val ctxt' = put_simpset HOL_basic_ss ctxt addsimps @{thms intyness_simps} val conv = Simplifier.rewrite ctxt then_conv Simplifier.rewrite ctxt' fun flip (Nat thm) = Neg_Nat (thm RS @{thm intyness_uminus}) | flip _ = No_Nat fun get_intyness' ct = case Thm.term_of ct of \<^term>\0::real\ => Nat @{thm intyness_0} | \<^term>\1::real\ => Nat @{thm intyness_1} | Const (\<^const_name>\numeral\, _) $ _ => Nat (Thm.reflexive (Thm.dest_arg ct) RS @{thm intyness_numeral}) | Const (\<^const_name>\uminus\, _) $ _ => flip (get_intyness' (Thm.dest_arg ct)) | Const (\<^const_name>\of_nat\, _) $ _ => Nat (Thm.reflexive (Thm.dest_arg ct) RS @{thm intyness_of_nat}) | _ => No_Nat val thm = conv ct val ct' = thm |> Thm.cprop_of |> Thm.dest_equals_rhs in case get_intyness' ct' of Nat thm' => Nat (Thm.transitive thm thm' RS @{thm HOL.meta_eq_to_obj_eq}) | Neg_Nat thm' => Neg_Nat (Thm.transitive thm thm' RS @{thm HOL.meta_eq_to_obj_eq}) | No_Nat => No_Nat end handle CTERM _ => No_Nat else No_Nat datatype parity = Even of thm | Odd of thm | Unknown_Parity (* TODO: powers *) fun get_parity ct = let fun inst thm cts = let val tvs = Term.add_tvars (Thm.concl_of thm) [] in case tvs of [v] => let - val thm' = Thm.instantiate ([(v, Thm.ctyp_of_cterm ct)], []) thm + val thm' = Thm.instantiate (TVars.make [(v, Thm.ctyp_of_cterm ct)], Vars.empty) thm val vs = take (length cts) (rev (Term.add_vars (Thm.concl_of thm') [])) in - Thm.instantiate ([], vs ~~ cts) thm' + Thm.instantiate (TVars.empty, Vars.make (vs ~~ cts)) thm' end | _ => raise THM ("get_parity", 0, [thm]) end val get_num = Thm.dest_arg o Thm.dest_arg in case Thm.term_of ct of Const (\<^const_name>\Groups.zero\, _) => Even (inst @{thm even_zero} []) | Const (\<^const_name>\Groups.one\, _) => Odd (inst @{thm odd_one} []) | Const (\<^const_name>\Num.numeral_class.numeral\, _) $ \<^term>\Num.One\ => Odd (inst @{thm odd_Numeral1} []) | Const (\<^const_name>\Num.numeral_class.numeral\, _) $ (\<^term>\Num.Bit0\ $ _) => Even (inst @{thm even_numeral} [get_num ct]) | Const (\<^const_name>\Num.numeral_class.numeral\, _) $ (\<^term>\Num.Bit1\ $ _) => Odd (inst @{thm odd_numeral} [get_num ct]) | Const (\<^const_name>\Groups.uminus\, _) $ _ => ( case get_parity (Thm.dest_arg ct) of Even thm => Even (@{thm even_uminusI} OF [thm]) | Odd thm => Odd (@{thm odd_uminusI} OF [thm]) | _ => Unknown_Parity) | Const (\<^const_name>\Groups.plus\, _) $ _ $ _ => ( case apply2 get_parity (Thm.dest_binop ct) of (Even thm1, Even thm2) => Even (@{thm even_addI(1)} OF [thm1, thm2]) | (Odd thm1, Odd thm2) => Even (@{thm even_addI(2)} OF [thm1, thm2]) | (Even thm1, Odd thm2) => Odd (@{thm odd_addI(1)} OF [thm1, thm2]) | (Odd thm1, Even thm2) => Odd (@{thm odd_addI(2)} OF [thm1, thm2]) | _ => Unknown_Parity) | Const (\<^const_name>\Groups.minus\, _) $ _ $ _ => ( case apply2 get_parity (Thm.dest_binop ct) of (Even thm1, Even thm2) => Even (@{thm even_diffI(1)} OF [thm1, thm2]) | (Odd thm1, Odd thm2) => Even (@{thm even_diffI(2)} OF [thm1, thm2]) | (Even thm1, Odd thm2) => Odd (@{thm odd_diffI(1)} OF [thm1, thm2]) | (Odd thm1, Even thm2) => Odd (@{thm odd_diffI(2)} OF [thm1, thm2]) | _ => Unknown_Parity) | Const (\<^const_name>\Groups.times\, _) $ _ $ _ => ( case apply2 get_parity (Thm.dest_binop ct) of (Even thm1, _) => Even (@{thm even_multI(1)} OF [thm1]) | (_, Even thm2) => Even (@{thm even_multI(2)} OF [thm2]) | (Odd thm1, Odd thm2) => Odd (@{thm odd_multI} OF [thm1, thm2]) | _ => Unknown_Parity) | Const (\<^const_name>\Power.power\, _) $ _ $ _ => let val (a, n) = Thm.dest_binop ct in case get_parity a of Odd thm => Odd (inst @{thm odd_powerI} [a, n] OF [thm]) | _ => Unknown_Parity end | _ => Unknown_Parity end fun simplify_term' facts ctxt = let val ctxt = Simplifier.add_prems facts ctxt in Thm.cterm_of ctxt #> Simplifier.rewrite ctxt #> Thm.concl_of #> Logic.dest_equals #> snd end fun simplify_term ectxt = simplify_term' (get_facts ectxt) (get_ctxt ectxt) fun simplify_eval ctxt = simplify_term' [] (put_simpset HOL_basic_ss ctxt addsimps @{thms eval_simps}) datatype zeroness = IsZero | IsNonZero | IsPos | IsNeg (* Caution: The following functions assume that the given expansion is in normal form already as far as needed. *) (* Returns the leading coefficient of the given expansion. This coefficient is a multiseries. *) fun try_get_coeff expr = case expr of Const (\<^const_name>\MS\, _) $ ( Const (\<^const_name>\MSLCons\, _) $ ( Const (\<^const_name>\Pair\, _) $ c $ _) $ _) $ _ => SOME c | _ => NONE fun get_coeff expr = expr |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> fst |> dest_comb |> snd (* Returns the coefficient of the leading term in the expansion (i.e. a real number) *) fun get_lead_coeff expr = case try_get_coeff expr of NONE => expr | SOME c => get_lead_coeff c (* Returns the exponent (w.r.t. the fastest-growing basis element) of the leading term *) fun get_exponent expr = expr |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> fst |> dest_comb |> snd |> dest_comb |> snd (* Returns the list of exponents of the leading term *) fun get_exponents exp = if fastype_of exp = \<^typ>\real\ then [] else get_exponent exp :: get_exponents (get_coeff exp) (* Returns the function that the expansion corresponds to *) fun get_eval expr = if fastype_of expr = \<^typ>\real\ then Abs ("x", \<^typ>\real\, expr) else expr |> dest_comb |> snd val eval_simps = @{thms eval_simps [THEN eq_reflection]} (* Tries to prove that the given function is eventually zero *) fun ev_zeroness_oracle ectxt t = let val ctxt = Lazy_Eval.get_ctxt ectxt val goal = betapply (\<^term>\\f::real \ real. eventually (\x. f x = 0) at_top\, t) |> HOLogic.mk_Trueprop fun tac {context = ctxt, ...} = HEADGOAL (Method.insert_tac ctxt (get_facts ectxt)) THEN Local_Defs.unfold_tac ctxt eval_simps THEN HEADGOAL (Simplifier.asm_full_simp_tac ctxt) in try (Goal.prove ctxt [] [] goal) tac end (* Encodes the kind of trimming/zeroness checking operation to be performed. Simple_Trim only checks for zeroness/non-zeroness. Pos_Trim/Neg_Trim try to prove either zeroness or positivity (resp. negativity). Sgn_Trim tries all three possibilities (positive, negative, zero). *) datatype trim_mode = Simple_Trim | Pos_Trim | Neg_Trim | Sgn_Trim (* Checks (and proves) whether the given term (assumed to be a real number) is zero, positive, or negative, depending on given flags. The "fail" flag determines whether an exception is thrown if this fails. *) fun zeroness_oracle fail mode ectxt exp = let val ctxt = Lazy_Eval.get_ctxt ectxt val eq = (exp, \<^term>\0::real\) |> HOLogic.mk_eq val goal1 = (IsZero, eq |> HOLogic.mk_Trueprop) val goal2 = case mode of SOME Pos_Trim => (IsPos, \<^term>\(<) (0::real)\ $ exp |> HOLogic.mk_Trueprop) | SOME Sgn_Trim => (IsPos, \<^term>\(<) (0::real)\ $ exp |> HOLogic.mk_Trueprop) | SOME Neg_Trim => (IsNeg, betapply (\<^term>\\x. x < (0::real)\, exp) |> HOLogic.mk_Trueprop) | _ => (IsNonZero, eq |> HOLogic.mk_not |> HOLogic.mk_Trueprop) val goals = (if mode = SOME Sgn_Trim then [(IsNeg, betapply (\<^term>\\x. x < (0::real)\, exp) |> HOLogic.mk_Trueprop)] else []) val goals = goal2 :: goals fun tac {context = ctxt, ...} = HEADGOAL (Method.insert_tac ctxt (get_facts ectxt)) THEN Local_Defs.unfold_tac ctxt eval_simps THEN HEADGOAL (apply_sign_oracles ctxt (Simplifier.asm_full_simp_tac ctxt)) fun prove (res, goal) = try (fn goal => (res, SOME (Goal.prove ctxt [] [] goal tac))) goal fun err () = let val mode_msg = case mode of SOME Simple_Trim => "whether the following constant is zero" | SOME Pos_Trim => "whether the following constant is zero or positive" | SOME Neg_Trim => "whether the following constant is zero or negative" | SOME Sgn_Trim => "the sign of the following constant" | _ => raise Match val t = simplify_term' (get_facts ectxt) ctxt exp val _ = if #verbose (#ctxt ectxt) then let val p = Pretty.str ("real_asymp failed to determine " ^ mode_msg ^ ":") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)] in Pretty.writeln p end else () in raise TERM ("zeroness_oracle", [t]) end in case prove goal1 of SOME res => res | NONE => if mode = NONE then (IsNonZero, NONE) else case get_first prove (goal2 :: goals) of NONE => if fail then err () else (IsNonZero, NONE) | SOME res => res end (* Tries to prove a given equality of real numbers. *) fun try_prove_real_eq fail ectxt (lhs, rhs) = case zeroness_oracle false NONE ectxt (\<^term>\(-) :: real => _\ $ lhs $ rhs) of (IsZero, SOME thm) => SOME (thm RS @{thm real_eqI}) | _ => if not fail then NONE else let val ctxt = get_ctxt ectxt val ts = map (simplify_term' (get_facts ectxt) ctxt) [lhs, rhs] val _ = if #verbose (#ctxt ectxt) then let val p = Pretty.str ("real_asymp failed to prove that the following two numbers are equal:") val p = Pretty.chunks (p :: map (Pretty.indent 2 o Syntax.pretty_term ctxt) ts) in Pretty.writeln p end else () in raise TERM ("try_prove_real_eq", [lhs, rhs]) end (* Tries to prove a given eventual equality of real functions. *) fun try_prove_ev_eq ectxt (f, g) = let val t = Envir.beta_eta_contract (\<^term>\\(f::real=>real) g x. f x - g x\ $ f $ g) in Option.map (fn thm => thm RS @{thm eventually_diff_zero_imp_eq}) (ev_zeroness_oracle ectxt t) end fun real_less a b = \<^term>\(<) :: real \ real \ bool\ $ a $ b fun real_eq a b = \<^term>\(=) :: real \ real \ bool\ $ a $ b fun real_neq a b = \<^term>\(\) :: real \ real \ bool\ $ a $ b (* The hook that is called by the Lazy_Eval module whenever two real numbers have to be compared *) fun real_sgn_hook ({pctxt = ctxt, facts, verbose, ...}) t = let val get_rhs = Thm.concl_of #> Logic.dest_equals #> snd fun tac {context = ctxt, ...} = HEADGOAL (Method.insert_tac ctxt (Net.content facts) THEN' (apply_sign_oracles ctxt (Simplifier.asm_full_simp_tac ctxt))) fun prove_first err [] [] = if not verbose then raise TERM ("real_sgn_hook", [t]) else let val _ = err () in raise TERM ("real_sgn_hook", [t]) end | prove_first err (goal :: goals) (thm :: thms) = (case try (Goal.prove ctxt [] [] goal) tac of SOME thm' => let val thm'' = thm' RS thm in SOME (get_rhs thm'', Conv.rewr_conv thm'') end | NONE => prove_first err goals thms) | prove_first _ _ _ = raise Match in case t of \<^term>\(=) :: real => _\ $ a $ \<^term>\0 :: real\ => let val goals = map (fn c => HOLogic.mk_Trueprop (c a \<^term>\0 :: real\)) [real_neq, real_eq] fun err () = let val facts' = Net.content facts val a' = simplify_term' facts' ctxt a val p = Pretty.str ("real_asymp failed to determine whether the following " ^ "constant is zero: ") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt a')] in Pretty.writeln p end in prove_first err goals @{thms Eq_FalseI Eq_TrueI} end | Const (\<^const_name>\COMPARE\, _) $ a $ b => let val goals = map HOLogic.mk_Trueprop [real_less a b, real_less b a, real_eq a b] fun err () = let val facts' = Net.content facts val (a', b') = apply2 (simplify_term' facts' ctxt) (a, b) val p = Pretty.str ("real_asymp failed to compare" ^ "the following two constants: ") val p = Pretty.chunks (p :: map (Pretty.indent 2 o Syntax.pretty_term ctxt) [a', b']) in Pretty.writeln p end in prove_first err goals @{thms COMPARE_intros} end | _ => NONE end (* Returns the datatype constructors registered for use with the Lazy_Eval package. All constructors on which pattern matching is performed need to be registered for evaluation to work. It should be rare for users to add additional ones. *) fun get_constructors ctxt = let val thms = Named_Theorems.get ctxt \<^named_theorems>\exp_log_eval_constructor\ fun go _ [] acc = rev acc | go f (x :: xs) acc = case f x of NONE => go f xs acc | SOME y => go f xs (y :: acc) fun map_option f xs = go f xs [] fun dest_constructor thm = case Thm.concl_of thm of Const (\<^const_name>\HOL.Trueprop\, _) $ (Const (\<^const_name>\REAL_ASYMP_EVAL_CONSTRUCTOR\, _) $ Const (c, T)) => SOME (c, length (fst (strip_type T))) | _ => NONE in thms |> map_option dest_constructor end (* Creates an evaluation context with the correct setup of constructors, equations, and hooks. *) fun mk_eval_ctxt ctxt = let val eval_eqs = (Named_Theorems.get ctxt \<^named_theorems>\real_asymp_eval_eqs\) val constructors = get_constructors ctxt in Lazy_Eval.mk_eval_ctxt ctxt constructors eval_eqs |> add_hook real_sgn_hook end (* A pattern for determining the leading coefficient of a multiseries *) val exp_pat = let val anypat = AnyPat ("_", 0) in ConsPat (\<^const_name>\MS\, [ConsPat (\<^const_name>\MSLCons\, [ConsPat (\<^const_name>\Pair\, [anypat, anypat]), anypat]), anypat]) end (* Evaluates an expansion to (weak) head normal form, so that the leading coefficient and exponent can be read off. *) fun whnf_expansion ectxt thm = let val ctxt = get_ctxt ectxt val exp = get_expansion thm val (_, _, conv) = match ectxt exp_pat exp (SOME []) val eq_thm = conv (Thm.cterm_of ctxt exp) val exp' = eq_thm |> Thm.concl_of |> Logic.dest_equals |> snd in case exp' of Const (\<^const_name>\MS\, _) $ (Const (\<^const_name>\MSLCons\, _) $ (Const (\<^const_name>\Pair\, _) $ c $ _) $ _) $ _ => (SOME c, @{thm expands_to_meta_eq_cong} OF [thm, eq_thm], eq_thm) | Const (\<^const_name>\MS\, _) $ Const (\<^const_name>\MSLNil\, _) $ _ => (NONE, @{thm expands_to_meta_eq_cong} OF [thm, eq_thm], eq_thm) | _ => raise TERM ("whnf_expansion", [exp']) end fun try_lift_function ectxt (thm, SEmpty) _ = (NONE, thm) | try_lift_function ectxt (thm, basis) cont = case whnf_expansion ectxt thm of (SOME c, thm, _) => let val f = get_expanded_fun thm val T = fastype_of c val t = Const (\<^const_name>\eval\, T --> \<^typ>\real \ real\) $ c val t = Term.betapply (Term.betapply (\<^term>\\(f::real\real) g x. f x - g x\, f), t) in case ev_zeroness_oracle ectxt t of NONE => (NONE, thm) | SOME zero_thm => let val thm' = cont ectxt (thm RS @{thm expands_to_hd''}, tl_basis basis) val thm'' = @{thm expands_to_lift_function} OF [zero_thm, thm'] in (SOME (lift basis thm''), thm) end end | _ => (NONE, thm) (* Turns an expansion theorem into an expansion theorem for the leading coefficient. *) fun expands_to_hd thm = thm RS (if fastype_of (get_expansion thm) = \<^typ>\real ms\ then @{thm expands_to_hd'} else @{thm expands_to_hd}) fun simplify_expansion ectxt thm = let val exp = get_expansion thm val ctxt = get_ctxt ectxt val eq_thm = Simplifier.rewrite ctxt (Thm.cterm_of ctxt exp) in @{thm expands_to_meta_eq_cong} OF [thm, eq_thm] end (* Simplifies a trimmed expansion and returns the simplified expansion theorem and the trimming theorem for that simplified expansion. *) fun simplify_trimmed_expansion ectxt (thm, trimmed_thm) = let val exp = get_expansion thm val ctxt = get_ctxt ectxt val eq_thm = Simplifier.rewrite ctxt (Thm.cterm_of ctxt exp) val trimmed_cong_thm = case trimmed_thm |> concl_of' |> dest_fun of Const (\<^const_name>\trimmed\, _) => @{thm trimmed_eq_cong} | Const (\<^const_name>\trimmed_pos\, _) => @{thm trimmed_pos_eq_cong} | Const (\<^const_name>\trimmed_neg\, _) => @{thm trimmed_neg_eq_cong} | _ => raise THM ("simplify_trimmed_expansion", 2, [thm, trimmed_thm]) in (@{thm expands_to_meta_eq_cong} OF [thm, eq_thm], trimmed_cong_thm OF [trimmed_thm, eq_thm]) end (* Re-normalises a trimmed expansion (so that the leading term with its (real) coefficient and all exponents can be read off. This may be necessary after lifting a trimmed expansion to a larger basis. *) fun retrim_expansion ectxt (thm, basis) = let val (c, thm, eq_thm) = whnf_expansion ectxt thm in case c of NONE => (thm, eq_thm) | SOME c => if fastype_of c = \<^typ>\real\ then (thm, eq_thm) else let val c_thm = thm RS @{thm expands_to_hd''} val (c_thm', eq_thm') = retrim_expansion ectxt (c_thm, tl_basis basis) val thm = @{thm expands_to_trim_cong} OF [thm, c_thm'] in (thm, @{thm trim_lift_eq} OF [eq_thm, eq_thm']) end end fun retrim_pos_expansion ectxt (thm, basis, trimmed_thm) = let val (thm', eq_thm) = retrim_expansion ectxt (thm, basis) in (thm', eq_thm, @{thm trimmed_pos_eq_cong} OF [trimmed_thm, eq_thm]) end (* Tries to determine whether the leading term is (identically) zero and drops it if it is. If "fail" is set, an exception is thrown when that term is a real number and zeroness cannot be determined. (Which typically indicates missing facts or case distinctions) *) fun try_drop_leading_term_ex fail ectxt thm = let val exp = get_expansion thm in if fastype_of exp = \<^typ>\real\ then NONE else if fastype_of (get_coeff exp) = \<^typ>\real\ then case zeroness_oracle fail (SOME Simple_Trim) ectxt (get_coeff exp) of (IsZero, SOME zero_thm) => SOME (@{thm drop_zero_ms'} OF [zero_thm, thm]) | _ => NONE else let val c = get_coeff exp val T = fastype_of c val t = Const (\<^const_name>\eval\, T --> \<^typ>\real \ real\) $ c in case ev_zeroness_oracle ectxt t of SOME zero_thm => SOME (@{thm expands_to_drop_zero} OF [zero_thm, thm]) | _ => NONE end end (* Tries to drop the leading term of an expansion. If this is not possible, an exception is thrown and an informative error message is printed. *) fun try_drop_leading_term ectxt thm = let fun err () = let val ctxt = get_ctxt ectxt val exp = get_expansion thm val c = get_coeff exp val t = if fastype_of c = \<^typ>\real\ then c else c |> dest_arg val t = simplify_term' (get_facts ectxt) ctxt t val _ = if #verbose (#ctxt ectxt) then let val p = Pretty.str ("real_asymp failed to prove that the following term is zero: ") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)] in Pretty.writeln p end else () in raise TERM ("try_drop_leading_term", [t]) end in case try_drop_leading_term_ex true ectxt thm of NONE => err () | SOME thm => thm end datatype trim_result = Trimmed of zeroness * trimmed_thm option | Aborted of order fun cstrip_assms ct = case Thm.term_of ct of \<^term>\(==>)\ $ _ $ _ => cstrip_assms (snd (Thm.dest_implies ct)) | _ => ct (* Trims an expansion (i.e. drops leading zero terms) and provides a trimmedness theorem. Optionally, a list of exponents can be given to instruct the function to only trim until the exponents of the leading term are lexicographically less than (or less than or equal) than the given ones. This is useful to avoid unnecessary trimming. The "strict" flag indicates whether the trimming should already be aborted when the exponents are lexicographically equal or not. The "fail" flag is passed on to the zeroness oracle and determines whether a failure to determine the sign of a real number leads to an exception. "mode" indicates what kind of trimmedness theorem will be returned: Simple_Trim only gives the default trimmedness theorem, whereas Pos_Trim/Neg_Trim/Sgn_Trim will give trimmed_pos or trimmed_neg. Giving "None" as mode will produce no trimmedness theorem; it will only drop leading zero terms until zeroness cannot be proven anymore, upon which it will stop. The main result of the function is the trimmed expansion theorem. The function returns whether the trimming has been aborted or not. If was aborted, either LESS or EQUAL will be returned, indicating whether the exponents of the leading term are now lexicographically smaller or equal to the given ones. In the other case, the zeroness of the leading coefficient is returned (zero, non-zero, positive, negative) together with a trimmedness theorem. Lastly, a list of the exponent comparison results and associated theorems is also returned, so that the caller can reconstruct the result of the lexicographic ordering without doing the exponent comparisons again. *) fun trim_expansion_while_greater strict es fail mode ectxt (thm, basis) = let val (_, thm, _) = whnf_expansion ectxt thm val thm = simplify_expansion ectxt thm val cexp = thm |> Thm.cprop_of |> cstrip_assms |> Thm.dest_arg |> Thm.dest_fun |> Thm.dest_arg val c = try_get_coeff (get_expansion thm) fun lift_trimmed_thm nz thm = let val cexp = thm |> Thm.cprop_of |> cstrip_assms |> Thm.dest_arg |> Thm.dest_fun |> Thm.dest_arg val lift_thm = case nz of IsNonZero => @{thm trimmed_eq_cong[rotated, OF _ lift_trimmed]} | IsPos => @{thm trimmed_pos_eq_cong[rotated, OF _ lift_trimmed_pos]} | IsNeg => @{thm trimmed_neg_eq_cong[rotated, OF _ lift_trimmed_neg]} | _ => raise TERM ("Unexpected zeroness result in trim_expansion", []) in Thm.reflexive cexp RS lift_thm end fun trimmed_real_thm nz = Thm.reflexive cexp RS ( case nz of IsNonZero => @{thm trimmed_eq_cong[rotated, OF _ lift_trimmed[OF trimmed_realI]]} | IsPos => @{thm trimmed_pos_eq_cong[rotated, OF _ lift_trimmed_pos[OF trimmed_pos_realI]]} | IsNeg => @{thm trimmed_neg_eq_cong[rotated, OF _ lift_trimmed_neg[OF trimmed_neg_realI]]} | _ => raise TERM ("Unexpected zeroness result in trim_expansion", [])) fun do_trim es = let val c = the c val T = fastype_of c val t = Const (\<^const_name>\eval\, T --> \<^typ>\real \ real\) $ c in if T = \<^typ>\real\ then ( case zeroness_oracle fail mode ectxt c of (IsZero, SOME zero_thm) => trim_expansion_while_greater strict es fail mode ectxt (@{thm drop_zero_ms'} OF [zero_thm, thm], basis) | (nz, SOME nz_thm) => (thm, Trimmed (nz, SOME (nz_thm RS trimmed_real_thm nz)), []) | (nz, NONE) => (thm, Trimmed (nz, NONE), [])) else case trim_expansion_while_greater strict (Option.map tl es) fail mode ectxt (thm RS @{thm expands_to_hd''}, tl_basis basis) of (c_thm', Aborted ord, thms) => (@{thm expands_to_trim_cong} OF [thm, c_thm'], Aborted ord, thms) | (c_thm', Trimmed (nz, trimmed_thm), thms) => let val thm = (@{thm expands_to_trim_cong} OF [thm, c_thm']) fun err () = raise TERM ("trim_expansion: zero coefficient should have been trimmed", [c]) in case (nz, trimmed_thm) of (IsZero, _) => if #verbose (#ctxt ectxt) then let val ctxt = get_ctxt ectxt val t' = t |> simplify_eval ctxt |> simplify_term' (get_facts ectxt) ctxt val p = Pretty.str ("trim_expansion failed to recognise zeroness of " ^ "the following term:") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t')] val _ = Pretty.writeln p in err () end else err () | (_, SOME trimmed_thm) => (thm, Trimmed (nz, SOME (trimmed_thm RS lift_trimmed_thm nz thm)), thms) | (_, NONE) => (thm, Trimmed (nz, NONE), thms) end end val minus = \<^term>\(-) :: real => real => real\ in case (c, es) of (NONE, _) => (thm, Trimmed (IsZero, NONE), []) | (SOME c, SOME (e' :: _)) => let val e = get_exponent (get_expansion thm) in case zeroness_oracle true (SOME Sgn_Trim) ectxt (minus $ e $ e') of (IsPos, SOME pos_thm) => ( case try_drop_leading_term_ex false ectxt thm of SOME thm => trim_expansion_while_greater strict es fail mode ectxt (thm, basis) | NONE => do_trim NONE |> @{apply 3(3)} (fn thms => (IsPos, pos_thm) :: thms)) | (IsNeg, SOME neg_thm) => (thm, Aborted LESS, [(IsNeg, neg_thm)]) | (IsZero, SOME zero_thm) => if not strict andalso fastype_of c = \<^typ>\real\ then (thm, Aborted EQUAL, [(IsZero, zero_thm)]) else ( case try_drop_leading_term_ex false ectxt thm of SOME thm => trim_expansion_while_greater strict es fail mode ectxt (thm, basis) | NONE => (do_trim es |> @{apply 3(3)} (fn thms => (IsZero, zero_thm) :: thms))) | _ => do_trim NONE end | _ => ( case try_drop_leading_term_ex false ectxt thm of SOME thm => trim_expansion_while_greater strict es fail mode ectxt (thm, basis) | NONE => do_trim NONE) end (* Trims an expansion without any stopping criterion. *) fun trim_expansion fail mode ectxt (thm, basis) = case trim_expansion_while_greater false NONE fail mode ectxt (thm, basis) of (thm, Trimmed (zeroness, trimmed_thm), _) => (thm, zeroness, trimmed_thm) | _ => raise Match (* Determines the sign of an expansion that has already been trimmed. *) fun determine_trimmed_sgn ectxt exp = if fastype_of exp = \<^typ>\real\ then (case zeroness_oracle true (SOME Sgn_Trim) ectxt exp of (IsPos, SOME thm) => (IsPos, thm RS @{thm trimmed_pos_realI}) | (IsNeg, SOME thm) => (IsNeg, thm RS @{thm trimmed_neg_realI}) | _ => raise TERM ("determine_trimmed_sgn", [])) else let val ct = Thm.cterm_of (get_ctxt ectxt) exp in (case determine_trimmed_sgn ectxt (get_coeff exp) of (IsPos, thm) => (IsPos, @{thm lift_trimmed_pos'} OF [thm, Thm.reflexive ct]) | (IsNeg, thm) => (IsNeg, @{thm lift_trimmed_neg'} OF [thm, Thm.reflexive ct]) | _ => raise TERM ("determine_trimmed_sgn", [])) end fun mk_compare_expansions_const T = Const (\<^const_name>\compare_expansions\, T --> T --> \<^typ>\cmp_result \ real \ real\) datatype comparison_result = Cmp_Dominated of order * thm list * zeroness * trimmed_thm * expansion_thm * expansion_thm | Cmp_Asymp_Equiv of thm * thm fun compare_expansions' _ (thm1, thm2, SEmpty) = Cmp_Asymp_Equiv (thm1, thm2) | compare_expansions' ectxt (thm1, thm2, basis) = let fun lift_trimmed_thm nz = case nz of IsPos => @{thm lift_trimmed_pos} | IsNeg => @{thm lift_trimmed_neg} | _ => raise TERM ("Unexpected zeroness result in compare_expansions'", []) val (e1, e2) = apply2 (get_expansion #> get_exponent) (thm1, thm2) val e = \<^term>\(-) :: real => _\ $ e1 $ e2 fun trim thm = trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) val try_drop = Option.map (whnf_expansion ectxt #> #2) o try_drop_leading_term_ex false ectxt fun handle_result ord zeroness trimmed_thm thm1 thm2 = let val (e1, e2) = apply2 (get_expansion #> get_exponent) (thm1, thm2) val e = \<^term>\(-) :: real => _\ $ e1 $ e2 val mode = if ord = LESS then Neg_Trim else Pos_Trim in case zeroness_oracle true (SOME mode) ectxt e of (_, SOME e_thm) => Cmp_Dominated (ord, [e_thm], zeroness, trimmed_thm, thm1, thm2) | _ => raise Match end fun recurse e_zero_thm = case basis of SNE (SSng _) => Cmp_Asymp_Equiv (thm1, thm2) | _ => let val (thm1', thm2') = apply2 (fn thm => thm RS @{thm expands_to_hd''}) (thm1, thm2) val (thm1', thm2') = apply2 (whnf_expansion ectxt #> #2) (thm1', thm2') in case compare_expansions' ectxt (thm1', thm2', tl_basis basis) of Cmp_Dominated (order, e_thms, zeroness, trimmed_thm, thm1', thm2') => Cmp_Dominated (order, e_zero_thm :: e_thms, zeroness, trimmed_thm RS lift_trimmed_thm zeroness, @{thm expands_to_trim_cong} OF [thm1, thm1'], @{thm expands_to_trim_cong} OF [thm2, thm2']) | Cmp_Asymp_Equiv (thm1', thm2') => Cmp_Asymp_Equiv (@{thm expands_to_trim_cong} OF [thm1, thm1'], @{thm expands_to_trim_cong} OF [thm2, thm2']) end in case zeroness_oracle false (SOME Sgn_Trim) ectxt e of (IsPos, SOME _) => ( case try_drop thm1 of SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => ( case trim thm1 of (thm1, zeroness, SOME trimmed_thm) => handle_result GREATER zeroness trimmed_thm thm1 thm2 | _ => raise TERM ("compare_expansions", map get_expansion [thm1, thm2]))) | (IsNeg, SOME _) => ( case try_drop thm2 of SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => ( case trim thm2 of (thm2, zeroness, SOME trimmed_thm) => handle_result LESS zeroness trimmed_thm thm1 thm2 | _ => raise TERM ("compare_expansions", map get_expansion [thm1, thm2]))) | (IsZero, SOME e_zero_thm) => ( case try_drop thm1 of SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => ( case try_drop thm2 of SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => recurse e_zero_thm)) | _ => case try_drop thm1 of SOME thm1 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => ( case try_drop thm2 of SOME thm2 => compare_expansions' ectxt (thm1, thm2, basis) | NONE => raise TERM ("compare_expansions", [e1, e2])) end (* Uses a list of exponent comparison results to show that compare_expansions has a given result.*) fun prove_compare_expansions ord [thm] = ( case ord of LESS => @{thm compare_expansions_LT_I} OF [thm] | GREATER => @{thm compare_expansions_GT_I} OF [thm] | EQUAL => @{thm compare_expansions_same_exp[OF _ compare_expansions_real]} OF [thm]) | prove_compare_expansions ord (thm :: thms) = @{thm compare_expansions_same_exp} OF [thm, prove_compare_expansions ord thms] | prove_compare_expansions _ [] = raise Match val ev_zero_pos_thm = Eventuallize.eventuallize \<^context> @{lemma "\x::real. f x = 0 \ g x > 0 \ f x < g x" by auto} NONE OF @{thms _ expands_to_imp_eventually_pos} val ev_zero_neg_thm = Eventuallize.eventuallize \<^context> @{lemma "\x::real. f x = 0 \ g x < 0 \ f x > g x" by auto} NONE OF @{thms _ expands_to_imp_eventually_neg} val ev_zero_zero_thm = Eventuallize.eventuallize \<^context> @{lemma "\x::real. f x = 0 \ g x = 0 \ f x = g x" by auto} NONE fun compare_expansions_trivial ectxt (thm1, thm2, basis) = case try_prove_ev_eq ectxt (apply2 get_expanded_fun (thm1, thm2)) of SOME thm => SOME (EQUAL, thm, thm1, thm2) | NONE => case apply2 (ev_zeroness_oracle ectxt o get_expanded_fun) (thm1, thm2) of (NONE, NONE) => NONE | (SOME zero1_thm, NONE) => ( case trim_expansion true (SOME Sgn_Trim) ectxt (thm2, basis) of (thm2, IsPos, SOME trimmed2_thm) => SOME (LESS, ev_zero_pos_thm OF [zero1_thm, get_basis_wf_thm basis, thm2, trimmed2_thm], thm1, thm2) | (thm2, IsNeg, SOME trimmed2_thm) => SOME (GREATER, ev_zero_neg_thm OF [zero1_thm, get_basis_wf_thm basis, thm2, trimmed2_thm], thm1, thm2) | _ => raise TERM ("Unexpected zeroness result in compare_expansions", [])) | (NONE, SOME zero2_thm) => ( case trim_expansion true (SOME Sgn_Trim) ectxt (thm1, basis) of (thm1, IsPos, SOME trimmed1_thm) => SOME (GREATER, ev_zero_pos_thm OF [zero2_thm, get_basis_wf_thm basis, thm1, trimmed1_thm], thm1, thm2) | (thm1, IsNeg, SOME trimmed1_thm) => SOME (LESS, ev_zero_neg_thm OF [zero2_thm, get_basis_wf_thm basis, thm1, trimmed1_thm], thm1, thm2) | _ => raise TERM ("Unexpected zeroness result in compare_expansions", [])) | (SOME zero1_thm, SOME zero2_thm) => SOME (EQUAL, ev_zero_zero_thm OF [zero1_thm, zero2_thm] , thm1, thm2) fun compare_expansions ectxt (thm1, thm2, basis) = case compare_expansions_trivial ectxt (thm1, thm2, basis) of SOME res => res | NONE => let val (_, thm1, _) = whnf_expansion ectxt thm1 val (_, thm2, _) = whnf_expansion ectxt thm2 in case compare_expansions' ectxt (thm1, thm2, basis) of Cmp_Dominated (order, e_thms, zeroness, trimmed_thm, thm1, thm2) => let val wf_thm = get_basis_wf_thm basis val cmp_thm = prove_compare_expansions order e_thms val trimmed_thm' = trimmed_thm RS (if zeroness = IsPos then @{thm trimmed_pos_imp_trimmed} else @{thm trimmed_neg_imp_trimmed}) val smallo_thm = (if order = LESS then @{thm compare_expansions_LT} else @{thm compare_expansions_GT}) OF [cmp_thm, trimmed_thm', thm1, thm2, wf_thm] val thm' = if zeroness = IsPos then @{thm smallo_trimmed_imp_eventually_less} else @{thm smallo_trimmed_imp_eventually_greater} val result_thm = thm' OF [smallo_thm, if order = LESS then thm2 else thm1, wf_thm, trimmed_thm] in (order, result_thm, thm1, thm2) end | Cmp_Asymp_Equiv (thm1, thm2) => let val thm = @{thm expands_to_minus} OF [get_basis_wf_thm basis, thm1, thm2] val (order, result_thm) = case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of (thm, IsPos, SOME pos_thm) => (GREATER, @{thm expands_to_imp_eventually_gt} OF [get_basis_wf_thm basis, thm, pos_thm]) | (thm, IsNeg, SOME neg_thm) => (LESS, @{thm expands_to_imp_eventually_lt} OF [get_basis_wf_thm basis, thm, neg_thm]) | _ => raise TERM ("Unexpected zeroness result in prove_eventually_less", []) in (order, result_thm, thm1, thm2) end end (* Throws an exception and prints an error message indicating that the leading term could not be determined to be either zero or non-zero. *) fun raise_trimming_error ectxt thm = let val ctxt = get_ctxt ectxt fun lead_coeff exp = if fastype_of exp = \<^typ>\real\ then exp else lead_coeff (get_coeff exp) val c = lead_coeff (get_expansion thm) fun err () = let val t = simplify_term' (get_facts ectxt) ctxt c val _ = if #verbose (#ctxt ectxt) then let val p = Pretty.str ("real_asymp failed to determine whether the following constant is zero:") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)] in Pretty.writeln p end else () in raise TERM ("zeroness_oracle", [t]) end in err () end (* TODO Here be dragons *) fun solve_eval_eq thm = case try (fn _ => @{thm refl} RS thm) () of SOME thm' => thm' | NONE => case try (fn _ => @{thm eval_real_def} RS thm) () of SOME thm' => thm' | NONE => @{thm eval_ms.simps} RS thm (* Returns an expansion theorem for the logarithm of the given expansion. May add one additional element to the basis at the end. *) fun ln_expansion _ _ _ SEmpty = raise TERM ("ln_expansion: empty basis", []) | ln_expansion ectxt trimmed_thm thm (SNE basis) = let fun trailing_exponent expr (SSng _) = get_exponent expr | trailing_exponent expr (SCons (_, _, tl)) = trailing_exponent (get_coeff expr) tl val e = trailing_exponent (get_expansion thm) basis fun ln_expansion_aux trimmed_thm zero_thm thm basis = let val t = betapply (\<^term>\\(f::real \ real) x. f x - 1 :: real\, get_expanded_fun thm) in case ev_zeroness_oracle ectxt t of NONE => ln_expansion_aux' trimmed_thm zero_thm thm basis | SOME zero_thm => @{thm expands_to_ln_eventually_1} OF [get_basis_wf_thm' basis, mk_expansion_level_eq_thm' basis, zero_thm] end and ln_expansion_aux' trimmed_thm zero_thm thm (SSng {wf_thm, ...}) = ( @{thm expands_to_ln} OF [trimmed_thm, wf_thm, thm, @{thm expands_to_ln_aux_0} OF [zero_thm, @{thm expands_to_ln_const}]]) |> solve_eval_eq | ln_expansion_aux' trimmed_thm zero_thm thm (SCons ({wf_thm, ...}, {ln_thm, ...}, basis')) = let val c_thm = ln_expansion_aux (trimmed_thm RS @{thm trimmed_pos_hd_coeff}) zero_thm (expands_to_hd thm) basis' val e = get_exponent (get_expansion thm) val c_thm' = case zeroness_oracle true NONE ectxt e of (IsZero, SOME thm) => @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux_0]} OF [thm,c_thm] | _ => case try_prove_real_eq false ectxt (e, \<^term>\1::real\) of SOME thm => @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux_1]} OF [thm, wf_thm, c_thm, ln_thm] | NONE => @{thm expands_to_ln_to_expands_to_ln_eval [OF expands_to_ln_aux]} OF [wf_thm, c_thm, ln_thm] in (@{thm expands_to_ln} OF [trimmed_thm, wf_thm, thm, c_thm']) |> solve_eval_eq end in case zeroness_oracle true NONE ectxt e of (IsZero, SOME zero_thm) => (ln_expansion_aux trimmed_thm zero_thm thm basis, SNE basis) | _ => let val basis' = insert_ln (SNE basis) val lifting = mk_lifting (get_basis_list' basis) basis' val thm' = lift_expands_to_thm lifting thm val trimmed_thm' = lift_trimmed_pos_thm lifting trimmed_thm val (thm'', eq_thm) = retrim_expansion ectxt (thm', basis') val trimmed_thm'' = @{thm trimmed_pos_eq_cong} OF [trimmed_thm', eq_thm] in ln_expansion ectxt trimmed_thm'' thm'' basis' end end (* Handles a possible basis change after expanding exp(c(x)) for an expansion of the form f(x) = c(x) + g(x). Expanding exp(c(x)) may have inserted an additional basis element. If the old basis was b :: bs (i.e. c is an expansion w.r.t. bs) and the updated one is bs' (which agrees with bs except for one additional element b'), we need to argue that b :: bs' is still well-formed. This may require us to show that ln(b') is o(ln(b)), which the function takes as an argument. *) fun adjust_exp_basis basis basis' ln_smallo_thm = if length (get_basis_list basis) = length (get_basis_list basis') + 1 then basis else let val SNE (SCons (info, ln_info, tail)) = basis val SNE tail' = basis' val wf_thms = map get_basis_wf_thm [basis, basis'] val wf_thm' = case get_first (fn f => try f ()) [fn _ => @{thm basis_wf_lift_modification} OF wf_thms, fn _ => @{thm basis_wf_insert_exp_near} OF (wf_thms @ [ln_smallo_thm]), fn _ => @{thm basis_wf_insert_exp_near} OF (wf_thms @ [ln_smallo_thm RS @{thm basis_wf_insert_exp_uminus'}])] of SOME wf_thm => wf_thm | _ => raise TERM ("Lifting basis modification in exp_expansion failed.", map Thm.concl_of (wf_thms @ [ln_smallo_thm])) val info' = {wf_thm = wf_thm', head = #head info} val lifting = mk_lifting (get_basis_list' tail) basis' val ln_info' = {trimmed_thm = lift_trimmed_pos_thm lifting (#trimmed_thm ln_info), ln_thm = lift_expands_to_thm lifting (#ln_thm ln_info)} in SNE (SCons (info', ln_info', tail')) end (* inserts the exponential of a given function at the beginning of the given basis *) fun insert_exp _ _ _ _ _ SEmpty = raise TERM ("insert_exp", []) | insert_exp t ln_thm ln_smallo_thm ln_trimmed_thm lim_thm (SNE basis) = let val head = Envir.beta_eta_contract (\<^term>\\(f::real\real) x. exp (f x)\ $ t) val ln_smallo_thm = ln_smallo_thm RS @{thm ln_smallo_ln_exp} val wf_thm = @{thm basis_wf_manyI} OF [lim_thm, ln_smallo_thm, get_basis_wf_thm' basis] val basis' = SNE (SCons ({wf_thm = wf_thm, head = head}, {ln_thm = ln_thm, trimmed_thm = ln_trimmed_thm} , basis)) in check_basis basis' end (* Returns an expansion of the exponential of the given expansion. This may add several new basis elements at any position of the basis (except at the very end *) fun exp_expansion _ thm SEmpty = (thm RS @{thm expands_to_exp_real}, SEmpty) | exp_expansion ectxt thm basis = let val (_, thm, _) = whnf_expansion ectxt thm in case ev_zeroness_oracle ectxt (get_eval (get_expansion thm)) of SOME zero_thm => (@{thm expands_to_exp_zero} OF [thm, zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis) | NONE => let val ln = Option.map (fn x => (#ln_thm x, #trimmed_thm x)) (get_ln_info basis) val ln = Option.map (fn (x, y) => retrim_pos_expansion ectxt (x, basis, y)) ln val es' = \<^term>\0::real\ :: ( case ln of NONE => [] | SOME (ln_thm, _, _) => get_exponents (get_expansion ln_thm)) val trim_result = trim_expansion_while_greater true (SOME es') false (SOME Simple_Trim) ectxt (thm, basis) in exp_expansion' ectxt trim_result ln basis end end and exp_expansion' _ (thm, _, _) _ SEmpty = (thm RS @{thm expands_to_exp_real}, SEmpty) | exp_expansion' ectxt (thm, trim_result, e_thms) ln basis = let val exp = get_expansion thm val wf_thm = get_basis_wf_thm basis val f = get_expanded_fun thm fun exp_expansion_insert ln_smallo_thm = ( case determine_trimmed_sgn ectxt exp of (IsPos, trimmed_thm) => let val [lim_thm, ln_thm', thm'] = @{thms expands_to_exp_insert_pos} |> map (fn thm' => thm' OF [thm, wf_thm, trimmed_thm, ln_smallo_thm]) val basis' = insert_exp f ln_thm' ln_smallo_thm trimmed_thm lim_thm basis in (thm', basis') end | (IsNeg, trimmed_thm) => let val [lim_thm, ln_thm', ln_trimmed_thm, thm'] = @{thms expands_to_exp_insert_neg} |> map (fn thm' => thm' OF [thm, wf_thm, trimmed_thm, ln_smallo_thm]) val ln_smallo_thm = ln_smallo_thm RS @{thm basis_wf_insert_exp_uminus} val f' = Envir.beta_eta_contract (\<^term>\\(f::real\real) x. -f x\ $ f) val basis' = insert_exp f' ln_thm' ln_smallo_thm ln_trimmed_thm lim_thm basis in (thm', basis') end | _ => raise TERM ("Unexpected zeroness result in exp_expansion", [])) fun lexord (IsNeg :: _) = LESS | lexord (IsPos :: _) = GREATER | lexord (IsZero :: xs) = lexord xs | lexord [] = EQUAL | lexord _ = raise Match val compare_result = lexord (map fst e_thms) in case (trim_result, e_thms, compare_result) of (Aborted _, (IsNeg, e_neg_thm) :: _, _) => (* leading exponent is negative; we can simply Taylor-expand exp(x) around 0 *) (@{thm expands_to_exp_neg} OF [thm, get_basis_wf_thm basis, e_neg_thm], basis) | (Trimmed (_, SOME trimmed_thm), (IsPos, e_pos_thm) :: _, GREATER) => (* leading exponent is positive; exp(f(x)) or exp(-f(x)) is new basis element *) let val ln_smallo_thm = @{thm basis_wf_insert_exp_pos} OF [thm, get_basis_wf_thm basis, trimmed_thm, e_pos_thm] in exp_expansion_insert ln_smallo_thm end | (Trimmed (_, SOME trimmed_thm), _, GREATER) => (* leading exponent is zero, but f(x) grows faster than ln(b(x)), so exp(f(x)) or exp(-f(x)) must still be new basis elements *) let val ln_thm = case ln of SOME (ln_thm, _, _) => ln_thm | NONE => raise TERM ("TODO blubb", []) val ln_thm = @{thm expands_to_lift''} OF [get_basis_wf_thm basis, ln_thm] val ln_smallo_thm = @{thm compare_expansions_GT} OF [prove_compare_expansions GREATER (map snd e_thms), trimmed_thm, thm, ln_thm, get_basis_wf_thm basis] in exp_expansion_insert ln_smallo_thm end | (Aborted LESS, (IsZero, e_zero_thm) :: e_thms', _) => (* leading exponent is zero and f(x) grows more slowly than ln(b(x)), so we can write f(x) = c(x) + g(x) and therefore exp(f(x)) = exp(c(x)) * exp(g(x)). The former is treated by a recursive call; the latter by Taylor expansion. *) let val (ln_thm, trimmed_thm) = case ln of SOME (ln_thm, _, trimmed_thm) => (ln_thm, trimmed_thm RS @{thm trimmed_pos_imp_trimmed}) | NONE => raise TERM ("TODO foo", []) val c_thm = expands_to_hd thm val ln_smallo_thm = @{thm compare_expansions_LT} OF [prove_compare_expansions LESS (map snd e_thms'), trimmed_thm, c_thm, ln_thm, get_basis_wf_thm (tl_basis basis)] val (c_thm, c_basis) = exp_expansion ectxt c_thm (tl_basis basis) val basis' = adjust_exp_basis basis c_basis ln_smallo_thm val wf_thm = get_basis_wf_thm basis' val thm' = lift basis' thm val (thm'', _) = retrim_expansion ectxt (thm', basis') in (@{thm expands_to_exp_0} OF [thm'', wf_thm, e_zero_thm, c_thm], basis') end | (Trimmed _, [(IsZero, e_zero_thm)], EQUAL) => (* f(x) can be written as c + g(x) where c is just a real constant. We can therefore write exp(f(x)) = exp(c) * exp(g(x)), where the latter is a simple Taylor expansion. *) (@{thm expands_to_exp_0_real} OF [thm, wf_thm, e_zero_thm], basis) | (Trimmed _, (_, e_zero_thm) :: _, EQUAL) => (* f(x) is asymptotically equivalent to c * ln(b(x)), so we can write f(x) as c * ln(b(x)) + g(x) and therefore exp(f(x)) = b(x)^c * exp(g(x)). The second factor is handled by a recursive call *) let val ln_thm = case ln of SOME (ln_thm, _, _) => ln_thm | NONE => raise TERM ("TODO blargh", []) val c = case (thm, ln_thm) |> apply2 (get_expansion #> get_lead_coeff) of (c1, c2) => \<^term>\(/) :: real => _\ $ c1 $ c2 val c = Thm.cterm_of (get_ctxt ectxt) c val thm' = @{thm expands_to_exp_0_pull_out1} OF [thm, ln_thm, wf_thm, e_zero_thm, Thm.reflexive c] val (thm'', basis') = exp_expansion ectxt thm' basis val pat = ConsPat ("MS", [AnyPat ("_", 0), AnyPat ("_", 0)]) val (_, _, conv) = match ectxt pat (get_expansion thm'') (SOME []) val eq_thm = conv (Thm.cterm_of (get_ctxt ectxt) (get_expansion thm'')) val thm''' = @{thm expands_to_meta_eq_cong} OF [thm'', eq_thm] val thm'''' = case get_intyness (get_ctxt ectxt) c of No_Nat => @{thm expands_to_exp_0_pull_out2} OF [thm''', get_basis_wf_thm basis'] | Nat nat_thm => @{thm expands_to_exp_0_pull_out2_nat} OF [thm''', get_basis_wf_thm basis', nat_thm] | Neg_Nat nat_thm => @{thm expands_to_exp_0_pull_out2_neg_nat} OF [thm''', get_basis_wf_thm basis', nat_thm] in (thm'''', basis') end | (Trimmed (IsZero, _), [], _) => (* Expansion is empty, i.e. f(x) is identically zero *) (@{thm expands_to_exp_MSLNil} OF [thm, get_basis_wf_thm basis], basis) | (Trimmed (_, NONE), _, GREATER) => (* We could not determine whether f(x) grows faster than ln(b(x)) or not. *) raise_trimming_error ectxt thm | _ => raise Match end fun powr_expansion ectxt (thm1, thm2, basis) = case ev_zeroness_oracle ectxt (get_expanded_fun thm1) of SOME zero_thm => (@{thm expands_to_powr_0} OF [zero_thm, Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) (get_expanded_fun thm2)), get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis) | NONE => let val (thm1, _, SOME trimmed_thm) = trim_expansion true (SOME Pos_Trim) ectxt (thm1, basis) val (ln_thm, basis') = ln_expansion ectxt trimmed_thm thm1 basis val thm2' = lift basis' thm2 |> simplify_expansion ectxt val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis', ln_thm, thm2'] val (exp_thm, basis'') = exp_expansion ectxt mult_thm basis' val thm = @{thm expands_to_powr} OF [trimmed_thm, get_basis_wf_thm basis, thm1, exp_thm] in (thm, basis'') end fun powr_nat_expansion ectxt (thm1, thm2, basis) = case ev_zeroness_oracle ectxt (get_expanded_fun thm1) of SOME zero_thm => ( case ev_zeroness_oracle ectxt (get_expanded_fun thm2) of SOME zero'_thm => (@{thm expands_to_powr_nat_0_0} OF [zero_thm, zero'_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis) | NONE => ( case trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis) of (thm2, _, SOME trimmed_thm) => (@{thm expands_to_powr_nat_0} OF [zero_thm, thm2, trimmed_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis], basis))) | NONE => let val (thm1, _, SOME trimmed_thm) = trim_expansion true (SOME Pos_Trim) ectxt (thm1, basis) val (ln_thm, basis') = ln_expansion ectxt trimmed_thm thm1 basis val thm2' = lift basis' thm2 |> simplify_expansion ectxt val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis', ln_thm, thm2'] val (exp_thm, basis'') = exp_expansion ectxt mult_thm basis' val thm = @{thm expands_to_powr_nat} OF [trimmed_thm, get_basis_wf_thm basis, thm1, exp_thm] in (thm, basis'') end fun is_numeral t = let val _ = HOLogic.dest_number t in true end handle TERM _ => false fun power_expansion ectxt (thm, n, basis) = case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => @{thm expands_to_power_0} OF [zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis, Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) n)] | NONE => ( case trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) of (thm', _, SOME trimmed_thm) => let val ctxt = get_ctxt ectxt val thm = if is_numeral n then @{thm expands_to_power[where abort = True]} else @{thm expands_to_power[where abort = False]} val thm = Drule.infer_instantiate' ctxt [NONE, NONE, NONE, SOME (Thm.cterm_of ctxt n)] thm in thm OF [trimmed_thm, get_basis_wf_thm basis, thm'] end | _ => raise TERM ("Unexpected zeroness result in power_expansion", [])) fun powr_const_expansion ectxt (thm, p, basis) = let val pthm = Thm.reflexive (Thm.cterm_of (get_ctxt ectxt) p) in case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => @{thm expands_to_powr_const_0} OF [zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis, pthm] | NONE => case trim_expansion true (SOME Pos_Trim) ectxt (thm, basis) of (_, _, NONE) => raise TERM ("Unexpected zeroness result for powr", []) | (thm, _, SOME trimmed_thm) => (if is_numeral p then @{thm expands_to_powr_const[where abort = True]} else @{thm expands_to_powr_const[where abort = False]}) OF [trimmed_thm, get_basis_wf_thm basis, thm, pthm] end fun sgn_expansion ectxt (thm, basis) = let val thms = [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis] in case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => @{thm expands_to_sgn_zero} OF (zero_thm :: thms) | NONE => case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of (thm, IsPos, SOME trimmed_thm) => @{thm expands_to_sgn_pos} OF ([trimmed_thm, thm] @ thms) | (thm, IsNeg, SOME trimmed_thm) => @{thm expands_to_sgn_neg} OF ([trimmed_thm, thm] @ thms) | _ => raise TERM ("Unexpected zeroness result in sgn_expansion", []) end (* Returns an expansion of the sine and cosine of the given expansion. Fails if that function goes to infinity. *) fun sin_cos_expansion _ thm SEmpty = (thm RS @{thm expands_to_sin_real}, thm RS @{thm expands_to_cos_real}) | sin_cos_expansion ectxt thm basis = let val exp = get_expansion thm val e = get_exponent exp in case zeroness_oracle true (SOME Sgn_Trim) ectxt e of (IsPos, _) => raise THM ("sin_cos_expansion", 0, [thm]) | (IsNeg, SOME e_thm) => let val [thm1, thm2] = map (fn thm' => thm' OF [e_thm, get_basis_wf_thm basis, thm]) @{thms expands_to_sin_ms_neg_exp expands_to_cos_ms_neg_exp} in (thm1, thm2) end | (IsZero, SOME e_thm) => let val (sin_thm, cos_thm) = (sin_cos_expansion ectxt (expands_to_hd thm) (tl_basis basis)) fun mk_thm thm' = (thm' OF [e_thm, get_basis_wf_thm basis, thm, sin_thm, cos_thm]) |> solve_eval_eq val [thm1, thm2] = map mk_thm @{thms expands_to_sin_ms_zero_exp expands_to_cos_ms_zero_exp} in (thm1, thm2) end | _ => raise TERM ("Unexpected zeroness result in sin_exp_expansion", []) end fun abconv (t, t') = Envir.beta_eta_contract t aconv Envir.beta_eta_contract t' (* Makes sure that an expansion theorem really talks about the right function. This is basically a sanity check to make things fail early and in the right place. *) fun check_expansion e thm = if abconv (expr_to_term e, get_expanded_fun thm) then thm else (* TODO Remove Debugging stuff *) let val _ = \<^print> e val _ = \<^print> (get_expanded_fun thm) in raise TERM ("check_expansion", [Thm.concl_of thm, expr_to_term e]) end fun minmax_expansion max [less_thm, eq_thm, gt_thm] ectxt (thm1, thm2, basis) = ( case compare_expansions ectxt (thm1, thm2, basis) of (LESS, less_thm', thm1, thm2) => less_thm OF [if max then thm2 else thm1, less_thm'] | (GREATER, gt_thm', thm1, thm2) => gt_thm OF [if max then thm1 else thm2, gt_thm'] | (EQUAL, eq_thm', thm1, _) => eq_thm OF [thm1, eq_thm']) | minmax_expansion _ _ _ _ = raise Match val min_expansion = minmax_expansion false @{thms expands_to_min_lt expands_to_min_eq expands_to_min_gt} val max_expansion = minmax_expansion true @{thms expands_to_max_lt expands_to_max_eq expands_to_max_gt} fun zero_expansion basis = @{thm expands_to_zero} OF [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis] fun const_expansion _ basis \<^term>\0 :: real\ = zero_expansion basis | const_expansion ectxt basis t = let val ctxt = get_ctxt ectxt val thm = Drule.infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt t)] @{thm expands_to_const} in thm OF [get_basis_wf_thm basis, mk_expansion_level_eq_thm basis] end fun root_expansion ectxt (thm, n, basis) = let val ctxt = get_ctxt ectxt fun tac {context = ctxt, ...} = HEADGOAL (Method.insert_tac ctxt (get_facts ectxt)) THEN Local_Defs.unfold_tac ctxt eval_simps THEN HEADGOAL (Simplifier.asm_full_simp_tac ctxt) fun prove goal = try (Goal.prove ctxt [] [] (HOLogic.mk_Trueprop (Term.betapply (goal, n)))) tac fun err () = let val t = simplify_term' (get_facts ectxt) ctxt n val _ = if #verbose (#ctxt ectxt) then let val p = Pretty.str ("real_asymp failed to determine whether the following constant " ^ "is zero or not:") val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)] in Pretty.writeln p end else () in raise TERM ("zeroness_oracle", [n]) end fun aux nz_thm = case trim_expansion true (SOME Sgn_Trim) ectxt (thm, basis) of (thm, IsPos, SOME trimmed_thm) => @{thm expands_to_root} OF [nz_thm, trimmed_thm, get_basis_wf_thm basis, thm] | (thm, IsNeg, SOME trimmed_thm) => @{thm expands_to_root_neg} OF [nz_thm, trimmed_thm, get_basis_wf_thm basis, thm] | _ => raise TERM ("Unexpected zeroness result in root_expansion", []) in case prove \<^term>\\n::nat. n = 0\ of SOME zero_thm => @{thm expands_to_0th_root} OF [zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis, Thm.reflexive (Thm.cterm_of ctxt (get_expanded_fun thm))] | NONE => case prove \<^term>\\n::nat. n > 0\ of NONE => err () | SOME nz_thm => case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => @{thm expands_to_root_0} OF [nz_thm, zero_thm, get_basis_wf_thm basis, mk_expansion_level_eq_thm basis] | NONE => aux nz_thm end fun arctan_expansion _ SEmpty thm = @{thm expands_to_real_compose[where g = arctan]} OF [thm] | arctan_expansion ectxt basis thm = case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => @{thm expands_to_arctan_zero} OF [zero_expansion basis, zero_thm] | NONE => let val (thm, _, _) = trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) val e = get_exponent (get_expansion thm) fun cont ectxt (thm, basis) = arctan_expansion ectxt basis thm in case zeroness_oracle true (SOME Sgn_Trim) ectxt e of (IsNeg, SOME neg_thm) => @{thm expands_to_arctan_ms_neg_exp} OF [neg_thm, get_basis_wf_thm basis, thm] | (IsPos, SOME e_pos_thm) => ( case determine_trimmed_sgn ectxt (get_expansion thm) of (IsPos, trimmed_thm) => @{thm expands_to_arctan_ms_pos_exp_pos} OF [e_pos_thm, trimmed_thm, get_basis_wf_thm basis, thm] | (IsNeg, trimmed_thm) => @{thm expands_to_arctan_ms_pos_exp_neg} OF [e_pos_thm, trimmed_thm, get_basis_wf_thm basis, thm] | _ => raise TERM ("Unexpected trim result during expansion of arctan", [])) | (IsZero, _) => ( case try_lift_function ectxt (thm, basis) cont of (SOME thm', _) => thm' | _ => let val _ = if get_verbose ectxt then writeln "Unsupported occurrence of arctan" else () in raise TERM ("Unsupported occurence of arctan", []) end) | _ => raise TERM ("Unexpected trim result during expansion of arctan", []) end (* Returns an expansion theorem for a function that is already a basis element *) fun expand_basic _ t SEmpty = raise TERM ("expand_basic", [t]) | expand_basic thm t basis = if abconv (get_basis_head basis, t) then thm (get_basis_wf_thm basis) (mk_expansion_level_eq_thm (tl_basis basis)) else @{thm expands_to_lift'} OF [get_basis_wf_thm basis, expand_basic thm t (tl_basis basis)] fun expand_unary ectxt thm e basis = let val (thm', basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt) in (thm OF [get_basis_wf_thm basis', thm'], basis') end and expand_binary ectxt thm (e1, e2) basis = let val (thm1, basis') = expand' ectxt e1 basis |> apfst (simplify_expansion ectxt) val (thm2, basis'') = expand' ectxt e2 basis' |> apfst (simplify_expansion ectxt) val thm1 = lift basis'' thm1 |> simplify_expansion ectxt in (thm OF [get_basis_wf_thm basis'', thm1, thm2], basis'') end and trim_nz mode ectxt e basis = let val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt) val (thm', nz, trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis') in case trimmed_thm of NONE => raise TERM ("expand: zero denominator", [get_expansion thm]) | SOME trimmed_thm => (thm', basis', nz, trimmed_thm) end and expand'' ectxt (ConstExpr c) basis = (const_expansion ectxt basis c, basis) | expand'' _ X basis = (lift basis @{thm expands_to_X}, basis) | expand'' ectxt (Uminus e) basis = expand_unary ectxt @{thm expands_to_uminus} e basis | expand'' ectxt (Add e12) basis = expand_binary ectxt @{thm expands_to_add} e12 basis | expand'' ectxt (Minus e12) basis = expand_binary ectxt @{thm expands_to_minus} e12 basis | expand'' ectxt (Mult e12) basis = expand_binary ectxt @{thm expands_to_mult} e12 basis | expand'' ectxt (Powr' (e, p)) basis = (* TODO zero basis *) let val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt) in (powr_const_expansion ectxt (thm, p, basis'), basis') end | expand'' ectxt (Powr (e1, e2)) basis = let val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt) val (thm1, basis2) = expand' ectxt e1 basis1 |> apfst (simplify_expansion ectxt) in powr_expansion ectxt (thm1, thm2, basis2) end | expand'' ectxt (Powr_Nat (e1, e2)) basis = let val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt) val (thm1, basis2) = expand' ectxt e1 basis1 |> apfst (simplify_expansion ectxt) in powr_nat_expansion ectxt (thm1, thm2, basis2) end | expand'' ectxt (LnPowr (e1, e2)) basis = let (* TODO zero base *) val (thm2, basis1) = expand' ectxt e2 basis |> apfst (simplify_expansion ectxt) val (thm1, basis2, _, trimmed_thm) = trim_nz Pos_Trim ectxt e1 basis1 val (ln_thm, basis3) = ln_expansion ectxt trimmed_thm thm1 basis2 val thm2' = lift basis3 thm2 |> simplify_expansion ectxt val mult_thm = @{thm expands_to_mult} OF [get_basis_wf_thm basis3, ln_thm, thm2'] val thm = @{thm expands_to_ln_powr} OF [trimmed_thm, get_basis_wf_thm basis2, thm1, mult_thm] in (thm, basis3) end | expand'' ectxt (ExpLn e) basis = let val (thm, basis', _, trimmed_thm) = trim_nz Pos_Trim ectxt e basis val thm = @{thm expands_to_exp_ln} OF [trimmed_thm, get_basis_wf_thm basis', thm] in (thm, basis') end | expand'' ectxt (Power (e, n)) basis = let val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt) in (power_expansion ectxt (thm, n, basis'), basis') end | expand'' ectxt (Root (e, n)) basis = let val (thm, basis') = expand' ectxt e basis |> apfst (simplify_expansion ectxt) in (root_expansion ectxt (thm, n, basis'), basis') end | expand'' ectxt (Inverse e) basis = (case trim_nz Simple_Trim ectxt e basis of (thm, basis', _, trimmed_thm) => (@{thm expands_to_inverse} OF [trimmed_thm, get_basis_wf_thm basis', thm], basis')) | expand'' ectxt (Div (e1, e2)) basis = let val (thm1, basis') = expand' ectxt e1 basis val (thm2, basis'', _, trimmed_thm) = trim_nz Simple_Trim ectxt e2 basis' val thm1 = lift basis'' thm1 in (@{thm expands_to_divide} OF [trimmed_thm, get_basis_wf_thm basis'', thm1, thm2], basis'') end | expand'' ectxt (Ln e) basis = let val (thm, basis', _, trimmed_thm) = trim_nz Pos_Trim ectxt e basis in ln_expansion ectxt trimmed_thm thm basis' end | expand'' ectxt (Exp e) basis = let val (thm, basis') = expand' ectxt e basis in exp_expansion ectxt thm basis' end | expand'' ectxt (Absolute e) basis = let val (thm, basis', nz, trimmed_thm) = trim_nz Sgn_Trim ectxt e basis val thm' = case nz of IsPos => @{thm expands_to_abs_pos} | IsNeg => @{thm expands_to_abs_neg} | _ => raise TERM ("Unexpected trim result during expansion of abs", []) in (thm' OF [trimmed_thm, get_basis_wf_thm basis', thm], basis') end | expand'' ectxt (Sgn e) basis = let val (thm, basis') = expand' ectxt e basis in (sgn_expansion ectxt (thm, basis'), basis') end | expand'' ectxt (Min (e1, e2)) basis = ( case try_prove_ev_eq ectxt (apply2 expr_to_term (e1, e2)) of SOME eq_thm => expand' ectxt e1 basis |> apfst (fn thm => @{thm expands_to_min_eq} OF [thm, eq_thm]) | NONE => let val (thm1, basis') = expand' ectxt e1 basis val (thm2, basis'') = expand' ectxt e2 basis' val thm1' = lift basis'' thm1 in (min_expansion ectxt (thm1', thm2, basis''), basis'') end) | expand'' ectxt (Max (e1, e2)) basis = ( case try_prove_ev_eq ectxt (apply2 expr_to_term (e1, e2)) of SOME eq_thm => expand' ectxt e1 basis |> apfst (fn thm => @{thm expands_to_max_eq} OF [thm, eq_thm]) | NONE => let val (thm1, basis') = expand' ectxt e1 basis val (thm2, basis'') = expand' ectxt e2 basis' val thm1' = lift basis'' thm1 in (max_expansion ectxt (thm1', thm2, basis''), basis'') end) | expand'' ectxt (Sin e) basis = let val (thm, basis', _, _) = trim_nz Simple_Trim ectxt e basis (* TODO could be relaxed *) in (sin_cos_expansion ectxt thm basis' |> fst, basis') end | expand'' ectxt (Cos e) basis = let val (thm, basis', _, _) = trim_nz Simple_Trim ectxt e basis (* TODO could be relaxed *) in (sin_cos_expansion ectxt thm basis' |> snd, basis') end | expand'' _ (Floor _) _ = raise TERM ("floor not supported.", []) | expand'' _ (Ceiling _) _ = raise TERM ("ceiling not supported.", []) | expand'' _ (Frac _) _ = raise TERM ("frac not supported.", []) | expand'' _ (NatMod _) _ = raise TERM ("mod not supported.", []) | expand'' ectxt (ArcTan e) basis = let (* TODO: what if it's zero *) val (thm, basis') = expand' ectxt e basis in (arctan_expansion ectxt basis' thm, basis') end | expand'' ectxt (Custom (name, t, args)) basis = let fun expand_args acc basis [] = (rev acc, basis) | expand_args acc basis (arg :: args) = case expand' ectxt arg basis of (thm, basis') => expand_args (thm :: acc) basis' args in case expand_custom (get_ctxt ectxt) name of NONE => raise TERM ("Unsupported custom function: " ^ name, [t]) | SOME e => e ectxt t (expand_args [] basis args) end and expand' ectxt (e' as (Inverse e)) basis = let val t = expr_to_term e fun thm wf_thm len_thm = @{thm expands_to_basic_inverse} OF [wf_thm, len_thm] in if member abconv (get_basis_list basis) t then (expand_basic thm t basis, basis) |> apfst (check_expansion e') else expand'' ectxt e' basis |> apfst (check_expansion e') end | expand' ectxt (Div (e1, e2)) basis = let val (thm1, basis') = expand' ectxt e1 basis val t = expr_to_term e2 fun thm wf_thm len_thm = @{thm expands_to_basic_inverse} OF [wf_thm, len_thm] in if member abconv (get_basis_list basis') t then (@{thm expands_to_div'} OF [get_basis_wf_thm basis', thm1, expand_basic thm t basis'], basis') else let val (thm2, basis'', _, trimmed_thm) = trim_nz Simple_Trim ectxt e2 basis' val thm1 = lift basis'' thm1 in (@{thm expands_to_divide} OF [trimmed_thm, get_basis_wf_thm basis'', thm1, thm2], basis'') end end | expand' ectxt (e' as (Powr' (e, p))) basis = let val t = expr_to_term e val ctxt = get_ctxt ectxt fun thm wf_thm len_thm = (Drule.infer_instantiate' ctxt [NONE, NONE, SOME (Thm.cterm_of ctxt p)] @{thm expands_to_basic_powr}) OF [wf_thm, len_thm] in if member abconv (get_basis_list basis) t then (expand_basic thm t basis, basis) |> apfst (check_expansion e') else expand'' ectxt e' basis |> apfst (check_expansion e') end | expand' ectxt e basis = let val t = expr_to_term e fun thm wf_thm len_thm = @{thm expands_to_basic} OF [wf_thm, len_thm] in if member abconv (get_basis_list basis) t then (expand_basic thm t basis, basis) |> apfst (check_expansion e) else expand'' ectxt e basis |> apfst (check_expansion e) end fun expand ectxt e basis = expand' ectxt e basis |> apfst (simplify_expansion ectxt) |> apfst (check_expansion e) fun expand_term ectxt t basis = let val ctxt = get_ctxt ectxt val (e, eq_thm) = reify ctxt t val (thm, basis) = expand ectxt e basis in (@{thm expands_to_meta_eq_cong'} OF [thm, eq_thm], basis) end fun expand_terms ectxt ts basis = let val ctxt = get_ctxt ectxt val e_eq_thms = map (reify ctxt) ts fun step (e, eq_thm) (thms, basis) = let val (thm, basis) = expand' ectxt e basis val thm = @{thm expands_to_meta_eq_cong'} OF [simplify_expansion ectxt thm, eq_thm] in (thm :: thms, basis) end val (thms, basis) = fold step e_eq_thms ([], basis) fun lift thm = lift_expands_to_thm (mk_lifting (extract_basis_list thm) basis) thm in (map lift (rev thms), basis) end datatype limit = Zero_Limit of bool option | Finite_Limit of term | Infinite_Limit of bool option fun is_empty_expansion (Const (\<^const_name>\MS\, _) $ Const (\<^const_name>\MSLNil\, _) $ _) = true | is_empty_expansion _ = false fun limit_of_expansion_aux ectxt basis thm = let val n = length (get_basis_list basis) val (thm, res, e_thms) = trim_expansion_while_greater false (SOME (replicate n \<^term>\0::real\)) true (SOME Simple_Trim) ectxt (thm, basis) val trimmed_thm = case res of Trimmed (_, trimmed_thm) => trimmed_thm | _ => NONE val res = case res of Trimmed _ => GREATER | Aborted res => res val exp = get_expansion thm val _ = if res = GREATER andalso is_none trimmed_thm andalso not (is_empty_expansion exp) then raise TERM ("limit_of_expansion", [get_expansion thm]) else () fun go thm _ _ [] = ( case zeroness_oracle false (SOME Simple_Trim) ectxt (get_expansion thm) of (IsZero, _) => (Zero_Limit NONE, @{thm expands_to_real_imp_filterlim} OF [thm]) | _ => (Finite_Limit \<^term>\0::real\, @{thm expands_to_real_imp_filterlim} OF [thm])) | go thm _ basis ((IsNeg, neg_thm) :: _) = (Zero_Limit NONE, @{thm expands_to_neg_exponent_imp_filterlim} OF [thm, get_basis_wf_thm basis, neg_thm RS @{thm compare_reals_diff_sgnD(1)}]) | go thm trimmed_thm basis ((IsPos, pos_thm) :: _) = (Infinite_Limit NONE, @{thm expands_to_pos_exponent_imp_filterlim} OF [thm, the trimmed_thm, get_basis_wf_thm basis, pos_thm RS @{thm compare_reals_diff_sgnD(3)}]) | go thm trimmed_thm basis ((IsZero, zero_thm) :: e_thms) = let val thm' = thm RS @{thm expands_to_hd''} val trimmed_thm' = Option.map (fn thm => thm RS @{thm trimmed_hd}) trimmed_thm val (lim, lim_thm) = go thm' trimmed_thm' (tl_basis basis) e_thms val lim_lift_thm = case lim of Infinite_Limit _ => @{thm expands_to_zero_exponent_imp_filterlim(1)} | _ => @{thm expands_to_zero_exponent_imp_filterlim(2)} val lim_thm' = lim_lift_thm OF [thm, get_basis_wf_thm basis, zero_thm RS @{thm compare_reals_diff_sgnD(2)}, lim_thm] in (lim, lim_thm') end | go _ _ _ _ = raise Match in if is_empty_expansion exp then (Zero_Limit NONE, thm RS @{thm expands_to_MSLNil_imp_filterlim}, thm) else case go thm trimmed_thm basis e_thms of (lim, lim_thm) => (lim, lim_thm, thm) end (* Determines the limit of a function from its expansion. The two flags control whether the the sign of the approach should be determined for the infinite case (i.e. at_top/at_bot instead of just at_infinity) and the zero case (i.e. at_right 0/at_left 0 instead of just nhds 0) *) fun limit_of_expansion (sgn_zero, sgn_inf) ectxt (thm, basis) = let val (lim, lim_thm, thm) = limit_of_expansion_aux ectxt basis thm in case lim of Zero_Limit _ => ( if sgn_zero then case trim_expansion false (SOME Sgn_Trim) ectxt (thm, basis) of (thm, IsPos, SOME pos_thm) => (Zero_Limit (SOME true), @{thm tendsto_imp_filterlim_at_right[OF _ expands_to_imp_eventually_pos]} OF [lim_thm, get_basis_wf_thm basis, thm, pos_thm]) | (thm, IsNeg, SOME neg_thm) => (Zero_Limit (SOME false), @{thm tendsto_imp_filterlim_at_left[OF _ expands_to_imp_eventually_neg]} OF [lim_thm, get_basis_wf_thm basis, thm, neg_thm]) | _ => (Zero_Limit NONE, lim_thm) else (Zero_Limit NONE, lim_thm)) | Infinite_Limit _ => ( if sgn_inf then case trim_expansion false (SOME Sgn_Trim) ectxt (thm, basis) of (thm, IsPos, SOME pos_thm) => (Infinite_Limit (SOME true), (@{thm filterlim_at_infinity_imp_filterlim_at_top[OF _ expands_to_imp_eventually_pos]} OF [lim_thm, get_basis_wf_thm basis, thm, pos_thm])) | (thm, IsNeg, SOME neg_thm) => (Infinite_Limit (SOME false), @{thm filterlim_at_infinity_imp_filterlim_at_bot[OF _ expands_to_imp_eventually_neg]} OF [lim_thm, get_basis_wf_thm basis, thm, neg_thm]) | _ => (Infinite_Limit NONE, lim_thm) else (Infinite_Limit NONE, lim_thm)) | Finite_Limit c => (Finite_Limit c, lim_thm) end fun compute_limit ectxt t = case expand_term ectxt t default_basis of (thm, basis) => limit_of_expansion (true, true) ectxt (thm, basis) fun prove_at_infinity ectxt (thm, basis) = let fun err () = raise TERM ("prove_at_infinity", [get_expanded_fun thm]) val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) fun go basis thm trimmed_thm = if fastype_of (get_expansion thm) = \<^typ>\real\ then err () else case zeroness_oracle true (SOME Pos_Trim) ectxt (get_exponent (get_expansion thm)) of (IsPos, SOME pos_thm) => @{thm expands_to_pos_exponent_imp_filterlim} OF [thm, trimmed_thm, get_basis_wf_thm basis, pos_thm] | (IsZero, SOME zero_thm) => @{thm expands_to_zero_exponent_imp_filterlim(1)} OF [thm, get_basis_wf_thm basis, zero_thm, go (tl_basis basis) (thm RS @{thm expands_to_hd''}) (trimmed_thm RS @{thm trimmed_hd})] | _ => err () in go basis thm trimmed_thm end fun prove_at_top_at_bot mode ectxt (thm, basis) = let val s = if mode = Pos_Trim then "prove_at_top" else "prove_at_bot" fun err () = raise TERM (s, [get_expanded_fun thm]) val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis) val trimmed_thm' = trimmed_thm RS (if mode = Pos_Trim then @{thm trimmed_pos_imp_trimmed} else @{thm trimmed_neg_imp_trimmed}) fun go basis thm trimmed_thm = if fastype_of (get_expansion thm) = \<^typ>\real\ then err () else case zeroness_oracle true (SOME Pos_Trim) ectxt (get_exponent (get_expansion thm)) of (IsPos, SOME pos_thm) => @{thm expands_to_pos_exponent_imp_filterlim} OF [thm, trimmed_thm, get_basis_wf_thm basis, pos_thm] | (IsZero, SOME zero_thm) => @{thm expands_to_zero_exponent_imp_filterlim(1)} OF [thm, get_basis_wf_thm basis, zero_thm, go (tl_basis basis) (thm RS @{thm expands_to_hd''}) (trimmed_thm RS @{thm trimmed_hd})] | _ => err () val lim_thm = go basis thm trimmed_thm' val add_sign_thm = if mode = Pos_Trim then @{thm filterlim_at_infinity_imp_filterlim_at_top[OF _ expands_to_imp_eventually_pos]} else @{thm filterlim_at_infinity_imp_filterlim_at_bot[OF _ expands_to_imp_eventually_neg]} in add_sign_thm OF [lim_thm, get_basis_wf_thm basis, thm, trimmed_thm] end val prove_at_top = prove_at_top_at_bot Pos_Trim val prove_at_bot = prove_at_top_at_bot Neg_Trim fun prove_at_aux mode ectxt (thm, basis) = let val (s, add_sign_thm) = case mode of Simple_Trim => ("prove_at_0", @{thm Topological_Spaces.filterlim_atI[OF _ expands_to_imp_eventually_nz]}) | Pos_Trim => ("prove_at_right_0", @{thm tendsto_imp_filterlim_at_right[OF _ expands_to_imp_eventually_pos]}) | Neg_Trim => ("prove_at_left_0", @{thm tendsto_imp_filterlim_at_left[OF _ expands_to_imp_eventually_neg]}) fun err () = raise TERM (s, [get_expanded_fun thm]) val (thm, _, SOME trimmed_thm) = trim_expansion true (SOME mode) ectxt (thm, basis) fun go basis thm = if fastype_of (get_expansion thm) = \<^typ>\real\ then err () else case zeroness_oracle true (SOME Neg_Trim) ectxt (get_exponent (get_expansion thm)) of (IsNeg, SOME neg_thm) => @{thm expands_to_neg_exponent_imp_filterlim} OF [thm, get_basis_wf_thm basis, neg_thm] | (IsZero, SOME zero_thm) => @{thm expands_to_zero_exponent_imp_filterlim(2)} OF [thm, get_basis_wf_thm basis, zero_thm, go (tl_basis basis) (thm RS @{thm expands_to_hd''})] | _ => err () val lim_thm = go basis thm in add_sign_thm OF [lim_thm, get_basis_wf_thm basis, thm, trimmed_thm] end val prove_at_0 = prove_at_aux Simple_Trim val prove_at_left_0 = prove_at_aux Neg_Trim val prove_at_right_0 = prove_at_aux Pos_Trim fun prove_nhds ectxt (thm, basis) = let fun simplify (a, b, c) = (a, simplify_expansion ectxt b, c) fun go thm basis = if fastype_of (get_expansion thm) = \<^typ>\real\ then @{thm expands_to_real_imp_filterlim} OF [thm] else case whnf_expansion ectxt thm |> simplify of (NONE, thm, _) => @{thm expands_to_MSLNil_imp_filterlim} OF [thm] | (SOME _, thm, _) => ( case zeroness_oracle true (SOME Sgn_Trim) ectxt (get_exponent (get_expansion thm)) of (IsZero, SOME zero_thm) => @{thm expands_to_zero_exponent_imp_filterlim(2)} OF [thm, get_basis_wf_thm basis, zero_thm, go (thm RS @{thm expands_to_hd''}) (tl_basis basis)] | (IsNeg, SOME neg_thm) => @{thm expands_to_neg_exponent_imp_filterlim} OF [thm, get_basis_wf_thm basis, neg_thm] | (IsPos, _) => go (try_drop_leading_term ectxt thm) basis | _ => raise TERM ("Unexpected zeroness result in prove_nhds", [get_exponent (get_expansion thm)])) in go thm basis end fun prove_equivalent theta ectxt (thm1, thm2, basis) = let val ((thm1, _, SOME trimmed_thm1), (thm2, _, SOME trimmed_thm2)) = apply2 (trim_expansion true (SOME Simple_Trim) ectxt) ((thm1, basis), (thm2, basis)) val pat = ConsPat (\<^const_name>\Pair\, [ConsPat (\<^const_name>\Lazy_Eval.cmp_result.EQ\, []), ConsPat (\<^const_name>\Pair\, [AnyPat ("_", 0), AnyPat ("_", 0)])]) val (exp1, exp2) = apply2 get_expansion (thm1, thm2) val T = fastype_of exp1 val t = mk_compare_expansions_const T $ exp1 $ exp2 fun eq_thm conv = HOLogic.mk_obj_eq (conv (Thm.cterm_of (get_ctxt ectxt) t)) val imp_thm = if theta then @{thm compare_expansions_EQ_imp_bigtheta} else @{thm compare_expansions_EQ_same} in case match ectxt pat t (SOME []) of (SOME _, t, conv) => let val [_, c1, c2] = HOLogic.strip_tuple t val c12_thm = if theta then [] else [the (try_prove_real_eq true ectxt (c1, c2))] in imp_thm OF ([eq_thm conv, trimmed_thm1, trimmed_thm2, thm1, thm2, get_basis_wf_thm basis] @ c12_thm) end | _ => raise TERM ("prove_equivalent", map get_expanded_fun [thm1, thm2]) end val prove_bigtheta = prove_equivalent true val prove_asymp_equiv = prove_equivalent false fun print_trimming_error s ectxt exp = let val c = get_coeff exp val t = if fastype_of c = \<^typ>\real\ then c else get_eval c in if #verbose (#ctxt ectxt) then let val ctxt = get_ctxt ectxt val p = Pretty.str "real_asymp failed to show zeroness of the following expression:" val p = Pretty.chunks [p, Pretty.indent 2 (Syntax.pretty_term ctxt t)] val _ = Pretty.writeln p in raise TERM (s, [t]) end else raise TERM (s, [t]) end fun prove_smallo ectxt (thm1, thm2, basis) = let val (thm2, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis) val es = get_exponents (get_expansion thm2) in case trim_expansion_while_greater true (SOME es) false NONE ectxt (thm1, basis) of (thm1, Aborted LESS, thms) => @{thm compare_expansions_LT} OF [prove_compare_expansions LESS (map snd thms), trimmed_thm, thm1, thm2, get_basis_wf_thm basis] | (thm1, Aborted _, _) => print_trimming_error "prove_smallo" ectxt (get_expansion thm1) | (thm1, Trimmed _, _) => print_trimming_error "prove_smallo" ectxt (get_expansion thm1) end fun prove_bigo ectxt (thm1, thm2, basis) = let val (thm2, _, SOME trimmed_thm) = trim_expansion true (SOME Simple_Trim) ectxt (thm2, basis) val es = get_exponents (get_expansion thm2) in case trim_expansion_while_greater false (SOME es) false NONE ectxt (thm1, basis) of (thm1, Aborted LESS, thms) => @{thm landau_o.small_imp_big[OF compare_expansions_LT]} OF [prove_compare_expansions LESS (map snd thms), trimmed_thm, thm1, thm2, get_basis_wf_thm basis] | (thm1, Aborted EQ, thms) => @{thm compare_expansions_EQ_imp_bigo} OF [prove_compare_expansions EQ (map snd thms), trimmed_thm, thm1, thm2, get_basis_wf_thm basis] | (thm1, Trimmed _, _) => print_trimming_error "prove_bigo" ectxt (get_expansion thm1) end fun prove_asymptotic_relation_aux mode f ectxt (thm1, thm2, basis) = f ( let val thm = @{thm expands_to_minus} OF [get_basis_wf_thm basis, thm1, thm2] in case ev_zeroness_oracle ectxt (get_expanded_fun thm) of SOME zero_thm => (EQUAL, zero_thm RS @{thm eventually_diff_zero_imp_eq}) | _ => ( case trim_expansion true (SOME mode) ectxt (thm, basis) of (thm, IsPos, SOME pos_thm) => (GREATER, @{thm expands_to_imp_eventually_gt} OF [get_basis_wf_thm basis, thm, pos_thm]) | (thm, IsNeg, SOME neg_thm) => (LESS, @{thm expands_to_imp_eventually_lt} OF [get_basis_wf_thm basis, thm, neg_thm]) | _ => raise TERM ("Unexpected zeroness result in prove_asymptotic_relation", [])) end) val prove_eventually_greater = prove_asymptotic_relation_aux Pos_Trim snd val prove_eventually_less = prove_asymptotic_relation_aux Neg_Trim snd val prove_asymptotic_relation = prove_asymptotic_relation_aux Sgn_Trim I fun prove_eventually_nonzero ectxt (thm, basis) = case trim_expansion true (SOME Simple_Trim) ectxt (thm, basis) of (thm, _, SOME trimmed_thm) => @{thm expands_to_imp_eventually_nz} OF [get_basis_wf_thm basis, thm, trimmed_thm] | _ => raise TERM ("prove_eventually_nonzero", [get_expanded_fun thm]) fun extract_terms (n, strict) ectxt basis t = let val bs = get_basis_list basis fun mk_constfun c = (Abs ("x", \<^typ>\real\, c)) val const_0 = mk_constfun \<^term>\0 :: real\ val const_1 = mk_constfun \<^term>\1 :: real\ fun uminus t = Term.betapply (\<^term>\\(f::real\real) x. -f x\, t) fun betapply2 a b c = Term.betapply (Term.betapply (a, b), c) fun mk_sum' [] acc = acc | mk_sum' ((t, sgn) :: ts) acc = mk_sum' ts ( if sgn then betapply2 \<^term>\%(f::real=>real) g x. f x - g x\ acc t else betapply2 \<^term>\%(f::real=>real) g x. f x + g x\ acc t) fun mk_sum [] = const_0 | mk_sum ((t, sgn) :: ts) = mk_sum' ts (if sgn then uminus t else t) fun mk_mult a b = if a aconv const_0 then const_0 else if b aconv const_0 then const_0 else if a aconv \<^term>\\_::real. 1 :: real\ then b else if b aconv \<^term>\\_::real. 1 :: real\ then a else if a aconv \<^term>\\_::real. -1 :: real\ then Term.betapply (\<^term>\\(f::real\real) x. -f x\, b) else if b aconv \<^term>\\_::real. -1 :: real\ then Term.betapply (\<^term>\\(f::real\real) x. -f x\, a) else Abs ("x", \<^typ>\real\, \<^term>\(*) :: real => _\ $ (Term.betapply (a, Bound 0)) $ (Term.betapply (b, Bound 0))) fun mk_powr b e = if e = \<^term>\0 :: real\ then const_1 else let val n = HOLogic.dest_number e |> snd in if n >= 0 then Term.betapply (Term.betapply (\<^term>\%(b::real=>real) e x. b x ^ e\, b), HOLogic.mk_number \<^typ>\nat\ n) else Term.betapply (Term.betapply (\<^term>\%(b::real=>real) e x. b x powr e\, b), e) end handle TERM _ => Term.betapply (Term.betapply (\<^term>\%(b::real=>real) e x. b x powr e\, b), e) fun mk_scale_elem b e acc = let val e = simplify_term ectxt e in if e = \<^term>\0 :: real\ then acc else if e = \<^term>\1 :: real\ then mk_mult acc b else mk_mult acc (mk_powr b e) end fun mk_scale_elems [] _ acc = acc | mk_scale_elems (b :: bs) (e :: es) acc = mk_scale_elems bs es (mk_scale_elem b e acc) | mk_scale_elems _ _ _ = raise Match fun mk_summand c es = let val es = mk_scale_elems bs es \<^term>\\_::real. 1 :: real\ in case c of Const (\<^const_name>\uminus\, _) $ c => ((c, true), es) | _ => ((c, false), es) end fun go _ _ _ acc 0 = (acc, 0) | go 0 es t acc n = let val c = simplify_term ectxt t in if strict andalso c = \<^term>\0 :: real\ then (acc, n) else (mk_summand c (rev es) :: acc, n - 1) end | go m es t acc n = case Lazy_Eval.whnf ectxt t |> fst of Const (\<^const_name>\MS\, _) $ cs $ _ => go' m es (simplify_term ectxt cs) acc n | _ => raise TERM("extract_terms", [t]) and go' _ _ _ acc 0 = (acc, 0) | go' m es cs acc n = case Lazy_Eval.whnf ectxt cs |> fst of Const (\<^const_name>\MSLNil\, _) => (acc, n) | Const (\<^const_name>\MSLCons\, _) $ c $ cs => ( case Lazy_Eval.whnf ectxt c |> fst |> HOLogic.dest_prod of (c, e) => case go (m - 1) (e :: es) c acc n of (acc, n) => go' m es (simplify_term ectxt cs) acc n) | _ => raise TERM("extract_terms", [t]) val (summands, remaining) = go (basis_size basis) [] t [] (n + 1) val (summands, error) = if remaining = 0 then (rev (tl summands), SOME (snd (hd summands))) else (rev summands, NONE) val summands = map (fn ((c, sgn), es) => (mk_mult (mk_constfun c) es, sgn)) summands val error = Option.map (fn err => Term.betapply (\<^term>\\f::real\real. O(f)\, err)) error val expansion = mk_sum summands in (expansion, error) end end structure Multiseries_Expansion_Basic : EXPANSION_INTERFACE = struct open Multiseries_Expansion; type T = expansion_thm val expand_term = expand_term val expand_terms = expand_terms val prove_nhds = prove_nhds val prove_at_infinity = prove_at_infinity val prove_at_top = prove_at_top val prove_at_bot = prove_at_bot val prove_at_0 = prove_at_0 val prove_at_right_0 = prove_at_right_0 val prove_at_left_0 = prove_at_left_0 val prove_eventually_nonzero = prove_eventually_nonzero val prove_eventually_less = prove_eventually_less val prove_eventually_greater = prove_eventually_greater val prove_smallo = prove_smallo val prove_bigo = prove_bigo val prove_bigtheta = prove_bigtheta val prove_asymp_equiv = prove_asymp_equiv end diff --git a/src/HOL/Statespace/distinct_tree_prover.ML b/src/HOL/Statespace/distinct_tree_prover.ML --- a/src/HOL/Statespace/distinct_tree_prover.ML +++ b/src/HOL/Statespace/distinct_tree_prover.ML @@ -1,367 +1,367 @@ (* Title: HOL/Statespace/distinct_tree_prover.ML Author: Norbert Schirmer, TU Muenchen *) signature DISTINCT_TREE_PROVER = sig datatype direction = Left | Right val mk_tree : ('a -> term) -> typ -> 'a list -> term val dest_tree : term -> term list val find_tree : term -> term -> direction list option val neq_to_eq_False : thm val distinctTreeProver : Proof.context -> thm -> direction list -> direction list -> thm val neq_x_y : Proof.context -> term -> term -> string -> thm option val distinctFieldSolver : string list -> solver val distinctTree_tac : string list -> Proof.context -> int -> tactic val distinct_implProver : Proof.context -> thm -> cterm -> thm val subtractProver : Proof.context -> term -> cterm -> thm -> thm val distinct_simproc : string list -> simproc val discharge : Proof.context -> thm list -> thm -> thm end; structure DistinctTreeProver : DISTINCT_TREE_PROVER = struct val neq_to_eq_False = @{thm neq_to_eq_False}; datatype direction = Left | Right; fun treeT T = Type (\<^type_name>\tree\, [T]); fun mk_tree' e T n [] = Const (\<^const_name>\Tip\, treeT T) | mk_tree' e T n xs = let val m = (n - 1) div 2; val (xsl,x::xsr) = chop m xs; val l = mk_tree' e T m xsl; val r = mk_tree' e T (n-(m+1)) xsr; in Const (\<^const_name>\Node\, treeT T --> T --> HOLogic.boolT--> treeT T --> treeT T) $ l $ e x $ \<^term>\False\ $ r end fun mk_tree e T xs = mk_tree' e T (length xs) xs; fun dest_tree (Const (\<^const_name>\Tip\, _)) = [] | dest_tree (Const (\<^const_name>\Node\, _) $ l $ e $ _ $ r) = dest_tree l @ e :: dest_tree r | dest_tree t = raise TERM ("dest_tree", [t]); fun lin_find_tree e (Const (\<^const_name>\Tip\, _)) = NONE | lin_find_tree e (Const (\<^const_name>\Node\, _) $ l $ x $ _ $ r) = if e aconv x then SOME [] else (case lin_find_tree e l of SOME path => SOME (Left :: path) | NONE => (case lin_find_tree e r of SOME path => SOME (Right :: path) | NONE => NONE)) | lin_find_tree e t = raise TERM ("find_tree: input not a tree", [t]) fun bin_find_tree order e (Const (\<^const_name>\Tip\, _)) = NONE | bin_find_tree order e (Const (\<^const_name>\Node\, _) $ l $ x $ _ $ r) = (case order (e, x) of EQUAL => SOME [] | LESS => Option.map (cons Left) (bin_find_tree order e l) | GREATER => Option.map (cons Right) (bin_find_tree order e r)) | bin_find_tree order e t = raise TERM ("find_tree: input not a tree", [t]) fun find_tree e t = (case bin_find_tree Term_Ord.fast_term_ord e t of NONE => lin_find_tree e t | x => x); fun split_common_prefix xs [] = ([], xs, []) | split_common_prefix [] ys = ([], [], ys) | split_common_prefix (xs as (x :: xs')) (ys as (y :: ys')) = if x = y then let val (ps, xs'', ys'') = split_common_prefix xs' ys' in (x :: ps, xs'', ys'') end else ([], xs, ys) (* Wrapper around Thm.instantiate. The type instiations of instTs are applied to * the right hand sides of insts *) fun instantiate ctxt instTs insts = let val instTs' = map (fn (T, U) => (dest_TVar (Thm.typ_of T), Thm.typ_of U)) instTs; fun substT x = (case AList.lookup (op =) instTs' x of NONE => TVar x | SOME T' => T'); fun mapT_and_recertify ct = (Thm.cterm_of ctxt (Term.map_types (Term.map_type_tvar substT) (Thm.term_of ct))); val insts' = map (apfst mapT_and_recertify) insts; in Thm.instantiate - (map (apfst (dest_TVar o Thm.typ_of)) instTs, - map (apfst (dest_Var o Thm.term_of)) insts') + (TVars.make (map (apfst (dest_TVar o Thm.typ_of)) instTs), + Vars.make (map (apfst (dest_Var o Thm.term_of)) insts')) end; fun tvar_clash ixn S S' = raise TYPE ("Type variable has two distinct sorts", [TVar (ixn, S), TVar (ixn, S')], []); fun lookup (tye, (ixn, S)) = (case AList.lookup (op =) tye ixn of NONE => NONE | SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S'); val naive_typ_match = let fun match (TVar (v, S), T) subs = (case lookup (subs, (v, S)) of NONE => ((v, (S, T))::subs) | SOME _ => subs) | match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise Type.TYPE_MATCH else matches (Ts, Us) subs | match (TFree x, TFree y) subs = if x = y then subs else raise Type.TYPE_MATCH | match _ _ = raise Type.TYPE_MATCH and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs) | matches _ subs = subs; in match end; (* expects that relevant type variables are already contained in * term variables. First instantiation of variables is returned without further * checking. *) fun naive_cterm_first_order_match (t, ct) env = let fun mtch (env as (tyinsts, insts)) = fn (Var (ixn, T), ct) => (case AList.lookup (op =) insts ixn of NONE => (naive_typ_match (T, Thm.typ_of_cterm ct) tyinsts, (ixn, ct) :: insts) | SOME _ => env) | (f $ t, ct) => let val (cf, ct') = Thm.dest_comb ct; in mtch (mtch env (f, cf)) (t, ct') end | _ => env; in mtch env (t, ct) end; fun discharge ctxt prems rule = let val (tyinsts,insts) = fold naive_cterm_first_order_match (Thm.prems_of rule ~~ map Thm.cprop_of prems) ([], []); val tyinsts' = map (fn (v, (S, U)) => ((v, S), Thm.ctyp_of ctxt U)) tyinsts; val insts' = map (fn (idxn, ct) => ((idxn, Thm.typ_of_cterm ct), ct)) insts; - val rule' = Thm.instantiate (tyinsts', insts') rule; + val rule' = Thm.instantiate (TVars.make tyinsts', Vars.make insts') rule; in fold Thm.elim_implies prems rule' end; local val (l_in_set_root, x_in_set_root, r_in_set_root) = let val (Node_l_x_d, r) = Thm.cprop_of @{thm in_set_root} |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; val (Node_l, x) = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb; val l = Node_l |> Thm.dest_comb |> #2; in (l,x,r) end; val (x_in_set_left, r_in_set_left) = let val (Node_l_x_d, r) = Thm.cprop_of @{thm in_set_left} |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; val x = Node_l_x_d |> Thm.dest_comb |> #1 |> Thm.dest_comb |> #2; in (x, r) end; val (x_in_set_right, l_in_set_right) = let val (Node_l, x) = Thm.cprop_of @{thm in_set_right} |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #1 |> Thm.dest_comb |> #1 |> Thm.dest_comb; val l = Node_l |> Thm.dest_comb |> #2; in (x, l) end; in (* 1. First get paths x_path y_path of x and y in the tree. 2. For the common prefix descend into the tree according to the path and lemmas all_distinct_left/right 3. If one restpath is empty use distinct_left/right, otherwise all_distinct_left_right *) fun distinctTreeProver ctxt dist_thm x_path y_path = let fun dist_subtree [] thm = thm | dist_subtree (p :: ps) thm = let val rule = (case p of Left => @{thm all_distinct_left} | Right => @{thm all_distinct_right}) in dist_subtree ps (discharge ctxt [thm] rule) end; val (ps, x_rest, y_rest) = split_common_prefix x_path y_path; val dist_subtree_thm = dist_subtree ps dist_thm; val subtree = Thm.cprop_of dist_subtree_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val (_, [l, _, _, r]) = Drule.strip_comb subtree; fun in_set ps tree = let val (_, [l, x, _, r]) = Drule.strip_comb tree; val xT = Thm.ctyp_of_cterm x; in (case ps of [] => instantiate ctxt [(Thm.ctyp_of_cterm x_in_set_root, xT)] [(l_in_set_root, l), (x_in_set_root, x), (r_in_set_root, r)] @{thm in_set_root} | Left :: ps' => let val in_set_l = in_set ps' l; val in_set_left' = instantiate ctxt [(Thm.ctyp_of_cterm x_in_set_left, xT)] [(x_in_set_left, x), (r_in_set_left, r)] @{thm in_set_left}; in discharge ctxt [in_set_l] in_set_left' end | Right :: ps' => let val in_set_r = in_set ps' r; val in_set_right' = instantiate ctxt [(Thm.ctyp_of_cterm x_in_set_right, xT)] [(x_in_set_right, x), (l_in_set_right, l)] @{thm in_set_right}; in discharge ctxt [in_set_r] in_set_right' end) end; fun in_set' [] = raise TERM ("distinctTreeProver", []) | in_set' (Left :: ps) = in_set ps l | in_set' (Right :: ps) = in_set ps r; fun distinct_lr node_in_set Left = discharge ctxt [dist_subtree_thm, node_in_set] @{thm distinct_left} | distinct_lr node_in_set Right = discharge ctxt [dist_subtree_thm, node_in_set] @{thm distinct_right} val (swap, neq) = (case x_rest of [] => let val y_in_set = in_set' y_rest; in (false, distinct_lr y_in_set (hd y_rest)) end | xr :: xrs => (case y_rest of [] => let val x_in_set = in_set' x_rest; in (true, distinct_lr x_in_set (hd x_rest)) end | yr :: yrs => let val x_in_set = in_set' x_rest; val y_in_set = in_set' y_rest; in (case xr of Left => (false, discharge ctxt [dist_subtree_thm, x_in_set, y_in_set] @{thm distinct_left_right}) | Right => (true, discharge ctxt [dist_subtree_thm, y_in_set, x_in_set] @{thm distinct_left_right})) end)); in if swap then discharge ctxt [neq] @{thm swap_neq} else neq end; fun deleteProver _ dist_thm [] = @{thm delete_root} OF [dist_thm] | deleteProver ctxt dist_thm (p::ps) = let val dist_rule = (case p of Left => @{thm all_distinct_left} | Right => @{thm all_distinct_right}); val dist_thm' = discharge ctxt [dist_thm] dist_rule; val del_rule = (case p of Left => @{thm delete_left} | Right => @{thm delete_right}); val del = deleteProver ctxt dist_thm' ps; in discharge ctxt [dist_thm, del] del_rule end; local val (alpha, v) = let val ct = @{thm subtract_Tip} |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val [alpha] = ct |> Thm.ctyp_of_cterm |> Thm.dest_ctyp; in (dest_TVar (Thm.typ_of alpha), #1 (dest_Var (Thm.term_of ct))) end; in fun subtractProver ctxt (Const (\<^const_name>\Tip\, T)) ct dist_thm = let val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val [alphaI] = #2 (dest_Type T); in Thm.instantiate - ([(alpha, Thm.ctyp_of ctxt alphaI)], - [((v, treeT alphaI), ct')]) @{thm subtract_Tip} + (TVars.make [(alpha, Thm.ctyp_of ctxt alphaI)], + Vars.make [((v, treeT alphaI), ct')]) @{thm subtract_Tip} end | subtractProver ctxt (Const (\<^const_name>\Node\, nT) $ l $ x $ d $ r) ct dist_thm = let val ct' = dist_thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val (_, [cl, _, _, cr]) = Drule.strip_comb ct; val ps = the (find_tree x (Thm.term_of ct')); val del_tree = deleteProver ctxt dist_thm ps; val dist_thm' = discharge ctxt [del_tree, dist_thm] @{thm delete_Some_all_distinct}; val sub_l = subtractProver ctxt (Thm.term_of cl) cl (dist_thm'); val sub_r = subtractProver ctxt (Thm.term_of cr) cr (discharge ctxt [sub_l, dist_thm'] @{thm subtract_Some_all_distinct_res}); in discharge ctxt [del_tree, sub_l, sub_r] @{thm subtract_Node} end; end; fun distinct_implProver ctxt dist_thm ct = let val ctree = ct |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val sub = subtractProver ctxt (Thm.term_of ctree) ctree dist_thm; in @{thm subtract_Some_all_distinct} OF [sub, dist_thm] end; fun get_fst_success f [] = NONE | get_fst_success f (x :: xs) = (case f x of NONE => get_fst_success f xs | SOME v => SOME v); fun neq_x_y ctxt x y name = (let val dist_thm = the (try (Proof_Context.get_thm ctxt) name); val ctree = Thm.cprop_of dist_thm |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2; val tree = Thm.term_of ctree; val x_path = the (find_tree x tree); val y_path = the (find_tree y tree); val thm = distinctTreeProver ctxt dist_thm x_path y_path; in SOME thm end handle Option.Option => NONE); fun distinctTree_tac names ctxt = SUBGOAL (fn (goal, i) => (case goal of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ x $ y)) => (case get_fst_success (neq_x_y ctxt x y) names of SOME neq => resolve_tac ctxt [neq] i | NONE => no_tac) | _ => no_tac)) fun distinctFieldSolver names = mk_solver "distinctFieldSolver" (distinctTree_tac names); fun distinct_simproc names = Simplifier.make_simproc \<^context> "DistinctTreeProver.distinct_simproc" {lhss = [\<^term>\x = y\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ x $ y => Option.map (fn neq => @{thm neq_to_eq_False} OF [neq]) (get_fst_success (neq_x_y ctxt x y) names) | _ => NONE)}; end; end; diff --git a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML --- a/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML +++ b/src/HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML @@ -1,784 +1,784 @@ (* Title: HOL/TPTP/TPTP_Parser/tptp_reconstruct_library.ML Author: Nik Sultana, Cambridge University Computer Laboratory Collection of general functions used in the reconstruction module. *) signature TPTP_RECONSTRUCT_LIBRARY = sig exception BREAK_LIST val break_list : 'a list -> 'a * 'a list val break_seq : 'a Seq.seq -> 'a * 'a Seq.seq exception MULTI_ELEMENT_LIST val cascaded_filter_single : bool -> ('a list -> 'a list) list -> 'a list -> 'a option val concat_between : 'a list list -> ('a option * 'a option) -> 'a list exception DIFF_TYPE of typ * typ exception DIFF of term * term val diff : theory -> term * term -> (term * term) list * (typ * typ) list exception DISPLACE_KV val displace_kv : ''a -> (''a * 'b) list -> (''a * 'b) list val enumerate : int -> 'a list -> (int * 'a) list val fold_options : 'a option list -> 'a list val find_and_remove : ('a -> bool) -> 'a list -> 'a * 'a list val lift_option : ('a -> 'b) -> 'a option -> 'b option val list_diff : ''a list -> ''a list -> ''a list val list_prod : 'a list list -> 'a list -> 'a list -> 'a list list val permute : ''a list -> ''a list list val prefix_intersection_list : ''a list -> ''a list -> ''a list val repeat_until_fixpoint : (''a -> ''a) -> ''a -> ''a val switch : ('a -> 'b -> 'c) -> 'b -> 'a -> 'c val zip_amap : 'a list -> 'b list -> ('a * 'b) list -> ('a * 'b) list * ('a list * 'b list) val consts_in : term -> term list val head_quantified_variable : Proof.context -> int -> thm -> (string * typ) option val push_allvar_in : string -> term -> term val strip_top_All_var : term -> (string * typ) * term val strip_top_All_vars : term -> (string * typ) list * term val strip_top_all_vars : (string * typ) list -> term -> (string * typ) list * term val trace_tac' : Proof.context -> string -> ('a -> thm -> 'b Seq.seq) -> 'a -> thm -> 'b Seq.seq val try_dest_Trueprop : term -> term val type_devar : typ TVars.table -> term -> term val diff_and_instantiate : Proof.context -> thm -> term -> term -> thm val batter_tac : Proof.context -> int -> tactic val break_hypotheses_tac : Proof.context -> int -> tactic val clause_breaker_tac : Proof.context -> int -> tactic (* val dist_all_and_tac : Proof.context -> int -> tactic *)(*FIXME unused*) val reassociate_conjs_tac : Proof.context -> int -> tactic val ASAP : (int -> tactic) -> (int -> tactic) -> int -> tactic val COND' : ('a -> thm -> bool) -> ('a -> tactic) -> ('a -> tactic) -> 'a -> tactic val TERMFUN : (term list * term -> 'a) -> int option -> thm -> 'a list val TERMPRED : (term -> bool) -> (term -> bool) -> int option -> thm -> bool val guided_abstract : bool -> term -> term -> ((string * typ) * term) * term list val abstract : term list -> term -> ((string * typ) * term) list * term end structure TPTP_Reconstruct_Library : TPTP_RECONSTRUCT_LIBRARY = struct (*zip as much as possible*) fun zip_amap [] ys acc = (acc, ([], ys)) | zip_amap xs [] acc = (acc, (xs, [])) | zip_amap (x :: xs) (y :: ys) acc = zip_amap xs ys ((x, y) :: acc); (*Pair a list up with the position number of each element, starting from n*) fun enumerate n ls = let fun enumerate' [] _ acc = acc | enumerate' (x :: xs) n acc = enumerate' xs (n + 1) ((n, x) :: acc) in enumerate' ls n [] |> rev end (* enumerate 0 []; enumerate 0 ["a", "b", "c"]; *) (*List subtraction*) fun list_diff l1 l2 = filter (fn x => forall (fn y => x <> y) l2) l1 val _ = \<^assert> (list_diff [1,2,3] [2,4] = [1, 3]) (* [a,b] times_list [c,d] gives [[a,c,d], [b,c,d]] *) fun list_prod acc [] _ = rev acc | list_prod acc (x :: xs) ys = list_prod ((x :: ys) :: acc) xs ys fun repeat_until_fixpoint f x = let val x' = f x in if x = x' then x else repeat_until_fixpoint f x' end (*compute all permutations of a list*) fun permute l = let fun permute' (l, []) = [(l, [])] | permute' (l, xs) = map (fn x => (x :: l, filter (fn y => y <> x) xs)) xs |> maps permute' in permute' ([], l) |> map fst end (* permute [1,2,3]; permute ["A", "B"] *) (*this exception is raised when the pair we wish to displace isn't found in the association list*) exception DISPLACE_KV; (*move a key-value pair, determined by the k, to the beginning of an association list. it moves the first occurrence of a pair keyed by "k"*) local fun fold_fun k (kv as (k', v)) (l, buff) = if is_some buff then (kv :: l, buff) else if k = k' then (l, SOME kv) else (kv :: l, buff) in (*"k" is the key value of the pair we wish to displace*) fun displace_kv k alist = let val (pre_alist, kv) = fold (fold_fun k) alist ([], NONE) in if is_some kv then the kv :: rev pre_alist else raise DISPLACE_KV end end (*Given two lists, it generates a new list where the intersection of the lists forms the prefix of the new list.*) local fun prefix_intersection_list' (acc_pre, acc_pro) l1 l2 = if null l1 then List.rev acc_pre @ List.rev acc_pro else if null l2 then List.rev acc_pre @ l1 @ List.rev acc_pro else let val l1_hd = hd l1 in prefix_intersection_list' (if member (op =) l2 l1_hd then (l1_hd :: acc_pre, acc_pro) else (acc_pre, l1_hd :: acc_pro)) (tl l1) l2 end in fun prefix_intersection_list l1 l2 = prefix_intersection_list' ([], []) l1 l2 end; val _ = \<^assert> (prefix_intersection_list [1,2,3,4,5] [1,3,5] = [1, 3, 5, 2, 4]); val _ = \<^assert> (prefix_intersection_list [1,2,3,4,5] [] = [1,2,3,4,5]); val _ = \<^assert> (prefix_intersection_list [] [1,3,5] = []) fun switch f y x = f x y (*Given a value of type "'a option list", produce a value of type "'a list" by dropping the NONE elements and projecting the SOME elements.*) fun fold_options opt_list = fold (fn x => fn l => if is_some x then the x :: l else l) opt_list []; val _ = \<^assert> ([2,0,1] = fold_options [NONE, SOME 1, NONE, SOME 0, NONE, NONE, SOME 2]); fun lift_option (f : 'a -> 'b) (x_opt : 'a option) : 'b option = case x_opt of NONE => NONE | SOME x => SOME (f x) fun break_seq x = (Seq.hd x, Seq.tl x) exception BREAK_LIST fun break_list (x :: xs) = (x, xs) | break_list _ = raise BREAK_LIST exception MULTI_ELEMENT_LIST (*Try a number of predicates, in order, to find a single element. Predicates are expected to either return an empty list or a singleton list. If strict=true and list has more than one element, then raise an exception. Otherwise try a new predicate.*) fun cascaded_filter_single strict preds l = case preds of [] => NONE | (p :: ps) => case p l of [] => cascaded_filter_single strict ps l | [x] => SOME x | l => if strict then raise MULTI_ELEMENT_LIST else cascaded_filter_single strict ps l (*concat but with optional before-and-after delimiters*) fun concat_between [] _ = [] | concat_between [l] _ = l | concat_between (l :: ls) (seps as (bef, aft)) = let val pre = if is_some bef then the bef :: l else l val mid = if is_some aft then [the aft] else [] val post = concat_between ls seps in pre @ mid @ post end (*Given a list, find an element satisfying pred, and return a pair consisting of that element and the list minus the element.*) fun find_and_remove pred l = find_index pred l |> switch chop l |> apsnd break_list |> (fn (xs, (y, ys)) => (y, xs @ ys)) val _ = \<^assert> (find_and_remove (curry (op =) 3) [0,1,2,3,4,5] = (3, [0,1,2,4,5])) (** Functions on terms **) (*Extract the forall-prefix of a term, and return a pair consisting of the prefix and the body*) local (*Strip off HOL's All combinator if it's at the toplevel*) fun try_dest_All (Const (\<^const_name>\HOL.All\, _) $ t) = t | try_dest_All (Const (\<^const_name>\HOL.Trueprop\, _) $ t) = try_dest_All t | try_dest_All t = t val _ = \<^assert> ((\<^term>\\x. (\y. P) = True\ |> try_dest_All |> Term.strip_abs_vars) = [("x", \<^typ>\'a\)]) val _ = \<^assert> ((\<^prop>\\x. (\y. P) = True\ |> try_dest_All |> Term.strip_abs_vars) = [("x", \<^typ>\'a\)]) fun strip_top_All_vars' once acc t = let val t' = try_dest_All t val var = try (Term.strip_abs_vars #> hd) t' fun strip v t = (v, subst_bounds ([Free v], Term.strip_abs_body t)) in if t' = t orelse is_none var then (acc, t) else let val (v, t) = strip (the var) t' val acc' = v :: acc in if once then (acc', t) else strip_top_All_vars' once acc' t end end in fun strip_top_All_vars t = strip_top_All_vars' false [] t val _ = let val answer = ([("x", \<^typ>\'a\)], HOLogic.all_const \<^typ>\'a\ $ (HOLogic.eq_const \<^typ>\'a\ $ Free ("x", \<^typ>\'a\))) in \<^assert> ((\<^term>\\x. All ((=) x)\ |> strip_top_All_vars) = answer) end (*like strip_top_All_vars, but peels a single variable off, instead of all of them*) fun strip_top_All_var t = strip_top_All_vars' true [] t |> apfst the_single end (*like strip_top_All_vars but for "Pure.all" instead of "HOL.All"*) fun strip_top_all_vars acc t = if Logic.is_all t then let val (v, t') = Logic.dest_all t (*bound instances in t' are replaced with free vars*) in strip_top_all_vars (v :: acc) t' end else (acc, (*variables are returned in FILO order*) t) (*given a term "t" ! X Y Z. t' then then "push_allvar_in "X" t" will give ! Y Z X. t' *) fun push_allvar_in v t = let val (vs, t') = strip_top_All_vars t val vs' = displace_kv v vs in fold (fn (v, ty) => fn t => HOLogic.mk_all (v, ty, t)) vs' t' end (*Lists all consts in a term, uniquely*) fun consts_in (Const c) = [Const c] | consts_in (Free _) = [] | consts_in (Var _) = [] | consts_in (Bound _) = [] | consts_in (Abs (_, _, t)) = consts_in t | consts_in (t1 $ t2) = union (op =) (consts_in t1) (consts_in t2); exception DIFF of term * term exception DIFF_TYPE of typ * typ (*This carries out naive form of matching. It "diffs" two formulas, to create a function which maps (schematic or non-schematic) variables to terms. The first argument is the more "general" term. The second argument is used to find the "image" for the variables in the first argument which don't appear in the second argument. Note that the list that is returned might have duplicate entries. It's not checked to see if the same variable maps to different values -- that should be regarded as an error.*) fun diff thy (initial as (t_gen, t)) = let fun diff_ty acc [] = acc | diff_ty acc ((pair as (ty_gen, ty)) :: ts) = case pair of (Type (s1, ty_gens1), Type (s2, ty_gens2)) => if s1 <> s2 orelse length ty_gens1 <> length ty_gens2 then raise (DIFF (t_gen, t)) else diff_ty acc (ts @ ListPair.zip (ty_gens1, ty_gens2)) | (TFree (s1, so1), TFree (s2, so2)) => if s1 <> s2 orelse not (Sign.subsort thy (so2, so1)) then raise (DIFF (t_gen, t)) else diff_ty acc ts | (TVar (idx1, so1), TVar (idx2, so2)) => if idx1 <> idx2 orelse not (Sign.subsort thy (so2, so1)) then raise (DIFF (t_gen, t)) else diff_ty acc ts | (TFree _, _) => diff_ty (pair :: acc) ts | (TVar _, _) => diff_ty (pair :: acc) ts | _ => raise (DIFF_TYPE pair) fun diff' (acc as (acc_t, acc_ty)) (pair as (t_gen, t)) ts = case pair of (Const (s1, ty1), Const (s2, ty2)) => if s1 <> s2 orelse not (Sign.typ_instance thy (ty2, ty1)) then raise (DIFF (t_gen, t)) else diff_probs acc ts | (Free (s1, ty1), Free (s2, ty2)) => if s1 <> s2 orelse not (Sign.typ_instance thy (ty2, ty1)) then raise (DIFF (t_gen, t)) else diff_probs acc ts | (Var (idx1, ty1), Var (idx2, ty2)) => if idx1 <> idx2 orelse not (Sign.typ_instance thy (ty2, ty1)) then raise (DIFF (t_gen, t)) else diff_probs acc ts | (Bound i1, Bound i2) => if i1 <> i2 then raise (DIFF (t_gen, t)) else diff_probs acc ts | (Abs (s1, ty1, t1), Abs (s2, ty2, t2)) => if s1 <> s2 orelse not (Sign.typ_instance thy (ty2, ty1)) then raise (DIFF (t_gen, t)) else diff' acc (t1, t2) ts | (ta1 $ ta2, tb1 $ tb2) => diff_probs acc ((ta1, tb1) :: (ta2, tb2) :: ts) (*the particularly important bit*) | (Free (_, ty), _) => diff_probs (pair :: acc_t, diff_ty acc_ty [(ty, Term.fastype_of t)]) ts | (Var (_, ty), _) => diff_probs (pair :: acc_t, diff_ty acc_ty [(ty, Term.fastype_of t)]) ts (*everything else is problematic*) | _ => raise (DIFF (t_gen, t)) and diff_probs acc ts = case ts of [] => acc | (pair :: ts') => diff' acc pair ts' in diff_probs ([], []) [initial] end (*Abstracts occurrences of "t_sub" in "t", returning a list of abstractions of "t" with a Var at each occurrence of "t_sub". If "strong=true" then it uses strong abstraction (i.e., replaces all occurrnces of "t_sub"), otherwise it uses weak abstraction (i.e., replaces the occurrences one at a time). NOTE there are many more possibilities between strong and week. These can be enumerated by abstracting based on the powerset of occurrences (minus the null element, which would correspond to "t"). *) fun guided_abstract strong t_sub t = let val varnames = Term.add_frees t [] |> map #1 val prefixK = "v" val freshvar = let fun find_fresh i = let val varname = prefixK ^ Int.toString i in if member (op =) varnames varname then find_fresh (i + 1) else (varname, fastype_of t_sub) end in find_fresh 0 end fun guided_abstract' t = case t of Abs (s, ty, t') => if t = t_sub then [Free freshvar] else (map (fn t' => Abs (s, ty, t')) (guided_abstract' t')) | t1 $ t2 => if t = t_sub then [Free freshvar] else (map (fn t' => t' $ t2) (guided_abstract' t1)) @ (map (fn t' => t1 $ t') (guided_abstract' t2)) | _ => if t = t_sub then [Free freshvar] else [t] fun guided_abstract_strong' t = let fun continue t = guided_abstract_strong' t |> (fn x => if null x then t else the_single x) in case t of Abs (s, ty, t') => if t = t_sub then [Free freshvar] else [Abs (s, ty, continue t')] | t1 $ t2 => if t = t_sub then [Free freshvar] else [continue t1 $ continue t2] | _ => if t = t_sub then [Free freshvar] else [t] end in ((freshvar, t_sub), if strong then guided_abstract_strong' t else guided_abstract' t) end (*Carries out strong abstraction of a term guided by a list of other terms. In case some of the latter terms happen to be the same, it only abstracts them once. It returns the abstracted term, together with a map from the fresh names to the terms.*) fun abstract ts t = fold_map (apsnd the_single oo (guided_abstract true)) ts t |> (fn (v_and_ts, t') => let val (vs, ts) = ListPair.unzip v_and_ts val vs' = (* list_diff vs (list_diff (Term.add_frees t' []) vs) *) Term.add_frees t' [] |> list_diff vs |> list_diff vs val v'_and_ts = map (fn v => (v, AList.lookup (op =) v_and_ts v |> the)) vs' in (v'_and_ts, t') end) (*Instantiate type variables in a term, based on a type environment*) fun type_devar tyenv (t : term) : term = case t of Const (s, ty) => Const (s, Term_Subst.instantiateT tyenv ty) | Free (s, ty) => Free (s, Term_Subst.instantiateT tyenv ty) | Var (idx, ty) => Var (idx, Term_Subst.instantiateT tyenv ty) | Bound _ => t | Abs (s, ty, t') => Abs (s, Term_Subst.instantiateT tyenv ty, type_devar tyenv t') | t1 $ t2 => type_devar tyenv t1 $ type_devar tyenv t2 (*Take a "diff" between an (abstract) thm's term, and another term (the latter is an instance of the form), then instantiate the abstract theorem. This is a way of turning the latter term into a theorem, but without exposing the proof-search functions to complex terms. In addition to the abstract thm ("scheme_thm"), this function is also supplied with the (sub)term of the abstract thm ("scheme_t") we want to use in the diff, in case only part of "scheme_t" might be needed (not the whole "Thm.prop_of scheme_thm")*) fun diff_and_instantiate ctxt scheme_thm scheme_t instance_t = let val (term_pairing, type_pairing) = diff (Proof_Context.theory_of ctxt) (scheme_t, instance_t) (*valuation of type variables*) val typeval = map (fn (v, T) => (dest_TVar v, Thm.ctyp_of ctxt T)) type_pairing val typeval_env = TVars.make (map (apfst dest_TVar) type_pairing) (*valuation of term variables*) val termval = map (apfst (dest_Var o type_devar typeval_env)) term_pairing |> map (apsnd (Thm.cterm_of ctxt)) in - Thm.instantiate (typeval, termval) scheme_thm + Thm.instantiate (TVars.make typeval, Vars.make termval) scheme_thm end (*FIXME this is bad form?*) val try_dest_Trueprop = perhaps (try HOLogic.dest_Trueprop) (** Some tacticals **) (*Lift COND to be parametrised by subgoal number*) fun COND' sat' tac'1 tac'2 i = COND (sat' i) (tac'1 i) (tac'2 i) (*Apply simplification ("wittler") as few times as possible before being able to apply a tactic ("tac"). This is like a lazy version of REPEAT, since it attempts to REPEAT a tactic the smallest number times as possible, to make some other tactic succeed subsequently.*) fun ASAP wittler (tac : int -> tactic) (i : int) = fn st => let val tac_result = tac i st val pulled_tac_result = Seq.pull tac_result val tac_failed = is_none pulled_tac_result orelse not (has_fewer_prems 1 (fst (the pulled_tac_result))) in if tac_failed then (wittler THEN' ASAP wittler tac) i st else tac_result end (** Some tactics **) fun break_hypotheses_tac ctxt = CHANGED o ((REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}) THEN' (REPEAT_DETERM o eresolve_tac ctxt @{thms disjE})) (*Prove subgoals of form A ==> B1 | ... | A | ... | Bn*) fun clause_breaker_tac ctxt = (REPEAT o resolve_tac ctxt @{thms disjI1 disjI2 conjI}) THEN' assume_tac ctxt (* Refines a subgoal have the form: A1 ... An ==> B1 | ... | Aj | ... | Bi | ... | Ak | ... into multiple subgoals of the form: A'1 ==> B1 | ... | Aj | ... | Bi | ... | Ak | ... : A'm ==> B1 | ... | Aj | ... | Bi | ... | Ak | ... where {A'1 .. A'm} is disjoint from {B1, ..., Aj, ..., Bi, ..., Ak, ...} (and solves the subgoal completely if the first set is empty) *) fun batter_tac ctxt i = break_hypotheses_tac ctxt i THEN ALLGOALS (TRY o clause_breaker_tac ctxt) (*Same idiom as ex_expander_tac*) fun dist_all_and_tac ctxt i = let val simpset = empty_simpset ctxt |> Simplifier.add_simp @{lemma "\x. P x \ Q x \ (\x. P x) \ (\x. Q x)" by (rule eq_reflection, auto)} in CHANGED (asm_full_simp_tac simpset i) end fun reassociate_conjs_tac ctxt = asm_full_simp_tac (Simplifier.add_simp @{lemma "(A & B) & C == A & B & C" by auto} (*FIXME duplicates @{thm simp_meta(3)}*) (Raw_Simplifier.empty_simpset ctxt)) #> CHANGED #> REPEAT_DETERM (** Subgoal analysis **) (*Given an inference C ----- D This function returns "SOME X" if C = "! X. C'". If C has no quantification prefix, then returns NONE.*) fun head_quantified_variable ctxt i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val hypos = if null gls then [] else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> snd |> Logic.strip_horn |> fst fun foralls_of_hd_hypos () = hd hypos |> try_dest_Trueprop |> strip_top_All_vars |> #1 |> rev val quantified_variables = foralls_of_hd_hypos () in if null hypos orelse null quantified_variables then NONE else SOME (hd quantified_variables) end (** Builders for goal analysers or transformers **) (*Lifts function over terms to apply it to subgoals. "fun_over_terms" has type (term list * term -> 'a), where (term list * term) will be the term representations of the hypotheses and conclusion. if i_opt=SOME i then applies fun_over_terms to that subgoal and returns singleton result. otherwise applies fun_over_terms to all subgoals and return list of results.*) fun TERMFUN (fun_over_terms : term list * term -> 'a) (i_opt : int option) : thm -> 'a list = fn st => let val t_raws = Thm.prop_of st |> strip_top_all_vars [] |> snd |> Logic.strip_horn |> fst in if null t_raws then [] else let val ts = let val stripper = strip_top_all_vars [] #> snd #> Logic.strip_horn #> apsnd try_dest_Trueprop #> apfst (map try_dest_Trueprop) in map stripper t_raws end in case i_opt of NONE => map fun_over_terms ts | SOME i => nth ts (i - 1) |> fun_over_terms |> single end end (*Applies a predicate to subgoal(s) conclusion(s)*) fun TERMPRED (hyp_pred_over_terms : term -> bool) (conc_pred_over_terms : term -> bool) (i_opt : int option) : thm -> bool = fn st => let val hyp_results = TERMFUN (fst (*discard hypotheses*) #> map hyp_pred_over_terms) i_opt st val conc_results = TERMFUN (snd (*discard hypotheses*) #> conc_pred_over_terms) i_opt st val _ = \<^assert> (length hyp_results = length conc_results) in if null hyp_results then true else let val hyps_conjoined = fold (fn a => fn b => b andalso (forall (fn x => x) a)) hyp_results true val concs_conjoined = fold (fn a => fn b => b andalso a) conc_results true in hyps_conjoined andalso concs_conjoined end end (** Tracing **) (*If "tac i st" succeeds then msg is printed to "trace" channel*) fun trace_tac' ctxt msg tac i st = let val result = tac i st in if Config.get ctxt tptp_trace_reconstruction andalso not (is_none (Seq.pull result)) then (tracing msg; result) else result end end diff --git a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy --- a/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy +++ b/src/HOL/TPTP/TPTP_Proof_Reconstruction.thy @@ -1,2285 +1,2285 @@ (* Title: HOL/TPTP/TPTP_Proof_Reconstruction.thy Author: Nik Sultana, Cambridge University Computer Laboratory Proof reconstruction for Leo-II. USAGE: * Simple call the "reconstruct_leo2" function. * For more advanced use, you could use the component functions used in "reconstruct_leo2" -- see TPTP_Proof_Reconstruction_Test.thy for examples of this. This file contains definitions describing how to interpret LEO-II's calculus in Isabelle/HOL, as well as more general proof-handling functions. The definitions in this file serve to build an intermediate proof script which is then evaluated into a tactic -- both these steps are independent of LEO-II, and are defined in the TPTP_Reconstruct SML module. CONFIG: The following attributes are mainly useful for debugging: tptp_unexceptional_reconstruction | unexceptional_reconstruction |-- when these are true, a low-level exception is allowed to float to the top (instead of triggering a higher-level exception, or simply indicating that the reconstruction failed). tptp_max_term_size --- fail of a term exceeds this size. "0" is taken to mean infinity. tptp_informative_failure | informative_failure |-- produce more output during reconstruction. tptp_trace_reconstruction | There are also two attributes, independent of the code here, that influence the success of reconstruction: blast_depth_limit and unify_search_bound. These are documented in their respective modules, but in summary, if unify_search_bound is increased then we can handle larger terms (at the cost of performance), since the unification engine takes longer to give up the search; blast_depth_limit is a limit on proof search performed by Blast. Blast is used for the limited proof search that needs to be done to interpret instances of LEO-II's inference rules. TODO: use RemoveRedundantQuantifications instead of the ad hoc use of remove_redundant_quantification_in_lit and remove_redundant_quantification *) theory TPTP_Proof_Reconstruction imports TPTP_Parser TPTP_Interpret (* keywords "import_leo2_proof" :: thy_decl *) (*FIXME currently unused*) begin section "Setup" ML \ val tptp_unexceptional_reconstruction = Attrib.setup_config_bool \<^binding>\tptp_unexceptional_reconstruction\ (K false) fun unexceptional_reconstruction ctxt = Config.get ctxt tptp_unexceptional_reconstruction val tptp_informative_failure = Attrib.setup_config_bool \<^binding>\tptp_informative_failure\ (K false) fun informative_failure ctxt = Config.get ctxt tptp_informative_failure val tptp_trace_reconstruction = Attrib.setup_config_bool \<^binding>\tptp_trace_reconstruction\ (K false) val tptp_max_term_size = Attrib.setup_config_int \<^binding>\tptp_max_term_size\ (K 0) (*0=infinity*) fun exceeds_tptp_max_term_size ctxt size = let val max = Config.get ctxt tptp_max_term_size in if max = 0 then false else size > max end \ (*FIXME move to TPTP_Proof_Reconstruction_Test_Units*) declare [[ tptp_unexceptional_reconstruction = false, (*NOTE should be "false" while testing*) tptp_informative_failure = true ]] ML_file \TPTP_Parser/tptp_reconstruct_library.ML\ ML "open TPTP_Reconstruct_Library" ML_file \TPTP_Parser/tptp_reconstruct.ML\ (*FIXME fudge*) declare [[ blast_depth_limit = 10, unify_search_bound = 5 ]] section "Proof reconstruction" text \There are two parts to proof reconstruction: \begin{itemize} \item interpreting the inferences \item building the skeleton, which indicates how to compose individual inferences into subproofs, and then compose the subproofs to give the proof). \end{itemize} One step detects unsound inferences, and the other step detects unsound composition of inferences. The two parts can be weakly coupled. They rely on a "proof index" which maps nodes to the inference information. This information consists of the (usually prover-specific) name of the inference step, and the Isabelle formalisation of the inference as a term. The inference interpretation then maps these terms into meta-theorems, and the skeleton is used to compose the inference-level steps into a proof. Leo2 operates on conjunctions of clauses. Each Leo2 inference has the following form, where Cx are clauses: C1 && ... && Cn ----------------- C'1 && ... && C'n Clauses consist of disjunctions of literals (shown as Px below), and might have a prefix of !-bound variables, as shown below. ! X... { P1 || ... || Pn} Literals are usually assigned a polarity, but this isn't always the case; you can come across inferences looking like this (where A is an object-level formula): F -------- F = true The symbol "||" represents literal-level disjunction and "&&" is clause-level conjunction. Rules will typically lift formula-level conjunctions; for instance the following rule lifts object-level disjunction: { (A | B) = true || ... } && ... -------------------------------------- { A = true || B = true || ... } && ... Using this setup, efficiency might be gained by only interpreting inferences once, merging identical inference steps, and merging identical subproofs into single inferences thus avoiding some effort. We can also attempt to minimising proof search when interpreting inferences. It is hoped that this setup can target other provers by modifying the clause representation to fit them, and adapting the inference interpretation to handle the rules used by the prover. It should also facilitate composing together proofs found by different provers. \ subsection "Instantiation" lemma polar_allE [rule_format]: "\(\x. P x) = True; (P x) = True \ R\ \ R" "\(\x. P x) = False; (P x) = False \ R\ \ R" by auto lemma polar_exE [rule_format]: "\(\x. P x) = True; \x. (P x) = True \ R\ \ R" "\(\x. P x) = False; \x. (P x) = False \ R\ \ R" by auto ML \ (*This carries out an allE-like rule but on (polarised) literals. Instead of yielding a free variable (which is a hell for the matcher) it seeks to use one of the subgoals' parameters. This ought to be sufficient for emulating extcnf_combined, but note that the complexity of the problem can be enormous.*) fun inst_parametermatch_tac ctxt thms i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val parameters = if null gls then [] else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> fst |> map fst (*just get the parameter names*) in if null parameters then no_tac st else let fun instantiate param = (map (Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), param)] []) thms |> FIRST') val attempts = map instantiate parameters in (fold (curry (op APPEND')) attempts (K no_tac)) i st end end (*Attempts to use the polar_allE theorems on a specific subgoal.*) fun forall_pos_tac ctxt = inst_parametermatch_tac ctxt @{thms polar_allE} \ ML \ (*This is similar to inst_parametermatch_tac, but prefers to match variables having identical names. Logically, this is a hack. But it reduces the complexity of the problem.*) fun nominal_inst_parametermatch_tac ctxt thm i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val parameters = if null gls then [] else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> fst |> map fst (*just get the parameter names*) in if null parameters then no_tac st else let fun instantiates param = Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), param)] [] thm val quantified_var = head_quantified_variable ctxt i st in if is_none quantified_var then no_tac st else if member (op =) parameters (the quantified_var |> fst) then instantiates (the quantified_var |> fst) i st else K no_tac i st end end \ subsection "Prefix massaging" ML \ exception NO_GOALS (*Get quantifier prefix of the hypothesis and conclusion, reorder the hypothesis' quantifiers to have the ones appearing in the conclusion first.*) fun canonicalise_qtfr_order ctxt i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst in if null gls then raise NO_GOALS else let val (params, (hyp_clause, conc_clause)) = rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> apsnd Logic.dest_implies val (hyp_quants, hyp_body) = HOLogic.dest_Trueprop hyp_clause |> strip_top_All_vars |> apfst rev val conc_quants = HOLogic.dest_Trueprop conc_clause |> strip_top_All_vars |> fst val new_hyp = (* fold absfree new_hyp_prefix hyp_body *) (*HOLogic.list_all*) fold_rev (fn (v, ty) => fn t => HOLogic.mk_all (v, ty, t)) (prefix_intersection_list hyp_quants conc_quants) hyp_body |> HOLogic.mk_Trueprop val thm = Goal.prove ctxt [] [] (Logic.mk_implies (hyp_clause, new_hyp)) (fn _ => (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI}))) THEN (REPEAT_DETERM (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE}))) THEN HEADGOAL (assume_tac ctxt)) in dresolve_tac ctxt [thm] i st end end \ subsection "Some general rules and congruences" (*this isn't an actual rule used in Leo2, but it seems to be applied implicitly during some Leo2 inferences.*) lemma polarise: "P ==> P = True" by auto ML \ fun is_polarised t = (TPTP_Reconstruct.remove_polarity true t; true) handle TPTP_Reconstruct.UNPOLARISED _ => false fun polarise_subgoal_hyps ctxt = COND' (SOME #> TERMPRED is_polarised (fn _ => true)) (K no_tac) (dresolve_tac ctxt @{thms polarise}) \ lemma simp_meta [rule_format]: "(A --> B) == (~A | B)" "(A | B) | C == A | B | C" "(A & B) & C == A & B & C" "(~ (~ A)) == A" (* "(A & B) == (~ (~A | ~B))" *) "~ (A & B) == (~A | ~B)" "~(A | B) == (~A) & (~B)" by auto subsection "Emulation of Leo2's inference rules" (*this is not included in simp_meta since it would make a mess of the polarities*) lemma expand_iff [rule_format]: "((A :: bool) = B) \ (~ A | B) & (~ B | A)" by (rule eq_reflection, auto) lemma polarity_switch [rule_format]: "(\ P) = True \ P = False" "(\ P) = False \ P = True" "P = False \ (\ P) = True" "P = True \ (\ P) = False" by auto lemma solved_all_splits: "False = True \ False" by simp ML \ fun solved_all_splits_tac ctxt = TRY (eresolve_tac ctxt @{thms conjE} 1) THEN resolve_tac ctxt @{thms solved_all_splits} 1 THEN assume_tac ctxt 1 \ lemma lots_of_logic_expansions_meta [rule_format]: "(((A :: bool) = B) = True) == (((A \ B) = True) & ((B \ A) = True))" "((A :: bool) = B) = False == (((~A) | B) = False) | (((~B) | A) = False)" "((F = G) = True) == (\x. (F x = G x)) = True" "((F = G) = False) == (\x. (F x = G x)) = False" "(A | B) = True == (A = True) | (B = True)" "(A & B) = False == (A = False) | (B = False)" "(A | B) = False == (A = False) & (B = False)" "(A & B) = True == (A = True) & (B = True)" "(~ A) = True == A = False" "(~ A) = False == A = True" "~ (A = True) == A = False" "~ (A = False) == A = True" by (rule eq_reflection, auto)+ (*this is used in extcnf_combined handler*) lemma eq_neg_bool: "((A :: bool) = B) = False ==> ((~ (A | B)) | ~ ((~ A) | (~ B))) = False" by auto lemma eq_pos_bool: "((A :: bool) = B) = True ==> ((~ (A | B)) | ~ (~ A | ~ B)) = True" "(A = B) = True \ A = True \ B = False" "(A = B) = True \ A = False \ B = True" by auto (*next formula is more versatile than "(F = G) = True \ \x. ((F x = G x) = True)" since it doesn't assume that clause is singleton. After splitqtfr, and after applying allI exhaustively to the conclusion, we can use the existing functions to find the "(F x = G x) = True" disjunct in the conclusion*) lemma eq_pos_func: "\ x. (F = G) = True \ (F x = G x) = True" by auto (*make sure the conclusion consists of just "False"*) lemma flip: "((A = True) ==> False) ==> A = False" "((A = False) ==> False) ==> A = True" by auto (*FIXME try to use Drule.equal_elim_rule1 directly for this*) lemma equal_elim_rule1: "(A \ B) \ A \ B" by auto lemmas leo2_rules = lots_of_logic_expansions_meta[THEN equal_elim_rule1] (*FIXME is there any overlap with lots_of_logic_expansions_meta or leo2_rules?*) lemma extuni_bool2 [rule_format]: "(A = B) = False \ (A = True) | (B = True)" by auto lemma extuni_bool1 [rule_format]: "(A = B) = False \ (A = False) | (B = False)" by auto lemma extuni_triv [rule_format]: "(A = A) = False \ R" by auto (*Order (of A, B, C, D) matters*) lemma dec_commut_eq [rule_format]: "((A = B) = (C = D)) = False \ (B = C) = False | (A = D) = False" "((A = B) = (C = D)) = False \ (B = D) = False | (A = C) = False" by auto lemma dec_commut_disj [rule_format]: "((A \ B) = (C \ D)) = False \ (B = C) = False \ (A = D) = False" by auto lemma extuni_func [rule_format]: "(F = G) = False \ (\X. (F X = G X)) = False" by auto subsection "Emulation: tactics" ML \ (*Instantiate a variable according to the info given in the proof annotation. Through this we avoid having to come up with instantiations during reconstruction.*) fun bind_tac ctxt prob_name ordered_binds = let val thy = Proof_Context.theory_of ctxt fun term_to_string t = Print_Mode.with_modes [""] (fn () => Output.output (Syntax.string_of_term ctxt t)) () val ordered_instances = TPTP_Reconstruct.interpret_bindings prob_name thy ordered_binds [] |> map (snd #> term_to_string) |> permute (*instantiate a list of variables, order matters*) fun instantiate_vars ctxt vars : tactic = map (fn var => Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), var)] [] @{thm allE} 1) vars |> EVERY fun instantiate_tac vars = instantiate_vars ctxt vars THEN (HEADGOAL (assume_tac ctxt)) in HEADGOAL (canonicalise_qtfr_order ctxt) THEN (REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI}))) THEN REPEAT_DETERM (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE})) (*now only the variable to instantiate should be left*) THEN FIRST (map instantiate_tac ordered_instances) end \ ML \ (*Simplification tactics*) local fun rew_goal_tac thms ctxt i = rewrite_goal_tac ctxt thms i |> CHANGED in val expander_animal = rew_goal_tac (@{thms simp_meta} @ @{thms lots_of_logic_expansions_meta}) val simper_animal = rew_goal_tac @{thms simp_meta} end \ lemma prop_normalise [rule_format]: "(A | B) | C == A | B | C" "(A & B) & C == A & B & C" "A | B == ~(~A & ~B)" "~~ A == A" by auto ML \ (*i.e., break_conclusion*) fun flip_conclusion_tac ctxt = let val default_tac = (TRY o CHANGED o (rewrite_goal_tac ctxt @{thms prop_normalise})) THEN' resolve_tac ctxt @{thms notI} THEN' (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}) THEN' (TRY o (expander_animal ctxt)) in default_tac ORELSE' resolve_tac ctxt @{thms flip} end \ subsection "Skolemisation" lemma skolemise [rule_format]: "\ P. (\ (\x. P x)) \ \ (P (SOME x. ~ P x))" proof - have "\ P. (\ (\x. P x)) \ \ (P (SOME x. ~ P x))" proof - fix P assume ption: "\ (\x. P x)" hence a: "\x. \ P x" by force have hilbert : "\P. (\x. P x) \ (P (SOME x. P x))" proof - fix P assume "(\x. P x)" thus "(P (SOME x. P x))" apply auto apply (rule someI) apply auto done qed from a show "\ P (SOME x. \ P x)" proof - assume "\x. \ P x" hence "\ P (SOME x. \ P x)" by (rule hilbert) thus ?thesis . qed qed thus ?thesis by blast qed lemma polar_skolemise [rule_format]: "\P. (\x. P x) = False \ (P (SOME x. \ P x)) = False" proof - have "\P. (\x. P x) = False \ (P (SOME x. \ P x)) = False" proof - fix P assume ption: "(\x. P x) = False" hence "\ (\x. P x)" by force hence "\ All P" by force hence "\ (P (SOME x. \ P x))" by (rule skolemise) thus "(P (SOME x. \ P x)) = False" by force qed thus ?thesis by blast qed lemma leo2_skolemise [rule_format]: "\P sk. (\x. P x) = False \ (sk = (SOME x. \ P x)) \ (P sk) = False" by (clarify, rule polar_skolemise) lemma lift_forall [rule_format]: "\x. (\x. A x) = True \ (A x) = True" "\x. (\x. A x) = False \ (A x) = False" by auto lemma lift_exists [rule_format]: "\(All P) = False; sk = (SOME x. \ P x)\ \ P sk = False" "\(Ex P) = True; sk = (SOME x. P x)\ \ P sk = True" apply (drule polar_skolemise, simp) apply (simp, drule someI_ex, simp) done ML \ (*FIXME LHS should be constant. Currently allow variables for testing. Probably should still allow Vars (but not Frees) since they'll act as intermediate values*) fun conc_is_skolem_def t = case t of Const (\<^const_name>\HOL.eq\, _) $ t' $ (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ _) => let val (h, args) = strip_comb t' |> apfst (strip_abs #> snd #> strip_comb #> fst) val get_const_name = dest_Const #> fst val h_property = is_Free h orelse is_Var h orelse (is_Const h andalso (get_const_name h <> get_const_name \<^term>\HOL.Ex\) andalso (get_const_name h <> get_const_name \<^term>\HOL.All\) andalso (h <> \<^term>\Hilbert_Choice.Eps\) andalso (h <> \<^term>\HOL.conj\) andalso (h <> \<^term>\HOL.disj\) andalso (h <> \<^term>\HOL.eq\) andalso (h <> \<^term>\HOL.implies\) andalso (h <> \<^term>\HOL.The\) andalso (h <> \<^term>\HOL.Ex1\) andalso (h <> \<^term>\HOL.Not\) andalso (h <> \<^term>\HOL.iff\) andalso (h <> \<^term>\HOL.not_equal\)) val args_property = fold (fn t => fn b => b andalso is_Free t) args true in h_property andalso args_property end | _ => false \ ML \ (*Hack used to detect if a Skolem definition, with an LHS Var, has had the LHS instantiated into an unacceptable term.*) fun conc_is_bad_skolem_def t = case t of Const (\<^const_name>\HOL.eq\, _) $ t' $ (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ _) => let val (h, args) = strip_comb t' val get_const_name = dest_Const #> fst val const_h_test = if is_Const h then (get_const_name h = get_const_name \<^term>\HOL.Ex\) orelse (get_const_name h = get_const_name \<^term>\HOL.All\) orelse (h = \<^term>\Hilbert_Choice.Eps\) orelse (h = \<^term>\HOL.conj\) orelse (h = \<^term>\HOL.disj\) orelse (h = \<^term>\HOL.eq\) orelse (h = \<^term>\HOL.implies\) orelse (h = \<^term>\HOL.The\) orelse (h = \<^term>\HOL.Ex1\) orelse (h = \<^term>\HOL.Not\) orelse (h = \<^term>\HOL.iff\) orelse (h = \<^term>\HOL.not_equal\) else true val h_property = not (is_Free h) andalso not (is_Var h) andalso const_h_test val args_property = fold (fn t => fn b => b andalso is_Free t) args true in h_property andalso args_property end | _ => false \ ML \ fun get_skolem_conc t = let val t' = strip_top_all_vars [] t |> snd |> try_dest_Trueprop in case t' of Const (\<^const_name>\HOL.eq\, _) $ t' $ (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ _) => SOME t' | _ => NONE end fun get_skolem_conc_const t = lift_option (fn t' => head_of t' |> strip_abs_body |> head_of |> dest_Const) (get_skolem_conc t) \ (* Technique for handling quantifiers: Principles: * allE should always match with a !! * exE should match with a constant, or bind a fresh !! -- currently not doing the latter since it never seems to arised in normal Leo2 proofs. *) ML \ fun forall_neg_tac candidate_consts ctxt i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val parameters = if null gls then "" else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> fst |> map fst (*just get the parameter names*) |> (fn l => if null l then "" else space_implode " " l |> pair " " |> (op ^)) in if null gls orelse null candidate_consts then no_tac st else let fun instantiate const_name = Rule_Insts.dres_inst_tac ctxt [((("sk", 0), Position.none), const_name ^ parameters)] [] @{thm leo2_skolemise} val attempts = map instantiate candidate_consts in (fold (curry (op APPEND')) attempts (K no_tac)) i st end end \ ML \ exception SKOLEM_DEF of term (*The tactic wasn't pointed at a skolem definition*) exception NO_SKOLEM_DEF of (*skolem const name*)string * Binding.binding * term (*The tactic could not find a skolem definition in the theory*) fun absorb_skolem_def ctxt prob_name_opt i = fn st => let val thy = Proof_Context.theory_of ctxt val gls = Thm.prop_of st |> Logic.strip_horn |> fst val conclusion = if null gls then (*this should never be thrown*) raise NO_GOALS else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> snd |> Logic.strip_horn |> snd fun skolem_const_info_of t = case t of Const (\<^const_name>\HOL.Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ t' $ (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ _)) => head_of t' |> strip_abs_body (*since in general might have a skolem term, so we want to rip out the prefixing lambdas to get to the constant (which should be at head position)*) |> head_of |> dest_Const | _ => raise SKOLEM_DEF t val const_name = skolem_const_info_of conclusion |> fst val def_name = const_name ^ "_def" val bnd_def = (*FIXME consts*) const_name |> Long_Name.implode o tl o Long_Name.explode (*FIXME hack to drop theory-name prefix*) |> Binding.qualified_name |> Binding.suffix_name "_def" val bnd_name = case prob_name_opt of NONE => bnd_def | SOME prob_name => (* Binding.qualify false (TPTP_Problem_Name.mangle_problem_name prob_name) *) bnd_def val thm = (case try (Thm.axiom thy) def_name of SOME thm => thm | NONE => if is_none prob_name_opt then (*This mode is for testing, so we can be a bit looser with theories*) (* FIXME bad theory context!? *) Thm.add_axiom_global (bnd_name, conclusion) thy |> fst |> snd else raise (NO_SKOLEM_DEF (def_name, bnd_name, conclusion))) in resolve_tac ctxt [Drule.export_without_context thm] i st end handle SKOLEM_DEF _ => no_tac st \ ML \ (* In current system, there should only be 2 subgoals: the one where the skolem definition is being built (with a Var in the LHS), and the other subgoal using Var. *) (*arity must be greater than 0. if arity=0 then there's no need to use this expensive matching.*) fun find_skolem_term ctxt consts_candidate arity = fn st => let val _ = \<^assert> (arity > 0) val gls = Thm.prop_of st |> Logic.strip_horn |> fst (*extract the conclusion of each subgoal*) val conclusions = if null gls then raise NO_GOALS else map (strip_top_all_vars [] #> snd #> Logic.strip_horn #> snd) gls (*Remove skolem-definition conclusion, to avoid wasting time analysing it*) |> filter (try_dest_Trueprop #> conc_is_skolem_def #> not) (*There should only be a single goal*) (*FIXME this might not always be the case, in practice*) (* |> tap (fn x => @{assert} (is_some (try the_single x))) *) (*look for subterms headed by a skolem constant, and whose arguments are all parameter Vars*) fun get_skolem_terms args (acc : term list) t = case t of (c as Const _) $ (v as Free _) => if c = consts_candidate andalso arity = length args + 1 then (list_comb (c, v :: args)) :: acc else acc | t1 $ (v as Free _) => get_skolem_terms (v :: args) acc t1 @ get_skolem_terms [] acc t1 | t1 $ t2 => get_skolem_terms [] acc t1 @ get_skolem_terms [] acc t2 | Abs (_, _, t') => get_skolem_terms [] acc t' | _ => acc in map (strip_top_All_vars #> snd) conclusions |> maps (get_skolem_terms [] []) |> distinct (op =) end \ ML \ fun instantiate_skols ctxt consts_candidates i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val (params, conclusion) = if null gls then raise NO_GOALS else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> apsnd (Logic.strip_horn #> snd) fun skolem_const_info_of t = case t of Const (\<^const_name>\HOL.Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ lhs $ (Const (\<^const_name>\Hilbert_Choice.Eps\, _) $ rhs)) => let (*the parameters we will concern ourselves with*) val params' = Term.add_frees lhs [] |> distinct (op =) (*check to make sure that params' <= params*) val _ = \<^assert> (forall (member (op =) params) params') val skolem_const_ty = let val (skolem_const_prety, no_params) = Term.strip_comb lhs |> apfst (dest_Var #> snd) (*head of lhs consists of a logical variable. we just want its type.*) |> apsnd length val _ = \<^assert> (length params = no_params) (*get value type of a function type after n arguments have been supplied*) fun get_val_ty n ty = if n = 0 then ty else get_val_ty (n - 1) (dest_funT ty |> snd) in get_val_ty no_params skolem_const_prety end in (skolem_const_ty, params') end | _ => raise (SKOLEM_DEF t) (* find skolem const candidates which, after applying distinct members of params' we end up with, give us something of type skolem_const_ty. given a candidate's type, skolem_const_ty, and params', we get some pemutations of params' (i.e. the order in which they can be given to the candidate in order to get skolem_const_ty). If the list of permutations is empty, then we cannot use that candidate. *) (* only returns a single matching -- since terms are linear, and variable arguments are Vars, order shouldn't matter, so we can ignore permutations. doesn't work with polymorphism (for which we'd need to use type unification) -- this is OK since no terms should be polymorphic, since Leo2 proofs aren't. *) fun use_candidate target_ty params acc cur_ty = if null params then if cur_ty = target_ty then SOME (rev acc) else NONE else let val (arg_ty, val_ty) = Term.dest_funT cur_ty (*now find a param of type arg_ty*) val (candidate_param, params') = find_and_remove (snd #> pair arg_ty #> (op =)) params in use_candidate target_ty params' (candidate_param :: acc) val_ty end handle TYPE ("dest_funT", _, _) => NONE (* FIXME fragile *) | _ => NONE (* FIXME avoid catch-all handler *) val (skolem_const_ty, params') = skolem_const_info_of conclusion (* For each candidate, build a term and pass it to Thm.instantiate, whic in turn is chained with PRIMITIVE to give us this_tactic. Big picture: we run the following: drule leo2_skolemise THEN' this_tactic NOTE: remember to APPEND' instead of ORELSE' the two tactics relating to skolemisation *) val filtered_candidates = map (dest_Const #> snd #> use_candidate skolem_const_ty params' []) consts_candidates (* prefiltered_candidates *) |> pair consts_candidates (* prefiltered_candidates *) |> ListPair.zip |> filter (snd #> is_none #> not) |> map (apsnd the) val skolem_terms = let fun make_result_t (t, args) = (* list_comb (t, map Free args) *) if length args > 0 then hd (find_skolem_term ctxt t (length args) st) else t in map make_result_t filtered_candidates end (*prefix a skolem term with bindings for the parameters*) (* val contextualise = fold absdummy (map snd params) *) val contextualise = fold absfree params val skolem_cts = map (contextualise #> Thm.cterm_of ctxt) skolem_terms (*now the instantiation code*) (*there should only be one Var -- that is from the previous application of drule leo2_skolemise. We look for it at the head position in some equation at a conclusion of a subgoal.*) val var_opt = let val pre_var = gls |> map (strip_top_all_vars [] #> snd #> Logic.strip_horn #> snd #> get_skolem_conc) |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) [] |> maps (switch Term.add_vars []) fun make_var pre_var = the_single pre_var |> Var |> Thm.cterm_of ctxt |> SOME in if null pre_var then NONE else make_var pre_var end fun instantiate_tac from to = - PRIMITIVE (Thm.instantiate ([], [(from, to)])) + PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(from, to)])) val tactic = if is_none var_opt then no_tac else fold (curry (op APPEND)) (map (instantiate_tac (dest_Var (Thm.term_of (the var_opt)))) skolem_cts) no_tac in tactic st end \ ML \ fun new_skolem_tac ctxt consts_candidates = let fun tac thm = dresolve_tac ctxt [thm] THEN' instantiate_skols ctxt consts_candidates in if null consts_candidates then K no_tac else FIRST' (map tac @{thms lift_exists}) end \ (* need a tactic to expand "? x . P" to "~ ! x. ~ P" *) ML \ fun ex_expander_tac ctxt i = let val simpset = empty_simpset ctxt (*NOTE for some reason, Bind exception gets raised if ctxt's simpset isn't emptied*) |> Simplifier.add_simp @{lemma "Ex P == (\ (\x. \ P x))" by auto} in CHANGED (asm_full_simp_tac simpset i) end \ subsubsection "extuni_dec" ML \ (*n-ary decomposition. Code is based on the n-ary arg_cong generator*) fun extuni_dec_n ctxt arity = let val _ = \<^assert> (arity > 0) val is = 1 upto arity |> map Int.toString val arg_tys = map (fn i => TFree ("arg" ^ i ^ "_ty", \<^sort>\type\)) is val res_ty = TFree ("res" ^ "_ty", \<^sort>\type\) val f_ty = arg_tys ---> res_ty val f = Free ("f", f_ty) val xs = map (fn i => Free ("x" ^ i, TFree ("arg" ^ i ^ "_ty", \<^sort>\type\))) is (*FIXME DRY principle*) val ys = map (fn i => Free ("y" ^ i, TFree ("arg" ^ i ^ "_ty", \<^sort>\type\))) is val hyp_lhs = list_comb (f, xs) val hyp_rhs = list_comb (f, ys) val hyp_eq = HOLogic.eq_const res_ty $ hyp_lhs $ hyp_rhs val hyp = HOLogic.eq_const HOLogic.boolT $ hyp_eq $ \<^term>\False\ |> HOLogic.mk_Trueprop fun conc_eq i = let val ty = TFree ("arg" ^ i ^ "_ty", \<^sort>\type\) val x = Free ("x" ^ i, ty) val y = Free ("y" ^ i, ty) val eq = HOLogic.eq_const ty $ x $ y in HOLogic.eq_const HOLogic.boolT $ eq $ \<^term>\False\ end val conc_disjs = map conc_eq is val conc = if length conc_disjs = 1 then the_single conc_disjs else fold (fn t => fn t_conc => HOLogic.mk_disj (t_conc, t)) (tl conc_disjs) (hd conc_disjs) val t = Logic.mk_implies (hyp, HOLogic.mk_Trueprop conc) in Goal.prove ctxt [] [] t (fn _ => auto_tac ctxt) |> Drule.export_without_context end \ ML \ (*Determine the arity of a function which the "dec" unification rule is about to be applied. NOTE: * Assumes that there is a single hypothesis *) fun find_dec_arity i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst in if null gls then raise NO_GOALS else let val (params, (literal, conc_clause)) = rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] |> apsnd Logic.strip_horn |> apsnd (apfst the_single) val get_ty = HOLogic.dest_Trueprop #> strip_top_All_vars #> snd #> HOLogic.dest_eq (*polarity's "="*) #> fst #> HOLogic.dest_eq (*the unification constraint's "="*) #> fst #> head_of #> dest_Const #> snd fun arity_of ty = let val (_, res_ty) = dest_funT ty in 1 + arity_of res_ty end handle (TYPE ("dest_funT", _, _)) => 0 in arity_of (get_ty literal) end end (*given an inference, it returns the parameters (i.e., we've already matched the leading & shared quantification in the hypothesis & conclusion clauses), and the "raw" inference*) fun breakdown_inference i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst in if null gls then raise NO_GOALS else rpair (i - 1) gls |> uncurry nth |> strip_top_all_vars [] end (*build a custom elimination rule for extuni_dec, and instantiate it to match a specific subgoal*) fun extuni_dec_elim_rule ctxt arity i = fn st => let val rule = extuni_dec_n ctxt arity val rule_hyp = Thm.prop_of rule |> Logic.dest_implies |> fst (*assuming that rule has single hypothesis*) (*having run break_hypothesis earlier, we know that the hypothesis now consists of a single literal. We can (and should) disregard the conclusion, since it hasn't been "broken", and it might include some unwanted literals -- the latter could cause "diff" to fail (since they won't agree with the rule we have generated.*) val inference_hyp = snd (breakdown_inference i st) |> Logic.dest_implies |> fst (*assuming that inference has single hypothesis, as explained above.*) in TPTP_Reconstruct_Library.diff_and_instantiate ctxt rule rule_hyp inference_hyp end fun extuni_dec_tac ctxt i = fn st => let val arity = find_dec_arity i st fun elim_tac i st = let val rule = extuni_dec_elim_rule ctxt arity i st (*in case we itroduced free variables during instantiation, we generalise the rule to make those free variables into logical variables.*) |> Thm.forall_intr_frees |> Drule.export_without_context in dresolve_tac ctxt [rule] i st end handle NO_GOALS => no_tac st fun closure tac = (*batter fails if there's no toplevel disjunction in the hypothesis, so we also try atac*) SOLVE o (tac THEN' (batter_tac ctxt ORELSE' assume_tac ctxt)) val search_tac = ASAP (resolve_tac ctxt @{thms disjI1} APPEND' resolve_tac ctxt @{thms disjI2}) (FIRST' (map closure [dresolve_tac ctxt @{thms dec_commut_eq}, dresolve_tac ctxt @{thms dec_commut_disj}, elim_tac])) in (CHANGED o search_tac) i st end \ subsubsection "standard_cnf" (*Given a standard_cnf inference, normalise it e.g. ((A & B) & C \ D & E \ F \ G) = False is changed to (A & B & C & D & E & F \ G) = False then custom-build a metatheorem which validates this: (A & B & C & D & E & F \ G) = False ------------------------------------------- (A = True) & (B = True) & (C = True) & (D = True) & (E = True) & (F = True) & (G = False) and apply this metatheorem. There aren't any "positive" standard_cnfs in Leo2's calculus: e.g., "(A \ B) = True \ A = False | (A = True & B = True)" since "standard_cnf" seems to be applied at the preprocessing stage, together with splitting. *) ML \ (*Conjunctive counterparts to Term.disjuncts_aux and Term.disjuncts*) fun conjuncts_aux (Const (\<^const_name>\HOL.conj\, _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs) | conjuncts_aux t conjs = t :: conjs fun conjuncts t = conjuncts_aux t [] (*HOL equivalent of Logic.strip_horn*) local fun imp_strip_horn' acc (Const (\<^const_name>\HOL.implies\, _) $ A $ B) = imp_strip_horn' (A :: acc) B | imp_strip_horn' acc t = (acc, t) in fun imp_strip_horn t = imp_strip_horn' [] t |> apfst rev end \ ML \ (*Returns whether the antecedents are separated by conjunctions or implications; the number of antecedents; and the polarity of the original clause -- I think this will always be "false".*) fun standard_cnf_type ctxt i : thm -> (TPTP_Reconstruct.formula_kind * int * bool) option = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst val hypos = if null gls then raise NO_GOALS else rpair (i - 1) gls |> uncurry nth |> TPTP_Reconstruct.strip_top_all_vars [] |> snd |> Logic.strip_horn |> fst (*hypothesis clause should be singleton*) val _ = \<^assert> (length hypos = 1) val (t, pol) = the_single hypos |> try_dest_Trueprop |> TPTP_Reconstruct.strip_top_All_vars |> snd |> TPTP_Reconstruct.remove_polarity true (*literal is negative*) val _ = \<^assert> (not pol) val (antes, conc) = imp_strip_horn t val (ante_type, antes') = if length antes = 1 then let val conjunctive_antes = the_single antes |> conjuncts in if length conjunctive_antes > 1 then (TPTP_Reconstruct.Conjunctive NONE, conjunctive_antes) else (TPTP_Reconstruct.Implicational NONE, antes) end else (TPTP_Reconstruct.Implicational NONE, antes) in if null antes then NONE else SOME (ante_type, length antes', pol) end \ ML \ (*Given a certain standard_cnf type, build a metatheorem that would validate it*) fun mk_standard_cnf ctxt kind arity = let val _ = \<^assert> (arity > 0) val vars = 1 upto (arity + 1) |> map (fn i => Free ("x" ^ Int.toString i, HOLogic.boolT)) val consequent = hd vars val antecedents = tl vars val conc = fold (curry HOLogic.mk_conj) (map (fn var => HOLogic.mk_eq (var, \<^term>\True\)) antecedents) (HOLogic.mk_eq (consequent, \<^term>\False\)) val pre_hyp = case kind of TPTP_Reconstruct.Conjunctive NONE => curry HOLogic.mk_imp (if length antecedents = 1 then the_single antecedents else fold (curry HOLogic.mk_conj) (tl antecedents) (hd antecedents)) (hd vars) | TPTP_Reconstruct.Implicational NONE => fold (curry HOLogic.mk_imp) antecedents consequent val hyp = HOLogic.mk_eq (pre_hyp, \<^term>\False\) val t = Logic.mk_implies (HOLogic.mk_Trueprop hyp, HOLogic.mk_Trueprop conc) in Goal.prove ctxt [] [] t (fn _ => HEADGOAL (blast_tac ctxt)) |> Drule.export_without_context end \ ML \ (*Applies a d-tactic, then breaks it up conjunctively. This can be used to transform subgoals as follows: (A \ B) = False \ R | v \A = True; B = False\ \ R *) fun weak_conj_tac ctxt drule = dresolve_tac ctxt [drule] THEN' (REPEAT_DETERM o eresolve_tac ctxt @{thms conjE}) \ ML \ fun uncurry_lit_neg_tac ctxt = REPEAT_DETERM o dresolve_tac ctxt [@{lemma "(A \ B \ C) = False \ (A & B \ C) = False" by auto}] \ ML \ fun standard_cnf_tac ctxt i = fn st => let fun core_tactic i = fn st => case standard_cnf_type ctxt i st of NONE => no_tac st | SOME (kind, arity, _) => let val rule = mk_standard_cnf ctxt kind arity; in (weak_conj_tac ctxt rule THEN' assume_tac ctxt) i st end in (uncurry_lit_neg_tac ctxt THEN' TPTP_Reconstruct_Library.reassociate_conjs_tac ctxt THEN' core_tactic) i st end \ subsubsection "Emulator prep" ML \ datatype cleanup_feature = RemoveHypothesesFromSkolemDefs | RemoveDuplicates datatype loop_feature = Close_Branch | ConjI | King_Cong | Break_Hypotheses | Donkey_Cong (*simper_animal + ex_expander_tac*) | RemoveRedundantQuantifications | Assumption (*Closely based on Leo2 calculus*) | Existential_Free | Existential_Var | Universal | Not_pos | Not_neg | Or_pos | Or_neg | Equal_pos | Equal_neg | Extuni_Bool2 | Extuni_Bool1 | Extuni_Dec | Extuni_Bind | Extuni_Triv | Extuni_FlexRigid | Extuni_Func | Polarity_switch | Forall_special_pos datatype feature = ConstsDiff | StripQuantifiers | Flip_Conclusion | Loop of loop_feature list | LoopOnce of loop_feature list | InnerLoopOnce of loop_feature list | CleanUp of cleanup_feature list | AbsorbSkolemDefs \ ML \ fun can_feature x l = let fun sublist_of_clean_up el = case el of CleanUp l'' => SOME l'' | _ => NONE fun sublist_of_loop el = case el of Loop l'' => SOME l'' | _ => NONE fun sublist_of_loop_once el = case el of LoopOnce l'' => SOME l'' | _ => NONE fun sublist_of_inner_loop_once el = case el of InnerLoopOnce l'' => SOME l'' | _ => NONE fun check_sublist sought_sublist opt_list = if forall is_none opt_list then false else fold_options opt_list |> flat |> pair sought_sublist |> subset (op =) in case x of CleanUp l' => map sublist_of_clean_up l |> check_sublist l' | Loop l' => map sublist_of_loop l |> check_sublist l' | LoopOnce l' => map sublist_of_loop_once l |> check_sublist l' | InnerLoopOnce l' => map sublist_of_inner_loop_once l |> check_sublist l' | _ => exists (curry (op =) x) l end; fun loop_can_feature loop_feats l = can_feature (Loop loop_feats) l orelse can_feature (LoopOnce loop_feats) l orelse can_feature (InnerLoopOnce loop_feats) l; \<^assert> (can_feature ConstsDiff [StripQuantifiers, ConstsDiff]); \<^assert> (can_feature (CleanUp [RemoveHypothesesFromSkolemDefs]) [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]]); \<^assert> (can_feature (Loop []) [Loop [Existential_Var]]); \<^assert> (not (can_feature (Loop []) [InnerLoopOnce [Existential_Var]])); \ ML \ exception NO_LOOP_FEATS fun get_loop_feats (feats : feature list) = let val loop_find = fold (fn x => fn loop_feats_acc => if is_some loop_feats_acc then loop_feats_acc else case x of Loop loop_feats => SOME loop_feats | LoopOnce loop_feats => SOME loop_feats | InnerLoopOnce loop_feats => SOME loop_feats | _ => NONE) feats NONE in if is_some loop_find then the loop_find else raise NO_LOOP_FEATS end; \<^assert> (get_loop_feats [Loop [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal]] = [King_Cong, Break_Hypotheses, Existential_Free, Existential_Var, Universal]) \ (*use as elim rule to remove premises*) lemma insa_prems: "\Q; P\ \ P" by auto ML \ fun cleanup_skolem_defs ctxt feats = let (*remove hypotheses from skolem defs, after testing that they look like skolem defs*) val dehypothesise_skolem_defs = COND' (SOME #> TERMPRED (fn _ => true) conc_is_skolem_def) (REPEAT_DETERM o eresolve_tac ctxt @{thms insa_prems}) (K no_tac) in if can_feature (CleanUp [RemoveHypothesesFromSkolemDefs]) feats then ALLGOALS (TRY o dehypothesise_skolem_defs) else all_tac end \ ML \ fun remove_duplicates_tac feats = (if can_feature (CleanUp [RemoveDuplicates]) feats then distinct_subgoals_tac else all_tac) \ ML \ (*given a goal state, indicates the skolem constants committed-to in it (i.e. appearing in LHS of a skolem definition)*) fun which_skolem_concs_used ctxt = fn st => let val feats = [CleanUp [RemoveHypothesesFromSkolemDefs, RemoveDuplicates]] val scrubup_tac = cleanup_skolem_defs ctxt feats THEN remove_duplicates_tac feats in scrubup_tac st |> break_seq |> tap (fn (_, rest) => \<^assert> (null (Seq.list_of rest))) |> fst |> TERMFUN (snd (*discard hypotheses*) #> get_skolem_conc_const) NONE |> switch (fold (fn x => fn l => if is_some x then the x :: l else l)) [] |> map Const end \ ML \ fun exists_tac ctxt feats consts_diff = let val ex_var = if loop_can_feature [Existential_Var] feats andalso consts_diff <> [] then new_skolem_tac ctxt consts_diff (*We're making sure that each skolem constant is used once in instantiations.*) else K no_tac val ex_free = if loop_can_feature [Existential_Free] feats andalso consts_diff = [] then eresolve_tac ctxt @{thms polar_exE} else K no_tac in ex_var APPEND' ex_free end fun forall_tac ctxt feats = if loop_can_feature [Universal] feats then forall_pos_tac ctxt else K no_tac \ subsubsection "Finite types" (*lift quantification from a singleton literal to a singleton clause*) lemma forall_pos_lift: "\(\X. P X) = True; \X. (P X = True) \ R\ \ R" by auto (*predicate over the type of the leading quantified variable*) ML \ fun extcnf_forall_special_pos_tac ctxt = let val bool = ["True", "False"] val bool_to_bool = ["% _ . True", "% _ . False", "% x . x", "Not"] val tacs = map (fn t_s => (* FIXME proper context!? *) Rule_Insts.eres_inst_tac \<^context> [((("x", 0), Position.none), t_s)] [] @{thm allE} THEN' assume_tac ctxt) in (TRY o eresolve_tac ctxt @{thms forall_pos_lift}) THEN' (assume_tac ctxt ORELSE' FIRST' (*FIXME could check the type of the leading quantified variable, instead of trying everything*) (tacs (bool @ bool_to_bool))) end \ subsubsection "Emulator" lemma efq: "[|A = True; A = False|] ==> R" by auto ML \ fun efq_tac ctxt = (eresolve_tac ctxt @{thms efq} THEN' assume_tac ctxt) ORELSE' assume_tac ctxt \ ML \ (*This is applied to all subgoals, repeatedly*) fun extcnf_combined_main ctxt feats consts_diff = let (*This is applied to subgoals which don't have a conclusion consisting of a Skolem definition*) fun extcnf_combined_tac' ctxt i = fn st => let val skolem_consts_used_so_far = which_skolem_concs_used ctxt st val consts_diff' = subtract (op =) skolem_consts_used_so_far consts_diff fun feat_to_tac feat = case feat of Close_Branch => trace_tac' ctxt "mark: closer" (efq_tac ctxt) | ConjI => trace_tac' ctxt "mark: conjI" (resolve_tac ctxt @{thms conjI}) | King_Cong => trace_tac' ctxt "mark: expander_animal" (expander_animal ctxt) | Break_Hypotheses => trace_tac' ctxt "mark: break_hypotheses" (break_hypotheses_tac ctxt) | RemoveRedundantQuantifications => K all_tac (* FIXME Building this into the loop instead.. maybe not the ideal choice | RemoveRedundantQuantifications => trace_tac' ctxt "mark: strip_unused_variable_hyp" (REPEAT_DETERM o remove_redundant_quantification_in_lit) *) | Assumption => assume_tac ctxt (*FIXME both Existential_Free and Existential_Var run same code*) | Existential_Free => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff') | Existential_Var => trace_tac' ctxt "mark: forall_neg" (exists_tac ctxt feats consts_diff') | Universal => trace_tac' ctxt "mark: forall_pos" (forall_tac ctxt feats) | Not_pos => trace_tac' ctxt "mark: not_pos" (dresolve_tac ctxt @{thms leo2_rules(9)}) | Not_neg => trace_tac' ctxt "mark: not_neg" (dresolve_tac ctxt @{thms leo2_rules(10)}) | Or_pos => trace_tac' ctxt "mark: or_pos" (dresolve_tac ctxt @{thms leo2_rules(5)}) (*could add (6) for negated conjunction*) | Or_neg => trace_tac' ctxt "mark: or_neg" (dresolve_tac ctxt @{thms leo2_rules(7)}) | Equal_pos => trace_tac' ctxt "mark: equal_pos" (dresolve_tac ctxt (@{thms eq_pos_bool} @ [@{thm leo2_rules(3)}, @{thm eq_pos_func}])) | Equal_neg => trace_tac' ctxt "mark: equal_neg" (dresolve_tac ctxt [@{thm eq_neg_bool}, @{thm leo2_rules(4)}]) | Donkey_Cong => trace_tac' ctxt "mark: donkey_cong" (simper_animal ctxt THEN' ex_expander_tac ctxt) | Extuni_Bool2 => trace_tac' ctxt "mark: extuni_bool2" (dresolve_tac ctxt @{thms extuni_bool2}) | Extuni_Bool1 => trace_tac' ctxt "mark: extuni_bool1" (dresolve_tac ctxt @{thms extuni_bool1}) | Extuni_Bind => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv}) | Extuni_Triv => trace_tac' ctxt "mark: extuni_triv" (eresolve_tac ctxt @{thms extuni_triv}) | Extuni_Dec => trace_tac' ctxt "mark: extuni_dec_tac" (extuni_dec_tac ctxt) | Extuni_FlexRigid => trace_tac' ctxt "mark: extuni_flex_rigid" (assume_tac ctxt ORELSE' asm_full_simp_tac ctxt) | Extuni_Func => trace_tac' ctxt "mark: extuni_func" (dresolve_tac ctxt @{thms extuni_func}) | Polarity_switch => trace_tac' ctxt "mark: polarity_switch" (eresolve_tac ctxt @{thms polarity_switch}) | Forall_special_pos => trace_tac' ctxt "mark: dorall_special_pos" (extcnf_forall_special_pos_tac ctxt) val core_tac = get_loop_feats feats |> map feat_to_tac |> FIRST' in core_tac i st end (*This is applied to all subgoals, repeatedly*) fun extcnf_combined_tac ctxt i = COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i)) no_tac (extcnf_combined_tac' ctxt i) val core_tac = CHANGED (ALLGOALS (IF_UNSOLVED o TRY o extcnf_combined_tac ctxt)) val full_tac = REPEAT core_tac in CHANGED (if can_feature (InnerLoopOnce []) feats then core_tac else full_tac) end val interpreted_consts = [\<^const_name>\HOL.All\, \<^const_name>\HOL.Ex\, \<^const_name>\Hilbert_Choice.Eps\, \<^const_name>\HOL.conj\, \<^const_name>\HOL.disj\, \<^const_name>\HOL.eq\, \<^const_name>\HOL.implies\, \<^const_name>\HOL.The\, \<^const_name>\HOL.Ex1\, \<^const_name>\HOL.Not\, (* @{const_name HOL.iff}, *) (*FIXME do these exist?*) (* @{const_name HOL.not_equal}, *) \<^const_name>\HOL.False\, \<^const_name>\HOL.True\, \<^const_name>\Pure.imp\] fun strip_qtfrs_tac ctxt = REPEAT_DETERM (HEADGOAL (resolve_tac ctxt @{thms allI})) THEN REPEAT_DETERM (HEADGOAL (eresolve_tac ctxt @{thms exE})) THEN HEADGOAL (canonicalise_qtfr_order ctxt) THEN ((REPEAT (HEADGOAL (nominal_inst_parametermatch_tac ctxt @{thm allE}))) APPEND (REPEAT (HEADGOAL (inst_parametermatch_tac ctxt [@{thm allE}])))) (*FIXME need to handle "@{thm exI}"?*) (*difference in constants between the hypothesis clause and the conclusion clause*) fun clause_consts_diff thm = let val t = Thm.prop_of thm |> Logic.dest_implies |> fst (*This bit should not be needed, since Leo2 inferences don't have parameters*) |> TPTP_Reconstruct.strip_top_all_vars [] |> snd val do_diff = Logic.dest_implies #> uncurry TPTP_Reconstruct.new_consts_between #> filter (fn Const (n, _) => not (member (op =) interpreted_consts n)) in if head_of t = Logic.implies then do_diff t else [] end \ ML \ (*remove quantification in hypothesis clause (! X. t), if X not free in t*) fun remove_redundant_quantification ctxt i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst in if null gls then raise NO_GOALS else let val (params, (hyp_clauses, conc_clause)) = rpair (i - 1) gls |> uncurry nth |> TPTP_Reconstruct.strip_top_all_vars [] |> apsnd Logic.strip_horn in (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*) if length hyp_clauses > 1 then no_tac st else let val hyp_clause = the_single hyp_clauses val sep_prefix = HOLogic.dest_Trueprop #> TPTP_Reconstruct.strip_top_All_vars #> apfst rev val (hyp_prefix, hyp_body) = sep_prefix hyp_clause val (conc_prefix, conc_body) = sep_prefix conc_clause in if null hyp_prefix orelse member (op =) conc_prefix (hd hyp_prefix) orelse member (op =) (Term.add_frees hyp_body []) (hd hyp_prefix) then no_tac st else Rule_Insts.eres_inst_tac ctxt [((("x", 0), Position.none), "(@X. False)")] [] @{thm allE} i st end end end \ ML \ fun remove_redundant_quantification_ignore_skolems ctxt i = COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i)) no_tac (remove_redundant_quantification ctxt i) \ lemma drop_redundant_literal_qtfr: "(\X. P) = True \ P = True" "(\X. P) = True \ P = True" "(\X. P) = False \ P = False" "(\X. P) = False \ P = False" by auto ML \ (*remove quantification in the literal "(! X. t) = True/False" in the singleton hypothesis clause, if X not free in t*) fun remove_redundant_quantification_in_lit ctxt i = fn st => let val gls = Thm.prop_of st |> Logic.strip_horn |> fst in if null gls then raise NO_GOALS else let val (params, (hyp_clauses, conc_clause)) = rpair (i - 1) gls |> uncurry nth |> TPTP_Reconstruct.strip_top_all_vars [] |> apsnd Logic.strip_horn in (*this is to fail gracefully in case this tactic is applied to a goal which doesn't have a single hypothesis*) if length hyp_clauses > 1 then no_tac st else let fun literal_content (Const (\<^const_name>\HOL.eq\, _) $ lhs $ (rhs as \<^term>\True\)) = SOME (lhs, rhs) | literal_content (Const (\<^const_name>\HOL.eq\, _) $ lhs $ (rhs as \<^term>\False\)) = SOME (lhs, rhs) | literal_content t = NONE val hyp_clause = the_single hyp_clauses |> HOLogic.dest_Trueprop |> literal_content in if is_none hyp_clause then no_tac st else let val (hyp_lit_prefix, hyp_lit_body) = the hyp_clause |> (fn (t, polarity) => TPTP_Reconstruct.strip_top_All_vars t |> apfst rev) in if null hyp_lit_prefix orelse member (op =) (Term.add_frees hyp_lit_body []) (hd hyp_lit_prefix) then no_tac st else dresolve_tac ctxt @{thms drop_redundant_literal_qtfr} i st end end end end \ ML \ fun remove_redundant_quantification_in_lit_ignore_skolems ctxt i = COND (TERMPRED (fn _ => true) conc_is_skolem_def (SOME i)) no_tac (remove_redundant_quantification_in_lit ctxt i) \ ML \ fun extcnf_combined_tac ctxt prob_name_opt feats skolem_consts = fn st => let val thy = Proof_Context.theory_of ctxt (*Initially, st consists of a single goal, showing the hypothesis clause implying the conclusion clause. There are no parameters.*) val consts_diff = union (=) skolem_consts (if can_feature ConstsDiff feats then clause_consts_diff st else []) val main_tac = if can_feature (LoopOnce []) feats orelse can_feature (InnerLoopOnce []) feats then extcnf_combined_main ctxt feats consts_diff else if can_feature (Loop []) feats then BEST_FIRST (TERMPRED (fn _ => true) conc_is_skolem_def NONE, size_of_thm) (*FIXME maybe need to weaken predicate to include "solved form"?*) (extcnf_combined_main ctxt feats consts_diff) else all_tac (*to allow us to use the cleaning features*) (*Remove hypotheses from Skolem definitions, then remove duplicate subgoals, then we should be left with skolem definitions: absorb them as axioms into the theory.*) val cleanup = cleanup_skolem_defs ctxt feats THEN remove_duplicates_tac feats THEN (if can_feature AbsorbSkolemDefs feats then ALLGOALS (absorb_skolem_def ctxt prob_name_opt) else all_tac) val have_loop_feats = (get_loop_feats feats; true) handle NO_LOOP_FEATS => false val tec = (if can_feature StripQuantifiers feats then (REPEAT (CHANGED (strip_qtfrs_tac ctxt))) else all_tac) THEN (if can_feature Flip_Conclusion feats then HEADGOAL (flip_conclusion_tac ctxt) else all_tac) (*after stripping the quantifiers any remaining quantifiers can be simply eliminated -- they're redundant*) (*FIXME instead of just using allE, instantiate to a silly term, to remove opportunities for unification.*) THEN (REPEAT_DETERM (eresolve_tac ctxt @{thms allE} 1)) THEN (REPEAT_DETERM (resolve_tac ctxt @{thms allI} 1)) THEN (if have_loop_feats then REPEAT (CHANGED ((ALLGOALS (TRY o clause_breaker_tac ctxt)) (*brush away literals which don't change*) THEN (*FIXME move this to a different level?*) (if loop_can_feature [Polarity_switch] feats then all_tac else (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_ignore_skolems ctxt)))) THEN (TRY (IF_UNSOLVED (HEADGOAL (remove_redundant_quantification_in_lit_ignore_skolems ctxt))))) THEN (TRY main_tac))) else all_tac) THEN IF_UNSOLVED cleanup in DEPTH_SOLVE (CHANGED tec) st end \ subsubsection "unfold_def" (*this is used when handling unfold_tac, because the skeleton includes the definitions conjoined with the goal. it turns out that, for my tactic, the definitions are harmful. instead of modifying the skeleton (which may be nontrivial) i'm just dropping the information using this lemma. obviously, and from the name, order matters here.*) lemma drop_first_hypothesis [rule_format]: "\A; B\ \ B" by auto (*Unfold_def works by reducing the goal to a meta equation, then working on it until it can be discharged by atac, or reflexive, or else turned back into an object equation and broken down further.*) lemma un_meta_polarise: "(X \ True) \ X" by auto lemma meta_polarise: "X \ X \ True" by auto ML \ fun unfold_def_tac ctxt depends_on_defs = fn st => let (*This is used when we end up with something like (A & B) \ True \ (B & A) \ True. It breaks down this subgoal until it can be trivially discharged. *) val kill_meta_eqs_tac = dresolve_tac ctxt @{thms un_meta_polarise} THEN' resolve_tac ctxt @{thms meta_polarise} THEN' (REPEAT_DETERM o (eresolve_tac ctxt @{thms conjE})) THEN' (REPEAT_DETERM o (resolve_tac ctxt @{thms conjI} ORELSE' assume_tac ctxt)) val continue_reducing_tac = resolve_tac ctxt @{thms meta_eq_to_obj_eq} 1 THEN (REPEAT_DETERM (ex_expander_tac ctxt 1)) THEN TRY (polarise_subgoal_hyps ctxt 1) (*no need to REPEAT_DETERM here, since there should only be one hypothesis*) THEN TRY (dresolve_tac ctxt @{thms eq_reflection} 1) THEN (TRY ((CHANGED o rewrite_goal_tac ctxt (@{thm expand_iff} :: @{thms simp_meta})) 1)) THEN HEADGOAL (resolve_tac ctxt @{thms reflexive} ORELSE' assume_tac ctxt ORELSE' kill_meta_eqs_tac) val tactic = (resolve_tac ctxt @{thms polarise} 1 THEN assume_tac ctxt 1) ORELSE (REPEAT_DETERM (eresolve_tac ctxt @{thms conjE} 1 THEN eresolve_tac ctxt @{thms drop_first_hypothesis} 1) THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion) THEN (REPEAT_DETERM (ex_expander_tac ctxt 1)) THEN (TRY ((CHANGED o rewrite_goal_tac ctxt @{thms simp_meta}) 1)) THEN PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion) THEN (HEADGOAL (assume_tac ctxt) ORELSE (unfold_tac ctxt depends_on_defs THEN IF_UNSOLVED continue_reducing_tac))) in tactic st end \ subsection "Handling split 'preprocessing'" lemma split_tranfs: "\x. P x \ Q x \ (\x. P x) \ (\x. Q x)" "\ (\ A) \ A" "\x. A \ A" "(A \ B) \ C \ A \ B \ C" "A = B \ (A \ B) \ (B \ A)" by (rule eq_reflection, auto)+ (*Same idiom as ex_expander_tac*) ML \ fun split_simp_tac (ctxt : Proof.context) i = let val simpset = fold Simplifier.add_simp @{thms split_tranfs} (empty_simpset ctxt) in CHANGED (asm_full_simp_tac simpset i) end \ subsection "Alternative reconstruction tactics" ML \ (*An "auto"-based proof reconstruction, where we attempt to reconstruct each inference using auto_tac. A realistic tactic would inspect the inference name and act accordingly.*) fun auto_based_reconstruction_tac ctxt prob_name n = let val thy = Proof_Context.theory_of ctxt val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name in TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) n |> the |> (fn {inference_fmla, ...} => Goal.prove ctxt [] [] inference_fmla (fn pdata => auto_tac (#context pdata))) end \ (*An oracle-based reconstruction, which is only used to test the shunting part of the system*) oracle oracle_iinterp = "fn t => t" ML \ fun oracle_based_reconstruction_tac ctxt prob_name n = let val thy = Proof_Context.theory_of ctxt val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name in TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) n |> the |> (fn {inference_fmla, ...} => Thm.cterm_of ctxt inference_fmla) |> oracle_iinterp end \ subsection "Leo2 reconstruction tactic" ML \ exception UNSUPPORTED_ROLE exception INTERPRET_INFERENCE (*Failure reports can be adjusted to avoid interrupting an overall reconstruction process*) fun fail ctxt x = if unexceptional_reconstruction ctxt then (warning x; raise INTERPRET_INFERENCE) else error x fun interpret_leo2_inference_tac ctxt prob_name node = let val thy = Proof_Context.theory_of ctxt val _ = if Config.get ctxt tptp_trace_reconstruction then tracing ("interpret_inference reconstructing node" ^ node ^ " of " ^ TPTP_Problem_Name.mangle_problem_name prob_name) else () val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name fun nonfull_extcnf_combined_tac feats = extcnf_combined_tac ctxt (SOME prob_name) [ConstsDiff, StripQuantifiers, InnerLoopOnce (Break_Hypotheses :: (*FIXME RemoveRedundantQuantifications :: *) feats), AbsorbSkolemDefs] [] val source_inf_opt = AList.lookup (op =) (#meta pannot) #> the #> #source_inf_opt (*FIXME integrate this with other lookup code, or in the early analysis*) local fun node_is_of_role role node = AList.lookup (op =) (#meta pannot) node |> the |> #role |> (fn role' => role = role') fun roled_dependencies_names role = let fun values () = case role of TPTP_Syntax.Role_Definition => map (apsnd Binding.name_of) (#defs pannot) | TPTP_Syntax.Role_Axiom => map (apsnd Binding.name_of) (#axs pannot) | _ => raise UNSUPPORTED_ROLE in if is_none (source_inf_opt node) then [] else case the (source_inf_opt node) of TPTP_Proof.Inference (_, _, parent_inf) => map TPTP_Proof.parent_name parent_inf |> filter (node_is_of_role role) |> (*FIXME currently definitions are not included in the proof annotations, so i'm using all the definitions available in the proof. ideally i should only use the ones in the proof annotation.*) (fn x => if role = TPTP_Syntax.Role_Definition then let fun values () = map (apsnd Binding.name_of) (#defs pannot) in map snd (values ()) end else map (fn node => AList.lookup (op =) (values ()) node |> the) x) | _ => [] end val roled_dependencies = roled_dependencies_names #> map (Global_Theory.get_thm thy) in val depends_on_defs = roled_dependencies TPTP_Syntax.Role_Definition val depends_on_axs = roled_dependencies TPTP_Syntax.Role_Axiom val depends_on_defs_names = roled_dependencies_names TPTP_Syntax.Role_Definition end fun get_binds source_inf_opt = case the source_inf_opt of TPTP_Proof.Inference (_, _, parent_inf) => maps (fn TPTP_Proof.Parent _ => [] | TPTP_Proof.ParentWithDetails (_, parent_details) => parent_details) parent_inf | _ => [] val inference_name = case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of NONE => fail ctxt "Cannot reconstruct rule: no information" | SOME {inference_name, ...} => inference_name val default_tac = HEADGOAL (blast_tac ctxt) in case inference_name of "fo_atp_e" => HEADGOAL (Metis_Tactic.metis_tac [] ATP_Problem_Generate.combs_or_liftingN ctxt []) (*NOTE To treat E as an oracle use the following line: HEADGOAL (etac (oracle_based_reconstruction_tac ctxt prob_name node)) *) | "copy" => HEADGOAL (assume_tac ctxt ORELSE' (resolve_tac ctxt @{thms polarise} THEN' assume_tac ctxt)) | "polarity_switch" => nonfull_extcnf_combined_tac [Polarity_switch] | "solved_all_splits" => solved_all_splits_tac ctxt | "extcnf_not_pos" => nonfull_extcnf_combined_tac [Not_pos] | "extcnf_forall_pos" => nonfull_extcnf_combined_tac [Universal] | "negate_conjecture" => fail ctxt "Should not handle negate_conjecture here" | "unfold_def" => unfold_def_tac ctxt depends_on_defs | "extcnf_not_neg" => nonfull_extcnf_combined_tac [Not_neg] | "extcnf_or_neg" => nonfull_extcnf_combined_tac [Or_neg] | "extcnf_equal_pos" => nonfull_extcnf_combined_tac [Equal_pos] | "extcnf_equal_neg" => nonfull_extcnf_combined_tac [Equal_neg] | "extcnf_forall_special_pos" => nonfull_extcnf_combined_tac [Forall_special_pos] ORELSE HEADGOAL (blast_tac ctxt) | "extcnf_or_pos" => nonfull_extcnf_combined_tac [Or_pos] | "extuni_bool2" => nonfull_extcnf_combined_tac [Extuni_Bool2] | "extuni_bool1" => nonfull_extcnf_combined_tac [Extuni_Bool1] | "extuni_dec" => HEADGOAL (assume_tac ctxt) ORELSE nonfull_extcnf_combined_tac [Extuni_Dec] | "extuni_bind" => nonfull_extcnf_combined_tac [Extuni_Bind] | "extuni_triv" => nonfull_extcnf_combined_tac [Extuni_Triv] | "extuni_flex_rigid" => nonfull_extcnf_combined_tac [Extuni_FlexRigid] | "prim_subst" => nonfull_extcnf_combined_tac [Assumption] | "bind" => let val ordered_binds = get_binds (source_inf_opt node) in bind_tac ctxt prob_name ordered_binds end | "standard_cnf" => HEADGOAL (standard_cnf_tac ctxt) | "extcnf_forall_neg" => nonfull_extcnf_combined_tac [Existential_Var(* , RemoveRedundantQuantifications *)] (*FIXME RemoveRedundantQuantifications*) | "extuni_func" => nonfull_extcnf_combined_tac [Extuni_Func, Existential_Var] | "replace_leibnizEQ" => nonfull_extcnf_combined_tac [Assumption] | "replace_andrewsEQ" => nonfull_extcnf_combined_tac [Assumption] | "split_preprocessing" => (REPEAT (HEADGOAL (split_simp_tac ctxt))) THEN TRY (PRIMITIVE (Conv.fconv_rule Thm.eta_long_conversion)) THEN HEADGOAL (assume_tac ctxt) (*FIXME some of these could eventually be handled specially*) | "fac_restr" => default_tac | "sim" => default_tac | "res" => default_tac | "rename" => default_tac | "flexflex" => default_tac | other => fail ctxt ("Unknown inference rule: " ^ other) end \ ML \ fun interpret_leo2_inference ctxt prob_name node = let val thy = Proof_Context.theory_of ctxt val pannot = TPTP_Reconstruct.get_pannot_of_prob thy prob_name val (inference_name, inference_fmla) = case TPTP_Reconstruct.inference_at_node thy prob_name (#meta pannot) node of NONE => fail ctxt "Cannot reconstruct rule: no information" | SOME {inference_name, inference_fmla, ...} => (inference_name, inference_fmla) val proof_outcome = let fun prove () = Goal.prove ctxt [] [] inference_fmla (fn pdata => interpret_leo2_inference_tac (#context pdata) prob_name node) in if informative_failure ctxt then SOME (prove ()) else try prove () end in case proof_outcome of NONE => fail ctxt (Pretty.string_of (Pretty.block [Pretty.str ("Failed inference reconstruction for '" ^ inference_name ^ "' at node " ^ node ^ ":\n"), Syntax.pretty_term ctxt inference_fmla])) | SOME thm => thm end \ ML \ (*filter a set of nodes based on which inference rule was used to derive a node*) fun nodes_by_inference (fms : TPTP_Reconstruct.formula_meaning list) inference_rule = let fun fold_fun n l = case TPTP_Reconstruct.node_info fms #source_inf_opt n of NONE => l | SOME (TPTP_Proof.File _) => l | SOME (TPTP_Proof.Inference (rule_name, _, _)) => if rule_name = inference_rule then n :: l else l in fold fold_fun (map fst fms) [] end \ section "Importing proofs and reconstructing theorems" ML \ (*Preprocessing carried out on a LEO-II proof.*) fun leo2_on_load (pannot : TPTP_Reconstruct.proof_annotation) thy = let val ctxt = Proof_Context.init_global thy val dud = ("", Binding.empty, \<^term>\False\) val pre_skolem_defs = nodes_by_inference (#meta pannot) "extcnf_forall_neg" @ nodes_by_inference (#meta pannot) "extuni_func" |> map (fn x => (interpret_leo2_inference ctxt (#problem_name pannot) x; dud) handle NO_SKOLEM_DEF (s, bnd, t) => (s, bnd, t)) |> filter (fn (x, _, _) => x <> "") (*In case no skolem constants were introduced in that inference*) val skolem_defs = map (fn (x, y, _) => (x, y)) pre_skolem_defs val thy' = fold (fn skolem_def => fn thy => let val ((s, thm), thy') = Thm.add_axiom_global skolem_def thy (* val _ = warning ("Added skolem definition " ^ s ^ ": " ^ @{make_string thm}) *) (*FIXME use of make_string*) in thy' end) (map (fn (_, y, z) => (y, z)) pre_skolem_defs) thy in ({problem_name = #problem_name pannot, skolem_defs = skolem_defs, defs = #defs pannot, axs = #axs pannot, meta = #meta pannot}, thy') end \ ML \ (*Imports and reconstructs a LEO-II proof.*) fun reconstruct_leo2 path thy = let val prob_file = Path.base path val dir = Path.dir path val thy' = TPTP_Reconstruct.import_thm true [dir, prob_file] path leo2_on_load thy val ctxt = Context.Theory thy' |> Context.proof_of val prob_name = Path.implode prob_file |> TPTP_Problem_Name.parse_problem_name val theorem = TPTP_Reconstruct.reconstruct ctxt (TPTP_Reconstruct.naive_reconstruct_tac ctxt interpret_leo2_inference) prob_name in (*NOTE we could return the theorem value alone, since users could get the thy value from the thm value.*) (thy', theorem) end \ end diff --git a/src/HOL/Tools/Argo/argo_real.ML b/src/HOL/Tools/Argo/argo_real.ML --- a/src/HOL/Tools/Argo/argo_real.ML +++ b/src/HOL/Tools/Argo/argo_real.ML @@ -1,331 +1,333 @@ (* Title: HOL/Tools/Argo/argo_real.ML Author: Sascha Boehme Extension of the Argo tactic for the reals. *) structure Argo_Real: sig end = struct (* translating input terms *) fun trans_type _ \<^typ>\Real.real\ tcx = SOME (Argo_Expr.Real, tcx) | trans_type _ _ _ = NONE fun trans_term f (@{const Groups.uminus_class.uminus (real)} $ t) tcx = tcx |> f t |>> Argo_Expr.mk_neg |> SOME | trans_term f (@{const Groups.plus_class.plus (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_add2 |> SOME | trans_term f (@{const Groups.minus_class.minus (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_sub |> SOME | trans_term f (@{const Groups.times_class.times (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_mul |> SOME | trans_term f (@{const Rings.divide_class.divide (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_div |> SOME | trans_term f (@{const Orderings.ord_class.min (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_min |> SOME | trans_term f (@{const Orderings.ord_class.max (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_max |> SOME | trans_term f (@{const Groups.abs_class.abs (real)} $ t) tcx = tcx |> f t |>> Argo_Expr.mk_abs |> SOME | trans_term f (@{const Orderings.ord_class.less (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_lt |> SOME | trans_term f (@{const Orderings.ord_class.less_eq (real)} $ t1 $ t2) tcx = tcx |> f t1 ||>> f t2 |>> uncurry Argo_Expr.mk_le |> SOME | trans_term _ t tcx = (case try HOLogic.dest_number t of SOME (\<^typ>\Real.real\, n) => SOME (Argo_Expr.mk_num (Rat.of_int n), tcx) | _ => NONE) (* reverse translation *) fun mk_plus t1 t2 = @{const Groups.plus_class.plus (real)} $ t1 $ t2 fun mk_sum ts = uncurry (fold_rev mk_plus) (split_last ts) fun mk_times t1 t2 = @{const Groups.times_class.times (real)} $ t1 $ t2 fun mk_divide t1 t2 = @{const Rings.divide_class.divide (real)} $ t1 $ t2 fun mk_le t1 t2 = @{const Orderings.ord_class.less_eq (real)} $ t1 $ t2 fun mk_lt t1 t2 = @{const Orderings.ord_class.less (real)} $ t1 $ t2 fun mk_real_num i = HOLogic.mk_number \<^typ>\Real.real\ i fun mk_number n = let val (p, q) = Rat.dest n in if q = 1 then mk_real_num p else mk_divide (mk_real_num p) (mk_real_num q) end fun term_of _ (Argo_Expr.E (Argo_Expr.Num n, _)) = SOME (mk_number n) | term_of f (Argo_Expr.E (Argo_Expr.Neg, [e])) = SOME (@{const Groups.uminus_class.uminus (real)} $ f e) | term_of f (Argo_Expr.E (Argo_Expr.Add, es)) = SOME (mk_sum (map f es)) | term_of f (Argo_Expr.E (Argo_Expr.Sub, [e1, e2])) = SOME (@{const Groups.minus_class.minus (real)} $ f e1 $ f e2) | term_of f (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) = SOME (mk_times (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) = SOME (mk_divide (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Min, [e1, e2])) = SOME (@{const Orderings.ord_class.min (real)} $ f e1 $ f e2) | term_of f (Argo_Expr.E (Argo_Expr.Max, [e1, e2])) = SOME (@{const Orderings.ord_class.max (real)} $ f e1 $ f e2) | term_of f (Argo_Expr.E (Argo_Expr.Abs, [e])) = SOME (@{const Groups.abs_class.abs (real)} $ f e) | term_of f (Argo_Expr.E (Argo_Expr.Le, [e1, e2])) = SOME (mk_le (f e1) (f e2)) | term_of f (Argo_Expr.E (Argo_Expr.Lt, [e1, e2])) = SOME (mk_lt (f e1) (f e2)) | term_of _ _ = NONE (* proof replay for rewrite steps *) fun mk_rewr thm = thm RS @{thm eq_reflection} fun by_simp ctxt t = let fun prove {context, ...} = HEADGOAL (Simplifier.simp_tac context) in Goal.prove ctxt [] [] (HOLogic.mk_Trueprop t) prove end fun prove_num_pred ctxt n = by_simp ctxt (uncurry mk_lt (apply2 mk_number (if @0 < n then (@0, n) else (n, @0)))) fun simp_conv ctxt t = Conv.rewr_conv (mk_rewr (by_simp ctxt t)) fun nums_conv mk f ctxt n m = simp_conv ctxt (HOLogic.mk_eq (mk (mk_number n) (mk_number m), mk_number (f (n, m)))) val add_nums_conv = nums_conv mk_plus (op +) val mul_nums_conv = nums_conv mk_times (op *) val div_nums_conv = nums_conv mk_divide (op /) fun inv_num_conv ctxt = nums_conv mk_divide (fn (_, n) => Rat.inv n) ctxt @1 fun cmp_nums_conv ctxt b ct = let val t = if b then \<^const>\HOL.True\ else \<^const>\HOL.False\ in simp_conv ctxt (HOLogic.mk_eq (Thm.term_of ct, t)) ct end local fun is_add2 (@{const Groups.plus_class.plus (real)} $ _ $ _) = true | is_add2 _ = false fun is_add3 (@{const Groups.plus_class.plus (real)} $ _ $ t) = is_add2 t | is_add3 _ = false val flatten_thm = mk_rewr @{lemma "(a::real) + b + c = a + (b + c)" by simp} fun flatten_conv ct = if is_add2 (Thm.term_of ct) then Argo_Tactic.flatten_conv flatten_conv flatten_thm ct else Conv.all_conv ct val swap_conv = Conv.rewrs_conv (map mk_rewr @{lemma "(a::real) + (b + c) = b + (a + c)" "(a::real) + b = b + a" by simp_all}) val assoc_conv = Conv.rewr_conv (mk_rewr @{lemma "(a::real) + (b + c) = (a + b) + c" by simp}) val norm_monom_thm = mk_rewr @{lemma "1 * (a::real) = a" by simp} fun norm_monom_conv n = if n = @1 then Conv.rewr_conv norm_monom_thm else Conv.all_conv val add2_thms = map mk_rewr @{lemma "n * (a::real) + m * a = (n + m) * a" "n * (a::real) + a = (n + 1) * a" "(a::real) + m * a = (1 + m) * a" "(a::real) + a = (1 + 1) * a" by algebra+} val add3_thms = map mk_rewr @{lemma "n * (a::real) + (m * a + b) = (n + m) * a + b" "n * (a::real) + (a + b) = (n + 1) * a + b" "(a::real) + (m * a + b) = (1 + m) * a + b" "(a::real) + (a + b) = (1 + 1) * a + b" by algebra+} fun choose_conv cv2 cv3 ct = if is_add3 (Thm.term_of ct) then cv3 ct else cv2 ct fun join_num_conv ctxt n m = let val conv = add_nums_conv ctxt n m in choose_conv conv (assoc_conv then_conv Conv.arg1_conv conv) end fun join_monom_conv ctxt n m = let val conv = Conv.arg1_conv (add_nums_conv ctxt n m) then_conv norm_monom_conv (n + m) fun seq_conv thms cv = Conv.rewrs_conv thms then_conv cv in choose_conv (seq_conv add2_thms conv) (seq_conv add3_thms (Conv.arg1_conv conv)) end fun join_conv NONE = join_num_conv | join_conv (SOME _) = join_monom_conv fun bubble_down_conv _ _ [] ct = Conv.no_conv ct | bubble_down_conv _ _ [_] ct = Conv.all_conv ct | bubble_down_conv ctxt i ((m1 as (n1, i1)) :: (m2 as (n2, i2)) :: ms) ct = if i1 = i then if i2 = i then (join_conv i ctxt n1 n2 then_conv bubble_down_conv ctxt i ((n1 + n2, i) :: ms)) ct else (swap_conv then_conv Conv.arg_conv (bubble_down_conv ctxt i (m1 :: ms))) ct else Conv.arg_conv (bubble_down_conv ctxt i (m2 :: ms)) ct fun drop_var i ms = filter_out (fn (_, i') => i' = i) ms fun permute_conv _ [] [] ct = Conv.all_conv ct | permute_conv ctxt (ms as ((_, i) :: _)) [] ct = (bubble_down_conv ctxt i ms then_conv permute_conv ctxt (drop_var i ms) []) ct | permute_conv ctxt ms1 ms2 ct = let val (ms2', (_, i)) = split_last ms2 in (bubble_down_conv ctxt i ms1 then_conv permute_conv ctxt (drop_var i ms1) ms2') ct end val no_monom_conv = Conv.rewr_conv (mk_rewr @{lemma "0 * (a::real) = 0" by simp}) val norm_sum_conv = Conv.rewrs_conv (map mk_rewr @{lemma "0 * (a::real) + b = b" "(a::real) + 0 * b = a" "0 + (a::real) = a" "(a::real) + 0 = a" by simp_all}) fun drop0_conv ct = if is_add2 (Thm.term_of ct) then ((norm_sum_conv then_conv drop0_conv) else_conv Conv.arg_conv drop0_conv) ct else Conv.try_conv no_monom_conv ct fun full_add_conv ctxt ms1 ms2 = if eq_list (op =) (ms1, ms2) then flatten_conv else flatten_conv then_conv permute_conv ctxt ms1 ms2 then_conv drop0_conv in fun add_conv ctxt (ms1, ms2 as [(n, NONE)]) = if n = @0 then full_add_conv ctxt ms1 [] else full_add_conv ctxt ms1 ms2 | add_conv ctxt (ms1, ms2) = full_add_conv ctxt ms1 ms2 end val mul_suml_thm = mk_rewr @{lemma "((x::real) + y) * z = x * z + y * z" by (rule ring_distribs)} val mul_sumr_thm = mk_rewr @{lemma "(x::real) * (y + z) = x * y + x * z" by (rule ring_distribs)} fun mul_sum_conv thm ct = Conv.try_conv (Conv.rewr_conv thm then_conv Conv.binop_conv (mul_sum_conv thm)) ct val div_sum_thm = mk_rewr @{lemma "((x::real) + y) / z = x / z + y / z" by (rule add_divide_distrib)} fun div_sum_conv ct = Conv.try_conv (Conv.rewr_conv div_sum_thm then_conv Conv.binop_conv div_sum_conv) ct fun var_of_add_cmp (_ $ _ $ (_ $ _ $ (_ $ _ $ Var v))) = v | var_of_add_cmp t = raise TERM ("var_of_add_cmp", [t]) fun add_cmp_conv ctxt n thm = - let val v = var_of_add_cmp (Thm.prop_of thm) - in Conv.rewr_conv (Thm.instantiate ([], [(v, Thm.cterm_of ctxt (mk_number n))]) thm) end + let val v = var_of_add_cmp (Thm.prop_of thm) in + Conv.rewr_conv + (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt (mk_number n))]) thm) + end fun mul_cmp_conv ctxt n pos_thm neg_thm = let val thm = if @0 < n then pos_thm else neg_thm in Conv.rewr_conv (prove_num_pred ctxt n RS thm) end val neg_thm = mk_rewr @{lemma "-(x::real) = -1 * x" by simp} val sub_thm = mk_rewr @{lemma "(x::real) - y = x + -1 * y" by simp} val mul_zero_thm = mk_rewr @{lemma "0 * (x::real) = 0" by (rule mult_zero_left)} val mul_one_thm = mk_rewr @{lemma "1 * (x::real) = x" by (rule mult_1)} val mul_comm_thm = mk_rewr @{lemma "(x::real) * y = y * x" by simp} val mul_assocl_thm = mk_rewr @{lemma "((x::real) * y) * z = x * (y * z)" by simp} val mul_assocr_thm = mk_rewr @{lemma "(x::real) * (y * z) = (x * y) * z" by simp} val mul_divl_thm = mk_rewr @{lemma "((x::real) / y) * z = (x * z) / y" by simp} val mul_divr_thm = mk_rewr @{lemma "(x::real) * (y / z) = (x * y) / z" by simp} val div_zero_thm = mk_rewr @{lemma "0 / (x::real) = 0" by (rule div_0)} val div_one_thm = mk_rewr @{lemma "(x::real) / 1 = x" by (rule div_by_1)} val div_numl_thm = mk_rewr @{lemma "(x::real) / y = x * (1 / y)" by simp} val div_numr_thm = mk_rewr @{lemma "(x::real) / y = (1 / y) * x" by simp} val div_mull_thm = mk_rewr @{lemma "((x::real) * y) / z = x * (y / z)" by simp} val div_mulr_thm = mk_rewr @{lemma "(x::real) / (y * z) = (1 / y) * (x / z)" by simp} val div_divl_thm = mk_rewr @{lemma "((x::real) / y) / z = x / (y * z)" by simp} val div_divr_thm = mk_rewr @{lemma "(x::real) / (y / z) = (x * z) / y" by simp} val min_eq_thm = mk_rewr @{lemma "min (x::real) x = x" by (simp add: min_def)} val min_lt_thm = mk_rewr @{lemma "min (x::real) y = (if x <= y then x else y)" by (rule min_def)} val min_gt_thm = mk_rewr @{lemma "min (x::real) y = (if y <= x then y else x)" by (simp add: min_def)} val max_eq_thm = mk_rewr @{lemma "max (x::real) x = x" by (simp add: max_def)} val max_lt_thm = mk_rewr @{lemma "max (x::real) y = (if x <= y then y else x)" by (rule max_def)} val max_gt_thm = mk_rewr @{lemma "max (x::real) y = (if y <= x then x else y)" by (simp add: max_def)} val abs_thm = mk_rewr @{lemma "abs (x::real) = (if 0 <= x then x else -x)" by simp} val sub_eq_thm = mk_rewr @{lemma "((x::real) = y) = (x - y = 0)" by simp} val eq_le_thm = mk_rewr @{lemma "((x::real) = y) = ((x \ y) \ (y \ x))" by (rule order_eq_iff)} val add_le_thm = mk_rewr @{lemma "((x::real) \ y) = (x + n \ y + n)" by simp} val add_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x + n < y + n)" by simp} val sub_le_thm = mk_rewr @{lemma "((x::real) <= y) = (x - y <= 0)" by simp} val sub_lt_thm = mk_rewr @{lemma "((x::real) < y) = (x - y < 0)" by simp} val pos_mul_le_thm = mk_rewr @{lemma "0 < n ==> ((x::real) <= y) = (n * x <= n * y)" by simp} val neg_mul_le_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) <= y) = (n * y <= n * x)" by simp} val pos_mul_lt_thm = mk_rewr @{lemma "0 < n ==> ((x::real) < y) = (n * x < n * y)" by simp} val neg_mul_lt_thm = mk_rewr @{lemma "n < 0 ==> ((x::real) < y) = (n * y < n * x)" by simp} val not_le_thm = mk_rewr @{lemma "(\((x::real) \ y)) = (y < x)" by auto} val not_lt_thm = mk_rewr @{lemma "(\((x::real) < y)) = (y \ x)" by auto} fun replay_rewr _ Argo_Proof.Rewr_Neg = Conv.rewr_conv neg_thm | replay_rewr ctxt (Argo_Proof.Rewr_Add ps) = add_conv ctxt ps | replay_rewr _ Argo_Proof.Rewr_Sub = Conv.rewr_conv sub_thm | replay_rewr _ Argo_Proof.Rewr_Mul_Zero = Conv.rewr_conv mul_zero_thm | replay_rewr _ Argo_Proof.Rewr_Mul_One = Conv.rewr_conv mul_one_thm | replay_rewr ctxt (Argo_Proof.Rewr_Mul_Nums (n, m)) = mul_nums_conv ctxt n m | replay_rewr _ Argo_Proof.Rewr_Mul_Comm = Conv.rewr_conv mul_comm_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Left) = Conv.rewr_conv mul_assocl_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Right) = Conv.rewr_conv mul_assocr_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Left) = mul_sum_conv mul_suml_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Right) = mul_sum_conv mul_sumr_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Left) = Conv.rewr_conv mul_divl_thm | replay_rewr _ (Argo_Proof.Rewr_Mul_Div Argo_Proof.Right) = Conv.rewr_conv mul_divr_thm | replay_rewr _ Argo_Proof.Rewr_Div_Zero = Conv.rewr_conv div_zero_thm | replay_rewr _ Argo_Proof.Rewr_Div_One = Conv.rewr_conv div_one_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Nums (n, m)) = div_nums_conv ctxt n m | replay_rewr _ (Argo_Proof.Rewr_Div_Num (Argo_Proof.Left, _)) = Conv.rewr_conv div_numl_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Num (Argo_Proof.Right, n)) = Conv.rewr_conv div_numr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | replay_rewr _ (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Left, _)) = Conv.rewr_conv div_mull_thm | replay_rewr ctxt (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Right, n)) = Conv.rewr_conv div_mulr_thm then_conv Conv.arg1_conv (inv_num_conv ctxt n) | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Left) = Conv.rewr_conv div_divl_thm | replay_rewr _ (Argo_Proof.Rewr_Div_Div Argo_Proof.Right) = Conv.rewr_conv div_divr_thm | replay_rewr _ Argo_Proof.Rewr_Div_Sum = div_sum_conv | replay_rewr _ Argo_Proof.Rewr_Min_Eq = Conv.rewr_conv min_eq_thm | replay_rewr _ Argo_Proof.Rewr_Min_Lt = Conv.rewr_conv min_lt_thm | replay_rewr _ Argo_Proof.Rewr_Min_Gt = Conv.rewr_conv min_gt_thm | replay_rewr _ Argo_Proof.Rewr_Max_Eq = Conv.rewr_conv max_eq_thm | replay_rewr _ Argo_Proof.Rewr_Max_Lt = Conv.rewr_conv max_lt_thm | replay_rewr _ Argo_Proof.Rewr_Max_Gt = Conv.rewr_conv max_gt_thm | replay_rewr _ Argo_Proof.Rewr_Abs = Conv.rewr_conv abs_thm | replay_rewr ctxt (Argo_Proof.Rewr_Eq_Nums b) = cmp_nums_conv ctxt b | replay_rewr _ Argo_Proof.Rewr_Eq_Sub = Conv.rewr_conv sub_eq_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Le = Conv.rewr_conv eq_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Nums (_, b)) = cmp_nums_conv ctxt b | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Le, n)) = add_cmp_conv ctxt n add_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Add (Argo_Proof.Lt, n)) = add_cmp_conv ctxt n add_lt_thm | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Le) = Conv.rewr_conv sub_le_thm | replay_rewr _ (Argo_Proof.Rewr_Ineq_Sub Argo_Proof.Lt) = Conv.rewr_conv sub_lt_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Le, n)) = mul_cmp_conv ctxt n pos_mul_le_thm neg_mul_le_thm | replay_rewr ctxt (Argo_Proof.Rewr_Ineq_Mul (Argo_Proof.Lt, n)) = mul_cmp_conv ctxt n pos_mul_lt_thm neg_mul_lt_thm | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Le) = Conv.rewr_conv not_le_thm | replay_rewr _ (Argo_Proof.Rewr_Not_Ineq Argo_Proof.Lt) = Conv.rewr_conv not_lt_thm | replay_rewr _ _ = Conv.no_conv (* proof replay *) val combine_thms = @{lemma "(a::real) < b ==> c < d ==> a + c < b + d" "(a::real) < b ==> c <= d ==> a + c < b + d" "(a::real) <= b ==> c < d ==> a + c < b + d" "(a::real) <= b ==> c <= d ==> a + c <= b + d" by auto} fun combine thm1 thm2 = hd (Argo_Tactic.discharges thm2 (Argo_Tactic.discharges thm1 combine_thms)) fun replay _ _ Argo_Proof.Linear_Comb prems = SOME (uncurry (fold_rev combine) (split_last prems)) | replay _ _ _ _ = NONE (* real extension of the Argo solver *) val _ = Theory.setup (Argo_Tactic.add_extension { trans_type = trans_type, trans_term = trans_term, term_of = term_of, replay_rewr = replay_rewr, replay = replay}) end diff --git a/src/HOL/Tools/Argo/argo_tactic.ML b/src/HOL/Tools/Argo/argo_tactic.ML --- a/src/HOL/Tools/Argo/argo_tactic.ML +++ b/src/HOL/Tools/Argo/argo_tactic.ML @@ -1,732 +1,733 @@ (* Title: HOL/Tools/Argo/argo_tactic.ML Author: Sascha Boehme HOL method and tactic for the Argo solver. *) signature ARGO_TACTIC = sig val trace: string Config.T val timeout: real Config.T (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} val add_extension: extension -> theory -> theory (* proof utilities *) val discharges: thm -> thm list -> thm list val flatten_conv: conv -> thm -> conv (* interface to the tactic as well as the underlying checker and prover *) datatype result = Satisfiable of term -> bool option | Unsatisfiable val check: term list -> Proof.context -> result * Proof.context val prove: thm list -> Proof.context -> thm option * Proof.context val argo_tac: Proof.context -> thm list -> int -> tactic end structure Argo_Tactic: ARGO_TACTIC = struct (* readable fresh names for terms *) fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n') fun fresh_type_name (Type (n, _)) = fresh_name n | fresh_type_name (TFree (n, _)) = fresh_name n | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) fun fresh_term_name (Const (n, _)) = fresh_name n | fresh_term_name (Free (n, _)) = fresh_name n | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i) | fresh_term_name _ = fresh_name "" (* tracing *) datatype mode = None | Basic | Full fun string_of_mode None = "none" | string_of_mode Basic = "basic" | string_of_mode Full = "full" fun requires_mode None = [] | requires_mode Basic = [Basic, Full] | requires_mode Full = [Full] val trace = Attrib.setup_config_string \<^binding>\argo_trace\ (K (string_of_mode None)) fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else () val tracing = output Basic val full_tracing = output Full fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else () val with_tracing = with_mode Basic val with_full_tracing = with_mode Full (* timeout *) val timeout = Attrib.setup_config_real \<^binding>\argo_timeout\ (K 10.0) val timeout_seconds = seconds o Config.apply timeout fun with_timeout ctxt f x = Timeout.apply (timeout_seconds ctxt) f x (* extending the tactic *) type trans_context = Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option type extension = { trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans', trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans', term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option, replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv, replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option} fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2) structure Extensions = Theory_Data ( type T = (serial * extension) list val empty = [] val extend = I val merge = merge eq_extension ) fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext)) fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt) fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt) fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx) val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type) val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term) fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e) fun ext_replay_rewr ctxt r = get_extensions ctxt |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r) |> Conv.first_conv fun ext_replay cprop_of ctxt rule prems = (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of SOME thm => thm | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, [])) (* translating input terms *) fun add_new_type T (names, types, terms) = let val (n, names') = fresh_type_name T names val ty = Argo_Expr.Type n in (ty, (names', Typtab.update (T, ty) types, terms)) end fun add_type T (tcx as (_, types, _)) = (case Typtab.lookup types T of SOME ty => (ty, tcx) | NONE => add_new_type T tcx) fun trans_type _ \<^typ>\HOL.bool\ = pair Argo_Expr.Bool | trans_type ctxt (Type (\<^type_name>\fun\, [T1, T2])) = trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func | trans_type ctxt T = (fn tcx => (case ext_trans_type ctxt (trans_type ctxt) T tcx of SOME result => result | NONE => add_type T tcx)) fun add_new_term ctxt t T tcx = let val (ty, (names, types, terms)) = trans_type ctxt T tcx val (n, names') = fresh_term_name t names val c = (n, ty) in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end fun add_term ctxt t (tcx as (_, _, terms)) = (case Termtab.lookup terms t of SOME c => (Argo_Expr.mk_con c, tcx) | NONE => add_new_term ctxt t (Term.fastype_of t) tcx) fun mk_eq \<^typ>\HOL.bool\ = Argo_Expr.mk_iff | mk_eq _ = Argo_Expr.mk_eq fun trans_term _ \<^const>\HOL.True\ = pair Argo_Expr.true_expr | trans_term _ \<^const>\HOL.False\ = pair Argo_Expr.false_expr | trans_term ctxt (\<^const>\HOL.Not\ $ t) = trans_term ctxt t #>> Argo_Expr.mk_not | trans_term ctxt (\<^const>\HOL.conj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2 | trans_term ctxt (\<^const>\HOL.disj\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2 | trans_term ctxt (\<^const>\HOL.implies\ $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp | trans_term ctxt (Const (\<^const_name>\HOL.If\, _) $ t1 $ t2 $ t3) = trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>> (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3) | trans_term ctxt (Const (\<^const_name>\HOL.eq\, T) $ t1 $ t2) = trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq (Term.domain_type T)) | trans_term ctxt (t as (t1 $ t2)) = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app)) | trans_term ctxt t = (fn tcx => (case ext_trans_term ctxt (trans_term ctxt) t tcx of SOME result => result | NONE => add_term ctxt t tcx)) fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop) (* invoking the solver *) type data = { names: Name.context, types: Argo_Expr.typ Typtab.table, terms: (string * Argo_Expr.typ) Termtab.table, argo: Argo_Solver.context} fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo} val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context structure Solver_Data = Proof_Data ( type T = data option fun init _ = SOME empty_data ) datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p fun raw_solve es argo = Model (Argo_Solver.assert es argo) handle Argo_Proof.UNSAT proof => Proof proof fun value_of terms model t = (case Termtab.lookup terms t of SOME c => model c | _ => NONE) fun trace_props props ctxt = tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:" (map (Syntax.pretty_term ctxt) props))) fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg = tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms") fun solve props ctxt = (case Solver_Data.get ctxt of NONE => error "bad Argo solver context" | SOME {names, types, terms, argo} => let val _ = with_tracing ctxt (trace_props props) val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms) val _ = tracing ctxt "starting the prover" in (case Timing.timing (raw_solve es) argo of (time, Proof proof) => let val _ = trace_result ctxt time "proof" in (Proof (terms', proof), Solver_Data.put NONE ctxt) end | (time, Model argo') => let val _ = trace_result ctxt time "model" val model = value_of terms' (Argo_Solver.model_of argo') in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end) end) (* reverse translation *) structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord) fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts) fun mk_nary' d _ [] = d | mk_nary' _ f ts = mk_nary f ts fun mk_ite t1 t2 t3 = let val T = Term.fastype_of t2 val ite = Const (\<^const_name>\HOL.If\, [\<^typ>\HOL.bool\, T, T] ---> T) in ite $ t1 $ t2 $ t3 end fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^const>\HOL.True\ | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^const>\HOL.False\ | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e) | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) = mk_nary' \<^const>\HOL.True\ HOLogic.mk_conj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) = mk_nary' \<^const>\HOL.False\ HOLogic.mk_disj (map (term_of cx) es) | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) = HOLogic.mk_imp (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) = mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3) | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) = HOLogic.mk_eq (term_of cx e1, term_of cx e2) | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) = term_of cx e1 $ term_of cx e2 | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) = (case Contab.lookup cons c of SOME t => t | NONE => error ("Unexpected expression named " ^ quote n)) | term_of (cx as (ctxt, _)) e = (case ext_term_of ctxt (term_of cx) e of SOME t => t | NONE => raise Fail "bad expression") fun as_prop ct = Thm.apply \<^cterm>\HOL.Trueprop\ ct fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e) fun cprop_of ctxt cons e = as_prop (cterm_of ctxt cons e) (* generic proof tools *) fun discharge thm rule = thm INCR_COMP rule fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule) fun discharges thm rules = [thm] RL rules fun under_assumption f ct = let val cprop = as_prop ct in Thm.implies_intr cprop (f (Thm.assume cprop)) end -fun instantiate cv ct = Thm.instantiate ([], [(Term.dest_Var (Thm.term_of cv), ct)]) +fun instantiate cv ct = + Thm.instantiate (TVars.empty, Vars.make [(Term.dest_Var (Thm.term_of cv), ct)]) (* proof replay for tautologies *) fun prove_taut ctxt ns t = Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t) (fn {context, ...} => HEADGOAL (Classical.fast_tac context)) fun with_frees ctxt n mk = let val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1)) val ts = map (Free o rpair \<^typ>\bool\) ns val t = mk_nary HOLogic.mk_disj (mk ts) in prove_taut ctxt ns t end fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i] fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)] fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts val iff_1_taut = @{lemma "P = Q \ P \ Q" by fast} val iff_2_taut = @{lemma "P = Q \ (\P) \ (\Q)" by fast} val iff_3_taut = @{lemma "\(P = Q) \ (\P) \ Q" by fast} val iff_4_taut = @{lemma "\(P = Q) \ P \ (\Q)" by fast} val ite_then_taut = @{lemma "\P \ (if P then t else u) = t" by auto} val ite_else_taut = @{lemma "P \ (if P then t else u) = u" by auto} fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i) | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut fun replay_taut ctxt k ct = let val rule = taut_rule_of ctxt k in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end (* proof replay for conjunct extraction *) fun replay_conjunct 0 1 thm = thm | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1} | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2} | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2}) (* proof replay for rewrite steps *) fun mk_rewr thm = thm RS @{thm eq_reflection} fun not_nary_conv rule i ct = if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct else Conv.all_conv ct val flatten_and_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} val flatten_or_thm = @{lemma "(P1 \ P2) \ P3 \ P1 \ P2 \ P3" by simp} fun flatten_conv cv rule ct = ( Conv.try_conv (Conv.arg_conv cv) then_conv Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct fun flat_conj_conv ct = (case Thm.term_of ct of \<^const>\HOL.conj\ $ _ $ _ => flatten_conv flat_conj_conv flatten_and_thm ct | _ => Conv.all_conv ct) fun flat_disj_conv ct = (case Thm.term_of ct of \<^const>\HOL.disj\ $ _ $ _ => flatten_conv flat_disj_conv flatten_or_thm ct | _ => Conv.all_conv ct) fun explode rule1 rule2 thm = explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm] val explode_conj = explode @{thm conjunct1} @{thm conjunct2} val explode_ndis = explode @{lemma "\(P \ Q) \ \P" by auto} @{lemma "\(P \ Q) \ \Q" by auto} fun pick_false i thms = nth thms i fun pick_dual rule (i1, i2) thms = rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1] val pick_dual_conj = pick_dual @{lemma "\P \ P \ False" by auto} val pick_dual_ndis = pick_dual @{lemma "\P \ P \ \True" by auto} fun join thm0 rule is thms = let val l = length thms val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is [] in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end val join_conj = join @{lemma "True" by auto} @{lemma "P \ Q \ P \ Q" by auto} val join_ndis = join @{lemma "\False" by auto} @{lemma "\P \ \Q \ \(P \ Q)" by auto} val false_thm = @{lemma "False \ P" by auto} val ntrue_thm = @{lemma "\True \ P" by auto} val iff_conj_thm = @{lemma "(P \ Q) \ (Q \ P) \ P = Q" by auto} val iff_ndis_thm = @{lemma "(\P \ \Q) \ (\Q \ \P) \ P = Q" by auto} fun iff_intro rule lf rf ct = let val lhs = under_assumption lf ct val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs)))) in mk_rewr (discharge2 lhs rhs rule) end fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\HOL.Not\ ct) fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1)) fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g)) val sort_conj = sort_nary with_conj join_conj explode_conj val sort_ndis = sort_nary with_ndis join_ndis explode_ndis val not_true_thm = mk_rewr @{lemma "(\True) = False" by auto} val not_false_thm = mk_rewr @{lemma "(\False) = True" by auto} val not_not_thm = mk_rewr @{lemma "(\\P) = P" by auto} val not_and_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_or_thm = mk_rewr @{lemma "(\(P \ Q)) = (\P \ \Q)" by auto} val not_iff_thms = map mk_rewr @{lemma "(\((\P) = Q)) = (P = Q)" "(\(P = (\Q))) = (P = Q)" "(\(P = Q)) = ((\P) = Q)" by auto} val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto} val iff_false_thms = map mk_rewr @{lemma "(False = P) = (\P)" "(P = False) = (\P)" by auto} val iff_not_not_thm = mk_rewr @{lemma "((\P) = (\Q)) = (P = Q)" by auto} val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto} val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto} val iff_dual_thms = map mk_rewr @{lemma "(P = (\P)) = False" "((\P) = P) = False" by auto} val imp_thm = mk_rewr @{lemma "(P \ Q) = (\P \ Q)" by auto} val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((\P \ Q) \ (P \ R) \ (Q \ R))" by auto} val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto} val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto} val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto} val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto} val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto} fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm) | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm) | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm | replay_rewr _ Argo_Proof.Rewr_Iff_Dual = Conv.rewrs_conv iff_dual_thms | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm | replay_rewr ctxt r = ext_replay_rewr ctxt r fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (Argo_Expr.App, [c1, c2])) ct = Conv.combination_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_conv ctxt (Argo_Proof.Args_Conv (_, cs)) ct = replay_args_conv ctxt cs ct | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct and replay_args_conv _ [] ct = Conv.all_conv ct | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct | replay_args_conv ctxt (c :: cs) ct = (case Term.head_of (Thm.term_of ct) of Const (\<^const_name>\HOL.If\, _) => let val (cs', c') = split_last cs in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct) fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm (* proof replay for clauses *) val prep_clause_rule = @{lemma "P \ \P \ False" by fast} val extract_lit_rule = @{lemma "(\(P \ Q) \ False) \ \P \ \Q \ False" by fast} fun add_lit i thm lits = let val ct = Thm.cprem_of thm 1 in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end fun extract_lits [] _ = error "Bad clause" | extract_lits [i] (thm, lits) = add_lit i thm lits | extract_lits (i :: is) (thm, lits) = extract_lits is (add_lit i (discharge thm extract_lit_rule) lits) fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2) fun replay_with_lits [] thm lits = (thm, lits) | replay_with_lits is thm lits = extract_lits is (discharge thm prep_clause_rule, lits) ||> Ord_List.make lit_ord fun replay_clause is thm = replay_with_lits is thm [] (* proof replay for unit resolution *) val unit_rule = @{lemma "(P \ False) \ (\P \ False) \ False" by fast} val unit_rule_var = Thm.dest_arg (Thm.dest_arg1 (Thm.cprem_of unit_rule 1)) val bogus_ct = \<^cterm>\HOL.True\ fun replay_unit_res lit (pthm, plits) (nthm, nlits) = let val plit = the (AList.lookup (op =) plits lit) val nlit = the (AList.lookup (op =) nlits (~lit)) val prune = Ord_List.remove lit_ord (lit, bogus_ct) in unit_rule |> instantiate unit_rule_var (Thm.dest_arg plit) |> Thm.elim_implies (Thm.implies_intr plit pthm) |> Thm.elim_implies (Thm.implies_intr nlit nthm) |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits)) end (* proof replay for hypothesis *) val dneg_rule = @{lemma "\\P \ P" by auto} fun replay_hyp i ct = if i < 0 then (Thm.assume ct, [(~i, ct)]) else let val cu = as_prop (Thm.apply \<^cterm>\HOL.Not\ (Thm.apply \<^cterm>\HOL.Not\ (Thm.dest_arg ct))) in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end (* proof replay for lemma *) fun replay_lemma is (thm, lits) = replay_with_lits is thm lits (* proof replay for reflexivity *) val refl_rule = @{thm refl} val refl_rule_var = Thm.dest_arg1 (Thm.dest_arg (Thm.cprop_of refl_rule)) fun replay_refl ct = Thm.instantiate (Thm.match (refl_rule_var, ct)) refl_rule (* proof replay for symmetry *) val symm_rules = @{lemma "a = b ==> b = a" "\(a = b) \ \(b = a)" by simp_all} fun replay_symm thm = hd (discharges thm symm_rules) (* proof replay for transitivity *) val trans_rules = @{lemma "\(a = b) \ b = c \ \(a = c)" "a = b \ \(b = c) \ \(a = c)" "a = b \ b = c ==> a = c" by simp_all} fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules)) (* proof replay for congruence *) fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong}) (* proof replay for substitution *) val subst_rule1 = @{lemma "\(p a) \ p = q \ a = b \ \(q b)" by simp} val subst_rule2 = @{lemma "p a \ p = q \ a = b \ q b" by simp} fun replay_subst thm1 thm2 thm3 = subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3] (* proof replay *) structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord) val unclausify_rule1 = @{lemma "(\P \ False) \ P" by auto} val unclausify_rule2 = @{lemma "(P \ False) \ \P" by auto} fun unclausify (thm, lits) ls = (case (Thm.prop_of thm, lits) of (\<^const>\HOL.Trueprop\ $ \<^const>\HOL.False\, [(_, ct)]) => let val thm = Thm.implies_intr ct thm in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end | _ => (thm, Ord_List.union lit_ord lits ls)) fun with_thms f tps = fold_map unclausify tps [] |>> f fun bad_premises () = raise Fail "bad number of premises" fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ()) fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ()) fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ()) fun replay_rule (ctxt, cons, facts) prems rule = (case rule of Argo_Proof.Axiom i => (nth facts i, []) | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), []) | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems | Argo_Proof.Clause is => replay_clause is (fst (hd prems)) | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems)) | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl) | Argo_Proof.Lemma is => replay_lemma is (hd prems) | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), []) | Argo_Proof.Symm => with_thms1 replay_symm prems | Argo_Proof.Trans => with_thms2 replay_trans prems | Argo_Proof.Cong => with_thms2 replay_cong prems | Argo_Proof.Subst => with_thms3 replay_subst prems | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems) fun with_cache f proof thm_cache = (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of SOME thm => (thm, thm_cache) | NONE => let val (thm, thm_cache') = f proof thm_cache in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end) fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' => let val id = Argo_Proof.string_of_proof_id proof_id val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs val rule_string = Argo_Proof.string_of_rule rule in full_tracing ctxt' (" " ^ id ^ " <- " ^ space_implode " " ids ^ " . " ^ rule_string) end) fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache = let val (proof_id, rule, proofs) = Argo_Proof.dest proof val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache val _ = trace_step ctxt proof_id rule proofs in (replay_rule env prems rule, thm_cache) end fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty fun replay ctxt terms facts proof = let val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts) val _ = tracing ctxt "replaying the proof" val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms") in thm end (* normalizing goals *) fun instantiate_elim_rule thm = let val ct = Drule.strip_imp_concl (Thm.cprop_of thm) in (case Thm.term_of ct of \<^const>\HOL.Trueprop\ $ Var (_, \<^typ>\HOL.bool\) => instantiate (Thm.dest_arg ct) \<^cterm>\HOL.False\ thm | Var _ => instantiate ct \<^cprop>\HOL.False\ thm | _ => thm) end fun atomize_conv ctxt ct = (case Thm.term_of ct of \<^const>\HOL.Trueprop\ $ _ => Conv.all_conv | \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct fun normalize ctxt thm = thm |> instantiate_elim_rule |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) |> Thm.forall_intr_vars |> Conv.fconv_rule (atomize_conv ctxt) (* prover, tactic and method *) datatype result = Satisfiable of term -> bool option | Unsatisfiable fun check props ctxt = (case solve props ctxt of (Proof _, ctxt') => (Unsatisfiable, ctxt') | (Model model, ctxt') => (Satisfiable model, ctxt')) fun prove thms ctxt = let val thms' = map (normalize ctxt) thms in (case solve (map Thm.prop_of thms') ctxt of (Model _, ctxt') => (NONE, ctxt') | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt')) end fun argo_tac ctxt thms = CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt) THEN' Tactic.resolve_tac ctxt [@{thm ccontr}] THEN' Subgoal.FOCUS (fn {context, prems, ...} => (case with_timeout context (prove (thms @ prems)) context of (SOME thm, _) => Tactic.resolve_tac context [thm] 1 | (NONE, _) => Tactical.no_tac)) ctxt val _ = Theory.setup (Method.setup \<^binding>\argo\ (Scan.optional Attrib.thms [] >> (fn thms => fn ctxt => METHOD (fn facts => HEADGOAL (argo_tac ctxt (thms @ facts))))) "Applies the Argo prover") end diff --git a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML --- a/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML +++ b/src/HOL/Tools/BNF/bnf_fp_def_sugar.ML @@ -1,2930 +1,2930 @@ (* Title: HOL/Tools/BNF/bnf_fp_def_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013, 2014 Sugared datatype and codatatype constructions. *) signature BNF_FP_DEF_SUGAR = sig type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list} type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list} type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list} type fp_sugar = {T: typ, BT: typ, X: typ, fp: BNF_Util.fp_kind, fp_res_index: int, fp_res: BNF_FP_Util.fp_result, pre_bnf: BNF_Def.bnf, fp_bnf: BNF_Def.bnf, absT_info: BNF_Comp.absT_info, fp_nesting_bnfs: BNF_Def.bnf list, live_nesting_bnfs: BNF_Def.bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option} val co_induct_of: 'a list -> 'a val strong_co_induct_of: 'a list -> 'a val morph_fp_bnf_sugar: morphism -> fp_bnf_sugar -> fp_bnf_sugar val morph_fp_co_induct_sugar: morphism -> fp_co_induct_sugar -> fp_co_induct_sugar val morph_fp_ctr_sugar: morphism -> fp_ctr_sugar -> fp_ctr_sugar val morph_fp_sugar: morphism -> fp_sugar -> fp_sugar val transfer_fp_sugar: theory -> fp_sugar -> fp_sugar val fp_sugar_of: Proof.context -> string -> fp_sugar option val fp_sugar_of_global: theory -> string -> fp_sugar option val fp_sugars_of: Proof.context -> fp_sugar list val fp_sugars_of_global: theory -> fp_sugar list val fp_sugars_interpretation: string -> (fp_sugar list -> local_theory -> local_theory) -> theory -> theory val interpret_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val register_fp_sugars_raw: fp_sugar list -> local_theory -> local_theory val register_fp_sugars: (string -> bool) -> fp_sugar list -> local_theory -> local_theory val merge_type_args: BNF_Util.fp_kind -> ''a list * ''a list -> ''a list val type_args_named_constrained_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'a val type_binding_of_spec: (((('a * 'b) * 'c) * 'd) * 'e) * 'f -> 'b val mixfix_of_spec: ((('a * 'b) * 'c) * 'd) * 'e -> 'b val mixfixed_ctr_specs_of_spec: (('a * 'b) * 'c) * 'd -> 'b val map_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'b val rel_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'c val pred_binding_of_spec: ('a * ('b * 'c * 'd)) * 'e -> 'd val sel_default_eqs_of_spec: 'a * 'b -> 'b val mk_parametricity_goal: Proof.context -> term list -> term -> term -> term val flat_corec_preds_predsss_gettersss: 'a list -> 'a list list list -> 'a list list list -> 'a list val mk_ctor: typ list -> term -> term val mk_dtor: typ list -> term -> term val mk_bnf_sets: BNF_Def.bnf -> string * term list val liveness_of_fp_bnf: int -> BNF_Def.bnf -> bool list val nesting_bnfs: Proof.context -> typ list list list -> typ list -> BNF_Def.bnf list val massage_simple_notes: string -> (bstring * 'a list * (int -> 'b)) list -> ((binding * 'c list) * ('a list * 'b) list) list val massage_multi_notes: string list -> typ list -> (string * 'a list list * (string -> 'b)) list -> ((binding * 'b) * ('a list * 'c list) list) list val define_ctrs_dtrs_for_type: string -> typ -> term -> term -> thm -> thm -> int -> int list -> term -> binding list -> mixfix list -> typ list list -> local_theory -> (term list list * term list * thm * thm list) * local_theory val wrap_ctrs: (string -> bool) -> BNF_Util.fp_kind -> bool -> string -> thm -> int -> int list -> thm -> thm -> binding list -> binding list list -> term list -> term list -> thm -> thm list -> local_theory -> Ctr_Sugar.ctr_sugar * local_theory val derive_map_set_rel_pred_thms: (string -> bool) -> BNF_Util.fp_kind -> int -> typ list -> typ list -> typ -> typ -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> string -> BNF_Def.bnf -> BNF_Def.bnf list -> typ -> term -> thm -> thm -> thm -> thm list -> thm -> thm -> thm list -> thm -> thm list -> thm list -> thm list -> typ list list -> Ctr_Sugar.ctr_sugar -> local_theory -> (thm list * thm list * thm list list * thm list * thm list * thm list * thm list * thm list * thm list * thm list * thm list list list list * thm list list list list * thm list * thm list * thm list * thm list * thm list) * local_theory type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list) val morph_lfp_sugar_thms: morphism -> lfp_sugar_thms -> lfp_sugar_thms val transfer_lfp_sugar_thms: theory -> lfp_sugar_thms -> lfp_sugar_thms type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list) val morph_gfp_sugar_thms: morphism -> gfp_sugar_thms -> gfp_sugar_thms val transfer_gfp_sugar_thms: theory -> gfp_sugar_thms -> gfp_sugar_thms val mk_co_recs_prelims: Proof.context -> BNF_Util.fp_kind -> typ list list list -> typ list -> typ list -> typ list -> typ list -> int list -> int list list -> term list -> term list * (typ list list * typ list list list list * term list list * term list list list list) option * (string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list)) option val repair_nullary_single_ctr: typ list list -> typ list list val mk_corec_p_pred_types: typ list -> int list -> typ list list val mk_corec_fun_arg_types: typ list list list -> typ list -> typ list -> typ list -> int list -> int list list -> term -> typ list list * (typ list list list list * typ list list list * typ list list list list * typ list) val define_co_rec_as: BNF_Util.fp_kind -> typ list -> typ -> binding -> term -> local_theory -> (term * thm) * local_theory val define_rec: typ list list * typ list list list list * term list list * term list list list list -> (string -> binding) -> typ list -> typ list -> term list -> term -> Proof.context -> (term * thm) * Proof.context val define_corec: 'a * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> (string -> binding) -> 'b list -> typ list -> term list -> term -> local_theory -> (term * thm) * local_theory val mk_induct_raw_prem: (term -> term) -> Proof.context -> typ list list -> (string * term list) list -> term -> term -> typ list -> typ list -> term list * ((term * (term * term)) list * (int * term)) list * term val finish_induct_prem: Proof.context -> int -> term list -> term list * ((term * (term * term)) list * (int * term)) list * term -> term val mk_coinduct_prem: Proof.context -> typ list list -> typ list list -> term list -> term -> term -> term -> int -> term list -> term list list -> term list -> term list list -> typ list list -> term val mk_induct_attrs: term list list -> Token.src list val mk_coinduct_attrs: typ list -> term list list -> term list list -> int list list -> Token.src list * Token.src list val derive_induct_recs_thms_for_types: (string -> bool) -> BNF_Def.bnf list -> ('a * typ list list list list * term list list * 'b) option -> thm -> thm list -> BNF_Def.bnf list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> thm list -> thm list -> thm list -> term list list -> thm list list -> term list -> thm list -> Proof.context -> lfp_sugar_thms val derive_coinduct_thms_for_types: Proof.context -> bool -> (term -> term) -> BNF_Def.bnf list -> thm -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> (thm list * thm) list val derive_coinduct_corecs_thms_for_types: Proof.context -> BNF_Def.bnf list -> string * term list * term list list * (((term list list * term list list * term list list list list * term list list list list) * term list list list) * typ list) -> thm -> thm list -> thm list -> thm list -> BNF_Def.bnf list -> typ list -> typ list -> typ list -> typ list list list -> int list list -> int list list -> int list -> thm list -> thm list -> (thm -> thm) -> thm list list -> Ctr_Sugar.ctr_sugar list -> term list -> thm list -> gfp_sugar_thms val co_datatypes: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> Ctr_Sugar.ctr_options * ((((((binding option * (typ * sort)) list * binding) * mixfix) * ((binding, binding * typ) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * term list) list -> local_theory -> local_theory val co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * Proof.context) -> ((Proof.context -> Plugin_Name.filter) * bool) * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list -> Proof.context -> local_theory val parse_ctr_arg: (binding * string) parser val parse_ctr_specs: ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list parser val parse_spec: ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) parser val parse_co_datatype: (Ctr_Sugar.ctr_options_cmd * ((((((binding option * (string * string option)) list * binding) * mixfix) * ((binding, binding * string) Ctr_Sugar.ctr_spec * mixfix) list) * (binding * binding * binding)) * string list) list) parser val parse_co_datatype_cmd: BNF_Util.fp_kind -> (mixfix list -> binding list -> binding list -> binding list -> binding list list -> binding list -> (string * sort) list -> typ list * typ list list -> BNF_Def.bnf list -> BNF_Comp.absT_info list -> local_theory -> BNF_FP_Util.fp_result * local_theory) -> (local_theory -> local_theory) parser end; structure BNF_FP_Def_Sugar : BNF_FP_DEF_SUGAR = struct open Ctr_Sugar open BNF_FP_Rec_Sugar_Util open BNF_Util open BNF_Comp open BNF_Def open BNF_FP_Util open BNF_FP_Def_Sugar_Tactics val Eq_prefix = "Eq_"; val case_transferN = "case_transfer"; val ctor_iff_dtorN = "ctor_iff_dtor"; val ctr_transferN = "ctr_transfer"; val disc_transferN = "disc_transfer"; val sel_transferN = "sel_transfer"; val corec_codeN = "corec_code"; val corec_transferN = "corec_transfer"; val map_disc_iffN = "map_disc_iff"; val map_o_corecN = "map_o_corec"; val map_selN = "map_sel"; val pred_injectN = "pred_inject"; val rec_o_mapN = "rec_o_map"; val rec_transferN = "rec_transfer"; val set0N = "set0"; val set_casesN = "set_cases"; val set_introsN = "set_intros"; val set_inductN = "set_induct"; val set_selN = "set_sel"; type fp_ctr_sugar = {ctrXs_Tss: typ list list, ctor_iff_dtor: thm, ctr_defs: thm list, ctr_sugar: Ctr_Sugar.ctr_sugar, ctr_transfers: thm list, case_transfers: thm list, disc_transfers: thm list, sel_transfers: thm list}; type fp_bnf_sugar = {map_thms: thm list, map_disc_iffs: thm list, map_selss: thm list list, rel_injects: thm list, rel_distincts: thm list, rel_sels: thm list, rel_intros: thm list, rel_cases: thm list, pred_injects: thm list, set_thms: thm list, set_selssss: thm list list list list, set_introssss: thm list list list list, set_cases: thm list}; type fp_co_induct_sugar = {co_rec: term, common_co_inducts: thm list, co_inducts: thm list, co_rec_def: thm, co_rec_thms: thm list, co_rec_discs: thm list, co_rec_disc_iffs: thm list, co_rec_selss: thm list list, co_rec_codes: thm list, co_rec_transfers: thm list, co_rec_o_maps: thm list, common_rel_co_inducts: thm list, rel_co_inducts: thm list, common_set_inducts: thm list, set_inducts: thm list}; type fp_sugar = {T: typ, BT: typ, X: typ, fp: fp_kind, fp_res_index: int, fp_res: fp_result, pre_bnf: bnf, fp_bnf: bnf, absT_info: absT_info, fp_nesting_bnfs: bnf list, live_nesting_bnfs: bnf list, fp_ctr_sugar: fp_ctr_sugar, fp_bnf_sugar: fp_bnf_sugar, fp_co_induct_sugar: fp_co_induct_sugar option}; fun co_induct_of (i :: _) = i; fun strong_co_induct_of [_, s] = s; fun morph_fp_bnf_sugar phi ({map_thms, map_disc_iffs, map_selss, rel_injects, rel_distincts, rel_sels, rel_intros, rel_cases, pred_injects, set_thms, set_selssss, set_introssss, set_cases} : fp_bnf_sugar) = {map_thms = map (Morphism.thm phi) map_thms, map_disc_iffs = map (Morphism.thm phi) map_disc_iffs, map_selss = map (map (Morphism.thm phi)) map_selss, rel_injects = map (Morphism.thm phi) rel_injects, rel_distincts = map (Morphism.thm phi) rel_distincts, rel_sels = map (Morphism.thm phi) rel_sels, rel_intros = map (Morphism.thm phi) rel_intros, rel_cases = map (Morphism.thm phi) rel_cases, pred_injects = map (Morphism.thm phi) pred_injects, set_thms = map (Morphism.thm phi) set_thms, set_selssss = map (map (map (map (Morphism.thm phi)))) set_selssss, set_introssss = map (map (map (map (Morphism.thm phi)))) set_introssss, set_cases = map (Morphism.thm phi) set_cases}; fun morph_fp_co_induct_sugar phi ({co_rec, common_co_inducts, co_inducts, co_rec_def, co_rec_thms, co_rec_discs, co_rec_disc_iffs, co_rec_selss, co_rec_codes, co_rec_transfers, co_rec_o_maps, common_rel_co_inducts, rel_co_inducts, common_set_inducts, set_inducts} : fp_co_induct_sugar) = {co_rec = Morphism.term phi co_rec, common_co_inducts = map (Morphism.thm phi) common_co_inducts, co_inducts = map (Morphism.thm phi) co_inducts, co_rec_def = Morphism.thm phi co_rec_def, co_rec_thms = map (Morphism.thm phi) co_rec_thms, co_rec_discs = map (Morphism.thm phi) co_rec_discs, co_rec_disc_iffs = map (Morphism.thm phi) co_rec_disc_iffs, co_rec_selss = map (map (Morphism.thm phi)) co_rec_selss, co_rec_codes = map (Morphism.thm phi) co_rec_codes, co_rec_transfers = map (Morphism.thm phi) co_rec_transfers, co_rec_o_maps = map (Morphism.thm phi) co_rec_o_maps, common_rel_co_inducts = map (Morphism.thm phi) common_rel_co_inducts, rel_co_inducts = map (Morphism.thm phi) rel_co_inducts, common_set_inducts = map (Morphism.thm phi) common_set_inducts, set_inducts = map (Morphism.thm phi) set_inducts}; fun morph_fp_ctr_sugar phi ({ctrXs_Tss, ctor_iff_dtor, ctr_defs, ctr_sugar, ctr_transfers, case_transfers, disc_transfers, sel_transfers} : fp_ctr_sugar) = {ctrXs_Tss = map (map (Morphism.typ phi)) ctrXs_Tss, ctor_iff_dtor = Morphism.thm phi ctor_iff_dtor, ctr_defs = map (Morphism.thm phi) ctr_defs, ctr_sugar = morph_ctr_sugar phi ctr_sugar, ctr_transfers = map (Morphism.thm phi) ctr_transfers, case_transfers = map (Morphism.thm phi) case_transfers, disc_transfers = map (Morphism.thm phi) disc_transfers, sel_transfers = map (Morphism.thm phi) sel_transfers}; fun morph_fp_sugar phi ({T, BT, X, fp, fp_res, fp_res_index, pre_bnf, fp_bnf, absT_info, fp_nesting_bnfs, live_nesting_bnfs, fp_ctr_sugar, fp_bnf_sugar, fp_co_induct_sugar} : fp_sugar) = {T = Morphism.typ phi T, BT = Morphism.typ phi BT, X = Morphism.typ phi X, fp = fp, fp_res = morph_fp_result phi fp_res, fp_res_index = fp_res_index, pre_bnf = morph_bnf phi pre_bnf, fp_bnf = morph_bnf phi fp_bnf, absT_info = morph_absT_info phi absT_info, fp_nesting_bnfs = map (morph_bnf phi) fp_nesting_bnfs, live_nesting_bnfs = map (morph_bnf phi) live_nesting_bnfs, fp_ctr_sugar = morph_fp_ctr_sugar phi fp_ctr_sugar, fp_bnf_sugar = morph_fp_bnf_sugar phi fp_bnf_sugar, fp_co_induct_sugar = Option.map (morph_fp_co_induct_sugar phi) fp_co_induct_sugar}; val transfer_fp_sugar = morph_fp_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = fp_sugar Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun fp_sugar_of_generic context = Option.map (transfer_fp_sugar (Context.theory_of context)) o Symtab.lookup (Data.get context); fun fp_sugars_of_generic context = Symtab.fold (cons o transfer_fp_sugar (Context.theory_of context) o snd) (Data.get context) []; val fp_sugar_of = fp_sugar_of_generic o Context.Proof; val fp_sugar_of_global = fp_sugar_of_generic o Context.Theory; val fp_sugars_of = fp_sugars_of_generic o Context.Proof; val fp_sugars_of_global = fp_sugars_of_generic o Context.Theory; structure FP_Sugar_Plugin = Plugin(type T = fp_sugar list); fun fp_sugars_interpretation name f = FP_Sugar_Plugin.interpretation name (fn fp_sugars => fn lthy => f (map (transfer_fp_sugar (Proof_Context.theory_of lthy)) fp_sugars) lthy); val interpret_fp_sugars = FP_Sugar_Plugin.data; val register_fp_sugars_raw = fold (fn fp_sugar as {T = Type (s, _), ...} => Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Data.map (Symtab.update (s, morph_fp_sugar phi fp_sugar)))); fun register_fp_sugars plugins fp_sugars = register_fp_sugars_raw fp_sugars #> interpret_fp_sugars plugins fp_sugars; fun interpret_bnfs_register_fp_sugars plugins Ts BTs Xs fp pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars co_recs co_rec_defs map_thmss common_co_inducts co_inductss co_rec_thmss co_rec_discss co_rec_selsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss co_rec_disc_iffss co_rec_codess co_rec_transferss common_rel_co_inducts rel_co_inductss common_set_inducts set_inductss co_rec_o_mapss noted = let val fp_sugars = map_index (fn (kk, T) => {T = T, BT = nth BTs kk, X = nth Xs kk, fp = fp, fp_res = fp_res, fp_res_index = kk, pre_bnf = nth pre_bnfs kk, absT_info = nth absT_infos kk, fp_bnf = nth (#bnfs fp_res) kk, fp_nesting_bnfs = fp_nesting_bnfs, live_nesting_bnfs = live_nesting_bnfs, fp_ctr_sugar = {ctrXs_Tss = nth ctrXs_Tsss kk, ctor_iff_dtor = nth ctor_iff_dtors kk, ctr_defs = nth ctr_defss kk, ctr_sugar = nth ctr_sugars kk, ctr_transfers = nth ctr_transferss kk, case_transfers = nth case_transferss kk, disc_transfers = nth disc_transferss kk, sel_transfers = nth sel_transferss kk}, fp_bnf_sugar = {map_thms = nth map_thmss kk, map_disc_iffs = nth map_disc_iffss kk, map_selss = nth map_selsss kk, rel_injects = nth rel_injectss kk, rel_distincts = nth rel_distinctss kk, rel_sels = nth rel_selss kk, rel_intros = nth rel_intross kk, rel_cases = nth rel_casess kk, pred_injects = nth pred_injectss kk, set_thms = nth set_thmss kk, set_selssss = nth set_selsssss kk, set_introssss = nth set_introsssss kk, set_cases = nth set_casess kk}, fp_co_induct_sugar = SOME {co_rec = nth co_recs kk, common_co_inducts = common_co_inducts, co_inducts = nth co_inductss kk, co_rec_def = nth co_rec_defs kk, co_rec_thms = nth co_rec_thmss kk, co_rec_discs = nth co_rec_discss kk, co_rec_disc_iffs = nth co_rec_disc_iffss kk, co_rec_selss = nth co_rec_selsss kk, co_rec_codes = nth co_rec_codess kk, co_rec_transfers = nth co_rec_transferss kk, co_rec_o_maps = nth co_rec_o_mapss kk, common_rel_co_inducts = common_rel_co_inducts, rel_co_inducts = nth rel_co_inductss kk, common_set_inducts = common_set_inducts, set_inducts = nth set_inductss kk}} |> morph_fp_sugar (substitute_noted_thm noted)) Ts; in register_fp_sugars_raw fp_sugars #> fold (interpret_bnf plugins) (#bnfs fp_res) #> interpret_fp_sugars plugins fp_sugars end; fun quasi_unambiguous_case_names names = let val ps = map (`Long_Name.base_name) names; val dups = Library.duplicates (op =) (map fst ps); fun underscore s = let val ss = Long_Name.explode s in space_implode "_" (drop (length ss - 2) ss) end; in map (fn (base, full) => if member (op =) dups base then underscore full else base) ps |> Name.variant_list [] end; fun zipper_map f = let fun zed _ [] = [] | zed xs (y :: ys) = f (xs, y, ys) :: zed (xs @ [y]) ys; in zed [] end; fun cannot_merge_types fp = error ("Mutually " ^ co_prefix fp ^ "recursive types must have the same type parameters"); fun merge_type_arg fp T T' = if T = T' then T else cannot_merge_types fp; fun merge_type_args fp (As, As') = if length As = length As' then map2 (merge_type_arg fp) As As' else cannot_merge_types fp; fun type_args_named_constrained_of_spec (((((ncAs, _), _), _), _), _) = ncAs; fun type_binding_of_spec (((((_, b), _), _), _), _) = b; fun mixfix_of_spec ((((_, mx), _), _), _) = mx; fun mixfixed_ctr_specs_of_spec (((_, mx_ctr_specs), _), _) = mx_ctr_specs; fun map_binding_of_spec ((_, (b, _, _)), _) = b; fun rel_binding_of_spec ((_, (_, b, _)), _) = b; fun pred_binding_of_spec ((_, (_, _, b)), _) = b; fun sel_default_eqs_of_spec (_, ts) = ts; fun ctr_sugar_kind_of_fp_kind Least_FP = Datatype | ctr_sugar_kind_of_fp_kind Greatest_FP = Codatatype; fun uncurry_thm 0 thm = thm | uncurry_thm 1 thm = thm | uncurry_thm n thm = rotate_prems ~1 (uncurry_thm (n - 1) (rotate_prems 1 (conjI RS thm))); fun choose_binary_fun fs AB = find_first (fastype_of #> binder_types #> (fn [A, B] => AB = (A, B))) fs; fun build_binary_fun_app fs t u = Option.map (rapp u o rapp t) (choose_binary_fun fs (fastype_of t, fastype_of u)); fun build_the_rel ctxt Rs Ts A B = build_rel [] ctxt Ts [] (the o choose_binary_fun Rs) (A, B); fun build_rel_app ctxt Rs Ts t u = build_the_rel ctxt Rs Ts (fastype_of t) (fastype_of u) $ t $ u; fun build_set_app ctxt A t = Term.betapply (build_set ctxt A (fastype_of t), t); fun mk_parametricity_goal ctxt Rs t u = let val prem = build_the_rel ctxt Rs [] (fastype_of t) (fastype_of u) in HOLogic.mk_Trueprop (prem $ t $ u) end; val name_of_set = name_of_const "set function" domain_type; val fundefcong_attrs = @{attributes [fundef_cong]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; val lists_bmoc = fold (fn xs => fn t => Term.list_comb (t, xs)); fun flat_corec_predss_getterss qss gss = maps (op @) (qss ~~ gss); fun flat_corec_preds_predsss_gettersss [] [qss] [gss] = flat_corec_predss_getterss qss gss | flat_corec_preds_predsss_gettersss (p :: ps) (qss :: qsss) (gss :: gsss) = p :: flat_corec_predss_getterss qss gss @ flat_corec_preds_predsss_gettersss ps qsss gsss; fun mk_flip (x, Type (_, [T1, Type (_, [T2, T3])])) = Abs ("x", T1, Abs ("y", T2, Var (x, T2 --> T1 --> T3) $ Bound 0 $ Bound 1)); fun flip_rels ctxt n thm = let val Rs = Term.add_vars (Thm.prop_of thm) []; val Rs' = rev (drop (length Rs - n) Rs); in infer_instantiate ctxt (map (fn f => (fst f, Thm.cterm_of ctxt (mk_flip f))) Rs') thm end; fun mk_ctor_or_dtor get_T Ts t = let val Type (_, Ts0) = get_T (fastype_of t) in Term.subst_atomic_types (Ts0 ~~ Ts) t end; val mk_ctor = mk_ctor_or_dtor range_type; val mk_dtor = mk_ctor_or_dtor domain_type; fun mk_bnf_sets bnf = let val Type (T_name, Us) = T_of_bnf bnf; val lives = lives_of_bnf bnf; val sets = sets_of_bnf bnf; fun mk_set U = (case find_index (curry (op =) U) lives of ~1 => Term.dummy | i => nth sets i); in (T_name, map mk_set Us) end; fun mk_xtor_co_recs thy fp fpTs Cs ts0 = let val nn = length fpTs; val (fpTs0, Cs0) = map ((fp = Greatest_FP ? swap) o dest_funT o snd o strip_typeN nn o fastype_of) ts0 |> split_list; val rho = tvar_subst thy (fpTs0 @ Cs0) (fpTs @ Cs); in map (Term.subst_TVars rho) ts0 end; fun liveness_of_fp_bnf n bnf = (case T_of_bnf bnf of Type (_, Ts) => map (not o member (op =) (deads_of_bnf bnf)) Ts | _ => replicate n false); fun add_nesting_bnf_names Us = let fun add (Type (s, Ts)) ss = let val (needs, ss') = fold_map add Ts ss in if exists I needs then (true, insert (op =) s ss') else (false, ss') end | add T ss = (member (op =) Us T, ss); in snd oo add end; fun nesting_bnfs ctxt ctr_Tsss Us = map_filter (bnf_of ctxt) (fold (fold (fold (add_nesting_bnf_names Us))) ctr_Tsss []); fun indexify proj xs f p = f (find_index (curry (op =) (proj p)) xs) p; fun massage_simple_notes base = filter_out (null o #2) #> map (fn (thmN, thms, f_attrs) => ((Binding.qualify true base (Binding.name thmN), []), map_index (fn (i, thm) => ([thm], f_attrs i)) thms)); fun massage_multi_notes b_names Ts = maps (fn (thmN, thmss, attrs) => @{map 3} (fn b_name => fn Type (T_name, _) => fn thms => ((Binding.qualify true b_name (Binding.name thmN), attrs T_name), [(thms, [])])) b_names Ts thmss) #> filter_out (null o fst o hd o snd); fun define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy = let val ctor_absT = domain_type (fastype_of ctor); val (((w, xss), u'), _) = lthy |> yield_singleton (mk_Frees "w") ctor_absT ||>> mk_Freess "x" ctr_Tss ||>> yield_singleton Variable.variant_fixes fp_b_name; val u = Free (u', fpT); val ctor_iff_dtor_thm = let val goal = fold_rev Logic.all [w, u] (mk_Trueprop_eq (HOLogic.mk_eq (u, ctor $ w), HOLogic.mk_eq (dtor $ u, w))); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_ctor_iff_dtor_tac ctxt (map (SOME o Thm.ctyp_of lthy) [ctor_absT, fpT]) (Thm.cterm_of lthy ctor) (Thm.cterm_of lthy dtor) ctor_dtor dtor_ctor) |> Thm.close_derivation \<^here> end; val ctr_rhss = map2 (fn k => fn xs => fold_rev Term.lambda xs (ctor $ mk_absumprod ctor_absT abs n k xs)) ks xss; val ((raw_ctrs, raw_ctr_defs), (lthy, lthy_old)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn b => fn mx => fn rhs => Local_Theory.define ((b, mx), ((Thm.make_def_binding (Config.get lthy bnf_internals) b, []), rhs)) #>> apsnd snd) ctr_bindings ctr_mixfixes ctr_rhss ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val ctr_defs = map (Morphism.thm phi) raw_ctr_defs; val ctrs0 = map (Morphism.term phi) raw_ctrs; in ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) end; fun wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs lthy = let val sumEN_thm' = unfold_thms lthy @{thms unit_all_eq1} (mk_absumprodE type_definition ms); fun exhaust_tac {context = ctxt, prems = _} = mk_exhaust_tac ctxt n ctr_defs ctor_iff_dtor_thm sumEN_thm'; val inject_tacss = map2 (fn ctr_def => fn 0 => [] | _ => [fn {context = ctxt, ...} => mk_inject_tac ctxt ctr_def ctor_inject abs_inject]) ctr_defs ms; val half_distinct_tacss = map (map (fn (def, def') => fn {context = ctxt, ...} => mk_half_distinct_tac ctxt ctor_inject abs_inject [def, def'])) (mk_half_pairss (`I ctr_defs)); val tacss = [exhaust_tac] :: inject_tacss @ half_distinct_tacss; fun ctr_spec_of disc_b ctr0 sel_bs = ((disc_b, ctr0), sel_bs); val ctr_specs = @{map 3} ctr_spec_of disc_bindings ctrs0 sel_bindingss; val (ctr_sugar as {case_cong, ...}, lthy) = free_constructors (ctr_sugar_kind_of_fp_kind fp) tacss ((((plugins, discs_sels), standard_binding), ctr_specs), sel_default_eqs) lthy; val anonymous_notes = [([case_cong], fundefcong_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = if Config.get lthy bnf_internals then [(ctor_iff_dtorN, [ctor_iff_dtor_thm], K [])] |> massage_simple_notes fp_b_name else []; in (ctr_sugar, lthy |> Local_Theory.notes (anonymous_notes @ notes) |> snd) end; fun derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps fp_nested_rel_eq_onps fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm extra_unfolds_map extra_unfolds_set extra_unfolds_rel ctr_Tss ({casex, case_thms, discs, selss, sel_defs, ctrs, exhaust, exhaust_discs, disc_thmss, sel_thmss, injects, distincts, distinct_discsss, ...} : ctr_sugar) lthy = let val n = length ctr_Tss; val ms = map length ctr_Tss; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpBT = B_ify_T fpT; val live_AsBs = filter (op <>) (As ~~ Bs); val live_As = map fst live_AsBs; val fTs = map (op -->) live_AsBs; val ((((((((xss, yss), fs), Ps), Rs), ta), tb), thesis), names_lthy) = lthy |> fold (fold Variable.declare_typ) [As, Bs] |> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" (map (map B_ify_T) ctr_Tss) ||>> mk_Frees "f" fTs ||>> mk_Frees "P" (map mk_pred1T live_As) ||>> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> yield_singleton (mk_Frees "a") fpT ||>> yield_singleton (mk_Frees "b") fpBT ||>> apfst HOLogic.mk_Trueprop o yield_singleton (mk_Frees "thesis") HOLogic.boolT; val ctrAs = map (mk_ctr As) ctrs; val ctrBs = map (mk_ctr Bs) ctrs; val ctr_defs' = map2 (fn m => fn def => mk_unabs_def m (HOLogic.mk_obj_eq def)) ms ctr_defs; val ABfs = live_AsBs ~~ fs; fun derive_rel_case relAsBs rel_inject_thms rel_distinct_thms = let val rel_Rs_a_b = list_comb (relAsBs, Rs) $ ta $ tb; fun mk_assms ctrA ctrB ctxt = let val argA_Ts = binder_types (fastype_of ctrA); val argB_Ts = binder_types (fastype_of ctrB); val ((argAs, argBs), names_ctxt) = ctxt |> mk_Frees "x" argA_Ts ||>> mk_Frees "y" argB_Ts; val ctrA_args = list_comb (ctrA, argAs); val ctrB_args = list_comb (ctrB, argBs); in (fold_rev Logic.all (argAs @ argBs) (Logic.list_implies (mk_Trueprop_eq (ta, ctrA_args) :: mk_Trueprop_eq (tb, ctrB_args) :: map2 (HOLogic.mk_Trueprop oo build_rel_app lthy Rs []) argAs argBs, thesis)), names_ctxt) end; val (assms, names_lthy) = @{fold_map 2} mk_assms ctrAs ctrBs names_lthy; val goal = Logic.list_implies (HOLogic.mk_Trueprop rel_Rs_a_b :: assms, thesis); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_rel_case_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust injects rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; fun derive_case_transfer rel_case_thm = let val (S, names_lthy) = yield_singleton (mk_Frees "S") (mk_pred2T C E) names_lthy; val caseA = mk_case As C casex; val caseB = mk_case Bs E casex; val goal = mk_parametricity_goal names_lthy (S :: Rs) caseA caseB; in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_case_transfer_tac ctxt rel_case_thm case_thms) |> singleton (Proof_Context.export names_lthy lthy) |> Thm.close_derivation \<^here> end; in if live = 0 then if plugins transfer_plugin then let val relAsBs = HOLogic.eq_const fpT; val rel_case_thm = derive_rel_case relAsBs [] []; val case_transfer_thm = derive_case_transfer rel_case_thm; val notes = [(case_transferN, [case_transfer_thm], K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> Local_Theory.notes notes; val subst = Morphism.thm (substitute_noted_thm noted); in (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [subst case_transfer_thm], [], []), lthy') end else (([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []), lthy) else let val mapx = mk_map live As Bs (map_of_bnf fp_bnf); val relAsBs = mk_rel live As Bs (rel_of_bnf fp_bnf); val setAs = map (mk_set As) (sets_of_bnf fp_bnf); val discAs = map (mk_disc_or_sel As) discs; val discBs = map (mk_disc_or_sel Bs) discs; val selAss = map (map (mk_disc_or_sel As)) selss; val selBss = map (map (mk_disc_or_sel Bs)) selss; val map_ctor_thm = if fp = Least_FP then fp_map_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val (y as Free (y_s, _), _) = lthy |> yield_singleton (mk_Frees "y") y_T; val ctor_cong = infer_instantiate' lthy [NONE, NONE, SOME (Thm.cterm_of lthy ctorB)] arg_cong; val fp_map_thm' = fp_map_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y))]) |> unfold_thms lthy [dtor_ctor]; in (fp_map_thm' RS ctor_cong RS (ctor_dtor RS sym RS trans)) |> Drule.generalize (Names.empty, Names.make_set [y_s]) end; val map_thms = let fun mk_goal ctrA ctrB xs ys = let val fmap = list_comb (mapx, fs); fun mk_arg (x as Free (_, T)) (Free (_, U)) = if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x; val xs' = map2 mk_arg xs ys; in mk_Trueprop_eq (fmap $ list_comb (ctrA, xs), list_comb (ctrB, xs')) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_tac ctxt abs_inverses pre_map_def map_ctor_thm live_nesting_map_id0s ctr_defs' extra_unfolds_map) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val set0_thms = let fun mk_goal A setA ctrA xs = let val sets = map (build_set_app lthy A) (filter (exists_subtype_in [A] o fastype_of) xs); in mk_Trueprop_eq (setA $ list_comb (ctrA, xs), (if null sets then HOLogic.mk_set A [] else Library.foldl1 mk_union sets)) end; val goals = @{map 2} (fn live_A => fn setA => map2 (mk_goal live_A setA) ctrAs xss) live_As setAs |> flat; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set0_tac ctxt abs_inverses pre_set_defs dtor_ctor fp_set_thms fp_nesting_set_maps live_nesting_set_maps ctr_defs' extra_unfolds_set) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val set_thms = set0_thms |> map (unfold_thms lthy @{thms insert_is_Un[THEN sym] Un_empty_left Un_insert_left}); val rel_ctor_thm = if fp = Least_FP then fp_rel_thm else let val ctorA = mk_ctor As ctor; val ctorB = mk_ctor Bs ctor; val y_T = domain_type (fastype_of ctorA); val z_T = domain_type (fastype_of ctorB); val ((y as Free (y_s, _), z as Free (z_s, _)), _) = lthy |> yield_singleton (mk_Frees "y") y_T ||>> yield_singleton (mk_Frees "z") z_T; in fp_rel_thm |> infer_instantiate' lthy (replicate live NONE @ [SOME (Thm.cterm_of lthy (ctorA $ y)), SOME (Thm.cterm_of lthy (ctorB $ z))]) |> unfold_thms lthy [dtor_ctor] |> Drule.generalize (Names.empty, Names.make_set [y_s, z_s]) end; val rel_inject_thms = let fun mk_goal ctrA ctrB xs ys = let val lhs = list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys); val conjuncts = map2 (build_rel_app lthy Rs []) xs ys; in HOLogic.mk_Trueprop (if null conjuncts then lhs else HOLogic.mk_eq (lhs, Library.foldr1 HOLogic.mk_conj conjuncts)) end; val goals = @{map 4} mk_goal ctrAs ctrBs xss yss; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val half_rel_distinct_thmss = let fun mk_goal ((ctrA, xs), (ctrB, ys)) = HOLogic.mk_Trueprop (HOLogic.mk_not (list_comb (relAsBs, Rs) $ list_comb (ctrA, xs) $ list_comb (ctrB, ys))); val rel_infos = (ctrAs ~~ xss, ctrBs ~~ yss); val goalss = map (map mk_goal) (mk_half_pairss rel_infos); val goals = flat goalss; in unflat goalss (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_tac ctxt abs_inverses pre_rel_def rel_ctor_thm live_nesting_rel_eqs ctr_defs' extra_unfolds_rel) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_flip = rel_flip_of_bnf fp_bnf; fun mk_other_half_rel_distinct_thm thm = flip_rels lthy live thm RS (rel_flip RS sym RS @{thm arg_cong[of _ _ Not]} RS iffD2); val other_half_rel_distinct_thmss = map (map mk_other_half_rel_distinct_thm) half_rel_distinct_thmss; val (rel_distinct_thms, _) = join_halves n half_rel_distinct_thmss other_half_rel_distinct_thmss; fun mk_rel_intro_thm m thm = uncurry_thm m (thm RS iffD2) handle THM _ => thm; val rel_intro_thms = map2 mk_rel_intro_thm ms rel_inject_thms; val rel_code_thms = map (fn thm => thm RS @{thm eq_False[THEN iffD2]}) rel_distinct_thms @ map2 (fn thm => fn 0 => thm RS @{thm eq_True[THEN iffD2]} | _ => thm) rel_inject_thms ms; val ctr_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) ctrAs ctrBs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_ctr_transfer_tac ctxt rel_intro_thms live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; val (set_cases_thms, set_cases_attrss) = let fun mk_prems assms elem t ctxt = (case fastype_of t of Type (type_name, xs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val new_var = not (T = fastype_of elem); val (x, ctxt') = if new_var then yield_singleton (mk_Frees "x") T ctxt else (elem, ctxt); in mk_prems (mk_Trueprop_mem (x, set $ t) :: assms) elem x ctxt' |>> map (new_var ? Logic.all x) end) (map (mk_set xs) (sets_of_bnf bnf)) ctxt)) | T => rpair ctxt (if T = fastype_of elem then [fold (curry Logic.mk_implies) assms thesis] else [])); in split_list (map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)); val (elem, names_lthy) = yield_singleton (mk_Frees "e") A names_lthy; val premss = map (fn ctr => let val (args, names_lthy) = mk_Frees "z" (binder_types (fastype_of ctr)) names_lthy; in flat (zipper_map (fn (prev_args, arg, next_args) => let val (args_with_elem, args_without_elem) = if fastype_of arg = A then (prev_args @ [elem] @ next_args, prev_args @ next_args) else `I (prev_args @ [arg] @ next_args); in mk_prems [mk_Trueprop_eq (ta, Term.list_comb (ctr, args_with_elem))] elem arg names_lthy |> fst |> map (fold_rev Logic.all args_without_elem) end) args) end) ctrAs; val goal = Logic.mk_implies (mk_Trueprop_mem (elem, set $ ta), thesis); val vars = Variable.add_free_names lthy goal []; val thm = Goal.prove_sorry lthy vars (flat premss) goal (fn {context = ctxt, prems} => mk_set_cases_tac ctxt (Thm.cterm_of ctxt ta) prems exhaust set_thms) |> Thm.close_derivation \<^here> |> rotate_prems ~1; val cases_set_attr = Attrib.internal (K (Induct.cases_pred (name_of_set set))); val ctr_names = quasi_unambiguous_case_names (flat (map (uncurry mk_names o map_prod length name_of_ctr) (premss ~~ ctrAs))); in (* TODO: @{attributes [elim?]} *) (thm, [Attrib.consumes 1, cases_set_attr, Attrib.case_names ctr_names]) end) setAs) end; val (set_intros_thmssss, set_intros_thms) = let fun mk_goals A setA ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => apfst flat (fold_map (fn set => fn ctxt => let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (y, ctxt') = yield_singleton (mk_Frees "y") T ctxt; val assm = mk_Trueprop_mem (y, set $ t); in apfst (map (Logic.mk_implies o pair assm)) (mk_goals A setA ctr_args y ctxt') end) (map (mk_set innerTs) (sets_of_bnf bnf)) ctxt)) | T => (if T = A then [mk_Trueprop_mem (t, setA $ ctr_args)] else [], ctxt)); val (goalssss, _) = fold_map (fn set => let val A = HOLogic.dest_setT (range_type (fastype_of set)) in @{fold_map 2} (fn ctr => fn xs => fold_map (mk_goals A set (Term.list_comb (ctr, xs))) xs) ctrAs xss end) setAs lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_intros_tac ctxt set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val rel_sel_thms = let val n = length discAs; fun mk_conjunct n k discA selAs discB selBs = (if k = n then [] else [HOLogic.mk_eq (discA $ ta, discB $ tb)]) @ (if null selAs then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA $ ta, discB $ tb], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app names_lthy Rs []) (map (rapp ta) selAs) (map (rapp tb) selBs)))]); val goals = if n = 0 then [] else [mk_Trueprop_eq (build_rel_app names_lthy Rs [] ta tb, (case flat (@{map 5} (mk_conjunct n) (1 upto n) discAs selAss discBs selBss) of [] => \<^term>\True\ | conjuncts => Library.foldr1 HOLogic.mk_conj conjuncts))]; fun prove goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rel_sel_tac ctxt (Thm.cterm_of ctxt ta) (Thm.cterm_of ctxt tb) exhaust (flat disc_thmss) (flat sel_thmss) rel_inject_thms distincts rel_distinct_thms live_nesting_rel_eqs)) |> Thm.close_derivation \<^here>; in map prove goals end; val (rel_case_thm, rel_case_attrs) = let val thm = derive_rel_case relAsBs rel_inject_thms rel_distinct_thms; val ctr_names = quasi_unambiguous_case_names (map name_of_ctr ctrAs); in (thm, [Attrib.case_names ctr_names, Attrib.consumes 1] @ @{attributes [cases pred]}) end; val case_transfer_thm = derive_case_transfer rel_case_thm; val sel_transfer_thms = if null selAss then [] else let val shared_sels = foldl1 (uncurry (inter (op =))) (map (op ~~) (selAss ~~ selBss)); val goals = map (uncurry (mk_parametricity_goal names_lthy Rs)) shared_sels; in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_sel_transfer_tac ctxt n sel_defs case_transfer_thm) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val disc_transfer_thms = let val goals = map2 (mk_parametricity_goal names_lthy Rs) discAs discBs in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_disc_transfer_tac ctxt (the_single rel_sel_thms) (the_single exhaust_discs) (flat (flat distinct_discsss))) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val map_disc_iff_thms = let val discsB = map (mk_disc_or_sel Bs) discs; val discsA_t = map (fn disc1 => Term.betapply (disc1, ta)) discAs; fun mk_goal (discA_t, discB) = if head_of discA_t aconv HOLogic.Not orelse is_refl_bool discA_t then NONE else SOME (mk_Trueprop_eq (betapply (discB, (Term.list_comb (mapx, fs) $ ta)), discA_t)); val goals = map_filter mk_goal (discsA_t ~~ discsB); in if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_disc_iff_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end end; val (map_sel_thmss, map_sel_thms) = let fun mk_goal discA selA selB = let val prem = Term.betapply (discA, ta); val lhs = selB $ (Term.list_comb (mapx, fs) $ ta); val lhsT = fastype_of lhs; val map_rhsT = map_atyps (perhaps (AList.lookup (op =) (map swap live_AsBs))) lhsT; val map_rhs = build_map lthy [] [] (the o (AList.lookup (op =) (live_AsBs ~~ fs))) (map_rhsT, lhsT); val rhs = (case map_rhs of Const (\<^const_name>\id\, _) => selA $ ta | _ => map_rhs $ (selA $ ta)); val concl = mk_Trueprop_eq (lhs, rhs); in if is_refl_bool prem then concl else Logic.mk_implies (HOLogic.mk_Trueprop prem, concl) end; val goalss = @{map 3} (map2 o mk_goal) discAs selAss selBss; val goals = flat goalss; in `(unflat goalss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_map_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) map_thms (flat sel_thmss) live_nesting_map_id0s) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val (set_sel_thmssss, set_sel_thms) = let fun mk_goal setA discA selA ctxt = let val prem = Term.betapply (discA, ta); val sel_rangeT = range_type (fastype_of selA); val A = HOLogic.dest_setT (range_type (fastype_of setA)); fun travese_nested_types t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val T = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") T ctxt; val assm = mk_Trueprop_mem (x, set $ a); in travese_nested_types x ctxt' |>> map (Logic.mk_implies o pair assm) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([mk_Trueprop_mem (t, setA $ ta)], ctxt) else ([], ctxt)); val (concls, ctxt') = if sel_rangeT = A then ([mk_Trueprop_mem (selA $ ta, setA $ ta)], ctxt) else travese_nested_types (selA $ ta) ctxt; in if exists_subtype_in [A] sel_rangeT then if is_refl_bool prem then (concls, ctxt') else (map (Logic.mk_implies o pair (HOLogic.mk_Trueprop prem)) concls, ctxt') else ([], ctxt) end; val (goalssss, _) = fold_map (fn set => @{fold_map 2} (fold_map o mk_goal set) discAs selAss) setAs names_lthy; val goals = flat (flat (flat goalssss)); in `(unflattt goalssss) (if null goals then [] else let val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_set_sel_tac ctxt (Thm.cterm_of ctxt ta) exhaust (flat disc_thmss) (flat sel_thmss) set0_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end) end; val pred_injects = let fun top_sweep_rewr_conv rewrs = Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) \<^context>; val rel_eq_onp_with_tops_of = Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg1_conv (top_sweep_rewr_conv @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))); val eq_onps = map rel_eq_onp_with_tops_of (map rel_eq_onp_of_bnf fp_bnfs @ fp_nesting_rel_eq_onps @ live_nesting_rel_eq_onps @ fp_nested_rel_eq_onps); val cTs = map (SOME o Thm.ctyp_of lthy) (maps (replicate 2) live_As); val cts = map (SOME o Thm.cterm_of lthy) (map mk_eq_onp Ps); val get_rhs = Thm.concl_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq #> snd; val pred_eq_onp_conj = List.foldr (fn (_, thm) => thm RS @{thm eq_onp_live_step}) @{thm refl[of True]}; fun predify_rel_inject rel_inject = let val conjuncts = try (get_rhs #> HOLogic.dest_conj) rel_inject |> the_default []; fun postproc thm = if null conjuncts then thm RS (@{thm eq_onp_same_args} RS iffD1) else @{thm box_equals} OF [thm, @{thm eq_onp_same_args}, pred_eq_onp_conj conjuncts |> unfold_thms lthy @{thms simp_thms(21)}]; in rel_inject |> Thm.instantiate' cTs cts |> Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.arg_conv (Raw_Simplifier.rewrite lthy false @{thms eq_onp_top_eq_eq[symmetric, THEN eq_reflection]}))) |> unfold_thms lthy eq_onps |> postproc |> unfold_thms lthy @{thms top_conj} end; in rel_inject_thms |> map (unfold_thms lthy [@{thm conj_assoc}]) |> map predify_rel_inject |> Proof_Context.export names_lthy lthy end; val anonymous_notes = [(rel_code_thms, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = (if Config.get lthy bnf_internals then [(set0N, set0_thms, K [])] else []) @ [(case_transferN, [case_transfer_thm], K []), (ctr_transferN, ctr_transfer_thms, K []), (disc_transferN, disc_transfer_thms, K []), (sel_transferN, sel_transfer_thms, K []), (mapN, map_thms, K (nitpicksimp_attrs @ simp_attrs)), (map_disc_iffN, map_disc_iff_thms, K simp_attrs), (map_selN, map_sel_thms, K []), (pred_injectN, pred_injects, K simp_attrs), (rel_casesN, [rel_case_thm], K rel_case_attrs), (rel_distinctN, rel_distinct_thms, K simp_attrs), (rel_injectN, rel_inject_thms, K simp_attrs), (rel_introsN, rel_intro_thms, K []), (rel_selN, rel_sel_thms, K []), (setN, set_thms, K (case_fp fp nitpicksimp_attrs [] @ simp_attrs)), (set_casesN, set_cases_thms, nth set_cases_attrss), (set_introsN, set_intros_thms, K []), (set_selN, set_sel_thms, K [])] |> massage_simple_notes fp_b_name; val (noted, lthy') = lthy |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) map_thms) |> fp = Least_FP ? uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) rel_code_thms) |> uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational) (`(single o lhs_head_of o hd) set0_thms) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (rel_code_thms @ map_thms @ set_thms)) |> Local_Theory.notes (anonymous_notes @ notes); val subst = Morphism.thm (substitute_noted_thm noted); in ((map subst map_thms, map subst map_disc_iff_thms, map (map subst) map_sel_thmss, map subst rel_inject_thms, map subst rel_distinct_thms, map subst rel_sel_thms, map subst rel_intro_thms, [subst rel_case_thm], map subst pred_injects, map subst set_thms, map (map (map (map subst))) set_sel_thmssss, map (map (map (map subst))) set_intros_thmssss, map subst set_cases_thms, map subst ctr_transfer_thms, [subst case_transfer_thm], map subst disc_transfer_thms, map subst sel_transfer_thms), lthy') end end; type lfp_sugar_thms = (thm list * thm * Token.src list) * (thm list list * Token.src list); fun morph_lfp_sugar_thms phi ((inducts, induct, induct_attrs), (recss, rec_attrs)) = ((map (Morphism.thm phi) inducts, Morphism.thm phi induct, induct_attrs), (map (map (Morphism.thm phi)) recss, rec_attrs)) : lfp_sugar_thms; val transfer_lfp_sugar_thms = morph_lfp_sugar_thms o Morphism.transfer_morphism; type gfp_sugar_thms = ((thm list * thm) list * (Token.src list * Token.src list)) * thm list list * thm list list * (thm list list * Token.src list) * (thm list list list * Token.src list); fun morph_gfp_sugar_thms phi ((coinducts_pairs, coinduct_attrs_pair), corecss, corec_discss, (corec_disc_iffss, corec_disc_iff_attrs), (corec_selsss, corec_sel_attrs)) = ((map (apfst (map (Morphism.thm phi)) o apsnd (Morphism.thm phi)) coinducts_pairs, coinduct_attrs_pair), map (map (Morphism.thm phi)) corecss, map (map (Morphism.thm phi)) corec_discss, (map (map (Morphism.thm phi)) corec_disc_iffss, corec_disc_iff_attrs), (map (map (map (Morphism.thm phi))) corec_selsss, corec_sel_attrs)) : gfp_sugar_thms; val transfer_gfp_sugar_thms = morph_gfp_sugar_thms o Morphism.transfer_morphism; fun unzip_recT (Type (\<^type_name>\prod\, [_, TFree x])) (T as Type (\<^type_name>\prod\, Ts as [_, TFree y])) = if x = y then [T] else Ts | unzip_recT _ (Type (\<^type_name>\prod\, Ts as [_, TFree _])) = Ts | unzip_recT _ T = [T]; fun mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss ctor_rec_fun_Ts = let val Css = map2 replicate ns Cs; val x_Tssss = @{map 6} (fn absT => fn repT => fn n => fn ms => fn ctr_Tss => fn ctor_rec_fun_T => map2 (map2 unzip_recT) ctr_Tss (dest_absumprodT absT repT n ms (domain_type ctor_rec_fun_T))) absTs repTs ns mss ctr_Tsss ctor_rec_fun_Ts; val x_Tsss' = map (map flat_rec_arg_args) x_Tssss; val f_Tss = map2 (map2 (curry (op --->))) x_Tsss' Css; val ((fss, xssss), _) = ctxt |> mk_Freess "f" f_Tss ||>> mk_Freessss "x" x_Tssss; in (f_Tss, x_Tssss, fss, xssss) end; fun unzip_corecT (Type (\<^type_name>\sum\, _)) T = [T] | unzip_corecT _ (Type (\<^type_name>\sum\, Ts)) = Ts | unzip_corecT _ T = [T]; (*avoid "'a itself" arguments in corecursors*) fun repair_nullary_single_ctr [[]] = [[HOLogic.unitT]] | repair_nullary_single_ctr Tss = Tss; fun mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss fun_Ts = let val ctr_Tsss' = map repair_nullary_single_ctr ctr_Tsss; val g_absTs = map range_type fun_Ts; val g_Tsss = map repair_nullary_single_ctr (@{map 5} dest_absumprodT absTs repTs ns mss g_absTs); val g_Tssss = @{map 3} (fn C => map2 (map2 (map (curry (op -->) C) oo unzip_corecT))) Cs ctr_Tsss' g_Tsss; val q_Tssss = map (map (map (fn [_] => [] | [_, T] => [mk_pred1T (domain_type T)]))) g_Tssss; in (q_Tssss, g_Tsss, g_Tssss, g_absTs) end; fun mk_corec_p_pred_types Cs ns = map2 (fn n => replicate (Int.max (0, n - 1)) o mk_pred1T) ns Cs; fun mk_corec_fun_arg_types ctr_Tsss Cs absTs repTs ns mss dtor_corec = (mk_corec_p_pred_types Cs ns, mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss (binder_fun_types (fastype_of dtor_corec))); fun mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts = let val p_Tss = mk_corec_p_pred_types Cs ns; val (q_Tssss, g_Tsss, g_Tssss, corec_types) = mk_corec_fun_arg_types0 ctr_Tsss Cs absTs repTs ns mss dtor_corec_fun_Ts; val (((((Free (x, _), cs), pss), qssss), gssss), _) = ctxt |> yield_singleton (mk_Frees "x") dummyT ||>> mk_Frees "a" Cs ||>> mk_Freess "p" p_Tss ||>> mk_Freessss "q" q_Tssss ||>> mk_Freessss "g" g_Tssss; val cpss = map2 (map o rapp) cs pss; fun build_sum_inj mk_inj = build_map ctxt [] [] (uncurry mk_inj o dest_sumT o snd); fun build_dtor_corec_arg _ [] [cg] = cg | build_dtor_corec_arg T [cq] [cg, cg'] = mk_If cq (build_sum_inj Inl_const (fastype_of cg, T) $ cg) (build_sum_inj Inr_const (fastype_of cg', T) $ cg'); val pgss = @{map 3} flat_corec_preds_predsss_gettersss pss qssss gssss; val cqssss = map2 (map o map o map o rapp) cs qssss; val cgssss = map2 (map o map o map o rapp) cs gssss; val cqgsss = @{map 3} (@{map 3} (@{map 3} build_dtor_corec_arg)) g_Tsss cqssss cgssss; in (x, cs, cpss, (((pgss, pss, qssss, gssss), cqgsss), corec_types)) end; fun mk_co_recs_prelims ctxt fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0 = let val thy = Proof_Context.theory_of ctxt; val (xtor_co_rec_fun_Ts, xtor_co_recs) = mk_xtor_co_recs thy fp fpTs Cs xtor_co_recs0 |> `(binder_fun_types o fastype_of o hd); val (recs_args_types, corecs_args_types) = if fp = Least_FP then mk_recs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (rpair NONE o SOME) else mk_corecs_args_types ctxt ctr_Tsss Cs absTs repTs ns mss xtor_co_rec_fun_Ts |> (pair NONE o SOME); in (xtor_co_recs, recs_args_types, corecs_args_types) end; fun mk_preds_getterss_join c cps absT abs cqgss = let val n = length cqgss; val ts = map2 (mk_absumprod absT abs n) (1 upto n) cqgss; in Term.lambda c (mk_IfN absT cps ts) end; fun define_co_rec_as fp Cs fpT b rhs lthy0 = let val thy = Proof_Context.theory_of lthy0; val ((cst, (_, def)), (lthy', lthy)) = lthy0 |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((b, NoSyn), ((Thm.make_def_binding (Config.get lthy0 bnf_internals) b, []), rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val cst' = mk_co_rec thy fp Cs fpT (Morphism.term phi cst); val def' = Morphism.thm phi def; in ((cst', def'), lthy') end; fun define_rec (_, _, fss, xssss) mk_binding fpTs Cs reps ctor_rec = let val nn = length fpTs; val (ctor_rec_absTs, fpT) = strip_typeN nn (fastype_of ctor_rec) |>> map domain_type ||> domain_type; in define_co_rec_as Least_FP Cs fpT (mk_binding recN) (fold_rev (fold_rev Term.lambda) fss (Term.list_comb (ctor_rec, @{map 4} (fn ctor_rec_absT => fn rep => fn fs => fn xsss => mk_case_absumprod ctor_rec_absT rep fs (map (map HOLogic.mk_tuple) xsss) (map flat_rec_arg_args xsss)) ctor_rec_absTs reps fss xssss))) end; fun define_corec (_, cs, cpss, (((pgss, _, _, _), cqgsss), f_absTs)) mk_binding fpTs Cs abss dtor_corec = let val nn = length fpTs; val fpT = range_type (snd (strip_typeN nn (fastype_of dtor_corec))); in define_co_rec_as Greatest_FP Cs fpT (mk_binding corecN) (fold_rev (fold_rev Term.lambda) pgss (Term.list_comb (dtor_corec, @{map 5} mk_preds_getterss_join cs cpss f_absTs abss cqgsss))) end; fun mk_induct_raw_prem_prems names_ctxt Xss setss_fp_nesting (x as Free (s, Type (T_name, Ts0))) (Type (_, Xs_Ts0)) = (case AList.lookup (op =) setss_fp_nesting T_name of NONE => [] | SOME raw_sets0 => let val (Xs_Ts, (Ts, raw_sets)) = filter (exists_subtype_in (flat Xss) o fst) (Xs_Ts0 ~~ (Ts0 ~~ raw_sets0)) |> split_list ||> split_list; val sets = map (mk_set Ts0) raw_sets; val (ys, names_ctxt') = names_ctxt |> mk_Frees s Ts; val xysets = map (pair x) (ys ~~ sets); val ppremss = map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) ys Xs_Ts; in flat (map2 (map o apfst o cons) xysets ppremss) end) | mk_induct_raw_prem_prems _ Xss _ (x as Free (_, Type _)) X = [([], (find_index (fn Xs => member (op =) Xs X) Xss + 1, x))] | mk_induct_raw_prem_prems _ _ _ _ _ = []; fun mk_induct_raw_prem alter_x names_ctxt Xss setss_fp_nesting p ctr ctr_Ts ctrXs_Ts = let val (xs, names_ctxt') = names_ctxt |> mk_Frees "x" ctr_Ts; val pprems = flat (map2 (mk_induct_raw_prem_prems names_ctxt' Xss setss_fp_nesting) xs ctrXs_Ts); val y = Term.list_comb (ctr, map alter_x xs); val p' = enforce_type names_ctxt domain_type (fastype_of y) p; in (xs, pprems, HOLogic.mk_Trueprop (p' $ y)) end; fun close_induct_prem_prem nn ps xs t = fold_rev Logic.all (map Free (drop (nn + length xs) (rev (Term.add_frees t (map dest_Free xs @ map_filter (try dest_Free) ps))))) t; fun finish_induct_prem_prem ctxt nn ps xs (xysets, (j, x)) = let val p' = enforce_type ctxt domain_type (fastype_of x) (nth ps (j - 1)) in close_induct_prem_prem nn ps xs (Logic.list_implies (map (fn (x', (y, set)) => mk_Trueprop_mem (y, set $ x')) xysets, HOLogic.mk_Trueprop (p' $ x))) end; fun finish_induct_prem ctxt nn ps (xs, raw_pprems, concl) = fold_rev Logic.all xs (Logic.list_implies (map (finish_induct_prem_prem ctxt nn ps xs) raw_pprems, concl)); fun mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n k udisc usels vdisc vsels ctrXs_Ts = let fun build_the_rel T Xs_T = build_rel [] ctxt [] [] (fn (T, X) => nth rs' (find_index (fn Xs => member (op =) Xs X) Xss) |> enforce_type ctxt domain_type T) (T, Xs_T) |> Term.subst_atomic_types (flat Xss ~~ flat fpTss); fun build_rel_app usel vsel Xs_T = fold rapp [usel, vsel] (build_the_rel (fastype_of usel) Xs_T); in (if k = n then [] else [HOLogic.mk_eq (udisc, vdisc)]) @ (if null usels then [] else [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [udisc, vdisc], Library.foldr1 HOLogic.mk_conj (@{map 3} build_rel_app usels vsels ctrXs_Ts))]) end; fun mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss = @{map 6} (mk_coinduct_prem_ctr_concls ctxt Xss fpTss rs' n) (1 upto n) udiscs uselss vdiscs vselss ctrXs_Tss |> flat |> Library.foldr1 HOLogic.mk_conj handle List.Empty => \<^term>\True\; fun mk_coinduct_prem ctxt Xss fpTss rs' uvr u v n udiscs uselss vdiscs vselss ctrXs_Tss = fold_rev Logic.all [u, v] (Logic.mk_implies (HOLogic.mk_Trueprop uvr, HOLogic.mk_Trueprop (mk_coinduct_prem_concl ctxt Xss fpTss rs' n udiscs uselss vdiscs vselss ctrXs_Tss))); fun postproc_co_induct ctxt nn prop prop_conj = Drule.zero_var_indexes #> `(conj_dests nn) #>> map (fn thm => Thm.permute_prems 0 ~1 (thm RS prop)) ##> (fn thm => Thm.permute_prems 0 (~ nn) (if nn = 1 then thm RS prop else funpow nn (fn thm => unfold_thms ctxt @{thms conj_assoc} (thm RS prop_conj)) thm)); fun mk_induct_attrs ctrss = let val induct_cases = quasi_unambiguous_case_names (maps (map name_of_ctr) ctrss); in [Attrib.case_names induct_cases] end; fun derive_rel_induct_thms_for_types ctxt nn fpA_Ts As Bs ctrAss ctrAs_Tsss exhausts ctor_rel_induct ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val fpB_Ts = map B_ify_T fpA_Ts; val ctrBs_Tsss = map (map (map B_ify_T)) ctrAs_Tsss; val ctrBss = map (map B_ify) ctrAss; val ((((Rs, IRs), ctrAsss), ctrBsss), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> mk_Freesss "a" ctrAs_Tsss ||>> mk_Freesss "b" ctrBs_Tsss; val prems = let fun mk_prem ctrA ctrB argAs argBs = fold_rev Logic.all (argAs @ argBs) (fold_rev (curry Logic.mk_implies) (map2 (HOLogic.mk_Trueprop oo build_rel_app names_ctxt (Rs @ IRs) fpA_Ts) argAs argBs) (HOLogic.mk_Trueprop (build_rel_app names_ctxt (Rs @ IRs) fpA_Ts (Term.list_comb (ctrA, argAs)) (Term.list_comb (ctrB, argBs))))); in flat (@{map 4} (@{map 4} mk_prem) ctrAss ctrBss ctrAsss ctrBsss) end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts) IRs)); val vars = Variable.add_free_names ctxt goal []; val rel_induct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_induct0_tac ctxt ctor_rel_induct prems (map (Thm.cterm_of ctxt) IRs) exhausts ctor_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_induct0_thm, mk_induct_attrs ctrAss) end; fun derive_induct_recs_thms_for_types plugins pre_bnfs rec_args_typess ctor_induct ctor_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss pre_abs_inverses pre_type_definitions abs_inverses ctrss ctr_defss recs rec_defs ctxt = let val ctr_Tsss = map (map (binder_types o fastype_of)) ctrss; val nn = length pre_bnfs; val ns = map length ctr_Tsss; val mss = map (map length) ctr_Tsss; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_nesting_map_ident0s = map map_ident0_of_bnf fp_nesting_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val (((ps, xsss), us'), names_ctxt) = ctxt |> mk_Frees "P" (map mk_pred1T fpTs) ||>> mk_Freesss "x" ctr_Tsss ||>> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; val setss_fp_nesting = map mk_bnf_sets fp_nesting_bnfs; val (induct_thms, induct_thm) = let val raw_premss = @{map 4} (@{map 3} o mk_induct_raw_prem I names_ctxt (map single Xs) setss_fp_nesting) ps ctrss ctr_Tsss ctrXs_Tsss; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry (op $)) ps us)); val goal = Library.foldr (Logic.list_implies o apfst (map (finish_induct_prem ctxt nn ps))) (raw_premss, concl); val vars = Variable.add_free_names ctxt goal []; val kksss = map (map (map (fst o snd) o #2)) raw_premss; val ctor_induct' = ctor_induct OF (map2 mk_absumprodE pre_type_definitions mss); val thm = Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_induct_tac ctxt nn ns mss kksss (flat ctr_defss) ctor_induct' pre_abs_inverses abs_inverses fp_nesting_set_maps pre_set_defss) |> Thm.close_derivation \<^here>; in `(conj_dests nn) thm end; val xctrss = map2 (map2 (curry Term.list_comb)) ctrss xsss; fun mk_rec_thmss (_, x_Tssss, fss, _) recs rec_defs ctor_rec_thms = let val frecs = map (lists_bmoc fss) recs; fun mk_goal frec xctr f xs fxs = fold_rev (fold_rev Logic.all) (xs :: fss) (mk_Trueprop_eq (frec $ xctr, Term.list_comb (f, fxs))); fun maybe_tick (T, U) u f = if try (fst o HOLogic.dest_prodT) U = SOME T then Term.lambda u (HOLogic.mk_prod (u, f $ u)) else f; fun build_rec (x as Free (_, T)) U = if T = U then x else let val build_simple = indexify (perhaps (try (snd o HOLogic.dest_prodT)) o snd) Cs (fn kk => fn TU => maybe_tick TU (nth us kk) (nth frecs kk)); in build_map ctxt [] [] build_simple (T, U) $ x end; val fxsss = map2 (map2 (flat_rec_arg_args oo map2 (map o build_rec))) xsss x_Tssss; val goalss = @{map 5} (@{map 4} o mk_goal) frecs xctrss fss xsss fxsss; val tacss = @{map 4} (map ooo mk_rec_tac pre_map_defs (fp_nesting_map_ident0s @ live_nesting_map_ident0s) rec_defs) ctor_rec_thms pre_abs_inverses abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss end; val rec_thmss = mk_rec_thmss (the rec_args_typess) recs rec_defs ctor_rec_thms; in ((induct_thms, induct_thm, mk_induct_attrs ctrss), (rec_thmss, nitpicksimp_attrs @ simp_attrs)) end; fun mk_coinduct_attrs fpTs ctrss discss mss = let val fp_b_names = map base_name_of_typ fpTs; fun mk_coinduct_concls ms discs ctrs = let fun mk_disc_concl disc = [name_of_disc disc]; fun mk_ctr_concl 0 _ = [] | mk_ctr_concl _ ctr = [name_of_ctr ctr]; val disc_concls = map mk_disc_concl (fst (split_last discs)) @ [[]]; val ctr_concls = map2 mk_ctr_concl ms ctrs; in flat (map2 append disc_concls ctr_concls) end; val coinduct_cases = quasi_unambiguous_case_names (map (prefix Eq_prefix) fp_b_names); val coinduct_conclss = @{map 3} (quasi_unambiguous_case_names ooo mk_coinduct_concls) mss discss ctrss; val coinduct_case_names_attr = Attrib.case_names coinduct_cases; val coinduct_case_concl_attrs = map2 (fn casex => fn concls => Attrib.case_conclusion (casex, concls)) coinduct_cases coinduct_conclss; val common_coinduct_attrs = coinduct_case_names_attr :: coinduct_case_concl_attrs; val coinduct_attrs = Attrib.consumes 1 :: coinduct_case_names_attr :: coinduct_case_concl_attrs; in (coinduct_attrs, common_coinduct_attrs) end; fun derive_rel_coinduct_thms_for_types ctxt nn fpA_Ts ns As Bs mss (ctr_sugars : ctr_sugar list) abs_inverses abs_injects ctor_injects dtor_ctors rel_pre_defs ctor_defss dtor_rel_coinduct live_nesting_rel_eqs = let val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val fpB_Ts = map B_ify_T fpA_Ts; val (Rs, IRs, fpAs, fpBs, _) = let val fp_names = map base_name_of_typ fpA_Ts; val ((((Rs, IRs), fpAs_names), fpBs_names), names_ctxt) = ctxt |> mk_Frees "R" (map2 mk_pred2T As Bs) ||>> mk_Frees "IR" (map2 mk_pred2T fpA_Ts fpB_Ts) ||>> Variable.variant_fixes fp_names ||>> Variable.variant_fixes (map (suffix "'") fp_names); in (Rs, IRs, map2 (curry Free) fpAs_names fpA_Ts, map2 (curry Free) fpBs_names fpB_Ts, names_ctxt) end; val ((discA_tss, selA_tsss), (discB_tss, selB_tsss)) = let val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; fun mk_discss ts Ts = map2 (map o rapp) ts (map (map (mk_disc_or_sel Ts)) discss); fun mk_selsss ts Ts = map2 (map o map o rapp) ts (map (map (map (mk_disc_or_sel Ts))) selsss); in ((mk_discss fpAs As, mk_selsss fpAs As), (mk_discss fpBs Bs, mk_selsss fpBs Bs)) end; val prems = let fun mk_prem_ctr_concls n k discA_t selA_ts discB_t selB_ts = (if k = n then [] else [HOLogic.mk_eq (discA_t, discB_t)]) @ (case (selA_ts, selB_ts) of ([], []) => [] | (_ :: _, _ :: _) => [Library.foldr HOLogic.mk_imp (if n = 1 then [] else [discA_t, discB_t], Library.foldr1 HOLogic.mk_conj (map2 (build_rel_app ctxt (Rs @ IRs) fpA_Ts) selA_ts selB_ts))]); fun mk_prem_concl n discA_ts selA_tss discB_ts selB_tss = Library.foldr1 HOLogic.mk_conj (flat (@{map 5} (mk_prem_ctr_concls n) (1 upto n) discA_ts selA_tss discB_ts selB_tss)) handle List.Empty => \<^term>\True\; fun mk_prem IR tA tB n discA_ts selA_tss discB_ts selB_tss = fold_rev Logic.all [tA, tB] (Logic.mk_implies (HOLogic.mk_Trueprop (IR $ tA $ tB), HOLogic.mk_Trueprop (mk_prem_concl n discA_ts selA_tss discB_ts selB_tss))); in @{map 8} mk_prem IRs fpAs fpBs ns discA_tss selA_tsss discB_tss selB_tsss end; val goal = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 mk_leq IRs (map2 (build_the_rel ctxt (Rs @ IRs) []) fpA_Ts fpB_Ts))); val vars = Variable.add_free_names ctxt goal []; val rel_coinduct0_thm = Goal.prove_sorry ctxt vars prems goal (fn {context = ctxt, prems} => mk_rel_coinduct0_tac ctxt dtor_rel_coinduct (map (Thm.cterm_of ctxt) IRs) prems (map #exhaust ctr_sugars) (map (flat o #disc_thmss) ctr_sugars) (map (flat o #sel_thmss) ctr_sugars) ctor_defss dtor_ctors ctor_injects abs_injects rel_pre_defs abs_inverses live_nesting_rel_eqs) |> Thm.close_derivation \<^here>; in (postproc_co_induct ctxt nn @{thm predicate2D} @{thm predicate2D_conj} rel_coinduct0_thm, mk_coinduct_attrs fpA_Ts (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss) end; fun derive_set_induct_thms_for_types ctxt nn fpTs ctrss setss dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses = let fun mk_prems A Ps ctr_args t ctxt = (case fastype_of t of Type (type_name, innerTs) => (case bnf_of ctxt type_name of NONE => ([], ctxt) | SOME bnf => let fun seq_assm a set ctxt = let val X = HOLogic.dest_setT (range_type (fastype_of set)); val (x, ctxt') = yield_singleton (mk_Frees "x") X ctxt; val assm = mk_Trueprop_mem (x, set $ a); in (case build_binary_fun_app Ps x a of NONE => mk_prems A Ps ctr_args x ctxt' |>> map (Logic.all x o Logic.mk_implies o pair assm) | SOME f => ([Logic.all x (Logic.mk_implies (assm, Logic.mk_implies (HOLogic.mk_Trueprop f, HOLogic.mk_Trueprop (the (build_binary_fun_app Ps x ctr_args)))))], ctxt')) end; in fold_map (seq_assm t o mk_set innerTs) (sets_of_bnf bnf) ctxt |>> flat end) | T => if T = A then ([HOLogic.mk_Trueprop (the (build_binary_fun_app Ps t ctr_args))], ctxt) else ([], ctxt)); fun mk_prems_for_ctr A Ps ctr ctxt = let val (args, ctxt') = mk_Frees "z" (binder_types (fastype_of ctr)) ctxt; in fold_map (mk_prems A Ps (list_comb (ctr, args))) args ctxt' |>> map (fold_rev Logic.all args) o flat |>> (fn prems => (prems, mk_names (length prems) (name_of_ctr ctr))) end; fun mk_prems_and_concl_for_type A Ps ((fpT, ctrs), set) ctxt = let val ((x, fp), ctxt') = ctxt |> yield_singleton (mk_Frees "x") A ||>> yield_singleton (mk_Frees "a") fpT; val concl = mk_Ball (set $ fp) (Term.absfree (dest_Free x) (the (build_binary_fun_app Ps x fp))); in fold_map (mk_prems_for_ctr A Ps) ctrs ctxt' |>> split_list |>> map_prod flat flat |>> apfst (rpair concl) end; fun mk_thm ctxt fpTs ctrss sets = let val A = HOLogic.dest_setT (range_type (fastype_of (hd sets))); val (Ps, ctxt') = mk_Frees "P" (map (fn fpT => A --> fpT --> HOLogic.boolT) fpTs) ctxt; val (((prems, concl), case_names), ctxt'') = fold_map (mk_prems_and_concl_for_type A Ps) (fpTs ~~ ctrss ~~ sets) ctxt' |>> apfst split_list o split_list |>> apfst (apfst flat) |>> apfst (apsnd (Library.foldr1 HOLogic.mk_conj)) |>> apsnd flat; val vars = fold (Variable.add_free_names ctxt) (concl :: prems) []; val thm = Goal.prove_sorry ctxt vars prems (HOLogic.mk_Trueprop concl) (fn {context = ctxt, prems} => mk_set_induct0_tac ctxt (map (Thm.cterm_of ctxt'') Ps) prems dtor_set_inducts exhausts set_pre_defs ctor_defs dtor_ctors Abs_pre_inverses) |> Thm.close_derivation \<^here>; val case_names_attr = Attrib.case_names (quasi_unambiguous_case_names case_names); val induct_set_attrs = map (Attrib.internal o K o Induct.induct_pred o name_of_set) sets; in (thm, case_names_attr :: induct_set_attrs) end val consumes_attr = Attrib.consumes 1; in map (mk_thm ctxt fpTs ctrss #> nn = 1 ? map_prod (fn thm => rotate_prems ~1 (thm RS bspec)) (cons consumes_attr)) (transpose setss) end; fun mk_coinduct_strong_thm coind rel_eqs rel_monos mk_vimage2p ctxt = let val n = Thm.nprems_of coind; val m = Thm.nprems_of (hd rel_monos) - n; fun mk_inst phi = (phi, Thm.cterm_of ctxt (mk_union (Var phi, HOLogic.eq_const (fst (dest_pred2T (#2 phi)))))); val insts = Term.add_vars (Thm.prop_of coind) [] |> rev |> take n |> map mk_inst; fun mk_unfold rel_eq rel_mono = let val eq = iffD2 OF [rel_eq RS @{thm predicate2_eqD}, refl]; val mono = rel_mono OF (replicate m @{thm order_refl} @ replicate n @{thm eq_subset}); in mk_vimage2p (eq RS (mono RS @{thm predicate2D})) RS eqTrueI end; val unfolds = map2 mk_unfold rel_eqs rel_monos @ @{thms sup_fun_def sup_bool_def imp_disjL all_conj_distrib subst_eq_imp simp_thms(18,21,35)}; in - Thm.instantiate ([], insts) coind + Thm.instantiate (TVars.empty, Vars.make insts) coind |> unfold_thms ctxt unfolds end; fun derive_coinduct_thms_for_types ctxt strong alter_r pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) = let val nn = length pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val exhausts = map #exhaust ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val (((rs, us'), vs'), _) = ctxt |> mk_Frees "R" (map (fn T => mk_pred2T T T) fpTs) ||>> Variable.variant_fixes fp_b_names ||>> Variable.variant_fixes (map (suffix "'") fp_b_names); val us = map2 (curry Free) us' fpTs; val udiscss = map2 (map o rapp) us discss; val uselsss = map2 (map o map o rapp) us selsss; val vs = map2 (curry Free) vs' fpTs; val vdiscss = map2 (map o rapp) vs discss; val vselsss = map2 (map o map o rapp) vs selsss; val uvrs = @{map 3} (fn r => fn u => fn v => r $ u $ v) rs us vs; val uv_eqs = map2 (curry HOLogic.mk_eq) us vs; val strong_rs = @{map 4} (fn u => fn v => fn uvr => fn uv_eq => fold_rev Term.lambda [u, v] (HOLogic.mk_disj (uvr, uv_eq))) us vs uvrs uv_eqs; val concl = HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (@{map 3} (fn uvr => fn u => fn v => HOLogic.mk_imp (uvr, HOLogic.mk_eq (u, v))) uvrs us vs)) fun mk_goal rs0' = Logic.list_implies (@{map 9} (mk_coinduct_prem ctxt (map single Xs) (map single fpTs) (map alter_r rs0')) uvrs us vs ns udiscss uselsss vdiscss vselsss ctrXs_Tsss, concl); val goals = map mk_goal ([rs] @ (if strong then [strong_rs] else [])); fun prove dtor_coinduct' goal = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (fn {context = ctxt, ...} => mk_coinduct_tac ctxt live_nesting_rel_eqs nn ns dtor_coinduct' pre_rel_defs pre_abs_inverses abs_inverses dtor_ctors exhausts ctr_defss disc_thmsss sel_thmsss)) |> Thm.close_derivation \<^here>; val rel_eqs = map rel_eq_of_bnf pre_bnfs; val rel_monos = map rel_mono_of_bnf pre_bnfs; val dtor_coinducts = [dtor_coinduct] @ (if strong then [mk_coinduct_strong_thm dtor_coinduct rel_eqs rel_monos mk_vimage2p ctxt] else []); in map2 (postproc_co_induct ctxt nn mp @{thm conj_commute[THEN iffD1]} oo prove) dtor_coinducts goals end; fun derive_coinduct_corecs_thms_for_types ctxt pre_bnfs (x, cs, cpss, (((pgss, _, _, _), cqgsss), _)) dtor_coinduct dtor_injects dtor_ctors dtor_corec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss (ctr_sugars : ctr_sugar list) corecs corec_defs = let fun mk_ctor_dtor_corec_thm dtor_inject dtor_ctor corec = iffD1 OF [dtor_inject, trans OF [corec, dtor_ctor RS sym]]; val ctor_dtor_corec_thms = @{map 3} mk_ctor_dtor_corec_thm dtor_injects dtor_ctors dtor_corec_thms; val pre_map_defs = map map_def_of_bnf pre_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val fp_b_names = map base_name_of_typ fpTs; val ctrss = map #ctrs ctr_sugars; val discss = map #discs ctr_sugars; val selsss = map #selss ctr_sugars; val disc_thmsss = map #disc_thmss ctr_sugars; val discIss = map #discIs ctr_sugars; val sel_thmsss = map #sel_thmss ctr_sugars; val coinduct_thms_pairs = derive_coinduct_thms_for_types ctxt true I pre_bnfs dtor_coinduct dtor_ctors live_nesting_bnfs fpTs Xs ctrXs_Tsss ns pre_abs_inverses abs_inverses mk_vimage2p ctr_defss ctr_sugars; fun mk_maybe_not pos = not pos ? HOLogic.mk_not; val gcorecs = map (lists_bmoc pgss) corecs; val corec_thmss = let val (us', _) = ctxt |> Variable.variant_fixes fp_b_names; val us = map2 (curry Free) us' fpTs; fun mk_goal c cps gcorec n k ctr m cfs' = fold_rev (fold_rev Logic.all) ([c] :: pgss) (Logic.list_implies (seq_conds (HOLogic.mk_Trueprop oo mk_maybe_not) n k cps, mk_Trueprop_eq (gcorec $ c, Term.list_comb (ctr, take m cfs')))); val mk_U = typ_subst_nonatomic (map2 (fn C => fn fpT => (mk_sumT (fpT, C), fpT)) Cs fpTs); fun tack (c, u) f = let val x' = Free (x, mk_sumT (fastype_of u, fastype_of c)) in Term.lambda x' (mk_case_sum (Term.lambda u u, Term.lambda c (f $ c)) $ x') end; fun build_corec cqg = let val T = fastype_of cqg in if exists_subtype_in Cs T then let val U = mk_U T; val build_simple = indexify fst (map2 (curry mk_sumT) fpTs Cs) (fn kk => fn _ => tack (nth cs kk, nth us kk) (nth gcorecs kk)); in build_map ctxt [] [] build_simple (T, U) $ cqg end else cqg end; val cqgsss' = map (map (map build_corec)) cqgsss; val goalss = @{map 8} (@{map 4} oooo mk_goal) cs cpss gcorecs ns kss ctrss mss cqgsss'; val tacss = @{map 4} (map ooo mk_corec_tac corec_defs live_nesting_map_ident0s) ctor_dtor_corec_thms pre_map_defs abs_inverses ctr_defss; fun prove goal tac = Goal.prove_sorry ctxt [] [] goal (tac o #context) |> Thm.close_derivation \<^here>; in map2 (map2 prove) goalss tacss |> map (map (unfold_thms ctxt @{thms case_sum_if})) end; val corec_disc_iff_thmss = let fun mk_goal c cps gcorec n k disc = mk_Trueprop_eq (disc $ (gcorec $ c), if n = 1 then \<^const>\True\ else Library.foldr1 HOLogic.mk_conj (seq_conds mk_maybe_not n k cps)); val goalss = @{map 6} (map2 oooo mk_goal) cs cpss gcorecs ns kss discss; fun mk_case_split' cp = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt cp)] @{thm case_split}; val case_splitss' = map (map mk_case_split') cpss; val tacss = @{map 3} (map oo mk_corec_disc_iff_tac) case_splitss' corec_thmss disc_thmsss; fun prove goal tac = Variable.add_free_names ctxt goal [] |> (fn vars => Goal.prove_sorry ctxt vars [] goal (tac o #context)) |> Thm.close_derivation \<^here>; fun proves [_] [_] = [] | proves goals tacs = map2 prove goals tacs; in map2 proves goalss tacss end; fun mk_corec_disc_thms corecs discIs = map (op RS) (corecs ~~ discIs); val corec_disc_thmss = map2 mk_corec_disc_thms corec_thmss discIss; fun mk_corec_sel_thm corec_thm sel sel_thm = let val (domT, ranT) = dest_funT (fastype_of sel); val arg_cong' = Thm.instantiate' (map (SOME o Thm.ctyp_of ctxt) [domT, ranT]) [NONE, NONE, SOME (Thm.cterm_of ctxt sel)] arg_cong |> Thm.varifyT_global; val sel_thm' = sel_thm RSN (2, trans); in corec_thm RS arg_cong' RS sel_thm' end; fun mk_corec_sel_thms corec_thmss = @{map 3} (@{map 3} (map2 o mk_corec_sel_thm)) corec_thmss selsss sel_thmsss; val corec_sel_thmsss = mk_corec_sel_thms corec_thmss; in ((coinduct_thms_pairs, mk_coinduct_attrs fpTs (map #ctrs ctr_sugars) (map #discs ctr_sugars) mss), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, simp_attrs), (corec_sel_thmsss, simp_attrs)) end; fun define_co_datatypes prepare_plugins prepare_constraint prepare_typ prepare_term fp construct_fp ((raw_plugins, discs_sels0), specs) lthy = let val plugins = prepare_plugins lthy raw_plugins; val discs_sels = discs_sels0 orelse fp = Greatest_FP; val nn = length specs; val fp_bs = map type_binding_of_spec specs; val fp_b_names = map Binding.name_of fp_bs; val fp_common_name = mk_common_name fp_b_names; val map_bs = map map_binding_of_spec specs; val rel_bs = map rel_binding_of_spec specs; val pred_bs = map pred_binding_of_spec specs; fun prepare_type_arg (_, (ty, c)) = let val TFree (s, _) = prepare_typ lthy ty in TFree (s, prepare_constraint lthy c) end; val Ass0 = map (map prepare_type_arg o type_args_named_constrained_of_spec) specs; val unsorted_Ass0 = map (map (resort_tfree_or_tvar \<^sort>\type\)) Ass0; val unsorted_As = Library.foldr1 (merge_type_args fp) unsorted_Ass0; val num_As = length unsorted_As; val set_boss = map (map fst o type_args_named_constrained_of_spec) specs; val set_bss = map (map (the_default Binding.empty)) set_boss; fun add_fake_type spec = Typedecl.basic_typedecl {final = true} (type_binding_of_spec spec, num_As, Mixfix.reset_pos (mixfix_of_spec spec)); val (fake_T_names, fake_lthy) = fold_map add_fake_type specs lthy; val qsoty = quote o Syntax.string_of_typ fake_lthy; val _ = (case Library.duplicates (op =) unsorted_As of [] => () | A :: _ => error ("Duplicate type parameter " ^ qsoty A ^ " in " ^ co_prefix fp ^ "datatype specification")); val bad_args = map (Logic.type_map (singleton (Variable.polymorphic lthy))) unsorted_As |> filter_out Term.is_TVar; val _ = null bad_args orelse error ("Locally fixed type argument " ^ qsoty (hd bad_args) ^ " in " ^ co_prefix fp ^ "datatype specification"); val mixfixes = map mixfix_of_spec specs; val _ = (case Library.duplicates Binding.eq_name fp_bs of [] => () | b :: _ => error ("Duplicate type name declaration " ^ quote (Binding.name_of b))); val mx_ctr_specss = map mixfixed_ctr_specs_of_spec specs; val ctr_specss = map (map fst) mx_ctr_specss; val ctr_mixfixess = map (map snd) mx_ctr_specss; val disc_bindingss = map (map disc_of_ctr_spec) ctr_specss; val ctr_bindingss = map2 (fn fp_b_name => map (Binding.qualify false fp_b_name o ctr_of_ctr_spec)) fp_b_names ctr_specss; val ctr_argsss = map (map args_of_ctr_spec) ctr_specss; val sel_bindingsss = map (map (map fst)) ctr_argsss; val fake_ctr_Tsss0 = map (map (map (prepare_typ fake_lthy o snd))) ctr_argsss; val raw_sel_default_eqss = map sel_default_eqs_of_spec specs; val (As :: _) :: fake_ctr_Tsss = burrow (burrow (Syntax.check_typs fake_lthy)) (Ass0 :: fake_ctr_Tsss0); val As' = map dest_TFree As; val rhs_As' = fold (fold (fold Term.add_tfreesT)) fake_ctr_Tsss []; val _ = (case subtract (op =) As' rhs_As' of [] => () | extras => error ("Extra type variables on right-hand side: " ^ commas (map (qsoty o TFree) extras))); val fake_Ts = map (fn s => Type (s, As)) fake_T_names; val ((((Bs0, Cs), Es), Xs), _) = lthy |> fold (Variable.declare_typ o resort_tfree_or_tvar dummyS) unsorted_As |> mk_TFrees num_As ||>> mk_TFrees nn ||>> mk_TFrees nn ||>> variant_tfrees fp_b_names; fun eq_fpT_check (T as Type (s, Ts)) (T' as Type (s', Ts')) = s = s' andalso (Ts = Ts' orelse error ("Wrong type arguments in " ^ co_prefix fp ^ "recursive type " ^ qsoty T ^ " (expected " ^ qsoty T' ^ ")")) | eq_fpT_check _ _ = false; fun freeze_fp (T as Type (s, Ts)) = (case find_index (eq_fpT_check T) fake_Ts of ~1 => Type (s, map freeze_fp Ts) | kk => nth Xs kk) | freeze_fp T = T; val unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fake_Ts); val ctrXs_Tsss = map (map (map freeze_fp)) fake_ctr_Tsss; val ctrXs_repTs = map mk_sumprodT_balanced ctrXs_Tsss; val _ = let fun add_deps i = fold (fn T => fold_index (fn (j, X) => (i <> j andalso exists_subtype_in [X] T) ? insert (op =) (i, j)) Xs); val add_missing_nodes = fold (AList.default (op =) o rpair []) (0 upto nn - 1); val deps = fold_index (uncurry (fold o add_deps)) ctrXs_Tsss [] |> AList.group (op =) |> add_missing_nodes; val G = Int_Graph.make (map (apfst (rpair ())) deps); val sccs = map (sort int_ord) (Int_Graph.strong_conn G); val str_of_scc = prefix (co_prefix fp ^ "datatype ") o space_implode " and " o map (suffix " = \" o Long_Name.base_name); fun warn [_] = () | warn sccs = warning ("Defined types not fully mutually " ^ co_prefix fp ^ "recursive\n\ \Alternative specification:\n" ^ cat_lines (map (prefix " " o str_of_scc o map (nth fp_b_names)) sccs)); in warn (order_strong_conn (op =) Int_Graph.make Int_Graph.topological_order deps sccs) end; val killed_As = map_filter (fn (A, set_bos) => if exists is_none set_bos then SOME A else NONE) (As ~~ transpose set_boss); val (((pre_bnfs, absT_infos), _), (fp_res as {bnfs = fp_bnfs as any_fp_bnf :: _, ctors = ctors0, dtors = dtors0, xtor_co_recs = xtor_co_recs0, xtor_co_induct, dtor_ctors, ctor_dtors, ctor_injects, dtor_injects, xtor_maps, xtor_setss, xtor_rels, xtor_co_rec_thms, xtor_rel_co_induct, dtor_set_inducts, xtor_co_rec_transfers, xtor_co_rec_o_maps, ...}, lthy)) = fixpoint_bnf false I (construct_fp mixfixes map_bs rel_bs pred_bs set_bss) fp_bs (map dest_TFree As) (map dest_TFree killed_As) (map dest_TFree Xs) ctrXs_repTs empty_comp_cache lthy handle BAD_DEAD (X, X_backdrop) => (case X_backdrop of Type (bad_tc, _) => let val fake_T = qsoty (unfreeze_fp X); val fake_T_backdrop = qsoty (unfreeze_fp X_backdrop); fun register_hint () = "\nUse the " ^ quote (#1 \<^command_keyword>\bnf\) ^ " command to register " ^ quote bad_tc ^ " as a bounded natural functor to allow nested (co)recursion through \ \it"; in if is_some (bnf_of lthy bad_tc) orelse is_some (fp_sugar_of lthy bad_tc) then error ("Inadmissible " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " in type expression " ^ fake_T_backdrop) else if is_some (Old_Datatype_Data.get_info (Proof_Context.theory_of lthy) bad_tc) then error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via the old-style datatype " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) else error ("Unsupported " ^ co_prefix fp ^ "recursive occurrence of type " ^ fake_T ^ " via type constructor " ^ quote bad_tc ^ " in type expression " ^ fake_T_backdrop ^ register_hint ()) end); val time = time lthy; val timer = time (Timer.startRealTimer ()); val fp_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss Xs; val live_nesting_bnfs = nesting_bnfs lthy ctrXs_Tsss As; val pre_map_defs = map map_def_of_bnf pre_bnfs; val pre_set_defss = map set_defs_of_bnf pre_bnfs; val pre_rel_defs = map rel_def_of_bnf pre_bnfs; val fp_nesting_set_maps = maps set_map_of_bnf fp_nesting_bnfs; val fp_nesting_rel_eq_onps = map rel_eq_onp_of_bnf fp_nesting_bnfs; val live_nesting_map_id0s = map map_id0_of_bnf live_nesting_bnfs; val live_nesting_map_ident0s = map map_ident0_of_bnf live_nesting_bnfs; val live_nesting_set_maps = maps set_map_of_bnf live_nesting_bnfs; val live_nesting_rel_eqs = map rel_eq_of_bnf live_nesting_bnfs; val live_nesting_rel_eq_onps = map rel_eq_onp_of_bnf live_nesting_bnfs; val liveness = liveness_of_fp_bnf num_As any_fp_bnf; val live = live_of_bnf any_fp_bnf; val _ = if live = 0 andalso exists (not o Binding.is_empty) (map_bs @ rel_bs @ pred_bs) then warning "Map function, relator, and predicator names ignored" else (); val Bs = @{map 3} (fn alive => fn A as TFree (_, S) => fn B => if alive then resort_tfree_or_tvar S B else A) liveness As Bs0; val B_ify_T = Term.typ_subst_atomic (As ~~ Bs); val B_ify = Term.map_types B_ify_T; val live_AsBs = filter (op <>) (As ~~ Bs); val abss = map #abs absT_infos; val reps = map #rep absT_infos; val absTs = map #absT absT_infos; val repTs = map #repT absT_infos; val abs_injects = map #abs_inject absT_infos; val abs_inverses = map #abs_inverse absT_infos; val type_definitions = map #type_definition absT_infos; val ctors = map (mk_ctor As) ctors0; val dtors = map (mk_dtor As) dtors0; val fpTs = map (domain_type o fastype_of) dtors; val fpBTs = map B_ify_T fpTs; val real_unfreeze_fp = Term.typ_subst_atomic (Xs ~~ fpTs); val ctr_Tsss = map (map (map real_unfreeze_fp)) ctrXs_Tsss; val ns = map length ctr_Tsss; val kss = map (fn n => 1 upto n) ns; val mss = map (map length) ctr_Tsss; val (xtor_co_recs, recs_args_types, corecs_args_types) = mk_co_recs_prelims lthy fp ctr_Tsss fpTs Cs absTs repTs ns mss xtor_co_recs0; fun define_ctrs_dtrs_for_type_etc fp_bnf fp_b fpT C E ctor dtor xtor_co_rec ctor_dtor dtor_ctor ctor_inject pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm n ks ms abs abs_inject type_definition ctr_bindings ctr_mixfixes ctr_Tss disc_bindings sel_bindingss raw_sel_default_eqs lthy = let val fp_b_name = Binding.name_of fp_b; val ((xss, ctrs0, ctor_iff_dtor_thm, ctr_defs), lthy) = define_ctrs_dtrs_for_type fp_b_name fpT ctor dtor ctor_dtor dtor_ctor n ks abs ctr_bindings ctr_mixfixes ctr_Tss lthy; val ctrs = map (mk_ctr As) ctrs0; val sel_default_eqs = let val sel_Tss = map (map (curry (op -->) fpT)) ctr_Tss; val sel_bTs = flat sel_bindingss ~~ flat sel_Tss |> filter_out (Binding.is_empty o fst) |> distinct (Binding.eq_name o apply2 fst); val sel_default_lthy = fake_local_theory_for_sel_defaults sel_bTs lthy in map (prepare_term sel_default_lthy) raw_sel_default_eqs end; fun mk_binding pre = Binding.qualify false fp_b_name (Binding.prefix_name (pre ^ "_") fp_b); fun massage_res (ctr_sugar, maps_sets_rels) = (maps_sets_rels, (ctrs, xss, ctor_iff_dtor_thm, ctr_defs, ctr_sugar)); in (wrap_ctrs plugins fp discs_sels fp_b_name ctor_inject n ms abs_inject type_definition disc_bindings sel_bindingss sel_default_eqs ctrs0 ctor_iff_dtor_thm ctr_defs #> (fn (ctr_sugar, lthy) => derive_map_set_rel_pred_thms plugins fp live As Bs C E abs_inverses ctr_defs fp_nesting_set_maps fp_nesting_rel_eq_onps live_nesting_map_id0s live_nesting_set_maps live_nesting_rel_eqs live_nesting_rel_eq_onps [] fp_b_name fp_bnf fp_bnfs fpT ctor ctor_dtor dtor_ctor pre_map_def pre_set_defs pre_rel_def fp_map_thm fp_set_thms fp_rel_thm [] [] [] ctr_Tss ctr_sugar lthy |>> pair ctr_sugar) ##>> (if fp = Least_FP then define_rec (the recs_args_types) mk_binding fpTs Cs reps else define_corec (the corecs_args_types) mk_binding fpTs Cs abss) xtor_co_rec #>> apfst massage_res, lthy) end; fun wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types (wrap_one_etcs, lthy) = fold_map I wrap_one_etcs lthy |>> apsnd split_list o apfst (apsnd @{split_list 5} o apfst @{split_list 17} o split_list) o split_list; fun mk_simp_thms ({injects, distincts, case_thms, ...} : ctr_sugar) co_recs map_thms rel_injects rel_distincts set_thmss = injects @ distincts @ case_thms @ co_recs @ map_thms @ rel_injects @ rel_distincts @ set_thmss; fun mk_co_rec_transfer_goals lthy co_recs = let val BE_ify = Term.subst_atomic_types (live_AsBs @ (Cs ~~ Es)); val ((Rs, Ss), names_lthy) = lthy |> mk_Frees "R" (map (uncurry mk_pred2T) live_AsBs) ||>> mk_Frees "S" (map2 mk_pred2T Cs Es); val co_recBs = map BE_ify co_recs; in (Rs, Ss, map2 (mk_parametricity_goal lthy (Rs @ Ss)) co_recs co_recBs, names_lthy) end; fun derive_rec_transfer_thms lthy recs rec_defs (SOME (_, _, _, xsssss)) = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy recs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_rec_transfer_tac ctxt nn ns (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) xsssss rec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced nn end; fun derive_rec_o_map_thmss lthy recs rec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_gs = AList.find (op =) (fs ~~ liveness) true; val gmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_gs)) maps0; val rec_arg_Ts = binder_fun_types (hd (map fastype_of recs)); val num_rec_args = length rec_arg_Ts; val g_Ts = map B_ify_T rec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names g_Ts; val grecs = map (fn recx => Term.list_comb (Term.map_types B_ify_T recx, gs)) recs; val rec_o_map_lhss = map2 (curry HOLogic.mk_comp) grecs gmaps; val ABfs = (As ~~ Bs) ~~ fs; fun mk_rec_arg_arg (x as Free (_, T)) = let val U = B_ify_T T in if T = U then x else build_map lthy [] [] (the o AList.lookup (op =) ABfs) (T, U) $ x end; fun mk_rec_o_map_arg rec_arg_T h = let val x_Ts = binder_types rec_arg_T; val m = length x_Ts; val x_names = variant_names m "x"; val xs = map2 (curry Free) x_names x_Ts; val xs' = map mk_rec_arg_arg xs; in fold_rev Term.lambda xs (Term.list_comb (h, xs')) end; fun mk_rec_o_map_rhs recx = let val args = map2 mk_rec_o_map_arg rec_arg_Ts gs in Term.list_comb (recx, args) end; val rec_o_map_rhss = map mk_rec_o_map_rhs recs; val rec_o_map_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) rec_o_map_lhss rec_o_map_rhss; val rec_o_map_thms = @{map 3} (fn goal => fn rec_def => fn ctor_rec_o_map => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt rec_def pre_map_defs live_nesting_map_ident0s abs_inverses ctor_rec_o_map) |> Thm.close_derivation \<^here>) rec_o_map_goals rec_defs xtor_co_rec_o_maps; in map single rec_o_map_thms end; fun derive_note_induct_recs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (ctrss, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (recs, rec_defs)), lthy) = let val ((induct_thms, induct_thm, induct_attrs), (rec_thmss, rec_attrs)) = derive_induct_recs_thms_for_types plugins pre_bnfs recs_args_types xtor_co_induct xtor_co_rec_thms live_nesting_bnfs fp_nesting_bnfs fpTs Cs Xs ctrXs_Tsss abs_inverses type_definitions abs_inverses ctrss ctr_defss recs rec_defs lthy; val rec_transfer_thmss = map single (derive_rec_transfer_thms lthy recs rec_defs recs_args_types); val induct_type_attr = Attrib.internal o K o Induct.induct_type; val induct_pred_attr = Attrib.internal o K o Induct.induct_pred; val ((rel_induct_thmss, common_rel_induct_thms), (rel_induct_attrs, common_rel_induct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_induct_thms, common_rel_induct_thm), rel_induct_attrs) = derive_rel_induct_thms_for_types lthy nn fpTs As Bs ctrss ctr_Tsss (map #exhaust ctr_sugars) xtor_rel_co_induct ctr_defss ctor_injects pre_rel_defs abs_inverses live_nesting_rel_eqs; in ((map single rel_induct_thms, single common_rel_induct_thm), (rel_induct_attrs, rel_induct_attrs)) end; val rec_o_map_thmss = derive_rec_o_map_thmss lthy recs rec_defs; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars rec_thmss map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (if nn > 1 then [(inductN, [induct_thm], K induct_attrs), (rel_inductN, common_rel_induct_thms, K common_rel_induct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(inductN, map single induct_thms, fn T_name => induct_attrs @ [induct_type_attr T_name]), (recN, rec_thmss, K rec_attrs), (rec_o_mapN, rec_o_map_thmss, K []), (rec_transferN, rec_transfer_thmss, K []), (rel_inductN, rel_induct_thmss, K (rel_induct_attrs @ [induct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational recs (flat rec_thmss) |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) (flat rec_thmss)) |> Local_Theory.notes (common_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms (fst (dest_Type (hd fpTs)) ^ implode (map (prefix "_") (tl fp_b_names))) inductN |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Least_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars recs rec_defs map_thmss [induct_thm] (map single induct_thms) rec_thmss (replicate nn []) (replicate nn []) rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss (replicate nn []) (replicate nn []) rec_transfer_thmss common_rel_induct_thms rel_induct_thmss [] (replicate nn []) rec_o_map_thmss end; fun derive_corec_transfer_thms lthy corecs corec_defs = let val (Rs, Ss, goals, _) = mk_co_rec_transfer_goals lthy corecs; val (_, _, _, (((pgss, pss, qssss, gssss), _), _)) = the corecs_args_types; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_corec_transfer_tac ctxt (map (Thm.cterm_of ctxt) Ss) (map (Thm.cterm_of ctxt) Rs) type_definitions corec_defs xtor_co_rec_transfers pre_rel_defs live_nesting_rel_eqs (flat pgss) pss qssss gssss) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; fun derive_map_o_corec_thmss lthy0 lthy2 corecs corec_defs = if live = 0 then replicate nn [] else let fun variant_names n pre = fst (Variable.variant_fixes (replicate n pre) lthy0); val maps0 = map map_of_bnf fp_bnfs; val f_names = variant_names num_As "f"; val fs = map2 (curry Free) f_names (map (op -->) (As ~~ Bs)); val live_fs = AList.find (op =) (fs ~~ liveness) true; val fmaps = map (fn map0 => Term.list_comb (mk_map live As Bs map0, live_fs)) maps0; val corec_arg_Ts = binder_fun_types (hd (map fastype_of corecs)); val num_rec_args = length corec_arg_Ts; val g_names = variant_names num_rec_args "g"; val gs = map2 (curry Free) g_names corec_arg_Ts; val gcorecs = map (fn corecx => Term.list_comb (corecx, gs)) corecs; val map_o_corec_lhss = map2 (curry HOLogic.mk_comp) fmaps gcorecs; val ABfs = (As ~~ Bs) ~~ fs; fun mk_map_o_corec_arg corec_argB_T g = let val T = range_type (fastype_of g); val U = range_type corec_argB_T; in if T = U then g else HOLogic.mk_comp (build_map lthy2 [] [] (the o AList.lookup (op =) ABfs) (T, U), g) end; fun mk_map_o_corec_rhs corecx = let val args = map2 (mk_map_o_corec_arg o B_ify_T) corec_arg_Ts gs in Term.list_comb (B_ify corecx, args) end; val map_o_corec_rhss = map mk_map_o_corec_rhs corecs; val map_o_corec_goals = map2 (fold_rev (fold_rev Logic.all) [fs, gs] o HOLogic.mk_Trueprop oo curry HOLogic.mk_eq) map_o_corec_lhss map_o_corec_rhss; val map_o_corec_thms = @{map 3} (fn goal => fn corec_def => fn dtor_map_o_corec => Goal.prove_sorry lthy2 [] [] goal (fn {context = ctxt, ...} => mk_co_rec_o_map_tac ctxt corec_def pre_map_defs live_nesting_map_ident0s abs_inverses dtor_map_o_corec) |> Thm.close_derivation \<^here>) map_o_corec_goals corec_defs xtor_co_rec_o_maps; in map single map_o_corec_thms end; fun derive_note_coinduct_corecs_thms_for_types ((((map_thmss, map_disc_iffss, map_selsss, rel_injectss, rel_distinctss, rel_selss, rel_intross, rel_casess, pred_injectss, set_thmss, set_selsssss, set_introsssss, set_casess, ctr_transferss, case_transferss, disc_transferss, sel_transferss), (_, _, ctor_iff_dtors, ctr_defss, ctr_sugars)), (corecs, corec_defs)), lthy) = let val (([(coinduct_thms, coinduct_thm), (coinduct_strong_thms, coinduct_strong_thm)], (coinduct_attrs, common_coinduct_attrs)), corec_thmss, corec_disc_thmss, (corec_disc_iff_thmss, corec_disc_iff_attrs), (corec_sel_thmsss, corec_sel_attrs)) = derive_coinduct_corecs_thms_for_types lthy pre_bnfs (the corecs_args_types) xtor_co_induct dtor_injects dtor_ctors xtor_co_rec_thms live_nesting_bnfs fpTs Cs Xs ctrXs_Tsss kss mss ns abs_inverses abs_inverses I ctr_defss ctr_sugars corecs corec_defs; fun distinct_prems ctxt thm = Goal.prove (*no sorry*) ctxt [] [] (thm |> Thm.prop_of |> Logic.strip_horn |>> distinct (op aconv) |> Logic.list_implies) (fn _ => HEADGOAL (cut_tac thm THEN' assume_tac ctxt) THEN ALLGOALS eq_assume_tac); fun eq_ifIN _ [thm] = thm | eq_ifIN ctxt (thm :: thms) = distinct_prems ctxt (@{thm eq_ifI} OF (map (unfold_thms ctxt @{thms atomize_imp[of _ "t = u" for t u]}) [thm, eq_ifIN ctxt thms])); val corec_code_thms = map (eq_ifIN lthy) corec_thmss; val corec_sel_thmss = map flat corec_sel_thmsss; val coinduct_type_attr = Attrib.internal o K o Induct.coinduct_type; val coinduct_pred_attr = Attrib.internal o K o Induct.coinduct_pred; val flat_corec_thms = append oo append; val corec_transfer_thmss = map single (derive_corec_transfer_thms lthy corecs corec_defs); val ((rel_coinduct_thmss, common_rel_coinduct_thms), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = if live = 0 then ((replicate nn [], []), ([], [])) else let val ((rel_coinduct_thms, common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) = derive_rel_coinduct_thms_for_types lthy nn fpTs ns As Bs mss ctr_sugars abs_inverses abs_injects ctor_injects dtor_ctors pre_rel_defs ctr_defss xtor_rel_co_induct live_nesting_rel_eqs; in ((map single rel_coinduct_thms, single common_rel_coinduct_thm), (rel_coinduct_attrs, common_rel_coinduct_attrs)) end; val map_o_corec_thmss = derive_map_o_corec_thmss lthy lthy corecs corec_defs; val (set_induct_thms, set_induct_attrss) = derive_set_induct_thms_for_types lthy nn fpTs (map #ctrs ctr_sugars) (map (map (mk_set As)) (map sets_of_bnf fp_bnfs)) dtor_set_inducts (map #exhaust ctr_sugars) (flat pre_set_defss) (flat ctr_defss) dtor_ctors abs_inverses |> split_list; val simp_thmss = @{map 6} mk_simp_thms ctr_sugars (@{map 3} flat_corec_thms corec_disc_thmss corec_disc_iff_thmss corec_sel_thmss) map_thmss rel_injectss rel_distinctss set_thmss; val common_notes = (set_inductN, set_induct_thms, nth set_induct_attrss) :: (if nn > 1 then [(coinductN, [coinduct_thm], K common_coinduct_attrs), (coinduct_strongN, [coinduct_strong_thm], K common_coinduct_attrs), (rel_coinductN, common_rel_coinduct_thms, K common_rel_coinduct_attrs)] else []) |> massage_simple_notes fp_common_name; val notes = [(coinductN, map single coinduct_thms, fn T_name => coinduct_attrs @ [coinduct_type_attr T_name]), (coinduct_strongN, map single coinduct_strong_thms, K coinduct_attrs), (corecN, corec_thmss, K []), (corec_codeN, map single corec_code_thms, K (nitpicksimp_attrs)), (corec_discN, corec_disc_thmss, K []), (corec_disc_iffN, corec_disc_iff_thmss, K corec_disc_iff_attrs), (corec_selN, corec_sel_thmss, K corec_sel_attrs), (corec_transferN, corec_transfer_thmss, K []), (map_o_corecN, map_o_corec_thmss, K []), (rel_coinductN, rel_coinduct_thmss, K (rel_coinduct_attrs @ [coinduct_pred_attr ""])), (simpsN, simp_thmss, K [])] |> massage_multi_notes fp_b_names fpTs; in lthy |> fold (Spec_Rules.add Binding.empty Spec_Rules.equational corecs) [flat corec_sel_thmss, flat corec_thmss] |> plugins code_plugin ? Code.declare_default_eqns (map (rpair true) corec_code_thms) |> Local_Theory.notes (common_notes @ notes) |-> interpret_bnfs_register_fp_sugars plugins fpTs fpBTs Xs Greatest_FP pre_bnfs absT_infos fp_nesting_bnfs live_nesting_bnfs fp_res ctrXs_Tsss ctor_iff_dtors ctr_defss ctr_sugars corecs corec_defs map_thmss [coinduct_thm, coinduct_strong_thm] (transpose [coinduct_thms, coinduct_strong_thms]) corec_thmss corec_disc_thmss corec_sel_thmsss rel_injectss rel_distinctss map_disc_iffss map_selsss rel_selss rel_intross rel_casess pred_injectss set_thmss set_selsssss set_introsssss set_casess ctr_transferss case_transferss disc_transferss sel_transferss corec_disc_iff_thmss (map single corec_code_thms) corec_transfer_thmss common_rel_coinduct_thms rel_coinduct_thmss set_induct_thms (replicate nn (if nn = 1 then set_induct_thms else [])) map_o_corec_thmss end; val lthy = lthy |> live > 0 ? fold2 (fn Type (s, _) => fn bnf => register_bnf_raw s bnf) fpTs fp_bnfs |> @{fold_map 29} define_ctrs_dtrs_for_type_etc fp_bnfs fp_bs fpTs Cs Es ctors dtors xtor_co_recs ctor_dtors dtor_ctors ctor_injects pre_map_defs pre_set_defss pre_rel_defs xtor_maps xtor_setss xtor_rels ns kss mss abss abs_injects type_definitions ctr_bindingss ctr_mixfixess ctr_Tsss disc_bindingss sel_bindingsss raw_sel_default_eqss |> wrap_ctrs_derive_map_set_rel_pred_thms_define_co_rec_for_types |> case_fp fp derive_note_induct_recs_thms_for_types derive_note_coinduct_corecs_thms_for_types; val timer = time (timer ("Constructors, discriminators, selectors, etc., for the new " ^ co_prefix fp ^ "datatype")); in lthy end; fun co_datatypes fp = define_co_datatypes (K I) (K I) (K I) (K I) fp; fun co_datatype_cmd fp construct_fp options lthy = define_co_datatypes Plugin_Name.make_filter Typedecl.read_constraint Syntax.parse_typ Syntax.parse_term fp construct_fp options lthy handle EMPTY_DATATYPE s => error ("Cannot define empty datatype " ^ quote s); val parse_ctr_arg = \<^keyword>\(\ |-- parse_binding_colon -- Parse.typ --| \<^keyword>\)\ || Parse.typ >> pair Binding.empty; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.binding parse_ctr_arg -- Parse.opt_mixfix); val parse_spec = parse_type_args_named_constrained -- Parse.binding -- Parse.opt_mixfix -- (\<^keyword>\=\ |-- parse_ctr_specs) -- parse_map_rel_pred_bindings -- parse_sel_default_eqs; val parse_co_datatype = parse_ctr_options -- Parse.and_list1 parse_spec; fun parse_co_datatype_cmd fp construct_fp = parse_co_datatype >> co_datatype_cmd fp construct_fp; end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar.ML @@ -1,1273 +1,1273 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar.ML Author: Jasmin Blanchette, TU Muenchen Author: Martin Desharnais, TU Muenchen Copyright 2012, 2013 Wrapping existing freely generated type's constructors. *) signature CTR_SUGAR = sig datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; val morph_ctr_sugar: morphism -> ctr_sugar -> ctr_sugar val transfer_ctr_sugar: theory -> ctr_sugar -> ctr_sugar val ctr_sugar_of: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_global: theory -> string -> ctr_sugar option val ctr_sugars_of: Proof.context -> ctr_sugar list val ctr_sugars_of_global: theory -> ctr_sugar list val ctr_sugar_of_case: Proof.context -> string -> ctr_sugar option val ctr_sugar_of_case_global: theory -> string -> ctr_sugar option val ctr_sugar_interpretation: string -> (ctr_sugar -> local_theory -> local_theory) -> theory -> theory val interpret_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val register_ctr_sugar_raw: ctr_sugar -> local_theory -> local_theory val register_ctr_sugar: (string -> bool) -> ctr_sugar -> local_theory -> local_theory val default_register_ctr_sugar_global: (string -> bool) -> ctr_sugar -> theory -> theory val mk_half_pairss: 'a list * 'a list -> ('a * 'a) list list val join_halves: int -> 'a list list -> 'a list list -> 'a list * 'a list list list val mk_ctr: typ list -> term -> term val mk_case: typ list -> typ -> term -> term val mk_disc_or_sel: typ list -> term -> term val name_of_ctr: term -> string val name_of_disc: term -> string val dest_ctr: Proof.context -> string -> term -> term * term list val dest_case: Proof.context -> string -> typ list -> term -> (ctr_sugar * term list * term list) option type ('c, 'a) ctr_spec = (binding * 'c) * 'a list val disc_of_ctr_spec: ('c, 'a) ctr_spec -> binding val ctr_of_ctr_spec: ('c, 'a) ctr_spec -> 'c val args_of_ctr_spec: ('c, 'a) ctr_spec -> 'a list val code_plugin: string type ctr_options = (string -> bool) * bool type ctr_options_cmd = (Proof.context -> string -> bool) * bool val fake_local_theory_for_sel_defaults: (binding * typ) list -> Proof.context -> Proof.context val free_constructors: ctr_sugar_kind -> ({prems: thm list, context: Proof.context} -> tactic) list list -> ((ctr_options * binding) * (term, binding) ctr_spec list) * term list -> local_theory -> ctr_sugar * local_theory val free_constructors_cmd: ctr_sugar_kind -> ((((Proof.context -> Plugin_Name.filter) * bool) * binding) * ((binding * string) * binding list) list) * string list -> Proof.context -> Proof.state val default_ctr_options: ctr_options val default_ctr_options_cmd: ctr_options_cmd val parse_bound_term: (binding * string) parser val parse_ctr_options: ctr_options_cmd parser val parse_ctr_spec: 'c parser -> 'a parser -> ('c, 'a) ctr_spec parser val parse_sel_default_eqs: string list parser end; structure Ctr_Sugar : CTR_SUGAR = struct open Ctr_Sugar_Util open Ctr_Sugar_Tactics open Ctr_Sugar_Code datatype ctr_sugar_kind = Datatype | Codatatype | Record | Unknown; type ctr_sugar = {kind: ctr_sugar_kind, T: typ, ctrs: term list, casex: term, discs: term list, selss: term list list, exhaust: thm, nchotomy: thm, injects: thm list, distincts: thm list, case_thms: thm list, case_cong: thm, case_cong_weak: thm, case_distribs: thm list, split: thm, split_asm: thm, disc_defs: thm list, disc_thmss: thm list list, discIs: thm list, disc_eq_cases: thm list, sel_defs: thm list, sel_thmss: thm list list, distinct_discsss: thm list list list, exhaust_discs: thm list, exhaust_sels: thm list, collapses: thm list, expands: thm list, split_sels: thm list, split_sel_asms: thm list, case_eq_ifs: thm list}; fun morph_ctr_sugar phi ({kind, T, ctrs, casex, discs, selss, exhaust, nchotomy, injects, distincts, case_thms, case_cong, case_cong_weak, case_distribs, split, split_asm, disc_defs, disc_thmss, discIs, disc_eq_cases, sel_defs, sel_thmss, distinct_discsss, exhaust_discs, exhaust_sels, collapses, expands, split_sels, split_sel_asms, case_eq_ifs} : ctr_sugar) = {kind = kind, T = Morphism.typ phi T, ctrs = map (Morphism.term phi) ctrs, casex = Morphism.term phi casex, discs = map (Morphism.term phi) discs, selss = map (map (Morphism.term phi)) selss, exhaust = Morphism.thm phi exhaust, nchotomy = Morphism.thm phi nchotomy, injects = map (Morphism.thm phi) injects, distincts = map (Morphism.thm phi) distincts, case_thms = map (Morphism.thm phi) case_thms, case_cong = Morphism.thm phi case_cong, case_cong_weak = Morphism.thm phi case_cong_weak, case_distribs = map (Morphism.thm phi) case_distribs, split = Morphism.thm phi split, split_asm = Morphism.thm phi split_asm, disc_defs = map (Morphism.thm phi) disc_defs, disc_thmss = map (map (Morphism.thm phi)) disc_thmss, discIs = map (Morphism.thm phi) discIs, disc_eq_cases = map (Morphism.thm phi) disc_eq_cases, sel_defs = map (Morphism.thm phi) sel_defs, sel_thmss = map (map (Morphism.thm phi)) sel_thmss, distinct_discsss = map (map (map (Morphism.thm phi))) distinct_discsss, exhaust_discs = map (Morphism.thm phi) exhaust_discs, exhaust_sels = map (Morphism.thm phi) exhaust_sels, collapses = map (Morphism.thm phi) collapses, expands = map (Morphism.thm phi) expands, split_sels = map (Morphism.thm phi) split_sels, split_sel_asms = map (Morphism.thm phi) split_sel_asms, case_eq_ifs = map (Morphism.thm phi) case_eq_ifs}; val transfer_ctr_sugar = morph_ctr_sugar o Morphism.transfer_morphism; structure Data = Generic_Data ( type T = (Position.T * ctr_sugar) Symtab.table; val empty = Symtab.empty; val extend = I; fun merge data : T = Symtab.merge (K true) data; ); fun ctr_sugar_of_generic context = Option.map (transfer_ctr_sugar (Context.theory_of context) o #2) o Symtab.lookup (Data.get context); fun ctr_sugars_of_generic context = Symtab.fold (cons o transfer_ctr_sugar (Context.theory_of context) o #2 o #2) (Data.get context) []; fun ctr_sugar_of_case_generic context s = find_first (fn {casex = Const (s', _), ...} => s' = s | _ => false) (ctr_sugars_of_generic context); val ctr_sugar_of = ctr_sugar_of_generic o Context.Proof; val ctr_sugar_of_global = ctr_sugar_of_generic o Context.Theory; val ctr_sugars_of = ctr_sugars_of_generic o Context.Proof; val ctr_sugars_of_global = ctr_sugars_of_generic o Context.Theory; val ctr_sugar_of_case = ctr_sugar_of_case_generic o Context.Proof; val ctr_sugar_of_case_global = ctr_sugar_of_case_generic o Context.Theory; structure Ctr_Sugar_Plugin = Plugin(type T = ctr_sugar); fun ctr_sugar_interpretation name f = Ctr_Sugar_Plugin.interpretation name (fn ctr_sugar => fn lthy => f (transfer_ctr_sugar (Proof_Context.theory_of lthy) ctr_sugar) lthy); val interpret_ctr_sugar = Ctr_Sugar_Plugin.data; fun register_ctr_sugar_raw (ctr_sugar as {T = Type (name, _), ...}) = Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context => let val pos = Position.thread_data () in Data.map (Symtab.update (name, (pos, morph_ctr_sugar phi ctr_sugar))) context end); fun register_ctr_sugar plugins ctr_sugar = register_ctr_sugar_raw ctr_sugar #> interpret_ctr_sugar plugins ctr_sugar; fun default_register_ctr_sugar_global plugins (ctr_sugar as {T = Type (name, _), ...}) thy = let val tab = Data.get (Context.Theory thy); val pos = Position.thread_data (); in if Symtab.defined tab name then thy else thy |> Context.theory_map (Data.put (Symtab.update_new (name, (pos, ctr_sugar)) tab)) |> Named_Target.theory_map (Ctr_Sugar_Plugin.data plugins ctr_sugar) end; val is_prefix = "is_"; val un_prefix = "un_"; val not_prefix = "not_"; fun mk_unN 1 1 suf = un_prefix ^ suf | mk_unN _ l suf = un_prefix ^ suf ^ string_of_int l; val caseN = "case"; val case_congN = "case_cong"; val case_eq_ifN = "case_eq_if"; val collapseN = "collapse"; val discN = "disc"; val disc_eq_caseN = "disc_eq_case"; val discIN = "discI"; val distinctN = "distinct"; val distinct_discN = "distinct_disc"; val exhaustN = "exhaust"; val exhaust_discN = "exhaust_disc"; val expandN = "expand"; val injectN = "inject"; val nchotomyN = "nchotomy"; val selN = "sel"; val exhaust_selN = "exhaust_sel"; val splitN = "split"; val split_asmN = "split_asm"; val split_selN = "split_sel"; val split_sel_asmN = "split_sel_asm"; val splitsN = "splits"; val split_selsN = "split_sels"; val case_cong_weak_thmsN = "case_cong_weak"; val case_distribN = "case_distrib"; val cong_attrs = @{attributes [cong]}; val dest_attrs = @{attributes [dest]}; val safe_elim_attrs = @{attributes [elim!]}; val iff_attrs = @{attributes [iff]}; val inductsimp_attrs = @{attributes [induct_simp]}; val nitpicksimp_attrs = @{attributes [nitpick_simp]}; val simp_attrs = @{attributes [simp]}; fun unflat_lookup eq xs ys = map (fn xs' => permute_like_unique eq xs xs' ys); fun mk_half_pairss' _ ([], []) = [] | mk_half_pairss' indent (x :: xs, _ :: ys) = indent @ fold_rev (cons o single o pair x) ys (mk_half_pairss' ([] :: indent) (xs, ys)); fun mk_half_pairss p = mk_half_pairss' [[]] p; fun join_halves n half_xss other_half_xss = (splice (flat half_xss) (flat other_half_xss), map2 (map2 append) (Library.chop_groups n half_xss) (transpose (Library.chop_groups n other_half_xss))); fun mk_undefined T = Const (\<^const_name>\undefined\, T); fun mk_ctr Ts t = let val Type (_, Ts0) = body_type (fastype_of t) in subst_nonatomic_types (Ts0 ~~ Ts) t end; fun mk_case Ts T t = let val (Type (_, Ts0), body) = strip_type (fastype_of t) |>> List.last in subst_nonatomic_types ((body, T) :: (Ts0 ~~ Ts)) t end; fun mk_disc_or_sel Ts t = subst_nonatomic_types (snd (Term.dest_Type (domain_type (fastype_of t))) ~~ Ts) t; val name_of_ctr = name_of_const "constructor" body_type; fun name_of_disc t = (case head_of t of Abs (_, _, \<^const>\Not\ $ (t' $ Bound 0)) => Long_Name.map_base_name (prefix not_prefix) (name_of_disc t') | Abs (_, _, Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t') => Long_Name.map_base_name (prefix is_prefix) (name_of_disc t') | Abs (_, _, \<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, _) $ Bound 0 $ t')) => Long_Name.map_base_name (prefix (not_prefix ^ is_prefix)) (name_of_disc t') | t' => name_of_const "discriminator" (perhaps (try domain_type)) t'); val base_name_of_ctr = Long_Name.base_name o name_of_ctr; fun dest_ctr ctxt s t = let val (f, args) = Term.strip_comb t in (case ctr_sugar_of ctxt s of SOME {ctrs, ...} => (case find_first (can (fo_match ctxt f)) ctrs of SOME f' => (f', args) | NONE => raise Fail "dest_ctr") | NONE => raise Fail "dest_ctr") end; fun dest_case ctxt s Ts t = (case Term.strip_comb t of (Const (c, _), args as _ :: _) => (case ctr_sugar_of ctxt s of SOME (ctr_sugar as {casex = Const (case_name, _), discs = discs0, selss = selss0, ...}) => if case_name = c then let val n = length discs0 in if n < length args then let val (branches, obj :: leftovers) = chop n args; val discs = map (mk_disc_or_sel Ts) discs0; val selss = map (map (mk_disc_or_sel Ts)) selss0; val conds = map (rapp obj) discs; val branch_argss = map (fn sels => map (rapp obj) sels @ leftovers) selss; val branches' = map2 (curry Term.betapplys) branches branch_argss; in SOME (ctr_sugar, conds, branches') end else NONE end else NONE | _ => NONE) | _ => NONE); fun const_or_free_name (Const (s, _)) = Long_Name.base_name s | const_or_free_name (Free (s, _)) = s | const_or_free_name t = raise TERM ("const_or_free_name", [t]) fun extract_sel_default ctxt t = let fun malformed () = error ("Malformed selector default value equation: " ^ Syntax.string_of_term ctxt t); val ((sel, (ctr, vars)), rhs) = fst (Term.replace_dummy_patterns (Syntax.check_term ctxt t) 0) |> HOLogic.dest_eq |>> (Term.dest_comb #>> const_or_free_name ##> (Term.strip_comb #>> (Term.dest_Const #> fst))) handle TERM _ => malformed (); in if forall (is_Free orf is_Var) vars andalso not (has_duplicates (op aconv) vars) then ((ctr, sel), fold_rev Term.lambda vars rhs) else malformed () end; (* Ideally, we would enrich the context with constants rather than free variables. *) fun fake_local_theory_for_sel_defaults sel_bTs = Proof_Context.allow_dummies #> Proof_Context.add_fixes (map (fn (b, T) => (b, SOME T, NoSyn)) sel_bTs) #> snd; type ('c, 'a) ctr_spec = (binding * 'c) * 'a list; fun disc_of_ctr_spec ((disc, _), _) = disc; fun ctr_of_ctr_spec ((_, ctr), _) = ctr; fun args_of_ctr_spec (_, args) = args; val code_plugin = Plugin_Name.declare_setup \<^binding>\code\; fun prepare_free_constructors kind prep_plugins prep_term ((((raw_plugins, discs_sels), raw_case_binding), ctr_specs), sel_default_eqs) no_defs_lthy = let val plugins = prep_plugins no_defs_lthy raw_plugins; (* TODO: sanity checks on arguments *) val raw_ctrs = map ctr_of_ctr_spec ctr_specs; val raw_disc_bindings = map disc_of_ctr_spec ctr_specs; val raw_sel_bindingss = map args_of_ctr_spec ctr_specs; val n = length raw_ctrs; val ks = 1 upto n; val _ = n > 0 orelse error "No constructors specified"; val ctrs0 = map (prep_term no_defs_lthy) raw_ctrs; val (fcT_name, As0) = (case body_type (fastype_of (hd ctrs0)) of Type T' => T' | _ => error "Expected type constructor in body type of constructor"); val _ = forall ((fn Type (T_name, _) => T_name = fcT_name | _ => false) o body_type o fastype_of) (tl ctrs0) orelse error "Constructors not constructing same type"; val fc_b_name = Long_Name.base_name fcT_name; val fc_b = Binding.name fc_b_name; fun qualify mandatory = Binding.qualify mandatory fc_b_name; val (unsorted_As, [B, C]) = no_defs_lthy |> variant_tfrees (map (fst o dest_TFree_or_TVar) As0) ||> fst o mk_TFrees 2; val As = map2 (resort_tfree_or_tvar o snd o dest_TFree_or_TVar) As0 unsorted_As; val fcT = Type (fcT_name, As); val ctrs = map (mk_ctr As) ctrs0; val ctr_Tss = map (binder_types o fastype_of) ctrs; val ms = map length ctr_Tss; fun can_definitely_rely_on_disc k = not (Binding.is_empty (nth raw_disc_bindings (k - 1))) orelse nth ms (k - 1) = 0; fun can_rely_on_disc k = can_definitely_rely_on_disc k orelse (k = 1 andalso not (can_definitely_rely_on_disc 2)); fun should_omit_disc_binding k = n = 1 orelse (n = 2 andalso can_rely_on_disc (3 - k)); val equal_binding = \<^binding>\=\; fun is_disc_binding_valid b = not (Binding.is_empty b orelse Binding.eq_name (b, equal_binding)); val standard_disc_binding = Binding.name o prefix is_prefix o base_name_of_ctr; val disc_bindings = raw_disc_bindings |> @{map 4} (fn k => fn m => fn ctr => fn disc => qualify false (if Binding.is_empty disc then if m = 0 then equal_binding else if should_omit_disc_binding k then disc else standard_disc_binding ctr else if Binding.eq_name (disc, standard_binding) then standard_disc_binding ctr else disc)) ks ms ctrs0; fun standard_sel_binding m l = Binding.name o mk_unN m l o base_name_of_ctr; val sel_bindingss = @{map 3} (fn ctr => fn m => map2 (fn l => fn sel => qualify false (if Binding.is_empty sel orelse Binding.eq_name (sel, standard_binding) then standard_sel_binding m l ctr else sel)) (1 upto m) o pad_list Binding.empty m) ctrs0 ms raw_sel_bindingss; val add_bindings = Variable.add_fixes (distinct (op =) (filter Symbol_Pos.is_identifier (map Binding.name_of (disc_bindings @ flat sel_bindingss)))) #> snd; val case_Ts = map (fn Ts => Ts ---> B) ctr_Tss; val (((((((((u, exh_y), xss), yss), fs), gs), w), (p, p'))), _) = no_defs_lthy |> add_bindings |> yield_singleton (mk_Frees fc_b_name) fcT ||>> yield_singleton (mk_Frees "y") fcT (* for compatibility with "datatype_realizer.ML" *) ||>> mk_Freess "x" ctr_Tss ||>> mk_Freess "y" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "z") B ||>> yield_singleton (apfst (op ~~) oo mk_Frees' "P") HOLogic.boolT; val q = Free (fst p', mk_pred1T B); val xctrs = map2 (curry Term.list_comb) ctrs xss; val yctrs = map2 (curry Term.list_comb) ctrs yss; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; (* TODO: Eta-expension is for compatibility with the old datatype package (but it also provides nicer names). Consider removing. *) val eta_fs = map2 (fold_rev Term.lambda) xss xfs; val eta_gs = map2 (fold_rev Term.lambda) xss xgs; val case_binding = qualify false (if Binding.is_empty raw_case_binding orelse Binding.eq_name (raw_case_binding, standard_binding) then Binding.prefix_name (caseN ^ "_") fc_b else raw_case_binding); fun mk_case_disj xctr xf xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_eq (w, xf))); val case_rhs = fold_rev (fold_rev Term.lambda) [fs, [u]] (Const (\<^const_name>\The\, (B --> HOLogic.boolT) --> B) $ Term.lambda w (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_case_disj xctrs xfs xss))); val ((raw_case, (_, raw_case_def)), (lthy, lthy_old)) = no_defs_lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.define ((case_binding, NoSyn), ((Binding.concealed (Thm.def_binding case_binding), []), case_rhs)) ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy_old lthy; val case_def = Morphism.thm phi raw_case_def; val case0 = Morphism.term phi raw_case; val casex = mk_case As B case0; val casexC = mk_case As C case0; val casexBool = mk_case As HOLogic.boolT case0; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val exist_xs_u_eq_ctrs = map2 (fn xctr => fn xs => list_exists_free xs (HOLogic.mk_eq (u, xctr))) xctrs xss; val unique_disc_no_def = TrueI; (*arbitrary marker*) val alternate_disc_no_def = FalseE; (*arbitrary marker*) fun alternate_disc_lhs get_udisc k = HOLogic.mk_not (let val b = nth disc_bindings (k - 1) in if is_disc_binding_valid b then get_udisc b (k - 1) else nth exist_xs_u_eq_ctrs (k - 1) end); val no_discs_sels = not discs_sels andalso forall (forall Binding.is_empty) (raw_disc_bindings :: raw_sel_bindingss) andalso null sel_default_eqs; val (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy) = if no_discs_sels then (true, [], [], [], [], [], lthy) else let val all_sel_bindings = flat sel_bindingss; val num_all_sel_bindings = length all_sel_bindings; val uniq_sel_bindings = distinct Binding.eq_name all_sel_bindings; val all_sels_distinct = (length uniq_sel_bindings = num_all_sel_bindings); val sel_binding_index = if all_sels_distinct then 1 upto num_all_sel_bindings else map (fn b => find_index (curry Binding.eq_name b) uniq_sel_bindings) all_sel_bindings; val all_proto_sels = flat (@{map 3} (fn k => fn xs => map (pair k o pair xs)) ks xss xss); val sel_infos = AList.group (op =) (sel_binding_index ~~ all_proto_sels) |> sort (int_ord o apply2 fst) |> map snd |> curry (op ~~) uniq_sel_bindings; val sel_bindings = map fst sel_infos; val sel_defaults = if null sel_default_eqs then [] else let val sel_Ts = map (curry (op -->) fcT o fastype_of o snd o snd o hd o snd) sel_infos; val fake_lthy = fake_local_theory_for_sel_defaults (sel_bindings ~~ sel_Ts) no_defs_lthy; in map (extract_sel_default fake_lthy o prep_term fake_lthy) sel_default_eqs end; fun disc_free b = Free (Binding.name_of b, mk_pred1T fcT); fun disc_spec b exist_xs_u_eq_ctr = mk_Trueprop_eq (disc_free b $ u, exist_xs_u_eq_ctr); fun alternate_disc k = Term.lambda u (alternate_disc_lhs (K o rapp u o disc_free) (3 - k)); fun mk_sel_case_args b proto_sels T = @{map 3} (fn Const (c, _) => fn Ts => fn k => (case AList.lookup (op =) proto_sels k of NONE => (case filter (curry (op =) (c, Binding.name_of b) o fst) sel_defaults of [] => fold_rev (Term.lambda o curry Free Name.uu) Ts (mk_undefined T) | [(_, t)] => t | _ => error "Multiple default values for selector/constructor pair") | SOME (xs, x) => fold_rev Term.lambda xs x)) ctrs ctr_Tss ks; fun sel_spec b proto_sels = let val _ = (case duplicates (op =) (map fst proto_sels) of k :: _ => error ("Duplicate selector name " ^ quote (Binding.name_of b) ^ " for constructor " ^ quote (Syntax.string_of_term lthy (nth ctrs (k - 1)))) | [] => ()) val T = (case distinct (op =) (map (fastype_of o snd o snd) proto_sels) of [T] => T | T :: T' :: _ => error ("Inconsistent range type for selector " ^ quote (Binding.name_of b) ^ ": " ^ quote (Syntax.string_of_typ lthy T) ^ " vs. " ^ quote (Syntax.string_of_typ lthy T'))); in mk_Trueprop_eq (Free (Binding.name_of b, fcT --> T) $ u, Term.list_comb (mk_case As T case0, mk_sel_case_args b proto_sels T) $ u) end; fun unflat_selss xs = unflat_lookup Binding.eq_name sel_bindings xs sel_bindingss; val (((raw_discs, raw_disc_defs), (raw_sels, raw_sel_defs)), (lthy', lthy)) = lthy |> (snd o Local_Theory.begin_nested) |> apfst split_list o @{fold_map 3} (fn k => fn exist_xs_u_eq_ctr => fn b => if Binding.is_empty b then if n = 1 then pair (Term.lambda u (mk_uu_eq ()), unique_disc_no_def) else pair (alternate_disc k, alternate_disc_no_def) else if Binding.eq_name (b, equal_binding) then pair (Term.lambda u exist_xs_u_eq_ctr, refl) else Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), disc_spec b exist_xs_u_eq_ctr) #>> apsnd snd) ks exist_xs_u_eq_ctrs disc_bindings ||>> apfst split_list o fold_map (fn (b, proto_sels) => Specification.definition (SOME (b, NONE, NoSyn)) [] [] ((Thm.def_binding b, []), sel_spec b proto_sels) #>> apsnd snd) sel_infos ||> `Local_Theory.end_nested; val phi = Proof_Context.export_morphism lthy lthy'; val disc_defs = map (Morphism.thm phi) raw_disc_defs; val sel_defs = map (Morphism.thm phi) raw_sel_defs; val sel_defss = unflat_selss sel_defs; val discs0 = map (Morphism.term phi) raw_discs; val selss0 = unflat_selss (map (Morphism.term phi) raw_sels); val discs = map (mk_disc_or_sel As) discs0; val selss = map (map (mk_disc_or_sel As)) selss0; in (all_sels_distinct, discs, selss, disc_defs, sel_defs, sel_defss, lthy') end; fun mk_imp_p Qs = Logic.list_implies (Qs, HOLogic.mk_Trueprop p); val exhaust_goal = let fun mk_prem xctr xs = fold_rev Logic.all xs (mk_imp_p [mk_Trueprop_eq (exh_y, xctr)]) in fold_rev Logic.all [p, exh_y] (mk_imp_p (map2 mk_prem xctrs xss)) end; val inject_goalss = let fun mk_goal _ _ [] [] = [] | mk_goal xctr yctr xs ys = [fold_rev Logic.all (xs @ ys) (mk_Trueprop_eq (HOLogic.mk_eq (xctr, yctr), Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) xs ys)))]; in @{map 4} mk_goal xctrs yctrs xss yss end; val half_distinct_goalss = let fun mk_goal ((xs, xc), (xs', xc')) = fold_rev Logic.all (xs @ xs') (HOLogic.mk_Trueprop (HOLogic.mk_not (HOLogic.mk_eq (xc, xc')))); in map (map mk_goal) (mk_half_pairss (`I (xss ~~ xctrs))) end; val goalss = [exhaust_goal] :: inject_goalss @ half_distinct_goalss; fun after_qed ([exhaust_thm] :: thmss) lthy = let val ((((((((u, u'), (xss, xss')), fs), gs), h), v), p), _) = lthy |> add_bindings |> yield_singleton (apfst (op ~~) oo mk_Frees' fc_b_name) fcT ||>> mk_Freess' "x" ctr_Tss ||>> mk_Frees "f" case_Ts ||>> mk_Frees "g" case_Ts ||>> yield_singleton (mk_Frees "h") (B --> C) ||>> yield_singleton (mk_Frees (fc_b_name ^ "'")) fcT ||>> yield_singleton (mk_Frees "P") HOLogic.boolT; val xfs = map2 (curry Term.list_comb) fs xss; val xgs = map2 (curry Term.list_comb) gs xss; val fcase = Term.list_comb (casex, fs); val ufcase = fcase $ u; val vfcase = fcase $ v; val eta_fcase = Term.list_comb (casex, eta_fs); val eta_gcase = Term.list_comb (casex, eta_gs); val eta_ufcase = eta_fcase $ u; val eta_vgcase = eta_gcase $ v; fun mk_uu_eq () = HOLogic.mk_eq (u, u); val uv_eq = mk_Trueprop_eq (u, v); val ((inject_thms, inject_thmss), half_distinct_thmss) = chop n thmss |>> `flat; val rho_As = map (fn (T, U) => (dest_TVar T, Thm.ctyp_of lthy U)) (map Logic.varifyT_global As ~~ As); fun inst_thm t thm = Thm.instantiate' [] [SOME (Thm.cterm_of lthy t)] - (Thm.instantiate (rho_As, []) (Drule.zero_var_indexes thm)); + (Thm.instantiate (TVars.make rho_As, Vars.empty) (Drule.zero_var_indexes thm)); val uexhaust_thm = inst_thm u exhaust_thm; val exhaust_cases = map base_name_of_ctr ctrs; val other_half_distinct_thmss = map (map (fn thm => thm RS not_sym)) half_distinct_thmss; val (distinct_thms, (distinct_thmsss', distinct_thmsss)) = join_halves n half_distinct_thmss other_half_distinct_thmss ||> `transpose; val nchotomy_thm = let val goal = HOLogic.mk_Trueprop (HOLogic.mk_all (fst u', snd u', Library.foldr1 HOLogic.mk_disj exist_xs_u_eq_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_nchotomy_tac ctxt n exhaust_thm) |> Thm.close_derivation \<^here> end; val case_thms = let val goals = @{map 3} (fn xctr => fn xf => fn xs => fold_rev Logic.all (fs @ xs) (mk_Trueprop_eq (fcase $ xctr, xf))) xctrs xfs xss; in @{map 4} (fn k => fn goal => fn injects => fn distinctss => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_case_tac ctxt n k case_def injects distinctss) |> Thm.close_derivation \<^here>) ks goals inject_thmss distinct_thmsss end; val (case_cong_thm, case_cong_weak_thm) = let fun mk_prem xctr xs xf xg = fold_rev Logic.all xs (Logic.mk_implies (mk_Trueprop_eq (v, xctr), mk_Trueprop_eq (xf, xg))); val goal = Logic.list_implies (uv_eq :: @{map 4} mk_prem xctrs xss xfs xgs, mk_Trueprop_eq (eta_ufcase, eta_vgcase)); val weak_goal = Logic.mk_implies (uv_eq, mk_Trueprop_eq (ufcase, vfcase)); val vars = Variable.add_free_names lthy goal []; val weak_vars = Variable.add_free_names lthy weak_goal []; in (Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_case_cong_tac ctxt uexhaust_thm case_thms), Goal.prove_sorry lthy weak_vars [] weak_goal (fn {context = ctxt, prems = _} => etac ctxt arg_cong 1)) |> apply2 (Thm.close_derivation \<^here>) end; val split_lhs = q $ ufcase; fun mk_split_conjunct xctr xs f_xs = list_all_free xs (HOLogic.mk_imp (HOLogic.mk_eq (u, xctr), q $ f_xs)); fun mk_split_disjunct xctr xs f_xs = list_exists_free xs (HOLogic.mk_conj (HOLogic.mk_eq (u, xctr), HOLogic.mk_not (q $ f_xs))); fun mk_split_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, Library.foldr1 HOLogic.mk_conj (@{map 3} mk_split_conjunct xctrs xss xfs)); fun mk_split_asm_goal xctrs xss xfs = mk_Trueprop_eq (split_lhs, HOLogic.mk_not (Library.foldr1 HOLogic.mk_disj (@{map 3} mk_split_disjunct xctrs xss xfs))); fun prove_split selss goal = Variable.add_free_names lthy goal [] |> (fn vars => Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_split_tac ctxt uexhaust_thm case_thms selss inject_thmss distinct_thmsss)) |> Thm.close_derivation \<^here>; fun prove_split_asm asm_goal split_thm = Variable.add_free_names lthy asm_goal [] |> (fn vars => Goal.prove_sorry lthy vars [] asm_goal (fn {context = ctxt, ...} => mk_split_asm_tac ctxt split_thm)) |> Thm.close_derivation \<^here>; val (split_thm, split_asm_thm) = let val goal = mk_split_goal xctrs xss xfs; val asm_goal = mk_split_asm_goal xctrs xss xfs; val thm = prove_split (replicate n []) goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, exhaust_disc_thms, exhaust_sel_thms, all_collapse_thms, safe_collapse_thms, expand_thms, split_sel_thms, split_sel_asm_thms, case_eq_if_thms, disc_eq_case_thms) = if no_discs_sels then ([], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], [], []) else let val udiscs = map (rapp u) discs; val uselss = map (map (rapp u)) selss; val usel_ctrs = map2 (curry Term.list_comb) ctrs uselss; val usel_fs = map2 (curry Term.list_comb) fs uselss; val vdiscs = map (rapp v) discs; val vselss = map (map (rapp v)) selss; fun make_sel_thm xs' case_thm sel_def = zero_var_indexes (Variable.gen_all lthy (Drule.rename_bvars' (map (SOME o fst) xs') (Thm.forall_intr_vars (case_thm RS (sel_def RS trans))))); val sel_thmss = @{map 3} (map oo make_sel_thm) xss' case_thms sel_defss; fun has_undefined_rhs thm = (case snd (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of thm))) of Const (\<^const_name>\undefined\, _) => true | _ => false); val all_sel_thms = (if all_sels_distinct andalso null sel_default_eqs then flat sel_thmss else map_product (fn s => fn (xs', c) => make_sel_thm xs' c s) sel_defs (xss' ~~ case_thms)) |> filter_out has_undefined_rhs; fun mk_unique_disc_def () = let val m = the_single ms; val goal = mk_Trueprop_eq (mk_uu_eq (), the_single exist_xs_u_eq_ctrs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_unique_disc_def_tac ctxt m uexhaust_thm) |> Thm.close_derivation \<^here> end; fun mk_alternate_disc_def k = let val goal = mk_Trueprop_eq (alternate_disc_lhs (K (nth udiscs)) (3 - k), nth exist_xs_u_eq_ctrs (k - 1)); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_alternate_disc_def_tac ctxt k (nth disc_defs (2 - k)) (nth distinct_thms (2 - k)) uexhaust_thm) |> Thm.close_derivation \<^here> end; val has_alternate_disc_def = exists (fn def => Thm.eq_thm_prop (def, alternate_disc_no_def)) disc_defs; val nontriv_disc_defs = disc_defs |> filter_out (member Thm.eq_thm_prop [unique_disc_no_def, alternate_disc_no_def, refl]); val disc_defs' = map2 (fn k => fn def => if Thm.eq_thm_prop (def, unique_disc_no_def) then mk_unique_disc_def () else if Thm.eq_thm_prop (def, alternate_disc_no_def) then mk_alternate_disc_def k else def) ks disc_defs; val discD_thms = map (fn def => def RS iffD1) disc_defs'; val discI_thms = map2 (fn m => fn def => funpow m (fn thm => exI RS thm) (def RS iffD2)) ms disc_defs'; val not_discI_thms = map2 (fn m => fn def => funpow m (fn thm => allI RS thm) (unfold_thms lthy @{thms not_ex} (def RS @{thm ssubst[of _ _ Not]}))) ms disc_defs'; val (disc_thmss', disc_thmss) = let fun mk_thm discI _ [] = refl RS discI | mk_thm _ not_discI [distinct] = distinct RS not_discI; fun mk_thms discI not_discI distinctss = map (mk_thm discI not_discI) distinctss; in @{map 3} mk_thms discI_thms not_discI_thms distinct_thmsss' |> `transpose end; val nontriv_disc_thmss = map2 (fn b => if is_disc_binding_valid b then I else K []) disc_bindings disc_thmss; fun is_discI_triv b = (n = 1 andalso Binding.is_empty b) orelse Binding.eq_name (b, equal_binding); val nontriv_discI_thms = flat (map2 (fn b => if is_discI_triv b then K [] else single) disc_bindings discI_thms); val (distinct_disc_thms, (distinct_disc_thmsss', distinct_disc_thmsss)) = let fun mk_goal [] = [] | mk_goal [((_, udisc), (_, udisc'))] = [Logic.all u (Logic.mk_implies (HOLogic.mk_Trueprop udisc, HOLogic.mk_Trueprop (HOLogic.mk_not udisc')))]; fun prove tac goal = Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => tac ctxt) |> Thm.close_derivation \<^here>; val half_pairss = mk_half_pairss (`I (ms ~~ discD_thms ~~ udiscs)); val half_goalss = map mk_goal half_pairss; val half_thmss = @{map 3} (fn [] => K (K []) | [goal] => fn [(((m, discD), _), _)] => fn disc_thm => [prove (fn ctxt => mk_half_distinct_disc_tac ctxt m discD disc_thm) goal]) half_goalss half_pairss (flat disc_thmss'); val other_half_goalss = map (mk_goal o map swap) half_pairss; val other_half_thmss = map2 (map2 (fn thm => prove (fn ctxt => mk_other_half_distinct_disc_tac ctxt thm))) half_thmss other_half_goalss; in join_halves n half_thmss other_half_thmss ||> `transpose |>> has_alternate_disc_def ? K [] end; val exhaust_disc_thm = let fun mk_prem udisc = mk_imp_p [HOLogic.mk_Trueprop udisc]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem udiscs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_disc_tac ctxt n exhaust_thm discI_thms) |> Thm.close_derivation \<^here> end; val (safe_collapse_thms, all_collapse_thms) = let fun mk_goal m udisc usel_ctr = let val prem = HOLogic.mk_Trueprop udisc; val concl = mk_Trueprop_eq ((usel_ctr, u) |> m = 0 ? swap); in (prem aconv concl, Logic.all u (Logic.mk_implies (prem, concl))) end; val (trivs, goals) = @{map 3} mk_goal ms udiscs usel_ctrs |> split_list; val thms = @{map 5} (fn m => fn discD => fn sel_thms => fn triv => fn goal => Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, ...} => mk_collapse_tac ctxt m discD sel_thms ORELSE HEADGOAL (assume_tac ctxt)) |> Thm.close_derivation \<^here> |> not triv ? perhaps (try (fn thm => refl RS thm))) ms discD_thms sel_thmss trivs goals; in (map_filter (fn (true, _) => NONE | (false, thm) => SOME thm) (trivs ~~ thms), thms) end; val swapped_all_collapse_thms = map2 (fn m => fn thm => if m = 0 then thm else thm RS sym) ms all_collapse_thms; val exhaust_sel_thm = let fun mk_prem usel_ctr = mk_imp_p [mk_Trueprop_eq (u, usel_ctr)]; val goal = fold_rev Logic.all [p, u] (mk_imp_p (map mk_prem usel_ctrs)); in Goal.prove_sorry lthy [] [] goal (fn {context = ctxt, prems = _} => mk_exhaust_sel_tac ctxt n exhaust_disc_thm swapped_all_collapse_thms) |> Thm.close_derivation \<^here> end; val expand_thm = let fun mk_prems k udisc usels vdisc vsels = (if k = n then [] else [mk_Trueprop_eq (udisc, vdisc)]) @ (if null usels then [] else [Logic.list_implies (if n = 1 then [] else map HOLogic.mk_Trueprop [udisc, vdisc], HOLogic.mk_Trueprop (Library.foldr1 HOLogic.mk_conj (map2 (curry HOLogic.mk_eq) usels vsels)))]); val goal = Library.foldr Logic.list_implies (@{map 5} mk_prems ks udiscs uselss vdiscs vselss, uv_eq); val uncollapse_thms = map2 (fn thm => fn [] => thm | _ => thm RS sym) all_collapse_thms uselss; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, prems = _} => mk_expand_tac ctxt n ms (inst_thm u exhaust_disc_thm) (inst_thm v exhaust_disc_thm) uncollapse_thms distinct_disc_thmsss distinct_disc_thmsss') |> Thm.close_derivation \<^here> end; val (split_sel_thm, split_sel_asm_thm) = let val zss = map (K []) xss; val goal = mk_split_goal usel_ctrs zss usel_fs; val asm_goal = mk_split_asm_goal usel_ctrs zss usel_fs; val thm = prove_split sel_thmss goal; val asm_thm = prove_split_asm asm_goal thm; in (thm, asm_thm) end; val case_eq_if_thm = let val goal = mk_Trueprop_eq (ufcase, mk_IfN B udiscs usel_fs); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_eq_if_tac ctxt n uexhaust_thm case_thms disc_thmss' sel_thmss) |> Thm.close_derivation \<^here> end; val disc_eq_case_thms = let fun const_of_bool b = if b then \<^const>\True\ else \<^const>\False\; fun mk_case_args n = map_index (fn (k, argTs) => fold_rev Term.absdummy argTs (const_of_bool (n = k))) ctr_Tss; val goals = map_index (fn (n, udisc) => mk_Trueprop_eq (udisc, list_comb (casexBool, mk_case_args n) $ u)) udiscs; val goal = Logic.mk_conjunction_balanced goals; val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_disc_eq_case_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm (flat nontriv_disc_thmss) distinct_thms case_thms) |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) end; in (sel_defs, all_sel_thms, sel_thmss, nontriv_disc_defs, disc_thmss, nontriv_disc_thmss, discI_thms, nontriv_discI_thms, distinct_disc_thms, distinct_disc_thmsss, [exhaust_disc_thm], [exhaust_sel_thm], all_collapse_thms, safe_collapse_thms, [expand_thm], [split_sel_thm], [split_sel_asm_thm], [case_eq_if_thm], disc_eq_case_thms) end; val case_distrib_thm = let val args = @{map 2} (fn f => fn argTs => let val (args, _) = mk_Frees "x" argTs lthy in fold_rev Term.lambda args (h $ list_comb (f, args)) end) fs ctr_Tss; val goal = mk_Trueprop_eq (h $ ufcase, list_comb (casexC, args) $ u); val vars = Variable.add_free_names lthy goal []; in Goal.prove_sorry lthy vars [] goal (fn {context = ctxt, ...} => mk_case_distrib_tac ctxt (Thm.cterm_of ctxt u) exhaust_thm case_thms) |> Thm.close_derivation \<^here> end; val exhaust_case_names_attr = Attrib.internal (K (Rule_Cases.case_names exhaust_cases)); val cases_type_attr = Attrib.internal (K (Induct.cases_type fcT_name)); val nontriv_disc_eq_thmss = map (map (fn th => th RS @{thm eq_False[THEN iffD2]} handle THM _ => th RS @{thm eq_True[THEN iffD2]})) nontriv_disc_thmss; val anonymous_notes = [(map (fn th => th RS notE) distinct_thms, safe_elim_attrs), (flat nontriv_disc_eq_thmss, nitpicksimp_attrs)] |> map (fn (thms, attrs) => ((Binding.empty, attrs), [(thms, [])])); val notes = [(caseN, case_thms, nitpicksimp_attrs @ simp_attrs), (case_congN, [case_cong_thm], []), (case_cong_weak_thmsN, [case_cong_weak_thm], cong_attrs), (case_distribN, [case_distrib_thm], []), (case_eq_ifN, case_eq_if_thms, []), (collapseN, safe_collapse_thms, if ms = [0] then [] else simp_attrs), (discN, flat nontriv_disc_thmss, simp_attrs), (disc_eq_caseN, disc_eq_case_thms, []), (discIN, nontriv_discI_thms, []), (distinctN, distinct_thms, simp_attrs @ inductsimp_attrs), (distinct_discN, distinct_disc_thms, dest_attrs), (exhaustN, [exhaust_thm], [exhaust_case_names_attr, cases_type_attr]), (exhaust_discN, exhaust_disc_thms, [exhaust_case_names_attr]), (exhaust_selN, exhaust_sel_thms, [exhaust_case_names_attr]), (expandN, expand_thms, []), (injectN, inject_thms, iff_attrs @ inductsimp_attrs), (nchotomyN, [nchotomy_thm], []), (selN, all_sel_thms, nitpicksimp_attrs @ simp_attrs), (splitN, [split_thm], []), (split_asmN, [split_asm_thm], []), (split_selN, split_sel_thms, []), (split_sel_asmN, split_sel_asm_thms, []), (split_selsN, split_sel_thms @ split_sel_asm_thms, []), (splitsN, [split_thm, split_asm_thm], [])] |> filter_out (null o #2) |> map (fn (thmN, thms, attrs) => ((qualify true (Binding.name thmN), attrs), [(thms, [])])); val (noted, lthy') = lthy |> Spec_Rules.add Binding.empty Spec_Rules.equational [casex] case_thms |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (AList.group (eq_list (op aconv)) (map (`(single o lhs_head_of)) all_sel_thms)) |> fold (uncurry (Spec_Rules.add Binding.empty Spec_Rules.equational)) (filter_out (null o snd) (map single discs ~~ nontriv_disc_eq_thmss)) |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Case_Translation.register (Morphism.term phi casex) (map (Morphism.term phi) ctrs)) |> plugins code_plugin ? (Code.declare_default_eqns (map (rpair true) (flat nontriv_disc_eq_thmss @ case_thms @ all_sel_thms)) #> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (add_ctr_code fcT_name (map (Morphism.typ phi) As) (map (dest_Const o Morphism.term phi) ctrs) (Morphism.fact phi inject_thms) (Morphism.fact phi distinct_thms) (Morphism.fact phi case_thms)) I)) |> Local_Theory.notes (anonymous_notes @ notes) (* for "datatype_realizer.ML": *) |>> name_noted_thms fcT_name exhaustN; val ctr_sugar = {kind = kind, T = fcT, ctrs = ctrs, casex = casex, discs = discs, selss = selss, exhaust = exhaust_thm, nchotomy = nchotomy_thm, injects = inject_thms, distincts = distinct_thms, case_thms = case_thms, case_cong = case_cong_thm, case_cong_weak = case_cong_weak_thm, case_distribs = [case_distrib_thm], split = split_thm, split_asm = split_asm_thm, disc_defs = nontriv_disc_defs, disc_thmss = disc_thmss, discIs = discI_thms, disc_eq_cases = disc_eq_case_thms, sel_defs = sel_defs, sel_thmss = sel_thmss, distinct_discsss = distinct_disc_thmsss, exhaust_discs = exhaust_disc_thms, exhaust_sels = exhaust_sel_thms, collapses = all_collapse_thms, expands = expand_thms, split_sels = split_sel_thms, split_sel_asms = split_sel_asm_thms, case_eq_ifs = case_eq_if_thms} |> morph_ctr_sugar (substitute_noted_thm noted); in (ctr_sugar, lthy' |> register_ctr_sugar plugins ctr_sugar) end; in (goalss, after_qed, lthy) end; fun free_constructors kind tacss = (fn (goalss, after_qed, lthy) => map2 (map2 (Thm.close_derivation \<^here> oo Goal.prove_sorry lthy [] [])) goalss tacss |> (fn thms => after_qed thms lthy)) oo prepare_free_constructors kind (K I) (K I); fun free_constructors_cmd kind = (fn (goalss, after_qed, lthy) => Proof.theorem NONE (snd oo after_qed) (map (map (rpair [])) goalss) lthy) oo prepare_free_constructors kind Plugin_Name.make_filter Syntax.read_term; val parse_bound_term = Parse.binding --| \<^keyword>\:\ -- Parse.term; type ctr_options = Plugin_Name.filter * bool; type ctr_options_cmd = (Proof.context -> Plugin_Name.filter) * bool; val default_ctr_options : ctr_options = (Plugin_Name.default_filter, false); val default_ctr_options_cmd : ctr_options_cmd = (K Plugin_Name.default_filter, false); val parse_ctr_options = Scan.optional (\<^keyword>\(\ |-- Parse.list1 (Plugin_Name.parse_filter >> (apfst o K) || Parse.reserved "discs_sels" >> (apsnd o K o K true)) --| \<^keyword>\)\ >> (fn fs => fold I fs default_ctr_options_cmd)) default_ctr_options_cmd; fun parse_ctr_spec parse_ctr parse_arg = parse_opt_binding_colon -- parse_ctr -- Scan.repeat parse_arg; val parse_ctr_specs = Parse.enum1 "|" (parse_ctr_spec Parse.term Parse.binding); val parse_sel_default_eqs = Scan.optional (\<^keyword>\where\ |-- Parse.enum1 "|" Parse.prop) []; val _ = Outer_Syntax.local_theory_to_proof \<^command_keyword>\free_constructors\ "register an existing freely generated type's constructors" (parse_ctr_options -- Parse.binding --| \<^keyword>\for\ -- parse_ctr_specs -- parse_sel_default_eqs >> free_constructors_cmd Unknown); (** external views **) (* document antiquotations *) local fun antiquote_setup binding co = Document_Output.antiquotation_pretty_source_embedded binding ((Scan.ahead (Scan.lift Parse.not_eof) >> Token.pos_of) -- Args.type_name {proper = true, strict = true}) (fn ctxt => fn (pos, type_name) => let fun err () = error ("Bad " ^ Binding.name_of binding ^ ": " ^ quote type_name ^ Position.here pos); in (case ctr_sugar_of ctxt type_name of NONE => err () | SOME {kind, T = T0, ctrs = ctrs0, ...} => let val _ = if co = (kind = Codatatype) then () else err (); val T = Logic.unvarifyT_global T0; val ctrs = map Logic.unvarify_global ctrs0; val pretty_typ_bracket = Syntax.pretty_typ (Config.put pretty_priority 1001 ctxt); fun pretty_ctr ctr = Pretty.block (Pretty.breaks (Syntax.pretty_term ctxt ctr :: map pretty_typ_bracket (binder_types (fastype_of ctr)))); in Pretty.block (Pretty.keyword1 (Binding.name_of binding) :: Pretty.brk 1 :: Syntax.pretty_typ ctxt T :: Pretty.str " =" :: Pretty.brk 1 :: flat (separate [Pretty.brk 1, Pretty.str "| "] (map (single o pretty_ctr) ctrs))) end) end); in val _ = Theory.setup (antiquote_setup \<^binding>\datatype\ false #> antiquote_setup \<^binding>\codatatype\ true); end; (* theory export *) val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => if Export_Theory.export_enabled context then let val parents = map (Data.get o Context.Theory) (Theory.parents_of thy); val datatypes = (Data.get (Context.Theory thy), []) |-> Symtab.fold (fn (name, (pos, {kind, T, ctrs, ...})) => if kind = Record orelse exists (fn tab => Symtab.defined tab name) parents then I else let val pos_properties = Thy_Info.adjust_pos_properties context pos; val typ = Logic.unvarifyT_global T; val constrs = map Logic.unvarify_global ctrs; val typargs = rev (fold Term.add_tfrees (Logic.mk_type typ :: constrs) []); val constructors = map (fn t => (t, Term.type_of t)) constrs; in cons (pos_properties, (name, (kind = Codatatype, (typargs, (typ, constructors))))) end); in if null datatypes then () else Export_Theory.export_body thy "datatypes" let open XML.Encode Term_XML.Encode in list (pair properties (pair string (pair bool (pair (list (pair string sort)) (pair typ (list (pair (term (Sign.consts_of thy)) typ))))))) datatypes end end else ()); end; diff --git a/src/HOL/Tools/Function/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,323 +1,323 @@ (* 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 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; val extend = I; 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 ([], [(P_var, P_inst), (x_var, x_inst)]) + |> Thm.instantiate (TVars.empty, Vars.make [(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/Function/relation.ML b/src/HOL/Tools/Function/relation.ML --- a/src/HOL/Tools/Function/relation.ML +++ b/src/HOL/Tools/Function/relation.ML @@ -1,48 +1,49 @@ (* Title: HOL/Tools/Function/relation.ML Author: Alexander Krauss, TU Muenchen Tactics to start a termination proof using a user-specified relation. *) signature FUNCTION_RELATION = sig val relation_tac: Proof.context -> (typ -> term) -> int -> tactic val relation_infer_tac: Proof.context -> term -> int -> tactic end structure Function_Relation: FUNCTION_RELATION = struct (* tactic version *) fun inst_state_tac ctxt inst = SUBGOAL (fn (goal, _) => (case Term.add_vars goal [] of - [v as (_, T)] => PRIMITIVE (Thm.instantiate ([], [(v, Thm.cterm_of ctxt (inst T))])) + [v as (_, T)] => + PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt (inst T))])) | _ => no_tac)); fun relation_tac ctxt rel i = TRY (Function_Common.termination_rule_tac ctxt i) THEN inst_state_tac ctxt rel i; (* version with type inference *) fun inst_state_infer_tac ctxt rel = SUBGOAL (fn (goal, _) => (case Term.add_vars goal [] of [v as (_, T)] => let val rel' = singleton (Variable.polymorphic ctxt) rel |> map_types Type_Infer.paramify_vars |> Type.constraint T |> Syntax.check_term ctxt; - in PRIMITIVE (Thm.instantiate ([], [(v, Thm.cterm_of ctxt rel')])) end + in PRIMITIVE (Thm.instantiate (TVars.empty, Vars.make [(v, Thm.cterm_of ctxt rel')])) end | _ => no_tac)); fun relation_infer_tac ctxt rel i = TRY (Function_Common.termination_rule_tac ctxt i) THEN inst_state_infer_tac ctxt rel i; end diff --git a/src/HOL/Tools/Lifting/lifting_bnf.ML b/src/HOL/Tools/Lifting/lifting_bnf.ML --- a/src/HOL/Tools/Lifting/lifting_bnf.ML +++ b/src/HOL/Tools/Lifting/lifting_bnf.ML @@ -1,115 +1,115 @@ (* Title: HOL/Tools/Lifting/lifting_bnf.ML Author: Ondrej Kuncar, TU Muenchen Setup for Lifting for types that are BNF. *) signature LIFTING_BNF = sig end structure Lifting_BNF : LIFTING_BNF = struct open BNF_Util open BNF_Def open Transfer_BNF (* Quotient map theorem *) fun Quotient_tac bnf ctxt i = let val rel_Grp = rel_Grp_of_bnf bnf fun get_lhs thm = thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst val vars = get_lhs rel_Grp |> strip_comb |> snd |> map_filter (try (strip_comb #> snd #> hd)) val UNIVs = map (fn var => HOLogic.mk_UNIV (var |> dest_Var |> snd |> dest_Type |> snd |> hd)) vars val inst = map dest_Var vars ~~ map (Thm.cterm_of ctxt) UNIVs - val rel_Grp_UNIV_sym = rel_Grp |> Drule.instantiate_normalize ([], inst) + val rel_Grp_UNIV_sym = rel_Grp |> Drule.instantiate_normalize (TVars.empty, Vars.make inst) |> Local_Defs.unfold0 ctxt @{thms subset_UNIV[THEN eqTrueI] UNIV_def[symmetric] simp_thms(21)} |> (fn thm => thm RS sym) val rel_mono = rel_mono_of_bnf bnf val rel_conversep_sym = rel_conversep_of_bnf bnf RS sym in EVERY' [SELECT_GOAL (Local_Defs.unfold0_tac ctxt [@{thm Quotient_alt_def5}]), REPEAT_DETERM o (etac ctxt conjE), rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [rel_Grp_UNIV_sym]), rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt, rtac ctxt conjI, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [rel_conversep_sym, rel_Grp_UNIV_sym]), rtac ctxt rel_mono THEN_ALL_NEW assume_tac ctxt, SELECT_GOAL (Local_Defs.unfold0_tac ctxt [rel_conversep_sym, rel_OO_of_bnf bnf RS sym]), hyp_subst_tac ctxt, rtac ctxt refl] i end fun mk_Quotient args = let val argTs = map fastype_of args in list_comb (Const (\<^const_name>\Quotient\, argTs ---> HOLogic.boolT), args) end fun prove_Quotient_map bnf ctxt = let val live = live_of_bnf bnf val (((As, Bs), Ds), ctxt') = ctxt |> mk_TFrees live ||>> mk_TFrees live ||>> mk_TFrees (dead_of_bnf bnf) val argTss = map2 (fn a => fn b => [mk_pred2T a a, a --> b, b --> a,mk_pred2T a b]) As Bs val ((argss, argss'), ctxt'') = ctxt' |> @{fold_map 2} mk_Frees ["R", "Abs", "Rep", "T"] (transpose argTss) |>> `transpose val assms = map (mk_Quotient #> HOLogic.mk_Trueprop) argss val R_rel = list_comb (mk_rel_of_bnf Ds As As bnf, nth argss' 0) val Abs_map = list_comb (mk_map_of_bnf Ds As Bs bnf, nth argss' 1) val Rep_map = list_comb (mk_map_of_bnf Ds Bs As bnf, nth argss' 2) val T_rel = list_comb (mk_rel_of_bnf Ds As Bs bnf, nth argss' 3) val concl = mk_Quotient [R_rel, Abs_map, Rep_map, T_rel] |> HOLogic.mk_Trueprop val goal = Logic.list_implies (assms, concl) in Goal.prove_sorry ctxt'' [] [] goal (fn {context = goal_ctxt, ...} => Quotient_tac bnf goal_ctxt 1) |> Thm.close_derivation \<^here> |> singleton (Variable.export ctxt'' ctxt) |> Drule.zero_var_indexes end fun Quotient_map bnf ctxt = let val Quotient = prove_Quotient_map bnf ctxt val Quotient_thm_name = Binding.qualify_name true (base_name_of_bnf bnf) "Quotient" in [((Quotient_thm_name, []), [([Quotient], @{attributes [quot_map]})])] end (* relator_eq_onp *) fun relator_eq_onp bnf = [(Binding.empty_atts, [([rel_eq_onp_of_bnf bnf], @{attributes [relator_eq_onp]})])] (* relator_mono *) fun relator_mono bnf = [(Binding.empty_atts, [([rel_mono_of_bnf bnf], @{attributes [relator_mono]})])] (* relator_distr *) fun relator_distr bnf = [(Binding.empty_atts, [([rel_OO_of_bnf bnf RS sym], @{attributes [relator_distr]})])] (* interpretation *) fun lifting_bnf_interpretation bnf lthy = if dead_of_bnf bnf > 0 then lthy else let val notess = [relator_eq_onp bnf, Quotient_map bnf lthy, relator_mono bnf, relator_distr bnf] in lthy |> fold (perhaps o try o (snd oo Local_Theory.notes)) notess end val lifting_plugin = Plugin_Name.declare_setup \<^binding>\lifting\ val _ = Theory.setup (bnf_interpretation lifting_plugin (bnf_only_type_ctr lifting_bnf_interpretation)) end diff --git a/src/HOL/Tools/Lifting/lifting_def.ML b/src/HOL/Tools/Lifting/lifting_def.ML --- a/src/HOL/Tools/Lifting/lifting_def.ML +++ b/src/HOL/Tools/Lifting/lifting_def.ML @@ -1,676 +1,676 @@ (* Title: HOL/Tools/Lifting/lifting_def.ML Author: Ondrej Kuncar Definitions for constants on quotient types. *) signature LIFTING_DEF = sig datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ type lift_def val rty_of_lift_def: lift_def -> typ val qty_of_lift_def: lift_def -> typ val rhs_of_lift_def: lift_def -> term val lift_const_of_lift_def: lift_def -> term val def_thm_of_lift_def: lift_def -> thm val rsp_thm_of_lift_def: lift_def -> thm val abs_eq_of_lift_def: lift_def -> thm val rep_eq_of_lift_def: lift_def -> thm option val code_eq_of_lift_def: lift_def -> code_eq val transfer_rules_of_lift_def: lift_def -> thm list val morph_lift_def: morphism -> lift_def -> lift_def val inst_of_lift_def: Proof.context -> typ -> lift_def -> lift_def val mk_lift_const_of_lift_def: typ -> lift_def -> term type config = { notes: bool } val map_config: (bool -> bool) -> config -> config val default_config: config val generate_parametric_transfer_rule: Proof.context -> thm -> thm -> thm val add_lift_def: config -> binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory val prepare_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> Proof.context -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> thm list -> local_theory -> term option * (thm -> Proof.context -> lift_def * local_theory) val gen_lift_def: (binding * mixfix -> typ -> term -> thm -> thm list -> local_theory -> lift_def * local_theory) -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val lift_def: config -> binding * mixfix -> typ -> term -> (Proof.context -> tactic) -> thm list -> local_theory -> lift_def * local_theory val can_generate_code_cert: thm -> bool end structure Lifting_Def: LIFTING_DEF = struct open Lifting_Util infix 0 MRSL datatype code_eq = UNKNOWN_EQ | NONE_EQ | ABS_EQ | REP_EQ datatype lift_def = LIFT_DEF of { rty: typ, qty: typ, rhs: term, lift_const: term, def_thm: thm, rsp_thm: thm, abs_eq: thm, rep_eq: thm option, code_eq: code_eq, transfer_rules: thm list }; fun rep_lift_def (LIFT_DEF lift_def) = lift_def; val rty_of_lift_def = #rty o rep_lift_def; val qty_of_lift_def = #qty o rep_lift_def; val rhs_of_lift_def = #rhs o rep_lift_def; val lift_const_of_lift_def = #lift_const o rep_lift_def; val def_thm_of_lift_def = #def_thm o rep_lift_def; val rsp_thm_of_lift_def = #rsp_thm o rep_lift_def; val abs_eq_of_lift_def = #abs_eq o rep_lift_def; val rep_eq_of_lift_def = #rep_eq o rep_lift_def; val code_eq_of_lift_def = #code_eq o rep_lift_def; val transfer_rules_of_lift_def = #transfer_rules o rep_lift_def; fun mk_lift_def rty qty rhs lift_const def_thm rsp_thm abs_eq rep_eq code_eq transfer_rules = LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }; fun map_lift_def f1 f2 f3 f4 f5 f6 f7 f8 f9 f10 (LIFT_DEF {rty = rty, qty = qty, rhs = rhs, lift_const = lift_const, def_thm = def_thm, rsp_thm = rsp_thm, abs_eq = abs_eq, rep_eq = rep_eq, code_eq = code_eq, transfer_rules = transfer_rules }) = LIFT_DEF {rty = f1 rty, qty = f2 qty, rhs = f3 rhs, lift_const = f4 lift_const, def_thm = f5 def_thm, rsp_thm = f6 rsp_thm, abs_eq = f7 abs_eq, rep_eq = f8 rep_eq, code_eq = f9 code_eq, transfer_rules = f10 transfer_rules } fun morph_lift_def phi = let val mtyp = Morphism.typ phi val mterm = Morphism.term phi val mthm = Morphism.thm phi in map_lift_def mtyp mtyp mterm mterm mthm mthm mthm (Option.map mthm) I (map mthm) end fun mk_inst_of_lift_def qty lift_def = Vartab.empty |> Type.raw_match (qty_of_lift_def lift_def, qty) fun mk_lift_const_of_lift_def qty lift_def = Envir.subst_term_types (mk_inst_of_lift_def qty lift_def) (lift_const_of_lift_def lift_def) fun inst_of_lift_def ctxt qty lift_def = let val instT = Vartab.fold (fn (a, (S, T)) => cons ((a, S), Thm.ctyp_of ctxt T)) (mk_inst_of_lift_def qty lift_def) [] - val phi = Morphism.instantiate_morphism (instT, []) + val phi = Morphism.instantiate_morphism (TVars.make instT, Vars.empty) in morph_lift_def phi lift_def end (* Config *) type config = { notes: bool }; fun map_config f1 { notes = notes } = { notes = f1 notes } val default_config = { notes = true }; (* Reflexivity prover *) fun mono_eq_prover ctxt prop = let val refl_rules = Lifting_Info.get_reflexivity_rules ctxt val transfer_rules = Transfer.get_transfer_raw ctxt fun main_tac i = (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt refl_rules) THEN_ALL_NEW (REPEAT_ALL_NEW (DETERM o resolve_tac ctxt transfer_rules))) i in SOME (Goal.prove ctxt [] [] prop (K (main_tac 1))) handle ERROR _ => NONE end fun try_prove_refl_rel ctxt rel = let fun mk_ge_eq x = let val T = fastype_of x in Const (\<^const_name>\less_eq\, T --> T --> HOLogic.boolT) $ (Const (\<^const_name>\HOL.eq\, T)) $ x end; val goal = HOLogic.mk_Trueprop (mk_ge_eq rel); in mono_eq_prover ctxt goal end; fun try_prove_reflexivity ctxt prop = let val cprop = Thm.cterm_of ctxt prop val rule = @{thm ge_eq_refl} val concl_pat = Drule.strip_imp_concl (Thm.cprop_of rule) val insts = Thm.first_order_match (concl_pat, cprop) val rule = Drule.instantiate_normalize insts rule val prop = hd (Thm.prems_of rule) in case mono_eq_prover ctxt prop of SOME thm => SOME (thm RS rule) | NONE => NONE end (* Generates a parametrized transfer rule. transfer_rule - of the form T t f parametric_transfer_rule - of the form par_R t' t Result: par_T t' f, after substituing op= for relations in par_R that relate a type constructor to the same type constructor, it is a merge of (par_R' OO T) t' f using Lifting_Term.merge_transfer_relations *) fun generate_parametric_transfer_rule ctxt0 transfer_rule parametric_transfer_rule = let fun preprocess ctxt thm = let val tm = (strip_args 2 o HOLogic.dest_Trueprop o Thm.concl_of) thm; val param_rel = (snd o dest_comb o fst o dest_comb) tm; val free_vars = Term.add_vars param_rel []; fun make_subst (xi, typ) subst = let val [rty, rty'] = binder_types typ in if Term.is_TVar rty andalso Term.is_Type rty' then (xi, Thm.cterm_of ctxt (HOLogic.eq_const rty')) :: subst else subst end; val inst_thm = infer_instantiate ctxt (fold make_subst free_vars []) thm; in Conv.fconv_rule ((Conv.concl_conv (Thm.nprems_of inst_thm) o HOLogic.Trueprop_conv o Conv.fun2_conv o Conv.arg1_conv) (Raw_Simplifier.rewrite ctxt false (Transfer.get_sym_relator_eq ctxt))) inst_thm end fun inst_relcomppI ctxt ant1 ant2 = let val t1 = (HOLogic.dest_Trueprop o Thm.concl_of) ant1 val t2 = (HOLogic.dest_Trueprop o Thm.prop_of) ant2 val fun1 = Thm.cterm_of ctxt (strip_args 2 t1) val args1 = map (Thm.cterm_of ctxt) (get_args 2 t1) val fun2 = Thm.cterm_of ctxt (strip_args 2 t2) val args2 = map (Thm.cterm_of ctxt) (get_args 1 t2) val relcomppI = Drule.incr_indexes2 ant1 ant2 @{thm relcomppI} val vars = map #1 (rev (Term.add_vars (Thm.prop_of relcomppI) [])) in infer_instantiate ctxt (vars ~~ ([fun1] @ args1 @ [fun2] @ args2)) relcomppI end fun zip_transfer_rules ctxt thm = let fun mk_POS ty = Const (\<^const_name>\POS\, ty --> ty --> HOLogic.boolT) val rel = (Thm.dest_fun2 o Thm.dest_arg o Thm.cprop_of) thm val typ = Thm.typ_of_cterm rel val POS_const = Thm.cterm_of ctxt (mk_POS typ) val var = Thm.cterm_of ctxt (Var (("X", Thm.maxidx_of_cterm rel + 1), typ)) val goal = Thm.apply (Thm.cterm_of ctxt HOLogic.Trueprop) (Thm.apply (Thm.apply POS_const rel) var) in [Lifting_Term.merge_transfer_relations ctxt goal, thm] MRSL @{thm POS_apply} end val thm = inst_relcomppI ctxt0 parametric_transfer_rule transfer_rule OF [parametric_transfer_rule, transfer_rule] val preprocessed_thm = preprocess ctxt0 thm val (fixed_thm, ctxt1) = ctxt0 |> yield_singleton (apfst snd oo Variable.import true) preprocessed_thm val assms = cprems_of fixed_thm val add_transfer_rule = Thm.attribute_declaration Transfer.transfer_add val (prems, ctxt2) = ctxt1 |> fold_map Thm.assume_hyps assms val ctxt3 = ctxt2 |> Context.proof_map (fold add_transfer_rule prems) val zipped_thm = fixed_thm |> undisch_all |> zip_transfer_rules ctxt3 |> implies_intr_list assms |> singleton (Variable.export ctxt3 ctxt0) in zipped_thm end fun print_generate_transfer_info msg = let val error_msg = cat_lines ["Generation of a parametric transfer rule failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in error error_msg end fun map_ter _ x [] = x | map_ter f _ xs = map f xs fun generate_transfer_rules lthy quot_thm rsp_thm def_thm par_thms = let val transfer_rule = ([quot_thm, rsp_thm, def_thm] MRSL @{thm Quotient_to_transfer}) |> Lifting_Term.parametrize_transfer_rule lthy in (map_ter (generate_parametric_transfer_rule lthy transfer_rule) [transfer_rule] par_thms handle Lifting_Term.MERGE_TRANSFER_REL msg => (print_generate_transfer_info msg; [transfer_rule])) end (* Generation of the code certificate from the rsp theorem *) fun get_body_types (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_types (U, V) | get_body_types (U, V) = (U, V) fun get_binder_types (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types (U, W) | get_binder_types _ = [] fun get_binder_types_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [T, U]), Type ("fun", [V, W])) = (T, V) :: get_binder_types_by_rel S (U, W) | get_binder_types_by_rel _ _ = [] fun get_body_type_by_rel (Const (\<^const_name>\rel_fun\, _) $ _ $ S) (Type ("fun", [_, U]), Type ("fun", [_, V])) = get_body_type_by_rel S (U, V) | get_body_type_by_rel _ (U, V) = (U, V) fun get_binder_rels (Const (\<^const_name>\rel_fun\, _) $ R $ S) = R :: get_binder_rels S | get_binder_rels _ = [] fun force_rty_type ctxt rty rhs = let val thy = Proof_Context.theory_of ctxt val rhs_schematic = singleton (Variable.polymorphic ctxt) rhs val rty_schematic = fastype_of rhs_schematic val match = Sign.typ_match thy (rty_schematic, rty) Vartab.empty in Envir.subst_term_types match rhs_schematic end fun unabs_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) fun dest_abs (Abs (var_name, T, _)) = (var_name, T) | dest_abs tm = raise TERM("get_abs_var",[tm]) val (var_name, T) = dest_abs (Thm.term_of rhs) val (new_var_names, ctxt') = Variable.variant_fixes [var_name] ctxt val refl_thm = Thm.reflexive (Thm.cterm_of ctxt' (Free (hd new_var_names, T))) in Thm.combination def refl_thm |> singleton (Proof_Context.export ctxt' ctxt) end fun unabs_all_def ctxt def = let val (_, rhs) = Thm.dest_equals (Thm.cprop_of def) val xs = strip_abs_vars (Thm.term_of rhs) in fold (K (unabs_def ctxt)) xs def end val map_fun_unfolded = @{thm map_fun_def[abs_def]} |> unabs_def \<^context> |> unabs_def \<^context> |> Local_Defs.unfold0 \<^context> [@{thm comp_def}] fun unfold_fun_maps ctm = let fun unfold_conv ctm = case (Thm.term_of ctm) of Const (\<^const_name>\map_fun\, _) $ _ $ _ => (Conv.arg_conv unfold_conv then_conv Conv.rewr_conv map_fun_unfolded) ctm | _ => Conv.all_conv ctm in (Conv.fun_conv unfold_conv) ctm end fun unfold_fun_maps_beta ctm = let val try_beta_conv = Conv.try_conv (Thm.beta_conversion false) in (unfold_fun_maps then_conv try_beta_conv) ctm end fun prove_rel ctxt rsp_thm (rty, qty) = let val ty_args = get_binder_types (rty, qty) fun disch_arg args_ty thm = let val quot_thm = Lifting_Term.prove_quot_thm ctxt args_ty in [quot_thm, thm] MRSL @{thm apply_rsp''} end in fold disch_arg ty_args rsp_thm end exception CODE_CERT_GEN of string fun simplify_code_eq ctxt def_thm = Local_Defs.unfold0 ctxt [@{thm o_apply}, @{thm map_fun_def}, @{thm id_apply}] def_thm (* quot_thm - quotient theorem (Quotient R Abs Rep T). returns: whether the Lifting package is capable to generate code for the abstract type represented by quot_thm *) fun can_generate_code_cert quot_thm = case quot_thm_rel quot_thm of Const (\<^const_name>\HOL.eq\, _) => true | Const (\<^const_name>\eq_onp\, _) $ _ => true | _ => false fun generate_rep_eq ctxt def_thm rsp_thm (rty, qty) = let val unfolded_def = Conv.fconv_rule (Conv.arg_conv unfold_fun_maps_beta) def_thm val unabs_def = unabs_all_def ctxt unfolded_def in if body_type rty = body_type qty then SOME (simplify_code_eq ctxt (HOLogic.mk_obj_eq unabs_def)) else let val quot_thm = Lifting_Term.prove_quot_thm ctxt (get_body_types (rty, qty)) val rel_fun = prove_rel ctxt rsp_thm (rty, qty) val rep_abs_thm = [quot_thm, rel_fun] MRSL @{thm Quotient_rep_abs_eq} in case mono_eq_prover ctxt (hd (Thm.prems_of rep_abs_thm)) of SOME mono_eq_thm => let val rep_abs_eq = mono_eq_thm RS rep_abs_thm val rep = Thm.cterm_of ctxt (quot_thm_rep quot_thm) val rep_refl = HOLogic.mk_obj_eq (Thm.reflexive rep) val repped_eq = [rep_refl, HOLogic.mk_obj_eq unabs_def] MRSL @{thm cong} val code_cert = [repped_eq, rep_abs_eq] MRSL trans in SOME (simplify_code_eq ctxt code_cert) end | NONE => NONE end end fun generate_abs_eq ctxt def_thm rsp_thm quot_thm = let val abs_eq_with_assms = let val (rty, qty) = quot_thm_rty_qty quot_thm val rel = quot_thm_rel quot_thm val ty_args = get_binder_types_by_rel rel (rty, qty) val body_type = get_body_type_by_rel rel (rty, qty) val quot_ret_thm = Lifting_Term.prove_quot_thm ctxt body_type val rep_abs_folded_unmapped_thm = let val rep_id = [quot_thm, def_thm] MRSL @{thm Quotient_Rep_eq} val ctm = Thm.dest_equals_lhs (Thm.cprop_of rep_id) val unfolded_maps_eq = unfold_fun_maps ctm val t1 = [quot_thm, def_thm, rsp_thm] MRSL @{thm Quotient_rep_abs_fold_unmap} val prems_pat = (hd o Drule.cprems_of) t1 val insts = Thm.first_order_match (prems_pat, Thm.cprop_of unfolded_maps_eq) in unfolded_maps_eq RS (Drule.instantiate_normalize insts t1) end in rep_abs_folded_unmapped_thm |> fold (fn _ => fn thm => thm RS @{thm rel_funD2}) ty_args |> (fn x => x RS (@{thm Quotient_rel_abs2} OF [quot_ret_thm])) end val prem_rels = get_binder_rels (quot_thm_rel quot_thm); val proved_assms = prem_rels |> map (try_prove_refl_rel ctxt) |> map_index (apfst (fn x => x + 1)) |> filter (is_some o snd) |> map (apsnd the) |> map (apsnd (fn assm => assm RS @{thm ge_eq_refl})) val abs_eq = fold_rev (fn (i, assm) => fn thm => assm RSN (i, thm)) proved_assms abs_eq_with_assms in simplify_code_eq ctxt abs_eq end fun register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy = let fun no_abstr (t $ u) = no_abstr t andalso no_abstr u | no_abstr (Abs (_, _, t)) = no_abstr t | no_abstr (Const (name, _)) = not (Code.is_abstr thy name) | no_abstr _ = true fun is_valid_eq eqn = can (Code.assert_eqn thy) (mk_meta_eq eqn, true) andalso no_abstr (Thm.prop_of eqn) fun is_valid_abs_eq abs_eq = can (Code.assert_abs_eqn thy NONE) (mk_meta_eq abs_eq) in if is_valid_eq abs_eq_thm then (ABS_EQ, Code.declare_default_eqns_global [(abs_eq_thm, true)] thy) else let val (rty_body, qty_body) = get_body_types (rty, qty) in if rty_body = qty_body then (REP_EQ, Code.declare_default_eqns_global [(the opt_rep_eq_thm, true)] thy) else if is_some opt_rep_eq_thm andalso is_valid_abs_eq (the opt_rep_eq_thm) then (REP_EQ, Code.declare_abstract_eqn_global (the opt_rep_eq_thm) thy) else (NONE_EQ, thy) end end local fun no_no_code ctxt (rty, qty) = if same_type_constrs (rty, qty) then forall (no_no_code ctxt) (Targs rty ~~ Targs qty) else if Term.is_Type qty then if Lifting_Info.is_no_code_type ctxt (Tname qty) then false else let val (rty', rtyq) = Lifting_Term.instantiate_rtys ctxt (rty, qty) val (rty's, rtyqs) = (Targs rty', Targs rtyq) in forall (no_no_code ctxt) (rty's ~~ rtyqs) end else true fun encode_code_eq ctxt abs_eq opt_rep_eq (rty, qty) = let fun mk_type typ = typ |> Logic.mk_type |> Thm.cterm_of ctxt |> Drule.mk_term in Conjunction.intr_balanced [abs_eq, (the_default TrueI opt_rep_eq), mk_type rty, mk_type qty] end exception DECODE fun decode_code_eq thm = if Thm.nprems_of thm > 0 then raise DECODE else let val [abs_eq, rep_eq, rty, qty] = Conjunction.elim_balanced 4 thm val opt_rep_eq = if Thm.eq_thm_prop (rep_eq, TrueI) then NONE else SOME rep_eq fun dest_type typ = typ |> Drule.dest_term |> Thm.term_of |> Logic.dest_type in (abs_eq, opt_rep_eq, (dest_type rty, dest_type qty)) end structure Data = Generic_Data ( type T = code_eq option val empty = NONE val extend = I fun merge _ = NONE ); fun register_encoded_code_eq thm thy = let val (abs_eq_thm, opt_rep_eq_thm, (rty, qty)) = decode_code_eq thm val (code_eq, thy) = register_code_eq_thy abs_eq_thm opt_rep_eq_thm (rty, qty) thy in Context.theory_map (Data.put (SOME code_eq)) thy end handle DECODE => thy val register_code_eq_attribute = Thm.declaration_attribute (fn thm => Context.mapping (register_encoded_code_eq thm) I) val register_code_eq_attrib = Attrib.internal (K register_code_eq_attribute) in fun register_code_eq abs_eq_thm opt_rep_eq_thm (rty, qty) lthy = let val encoded_code_eq = encode_code_eq lthy abs_eq_thm opt_rep_eq_thm (rty, qty) in if no_no_code lthy (rty, qty) then let val lthy' = lthy |> (#2 oo Local_Theory.note) ((Binding.empty, [register_code_eq_attrib]), [encoded_code_eq]) val opt_code_eq = Data.get (Context.Theory (Proof_Context.theory_of lthy')) val code_eq = if is_some opt_code_eq then the opt_code_eq else UNKNOWN_EQ (* UNKNOWN_EQ means that we are in a locale and we do not know which code equation is going to be used. This is going to be resolved at the point when an interpretation of the locale is executed. *) val lthy'' = lthy' |> Local_Theory.declaration {syntax = false, pervasive = true} (K (Data.put NONE)) in (code_eq, lthy'') end else (NONE_EQ, lthy) end end (* Defines an operation on an abstract type in terms of a corresponding operation on a representation type. var - a binding and a mixfix of the new constant being defined qty - an abstract type of the new constant rhs - a term representing the new constant on the raw level rsp_thm - a respectfulness theorem in the internal tagged form (like '(R ===> R ===> R) f f'), i.e. "(Lifting_Term.equiv_relation (fastype_of rhs, qty)) $ rhs $ rhs" par_thms - a parametricity theorem for rhs *) fun add_lift_def (config: config) (binding, mx) qty rhs rsp_thm par_thms lthy0 = let val rty = fastype_of rhs val quot_thm = Lifting_Term.prove_quot_thm lthy0 (rty, qty) val absrep_trm = quot_thm_abs quot_thm val rty_forced = (domain_type o fastype_of) absrep_trm val forced_rhs = force_rty_type lthy0 rty_forced rhs val lhs = Free (Binding.name_of binding, qty) val prop = Logic.mk_equals (lhs, absrep_trm $ forced_rhs) val (_, prop') = Local_Defs.cert_def lthy0 (K []) prop val (_, newrhs) = Local_Defs.abs_def prop' val var = ((#notes config = false ? Binding.concealed) binding, mx) val def_name = Thm.make_def_binding (#notes config) (#1 var) val ((lift_const, (_ , def_thm)), lthy1) = lthy0 |> Local_Theory.define (var, ((def_name, []), newrhs)) val transfer_rules = generate_transfer_rules lthy1 quot_thm rsp_thm def_thm par_thms val abs_eq_thm = generate_abs_eq lthy1 def_thm rsp_thm quot_thm val opt_rep_eq_thm = generate_rep_eq lthy1 def_thm rsp_thm (rty_forced, qty) fun notes names = let val lhs_name = Binding.reset_pos (#1 var) val rsp_thmN = Binding.qualify_name true lhs_name "rsp" val abs_eq_thmN = Binding.qualify_name true lhs_name "abs_eq" val rep_eq_thmN = Binding.qualify_name true lhs_name "rep_eq" val transfer_ruleN = Binding.qualify_name true lhs_name "transfer" val notes = [(rsp_thmN, [], [rsp_thm]), (transfer_ruleN, @{attributes [transfer_rule]}, transfer_rules), (abs_eq_thmN, [], [abs_eq_thm])] @ (case opt_rep_eq_thm of SOME rep_eq_thm => [(rep_eq_thmN, [], [rep_eq_thm])] | NONE => []) in if names then map (fn (name, attrs, thms) => ((name, []), [(thms, attrs)])) notes else map_filter (fn (_, attrs, thms) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) notes end val (code_eq, lthy2) = lthy1 |> register_code_eq abs_eq_thm opt_rep_eq_thm (rty_forced, qty) val lift_def = mk_lift_def rty_forced qty newrhs lift_const def_thm rsp_thm abs_eq_thm opt_rep_eq_thm code_eq transfer_rules in lthy2 |> (snd o Local_Theory.begin_nested) |> Local_Theory.notes (notes (#notes config)) |> snd |> `(fn lthy => morph_lift_def (Local_Theory.target_morphism lthy) lift_def) ||> Local_Theory.end_nested end (* This is not very cheap way of getting the rules but we have only few active liftings in the current setting *) fun get_cr_pcr_eqs ctxt = let fun collect (data : Lifting_Info.quotient) l = if is_some (#pcr_info data) then ((Thm.symmetric o safe_mk_meta_eq o Thm.transfer' ctxt o #pcr_cr_eq o the o #pcr_info) data :: l) else l val table = Lifting_Info.get_quotients ctxt in Symtab.fold (fn (_, data) => fn l => collect data l) table [] end fun prepare_lift_def add_lift_def var qty rhs par_thms ctxt = let val rsp_rel = Lifting_Term.equiv_relation ctxt (fastype_of rhs, qty) val rty_forced = (domain_type o fastype_of) rsp_rel; val forced_rhs = force_rty_type ctxt rty_forced rhs; val cr_to_pcr_conv = HOLogic.Trueprop_conv (Conv.fun2_conv (Raw_Simplifier.rewrite ctxt false (get_cr_pcr_eqs ctxt))) val (prsp_tm, rsp_prsp_eq) = HOLogic.mk_Trueprop (rsp_rel $ forced_rhs $ forced_rhs) |> Thm.cterm_of ctxt |> cr_to_pcr_conv |> `Thm.concl_of |>> Logic.dest_equals |>> snd val to_rsp = rsp_prsp_eq RS Drule.equal_elim_rule2 val opt_proven_rsp_thm = try_prove_reflexivity ctxt prsp_tm fun after_qed internal_rsp_thm = add_lift_def var qty rhs (internal_rsp_thm RS to_rsp) par_thms in case opt_proven_rsp_thm of SOME thm => (NONE, K (after_qed thm)) | NONE => (SOME prsp_tm, after_qed) end fun gen_lift_def add_lift_def var qty rhs tac par_thms lthy = let val (goal, after_qed) = prepare_lift_def add_lift_def var qty rhs par_thms lthy in case goal of SOME goal => let val rsp_thm = Goal.prove_sorry lthy [] [] goal (tac o #context) |> Thm.close_derivation \<^here> in after_qed rsp_thm lthy end | NONE => after_qed Drule.dummy_thm lthy end val lift_def = gen_lift_def o add_lift_def end (* structure *) diff --git a/src/HOL/Tools/Lifting/lifting_term.ML b/src/HOL/Tools/Lifting/lifting_term.ML --- a/src/HOL/Tools/Lifting/lifting_term.ML +++ b/src/HOL/Tools/Lifting/lifting_term.ML @@ -1,592 +1,592 @@ (* Title: HOL/Tools/Lifting/lifting_term.ML Author: Ondrej Kuncar Proves Quotient theorem. *) signature LIFTING_TERM = sig exception QUOT_THM of typ * typ * Pretty.T exception PARAM_QUOT_THM of typ * Pretty.T exception MERGE_TRANSFER_REL of Pretty.T exception CHECK_RTY of typ * typ type 'a fold_quot_thm = { constr: typ -> thm * 'a -> thm * 'a, lift: typ -> thm * 'a -> thm * 'a, comp_lift: typ -> thm * 'a -> thm * 'a } type quotients = Lifting_Info.quotient Symtab.table val force_qty_type: Proof.context -> typ -> thm -> thm val prove_schematic_quot_thm: 'a fold_quot_thm -> quotients -> Proof.context -> typ * typ -> 'a -> thm * 'a val instantiate_rtys: Proof.context -> typ * typ -> typ * typ val prove_quot_thm: Proof.context -> typ * typ -> thm val abs_fun: Proof.context -> typ * typ -> term val equiv_relation: Proof.context -> typ * typ -> term val prove_param_quot_thm: Proof.context -> typ -> thm * (typ * thm) list * Proof.context val generate_parametrized_relator: Proof.context -> typ -> term * term list val merge_transfer_relations: Proof.context -> cterm -> thm val parametrize_transfer_rule: Proof.context -> thm -> thm end structure Lifting_Term: LIFTING_TERM = struct open Lifting_Util infix 0 MRSL exception QUOT_THM_INTERNAL of Pretty.T exception QUOT_THM of typ * typ * Pretty.T exception PARAM_QUOT_THM of typ * Pretty.T exception MERGE_TRANSFER_REL of Pretty.T exception CHECK_RTY of typ * typ type quotients = Lifting_Info.quotient Symtab.table fun match ctxt err ty_pat ty = Sign.typ_match (Proof_Context.theory_of ctxt) (ty_pat, ty) Vartab.empty handle Type.TYPE_MATCH => err ctxt ty_pat ty fun equiv_match_err ctxt ty_pat ty = let val ty_pat_str = Syntax.string_of_typ ctxt ty_pat val ty_str = Syntax.string_of_typ ctxt ty in raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("The quotient type " ^ quote ty_str), Pretty.brk 1, Pretty.str ("and the quotient type pattern " ^ quote ty_pat_str), Pretty.brk 1, Pretty.str "don't match."]) end fun get_quot_data (quotients: quotients) s = case Symtab.lookup quotients s of SOME qdata => qdata | NONE => raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("No quotient type " ^ quote s), Pretty.brk 1, Pretty.str "found."]) fun get_quot_thm quotients ctxt s = Thm.transfer' ctxt (#quot_thm (get_quot_data quotients s)) fun has_pcrel_info quotients s = is_some (#pcr_info (get_quot_data quotients s)) fun get_pcrel_info quotients s = case #pcr_info (get_quot_data quotients s) of SOME pcr_info => pcr_info | NONE => raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("No parametrized correspondce relation for " ^ quote s), Pretty.brk 1, Pretty.str "found."]) fun get_pcrel_def quotients ctxt s = Thm.transfer' ctxt (#pcrel_def (get_pcrel_info quotients s)) fun get_pcr_cr_eq quotients ctxt s = Thm.transfer' ctxt (#pcr_cr_eq (get_pcrel_info quotients s)) fun get_rel_quot_thm ctxt s = (case Lifting_Info.lookup_quot_maps ctxt s of SOME map_data => Thm.transfer' ctxt (#rel_quot_thm map_data) | NONE => raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("No relator for the type " ^ quote s), Pretty.brk 1, Pretty.str "found."])) fun get_rel_distr_rules ctxt s tm = (case Lifting_Info.lookup_relator_distr_data ctxt s of SOME rel_distr_thm => (case tm of Const (\<^const_name>\POS\, _) => map (Thm.transfer' ctxt) (#pos_distr_rules rel_distr_thm) | Const (\<^const_name>\NEG\, _) => map (Thm.transfer' ctxt) (#neg_distr_rules rel_distr_thm)) | NONE => raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("No relator distr. data for the type " ^ quote s), Pretty.brk 1, Pretty.str "found."])) fun is_id_quot thm = Thm.eq_thm_prop (thm, @{thm identity_quotient}) fun zip_Tvars ctxt0 type_name rty_Tvars qty_Tvars = case try (get_rel_quot_thm ctxt0) type_name of NONE => rty_Tvars ~~ qty_Tvars | SOME rel_quot_thm => let fun quot_term_absT quot_term = let val (_, abs, _, _) = (dest_Quotient o HOLogic.dest_Trueprop) quot_term in fastype_of abs end fun equiv_univ_err ctxt ty_pat ty = let val ty_pat_str = Syntax.string_of_typ ctxt ty_pat val ty_str = Syntax.string_of_typ ctxt ty in raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str ("The type " ^ quote ty_str), Pretty.brk 1, Pretty.str ("and the relator type pattern " ^ quote ty_pat_str), Pretty.brk 1, Pretty.str "don't unify."]) end fun raw_match (TVar (v, S), T) subs = (case Vartab.defined subs v of false => Vartab.update_new (v, (S, T)) subs | true => subs) | raw_match (Type (_, Ts), Type (_, Us)) subs = raw_matches (Ts, Us) subs | raw_match _ subs = subs and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs) | raw_matches _ subs = subs val rty = Type (type_name, rty_Tvars) val qty = Type (type_name, qty_Tvars) val rel_quot_thm_concl = (Logic.strip_imp_concl o Thm.prop_of) rel_quot_thm val schematic_rel_absT = quot_term_absT rel_quot_thm_concl; val absT = rty --> qty val schematic_absT = absT |> Logic.type_map (singleton (Variable.polymorphic ctxt0)) |> Logic.incr_tvar (maxidx_of_typ schematic_rel_absT + 1) (* because absT can already contain schematic variables from rty patterns *) val maxidx = Term.maxidx_of_typs [schematic_rel_absT, schematic_absT] val _ = Sign.typ_unify (Proof_Context.theory_of ctxt0) (schematic_rel_absT, schematic_absT) (Vartab.empty, maxidx) handle Type.TUNIFY => equiv_univ_err ctxt0 schematic_rel_absT schematic_absT val subs = raw_match (schematic_rel_absT, absT) Vartab.empty val rel_quot_thm_prems = (Logic.strip_imp_prems o Thm.prop_of) rel_quot_thm in map (dest_funT o Envir.subst_type subs o quot_term_absT) rel_quot_thm_prems end fun gen_rty_is_TVar quotients ctxt qty = qty |> Tname |> get_quot_thm quotients ctxt |> quot_thm_rty_qty |> fst |> is_TVar fun gen_instantiate_rtys quotients ctxt (rty, (qty as Type (qty_name, _))) = let val quot_thm = get_quot_thm quotients ctxt qty_name val (rty_pat, qty_pat) = quot_thm_rty_qty quot_thm fun inst_rty (Type (s, tys), Type (s', tys')) = if s = s' then Type (s', map inst_rty (tys ~~ tys')) else raise QUOT_THM_INTERNAL (Pretty.block [Pretty.str "The type", Pretty.brk 1, Syntax.pretty_typ ctxt rty, Pretty.brk 1, Pretty.str ("is not a raw type for the quotient type " ^ quote qty_name ^ ";"), Pretty.brk 1, Pretty.str "the correct raw type must be an instance of", Pretty.brk 1, Syntax.pretty_typ ctxt rty_pat]) | inst_rty (t as Type (_, _), TFree _) = t | inst_rty ((TVar _), rty) = rty | inst_rty ((TFree _), rty) = rty | inst_rty (_, _) = error "check_raw_types: we should not be here" val qtyenv = match ctxt equiv_match_err qty_pat qty in (inst_rty (rty_pat, rty), Envir.subst_type qtyenv rty_pat) end | gen_instantiate_rtys _ _ _ = error "gen_instantiate_rtys: not Type" fun instantiate_rtys ctxt (rty, qty) = gen_instantiate_rtys (Lifting_Info.get_quotients ctxt) ctxt (rty, qty) type 'a fold_quot_thm = { constr: typ -> thm * 'a -> thm * 'a, lift: typ -> thm * 'a -> thm * 'a, comp_lift: typ -> thm * 'a -> thm * 'a } fun prove_schematic_quot_thm (actions: 'a fold_quot_thm) quotients ctxt (rty, qty) fold_val = let fun lifting_step (rty, qty) = let val (rty', rtyq) = gen_instantiate_rtys quotients ctxt (rty, qty) val (rty's, rtyqs) = if gen_rty_is_TVar quotients ctxt qty then ([rty'],[rtyq]) else (Targs rty', Targs rtyq) val (args, fold_val) = fold_map (prove_schematic_quot_thm actions quotients ctxt) (rty's ~~ rtyqs) fold_val in if forall is_id_quot args then let val quot_thm = get_quot_thm quotients ctxt (Tname qty) in #lift actions qty (quot_thm, fold_val) end else let val quot_thm = get_quot_thm quotients ctxt (Tname qty) val rel_quot_thm = if gen_rty_is_TVar quotients ctxt qty then the_single args else args MRSL (get_rel_quot_thm ctxt (Tname rty)) val comp_quot_thm = [rel_quot_thm, quot_thm] MRSL @{thm Quotient_compose} in #comp_lift actions qty (comp_quot_thm, fold_val) end end in (case (rty, qty) of (Type (s, tys), Type (s', tys')) => if s = s' then let val (args, fold_val) = fold_map (prove_schematic_quot_thm actions quotients ctxt) (zip_Tvars ctxt s tys tys') fold_val in if forall is_id_quot args then (@{thm identity_quotient}, fold_val) else let val quot_thm = args MRSL (get_rel_quot_thm ctxt s) in #constr actions qty (quot_thm, fold_val) end end else lifting_step (rty, qty) | (_, Type (s', tys')) => (case try (get_quot_thm quotients ctxt) s' of SOME quot_thm => let val rty_pat = (fst o quot_thm_rty_qty) quot_thm in lifting_step (rty_pat, qty) end | NONE => let val rty_pat = Type (s', map (fn _ => TFree ("a",[])) tys') in prove_schematic_quot_thm actions quotients ctxt (rty_pat, qty) fold_val end) | _ => (@{thm identity_quotient}, fold_val) ) end handle QUOT_THM_INTERNAL pretty_msg => raise QUOT_THM (rty, qty, pretty_msg) fun force_qty_type ctxt qty quot_thm = let val (_, qty_schematic) = quot_thm_rty_qty quot_thm val match_env = Sign.typ_match (Proof_Context.theory_of ctxt) (qty_schematic, qty) Vartab.empty fun prep_ty (x, (S, ty)) = ((x, S), Thm.ctyp_of ctxt ty) val ty_inst = Vartab.fold (cons o prep_ty) match_env [] - in Thm.instantiate (ty_inst, []) quot_thm end + in Thm.instantiate (TVars.make ty_inst, Vars.empty) quot_thm end fun check_rty_type ctxt rty quot_thm = let val (rty_forced, _) = quot_thm_rty_qty quot_thm val rty_schematic = Logic.type_map (singleton (Variable.polymorphic ctxt)) rty val _ = Sign.typ_match (Proof_Context.theory_of ctxt) (rty_schematic, rty_forced) Vartab.empty handle Type.TYPE_MATCH => raise CHECK_RTY (rty_schematic, rty_forced) in () end (* The function tries to prove that rty and qty form a quotient. Returns: Quotient theorem; an abstract type of the theorem is exactly qty, a representation type of the theorem is an instance of rty in general. *) local val id_actions = { constr = K I, lift = K I, comp_lift = K I } in fun prove_quot_thm ctxt (rty, qty) = let val quotients = Lifting_Info.get_quotients ctxt val (schematic_quot_thm, _) = prove_schematic_quot_thm id_actions quotients ctxt (rty, qty) () val quot_thm = force_qty_type ctxt qty schematic_quot_thm val _ = check_rty_type ctxt rty quot_thm in quot_thm end end (* Computes the composed abstraction function for rty and qty. *) fun abs_fun ctxt (rty, qty) = quot_thm_abs (prove_quot_thm ctxt (rty, qty)) (* Computes the composed equivalence relation for rty and qty. *) fun equiv_relation ctxt (rty, qty) = quot_thm_rel (prove_quot_thm ctxt (rty, qty)) val get_fresh_Q_t = let val Q_t = \<^term>\Trueprop (Quotient R Abs Rep T)\ val frees_Q_t = Term.add_free_names Q_t [] val tfrees_Q_t = rev (Term.add_tfree_names Q_t []) in fn ctxt => let fun rename_free_var tab (Free (name, typ)) = Free (the_default name (AList.lookup op= tab name), typ) | rename_free_var _ t = t fun rename_free_vars tab = map_aterms (rename_free_var tab) fun rename_free_tvars tab = map_types (map_type_tfree (fn (name, sort) => TFree (the_default name (AList.lookup op= tab name), sort))) val (new_frees_Q_t, ctxt') = Variable.variant_fixes frees_Q_t ctxt val tab_frees = frees_Q_t ~~ new_frees_Q_t val (new_tfrees_Q_t, ctxt'') = Variable.invent_types (replicate (length tfrees_Q_t) []) ctxt' val tab_tfrees = tfrees_Q_t ~~ (fst o split_list) new_tfrees_Q_t val renamed_Q_t = rename_free_vars tab_frees Q_t val renamed_Q_t = rename_free_tvars tab_tfrees renamed_Q_t in (renamed_Q_t, ctxt'') end end (* For the given type, it proves a composed Quotient map theorem, where for each type variable extra Quotient assumption is generated. E.g., for 'a list it generates exactly the Quotient map theorem for the list type. The function generalizes this for the whole type universe. New fresh variables in the assumptions are fixed in the returned context. Returns: the composed Quotient map theorem and list mapping each type variable in ty to the corresponding assumption in the returned theorem. *) fun prove_param_quot_thm ctxt0 ty = let fun generate (ty as Type (s, tys)) (table, ctxt) = if null tys then let val instantiated_id_quot_thm = Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)] [] @{thm identity_quotient} in (instantiated_id_quot_thm, (table, ctxt)) end else let val (args, table_ctxt') = fold_map generate tys (table, ctxt) in (args MRSL (get_rel_quot_thm ctxt s), table_ctxt') end | generate ty (table, ctxt) = if AList.defined (op =) table ty then (the (AList.lookup (op =) table ty), (table, ctxt)) else let val (Q_t, ctxt') = get_fresh_Q_t ctxt val Q_thm = Thm.assume (Thm.cterm_of ctxt' Q_t) val table' = (ty, Q_thm) :: table in (Q_thm, (table', ctxt')) end val (param_quot_thm, (table, ctxt1)) = generate ty ([], ctxt0) in (param_quot_thm, rev table, ctxt1) end handle QUOT_THM_INTERNAL pretty_msg => raise PARAM_QUOT_THM (ty, pretty_msg) (* It computes a parametrized relator for the given type ty. E.g., for 'a dlist: list_all2 ?R OO cr_dlist with parameters [?R]. Returns: the definitional term and list of parameters (relations). *) fun generate_parametrized_relator ctxt ty = let val (quot_thm, table, ctxt') = prove_param_quot_thm ctxt ty val parametrized_relator = quot_thm_crel quot_thm val args = map (fn (_, q_thm) => quot_thm_crel q_thm) table val exported_terms = Variable.exportT_terms ctxt' ctxt (parametrized_relator :: args) in (hd exported_terms, tl exported_terms) end (* Parametrization *) local fun get_lhs rule = (Thm.dest_fun o Thm.dest_arg o strip_imp_concl o Thm.cprop_of) rule; fun no_imp _ = raise CTERM ("no implication", []); infix 0 else_imp fun (cv1 else_imp cv2) ct = (cv1 ct handle THM _ => cv2 ct | CTERM _ => cv2 ct | TERM _ => cv2 ct | TYPE _ => cv2 ct); fun first_imp cvs = fold_rev (curry op else_imp) cvs no_imp fun rewr_imp rule ct = let val rule1 = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) rule; val lhs_rule = get_lhs rule1; val rule2 = Thm.rename_boundvars (Thm.term_of lhs_rule) (Thm.term_of ct) rule1; val lhs_ct = Thm.dest_fun ct in Thm.instantiate (Thm.match (lhs_rule, lhs_ct)) rule2 handle Pattern.MATCH => raise CTERM ("rewr_imp", [lhs_rule, lhs_ct]) end fun rewrs_imp rules = first_imp (map rewr_imp rules) in fun gen_merge_transfer_relations quotients ctxt0 ctm = let val ctm = Thm.dest_arg ctm val tm = Thm.term_of ctm val rel = (hd o get_args 2) tm fun same_constants (Const (n1,_)) (Const (n2,_)) = n1 = n2 | same_constants _ _ = false fun prove_extra_assms ctxt ctm distr_rule = let fun prove_assm assm = try (Goal.prove ctxt [] [] (Thm.term_of assm)) (fn {context = goal_ctxt, ...} => SOLVED' (REPEAT_ALL_NEW (resolve_tac goal_ctxt (Transfer.get_transfer_raw goal_ctxt))) 1) fun is_POS_or_NEG ctm = case (head_of o Thm.term_of o Thm.dest_arg) ctm of Const (\<^const_name>\POS\, _) => true | Const (\<^const_name>\NEG\, _) => true | _ => false val inst_distr_rule = rewr_imp distr_rule ctm val extra_assms = filter_out is_POS_or_NEG (cprems_of inst_distr_rule) val proved_assms = map_interrupt prove_assm extra_assms in Option.map (curry op OF inst_distr_rule) proved_assms end handle CTERM _ => NONE fun cannot_merge_error_msg () = Pretty.block [Pretty.str "Rewriting (merging) of this term has failed:", Pretty.brk 1, Syntax.pretty_term ctxt0 rel] in case get_args 2 rel of [Const (\<^const_name>\HOL.eq\, _), _] => rewrs_imp @{thms neg_eq_OO pos_eq_OO} ctm | [_, Const (\<^const_name>\HOL.eq\, _)] => rewrs_imp @{thms neg_OO_eq pos_OO_eq} ctm | [_, trans_rel] => let val (rty', qty) = (relation_types o fastype_of) trans_rel in if same_type_constrs (rty', qty) then let val distr_rules = get_rel_distr_rules ctxt0 ((fst o dest_Type) rty') (head_of tm) val distr_rule = get_first (prove_extra_assms ctxt0 ctm) distr_rules in case distr_rule of NONE => raise MERGE_TRANSFER_REL (cannot_merge_error_msg ()) | SOME distr_rule => map (gen_merge_transfer_relations quotients ctxt0) (cprems_of distr_rule) MRSL distr_rule end else let val pcrel_def = get_pcrel_def quotients ctxt0 ((fst o dest_Type) qty) val pcrel_const = (head_of o fst o Logic.dest_equals o Thm.prop_of) pcrel_def in if same_constants pcrel_const (head_of trans_rel) then let val unfolded_ctm = Thm.rhs_of (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv pcrel_def)) ctm) val distr_rule = rewrs_imp @{thms POS_pcr_rule NEG_pcr_rule} unfolded_ctm val result = (map (gen_merge_transfer_relations quotients ctxt0) (cprems_of distr_rule)) MRSL distr_rule val fold_pcr_rel = Conv.rewr_conv (Thm.symmetric pcrel_def) in Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.combination_conv (Conv.arg_conv (Conv.arg_conv fold_pcr_rel)) fold_pcr_rel)) result end else raise MERGE_TRANSFER_REL (Pretty.str "Non-parametric correspondence relation used.") end end end handle QUOT_THM_INTERNAL pretty_msg => raise MERGE_TRANSFER_REL pretty_msg (* ctm - of the form "[POS|NEG] (par_R OO T) t f) ?X", where par_R is a parametricity transfer relation for t and T is a transfer relation between t and f, which consists only from parametrized transfer relations (i.e., pcr_?) and equalities op=. POS or NEG encodes co-variance or contra-variance. The function merges par_R OO T using definitions of parametrized correspondence relations (e.g., (rel_S R) OO (pcr_T op=) --> pcr_T R using the definition pcr_T R = (rel_S R) OO cr_T). *) fun merge_transfer_relations ctxt ctm = gen_merge_transfer_relations (Lifting_Info.get_quotients ctxt) ctxt ctm end fun gen_parametrize_transfer_rule quotients ctxt thm = let fun parametrize_relation_conv ctm = let val (rty, qty) = (relation_types o fastype_of) (Thm.term_of ctm) in if same_type_constrs (rty, qty) then if forall op= (Targs rty ~~ Targs qty) then Conv.all_conv ctm else all_args_conv parametrize_relation_conv ctm else if is_Type qty then let val q = (fst o dest_Type) qty in let val (rty', rtyq) = gen_instantiate_rtys quotients ctxt (rty, qty) val (rty's, rtyqs) = if gen_rty_is_TVar quotients ctxt qty then ([rty'],[rtyq]) else (Targs rty', Targs rtyq) in if forall op= (rty's ~~ rtyqs) then let val pcr_cr_eq = (Thm.symmetric o mk_meta_eq) (get_pcr_cr_eq quotients ctxt q) in Conv.rewr_conv pcr_cr_eq ctm end handle QUOT_THM_INTERNAL _ => Conv.all_conv ctm else if has_pcrel_info quotients q then let val pcrel_def = Thm.symmetric (get_pcrel_def quotients ctxt q) in (Conv.rewr_conv pcrel_def then_conv all_args_conv parametrize_relation_conv) ctm end else Conv.arg1_conv (all_args_conv parametrize_relation_conv) ctm end end else Conv.all_conv ctm end in Conv.fconv_rule (HOLogic.Trueprop_conv (Conv.fun2_conv parametrize_relation_conv)) thm end (* It replaces cr_T by pcr_T op= in the transfer relation. For composed abstract types, it replaces T_rel R OO cr_T by pcr_T R. If the parametrized correspondce relation does not exist, the original relation is kept. thm - a transfer rule *) fun parametrize_transfer_rule ctxt thm = gen_parametrize_transfer_rule (Lifting_Info.get_quotients ctxt) ctxt thm end diff --git a/src/HOL/Tools/Meson/meson.ML b/src/HOL/Tools/Meson/meson.ML --- a/src/HOL/Tools/Meson/meson.ML +++ b/src/HOL/Tools/Meson/meson.ML @@ -1,797 +1,798 @@ (* Title: HOL/Tools/Meson/meson.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen The MESON resolution proof procedure for HOL. When making clauses, avoids using the rewriter -- instead uses RS recursively. *) signature MESON = sig type simp_options = {if_simps : bool, let_simps : bool} val simp_options_all_true : simp_options val trace : bool Config.T val max_clauses : int Config.T val first_order_resolve : Proof.context -> thm -> thm -> thm val size_of_subgoals: thm -> int val has_too_many_clauses: Proof.context -> term -> bool val make_cnf: thm list -> thm -> Proof.context -> thm list * Proof.context val finish_cnf: thm list -> thm list val presimplified_consts : string list val presimplify: simp_options -> Proof.context -> thm -> thm val make_nnf: simp_options -> Proof.context -> thm -> thm val choice_theorems : theory -> thm list val skolemize_with_choice_theorems : simp_options -> Proof.context -> thm list -> thm -> thm val skolemize : simp_options -> Proof.context -> thm -> thm val cong_extensionalize_thm : Proof.context -> thm -> thm val abs_extensionalize_conv : Proof.context -> conv val abs_extensionalize_thm : Proof.context -> thm -> thm val make_clauses_unsorted: Proof.context -> thm list -> thm list val make_clauses: Proof.context -> thm list -> thm list val make_horns: thm list -> thm list val best_prolog_tac: Proof.context -> (thm -> int) -> thm list -> tactic val depth_prolog_tac: Proof.context -> thm list -> tactic val gocls: thm list -> thm list val skolemize_prems_tac : simp_options -> Proof.context -> thm list -> int -> tactic val MESON: tactic -> (thm list -> thm list) -> (thm list -> tactic) -> Proof.context -> int -> tactic val best_meson_tac: (thm -> int) -> Proof.context -> int -> tactic val safe_best_meson_tac: Proof.context -> int -> tactic val depth_meson_tac: Proof.context -> int -> tactic val prolog_step_tac': Proof.context -> thm list -> int -> tactic val iter_deepen_prolog_tac: Proof.context -> thm list -> tactic val iter_deepen_meson_tac: Proof.context -> thm list -> int -> tactic val make_meta_clause: Proof.context -> thm -> thm val make_meta_clauses: Proof.context -> thm list -> thm list val meson_tac: Proof.context -> thm list -> int -> tactic end structure Meson : MESON = struct type simp_options = {if_simps : bool, let_simps : bool} val simp_options_all_true = {if_simps = true, let_simps = true} val trace = Attrib.setup_config_bool \<^binding>\meson_trace\ (K false) fun trace_msg ctxt msg = if Config.get ctxt trace then tracing (msg ()) else () val max_clauses = Attrib.setup_config_int \<^binding>\meson_max_clauses\ (K 60) (*No known example (on 1-5-2007) needs even thirty*) val iter_deepen_limit = 50; val disj_forward = @{thm disj_forward}; val disj_forward2 = @{thm disj_forward2}; val make_pos_rule = @{thm make_pos_rule}; val make_pos_rule' = @{thm make_pos_rule'}; val make_pos_goal = @{thm make_pos_goal}; val make_neg_rule = @{thm make_neg_rule}; val make_neg_rule' = @{thm make_neg_rule'}; val make_neg_goal = @{thm make_neg_goal}; val conj_forward = @{thm conj_forward}; val all_forward = @{thm all_forward}; val ex_forward = @{thm ex_forward}; val not_conjD = @{thm not_conjD}; val not_disjD = @{thm not_disjD}; val not_notD = @{thm not_notD}; val not_allD = @{thm not_allD}; val not_exD = @{thm not_exD}; val imp_to_disjD = @{thm imp_to_disjD}; val not_impD = @{thm not_impD}; val iff_to_disjD = @{thm iff_to_disjD}; val not_iffD = @{thm not_iffD}; val conj_exD1 = @{thm conj_exD1}; val conj_exD2 = @{thm conj_exD2}; val disj_exD = @{thm disj_exD}; val disj_exD1 = @{thm disj_exD1}; val disj_exD2 = @{thm disj_exD2}; val disj_assoc = @{thm disj_assoc}; val disj_comm = @{thm disj_comm}; val disj_FalseD1 = @{thm disj_FalseD1}; val disj_FalseD2 = @{thm disj_FalseD2}; (**** Operators for forward proof ****) (** First-order Resolution **) (*FIXME: currently does not "rename variables apart"*) fun first_order_resolve ctxt thA thB = (case \<^try>\ let val thy = Proof_Context.theory_of ctxt val tmA = Thm.concl_of thA val Const(\<^const_name>\Pure.imp\,_) $ tmB $ _ = Thm.prop_of thB val tenv = Pattern.first_order_match thy (tmB, tmA) (Vartab.empty, Vartab.empty) |> snd val insts = Vartab.fold (fn (xi, (_, t)) => cons (xi, Thm.cterm_of ctxt t)) tenv []; in thA RS (infer_instantiate ctxt insts thB) end\ of SOME th => th | NONE => raise THM ("first_order_resolve", 0, [thA, thB])) (* Hack to make it less likely that we lose our precious bound variable names in "rename_bound_vars_RS" below, because of a clash. *) val protect_prefix = "Meson_xyzzy" fun protect_bound_var_names (t $ u) = protect_bound_var_names t $ protect_bound_var_names u | protect_bound_var_names (Abs (s, T, t')) = Abs (protect_prefix ^ s, T, protect_bound_var_names t') | protect_bound_var_names t = t fun fix_bound_var_names old_t new_t = let fun quant_of \<^const_name>\All\ = SOME true | quant_of \<^const_name>\Ball\ = SOME true | quant_of \<^const_name>\Ex\ = SOME false | quant_of \<^const_name>\Bex\ = SOME false | quant_of _ = NONE val flip_quant = Option.map not fun some_eq (SOME x) (SOME y) = x = y | some_eq _ _ = false fun add_names quant (Const (quant_s, _) $ Abs (s, _, t')) = add_names quant t' #> some_eq quant (quant_of quant_s) ? cons s | add_names quant (\<^const>\Not\ $ t) = add_names (flip_quant quant) t | add_names quant (\<^const>\implies\ $ t1 $ t2) = add_names (flip_quant quant) t1 #> add_names quant t2 | add_names quant (t1 $ t2) = fold (add_names quant) [t1, t2] | add_names _ _ = I fun lost_names quant = subtract (op =) (add_names quant new_t []) (add_names quant old_t []) fun aux ((t1 as Const (quant_s, _)) $ (Abs (s, T, t'))) = t1 $ Abs (s |> String.isPrefix protect_prefix s ? perhaps (try (fn _ => hd (lost_names (quant_of quant_s)))), T, aux t') | aux (t1 $ t2) = aux t1 $ aux t2 | aux t = t in aux new_t end (* Forward proof while preserving bound variables names *) fun rename_bound_vars_RS th rl = let val t = Thm.concl_of th val r = Thm.concl_of rl val th' = th RS Thm.rename_boundvars r (protect_bound_var_names r) rl val t' = Thm.concl_of th' in Thm.rename_boundvars t' (fix_bound_var_names t t') th' end (*raises exception if no rules apply*) fun tryres (th, rls) = let fun tryall [] = raise THM("tryres", 0, th::rls) | tryall (rl::rls) = (rename_bound_vars_RS th rl handle THM _ => tryall rls) in tryall rls end; (* Special version of "resolve_tac" that works around an explosion in the unifier. If the goal has the form "?P c", the danger is that resolving it against a property of the form "... c ... c ... c ..." will lead to a huge unification problem, due to the (spurious) choices between projection and imitation. The workaround is to instantiate "?P := (%c. ... c ... c ... c ...)" manually. *) fun quant_resolve_tac ctxt th i st = case (Thm.concl_of st, Thm.prop_of th) of (\<^const>\Trueprop\ $ (Var _ $ (c as Free _)), \<^const>\Trueprop\ $ _) => let val cc = Thm.cterm_of ctxt c val ct = Thm.dest_arg (Thm.cprop_of th) in resolve_tac ctxt [th] i (Thm.instantiate' [] [SOME (Thm.lambda cc ct)] st) end | _ => resolve_tac ctxt [th] i st (*Permits forward proof from rules that discharge assumptions. The supplied proof state st, e.g. from conj_forward, should have the form "[| P' ==> ?P; Q' ==> ?Q |] ==> ?P & ?Q" and the effect should be to instantiate ?P and ?Q with normalized versions of P' and Q'.*) fun forward_res ctxt nf st = let fun tacf [prem] = quant_resolve_tac ctxt (nf prem) 1 | tacf prems = error (cat_lines ("Bad proof state in forward_res, please inform lcp@cl.cam.ac.uk:" :: Thm.string_of_thm ctxt st :: "Premises:" :: map (Thm.string_of_thm ctxt) prems)) in case Seq.pull (ALLGOALS (Misc_Legacy.METAHYPS ctxt tacf) st) of SOME (th, _) => th | NONE => raise THM ("forward_res", 0, [st]) end; (*Are any of the logical connectives in "bs" present in the term?*) fun has_conns bs = let fun has (Const _) = false | has (Const(\<^const_name>\Trueprop\,_) $ p) = has p | has (Const(\<^const_name>\Not\,_) $ p) = has p | has (Const(\<^const_name>\HOL.disj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.disj\ orelse has p orelse has q | has (Const(\<^const_name>\HOL.conj\,_) $ p $ q) = member (op =) bs \<^const_name>\HOL.conj\ orelse has p orelse has q | has (Const(\<^const_name>\All\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\All\ orelse has p | has (Const(\<^const_name>\Ex\,_) $ Abs(_,_,p)) = member (op =) bs \<^const_name>\Ex\ orelse has p | has _ = false in has end; (**** Clause handling ****) fun literals (Const(\<^const_name>\Trueprop\,_) $ P) = literals P | literals (Const(\<^const_name>\HOL.disj\,_) $ P $ Q) = literals P @ literals Q | literals (Const(\<^const_name>\Not\,_) $ P) = [(false,P)] | literals P = [(true,P)]; (*number of literals in a term*) val nliterals = length o literals; (*** Tautology Checking ***) fun signed_lits_aux (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) (poslits, neglits) = signed_lits_aux Q (signed_lits_aux P (poslits, neglits)) | signed_lits_aux (Const(\<^const_name>\Not\,_) $ P) (poslits, neglits) = (poslits, P::neglits) | signed_lits_aux P (poslits, neglits) = (P::poslits, neglits); fun signed_lits th = signed_lits_aux (HOLogic.dest_Trueprop (Thm.concl_of th)) ([],[]); (*Literals like X=X are tautologous*) fun taut_poslit (Const(\<^const_name>\HOL.eq\,_) $ t $ u) = t aconv u | taut_poslit (Const(\<^const_name>\True\,_)) = true | taut_poslit _ = false; fun is_taut th = let val (poslits,neglits) = signed_lits th in exists taut_poslit poslits orelse exists (member (op aconv) neglits) (\<^term>\False\ :: poslits) end handle TERM _ => false; (*probably dest_Trueprop on a weird theorem*) (*** To remove trivial negated equality literals from clauses ***) (*They are typically functional reflexivity axioms and are the converses of injectivity equivalences*) val not_refl_disj_D = @{thm not_refl_disj_D}; (*Is either term a Var that does not properly occur in the other term?*) fun eliminable (t as Var _, u) = t aconv u orelse not (Logic.occs(t,u)) | eliminable (u, t as Var _) = t aconv u orelse not (Logic.occs(t,u)) | eliminable _ = false; fun refl_clause_aux 0 th = th | refl_clause_aux n th = case HOLogic.dest_Trueprop (Thm.concl_of th) of (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _) => refl_clause_aux n (th RS disj_assoc) (*isolate an atom as first disjunct*) | (Const (\<^const_name>\HOL.disj\, _) $ (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ t $ u)) $ _) => if eliminable(t,u) then refl_clause_aux (n-1) (th RS not_refl_disj_D) (*Var inequation: delete*) else refl_clause_aux (n-1) (th RS disj_comm) (*not between Vars: ignore*) | (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) => refl_clause_aux n (th RS disj_comm) | _ => (*not a disjunction*) th; fun notequal_lits_count (Const (\<^const_name>\HOL.disj\, _) $ P $ Q) = notequal_lits_count P + notequal_lits_count Q | notequal_lits_count (Const(\<^const_name>\Not\,_) $ (Const(\<^const_name>\HOL.eq\,_) $ _ $ _)) = 1 | notequal_lits_count _ = 0; (*Simplify a clause by applying reflexivity to its negated equality literals*) fun refl_clause th = let val neqs = notequal_lits_count (HOLogic.dest_Trueprop (Thm.concl_of th)) in zero_var_indexes (refl_clause_aux neqs th) end handle TERM _ => th; (*probably dest_Trueprop on a weird theorem*) (*** Removal of duplicate literals ***) (*Forward proof, passing extra assumptions as theorems to the tactic*) fun forward_res2 ctxt nf hyps st = case Seq.pull (REPEAT (Misc_Legacy.METAHYPS ctxt (fn major::minors => resolve_tac ctxt [nf (minors @ hyps) major] 1) 1) st) of SOME(th,_) => th | NONE => raise THM("forward_res2", 0, [st]); (*Remove duplicates in P|Q by assuming ~P in Q rls (initially []) accumulates assumptions of the form P==>False*) fun nodups_aux ctxt rls th = nodups_aux ctxt rls (th RS disj_assoc) handle THM _ => tryres(th,rls) handle THM _ => tryres(forward_res2 ctxt (nodups_aux ctxt) rls (th RS disj_forward2), [disj_FalseD1, disj_FalseD2, asm_rl]) handle THM _ => th; (*Remove duplicate literals, if there are any*) fun nodups ctxt th = if has_duplicates (op =) (literals (Thm.prop_of th)) then nodups_aux ctxt [] th else th; (*** The basic CNF transformation ***) fun estimated_num_clauses bound t = let fun sum x y = if x < bound andalso y < bound then x+y else bound fun prod x y = if x < bound andalso y < bound then x*y else bound (*Estimate the number of clauses in order to detect infeasible theorems*) fun signed_nclauses b (Const(\<^const_name>\Trueprop\,_) $ t) = signed_nclauses b t | signed_nclauses b (Const(\<^const_name>\Not\,_) $ t) = signed_nclauses (not b) t | signed_nclauses b (Const(\<^const_name>\HOL.conj\,_) $ t $ u) = if b then sum (signed_nclauses b t) (signed_nclauses b u) else prod (signed_nclauses b t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.disj\,_) $ t $ u) = if b then prod (signed_nclauses b t) (signed_nclauses b u) else sum (signed_nclauses b t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.implies\,_) $ t $ u) = if b then prod (signed_nclauses (not b) t) (signed_nclauses b u) else sum (signed_nclauses (not b) t) (signed_nclauses b u) | signed_nclauses b (Const(\<^const_name>\HOL.eq\, Type ("fun", [T, _])) $ t $ u) = if T = HOLogic.boolT then (*Boolean equality is if-and-only-if*) if b then sum (prod (signed_nclauses (not b) t) (signed_nclauses b u)) (prod (signed_nclauses (not b) u) (signed_nclauses b t)) else sum (prod (signed_nclauses b t) (signed_nclauses b u)) (prod (signed_nclauses (not b) t) (signed_nclauses (not b) u)) else 1 | signed_nclauses b (Const(\<^const_name>\Ex\, _) $ Abs (_,_,t)) = signed_nclauses b t | signed_nclauses b (Const(\<^const_name>\All\,_) $ Abs (_,_,t)) = signed_nclauses b t | signed_nclauses _ _ = 1; (* literal *) in signed_nclauses true t end fun has_too_many_clauses ctxt t = let val max_cl = Config.get ctxt max_clauses in estimated_num_clauses (max_cl + 1) t > max_cl end (*Replaces universally quantified variables by FREE variables -- because assumptions may not contain scheme variables. Later, generalize using Variable.export. *) local val spec_var = Thm.dest_arg (Thm.dest_arg (#2 (Thm.dest_implies (Thm.cprop_of spec)))) |> Thm.term_of |> dest_Var; fun name_of (Const (\<^const_name>\All\, _) $ Abs(x, _, _)) = x | name_of _ = Name.uu; in fun freeze_spec th ctxt = let val ([x], ctxt') = Variable.variant_fixes [name_of (HOLogic.dest_Trueprop (Thm.concl_of th))] ctxt; val spec' = spec - |> Thm.instantiate ([], [(spec_var, Thm.cterm_of ctxt' (Free (x, snd spec_var)))]); + |> Thm.instantiate + (TVars.empty, Vars.make [(spec_var, Thm.cterm_of ctxt' (Free (x, snd spec_var)))]); in (th RS spec', ctxt') end end; fun apply_skolem_theorem ctxt (th, rls) = let fun tryall [] = raise THM ("apply_skolem_theorem", 0, th::rls) | tryall (rl :: rls) = first_order_resolve ctxt th rl handle THM _ => tryall rls in tryall rls end (* Conjunctive normal form, adding clauses from th in front of ths (for foldr). Strips universal quantifiers and breaks up conjunctions. Eliminates existential quantifiers using Skolemization theorems. *) fun cnf old_skolem_ths ctxt (th, ths) = let val ctxt_ref = Unsynchronized.ref ctxt (* FIXME ??? *) fun cnf_aux (th,ths) = if not (can HOLogic.dest_Trueprop (Thm.prop_of th)) then ths (*meta-level: ignore*) else if not (has_conns [\<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\HOL.conj\] (Thm.prop_of th)) then nodups ctxt th :: ths (*no work to do, terminate*) else case head_of (HOLogic.dest_Trueprop (Thm.concl_of th)) of Const (\<^const_name>\HOL.conj\, _) => (*conjunction*) cnf_aux (th RS conjunct1, cnf_aux (th RS conjunct2, ths)) | Const (\<^const_name>\All\, _) => (*universal quantifier*) let val (th', ctxt') = freeze_spec th (! ctxt_ref) in ctxt_ref := ctxt'; cnf_aux (th', ths) end | Const (\<^const_name>\Ex\, _) => (*existential quantifier: Insert Skolem functions*) cnf_aux (apply_skolem_theorem (! ctxt_ref) (th, old_skolem_ths), ths) | Const (\<^const_name>\HOL.disj\, _) => (*Disjunction of P, Q: Create new goal of proving ?P | ?Q and solve it using all combinations of converting P, Q to CNF.*) (*There is one assumption, which gets bound to prem and then normalized via cnf_nil. The normal form is given to resolve_tac, instantiate a Boolean variable created by resolution with disj_forward. Since (cnf_nil prem) returns a LIST of theorems, we can backtrack to get all combinations.*) let val tac = Misc_Legacy.METAHYPS ctxt (fn [prem] => resolve_tac ctxt (cnf_nil prem) 1) 1 in Seq.list_of ((tac THEN tac) (th RS disj_forward)) @ ths end | _ => nodups ctxt th :: ths (*no work to do*) and cnf_nil th = cnf_aux (th, []) val cls = if has_too_many_clauses ctxt (Thm.concl_of th) then (trace_msg ctxt (fn () => "cnf is ignoring: " ^ Thm.string_of_thm ctxt th); ths) else cnf_aux (th, ths) in (cls, !ctxt_ref) end fun make_cnf old_skolem_ths th ctxt = cnf old_skolem_ths ctxt (th, []) (*Generalization, removal of redundant equalities, removal of tautologies.*) fun finish_cnf ths = filter (not o is_taut) (map refl_clause ths); (**** Generation of contrapositives ****) fun is_left (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ (Const (\<^const_name>\HOL.disj\, _) $ _ $ _) $ _)) = true | is_left _ = false; (*Associate disjuctions to right -- make leftmost disjunct a LITERAL*) fun assoc_right th = if is_left (Thm.prop_of th) then assoc_right (th RS disj_assoc) else th; (*Must check for negative literal first!*) val clause_rules = [disj_assoc, make_neg_rule, make_pos_rule]; (*For ordinary resolution. *) val resolution_clause_rules = [disj_assoc, make_neg_rule', make_pos_rule']; (*Create a goal or support clause, conclusing False*) fun make_goal th = (*Must check for negative literal first!*) make_goal (tryres(th, clause_rules)) handle THM _ => tryres(th, [make_neg_goal, make_pos_goal]); fun rigid t = not (is_Var (head_of t)); fun ok4horn (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\HOL.disj\, _) $ t $ _)) = rigid t | ok4horn (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4horn _ = false; (*Create a meta-level Horn clause*) fun make_horn crules th = if ok4horn (Thm.concl_of th) then make_horn crules (tryres(th,crules)) handle THM _ => th else th; (*Generate Horn clauses for all contrapositives of a clause. The input, th, is a HOL disjunction.*) fun add_contras crules th hcs = let fun rots (0,_) = hcs | rots (k,th) = zero_var_indexes (make_horn crules th) :: rots(k-1, assoc_right (th RS disj_comm)) in case nliterals(Thm.prop_of th) of 1 => th::hcs | n => rots(n, assoc_right th) end; (*Use "theorem naming" to label the clauses*) fun name_thms label = let fun name1 th (k, ths) = (k-1, Thm.put_name_hint (label ^ string_of_int k) th :: ths) in fn ths => #2 (fold_rev name1 ths (length ths, [])) end; (*Is the given disjunction an all-negative support clause?*) fun is_negative th = forall (not o #1) (literals (Thm.prop_of th)); val neg_clauses = filter is_negative; (***** MESON PROOF PROCEDURE *****) fun rhyps (Const(\<^const_name>\Pure.imp\,_) $ (Const(\<^const_name>\Trueprop\,_) $ A) $ phi, As) = rhyps(phi, A::As) | rhyps (_, As) = As; (** Detecting repeated assumptions in a subgoal **) (*The stringtree detects repeated assumptions.*) fun ins_term t net = Net.insert_term (op aconv) (t, t) net; (*detects repetitions in a list of terms*) fun has_reps [] = false | has_reps [_] = false | has_reps [t,u] = (t aconv u) | has_reps ts = (fold ins_term ts Net.empty; false) handle Net.INSERT => true; (*Like TRYALL eq_assume_tac, but avoids expensive THEN calls*) fun TRYING_eq_assume_tac 0 st = Seq.single st | TRYING_eq_assume_tac i st = TRYING_eq_assume_tac (i-1) (Thm.eq_assumption i st) handle THM _ => TRYING_eq_assume_tac (i-1) st; fun TRYALL_eq_assume_tac st = TRYING_eq_assume_tac (Thm.nprems_of st) st; (*Loop checking: FAIL if trying to prove the same thing twice -- if *ANY* subgoal has repeated literals*) fun check_tac st = if exists (fn prem => has_reps (rhyps(prem,[]))) (Thm.prems_of st) then Seq.empty else Seq.single st; (* resolve_from_net_tac actually made it slower... *) fun prolog_step_tac ctxt horns i = (assume_tac ctxt i APPEND resolve_tac ctxt horns i) THEN check_tac THEN TRYALL_eq_assume_tac; (*Sums the sizes of the subgoals, ignoring hypotheses (ancestors)*) fun addconcl prem sz = size_of_term (Logic.strip_assums_concl prem) + sz; fun size_of_subgoals st = fold_rev addconcl (Thm.prems_of st) 0; (*Negation Normal Form*) val nnf_rls = [imp_to_disjD, iff_to_disjD, not_conjD, not_disjD, not_impD, not_iffD, not_allD, not_exD, not_notD]; fun ok4nnf (Const (\<^const_name>\Trueprop\,_) $ (Const (\<^const_name>\Not\, _) $ t)) = rigid t | ok4nnf (Const (\<^const_name>\Trueprop\,_) $ t) = rigid t | ok4nnf _ = false; fun make_nnf1 ctxt th = if ok4nnf (Thm.concl_of th) then make_nnf1 ctxt (tryres(th, nnf_rls)) handle THM ("tryres", _, _) => forward_res ctxt (make_nnf1 ctxt) (tryres(th, [conj_forward,disj_forward,all_forward,ex_forward])) handle THM ("tryres", _, _) => th else th (*The simplification removes defined quantifiers and occurrences of True and False. nnf_ss also includes the one-point simprocs, which are needed to avoid the various one-point theorems from generating junk clauses.*) val nnf_simps = @{thms simp_implies_def Ex1_def Ball_def Bex_def if_True if_False if_cancel if_eq_cancel cases_simp} fun nnf_extra_simps ({if_simps, ...} : simp_options) = (if if_simps then @{thms split_ifs} else []) @ @{thms ex_simps all_simps simp_thms} (* FIXME: "let_simp" is probably redundant now that we also rewrite with "Let_def [abs_def]". *) fun nnf_ss simp_options = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps (nnf_extra_simps simp_options) addsimprocs [\<^simproc>\defined_All\, \<^simproc>\defined_Ex\, \<^simproc>\neq\, \<^simproc>\let_simp\]) val presimplified_consts = [\<^const_name>\simp_implies\, \<^const_name>\False\, \<^const_name>\True\, \<^const_name>\Ex1\, \<^const_name>\Ball\, \<^const_name>\Bex\, \<^const_name>\If\, \<^const_name>\Let\] fun presimplify (simp_options as {let_simps, ...} : simp_options) ctxt = rewrite_rule ctxt (map safe_mk_meta_eq nnf_simps) #> simplify (put_simpset (nnf_ss simp_options) ctxt) #> let_simps ? rewrite_rule ctxt @{thms Let_def [abs_def]} fun make_nnf simp_options ctxt th = (case Thm.prems_of th of [] => th |> presimplify simp_options ctxt |> make_nnf1 ctxt | _ => raise THM ("make_nnf: premises in argument", 0, [th])); fun choice_theorems thy = try (Global_Theory.get_thm thy) "Hilbert_Choice.choice" |> the_list (* Pull existential quantifiers to front. This accomplishes Skolemization for clauses that arise from a subgoal. *) fun skolemize_with_choice_theorems simp_options ctxt choice_ths = let fun aux th = if not (has_conns [\<^const_name>\Ex\] (Thm.prop_of th)) then th else tryres (th, choice_ths @ [conj_exD1, conj_exD2, disj_exD, disj_exD1, disj_exD2]) |> aux handle THM ("tryres", _, _) => tryres (th, [conj_forward, disj_forward, all_forward]) |> forward_res ctxt aux |> aux handle THM ("tryres", _, _) => rename_bound_vars_RS th ex_forward |> forward_res ctxt aux in aux o make_nnf simp_options ctxt end fun skolemize simp_options ctxt = let val thy = Proof_Context.theory_of ctxt in skolemize_with_choice_theorems simp_options ctxt (choice_theorems thy) end exception NO_F_PATTERN of unit fun get_F_pattern T t u = let fun pat t u = let val ((head1, args1), (head2, args2)) = (t, u) |> apply2 strip_comb in if head1 = head2 then let val pats = map2 pat args1 args2 in case filter (is_some o fst) pats of [(SOME T, _)] => (SOME T, list_comb (head1, map snd pats)) | [] => (NONE, t) | _ => raise NO_F_PATTERN () end else let val T = fastype_of t in if can dest_funT T then (SOME T, Bound 0) else raise NO_F_PATTERN () end end in if T = \<^typ>\bool\ then NONE else case pat t u of (SOME T, p as _ $ _) => SOME (Abs (Name.uu, T, p)) | _ => NONE end handle NO_F_PATTERN () => NONE val ext_cong_neq = @{thm ext_cong_neq} (* Strengthens "f g ~= f h" to "f g ~= f h & (EX x. g x ~= h x)". *) fun cong_extensionalize_thm ctxt th = (case Thm.concl_of th of \<^const>\Trueprop\ $ (\<^const>\Not\ $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ (t as _ $ _) $ (u as _ $ _))) => (case get_F_pattern T t u of SOME p => th RS infer_instantiate ctxt [(("F", 0), Thm.cterm_of ctxt p)] ext_cong_neq | NONE => th) | _ => th) (* Removes the lambdas from an equation of the form "t = (%x1 ... xn. u)". It would be desirable to do this symmetrically but there's at least one existing proof in "Tarski" that relies on the current behavior. *) fun abs_extensionalize_conv ctxt ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _ => ct |> (Conv.rewr_conv @{thm fun_eq_iff [THEN eq_reflection]} then_conv abs_extensionalize_conv ctxt) | _ $ _ => Conv.comb_conv (abs_extensionalize_conv ctxt) ct | Abs _ => Conv.abs_conv (abs_extensionalize_conv o snd) ctxt ct | _ => Conv.all_conv ct) val abs_extensionalize_thm = Conv.fconv_rule o abs_extensionalize_conv fun try_skolemize_etc simp_options ctxt th = let val th = th |> cong_extensionalize_thm ctxt in [th] (* Extensionalize lambdas in "th", because that makes sense and that's what Sledgehammer does, but also keep an unextensionalized version of "th" for backward compatibility. *) |> insert Thm.eq_thm_prop (abs_extensionalize_thm ctxt th) |> map_filter (fn th => th |> try (skolemize simp_options ctxt) |> tap (fn NONE => trace_msg ctxt (fn () => "Failed to skolemize " ^ Thm.string_of_thm ctxt th) | _ => ())) end fun add_clauses ctxt th cls = let val (cnfs, ctxt') = ctxt |> Variable.declare_thm th |> make_cnf [] th; in Variable.export ctxt' ctxt cnfs @ cls end; (*Sort clauses by number of literals*) fun fewerlits (th1, th2) = nliterals (Thm.prop_of th1) < nliterals (Thm.prop_of th2) (*Make clauses from a list of theorems, previously Skolemized and put into nnf. The resulting clauses are HOL disjunctions.*) fun make_clauses_unsorted ctxt ths = fold_rev (add_clauses ctxt) ths []; val make_clauses = sort (make_ord fewerlits) oo make_clauses_unsorted; (*Convert a list of clauses (disjunctions) to Horn clauses (contrapositives)*) fun make_horns ths = name_thms "Horn#" (distinct Thm.eq_thm_prop (fold_rev (add_contras clause_rules) ths [])); (*Could simply use nprems_of, which would count remaining subgoals -- no discrimination as to their size! With BEST_FIRST, fails for problem 41.*) fun best_prolog_tac ctxt sizef horns = BEST_FIRST (has_fewer_prems 1, sizef) (prolog_step_tac ctxt horns 1); fun depth_prolog_tac ctxt horns = DEPTH_FIRST (has_fewer_prems 1) (prolog_step_tac ctxt horns 1); (*Return all negative clauses, as possible goal clauses*) fun gocls cls = name_thms "Goal#" (map make_goal (neg_clauses cls)); fun skolemize_prems_tac simp_options ctxt prems = cut_facts_tac (maps (try_skolemize_etc simp_options ctxt) prems) THEN' REPEAT o eresolve_tac ctxt [exE] (*Basis of all meson-tactics. Supplies cltac with clauses: HOL disjunctions. Function mkcl converts theorems to clauses.*) fun MESON preskolem_tac mkcl cltac ctxt i st = SELECT_GOAL (EVERY [Object_Logic.atomize_prems_tac ctxt 1, resolve_tac ctxt @{thms ccontr} 1, preskolem_tac, Subgoal.FOCUS (fn {context = ctxt', prems = negs, ...} => EVERY1 [skolemize_prems_tac simp_options_all_true ctxt' negs, Subgoal.FOCUS (cltac o mkcl o #prems) ctxt']) ctxt 1]) i st handle THM _ => no_tac st; (*probably from make_meta_clause, not first-order*) (** Best-first search versions **) (*ths is a list of additional clauses (HOL disjunctions) to use.*) fun best_meson_tac sizef ctxt = MESON all_tac (make_clauses ctxt) (fn cls => THEN_BEST_FIRST (resolve_tac ctxt (gocls cls) 1) (has_fewer_prems 1, sizef) (prolog_step_tac ctxt (make_horns cls) 1)) ctxt (*First, breaks the goal into independent units*) fun safe_best_meson_tac ctxt = SELECT_GOAL (TRY (safe_tac ctxt) THEN TRYALL (best_meson_tac size_of_subgoals ctxt)); (** Depth-first search version **) fun depth_meson_tac ctxt = MESON all_tac (make_clauses ctxt) (fn cls => EVERY [resolve_tac ctxt (gocls cls) 1, depth_prolog_tac ctxt (make_horns cls)]) ctxt (** Iterative deepening version **) (*This version does only one inference per call; having only one eq_assume_tac speeds it up!*) fun prolog_step_tac' ctxt horns = let val horn0s = (*0 subgoals vs 1 or more*) take_prefix Thm.no_prems horns val nrtac = resolve_from_net_tac ctxt (Tactic.build_net horns) in fn i => eq_assume_tac i ORELSE match_tac ctxt horn0s i ORELSE (*no backtracking if unit MATCHES*) ((assume_tac ctxt i APPEND nrtac i) THEN check_tac) end; fun iter_deepen_prolog_tac ctxt horns = ITER_DEEPEN iter_deepen_limit (has_fewer_prems 1) (prolog_step_tac' ctxt horns); fun iter_deepen_meson_tac ctxt ths = ctxt |> MESON all_tac (make_clauses ctxt) (fn cls => (case (gocls (cls @ ths)) of [] => no_tac (*no goal clauses*) | goes => let val horns = make_horns (cls @ ths) val _ = trace_msg ctxt (fn () => cat_lines ("meson method called:" :: map (Thm.string_of_thm ctxt) (cls @ ths) @ ["clauses:"] @ map (Thm.string_of_thm ctxt) horns)) in THEN_ITER_DEEPEN iter_deepen_limit (resolve_tac ctxt goes 1) (has_fewer_prems 1) (prolog_step_tac' ctxt horns) end)); fun meson_tac ctxt ths = SELECT_GOAL (TRY (safe_tac ctxt) THEN TRYALL (iter_deepen_meson_tac ctxt ths)); (**** Code to support ordinary resolution, rather than Model Elimination ****) (*Convert a list of clauses (disjunctions) to meta-level clauses (==>), with no contrapositives, for ordinary resolution.*) (*Rules to convert the head literal into a negated assumption. If the head literal is already negated, then using notEfalse instead of notEfalse' prevents a double negation.*) val notEfalse = @{lemma "\ P \ P \ False" by (rule notE)}; val notEfalse' = @{lemma "P \ \ P \ False" by (rule notE)}; fun negated_asm_of_head th = th RS notEfalse handle THM _ => th RS notEfalse'; (*Converting one theorem from a disjunction to a meta-level clause*) fun make_meta_clause ctxt th = let val (fth, thaw) = Misc_Legacy.freeze_thaw_robust ctxt th in (zero_var_indexes o Thm.varifyT_global o thaw 0 o negated_asm_of_head o make_horn resolution_clause_rules) fth end; fun make_meta_clauses ctxt ths = name_thms "MClause#" (distinct Thm.eq_thm_prop (map (make_meta_clause ctxt) ths)); end; diff --git a/src/HOL/Tools/Meson/meson_clausify.ML b/src/HOL/Tools/Meson/meson_clausify.ML --- a/src/HOL/Tools/Meson/meson_clausify.ML +++ b/src/HOL/Tools/Meson/meson_clausify.ML @@ -1,388 +1,389 @@ (* Title: HOL/Tools/Meson/meson_clausify.ML Author: Jia Meng, Cambridge University Computer Laboratory and NICTA Author: Jasmin Blanchette, TU Muenchen Transformation of HOL theorems into CNF forms. *) signature MESON_CLAUSIFY = sig val new_skolem_var_prefix : string val new_nonskolem_var_prefix : string val is_zapped_var_name : string -> bool val is_quasi_lambda_free : term -> bool val introduce_combinators_in_cterm : Proof.context -> cterm -> thm val introduce_combinators_in_theorem : Proof.context -> thm -> thm val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool val ss_only : thm list -> Proof.context -> Proof.context val cnf_axiom : Meson.simp_options -> Proof.context -> bool -> bool -> int -> thm -> (thm * term) option * thm list end; structure Meson_Clausify : MESON_CLAUSIFY = struct (* the extra "Meson" helps prevent clashes (FIXME) *) val new_skolem_var_prefix = "MesonSK" val new_nonskolem_var_prefix = "MesonV" fun is_zapped_var_name s = exists (fn prefix => String.isPrefix prefix s) [new_skolem_var_prefix, new_nonskolem_var_prefix] (**** Transformation of Elimination Rules into First-Order Formulas****) val cfalse = Thm.cterm_of \<^theory_context>\HOL\ \<^term>\False\; val ctp_false = Thm.cterm_of \<^theory_context>\HOL\ (HOLogic.mk_Trueprop \<^term>\False\); (* Converts an elim-rule into an equivalent theorem that does not have the predicate variable. Leaves other theorems unchanged. We simply instantiate the conclusion variable to False. (Cf. "transform_elim_prop" in "Sledgehammer_Util".) *) fun transform_elim_theorem th = (case Thm.concl_of th of (*conclusion variable*) \<^const>\Trueprop\ $ (Var (v as (_, \<^typ>\bool\))) => - Thm.instantiate ([], [(v, cfalse)]) th + Thm.instantiate (TVars.empty, Vars.make [(v, cfalse)]) th | Var (v as (_, \<^typ>\prop\)) => - Thm.instantiate ([], [(v, ctp_false)]) th + Thm.instantiate (TVars.empty, Vars.make [(v, ctp_false)]) th | _ => th) (**** SKOLEMIZATION BY INFERENCE (lcp) ****) fun mk_old_skolem_term_wrapper t = let val T = fastype_of t in Const (\<^const_name>\Meson.skolem\, T --> T) $ t end fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') | beta_eta_in_abs_body t = Envir.beta_eta_contract t (*Traverse a theorem, accumulating Skolem function definitions.*) fun old_skolem_defs th = let fun dec_sko (Const (\<^const_name>\Ex\, _) $ (body as Abs (_, T, p))) rhss = (*Existential: declare a Skolem function, then insert into body and continue*) let val args = Misc_Legacy.term_frees body (* Forms a lambda-abstraction over the formal parameters *) val rhs = fold_rev (absfree o dest_Free) args (HOLogic.choice_const T $ beta_eta_in_abs_body body) |> mk_old_skolem_term_wrapper val comb = list_comb (rhs, args) in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end | dec_sko (Const (\<^const_name>\All\,_) $ Abs (a, T, p)) rhss = (*Universal quant: insert a free variable into body and continue*) let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a in dec_sko (subst_bound (Free(fname,T), p)) rhss end | dec_sko (\<^const>\conj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\disj\ $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q | dec_sko (\<^const>\Trueprop\ $ p) rhss = dec_sko p rhss | dec_sko _ rhss = rhss in dec_sko (Thm.prop_of th) [] end; (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) fun is_quasi_lambda_free (Const (\<^const_name>\Meson.skolem\, _) $ _) = true | is_quasi_lambda_free (t1 $ t2) = is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 | is_quasi_lambda_free (Abs _) = false | is_quasi_lambda_free _ = true fun abstract ctxt ct = let val Abs (_, _, body) = Thm.term_of ct val (x, cbody) = Thm.dest_abs NONE ct val (A, cbodyT) = Thm.dest_funT (Thm.ctyp_of_cterm ct) fun makeK () = Thm.instantiate' [SOME A, SOME cbodyT] [SOME cbody] @{thm abs_K} in case body of Const _ => makeK() | Free _ => makeK() | Var _ => makeK() (*though Var isn't expected*) | Bound 0 => Thm.instantiate' [SOME A] [] @{thm abs_I} (*identity: I*) | rator$rand => if Term.is_dependent rator then (*C or S*) if Term.is_dependent rand then (*S*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_S' = @{thm abs_S} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S') in Thm.transitive abs_S' (Conv.binop_conv (abstract ctxt) rhs) end else (*C*) let val crator = Thm.lambda x (Thm.dest_fun cbody) val crand = Thm.dest_arg cbody val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) val abs_C' = @{thm abs_C} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C') in Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv (abstract ctxt)) rhs) end else if Term.is_dependent rand then (*B or eta*) if rand = Bound 0 then Thm.eta_conversion ct else (*B*) let val crator = Thm.dest_fun cbody val crand = Thm.lambda x (Thm.dest_arg cbody) val (C, B) = Thm.dest_funT (Thm.ctyp_of_cterm crator) val abs_B' = @{thm abs_B} |> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') in Thm.transitive abs_B' (Conv.arg_conv (abstract ctxt) rhs) end else makeK () | _ => raise Fail "abstract: Bad term" end; (* Traverse a theorem, remplacing lambda-abstractions with combinators. *) fun introduce_combinators_in_cterm ctxt ct = if is_quasi_lambda_free (Thm.term_of ct) then Thm.reflexive ct else case Thm.term_of ct of Abs _ => let val (cv, cta) = Thm.dest_abs NONE ct val (v, _) = dest_Free (Thm.term_of cv) val u_th = introduce_combinators_in_cterm ctxt cta val cu = Thm.rhs_of u_th val comb_eq = abstract ctxt (Thm.lambda cv cu) in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end | _ $ _ => let val (ct1, ct2) = Thm.dest_comb ct in Thm.combination (introduce_combinators_in_cterm ctxt ct1) (introduce_combinators_in_cterm ctxt ct2) end fun introduce_combinators_in_theorem ctxt th = if is_quasi_lambda_free (Thm.prop_of th) then th else let val th = Drule.eta_contraction_rule th val eqth = introduce_combinators_in_cterm ctxt (Thm.cprop_of th) in Thm.equal_elim eqth th end handle THM (msg, _, _) => (warning ("Error in the combinator translation of " ^ Thm.string_of_thm ctxt th ^ "\nException message: " ^ msg); (* A type variable of sort "{}" will make "abstraction" fail. *) TrueI) (*cterms are used throughout for efficiency*) val cTrueprop = Thm.cterm_of \<^theory_context>\HOL\ HOLogic.Trueprop; (*Given an abstraction over n variables, replace the bound variables by free ones. Return the body, along with the list of free variables.*) fun c_variant_abs_multi (ct0, vars) = let val (cv,ct) = Thm.dest_abs NONE ct0 in c_variant_abs_multi (ct, cv::vars) end handle CTERM _ => (ct0, rev vars); (* Given the definition of a Skolem function, return a theorem to replace an existential formula by a use of that function. Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) fun old_skolem_theorem_of_def ctxt rhs0 = let val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of ctxt val rhs' = rhs |> Thm.dest_comb |> snd val (ch, frees) = c_variant_abs_multi (rhs', []) val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of val T = case hilbert of Const (_, Type (\<^type_name>\fun\, [_, T])) => T | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) val cex = Thm.cterm_of ctxt (HOLogic.exists_const T) val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) val conc = Drule.list_comb (rhs, frees) |> Drule.beta_conv cabs |> Thm.apply cTrueprop fun tacf [prem] = rewrite_goals_tac ctxt @{thms skolem_def [abs_def]} THEN resolve_tac ctxt [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]}) RS Global_Theory.get_thm (Proof_Context.theory_of ctxt) "Hilbert_Choice.someI_ex"] 1 in Goal.prove_internal ctxt [ex_tm] conc tacf |> forall_intr_list frees |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |> Thm.varifyT_global end fun to_definitional_cnf_with_quantifiers ctxt th = let val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (Thm.prop_of th)) val eqth = eqth RS @{thm eq_reflection} val eqth = eqth RS @{thm TruepropI} in Thm.equal_elim eqth th end fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ string_of_int index_no ^ "_" ^ Name.desymbolize (SOME false) s fun cluster_of_zapped_var_name s = let val get_int = the o Int.fromString o nth (space_explode "_" s) in ((get_int 1, (get_int 2, get_int 3)), String.isPrefix new_skolem_var_prefix s) end fun rename_bound_vars_to_be_zapped ax_no = let fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t = case t of (t1 as Const (s, _)) $ Abs (s', T, t') => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then let val skolem = (pos = (s = \<^const_name>\Ex\)) val (cluster, index_no) = if skolem = cluster_skolem then (cluster, index_no) else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) val s' = zapped_var_name cluster index_no s' in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end else t | (t1 as Const (s, _)) $ t2 $ t3 => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\HOL.implies\ then t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 else if s = \<^const_name>\HOL.conj\ orelse s = \<^const_name>\HOL.disj\ then t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3 else t | (t1 as Const (s, _)) $ t2 => if s = \<^const_name>\Trueprop\ then t1 $ aux cluster index_no pos t2 else if s = \<^const_name>\Not\ then t1 $ aux cluster index_no (not pos) t2 else t | _ => t in aux ((ax_no, 0), true) 0 true end fun zap pos ct = ct |> (case Thm.term_of ct of Const (s, _) $ Abs (s', _, _) => if s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\ then Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos else Conv.all_conv | Const (s, _) $ _ $ _ => if s = \<^const_name>\Pure.imp\ orelse s = \<^const_name>\implies\ then Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) else if s = \<^const_name>\conj\ orelse s = \<^const_name>\disj\ then Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) else Conv.all_conv | Const (s, _) $ _ => if s = \<^const_name>\Trueprop\ then Conv.arg_conv (zap pos) else if s = \<^const_name>\Not\ then Conv.arg_conv (zap (not pos)) else Conv.all_conv | _ => Conv.all_conv) fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths val cheat_choice = \<^prop>\\x. \y. Q x y \ \f. \x. Q x (f x)\ |> Logic.varify_global |> Skip_Proof.make_thm \<^theory> (* Converts an Isabelle theorem into NNF. *) fun nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt = let val thy = Proof_Context.theory_of ctxt val th = th |> transform_elim_theorem |> zero_var_indexes |> new_skolem ? Thm.forall_intr_vars val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single val th = th |> Conv.fconv_rule (Object_Logic.atomize ctxt) |> Meson.cong_extensionalize_thm ctxt |> Meson.abs_extensionalize_thm ctxt |> Meson.make_nnf simp_options ctxt in if new_skolem then let fun skolemize choice_ths = Meson.skolemize_with_choice_theorems simp_options ctxt choice_ths #> simplify (ss_only @{thms all_simps[symmetric]} ctxt) val no_choice = null choice_ths val pull_out = if no_choice then simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]} ctxt) else skolemize choice_ths val discharger_th = th |> pull_out val discharger_th = discharger_th |> Meson.has_too_many_clauses ctxt (Thm.concl_of discharger_th) ? (to_definitional_cnf_with_quantifiers ctxt #> pull_out) val zapped_th = discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no |> (if no_choice then Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of else Thm.cterm_of ctxt) |> zap true val fixes = [] |> Term.add_free_names (Thm.prop_of zapped_th) |> filter is_zapped_var_name val ctxt' = ctxt |> Variable.add_fixes_direct fixes val fully_skolemized_t = zapped_th |> singleton (Variable.export ctxt' ctxt) |> Thm.cprop_of |> Thm.dest_equals |> snd |> Thm.term_of in if exists_subterm (fn Var ((s, _), _) => String.isPrefix new_skolem_var_prefix s | _ => false) fully_skolemized_t then let val (fully_skolemized_ct, ctxt) = yield_singleton (Variable.import_terms true) fully_skolemized_t ctxt |>> Thm.cterm_of ctxt in (SOME (discharger_th, fully_skolemized_ct), (Thm.assume fully_skolemized_ct, ctxt)) end else (NONE, (th, ctxt)) end else (NONE, (th |> Meson.has_too_many_clauses ctxt (Thm.concl_of th) ? to_definitional_cnf_with_quantifiers ctxt, ctxt)) end (* Convert a theorem to CNF, with additional premises due to skolemization. *) fun cnf_axiom simp_options ctxt0 new_skolem combinators ax_no th = let val choice_ths = Meson.choice_theorems (Proof_Context.theory_of ctxt0) val (opt, (nnf_th, ctxt1)) = nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt0 fun clausify th = Meson.make_cnf (if new_skolem orelse null choice_ths then [] else map (old_skolem_theorem_of_def ctxt1) (old_skolem_defs th)) th ctxt1 val (cnf_ths, ctxt2) = clausify nnf_th fun intr_imp ct th = - Thm.instantiate ([], [((("i", 0), \<^typ>\nat\), Thm.cterm_of ctxt2 (HOLogic.mk_nat ax_no))]) + Thm.instantiate (TVars.empty, + Vars.make [((("i", 0), \<^typ>\nat\), Thm.cterm_of ctxt2 (HOLogic.mk_nat ax_no))]) (zero_var_indexes @{thm skolem_COMBK_D}) RS Thm.implies_intr ct th in (opt |> Option.map (I #>> singleton (Variable.export ctxt2 ctxt0) ##> (Thm.term_of #> HOLogic.dest_Trueprop #> singleton (Variable.export_terms ctxt2 ctxt0))), cnf_ths |> map (combinators ? introduce_combinators_in_theorem ctxt2 #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) |> Variable.export ctxt2 ctxt0 |> Meson.finish_cnf |> map (Thm.close_derivation \<^here>)) end handle THM _ => (NONE, []) end; diff --git a/src/HOL/Tools/Metis/metis_reconstruct.ML b/src/HOL/Tools/Metis/metis_reconstruct.ML --- a/src/HOL/Tools/Metis/metis_reconstruct.ML +++ b/src/HOL/Tools/Metis/metis_reconstruct.ML @@ -1,757 +1,758 @@ (* Title: HOL/Tools/Metis/metis_reconstruct.ML Author: Kong W. Susanto, Cambridge University Computer Laboratory Author: Lawrence C. Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Copyright Cambridge University 2007 Proof reconstruction for Metis. *) signature METIS_RECONSTRUCT = sig type type_enc = ATP_Problem_Generate.type_enc exception METIS_RECONSTRUCT of string * string val hol_clause_of_metis : Proof.context -> type_enc -> int Symtab.table -> (string * term) list * (string * term) list -> Metis_Thm.thm -> term val lookth : (Metis_Thm.thm * 'a) list -> Metis_Thm.thm -> 'a val replay_one_inference : Proof.context -> type_enc -> (string * term) list * (string * term) list -> int Symtab.table -> Metis_Thm.thm * Metis_Proof.inference -> (Metis_Thm.thm * thm) list -> (Metis_Thm.thm * thm) list val discharge_skolem_premises : Proof.context -> (thm * term) option list -> thm -> thm end; structure Metis_Reconstruct : METIS_RECONSTRUCT = struct open ATP_Problem open ATP_Problem_Generate open ATP_Proof_Reconstruct open Metis_Generate exception METIS_RECONSTRUCT of string * string val meta_not_not = @{thms not_not[THEN eq_reflection]} fun atp_name_of_metis type_enc s = (case find_first (fn (_, (f, _)) => f type_enc = s) metis_name_table of SOME ((s, _), (_, swap)) => (s, swap) | _ => (s, false)) fun atp_term_of_metis type_enc (Metis_Term.Fn (s, tms)) = let val (s, swap) = atp_name_of_metis type_enc (Metis_Name.toString s) in ATerm ((s, []), tms |> map (atp_term_of_metis type_enc) |> swap ? rev) end | atp_term_of_metis _ (Metis_Term.Var s) = ATerm ((Metis_Name.toString s, []), []) fun hol_term_of_metis ctxt type_enc sym_tab = atp_term_of_metis type_enc #> term_of_atp ctxt ATP_Problem.CNF type_enc false sym_tab NONE fun atp_literal_of_metis type_enc (pos, atom) = atom |> Metis_Term.Fn |> atp_term_of_metis type_enc |> AAtom |> not pos ? mk_anot fun atp_clause_of_metis _ [] = AAtom (ATerm ((tptp_false, []), [])) | atp_clause_of_metis type_enc lits = lits |> map (atp_literal_of_metis type_enc) |> mk_aconns AOr fun polish_hol_terms ctxt (lifted, old_skolems) = map (reveal_lam_lifted lifted #> reveal_old_skolem_terms old_skolems) #> Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_pattern ctxt) fun hol_clause_of_metis ctxt type_enc sym_tab concealed = Metis_Thm.clause #> Metis_LiteralSet.toList #> atp_clause_of_metis type_enc #> prop_of_atp ctxt ATP_Problem.CNF type_enc false sym_tab #> singleton (polish_hol_terms ctxt concealed) fun hol_terms_of_metis ctxt type_enc concealed sym_tab fol_tms = let val ts = map (hol_term_of_metis ctxt type_enc sym_tab) fol_tms val _ = trace_msg ctxt (fn () => " calling type inference:") val _ = List.app (fn t => trace_msg ctxt (fn () => Syntax.string_of_term ctxt t)) ts val ts' = ts |> polish_hol_terms ctxt concealed val _ = List.app (fn t => trace_msg ctxt (fn () => " final term: " ^ Syntax.string_of_term ctxt t ^ " of type " ^ Syntax.string_of_typ ctxt (type_of t))) ts' in ts' end (** FOL step Inference Rules **) fun lookth th_pairs fol_th = (case AList.lookup (uncurry Metis_Thm.equal) th_pairs fol_th of SOME th => th | NONE => raise Fail ("Failed to find Metis theorem " ^ Metis_Thm.toString fol_th)) fun cterm_incr_types ctxt idx = Thm.cterm_of ctxt o map_types (Logic.incr_tvar idx) (* INFERENCE RULE: AXIOM *) (*This causes variables to have an index of 1 by default. See also "term_of_atp" in "ATP_Proof_Reconstruct".*) val axiom_inference = Thm.incr_indexes 1 oo lookth (* INFERENCE RULE: ASSUME *) fun inst_excluded_middle i_atom = @{lemma "P \ \ P \ False" by (rule notE)} - |> instantiate_normalize ([], [((("P", 0), \<^typ>\bool\), i_atom)]) + |> instantiate_normalize (TVars.empty, Vars.make [((("P", 0), \<^typ>\bool\), i_atom)]) fun assume_inference ctxt type_enc concealed sym_tab atom = singleton (hol_terms_of_metis ctxt type_enc concealed sym_tab) (Metis_Term.Fn atom) |> Thm.cterm_of ctxt |> inst_excluded_middle (* INFERENCE RULE: INSTANTIATE (SUBST). *) (*Type instantiations are ignored. Trying to reconstruct them admits new possibilities of errors, e.g. concerning sorts. Instead we try to arrange hat new TVars are distinct and that types can be inferred from terms.*) fun inst_inference ctxt type_enc concealed sym_tab th_pairs fsubst th = let val i_th = lookth th_pairs th val i_th_vars = Term.add_vars (Thm.prop_of i_th) [] fun find_var x = the (List.find (fn ((a,_),_) => a=x) i_th_vars) fun subst_translation (x,y) = let val v = find_var x (*We call "polish_hol_terms" below.*) val t = hol_term_of_metis ctxt type_enc sym_tab y in SOME (Thm.cterm_of ctxt (Var v), t) end handle Option.Option => (trace_msg ctxt (fn () => "\"find_var\" failed for " ^ x ^ " in " ^ Thm.string_of_thm ctxt i_th); NONE) | TYPE _ => (trace_msg ctxt (fn () => "\"hol_term_of_metis\" failed for " ^ x ^ " in " ^ Thm.string_of_thm ctxt i_th); NONE) fun remove_typeinst (a, t) = let val a = Metis_Name.toString a in (case unprefix_and_unascii schematic_var_prefix a of SOME b => SOME (b, t) | NONE => (case unprefix_and_unascii tvar_prefix a of SOME _ => NONE (*type instantiations are forbidden*) | NONE => SOME (a, t) (*internal Metis var?*))) end val _ = trace_msg ctxt (fn () => " isa th: " ^ Thm.string_of_thm ctxt i_th) val substs = map_filter remove_typeinst (Metis_Subst.toList fsubst) val (vars, tms) = ListPair.unzip (map_filter subst_translation substs) ||> polish_hol_terms ctxt concealed val ctm_of = cterm_incr_types ctxt (Thm.maxidx_of i_th + 1) val substs' = ListPair.zip (vars, map ctm_of tms) val _ = trace_msg ctxt (fn () => cat_lines ("subst_translations:" :: (substs' |> map (fn (x, y) => Syntax.string_of_term ctxt (Thm.term_of x) ^ " |-> " ^ Syntax.string_of_term ctxt (Thm.term_of y))))) in infer_instantiate_types ctxt (map (apfst (dest_Var o Thm.term_of)) substs') i_th end handle THM (msg, _, _) => raise METIS_RECONSTRUCT ("inst_inference", msg) | ERROR msg => raise METIS_RECONSTRUCT ("inst_inference", msg) (* INFERENCE RULE: RESOLVE *) (*Increment the indexes of only the type variables*) fun incr_type_indexes ctxt inc th = let val tvs = Term.add_tvars (Thm.full_prop_of th) [] fun inc_tvar ((a, i), s) = (((a, i), s), Thm.ctyp_of ctxt (TVar ((a, i + inc), s))) in - Thm.instantiate (map inc_tvar tvs, []) th + Thm.instantiate (TVars.make (map inc_tvar tvs), Vars.empty) th end (*Like RSN, but we rename apart only the type variables. Vars here typically have an index of 1, and the use of RSN would increase this typically to 3. Instantiations of those Vars could then fail.*) fun resolve_inc_tyvars ctxt th1 i th2 = let val th1' = incr_type_indexes ctxt (Thm.maxidx_of th2 + 1) th1 fun res (tha, thb) = (case Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = true} (false, Thm.close_derivation \<^here> tha, Thm.nprems_of tha) i thb |> Seq.list_of |> distinct Thm.eq_thm of [th] => th | _ => let val thaa'bb' as [(tha', _), (thb', _)] = map (`(Local_Defs.unfold0 ctxt meta_not_not)) [tha, thb] in if forall Thm.eq_thm_prop thaa'bb' then raise THM ("resolve_inc_tyvars: unique result expected", i, [tha, thb]) else res (tha', thb') end) in res (th1', th2) handle TERM z => let val ps = [] |> fold (Term.add_vars o Thm.prop_of) [th1', th2] |> AList.group (op =) |> maps (fn ((s, _), T :: Ts) => map (fn T' => (Free (s, T), Free (s, T'))) Ts) |> rpair Envir.init |-> fold (Pattern.unify (Context.Proof ctxt)) |> Envir.type_env |> Vartab.dest |> map (fn (x, (S, T)) => ((x, S), Thm.ctyp_of ctxt T)) in (*The unifier, which is invoked from "Thm.bicompose", will sometimes refuse to unify "?a::?'a" with "?a::?'b" or "?a::nat" and throw a "TERM" exception (with "add_ffpair" as first argument). We then perform unification of the types of variables by hand and try again. We could do this the first time around but this error occurs seldom and we don't want to break existing proofs in subtle ways or slow them down.*) if null ps then raise TERM z - else res (apply2 (Drule.instantiate_normalize (ps, [])) (th1', th2)) + else res (apply2 (Drule.instantiate_normalize (TVars.make ps, Vars.empty)) (th1', th2)) end end fun s_not (\<^const>\Not\ $ t) = t | s_not t = HOLogic.mk_not t fun simp_not_not (\<^const>\Trueprop\ $ t) = \<^const>\Trueprop\ $ simp_not_not t | simp_not_not (\<^const>\Not\ $ t) = s_not (simp_not_not t) | simp_not_not t = t val normalize_literal = simp_not_not o Envir.eta_contract (*Find the relative location of an untyped term within a list of terms as a 1-based index. Returns 0 in case of failure.*) fun index_of_literal lit haystack = let fun match_lit normalize = HOLogic.dest_Trueprop #> normalize #> curry Term.aconv_untyped (lit |> normalize) in (case find_index (match_lit I) haystack of ~1 => find_index (match_lit (simp_not_not o Envir.eta_contract)) haystack | j => j) + 1 end (*Permute a rule's premises to move the i-th premise to the last position.*) fun make_last i th = let val n = Thm.nprems_of th in if i >= 1 andalso i <= n then Thm.permute_prems (i - 1) 1 th else raise THM ("select_literal", i, [th]) end (*Maps a rule that ends "... ==> P ==> False" to "... ==> ~ P" while avoiding to create double negations. The "select" wrapper is a trick to ensure that "P ==> ~ False ==> False" is rewritten to "P ==> False", not to "~ P". We don't use this trick in general because it makes the proof object uglier than necessary. FIXME.*) fun negate_head ctxt th = if exists (fn t => t aconv \<^prop>\\ False\) (Thm.prems_of th) then (th RS @{thm select_FalseI}) |> fold (rewrite_rule ctxt o single) @{thms not_atomize_select atomize_not_select} else th |> fold (rewrite_rule ctxt o single) @{thms not_atomize atomize_not} (* Maps the clause [P1,...Pn]==>False to [P1,...,P(i-1),P(i+1),...Pn] ==> ~P *) fun select_literal ctxt = negate_head ctxt oo make_last fun resolve_inference ctxt type_enc concealed sym_tab th_pairs atom th1 th2 = let val (i_th1, i_th2) = apply2 (lookth th_pairs) (th1, th2) val _ = trace_msg ctxt (fn () => " isa th1 (pos): " ^ Thm.string_of_thm ctxt i_th1 ^ "\n\ \ isa th2 (neg): " ^ Thm.string_of_thm ctxt i_th2) in (* Trivial cases where one operand is type info *) if Thm.eq_thm (TrueI, i_th1) then i_th2 else if Thm.eq_thm (TrueI, i_th2) then i_th1 else let val i_atom = singleton (hol_terms_of_metis ctxt type_enc concealed sym_tab) (Metis_Term.Fn atom) val _ = trace_msg ctxt (fn () => " atom: " ^ Syntax.string_of_term ctxt i_atom) in (case index_of_literal (s_not i_atom) (Thm.prems_of i_th1) of 0 => (trace_msg ctxt (fn () => "Failed to find literal in \"th1\""); i_th1) | j1 => (trace_msg ctxt (fn () => " index th1: " ^ string_of_int j1); (case index_of_literal i_atom (Thm.prems_of i_th2) of 0 => (trace_msg ctxt (fn () => "Failed to find literal in \"th2\""); i_th2) | j2 => (trace_msg ctxt (fn () => " index th2: " ^ string_of_int j2); resolve_inc_tyvars ctxt (select_literal ctxt j1 i_th1) j2 i_th2 handle TERM (s, _) => raise METIS_RECONSTRUCT ("resolve_inference", s))))) end end (* INFERENCE RULE: REFL *) val REFL_THM = Thm.incr_indexes 2 @{lemma "x \ x \ False" by (drule notE) (rule refl)} val [refl_x] = Term.add_vars (Thm.prop_of REFL_THM) []; fun refl_inference ctxt type_enc concealed sym_tab t = let val i_t = singleton (hol_terms_of_metis ctxt type_enc concealed sym_tab) t val _ = trace_msg ctxt (fn () => " term: " ^ Syntax.string_of_term ctxt i_t) val c_t = cterm_incr_types ctxt (Thm.maxidx_of REFL_THM + 1) i_t in infer_instantiate_types ctxt [(refl_x, c_t)] REFL_THM end (* INFERENCE RULE: EQUALITY *) val subst_em = @{lemma "s = t \ P s \ \ P t \ False" by (erule notE) (erule subst)} val ssubst_em = @{lemma "s = t \ P t \ \ P s \ False" by (erule notE) (erule ssubst)} fun equality_inference ctxt type_enc concealed sym_tab (pos, atom) fp fr = let val m_tm = Metis_Term.Fn atom val [i_atom, i_tm] = hol_terms_of_metis ctxt type_enc concealed sym_tab [m_tm, fr] val _ = trace_msg ctxt (fn () => "sign of the literal: " ^ Bool.toString pos) fun replace_item_list lx 0 (_::ls) = lx::ls | replace_item_list lx i (l::ls) = l :: replace_item_list lx (i-1) ls fun path_finder_fail tm ps t = raise METIS_RECONSTRUCT ("equality_inference (path_finder)", "path = " ^ space_implode " " (map string_of_int ps) ^ " isa-term: " ^ Syntax.string_of_term ctxt tm ^ (case t of SOME t => " fol-term: " ^ Metis_Term.toString t | NONE => "")) fun path_finder tm [] _ = (tm, Bound 0) | path_finder tm (p :: ps) (t as Metis_Term.Fn (s, ts)) = let val s = s |> Metis_Name.toString |> perhaps (try (unprefix_and_unascii const_prefix #> the #> unmangled_const_name #> hd)) in if s = metis_predicator orelse s = predicator_name orelse s = metis_systematic_type_tag orelse s = metis_ad_hoc_type_tag orelse s = type_tag_name then path_finder tm ps (nth ts p) else if s = metis_app_op orelse s = app_op_name then let val (tm1, tm2) = dest_comb tm val p' = p - (length ts - 2) in if p' = 0 then path_finder tm1 ps (nth ts p) ||> (fn y => y $ tm2) else path_finder tm2 ps (nth ts p) ||> (fn y => tm1 $ y) end else let val (tm1, args) = strip_comb tm val adjustment = length ts - length args val p' = if adjustment > p then p else p - adjustment val tm_p = nth args p' handle General.Subscript => path_finder_fail tm (p :: ps) (SOME t) val _ = trace_msg ctxt (fn () => "path_finder: " ^ string_of_int p ^ " " ^ Syntax.string_of_term ctxt tm_p) val (r, t) = path_finder tm_p ps (nth ts p) in (r, list_comb (tm1, replace_item_list t p' args)) end end | path_finder tm ps t = path_finder_fail tm ps (SOME t) val (tm_subst, body) = path_finder i_atom fp m_tm val tm_abs = Abs ("x", type_of tm_subst, body) val _ = trace_msg ctxt (fn () => "abstraction: " ^ Syntax.string_of_term ctxt tm_abs) val _ = trace_msg ctxt (fn () => "i_tm: " ^ Syntax.string_of_term ctxt i_tm) val _ = trace_msg ctxt (fn () => "located term: " ^ Syntax.string_of_term ctxt tm_subst) val maxidx = fold Term.maxidx_term [i_tm, tm_abs, tm_subst] ~1 val subst' = Thm.incr_indexes (maxidx + 1) (if pos then subst_em else ssubst_em) val _ = trace_msg ctxt (fn () => "subst' " ^ Thm.string_of_thm ctxt subst') val eq_terms = map (apply2 (Thm.cterm_of ctxt)) (ListPair.zip (Misc_Legacy.term_vars (Thm.prop_of subst'), [tm_abs, tm_subst, i_tm])) in infer_instantiate_types ctxt (map (apfst (dest_Var o Thm.term_of)) eq_terms) subst' end val factor = Seq.hd o distinct_subgoals_tac fun one_step ctxt type_enc concealed sym_tab th_pairs (fol_th, inference) = (case inference of Metis_Proof.Axiom _ => axiom_inference th_pairs fol_th |> factor | Metis_Proof.Assume atom => assume_inference ctxt type_enc concealed sym_tab atom | Metis_Proof.Metis_Subst (subst, th1) => inst_inference ctxt type_enc concealed sym_tab th_pairs subst th1 |> factor | Metis_Proof.Resolve (atom, th1, th2) => resolve_inference ctxt type_enc concealed sym_tab th_pairs atom th1 th2 |> factor | Metis_Proof.Refl tm => refl_inference ctxt type_enc concealed sym_tab tm | Metis_Proof.Equality (lit, p, r) => equality_inference ctxt type_enc concealed sym_tab lit p r) fun flexflex_first_order ctxt th = (case Thm.tpairs_of th of [] => th | pairs => let val thy = Proof_Context.theory_of ctxt val (tyenv, tenv) = fold (Pattern.first_order_match thy) pairs (Vartab.empty, Vartab.empty) fun mkT (v, (S, T)) = ((v, S), Thm.ctyp_of ctxt T) fun mk (v, (T, t)) = ((v, Envir.subst_type tyenv T), Thm.cterm_of ctxt t) val instsT = Vartab.fold (cons o mkT) tyenv [] val insts = Vartab.fold (cons o mk) tenv [] in - Thm.instantiate (instsT, insts) th + Thm.instantiate (TVars.make instsT, Vars.make insts) th end handle THM _ => th) fun is_metis_literal_genuine (_, (s, _)) = not (String.isPrefix class_prefix (Metis_Name.toString s)) fun is_isabelle_literal_genuine t = (case t of _ $ (Const (\<^const_name>\Meson.skolem\, _) $ _) => false | _ => true) fun count p xs = fold (fn x => if p x then Integer.add 1 else I) xs 0 (*Seldomly needed hack. A Metis clause is represented as a set, so duplicate disjuncts are impossible. In the Isabelle proof, in spite of efforts to eliminate them, duplicates sometimes appear with slightly different (but unifiable) types.*) fun resynchronize ctxt fol_th th = let val num_metis_lits = count is_metis_literal_genuine (Metis_LiteralSet.toList (Metis_Thm.clause fol_th)) val num_isabelle_lits = count is_isabelle_literal_genuine (Thm.prems_of th) in if num_metis_lits >= num_isabelle_lits then th else let val (prems0, concl) = th |> Thm.prop_of |> Logic.strip_horn val prems = prems0 |> map normalize_literal |> distinct Term.aconv_untyped val goal = Logic.list_implies (prems, concl) val ctxt' = fold Thm.declare_hyps (Thm.chyps_of th) ctxt val tac = cut_tac th 1 THEN rewrite_goals_tac ctxt' meta_not_not THEN ALLGOALS (assume_tac ctxt') in if length prems = length prems0 then raise METIS_RECONSTRUCT ("resynchronize", "Out of sync") else Goal.prove ctxt' [] [] goal (K tac) |> resynchronize ctxt' fol_th end end fun replay_one_inference ctxt type_enc concealed sym_tab (fol_th, inf) th_pairs = if not (null th_pairs) andalso Thm.prop_of (snd (hd th_pairs)) aconv \<^prop>\False\ then (*Isabelle sometimes identifies literals (premises) that are distinct in Metis (e.g., because of type variables). We give the Isabelle proof the benefice of the doubt.*) th_pairs else let val _ = trace_msg ctxt (fn () => "=============================================") val _ = trace_msg ctxt (fn () => "METIS THM: " ^ Metis_Thm.toString fol_th) val _ = trace_msg ctxt (fn () => "INFERENCE: " ^ Metis_Proof.inferenceToString inf) val th = one_step ctxt type_enc concealed sym_tab th_pairs (fol_th, inf) |> flexflex_first_order ctxt |> resynchronize ctxt fol_th val _ = trace_msg ctxt (fn () => "ISABELLE THM: " ^ Thm.string_of_thm ctxt th) val _ = trace_msg ctxt (fn () => "=============================================") in (fol_th, th) :: th_pairs end (*It is normally sufficient to apply "assume_tac" to unify the conclusion with one of the premises. Unfortunately, this sometimes yields "Variable has two distinct types" errors. To avoid this, we instantiate the variables before applying "assume_tac". Typical constraints are of the form ?SK_a_b_c_x SK_d_e_f_y ... SK_a_b_c_x ... SK_g_h_i_z \\<^sup>? SK_a_b_c_x, where the nonvariables are goal parameters.*) fun unify_first_prem_with_concl ctxt i th = let val goal = Logic.get_goal (Thm.prop_of th) i |> Envir.beta_eta_contract val prem = goal |> Logic.strip_assums_hyp |> hd val concl = goal |> Logic.strip_assums_concl fun pair_untyped_aconv (t1, t2) (u1, u2) = Term.aconv_untyped (t1, u1) andalso Term.aconv_untyped (t2, u2) fun add_terms tp inst = if exists (pair_untyped_aconv tp) inst then inst else tp :: map (apsnd (subst_atomic [tp])) inst fun is_flex t = (case strip_comb t of (Var _, args) => forall is_Bound args | _ => false) fun unify_flex flex rigid = (case strip_comb flex of (Var (z as (_, T)), args) => add_terms (Var z, fold_rev absdummy (take (length args) (binder_types T)) rigid) | _ => I) fun unify_potential_flex comb atom = if is_flex comb then unify_flex comb atom else if is_Var atom then add_terms (atom, comb) else I fun unify_terms (t, u) = (case (t, u) of (t1 $ t2, u1 $ u2) => if is_flex t then unify_flex t u else if is_flex u then unify_flex u t else fold unify_terms [(t1, u1), (t2, u2)] | (_ $ _, _) => unify_potential_flex t u | (_, _ $ _) => unify_potential_flex u t | (Var _, _) => add_terms (t, u) | (_, Var _) => add_terms (u, t) | _ => I) val t_inst = [] |> try (unify_terms (prem, concl) #> map (apply2 (Thm.cterm_of ctxt))) |> the_default [] (* FIXME *) in infer_instantiate_types ctxt (map (apfst (dest_Var o Thm.term_of)) t_inst) th end val copy_prem = @{lemma "P \ (P \ P \ Q) \ Q" by assumption} fun copy_prems_tac ctxt [] ns i = if forall (curry (op =) 1) ns then all_tac else copy_prems_tac ctxt (rev ns) [] i | copy_prems_tac ctxt (1 :: ms) ns i = rotate_tac 1 i THEN copy_prems_tac ctxt ms (1 :: ns) i | copy_prems_tac ctxt (m :: ms) ns i = eresolve_tac ctxt [copy_prem] i THEN copy_prems_tac ctxt ms (m div 2 :: (m + 1) div 2 :: ns) i (*Metis generates variables of the form _nnn.*) val is_metis_fresh_variable = String.isPrefix "_" fun instantiate_forall_tac ctxt t i st = let val params = Logic.strip_params (Logic.get_goal (Thm.prop_of st) i) |> rev fun repair (t as (Var ((s, _), _))) = (case find_index (fn (s', _) => s' = s) params of ~1 => t | j => Bound j) | repair (t $ u) = (case (repair t, repair u) of (t as Bound j, u as Bound k) => (*This is a trick to repair the discrepancy between the fully skolemized term that MESON gives us (where existentials were pulled out) and the reality.*) if k > j then t else t $ u | (t, u) => t $ u) | repair t = t val t' = t |> repair |> fold (absdummy o snd) params fun do_instantiate th = (case Term.add_vars (Thm.prop_of th) [] |> filter_out ((Meson_Clausify.is_zapped_var_name orf is_metis_fresh_variable) o fst o fst) of [] => th | [var as (_, T)] => let val var_binder_Ts = T |> binder_types |> take (length params) |> rev val var_body_T = T |> funpow (length params) range_type val tyenv = Vartab.empty |> Type.raw_unifys (fastype_of t :: map snd params, var_body_T :: var_binder_Ts) val env = Envir.Envir {maxidx = Vartab.fold (Integer.max o snd o fst) tyenv 0, tenv = Vartab.empty, tyenv = tyenv} val ty_inst = Vartab.fold (fn (x, (S, T)) => cons (((x, S), Thm.ctyp_of ctxt T))) tyenv [] val t_inst = [apply2 (Thm.cterm_of ctxt o Envir.norm_term env) (Var var, t')] in - Drule.instantiate_normalize (ty_inst, map (apfst (dest_Var o Thm.term_of)) t_inst) th + Drule.instantiate_normalize + (TVars.make ty_inst, Vars.make (map (apfst (dest_Var o Thm.term_of)) t_inst)) th end | _ => raise Fail "expected a single non-zapped, non-Metis Var") in (DETERM (eresolve_tac ctxt @{thms allE} i THEN rotate_tac ~1 i) THEN PRIMITIVE do_instantiate) st end fun fix_exists_tac ctxt t = eresolve_tac ctxt [exE] THEN' rename_tac [t |> dest_Var |> fst |> fst] fun release_quantifier_tac ctxt (skolem, t) = (if skolem then fix_exists_tac ctxt else instantiate_forall_tac ctxt) t fun release_clusters_tac _ _ _ [] = K all_tac | release_clusters_tac ctxt ax_counts substs ((ax_no, cluster_no) :: clusters) = let val cluster_of_var = Meson_Clausify.cluster_of_zapped_var_name o fst o fst o dest_Var fun in_right_cluster ((_, (cluster_no', _)), _) = cluster_no' = cluster_no val cluster_substs = substs |> map_filter (fn (ax_no', (_, (_, tsubst))) => if ax_no' = ax_no then tsubst |> map (apfst cluster_of_var) |> filter (in_right_cluster o fst) |> map (apfst snd) |> SOME else NONE) fun do_cluster_subst cluster_subst = map (release_quantifier_tac ctxt) cluster_subst @ [rotate_tac 1] val first_prem = find_index (fn (ax_no', _) => ax_no' = ax_no) substs in rotate_tac first_prem THEN' (EVERY' (maps do_cluster_subst cluster_substs)) THEN' rotate_tac (~ first_prem - length cluster_substs) THEN' release_clusters_tac ctxt ax_counts substs clusters end fun cluster_key ((ax_no, (cluster_no, index_no)), skolem) = (ax_no, (cluster_no, (skolem, index_no))) fun cluster_ord p = prod_ord int_ord (prod_ord int_ord (prod_ord bool_ord int_ord)) (apply2 cluster_key p) val tysubst_ord = list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord)) structure Int_Tysubst_Table = Table ( type key = int * (indexname * (sort * typ)) list val ord = prod_ord int_ord tysubst_ord ) structure Int_Pair_Graph = Graph( type key = int * int val ord = prod_ord int_ord int_ord ) fun shuffle_key (((axiom_no, (_, index_no)), _), _) = (axiom_no, index_no) fun shuffle_ord p = prod_ord int_ord int_ord (apply2 shuffle_key p) (*Attempts to derive the theorem "False" from a theorem of the form "P1 ==> ... ==> Pn ==> False", where the "Pi"s are to be discharged using the specified axioms. The axioms have leading "All" and "Ex" quantifiers, which must be eliminated first.*) fun discharge_skolem_premises ctxt axioms prems_imp_false = if Thm.prop_of prems_imp_false aconv \<^prop>\False\ then prems_imp_false else let val thy = Proof_Context.theory_of ctxt fun match_term p = let val (tyenv, tenv) = Pattern.first_order_match thy p (Vartab.empty, Vartab.empty) val tsubst = tenv |> Vartab.dest |> filter (Meson_Clausify.is_zapped_var_name o fst o fst) |> sort (cluster_ord o apply2 (Meson_Clausify.cluster_of_zapped_var_name o fst o fst)) |> map (fn (xi, (T, t)) => apply2 (Envir.subst_term_types tyenv) (Var (xi, T), t)) val tysubst = tyenv |> Vartab.dest in (tysubst, tsubst) end fun subst_info_of_prem subgoal_no prem = (case prem of _ $ (Const (\<^const_name>\Meson.skolem\, _) $ (_ $ t $ num)) => let val ax_no = HOLogic.dest_nat num in (ax_no, (subgoal_no, match_term (nth axioms ax_no |> the |> snd, t))) end | _ => raise TERM ("discharge_skolem_premises: Malformed premise", [prem])) fun cluster_of_var_name skolem s = (case try Meson_Clausify.cluster_of_zapped_var_name s of NONE => NONE | SOME ((ax_no, (cluster_no, _)), skolem') => if skolem' = skolem andalso cluster_no > 0 then SOME (ax_no, cluster_no) else NONE) fun clusters_in_term skolem t = Term.add_var_names t [] |> map_filter (cluster_of_var_name skolem o fst) fun deps_of_term_subst (var, t) = (case clusters_in_term false var of [] => NONE | [(ax_no, cluster_no)] => SOME ((ax_no, cluster_no), clusters_in_term true t |> cluster_no > 1 ? cons (ax_no, cluster_no - 1)) | _ => raise TERM ("discharge_skolem_premises: Expected Var", [var])) val prems = Logic.strip_imp_prems (Thm.prop_of prems_imp_false) val substs = prems |> map2 subst_info_of_prem (1 upto length prems) |> sort (int_ord o apply2 fst) val depss = maps (map_filter deps_of_term_subst o snd o snd o snd) substs val clusters = maps (op ::) depss val ordered_clusters = Int_Pair_Graph.empty |> fold Int_Pair_Graph.default_node (map (rpair ()) clusters) |> fold Int_Pair_Graph.add_deps_acyclic depss |> Int_Pair_Graph.topological_order handle Int_Pair_Graph.CYCLES _ => error "Cannot replay Metis proof in Isabelle without \"Hilbert_Choice\"" val ax_counts = Int_Tysubst_Table.empty |> fold (fn (ax_no, (_, (tysubst, _))) => Int_Tysubst_Table.map_default ((ax_no, tysubst), 0) (Integer.add 1)) substs |> Int_Tysubst_Table.dest val needed_axiom_props = 0 upto length axioms - 1 ~~ axioms |> map_filter (fn (_, NONE) => NONE | (ax_no, SOME (_, t)) => if exists (fn ((ax_no', _), n) => ax_no' = ax_no andalso n > 0) ax_counts then SOME t else NONE) val outer_param_names = [] |> fold Term.add_var_names needed_axiom_props |> filter (Meson_Clausify.is_zapped_var_name o fst) |> map (`(Meson_Clausify.cluster_of_zapped_var_name o fst)) |> filter (fn (((_, (cluster_no, _)), skolem), _) => cluster_no = 0 andalso skolem) |> sort shuffle_ord |> map (fst o snd) (* for debugging only: fun string_of_subst_info (ax_no, (subgoal_no, (tysubst, tsubst))) = "ax: " ^ string_of_int ax_no ^ "; asm: " ^ string_of_int subgoal_no ^ "; tysubst: " ^ @{make_string} tysubst ^ "; tsubst: {" ^ commas (map ((fn (s, t) => s ^ " |-> " ^ t) o apply2 (Syntax.string_of_term ctxt)) tsubst) ^ "}" val _ = tracing ("ORDERED CLUSTERS: " ^ @{make_string} ordered_clusters) val _ = tracing ("AXIOM COUNTS: " ^ @{make_string} ax_counts) val _ = tracing ("OUTER PARAMS: " ^ @{make_string} outer_param_names) val _ = tracing ("SUBSTS (" ^ string_of_int (length substs) ^ "):\n" ^ cat_lines (map string_of_subst_info substs)) *) val ctxt' = fold Thm.declare_hyps (Thm.chyps_of prems_imp_false) ctxt fun cut_and_ex_tac axiom = cut_tac axiom 1 THEN TRY (REPEAT_ALL_NEW (eresolve_tac ctxt' @{thms exE}) 1) fun rotation_of_subgoal i = find_index (fn (_, (subgoal_no, _)) => subgoal_no = i) substs in Goal.prove ctxt' [] [] \<^prop>\False\ (K (DETERM (EVERY (map (cut_and_ex_tac o fst o the o nth axioms o fst o fst) ax_counts) THEN rename_tac outer_param_names 1 THEN copy_prems_tac ctxt' (map snd ax_counts) [] 1) THEN release_clusters_tac ctxt' ax_counts substs ordered_clusters 1 THEN match_tac ctxt' [prems_imp_false] 1 THEN ALLGOALS (fn i => resolve_tac ctxt' @{thms Meson.skolem_COMBK_I} i THEN rotate_tac (rotation_of_subgoal i) i THEN PRIMITIVE (unify_first_prem_with_concl ctxt' i) THEN assume_tac ctxt' i THEN flexflex_tac ctxt'))) handle ERROR msg => cat_error msg "Cannot replay Metis proof in Isabelle: error when discharging Skolem assumptions" end end; diff --git a/src/HOL/Tools/Predicate_Compile/core_data.ML b/src/HOL/Tools/Predicate_Compile/core_data.ML --- a/src/HOL/Tools/Predicate_Compile/core_data.ML +++ b/src/HOL/Tools/Predicate_Compile/core_data.ML @@ -1,480 +1,481 @@ (* Title: HOL/Tools/Predicate_Compile/core_data.ML Author: Lukas Bulwahn, TU Muenchen Data of the predicate compiler core. *) signature CORE_DATA = sig type mode = Predicate_Compile_Aux.mode type compilation = Predicate_Compile_Aux.compilation type compilation_funs = Predicate_Compile_Aux.compilation_funs datatype predfun_data = PredfunData of { definition : thm, intro : thm, elim : thm, neg_intro : thm option }; datatype pred_data = PredData of { pos : Position.T, intros : (string option * thm) list, elim : thm option, preprocessed : bool, function_names : (compilation * (mode * string) list) list, predfun_data : (mode * predfun_data) list, needs_random : mode list }; structure PredData : THEORY_DATA (* FIXME keep data private *) (* queries *) val defined_functions : compilation -> Proof.context -> string -> bool val is_registered : Proof.context -> string -> bool val function_name_of : compilation -> Proof.context -> string -> mode -> string val the_elim_of : Proof.context -> string -> thm val has_elim : Proof.context -> string -> bool val needs_random : Proof.context -> string -> mode -> bool val predfun_intro_of : Proof.context -> string -> mode -> thm val predfun_neg_intro_of : Proof.context -> string -> mode -> thm option val predfun_elim_of : Proof.context -> string -> mode -> thm val predfun_definition_of : Proof.context -> string -> mode -> thm val all_preds_of : Proof.context -> string list val modes_of: compilation -> Proof.context -> string -> mode list val all_modes_of : compilation -> Proof.context -> (string * mode list) list val all_random_modes_of : Proof.context -> (string * mode list) list val intros_of : Proof.context -> string -> thm list val names_of : Proof.context -> string -> string option list val intros_graph_of : Proof.context -> thm list Graph.T (* updaters *) val register_predicate : (string * thm list * thm) -> theory -> theory val register_intros : string * thm list -> theory -> theory (* FIXME: naming of function is strange *) val defined_function_of : compilation -> string -> theory -> theory val add_intro : string option * thm -> theory -> theory val set_elim : thm -> theory -> theory val set_function_name : compilation -> string -> mode -> string -> theory -> theory val add_predfun_data : string -> mode -> thm * ((thm * thm) * thm option) -> theory -> theory val set_needs_random : string -> mode list -> theory -> theory (* sophisticated updaters *) val extend_intro_graph : string list -> theory -> theory val preprocess_intros : string -> theory -> theory (* alternative function definitions *) val register_alternative_function : string -> mode -> string -> theory -> theory val alternative_compilation_of_global : theory -> string -> mode -> (compilation_funs -> typ -> term) option val alternative_compilation_of : Proof.context -> string -> mode -> (compilation_funs -> typ -> term) option val functional_compilation : string -> mode -> compilation_funs -> typ -> term val force_modes_and_functions : string -> (mode * (string * bool)) list -> theory -> theory val force_modes_and_compilations : string -> (mode * ((compilation_funs -> typ -> term) * bool)) list -> theory -> theory end; structure Core_Data : CORE_DATA = struct open Predicate_Compile_Aux; (* book-keeping *) datatype predfun_data = PredfunData of { definition : thm, intro : thm, elim : thm, neg_intro : thm option }; fun rep_predfun_data (PredfunData data) = data; fun mk_predfun_data (definition, ((intro, elim), neg_intro)) = PredfunData {definition = definition, intro = intro, elim = elim, neg_intro = neg_intro} datatype pred_data = PredData of { pos: Position.T, intros : (string option * thm) list, elim : thm option, preprocessed : bool, function_names : (compilation * (mode * string) list) list, predfun_data : (mode * predfun_data) list, needs_random : mode list }; fun rep_pred_data (PredData data) = data; val pos_of = #pos o rep_pred_data; fun mk_pred_data (pos, (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random)))) = PredData {pos = pos, intros = intros, elim = elim, preprocessed = preprocessed, function_names = function_names, predfun_data = predfun_data, needs_random = needs_random} fun map_pred_data f (PredData {pos, intros, elim, preprocessed, function_names, predfun_data, needs_random}) = mk_pred_data (f (pos, (((intros, elim), preprocessed), (function_names, (predfun_data, needs_random))))) fun eq_pred_data (PredData d1, PredData d2) = eq_list (eq_pair (op =) Thm.eq_thm) (#intros d1, #intros d2) andalso eq_option Thm.eq_thm (#elim d1, #elim d2) structure PredData = Theory_Data ( type T = pred_data Graph.T; val empty = Graph.empty; val extend = I; val merge = Graph.join (fn key => fn (x, y) => if eq_pred_data (x, y) then raise Graph.SAME else error ("Duplicate predicate declarations for " ^ quote key ^ Position.here (pos_of x) ^ Position.here (pos_of y))); ); (* queries *) fun lookup_pred_data ctxt name = Option.map rep_pred_data (try (Graph.get_node (PredData.get (Proof_Context.theory_of ctxt))) name) fun the_pred_data ctxt name = (case lookup_pred_data ctxt name of NONE => error ("No such predicate: " ^ quote name) | SOME data => data) val is_registered = is_some oo lookup_pred_data val all_preds_of = Graph.keys o PredData.get o Proof_Context.theory_of val intros_of = map snd o #intros oo the_pred_data val names_of = map fst o #intros oo the_pred_data fun the_elim_of ctxt name = (case #elim (the_pred_data ctxt name) of NONE => error ("No elimination rule for predicate " ^ quote name) | SOME thm => thm) val has_elim = is_some o #elim oo the_pred_data fun function_names_of compilation ctxt name = (case AList.lookup (op =) (#function_names (the_pred_data ctxt name)) compilation of NONE => error ("No " ^ string_of_compilation compilation ^ " functions defined for predicate " ^ quote name) | SOME fun_names => fun_names) fun function_name_of compilation ctxt name mode = (case AList.lookup eq_mode (function_names_of compilation ctxt name) mode of NONE => error ("No " ^ string_of_compilation compilation ^ " function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ quote name) | SOME function_name => function_name) fun modes_of compilation ctxt name = map fst (function_names_of compilation ctxt name) fun all_modes_of compilation ctxt = map_filter (fn name => Option.map (pair name) (try (modes_of compilation ctxt) name)) (all_preds_of ctxt) val all_random_modes_of = all_modes_of Random fun defined_functions compilation ctxt name = (case lookup_pred_data ctxt name of NONE => false | SOME data => AList.defined (op =) (#function_names data) compilation) fun needs_random ctxt s m = member (op =) (#needs_random (the_pred_data ctxt s)) m fun lookup_predfun_data ctxt name mode = Option.map rep_predfun_data (AList.lookup eq_mode (#predfun_data (the_pred_data ctxt name)) mode) fun the_predfun_data ctxt name mode = (case lookup_predfun_data ctxt name mode of NONE => error ("No function defined for mode " ^ string_of_mode mode ^ " of predicate " ^ name) | SOME data => data) val predfun_definition_of = #definition ooo the_predfun_data val predfun_intro_of = #intro ooo the_predfun_data val predfun_elim_of = #elim ooo the_predfun_data val predfun_neg_intro_of = #neg_intro ooo the_predfun_data val intros_graph_of = Graph.map (K (map snd o #intros o rep_pred_data)) o PredData.get o Proof_Context.theory_of fun prove_casesrule ctxt (pred, (pre_cases_rule, nparams)) cases_rule = let val thy = Proof_Context.theory_of ctxt val nargs = length (binder_types (fastype_of pred)) fun meta_eq_of th = th RS @{thm eq_reflection} val tuple_rew_rules = map meta_eq_of [@{thm fst_conv}, @{thm snd_conv}, @{thm prod.inject}] fun instantiate i n ({context = ctxt2, prems, ...}: Subgoal.focus) = let fun inst_pair_of (ix, (ty, t)) = ((ix, ty), t) fun inst_of_matches tts = fold (Pattern.match thy) tts (Vartab.empty, Vartab.empty) |> snd |> Vartab.dest |> map (apsnd (Thm.cterm_of ctxt2) o inst_pair_of) val (cases, (eqs, prems1)) = apsnd (chop (nargs - nparams)) (chop n prems) val case_th = rewrite_rule ctxt2 (@{thm Predicate.eq_is_eq} :: map meta_eq_of eqs) (nth cases (i - 1)) val prems2 = maps (dest_conjunct_prem o rewrite_rule ctxt2 tuple_rew_rules) prems1 val pats = map (swap o HOLogic.dest_eq o HOLogic.dest_Trueprop) (take nargs (Thm.prems_of case_th)) val case_th' = - Thm.instantiate ([], inst_of_matches pats) case_th + Thm.instantiate (TVars.empty, Vars.make (inst_of_matches pats)) case_th OF replicate nargs @{thm refl} val thesis = - Thm.instantiate ([], inst_of_matches (Thm.prems_of case_th' ~~ map Thm.prop_of prems2)) + Thm.instantiate (TVars.empty, + Vars.make (inst_of_matches (Thm.prems_of case_th' ~~ map Thm.prop_of prems2))) case_th' OF prems2 in resolve_tac ctxt2 [thesis] 1 end in Goal.prove ctxt (Term.add_free_names cases_rule []) [] cases_rule (fn {context = ctxt1, ...} => eresolve_tac ctxt1 [pre_cases_rule] 1 THEN (fn st => let val n = Thm.nprems_of st in st |> ALLGOALS (fn i => rewrite_goal_tac ctxt1 @{thms split_paired_all} i THEN SUBPROOF (instantiate i n) ctxt1 i) end)) end (* updaters *) (* fetching introduction rules or registering introduction rules *) val no_compilation = ([], ([], [])) fun fetch_pred_data ctxt name = (case try (Inductive.the_inductive_global ctxt) name of SOME (info as (_, result)) => let val thy = Proof_Context.theory_of ctxt val pos = Position.thread_data () fun is_intro_of intro = let val (const, _) = strip_comb (HOLogic.dest_Trueprop (Thm.concl_of intro)) in (fst (dest_Const const) = name) end; val intros = map (preprocess_intro thy) (filter is_intro_of (#intrs result)) val index = find_index (fn s => s = name) (#names (fst info)) val pre_elim = nth (#elims result) index val pred = nth (#preds result) index val elim_t = mk_casesrule ctxt pred intros val nparams = length (Inductive.params_of (#raw_induct result)) val elim = prove_casesrule ctxt (pred, (pre_elim, nparams)) elim_t in mk_pred_data (pos, (((map (pair NONE) intros, SOME elim), true), no_compilation)) end | NONE => error ("No such predicate: " ^ quote name)) fun add_predfun_data name mode data = let val add = (apsnd o apsnd o apsnd o apfst) (cons (mode, mk_predfun_data data)) in PredData.map (Graph.map_node name (map_pred_data add)) end fun is_inductive_predicate ctxt name = is_some (try (Inductive.the_inductive_global ctxt) name) fun depending_preds_of ctxt (key, value) = let val intros = map (Thm.prop_of o snd) ((#intros o rep_pred_data) value) in fold Term.add_const_names intros [] |> (fn cs => if member (op =) cs \<^const_name>\HOL.eq\ then insert (op =) \<^const_name>\Predicate.eq\ cs else cs) |> filter (fn c => (not (c = key)) andalso (is_inductive_predicate ctxt c orelse is_registered ctxt c)) end; fun add_intro (opt_case_name, thm) thy = let val (name, _) = dest_Const (fst (strip_intro_concl thm)) fun cons_intro gr = (case try (Graph.get_node gr) name of SOME _ => Graph.map_node name (map_pred_data (apsnd (apfst (apfst (apfst (fn intros => intros @ [(opt_case_name, thm)])))))) gr | NONE => Graph.new_node (name, mk_pred_data (Position.thread_data (), (((([(opt_case_name, thm)], NONE), false), no_compilation)))) gr) in PredData.map cons_intro thy end fun set_elim thm = let val (name, _) = dest_Const (fst (strip_comb (HOLogic.dest_Trueprop (hd (Thm.prems_of thm))))) in PredData.map (Graph.map_node name (map_pred_data (apsnd (apfst (apfst (apsnd (K (SOME thm)))))))) end fun register_predicate (constname, intros, elim) thy = let val named_intros = map (pair NONE) intros in if not (member (op =) (Graph.keys (PredData.get thy)) constname) then PredData.map (Graph.new_node (constname, mk_pred_data (Position.thread_data (), (((named_intros, SOME elim), false), no_compilation)))) thy else thy end fun register_intros (constname, pre_intros) thy = let val T = Sign.the_const_type thy constname fun constname_of_intro intr = fst (dest_Const (fst (strip_intro_concl intr))) val _ = if not (forall (fn intr => constname_of_intro intr = constname) pre_intros) then error ("register_intros: Introduction rules of different constants are used\n" ^ "expected rules for " ^ constname ^ ", but received rules for " ^ commas (map constname_of_intro pre_intros)) else () val pred = Const (constname, T) val pre_elim = (Drule.export_without_context o Skip_Proof.make_thm thy) (mk_casesrule (Proof_Context.init_global thy) pred pre_intros) in register_predicate (constname, pre_intros, pre_elim) thy end fun defined_function_of compilation pred = let val set = (apsnd o apsnd o apfst) (cons (compilation, [])) in PredData.map (Graph.map_node pred (map_pred_data set)) end fun set_function_name compilation pred mode name = let val set = (apsnd o apsnd o apfst) (AList.map_default (op =) (compilation, [(mode, name)]) (cons (mode, name))) in PredData.map (Graph.map_node pred (map_pred_data set)) end fun set_needs_random name modes = let val set = (apsnd o apsnd o apsnd o apsnd) (K modes) in PredData.map (Graph.map_node name (map_pred_data set)) end fun extend' value_of edges_of key (G, visited) = let val (G', v) = (case try (Graph.get_node G) key of SOME v => (G, v) | NONE => (Graph.new_node (key, value_of key) G, value_of key)) val (G'', visited') = fold (extend' value_of edges_of) (subtract (op =) visited (edges_of (key, v))) (G', key :: visited) in (fold (Graph.add_edge o (pair key)) (edges_of (key, v)) G'', visited') end; fun extend value_of edges_of key G = fst (extend' value_of edges_of key (G, [])) fun extend_intro_graph names thy = let val ctxt = Proof_Context.init_global thy in PredData.map (fold (extend (fetch_pred_data ctxt) (depending_preds_of ctxt)) names) thy end fun preprocess_intros name thy = PredData.map (Graph.map_node name (map_pred_data (apsnd (apfst (fn (rules, preprocessed) => if preprocessed then (rules, preprocessed) else let val (named_intros, SOME elim) = rules val named_intros' = map (apsnd (preprocess_intro thy)) named_intros val pred = Const (name, Sign.the_const_type thy name) val ctxt = Proof_Context.init_global thy val elim_t = mk_casesrule ctxt pred (map snd named_intros') val elim' = prove_casesrule ctxt (pred, (elim, 0)) elim_t in ((named_intros', SOME elim'), true) end))))) thy (* registration of alternative function names *) structure Alt_Compilations_Data = Theory_Data ( type T = (mode * (compilation_funs -> typ -> term)) list Symtab.table val empty = Symtab.empty val extend = I fun merge data : T = Symtab.merge (K true) data ); fun alternative_compilation_of_global thy pred_name mode = AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get thy) pred_name) mode fun alternative_compilation_of ctxt pred_name mode = AList.lookup eq_mode (Symtab.lookup_list (Alt_Compilations_Data.get (Proof_Context.theory_of ctxt)) pred_name) mode fun force_modes_and_compilations pred_name compilations thy = let (* thm refl is a dummy thm *) val modes = map fst compilations val (needs_random, non_random_modes) = apply2 (map fst) (List.partition (fn (_, (_, random)) => random) compilations) val non_random_dummys = map (rpair "dummy") non_random_modes val all_dummys = map (rpair "dummy") modes val dummy_function_names = map (rpair all_dummys) Predicate_Compile_Aux.random_compilations @ map (rpair non_random_dummys) Predicate_Compile_Aux.non_random_compilations val alt_compilations = map (apsnd fst) compilations in thy |> PredData.map (Graph.new_node (pred_name, mk_pred_data (Position.thread_data (), ((([], SOME @{thm refl}), true), (dummy_function_names, ([], needs_random)))))) |> Alt_Compilations_Data.map (Symtab.insert (K false) (pred_name, alt_compilations)) end fun functional_compilation fun_name mode compfuns T = let val (inpTs, outpTs) = split_map_modeT (fn _ => fn T => (SOME T, NONE)) mode (binder_types T) val bs = map (pair "x") inpTs val bounds = map Bound (rev (0 upto (length bs) - 1)) val f = Const (fun_name, inpTs ---> HOLogic.mk_tupleT outpTs) in fold_rev Term.abs bs (mk_single compfuns (list_comb (f, bounds))) end fun register_alternative_function pred_name mode fun_name = Alt_Compilations_Data.map (Symtab.insert_list (eq_pair eq_mode (K false)) (pred_name, (mode, functional_compilation fun_name mode))) fun force_modes_and_functions pred_name fun_names = force_modes_and_compilations pred_name (map (fn (mode, (fun_name, random)) => (mode, (functional_compilation fun_name mode, random))) fun_names) end diff --git a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_aux.ML @@ -1,1223 +1,1226 @@ (* Title: HOL/Tools/Predicate_Compile/predicate_compile_aux.ML Author: Lukas Bulwahn, TU Muenchen Auxilary functions for predicate compiler. *) signature PREDICATE_COMPILE_AUX = sig val find_indices : ('a -> bool) -> 'a list -> int list (* mode *) datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode val eq_mode : mode * mode -> bool val mode_ord: mode ord val list_fun_mode : mode list -> mode val strip_fun_mode : mode -> mode list val dest_fun_mode : mode -> mode list val dest_tuple_mode : mode -> mode list val all_modes_of_typ : typ -> mode list val all_smodes_of_typ : typ -> mode list val fold_map_aterms_prodT : ('a -> 'a -> 'a) -> (typ -> 'b -> 'a * 'b) -> typ -> 'b -> 'a * 'b val map_filter_prod : (term -> term option) -> term -> term option val replace_ho_args : mode -> term list -> term list -> term list val ho_arg_modes_of : mode -> mode list val ho_argsT_of : mode -> typ list -> typ list val ho_args_of : mode -> term list -> term list val ho_args_of_typ : typ -> term list -> term list val ho_argsT_of_typ : typ list -> typ list val split_map_mode : (mode -> term -> term option * term option) -> mode -> term list -> term list * term list val split_map_modeT : (mode -> typ -> typ option * typ option) -> mode -> typ list -> typ list * typ list val split_mode : mode -> term list -> term list * term list val split_modeT : mode -> typ list -> typ list * typ list val string_of_mode : mode -> string val ascii_string_of_mode : mode -> string (* premises *) datatype indprem = Prem of term | Negprem of term | Sidecond of term | Generator of (string * typ) val dest_indprem : indprem -> term val map_indprem : (term -> term) -> indprem -> indprem (* general syntactic functions *) val is_equationlike : thm -> bool val is_pred_equation : thm -> bool val is_intro : string -> thm -> bool val is_predT : typ -> bool val lookup_constr : Proof.context -> (string * typ) -> int option val is_constrt : Proof.context -> term -> bool val strip_ex : term -> (string * typ) list * term val focus_ex : term -> Name.context -> ((string * typ) list * term) * Name.context val strip_all : term -> (string * typ) list * term val strip_intro_concl : thm -> term * term list (* introduction rule combinators *) val map_atoms : (term -> term) -> term -> term val fold_atoms : (term -> 'a -> 'a) -> term -> 'a -> 'a val fold_map_atoms : (term -> 'a -> term * 'a) -> term -> 'a -> term * 'a val maps_premises : (term -> term list) -> term -> term val map_concl : (term -> term) -> term -> term val map_term : theory -> (term -> term) -> thm -> thm (* split theorems of case expressions *) val prepare_split_thm : Proof.context -> thm -> thm val find_split_thm : theory -> term -> thm option (* datastructures and setup for generic compilation *) datatype compilation_funs = CompilationFuns of { mk_monadT : typ -> typ, dest_monadT : typ -> typ, mk_empty : typ -> term, mk_single : term -> term, mk_bind : term * term -> term, mk_plus : term * term -> term, mk_if : term -> term, mk_iterate_upto : typ -> term * term * term -> term, mk_not : term -> term, mk_map : typ -> typ -> term -> term -> term }; val mk_monadT : compilation_funs -> typ -> typ val dest_monadT : compilation_funs -> typ -> typ val mk_empty : compilation_funs -> typ -> term val mk_single : compilation_funs -> term -> term val mk_bind : compilation_funs -> term * term -> term val mk_plus : compilation_funs -> term * term -> term val mk_if : compilation_funs -> term -> term val mk_iterate_upto : compilation_funs -> typ -> term * term * term -> term val mk_not : compilation_funs -> term -> term val mk_map : compilation_funs -> typ -> typ -> term -> term -> term val funT_of : compilation_funs -> mode -> typ -> typ (* Different compilations *) datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq | Pos_Generator_DSeq | Neg_Generator_DSeq | Pos_Generator_CPS | Neg_Generator_CPS val negative_compilation_of : compilation -> compilation val compilation_for_polarity : bool -> compilation -> compilation val is_depth_limited_compilation : compilation -> bool val string_of_compilation : compilation -> string val compilation_names : (string * compilation) list val non_random_compilations : compilation list val random_compilations : compilation list (* Different options for compiler *) datatype options = Options of { expected_modes : (string * mode list) option, proposed_modes : (string * mode list) list, proposed_names : ((string * mode) * string) list, show_steps : bool, show_proof_trace : bool, show_intermediate_results : bool, show_mode_inference : bool, show_modes : bool, show_compilation : bool, show_caught_failures : bool, show_invalid_clauses : bool, skip_proof : bool, no_topmost_reordering : bool, function_flattening : bool, fail_safe_function_flattening : bool, specialise : bool, no_higher_order_predicate : string list, inductify : bool, detect_switches : bool, smart_depth_limiting : bool, compilation : compilation }; val expected_modes : options -> (string * mode list) option val proposed_modes : options -> string -> mode list option val proposed_names : options -> string -> mode -> string option val show_steps : options -> bool val show_proof_trace : options -> bool val show_intermediate_results : options -> bool val show_mode_inference : options -> bool val show_modes : options -> bool val show_compilation : options -> bool val show_caught_failures : options -> bool val show_invalid_clauses : options -> bool val skip_proof : options -> bool val no_topmost_reordering : options -> bool val function_flattening : options -> bool val fail_safe_function_flattening : options -> bool val specialise : options -> bool val no_higher_order_predicate : options -> string list val is_inductify : options -> bool val detect_switches : options -> bool val smart_depth_limiting : options -> bool val compilation : options -> compilation val default_options : options val bool_options : string list val print_step : options -> string -> unit (* conversions *) val imp_prems_conv : conv -> conv (* simple transformations *) val split_conjuncts_in_assms : Proof.context -> thm -> thm val dest_conjunct_prem : thm -> thm list val expand_tuples : theory -> thm -> thm val case_betapply : theory -> term -> term val eta_contract_ho_arguments : theory -> thm -> thm val remove_equalities : theory -> thm -> thm val remove_pointless_clauses : thm -> thm list val peephole_optimisation : theory -> thm -> thm option (* auxillary *) val unify_consts : theory -> term list -> term list -> (term list * term list) val mk_casesrule : Proof.context -> term -> thm list -> term val preprocess_intro : theory -> thm -> thm val define_quickcheck_predicate : term -> theory -> (((string * typ) * (string * typ) list) * thm) * theory end structure Predicate_Compile_Aux : PREDICATE_COMPILE_AUX = struct (* general functions *) fun comb_option f (SOME x1, SOME x2) = SOME (f (x1, x2)) | comb_option f (NONE, SOME x2) = SOME x2 | comb_option f (SOME x1, NONE) = SOME x1 | comb_option f (NONE, NONE) = NONE fun map2_optional f (x :: xs) (y :: ys) = f x (SOME y) :: (map2_optional f xs ys) | map2_optional f (x :: xs) [] = (f x NONE) :: (map2_optional f xs []) | map2_optional f [] [] = [] fun find_indices f xs = map_filter (fn (i, true) => SOME i | (_, false) => NONE) (map_index (apsnd f) xs) (* mode *) datatype mode = Bool | Input | Output | Pair of mode * mode | Fun of mode * mode (* equality of instantiatedness with respect to equivalences: Pair Input Input == Input and Pair Output Output == Output *) fun eq_mode (Fun (m1, m2), Fun (m3, m4)) = eq_mode (m1, m3) andalso eq_mode (m2, m4) | eq_mode (Pair (m1, m2), Pair (m3, m4)) = eq_mode (m1, m3) andalso eq_mode (m2, m4) | eq_mode (Pair (m1, m2), Input) = eq_mode (m1, Input) andalso eq_mode (m2, Input) | eq_mode (Pair (m1, m2), Output) = eq_mode (m1, Output) andalso eq_mode (m2, Output) | eq_mode (Input, Pair (m1, m2)) = eq_mode (Input, m1) andalso eq_mode (Input, m2) | eq_mode (Output, Pair (m1, m2)) = eq_mode (Output, m1) andalso eq_mode (Output, m2) | eq_mode (Input, Input) = true | eq_mode (Output, Output) = true | eq_mode (Bool, Bool) = true | eq_mode _ = false fun mode_ord (Input, Output) = LESS | mode_ord (Output, Input) = GREATER | mode_ord (Input, Input) = EQUAL | mode_ord (Output, Output) = EQUAL | mode_ord (Bool, Bool) = EQUAL | mode_ord (Pair (m1, m2), Pair (m3, m4)) = prod_ord mode_ord mode_ord ((m1, m2), (m3, m4)) | mode_ord (Fun (m1, m2), Fun (m3, m4)) = prod_ord mode_ord mode_ord ((m1, m2), (m3, m4)) fun list_fun_mode [] = Bool | list_fun_mode (m :: ms) = Fun (m, list_fun_mode ms) (* name: binder_modes? *) fun strip_fun_mode (Fun (mode, mode')) = mode :: strip_fun_mode mode' | strip_fun_mode Bool = [] | strip_fun_mode _ = raise Fail "Bad mode for strip_fun_mode" (* name: strip_fun_mode? *) fun dest_fun_mode (Fun (mode, mode')) = mode :: dest_fun_mode mode' | dest_fun_mode mode = [mode] fun dest_tuple_mode (Pair (mode, mode')) = mode :: dest_tuple_mode mode' | dest_tuple_mode _ = [] fun all_modes_of_typ' (T as Type ("fun", _)) = let val (S, U) = strip_type T in if U = HOLogic.boolT then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_modes_of_typ' S) [Bool] else [Input, Output] end | all_modes_of_typ' (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = map_product (curry Pair) (all_modes_of_typ' T1) (all_modes_of_typ' T2) | all_modes_of_typ' _ = [Input, Output] fun all_modes_of_typ (T as Type ("fun", _)) = let val (S, U) = strip_type T in if U = \<^typ>\bool\ then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_modes_of_typ' S) [Bool] else raise Fail "Invocation of all_modes_of_typ with a non-predicate type" end | all_modes_of_typ \<^typ>\bool\ = [Bool] | all_modes_of_typ _ = raise Fail "Invocation of all_modes_of_typ with a non-predicate type" fun all_smodes_of_typ (T as Type ("fun", _)) = let val (S, U) = strip_type T fun all_smodes (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = map_product (curry Pair) (all_smodes T1) (all_smodes T2) | all_smodes _ = [Input, Output] in if U = HOLogic.boolT then fold_rev (fn m1 => fn m2 => map_product (curry Fun) m1 m2) (map all_smodes S) [Bool] else raise Fail "invalid type for predicate" end fun ho_arg_modes_of mode = let fun ho_arg_mode (m as Fun _) = [m] | ho_arg_mode (Pair (m1, m2)) = ho_arg_mode m1 @ ho_arg_mode m2 | ho_arg_mode _ = [] in maps ho_arg_mode (strip_fun_mode mode) end fun ho_args_of mode ts = let fun ho_arg (Fun _) (SOME t) = [t] | ho_arg (Fun _) NONE = raise Fail "mode and term do not match" | ho_arg (Pair (m1, m2)) (SOME (Const (\<^const_name>\Pair\, _) $ t1 $ t2)) = ho_arg m1 (SOME t1) @ ho_arg m2 (SOME t2) | ho_arg (Pair (m1, m2)) NONE = ho_arg m1 NONE @ ho_arg m2 NONE | ho_arg _ _ = [] in flat (map2_optional ho_arg (strip_fun_mode mode) ts) end fun ho_args_of_typ T ts = let fun ho_arg (T as Type ("fun", [_, _])) (SOME t) = if body_type T = \<^typ>\bool\ then [t] else [] | ho_arg (Type ("fun", [_, _])) NONE = raise Fail "mode and term do not match" | ho_arg (Type(\<^type_name>\Product_Type.prod\, [T1, T2])) (SOME (Const (\<^const_name>\Pair\, _) $ t1 $ t2)) = ho_arg T1 (SOME t1) @ ho_arg T2 (SOME t2) | ho_arg (Type(\<^type_name>\Product_Type.prod\, [T1, T2])) NONE = ho_arg T1 NONE @ ho_arg T2 NONE | ho_arg _ _ = [] in flat (map2_optional ho_arg (binder_types T) ts) end fun ho_argsT_of_typ Ts = let fun ho_arg (T as Type("fun", [_,_])) = if body_type T = \<^typ>\bool\ then [T] else [] | ho_arg (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = ho_arg T1 @ ho_arg T2 | ho_arg _ = [] in maps ho_arg Ts end (* temporary function should be replaced by unsplit_input or so? *) fun replace_ho_args mode hoargs ts = let fun replace (Fun _, _) (arg' :: hoargs') = (arg', hoargs') | replace (Pair (m1, m2), Const (\<^const_name>\Pair\, T) $ t1 $ t2) hoargs = let val (t1', hoargs') = replace (m1, t1) hoargs val (t2', hoargs'') = replace (m2, t2) hoargs' in (Const (\<^const_name>\Pair\, T) $ t1' $ t2', hoargs'') end | replace (_, t) hoargs = (t, hoargs) in fst (fold_map replace (strip_fun_mode mode ~~ ts) hoargs) end fun ho_argsT_of mode Ts = let fun ho_arg (Fun _) T = [T] | ho_arg (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = ho_arg m1 T1 @ ho_arg m2 T2 | ho_arg _ _ = [] in flat (map2 ho_arg (strip_fun_mode mode) Ts) end (* splits mode and maps function to higher-order argument types *) fun split_map_mode f mode ts = let fun split_arg_mode' (m as Fun _) t = f m t | split_arg_mode' (Pair (m1, m2)) (Const (\<^const_name>\Pair\, _) $ t1 $ t2) = let val (i1, o1) = split_arg_mode' m1 t1 val (i2, o2) = split_arg_mode' m2 t2 in (comb_option HOLogic.mk_prod (i1, i2), comb_option HOLogic.mk_prod (o1, o2)) end | split_arg_mode' m t = if eq_mode (m, Input) then (SOME t, NONE) else if eq_mode (m, Output) then (NONE, SOME t) else raise Fail "split_map_mode: mode and term do not match" in (apply2 (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) ts) end (* splits mode and maps function to higher-order argument types *) fun split_map_modeT f mode Ts = let fun split_arg_mode' (m as Fun _) T = f m T | split_arg_mode' (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = let val (i1, o1) = split_arg_mode' m1 T1 val (i2, o2) = split_arg_mode' m2 T2 in (comb_option HOLogic.mk_prodT (i1, i2), comb_option HOLogic.mk_prodT (o1, o2)) end | split_arg_mode' Input T = (SOME T, NONE) | split_arg_mode' Output T = (NONE, SOME T) | split_arg_mode' _ _ = raise Fail "split_modeT': mode and type do not match" in (apply2 (map_filter I) o split_list) (map2 split_arg_mode' (strip_fun_mode mode) Ts) end fun split_mode mode ts = split_map_mode (fn _ => fn _ => (NONE, NONE)) mode ts fun fold_map_aterms_prodT comb f (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) s = let val (x1, s') = fold_map_aterms_prodT comb f T1 s val (x2, s'') = fold_map_aterms_prodT comb f T2 s' in (comb x1 x2, s'') end | fold_map_aterms_prodT _ f T s = f T s fun map_filter_prod f (Const (\<^const_name>\Pair\, _) $ t1 $ t2) = comb_option HOLogic.mk_prod (map_filter_prod f t1, map_filter_prod f t2) | map_filter_prod f t = f t fun split_modeT mode Ts = let fun split_arg_mode (Fun _) _ = ([], []) | split_arg_mode (Pair (m1, m2)) (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) = let val (i1, o1) = split_arg_mode m1 T1 val (i2, o2) = split_arg_mode m2 T2 in (i1 @ i2, o1 @ o2) end | split_arg_mode Input T = ([T], []) | split_arg_mode Output T = ([], [T]) | split_arg_mode _ _ = raise Fail "split_modeT: mode and type do not match" in (apply2 flat o split_list) (map2 split_arg_mode (strip_fun_mode mode) Ts) end fun string_of_mode mode = let fun string_of_mode1 Input = "i" | string_of_mode1 Output = "o" | string_of_mode1 Bool = "bool" | string_of_mode1 mode = "(" ^ (string_of_mode3 mode) ^ ")" and string_of_mode2 (Pair (m1, m2)) = string_of_mode3 m1 ^ " * " ^ string_of_mode2 m2 | string_of_mode2 mode = string_of_mode1 mode and string_of_mode3 (Fun (m1, m2)) = string_of_mode2 m1 ^ " => " ^ string_of_mode3 m2 | string_of_mode3 mode = string_of_mode2 mode in string_of_mode3 mode end fun ascii_string_of_mode mode' = let fun ascii_string_of_mode' Input = "i" | ascii_string_of_mode' Output = "o" | ascii_string_of_mode' Bool = "b" | ascii_string_of_mode' (Pair (m1, m2)) = "P" ^ ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Pair m2 | ascii_string_of_mode' (Fun (m1, m2)) = "F" ^ ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Fun m2 ^ "B" and ascii_string_of_mode'_Fun (Fun (m1, m2)) = ascii_string_of_mode' m1 ^ (if m2 = Bool then "" else "_" ^ ascii_string_of_mode'_Fun m2) | ascii_string_of_mode'_Fun Bool = "B" | ascii_string_of_mode'_Fun m = ascii_string_of_mode' m and ascii_string_of_mode'_Pair (Pair (m1, m2)) = ascii_string_of_mode' m1 ^ ascii_string_of_mode'_Pair m2 | ascii_string_of_mode'_Pair m = ascii_string_of_mode' m in ascii_string_of_mode'_Fun mode' end (* premises *) datatype indprem = Prem of term | Negprem of term | Sidecond of term | Generator of (string * typ) fun dest_indprem (Prem t) = t | dest_indprem (Negprem t) = t | dest_indprem (Sidecond t) = t | dest_indprem (Generator _) = raise Fail "cannot destruct generator" fun map_indprem f (Prem t) = Prem (f t) | map_indprem f (Negprem t) = Negprem (f t) | map_indprem f (Sidecond t) = Sidecond (f t) | map_indprem f (Generator (v, T)) = Generator (dest_Free (f (Free (v, T)))) (* general syntactic functions *) fun is_equationlike_term (Const (\<^const_name>\Pure.eq\, _) $ _ $ _) = true | is_equationlike_term (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ _)) = true | is_equationlike_term _ = false val is_equationlike = is_equationlike_term o Thm.prop_of fun is_pred_equation_term (Const (\<^const_name>\Pure.eq\, _) $ u $ v) = (fastype_of u = \<^typ>\bool\) andalso (fastype_of v = \<^typ>\bool\) | is_pred_equation_term _ = false val is_pred_equation = is_pred_equation_term o Thm.prop_of fun is_intro_term constname t = the_default false (try (fn t => case fst (strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl t))) of Const (c, _) => c = constname | _ => false) t) fun is_intro constname t = is_intro_term constname (Thm.prop_of t) fun is_predT (T as Type("fun", [_, _])) = (body_type T = \<^typ>\bool\) | is_predT _ = false fun lookup_constr ctxt = let val tab = Ctr_Sugar.ctr_sugars_of ctxt |> maps (map_filter (try dest_Const) o #ctrs) |> map (fn (c, T) => ((c, (fst o dest_Type o body_type) T), BNF_Util.num_binder_types T)) in fn (c, T) => case body_type T of Type (Tname, _) => AList.lookup (op =) tab (c, Tname) | _ => NONE end; fun is_constrt ctxt = let val lookup_constr = lookup_constr ctxt fun check t = (case strip_comb t of (Var _, []) => true | (Free _, []) => true | (Const cT, ts) => (case lookup_constr cT of SOME i => length ts = i andalso forall check ts | _ => false) | _ => false) in check end fun strip_all t = (Term.strip_all_vars t, Term.strip_all_body t) fun strip_ex (Const (\<^const_name>\Ex\, _) $ Abs (x, T, t)) = let val (xTs, t') = strip_ex t in ((x, T) :: xTs, t') end | strip_ex t = ([], t) fun focus_ex t nctxt = let val ((xs, Ts), t') = apfst split_list (strip_ex t) val (xs', nctxt') = fold_map Name.variant xs nctxt; val ps' = xs' ~~ Ts; val vs = map Free ps'; val t'' = Term.subst_bounds (rev vs, t'); in ((ps', t''), nctxt') end val strip_intro_concl = strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of (* introduction rule combinators *) fun map_atoms f intro = let val (literals, head) = Logic.strip_horn intro fun appl t = (case t of (\<^term>\Not\ $ t') => HOLogic.mk_not (f t') | _ => f t) in Logic.list_implies (map (HOLogic.mk_Trueprop o appl o HOLogic.dest_Trueprop) literals, head) end fun fold_atoms f intro s = let val (literals, _) = Logic.strip_horn intro fun appl t s = (case t of (\<^term>\Not\ $ t') => f t' s | _ => f t s) in fold appl (map HOLogic.dest_Trueprop literals) s end fun fold_map_atoms f intro s = let val (literals, head) = Logic.strip_horn intro fun appl t s = (case t of (\<^term>\Not\ $ t') => apfst HOLogic.mk_not (f t' s) | _ => f t s) val (literals', s') = fold_map appl (map HOLogic.dest_Trueprop literals) s in (Logic.list_implies (map HOLogic.mk_Trueprop literals', head), s') end; fun map_filter_premises f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (map_filter f premises, head) end fun maps_premises f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (maps f premises, head) end fun map_concl f intro = let val (premises, head) = Logic.strip_horn intro in Logic.list_implies (premises, f head) end (* combinators to apply a function to all basic parts of nested products *) fun map_products f (Const (\<^const_name>\Pair\, T) $ t1 $ t2) = Const (\<^const_name>\Pair\, T) $ map_products f t1 $ map_products f t2 | map_products f t = f t (* split theorems of case expressions *) fun prepare_split_thm ctxt split_thm = (split_thm RS @{thm iffD2}) |> Local_Defs.unfold0 ctxt [@{thm atomize_conjL[symmetric]}, @{thm atomize_all[symmetric]}, @{thm atomize_imp[symmetric]}] fun find_split_thm thy (Const (name, _)) = Option.map #split (Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy) name) | find_split_thm _ _ = NONE (* lifting term operations to theorems *) fun map_term thy f th = Skip_Proof.make_thm thy (f (Thm.prop_of th)) (* fun equals_conv lhs_cv rhs_cv ct = case Thm.term_of ct of Const (@{const_name Pure.eq}, _) $ _ $ _ => Conv.arg_conv cv ct | _ => error "equals_conv" *) (* Different compilations *) datatype compilation = Pred | Depth_Limited | Random | Depth_Limited_Random | DSeq | Annotated | Pos_Random_DSeq | Neg_Random_DSeq | New_Pos_Random_DSeq | New_Neg_Random_DSeq | Pos_Generator_DSeq | Neg_Generator_DSeq | Pos_Generator_CPS | Neg_Generator_CPS fun negative_compilation_of Pos_Random_DSeq = Neg_Random_DSeq | negative_compilation_of Neg_Random_DSeq = Pos_Random_DSeq | negative_compilation_of New_Pos_Random_DSeq = New_Neg_Random_DSeq | negative_compilation_of New_Neg_Random_DSeq = New_Pos_Random_DSeq | negative_compilation_of Pos_Generator_DSeq = Neg_Generator_DSeq | negative_compilation_of Neg_Generator_DSeq = Pos_Generator_DSeq | negative_compilation_of Pos_Generator_CPS = Neg_Generator_CPS | negative_compilation_of Neg_Generator_CPS = Pos_Generator_CPS | negative_compilation_of c = c fun compilation_for_polarity false Pos_Random_DSeq = Neg_Random_DSeq | compilation_for_polarity false New_Pos_Random_DSeq = New_Neg_Random_DSeq | compilation_for_polarity _ c = c fun is_depth_limited_compilation c = (c = New_Pos_Random_DSeq) orelse (c = New_Neg_Random_DSeq) orelse (c = Pos_Generator_DSeq) orelse (c = Pos_Generator_DSeq) fun string_of_compilation c = (case c of Pred => "" | Random => "random" | Depth_Limited => "depth limited" | Depth_Limited_Random => "depth limited random" | DSeq => "dseq" | Annotated => "annotated" | Pos_Random_DSeq => "pos_random dseq" | Neg_Random_DSeq => "neg_random_dseq" | New_Pos_Random_DSeq => "new_pos_random dseq" | New_Neg_Random_DSeq => "new_neg_random_dseq" | Pos_Generator_DSeq => "pos_generator_dseq" | Neg_Generator_DSeq => "neg_generator_dseq" | Pos_Generator_CPS => "pos_generator_cps" | Neg_Generator_CPS => "neg_generator_cps") val compilation_names = [("pred", Pred), ("random", Random), ("depth_limited", Depth_Limited), ("depth_limited_random", Depth_Limited_Random), (*("annotated", Annotated),*) ("dseq", DSeq), ("random_dseq", Pos_Random_DSeq), ("new_random_dseq", New_Pos_Random_DSeq), ("generator_dseq", Pos_Generator_DSeq), ("generator_cps", Pos_Generator_CPS)] val non_random_compilations = [Pred, Depth_Limited, DSeq, Annotated] val random_compilations = [Random, Depth_Limited_Random, Pos_Random_DSeq, Neg_Random_DSeq, New_Pos_Random_DSeq, New_Neg_Random_DSeq, Pos_Generator_CPS, Neg_Generator_CPS] (* datastructures and setup for generic compilation *) datatype compilation_funs = CompilationFuns of { mk_monadT : typ -> typ, dest_monadT : typ -> typ, mk_empty : typ -> term, mk_single : term -> term, mk_bind : term * term -> term, mk_plus : term * term -> term, mk_if : term -> term, mk_iterate_upto : typ -> term * term * term -> term, mk_not : term -> term, mk_map : typ -> typ -> term -> term -> term } fun mk_monadT (CompilationFuns funs) = #mk_monadT funs fun dest_monadT (CompilationFuns funs) = #dest_monadT funs fun mk_empty (CompilationFuns funs) = #mk_empty funs fun mk_single (CompilationFuns funs) = #mk_single funs fun mk_bind (CompilationFuns funs) = #mk_bind funs fun mk_plus (CompilationFuns funs) = #mk_plus funs fun mk_if (CompilationFuns funs) = #mk_if funs fun mk_iterate_upto (CompilationFuns funs) = #mk_iterate_upto funs fun mk_not (CompilationFuns funs) = #mk_not funs fun mk_map (CompilationFuns funs) = #mk_map funs (** function types and names of different compilations **) fun funT_of compfuns mode T = let val Ts = binder_types T val (inTs, outTs) = split_map_modeT (fn m => fn T => (SOME (funT_of compfuns m T), NONE)) mode Ts in inTs ---> (mk_monadT compfuns (HOLogic.mk_tupleT outTs)) end (* Different options for compiler *) datatype options = Options of { expected_modes : (string * mode list) option, proposed_modes : (string * mode list) list, proposed_names : ((string * mode) * string) list, show_steps : bool, show_proof_trace : bool, show_intermediate_results : bool, show_mode_inference : bool, show_modes : bool, show_compilation : bool, show_caught_failures : bool, show_invalid_clauses : bool, skip_proof : bool, no_topmost_reordering : bool, function_flattening : bool, specialise : bool, fail_safe_function_flattening : bool, no_higher_order_predicate : string list, inductify : bool, detect_switches : bool, smart_depth_limiting : bool, compilation : compilation } fun expected_modes (Options opt) = #expected_modes opt fun proposed_modes (Options opt) = AList.lookup (op =) (#proposed_modes opt) fun proposed_names (Options opt) name mode = AList.lookup (eq_pair (op =) eq_mode) (#proposed_names opt) (name, mode) fun show_steps (Options opt) = #show_steps opt fun show_intermediate_results (Options opt) = #show_intermediate_results opt fun show_proof_trace (Options opt) = #show_proof_trace opt fun show_modes (Options opt) = #show_modes opt fun show_mode_inference (Options opt) = #show_mode_inference opt fun show_compilation (Options opt) = #show_compilation opt fun show_caught_failures (Options opt) = #show_caught_failures opt fun show_invalid_clauses (Options opt) = #show_invalid_clauses opt fun skip_proof (Options opt) = #skip_proof opt fun function_flattening (Options opt) = #function_flattening opt fun fail_safe_function_flattening (Options opt) = #fail_safe_function_flattening opt fun specialise (Options opt) = #specialise opt fun no_topmost_reordering (Options opt) = #no_topmost_reordering opt fun no_higher_order_predicate (Options opt) = #no_higher_order_predicate opt fun is_inductify (Options opt) = #inductify opt fun compilation (Options opt) = #compilation opt fun detect_switches (Options opt) = #detect_switches opt fun smart_depth_limiting (Options opt) = #smart_depth_limiting opt val default_options = Options { expected_modes = NONE, proposed_modes = [], proposed_names = [], show_steps = false, show_intermediate_results = false, show_proof_trace = false, show_modes = false, show_mode_inference = false, show_compilation = false, show_caught_failures = false, show_invalid_clauses = false, skip_proof = true, no_topmost_reordering = false, function_flattening = false, specialise = false, fail_safe_function_flattening = false, no_higher_order_predicate = [], inductify = false, detect_switches = true, smart_depth_limiting = false, compilation = Pred } val bool_options = ["show_steps", "show_intermediate_results", "show_proof_trace", "show_modes", "show_mode_inference", "show_compilation", "show_invalid_clauses", "skip_proof", "inductify", "no_function_flattening", "detect_switches", "specialise", "no_topmost_reordering", "smart_depth_limiting"] fun print_step options s = if show_steps options then tracing s else () (* simple transformations *) (** tuple processing **) fun rewrite_args [] (pats, intro_t, ctxt) = (pats, intro_t, ctxt) | rewrite_args (arg::args) (pats, intro_t, ctxt) = (case HOLogic.strip_tupleT (fastype_of arg) of (_ :: _ :: _) => let fun rewrite_arg' (Const (\<^const_name>\Pair\, _) $ _ $ t2, Type (\<^type_name>\Product_Type.prod\, [_, T2])) (args, (pats, intro_t, ctxt)) = rewrite_arg' (t2, T2) (args, (pats, intro_t, ctxt)) | rewrite_arg' (t, Type (\<^type_name>\Product_Type.prod\, [T1, T2])) (args, (pats, intro_t, ctxt)) = let val thy = Proof_Context.theory_of ctxt val ([x, y], ctxt') = Variable.variant_fixes ["x", "y"] ctxt val pat = (t, HOLogic.mk_prod (Free (x, T1), Free (y, T2))) val intro_t' = Pattern.rewrite_term thy [pat] [] intro_t val args' = map (Pattern.rewrite_term thy [pat] []) args in rewrite_arg' (Free (y, T2), T2) (args', (pat::pats, intro_t', ctxt')) end | rewrite_arg' _ (args, (pats, intro_t, ctxt)) = (args, (pats, intro_t, ctxt)) val (args', (pats, intro_t', ctxt')) = rewrite_arg' (arg, fastype_of arg) (args, (pats, intro_t, ctxt)) in rewrite_args args' (pats, intro_t', ctxt') end | _ => rewrite_args args (pats, intro_t, ctxt)) fun rewrite_prem atom = let val (_, args) = strip_comb atom in rewrite_args args end fun split_conjuncts_in_assms ctxt th = let val ((_, [fixed_th]), ctxt') = Variable.import false [th] ctxt fun split_conjs i nprems th = if i > nprems then th else (case try (op RSN) (@{thm conjI}, (i, th)) of SOME th' => split_conjs i (nprems + 1) th' | NONE => split_conjs (i + 1) nprems th) in singleton (Variable.export ctxt' ctxt) (split_conjs 1 (Thm.nprems_of fixed_th) fixed_th) end fun dest_conjunct_prem th = (case HOLogic.dest_Trueprop (Thm.prop_of th) of (Const (\<^const_name>\HOL.conj\, _) $ _ $ _) => dest_conjunct_prem (th RS @{thm conjunct1}) @ dest_conjunct_prem (th RS @{thm conjunct2}) | _ => [th]) fun expand_tuples thy intro = let val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) val (((T_insts, t_insts), [intro']), ctxt1) = Variable.import false [intro] ctxt val intro_t = Thm.prop_of intro' val concl = Logic.strip_imp_concl intro_t val (_, args) = strip_comb (HOLogic.dest_Trueprop concl) val (pats', intro_t', ctxt2) = rewrite_args args ([], intro_t, ctxt1) val (pats', _, ctxt3) = fold_atoms rewrite_prem intro_t' (pats', intro_t', ctxt2) fun rewrite_pat (ct1, ct2) = (ct1, Thm.cterm_of ctxt3 (Pattern.rewrite_term thy pats' [] (Thm.term_of ct2))) val t_insts' = map rewrite_pat (Vars.dest t_insts) - val intro'' = Thm.instantiate (TVars.dest T_insts, t_insts') intro + val intro'' = Thm.instantiate (T_insts, Vars.make t_insts') intro val [intro'''] = Variable.export ctxt3 ctxt [intro''] val intro'''' = Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}, @{thm prod.inject}]) intro''' (* splitting conjunctions introduced by prod.inject*) val intro''''' = split_conjuncts_in_assms ctxt intro'''' in intro''''' end (** making case distributivity rules **) (*** this should be part of the datatype package ***) fun datatype_name_of_case_name thy = Ctr_Sugar.ctr_sugar_of_case (Proof_Context.init_global thy) #> the #> #ctrs #> hd #> fastype_of #> body_type #> dest_Type #> fst fun make_case_comb thy Tcon = let val ctxt = Proof_Context.init_global thy val SOME {casex, ...} = Ctr_Sugar.ctr_sugar_of ctxt Tcon val casex' = Type.legacy_freeze casex val Ts = BNF_Util.binder_fun_types (fastype_of casex') in list_comb (casex', map_index (fn (j, T) => Free ("f" ^ string_of_int j, T)) Ts) end fun make_case_distrib thy Tcon = let val comb = make_case_comb thy Tcon; val Type ("fun", [T, T']) = fastype_of comb; val (Const (case_name, _), fs) = strip_comb comb val used = Term.add_tfree_names comb [] val U = TFree (singleton (Name.variant_list used) "'t", \<^sort>\type\) val x = Free ("x", T) val f = Free ("f", T' --> U) fun apply_f f' = let val Ts = binder_types (fastype_of f') val bs = map Bound ((length Ts - 1) downto 0) in fold_rev absdummy Ts (f $ (list_comb (f', bs))) end val fs' = map apply_f fs val case_c' = Const (case_name, (map fastype_of fs') @ [T] ---> U) in HOLogic.mk_eq (f $ (comb $ x), list_comb (case_c', fs') $ x) end fun case_rewrite thy Tcon = (Drule.export_without_context o Skip_Proof.make_thm thy o HOLogic.mk_Trueprop) (make_case_distrib thy Tcon) fun instantiated_case_rewrite thy Tcon = let val th = case_rewrite thy Tcon val ctxt = Proof_Context.init_global thy val f = fst (strip_comb (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th))))) val Type ("fun", [uninst_T, uninst_T']) = fastype_of f val ([yname], ctxt') = Variable.add_fixes ["y"] ctxt val T' = TFree ("'t'", \<^sort>\type\) val U = TFree ("'u", \<^sort>\type\) val y = Free (yname, U) val f' = absdummy (U --> T') (Bound 0 $ y) val th' = Thm.instantiate - ([(dest_TVar uninst_T, Thm.ctyp_of ctxt' (U --> T')), + (TVars.make [(dest_TVar uninst_T, Thm.ctyp_of ctxt' (U --> T')), (dest_TVar uninst_T', Thm.ctyp_of ctxt' T')], - [((fst (dest_Var f), (U --> T') --> T'), Thm.cterm_of ctxt' f')]) th + Vars.make [((fst (dest_Var f), (U --> T') --> T'), Thm.cterm_of ctxt' f')]) th val [th'] = Variable.export (Variable.declare_thm th' ctxt') ctxt [th'] in th' end fun case_betapply thy t = let val case_name = fst (dest_Const (fst (strip_comb t))) val Tcon = datatype_name_of_case_name thy case_name val th = instantiated_case_rewrite thy Tcon in Raw_Simplifier.rewrite_term thy [th RS @{thm eq_reflection}] [] t end (*** conversions ***) fun imp_prems_conv cv ct = (case Thm.term_of ct of Const (\<^const_name>\Pure.imp\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv) (imp_prems_conv cv) ct | _ => Conv.all_conv ct) (** eta contract higher-order arguments **) fun eta_contract_ho_arguments thy intro = let fun f atom = list_comb (apsnd ((map o map_products) Envir.eta_contract) (strip_comb atom)) in map_term thy (map_concl f o map_atoms f) intro end (** remove equalities **) fun remove_equalities thy intro = let fun remove_eqs intro_t = let val (prems, concl) = Logic.strip_horn intro_t fun remove_eq (prems, concl) = let fun removable_eq prem = (case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) prem of SOME (lhs, rhs) => (case lhs of Var _ => true | _ => (case rhs of Var _ => true | _ => false)) | NONE => false) in (case find_first removable_eq prems of NONE => (prems, concl) | SOME eq => let val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop eq) val prems' = remove (op =) eq prems val subst = (case lhs of (v as Var _) => (fn t => if t = v then rhs else t) | _ => (case rhs of (v as Var _) => (fn t => if t = v then lhs else t))) in remove_eq (map (map_aterms subst) prems', map_aterms subst concl) end) end in Logic.list_implies (remove_eq (prems, concl)) end in map_term thy remove_eqs intro end (* Some last processing *) fun remove_pointless_clauses intro = if Logic.strip_imp_prems (Thm.prop_of intro) = [\<^prop>\False\] then [] else [intro] (* some peephole optimisations *) fun peephole_optimisation thy intro = let val ctxt = Proof_Context.init_global thy (* FIXME proper context!? *) val process = rewrite_rule ctxt (Named_Theorems.get ctxt \<^named_theorems>\code_pred_simp\) fun process_False intro_t = if member (op =) (Logic.strip_imp_prems intro_t) \<^prop>\False\ then NONE else SOME intro_t fun process_True intro_t = map_filter_premises (fn p => if p = \<^prop>\True\ then NONE else SOME p) intro_t in Option.map (Skip_Proof.make_thm thy) (process_False (process_True (Thm.prop_of (process intro)))) end (* importing introduction rules *) fun import_intros inp_pred [] ctxt = let val (outp_pred, ctxt') = yield_singleton (Variable.import_terms true) inp_pred ctxt val T = fastype_of outp_pred val paramTs = ho_argsT_of_typ (binder_types T) val (param_names, _) = Variable.variant_fixes (map (fn i => "p" ^ (string_of_int i)) (1 upto (length paramTs))) ctxt' val params = map2 (curry Free) param_names paramTs in (((outp_pred, params), []), ctxt') end | import_intros inp_pred (th :: ths) ctxt = let val ((_, [th']), ctxt') = Variable.import true [th] ctxt val thy = Proof_Context.theory_of ctxt' val (pred, args) = strip_intro_concl th' val T = fastype_of pred val ho_args = ho_args_of_typ T args fun subst_of (pred', pred) = let val subst = Sign.typ_match thy (fastype_of pred', fastype_of pred) Vartab.empty handle Type.TYPE_MATCH => error ("Type mismatch of predicate " ^ fst (dest_Const pred) ^ " (trying to match " ^ Syntax.string_of_typ ctxt' (fastype_of pred') ^ " and " ^ Syntax.string_of_typ ctxt' (fastype_of pred) ^ ")" ^ " in " ^ Thm.string_of_thm ctxt' th) in TVars.build (Vartab.fold (fn (xi, (S, T)) => TVars.add ((xi, S), T)) subst) end fun instantiate_typ th = let val (pred', _) = strip_intro_concl th val _ = if not (fst (dest_Const pred) = fst (dest_Const pred')) then raise Fail "Trying to instantiate another predicate" else () val instT = TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) (subst_of (pred', pred)) []; - in Thm.instantiate (instT, []) th end + in Thm.instantiate (TVars.make instT, Vars.empty) th end fun instantiate_ho_args th = let val (_, args') = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl o Thm.prop_of) th val ho_args' = map dest_Var (ho_args_of_typ T args') - in Thm.instantiate ([], ho_args' ~~ map (Thm.cterm_of ctxt') ho_args) th end + in + Thm.instantiate (TVars.empty, Vars.make (ho_args' ~~ map (Thm.cterm_of ctxt') ho_args)) + th + end val outp_pred = Term_Subst.instantiate (subst_of (inp_pred, pred), Vars.empty) inp_pred val ((_, ths'), ctxt1) = Variable.import false (map (instantiate_typ #> instantiate_ho_args) ths) ctxt' in (((outp_pred, ho_args), th' :: ths'), ctxt1) end (* generation of case rules from user-given introduction rules *) fun mk_args2 (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) st = let val (t1, st') = mk_args2 T1 st val (t2, st'') = mk_args2 T2 st' in (HOLogic.mk_prod (t1, t2), st'') end (*| mk_args2 (T as Type ("fun", _)) (params, ctxt) = let val (S, U) = strip_type T in if U = HOLogic.boolT then (hd params, (tl params, ctxt)) else let val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt in (Free (x, T), (params, ctxt')) end end*) | mk_args2 T (params, ctxt) = let val ([x], ctxt') = Variable.variant_fixes ["x"] ctxt in (Free (x, T), (params, ctxt')) end fun mk_casesrule ctxt pred introrules = let (* TODO: can be simplified if parameters are not treated specially ? *) val (((pred, params), intros_th), ctxt1) = import_intros pred introrules ctxt (* TODO: distinct required ? -- test case with more than one parameter! *) val params = distinct (op aconv) params val intros = map Thm.prop_of intros_th val ([propname], ctxt2) = Variable.variant_fixes ["thesis"] ctxt1 val prop = HOLogic.mk_Trueprop (Free (propname, HOLogic.boolT)) val argsT = binder_types (fastype_of pred) (* TODO: can be simplified if parameters are not treated specially ? <-- see uncommented code! *) val (argvs, _) = fold_map mk_args2 argsT (params, ctxt2) fun mk_case intro = let val (_, args) = (strip_comb o HOLogic.dest_Trueprop o Logic.strip_imp_concl) intro val prems = Logic.strip_imp_prems intro val eqprems = map2 (HOLogic.mk_Trueprop oo (curry HOLogic.mk_eq)) argvs args val frees = map Free (fold Term.add_frees (args @ prems) []) in fold Logic.all frees (Logic.list_implies (eqprems @ prems, prop)) end val assm = HOLogic.mk_Trueprop (list_comb (pred, argvs)) val cases = map mk_case intros in Logic.list_implies (assm :: cases, prop) end; (* unifying constants to have the same type variables *) fun unify_consts thy cs intr_ts = let val add_term_consts_2 = fold_aterms (fn Const c => insert (op =) c | _ => I); fun varify (t, (i, ts)) = let val t' = map_types (Logic.incr_tvar (i + 1)) (#2 (Type.varify_global TFrees.empty t)) in (maxidx_of_term t', t' :: ts) end val (i, cs') = List.foldr varify (~1, []) cs val (i', intr_ts') = List.foldr varify (i, []) intr_ts val rec_consts = fold add_term_consts_2 cs' [] val intr_consts = fold add_term_consts_2 intr_ts' [] fun unify (cname, cT) = let val consts = map snd (filter (fn c => fst c = cname) intr_consts) in fold (Sign.typ_unify thy) ((replicate (length consts) cT) ~~ consts) end val (env, _) = fold unify rec_consts (Vartab.empty, i') val subst = map_types (Envir.norm_type env) in (map subst cs', map subst intr_ts') end handle Type.TUNIFY => (warning "Occurrences of recursive constant have non-unifiable types"; (cs, intr_ts)) (* preprocessing rules *) fun preprocess_equality thy rule = Conv.fconv_rule (imp_prems_conv (HOLogic.Trueprop_conv (Conv.try_conv (Conv.rewr_conv (Thm.symmetric @{thm Predicate.eq_is_eq}))))) (Thm.transfer thy rule) fun preprocess_intro thy = expand_tuples thy #> preprocess_equality thy (* defining a quickcheck predicate *) fun strip_imp_prems (Const(\<^const_name>\HOL.implies\, _) $ A $ B) = A :: strip_imp_prems B | strip_imp_prems _ = []; fun strip_imp_concl (Const(\<^const_name>\HOL.implies\, _) $ _ $ B) = strip_imp_concl B | strip_imp_concl A = A; fun strip_horn A = (strip_imp_prems A, strip_imp_concl A) fun define_quickcheck_predicate t thy = let val (vs, t') = strip_abs t val vs' = Variable.variant_frees (Proof_Context.init_global thy) [] vs (* FIXME proper context!? *) val t'' = subst_bounds (map Free (rev vs'), t') val (prems, concl) = strip_horn t'' val constname = "quickcheck" val full_constname = Sign.full_bname thy constname val constT = map snd vs' ---> \<^typ>\bool\ val thy1 = Sign.add_consts [(Binding.name constname, constT, NoSyn)] thy val const = Const (full_constname, constT) val t = Logic.list_implies (map HOLogic.mk_Trueprop (prems @ [HOLogic.mk_not concl]), HOLogic.mk_Trueprop (list_comb (const, map Free vs'))) val intro = Goal.prove (Proof_Context.init_global thy1) (map fst vs') [] t (fn {context = ctxt, ...} => ALLGOALS (Skip_Proof.cheat_tac ctxt)) in ((((full_constname, constT), vs'), intro), thy1) end end diff --git a/src/HOL/Tools/Qelim/cooper.ML b/src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML +++ b/src/HOL/Tools/Qelim/cooper.ML @@ -1,913 +1,918 @@ (* Title: HOL/Tools/Qelim/cooper.ML Author: Amine Chaieb, TU Muenchen Presburger arithmetic by Cooper's algorithm. *) signature COOPER = sig type entry val get: Proof.context -> entry val del: term list -> attribute val add: term list -> attribute exception COOPER of string val conv: Proof.context -> conv val tac: bool -> thm list -> thm list -> Proof.context -> int -> tactic end; structure Cooper: COOPER = struct type entry = simpset * term list; val allowed_consts = [\<^term>\(+) :: int => _\, \<^term>\(+) :: nat => _\, \<^term>\(-) :: int => _\, \<^term>\(-) :: nat => _\, \<^term>\(*) :: int => _\, \<^term>\(*) :: nat => _\, \<^term>\(div) :: int => _\, \<^term>\(div) :: nat => _\, \<^term>\(mod) :: int => _\, \<^term>\(mod) :: nat => _\, \<^term>\HOL.conj\, \<^term>\HOL.disj\, \<^term>\HOL.implies\, \<^term>\(=) :: int => _\, \<^term>\(=) :: nat => _\, \<^term>\(=) :: bool => _\, \<^term>\(<) :: int => _\, \<^term>\(<) :: nat => _\, \<^term>\(<=) :: int => _\, \<^term>\(<=) :: nat => _\, \<^term>\(dvd) :: int => _\, \<^term>\(dvd) :: nat => _\, \<^term>\abs :: int => _\, \<^term>\max :: int => _\, \<^term>\max :: nat => _\, \<^term>\min :: int => _\, \<^term>\min :: nat => _\, \<^term>\uminus :: int => _\, (*@ {term "uminus :: nat => _"},*) \<^term>\Not\, \<^term>\Suc\, \<^term>\Ex :: (int => _) => _\, \<^term>\Ex :: (nat => _) => _\, \<^term>\All :: (int => _) => _\, \<^term>\All :: (nat => _) => _\, \<^term>\nat\, \<^term>\int\, \<^term>\Num.One\, \<^term>\Num.Bit0\, \<^term>\Num.Bit1\, \<^term>\Num.numeral :: num => int\, \<^term>\Num.numeral :: num => nat\, \<^term>\0::int\, \<^term>\1::int\, \<^term>\0::nat\, \<^term>\1::nat\, \<^term>\True\, \<^term>\False\]; structure Data = Generic_Data ( type T = simpset * term list; val empty = (HOL_ss, allowed_consts); val extend = I; fun merge ((ss1, ts1), (ss2, ts2)) = (merge_ss (ss1, ss2), Library.merge (op aconv) (ts1, ts2)); ); val get = Data.get o Context.Proof; fun add ts = Thm.declaration_attribute (fn th => fn context => context |> Data.map (fn (ss, ts') => (simpset_map (Context.proof_of context) (fn ctxt => ctxt addsimps [th]) ss, merge (op aconv) (ts', ts)))) fun del ts = Thm.declaration_attribute (fn th => fn context => context |> Data.map (fn (ss, ts') => (simpset_map (Context.proof_of context) (fn ctxt => ctxt delsimps [th]) ss, subtract (op aconv) ts' ts))) fun simp_thms_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms}); val FWD = Drule.implies_elim_list; val true_tm = \<^cterm>\True\; val false_tm = \<^cterm>\False\; val presburger_ss = simpset_of (\<^context> addsimps @{thms zdvd1_eq}); val lin_ss = simpset_of (put_simpset presburger_ss \<^context> addsimps (@{thms dvd_eq_mod_eq_0 add.assoc [where 'a = int] add.commute [where 'a = int] add.left_commute [where 'a = int] mult.assoc [where 'a = int] mult.commute [where 'a = int] mult.left_commute [where 'a = int] })); val iT = HOLogic.intT val bT = HOLogic.boolT; val dest_number = HOLogic.dest_number #> snd; val perhaps_number = try dest_number; val is_number = can dest_number; val [miconj, midisj, mieq, mineq, milt, mile, migt, mige, midvd, mindvd, miP] = map (Thm.instantiate' [SOME \<^ctyp>\int\] []) @{thms "minf"}; val [infDconj, infDdisj, infDdvd,infDndvd,infDP] = map (Thm.instantiate' [SOME \<^ctyp>\int\] []) @{thms "inf_period"}; val [piconj, pidisj, pieq,pineq,pilt,pile,pigt,pige,pidvd,pindvd,piP] = map (Thm.instantiate' [SOME \<^ctyp>\int\] []) @{thms "pinf"}; val [miP, piP] = map (Thm.instantiate' [SOME \<^ctyp>\bool\] []) [miP, piP]; val infDP = Thm.instantiate' (map SOME [\<^ctyp>\int\, \<^ctyp>\bool\]) [] infDP; val [[asetconj, asetdisj, aseteq, asetneq, asetlt, asetle, asetgt, asetge, asetdvd, asetndvd,asetP], [bsetconj, bsetdisj, bseteq, bsetneq, bsetlt, bsetle, bsetgt, bsetge, bsetdvd, bsetndvd,bsetP]] = [@{thms "aset"}, @{thms "bset"}]; val [cpmi, cppi] = [@{thm "cpmi"}, @{thm "cppi"}]; val unity_coeff_ex = Thm.instantiate' [SOME \<^ctyp>\int\] [] @{thm "unity_coeff_ex"}; val [zdvd_mono,simp_from_to,all_not_ex] = [@{thm "zdvd_mono"}, @{thm "simp_from_to"}, @{thm "all_not_ex"}]; val [dvd_uminus, dvd_uminus'] = @{thms "uminus_dvd_conv"}; val eval_ss = simpset_of (put_simpset presburger_ss \<^context> addsimps [simp_from_to] delsimps [insert_iff, bex_triv]); fun eval_conv ctxt = Simplifier.rewrite (put_simpset eval_ss ctxt); (* recognising cterm without moving to terms *) datatype fm = And of cterm*cterm| Or of cterm*cterm| Eq of cterm | NEq of cterm | Lt of cterm | Le of cterm | Gt of cterm | Ge of cterm | Dvd of cterm*cterm | NDvd of cterm*cterm | Nox fun whatis x ct = ( case Thm.term_of ct of Const(\<^const_name>\HOL.conj\,_)$_$_ => And (Thm.dest_binop ct) | Const (\<^const_name>\HOL.disj\,_)$_$_ => Or (Thm.dest_binop ct) | Const (\<^const_name>\HOL.eq\,_)$y$_ => if Thm.term_of x aconv y then Eq (Thm.dest_arg ct) else Nox | Const (\<^const_name>\Not\,_) $ (Const (\<^const_name>\HOL.eq\,_)$y$_) => if Thm.term_of x aconv y then NEq (funpow 2 Thm.dest_arg ct) else Nox | Const (\<^const_name>\Orderings.less\, _) $ y$ z => if Thm.term_of x aconv y then Lt (Thm.dest_arg ct) else if Thm.term_of x aconv z then Gt (Thm.dest_arg1 ct) else Nox | Const (\<^const_name>\Orderings.less_eq\, _) $ y $ z => if Thm.term_of x aconv y then Le (Thm.dest_arg ct) else if Thm.term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox | Const (\<^const_name>\Rings.dvd\,_)$_$(Const(\<^const_name>\Groups.plus\,_)$y$_) => if Thm.term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox | Const (\<^const_name>\Not\,_) $ (Const (\<^const_name>\Rings.dvd\,_)$_$(Const(\<^const_name>\Groups.plus\,_)$y$_)) => if Thm.term_of x aconv y then NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox | _ => Nox) handle CTERM _ => Nox; fun get_pmi_term t = let val (x,eq) = (Thm.dest_abs NONE o Thm.dest_arg o snd o Thm.dest_abs NONE o Thm.dest_arg) (Thm.dest_arg t) in (Thm.lambda x o Thm.dest_arg o Thm.dest_arg) eq end; val get_pmi = get_pmi_term o Thm.cprop_of; val p_v' = (("P'", 0), \<^typ>\int \ bool\); val q_v' = (("Q'", 0), \<^typ>\int \ bool\); val p_v = (("P", 0), \<^typ>\int \ bool\); val q_v = (("Q", 0), \<^typ>\int \ bool\); fun myfwd (th1, th2, th3) p q [(th_1,th_2,th_3), (th_1',th_2',th_3')] = let val (mp', mq') = (get_pmi th_1, get_pmi th_1') - val mi_th = FWD (Drule.instantiate_normalize ([],[(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1) - [th_1, th_1'] - val infD_th = FWD (Drule.instantiate_normalize ([],[(p_v,mp'), (q_v, mq')]) th3) [th_3,th_3'] - val set_th = FWD (Drule.instantiate_normalize ([],[(p_v,p), (q_v,q)]) th2) [th_2, th_2'] + val mi_th = + FWD (Drule.instantiate_normalize + (TVars.empty, Vars.make [(p_v,p),(q_v,q), (p_v',mp'),(q_v',mq')]) th1) [th_1, th_1'] + val infD_th = + FWD (Drule.instantiate_normalize (TVars.empty, Vars.make [(p_v,mp'), (q_v, mq')]) th3) [th_3,th_3'] + val set_th = + FWD (Drule.instantiate_normalize (TVars.empty, Vars.make [(p_v,p), (q_v,q)]) th2) [th_2, th_2'] in (mi_th, set_th, infD_th) end; val inst' = fn cts => Thm.instantiate' [] (map SOME cts); val infDTrue = Thm.instantiate' [] [SOME true_tm] infDP; val infDFalse = Thm.instantiate' [] [SOME false_tm] infDP; val cadd = \<^cterm>\(+) :: int => _\ val cmulC = \<^cterm>\(*) :: int => _\ val cminus = \<^cterm>\(-) :: int => _\ val cone = \<^cterm>\1 :: int\ val [addC, mulC, subC] = map Thm.term_of [cadd, cmulC, cminus] val [zero, one] = [\<^term>\0 :: int\, \<^term>\1 :: int\]; fun numeral1 f n = HOLogic.mk_number iT (f (dest_number n)); fun numeral2 f m n = HOLogic.mk_number iT (f (dest_number m) (dest_number n)); val [minus1,plus1] = map (fn c => fn t => Thm.apply (Thm.apply c t) cone) [cminus,cadd]; fun decomp_pinf x dvd inS [aseteq, asetneq, asetlt, asetle, asetgt, asetge,asetdvd,asetndvd,asetP, infDdvd, infDndvd, asetconj, asetdisj, infDconj, infDdisj] cp = case (whatis x cp) of And (p,q) => ([p,q], myfwd (piconj, asetconj, infDconj) (Thm.lambda x p) (Thm.lambda x q)) | Or (p,q) => ([p,q], myfwd (pidisj, asetdisj, infDdisj) (Thm.lambda x p) (Thm.lambda x q)) | Eq t => ([], K (inst' [t] pieq, FWD (inst' [t] aseteq) [inS (plus1 t)], infDFalse)) | NEq t => ([], K (inst' [t] pineq, FWD (inst' [t] asetneq) [inS t], infDTrue)) | Lt t => ([], K (inst' [t] pilt, FWD (inst' [t] asetlt) [inS t], infDFalse)) | Le t => ([], K (inst' [t] pile, FWD (inst' [t] asetle) [inS (plus1 t)], infDFalse)) | Gt t => ([], K (inst' [t] pigt, (inst' [t] asetgt), infDTrue)) | Ge t => ([], K (inst' [t] pige, (inst' [t] asetge), infDTrue)) | Dvd (d,s) => ([],let val dd = dvd d in K (inst' [d,s] pidvd, FWD (inst' [d,s] asetdvd) [dd],FWD (inst' [d,s] infDdvd) [dd]) end) | NDvd(d,s) => ([],let val dd = dvd d in K (inst' [d,s] pindvd, FWD (inst' [d,s] asetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end) | _ => ([], K (inst' [cp] piP, inst' [cp] asetP, inst' [cp] infDP)); fun decomp_minf x dvd inS [bseteq,bsetneq,bsetlt, bsetle, bsetgt, bsetge,bsetdvd,bsetndvd,bsetP, infDdvd, infDndvd, bsetconj, bsetdisj, infDconj, infDdisj] cp = case (whatis x cp) of And (p,q) => ([p,q], myfwd (miconj, bsetconj, infDconj) (Thm.lambda x p) (Thm.lambda x q)) | Or (p,q) => ([p,q], myfwd (midisj, bsetdisj, infDdisj) (Thm.lambda x p) (Thm.lambda x q)) | Eq t => ([], K (inst' [t] mieq, FWD (inst' [t] bseteq) [inS (minus1 t)], infDFalse)) | NEq t => ([], K (inst' [t] mineq, FWD (inst' [t] bsetneq) [inS t], infDTrue)) | Lt t => ([], K (inst' [t] milt, (inst' [t] bsetlt), infDTrue)) | Le t => ([], K (inst' [t] mile, (inst' [t] bsetle), infDTrue)) | Gt t => ([], K (inst' [t] migt, FWD (inst' [t] bsetgt) [inS t], infDFalse)) | Ge t => ([], K (inst' [t] mige,FWD (inst' [t] bsetge) [inS (minus1 t)], infDFalse)) | Dvd (d,s) => ([],let val dd = dvd d in K (inst' [d,s] midvd, FWD (inst' [d,s] bsetdvd) [dd] , FWD (inst' [d,s] infDdvd) [dd]) end) | NDvd (d,s) => ([],let val dd = dvd d in K (inst' [d,s] mindvd, FWD (inst' [d,s] bsetndvd) [dd], FWD (inst' [d,s] infDndvd) [dd]) end) | _ => ([], K (inst' [cp] miP, inst' [cp] bsetP, inst' [cp] infDP)) (* Canonical linear form for terms, formulae etc.. *) fun provelin ctxt t = Goal.prove ctxt [] [] t (fn _ => EVERY [simp_tac (put_simpset lin_ss ctxt) 1, TRY (Lin_Arith.tac ctxt 1)]); fun linear_cmul 0 tm = zero | linear_cmul n tm = case tm of Const (\<^const_name>\Groups.plus\, _) $ a $ b => addC $ linear_cmul n a $ linear_cmul n b | Const (\<^const_name>\Groups.times\, _) $ c $ x => mulC $ numeral1 (fn m => n * m) c $ x | Const (\<^const_name>\Groups.minus\, _) $ a $ b => subC $ linear_cmul n a $ linear_cmul n b | (m as Const (\<^const_name>\Groups.uminus\, _)) $ a => m $ linear_cmul n a | _ => numeral1 (fn m => n * m) tm; fun earlier [] x y = false | earlier (h::t) x y = if h aconv y then false else if h aconv x then true else earlier t x y; fun linear_add vars tm1 tm2 = case (tm1, tm2) of (Const (\<^const_name>\Groups.plus\, _) $ (Const (\<^const_name>\Groups.times\, _) $ c1 $ x1) $ r1, Const (\<^const_name>\Groups.plus\, _) $ (Const (\<^const_name>\Groups.times\, _) $ c2 $ x2) $ r2) => if x1 = x2 then let val c = numeral2 Integer.add c1 c2 in if c = zero then linear_add vars r1 r2 else addC$(mulC$c$x1)$(linear_add vars r1 r2) end else if earlier vars x1 x2 then addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2 else addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2 | (Const (\<^const_name>\Groups.plus\, _) $ (Const (\<^const_name>\Groups.times\, _) $ c1 $ x1) $ r1, _) => addC $ (mulC $ c1 $ x1) $ linear_add vars r1 tm2 | (_, Const (\<^const_name>\Groups.plus\, _) $ (Const (\<^const_name>\Groups.times\, _) $ c2 $ x2) $ r2) => addC $ (mulC $ c2 $ x2) $ linear_add vars tm1 r2 | (_, _) => numeral2 Integer.add tm1 tm2; fun linear_neg tm = linear_cmul ~1 tm; fun linear_sub vars tm1 tm2 = linear_add vars tm1 (linear_neg tm2); exception COOPER of string; fun lint vars tm = if is_number tm then tm else case tm of Const (\<^const_name>\Groups.uminus\, _) $ t => linear_neg (lint vars t) | Const (\<^const_name>\Groups.plus\, _) $ s $ t => linear_add vars (lint vars s) (lint vars t) | Const (\<^const_name>\Groups.minus\, _) $ s $ t => linear_sub vars (lint vars s) (lint vars t) | Const (\<^const_name>\Groups.times\, _) $ s $ t => let val s' = lint vars s val t' = lint vars t in case perhaps_number s' of SOME n => linear_cmul n t' | NONE => (case perhaps_number t' of SOME n => linear_cmul n s' | NONE => raise COOPER "lint: not linear") end | _ => addC $ (mulC $ one $ tm) $ zero; fun lin (vs as _::_) (Const (\<^const_name>\Not\, _) $ (Const (\<^const_name>\Orderings.less\, T) $ s $ t)) = lin vs (Const (\<^const_name>\Orderings.less_eq\, T) $ t $ s) | lin (vs as _::_) (Const (\<^const_name>\Not\,_) $ (Const(\<^const_name>\Orderings.less_eq\, T) $ s $ t)) = lin vs (Const (\<^const_name>\Orderings.less\, T) $ t $ s) | lin vs (Const (\<^const_name>\Not\,T)$t) = Const (\<^const_name>\Not\,T)$ (lin vs t) | lin (vs as _::_) (Const(\<^const_name>\Rings.dvd\,_)$d$t) = HOLogic.mk_binrel \<^const_name>\Rings.dvd\ (numeral1 abs d, lint vs t) | lin (vs as x::_) ((b as Const(\<^const_name>\HOL.eq\,_))$s$t) = (case lint vs (subC$t$s) of (t as _$(m$c$y)$r) => if x <> y then b$zero$t else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r else b$(m$c$y)$(linear_neg r) | t => b$zero$t) | lin (vs as x::_) (b$s$t) = (case lint vs (subC$t$s) of (t as _$(m$c$y)$r) => if x <> y then b$zero$t else if dest_number c < 0 then b$(m$(numeral1 ~ c)$y)$r else b$(linear_neg r)$(m$c$y) | t => b$zero$t) | lin vs fm = fm; fun lint_conv ctxt vs ct = let val t = Thm.term_of ct in (provelin ctxt ((HOLogic.eq_const iT)$t$(lint vs t) |> HOLogic.mk_Trueprop)) RS eq_reflection end; fun is_intrel_type T = T = \<^typ>\int => int => bool\; fun is_intrel (b$_$_) = is_intrel_type (fastype_of b) | is_intrel (\<^term>\Not\$(b$_$_)) = is_intrel_type (fastype_of b) | is_intrel _ = false; fun linearize_conv ctxt vs ct = case Thm.term_of ct of Const(\<^const_name>\Rings.dvd\,_)$_$_ => let val th = Conv.binop_conv (lint_conv ctxt vs) ct val (d',t') = Thm.dest_binop (Thm.rhs_of th) val (dt',tt') = (Thm.term_of d', Thm.term_of t') in if is_number dt' andalso is_number tt' then Conv.fconv_rule (Conv.arg_conv (Simplifier.rewrite (put_simpset presburger_ss ctxt))) th else let val dth = case perhaps_number (Thm.term_of d') of SOME d => if d < 0 then (Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (lint_conv ctxt vs))) (Thm.transitive th (inst' [d',t'] dvd_uminus)) handle TERM _ => th) else th | NONE => raise COOPER "linearize_conv: not linear" val d'' = Thm.rhs_of dth |> Thm.dest_arg1 in case tt' of Const(\<^const_name>\Groups.plus\,_)$(Const(\<^const_name>\Groups.times\,_)$c$_)$_ => let val x = dest_number c in if x < 0 then Conv.fconv_rule (Conv.arg_conv (Conv.arg_conv (lint_conv ctxt vs))) (Thm.transitive dth (inst' [d'',t'] dvd_uminus')) else dth end | _ => dth end end | Const (\<^const_name>\Not\,_)$(Const(\<^const_name>\Rings.dvd\,_)$_$_) => Conv.arg_conv (linearize_conv ctxt vs) ct | t => if is_intrel t then (provelin ctxt ((HOLogic.eq_const bT)$t$(lin vs t) |> HOLogic.mk_Trueprop)) RS eq_reflection else Thm.reflexive ct; val dvdc = \<^cterm>\(dvd) :: int => _\; fun unify ctxt q = let val (e,(cx,p)) = q |> Thm.dest_comb ||> Thm.dest_abs NONE val x = Thm.term_of cx val ins = insert (op = : int * int -> bool) fun h (acc,dacc) t = case Thm.term_of t of Const(s,_)$(Const(\<^const_name>\Groups.times\,_)$c$y)$ _ => if x aconv y andalso member (op =) [\<^const_name>\HOL.eq\, \<^const_name>\Orderings.less\, \<^const_name>\Orderings.less_eq\] s then (ins (dest_number c) acc,dacc) else (acc,dacc) | Const(s,_)$_$(Const(\<^const_name>\Groups.times\,_)$c$y) => if x aconv y andalso member (op =) [\<^const_name>\Orderings.less\, \<^const_name>\Orderings.less_eq\] s then (ins (dest_number c) acc, dacc) else (acc,dacc) | Const(\<^const_name>\Rings.dvd\,_)$_$(Const(\<^const_name>\Groups.plus\,_)$(Const(\<^const_name>\Groups.times\,_)$c$y)$_) => if x aconv y then (acc,ins (dest_number c) dacc) else (acc,dacc) | Const(\<^const_name>\HOL.conj\,_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t) | Const(\<^const_name>\HOL.disj\,_)$_$_ => h (h (acc,dacc) (Thm.dest_arg1 t)) (Thm.dest_arg t) | Const (\<^const_name>\Not\,_)$_ => h (acc,dacc) (Thm.dest_arg t) | _ => (acc, dacc) val (cs,ds) = h ([],[]) p val l = Integer.lcms (union (op =) cs ds) fun cv k ct = let val (tm as b$s$t) = Thm.term_of ct in ((HOLogic.eq_const bT)$tm$(b$(linear_cmul k s)$(linear_cmul k t)) |> HOLogic.mk_Trueprop |> provelin ctxt) RS eq_reflection end fun nzprop x = let val th = Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply \<^cterm>\Not\ (Thm.apply (Thm.apply \<^cterm>\(=) :: int => _\ (Numeral.mk_cnumber \<^ctyp>\int\ x)) \<^cterm>\0::int\))) in Thm.equal_elim (Thm.symmetric th) TrueI end; val notz = let val tab = fold Inttab.update (ds ~~ (map (fn x => nzprop (l div x)) ds)) Inttab.empty in fn ct => the (Inttab.lookup tab (ct |> Thm.term_of |> dest_number)) handle Option.Option => (writeln ("noz: Theorems-Table contains no entry for " ^ Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option.Option) end fun unit_conv t = case Thm.term_of t of Const(\<^const_name>\HOL.conj\,_)$_$_ => Conv.binop_conv unit_conv t | Const(\<^const_name>\HOL.disj\,_)$_$_ => Conv.binop_conv unit_conv t | Const (\<^const_name>\Not\,_)$_ => Conv.arg_conv unit_conv t | Const(s,_)$(Const(\<^const_name>\Groups.times\,_)$c$y)$ _ => if x=y andalso member (op =) [\<^const_name>\HOL.eq\, \<^const_name>\Orderings.less\, \<^const_name>\Orderings.less_eq\] s then cv (l div dest_number c) t else Thm.reflexive t | Const(s,_)$_$(Const(\<^const_name>\Groups.times\,_)$c$y) => if x=y andalso member (op =) [\<^const_name>\Orderings.less\, \<^const_name>\Orderings.less_eq\] s then cv (l div dest_number c) t else Thm.reflexive t | Const(\<^const_name>\Rings.dvd\,_)$d$(r as (Const(\<^const_name>\Groups.plus\,_)$(Const(\<^const_name>\Groups.times\,_)$c$y)$_)) => if x=y then let val k = l div dest_number c val kt = HOLogic.mk_number iT k val th1 = inst' [Thm.dest_arg1 t, Thm.dest_arg t] ((Thm.dest_arg t |> funpow 2 Thm.dest_arg1 |> notz) RS zdvd_mono) val (d',t') = (mulC$kt$d, mulC$kt$r) val thc = (provelin ctxt ((HOLogic.eq_const iT)$d'$(lint [] d') |> HOLogic.mk_Trueprop)) RS eq_reflection val tht = (provelin ctxt ((HOLogic.eq_const iT)$t'$(linear_cmul k r) |> HOLogic.mk_Trueprop)) RS eq_reflection in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule dvdc thc) tht) end else Thm.reflexive t | _ => Thm.reflexive t val uth = unit_conv p val clt = Numeral.mk_cnumber \<^ctyp>\int\ l val ltx = Thm.apply (Thm.apply cmulC clt) cx val th = Drule.arg_cong_rule e (Thm.abstract_rule (fst (dest_Free x )) cx uth) val th' = inst' [Thm.lambda ltx (Thm.rhs_of uth), clt] unity_coeff_ex val thf = Thm.transitive th (Thm.transitive (Thm.symmetric (Thm.beta_conversion true (Thm.cprop_of th' |> Thm.dest_arg1))) th') val (lth,rth) = Thm.dest_comb (Thm.cprop_of thf) |>> Thm.dest_arg |>> Thm.beta_conversion true ||> Thm.beta_conversion true |>> Thm.symmetric in Thm.transitive (Thm.transitive lth thf) rth end; val emptyIS = \<^cterm>\{}::int set\; val insert_tm = \<^cterm>\insert :: int => _\; fun mkISet cts = fold_rev (Thm.apply insert_tm #> Thm.apply) cts emptyIS; val eqelem_imp_imp = @{thm eqelem_imp_iff} RS iffD1; val [A_v,B_v] = map (fn th => Thm.cprop_of th |> funpow 2 Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_arg1 |> Thm.dest_arg |> Thm.dest_abs NONE |> snd |> Thm.dest_fun |> Thm.dest_arg |> Thm.term_of |> dest_Var) [asetP, bsetP]; val D_v = (("D", 0), \<^typ>\int\); fun cooperex_conv ctxt vs q = let val uth = unify ctxt q val (x,p) = Thm.dest_abs NONE (Thm.dest_arg (Thm.rhs_of uth)) val ins = insert (op aconvc) fun h t (bacc,aacc,dacc) = case (whatis x t) of And (p,q) => h q (h p (bacc,aacc,dacc)) | Or (p,q) => h q (h p (bacc,aacc,dacc)) | Eq t => (ins (minus1 t) bacc, ins (plus1 t) aacc,dacc) | NEq t => (ins t bacc, ins t aacc, dacc) | Lt t => (bacc, ins t aacc, dacc) | Le t => (bacc, ins (plus1 t) aacc,dacc) | Gt t => (ins t bacc, aacc,dacc) | Ge t => (ins (minus1 t) bacc, aacc,dacc) | Dvd (d,_) => (bacc,aacc,insert (op =) (Thm.term_of d |> dest_number) dacc) | NDvd (d,_) => (bacc,aacc,insert (op =) (Thm.term_of d|> dest_number) dacc) | _ => (bacc, aacc, dacc) val (b0,a0,ds) = h p ([],[],[]) val d = Integer.lcms ds val cd = Numeral.mk_cnumber \<^ctyp>\int\ d fun divprop x = let val th = Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply (Thm.apply dvdc (Numeral.mk_cnumber \<^ctyp>\int\ x)) cd)) in Thm.equal_elim (Thm.symmetric th) TrueI end; val dvd = let val tab = fold Inttab.update (ds ~~ (map divprop ds)) Inttab.empty in fn ct => the (Inttab.lookup tab (Thm.term_of ct |> dest_number)) handle Option.Option => (writeln ("dvd: Theorems-Table contains no entry for" ^ Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option.Option) end val dp = let val th = Simplifier.rewrite (put_simpset lin_ss ctxt) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply (Thm.apply \<^cterm>\(<) :: int => _\ \<^cterm>\0::int\) cd)) in Thm.equal_elim (Thm.symmetric th) TrueI end; (* A and B set *) local val insI1 = Thm.instantiate' [SOME \<^ctyp>\int\] [] @{thm "insertI1"} val insI2 = Thm.instantiate' [SOME \<^ctyp>\int\] [] @{thm "insertI2"} in fun provein x S = case Thm.term_of S of Const(\<^const_name>\Orderings.bot\, _) => error "Unexpected error in Cooper, please email Amine Chaieb" | Const(\<^const_name>\insert\, _) $ y $ _ => let val (cy,S') = Thm.dest_binop S in if Thm.term_of x aconv y then Thm.instantiate' [] [SOME x, SOME S'] insI1 else Thm.implies_elim (Thm.instantiate' [] [SOME x, SOME S', SOME cy] insI2) (provein x S') end end val al = map (lint vs o Thm.term_of) a0 val bl = map (lint vs o Thm.term_of) b0 val (sl,s0,f,abths,cpth) = if length (distinct (op aconv) bl) <= length (distinct (op aconv) al) then (bl,b0,decomp_minf, - fn B => (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(B_v,B), (D_v,cd)]) th) dp) + fn B => (map (fn th => + Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make [(B_v,B), (D_v,cd)]) th) dp) [bseteq,bsetneq,bsetlt, bsetle, bsetgt,bsetge])@ - (map (Thm.instantiate ([],[(B_v,B), (D_v,cd)])) + (map (Thm.instantiate (TVars.empty, Vars.make [(B_v,B), (D_v,cd)])) [bsetdvd,bsetndvd,bsetP,infDdvd, infDndvd,bsetconj, bsetdisj,infDconj, infDdisj]), cpmi) else (al,a0,decomp_pinf,fn A => - (map (fn th => Thm.implies_elim (Thm.instantiate ([],[(A_v,A), (D_v,cd)]) th) dp) + (map (fn th => + Thm.implies_elim (Thm.instantiate (TVars.empty, Vars.make [(A_v,A), (D_v,cd)]) th) dp) [aseteq,asetneq,asetlt, asetle, asetgt,asetge])@ - (map (Thm.instantiate ([],[(A_v,A), (D_v,cd)])) + (map (Thm.instantiate (TVars.empty, Vars.make [(A_v,A), (D_v,cd)])) [asetdvd,asetndvd, asetP, infDdvd, infDndvd,asetconj, asetdisj,infDconj, infDdisj]),cppi) val cpth = let val sths = map (fn (tl,t0) => if tl = Thm.term_of t0 then Thm.instantiate' [SOME \<^ctyp>\int\] [SOME t0] refl else provelin ctxt ((HOLogic.eq_const iT)$tl$(Thm.term_of t0) |> HOLogic.mk_Trueprop)) (sl ~~ s0) val csl = distinct (op aconvc) (map (Thm.cprop_of #> Thm.dest_arg #> Thm.dest_arg1) sths) val S = mkISet csl val inStab = fold (fn ct => fn tab => Termtab.update (Thm.term_of ct, provein ct S) tab) csl Termtab.empty val eqelem_th = Thm.instantiate' [SOME \<^ctyp>\int\] [NONE,NONE, SOME S] eqelem_imp_imp val inS = let val tab = fold Termtab.update (map (fn eq => let val (s,t) = Thm.cprop_of eq |> Thm.dest_arg |> Thm.dest_binop val th = if s aconvc t then the (Termtab.lookup inStab (Thm.term_of s)) else FWD (Thm.instantiate' [] [SOME s, SOME t] eqelem_th) [eq, the (Termtab.lookup inStab (Thm.term_of s))] in (Thm.term_of t, th) end) sths) Termtab.empty in fn ct => the (Termtab.lookup tab (Thm.term_of ct)) handle Option.Option => (writeln ("inS: No theorem for " ^ Syntax.string_of_term ctxt (Thm.term_of ct)); raise Option.Option) end val (inf, nb, pd) = divide_and_conquer (f x dvd inS (abths S)) p in [dp, inf, nb, pd] MRS cpth end val cpth' = Thm.transitive uth (cpth RS eq_reflection) in Thm.transitive cpth' ((simp_thms_conv ctxt then_conv eval_conv ctxt) (Thm.rhs_of cpth')) end; fun literals_conv bops uops env cv = let fun h t = case Thm.term_of t of b$_$_ => if member (op aconv) bops b then Conv.binop_conv h t else cv env t | u$_ => if member (op aconv) uops u then Conv.arg_conv h t else cv env t | _ => cv env t in h end; fun integer_nnf_conv ctxt env = nnf_conv ctxt then_conv literals_conv [HOLogic.conj, HOLogic.disj] [] env (linearize_conv ctxt); val conv_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps (@{thms simp_thms} @ take 4 @{thms ex_simps} @ [not_all, all_not_ex, @{thm ex_disj_distrib}])); fun conv ctxt p = Qelim.gen_qelim_conv ctxt (Simplifier.rewrite (put_simpset conv_ss ctxt)) (Simplifier.rewrite (put_simpset presburger_ss ctxt)) (Simplifier.rewrite (put_simpset conv_ss ctxt)) (cons o Thm.term_of) (Misc_Legacy.term_frees (Thm.term_of p)) (linearize_conv ctxt) (integer_nnf_conv ctxt) (cooperex_conv ctxt) p handle CTERM _ => raise COOPER "bad cterm" | THM _ => raise COOPER "bad thm" | TYPE _ => raise COOPER "bad type" fun add_bools t = let val ops = [\<^term>\(=) :: int => _\, \<^term>\(<) :: int => _\, \<^term>\(<=) :: int => _\, \<^term>\HOL.conj\, \<^term>\HOL.disj\, \<^term>\HOL.implies\, \<^term>\(=) :: bool => _\, \<^term>\Not\, \<^term>\All :: (int => _) => _\, \<^term>\Ex :: (int => _) => _\, \<^term>\True\, \<^term>\False\]; val is_op = member (op =) ops; val skip = not (fastype_of t = HOLogic.boolT) in case t of (l as f $ a) $ b => if skip orelse is_op f then add_bools b o add_bools l else insert (op aconv) t | f $ a => if skip orelse is_op f then add_bools a o add_bools f else insert (op aconv) t | Abs p => add_bools (snd (Syntax_Trans.variant_abs p)) (* FIXME !? *) | _ => if skip orelse is_op t then I else insert (op aconv) t end; fun descend vs (abs as (_, xT, _)) = let val (xn', p') = Syntax_Trans.variant_abs abs; (* FIXME !? *) in ((xn', xT) :: vs, p') end; local structure Proc = Cooper_Procedure in fun num_of_term vs (Free vT) = Proc.Bound (Proc.nat_of_integer (find_index (fn vT' => vT' = vT) vs)) | num_of_term vs (Term.Bound i) = Proc.Bound (Proc.nat_of_integer i) | num_of_term vs \<^term>\0::int\ = Proc.C (Proc.Int_of_integer 0) | num_of_term vs \<^term>\1::int\ = Proc.C (Proc.Int_of_integer 1) | num_of_term vs (t as Const (\<^const_name>\numeral\, _) $ _) = Proc.C (Proc.Int_of_integer (dest_number t)) | num_of_term vs (Const (\<^const_name>\Groups.uminus\, _) $ t') = Proc.Neg (num_of_term vs t') | num_of_term vs (Const (\<^const_name>\Groups.plus\, _) $ t1 $ t2) = Proc.Add (num_of_term vs t1, num_of_term vs t2) | num_of_term vs (Const (\<^const_name>\Groups.minus\, _) $ t1 $ t2) = Proc.Sub (num_of_term vs t1, num_of_term vs t2) | num_of_term vs (Const (\<^const_name>\Groups.times\, _) $ t1 $ t2) = (case perhaps_number t1 of SOME n => Proc.Mul (Proc.Int_of_integer n, num_of_term vs t2) | NONE => (case perhaps_number t2 of SOME n => Proc.Mul (Proc.Int_of_integer n, num_of_term vs t1) | NONE => raise COOPER "reification: unsupported kind of multiplication")) | num_of_term _ _ = raise COOPER "reification: bad term"; fun fm_of_term ps vs (Const (\<^const_name>\True\, _)) = Proc.T | fm_of_term ps vs (Const (\<^const_name>\False\, _)) = Proc.F | fm_of_term ps vs (Const (\<^const_name>\HOL.conj\, _) $ t1 $ t2) = Proc.And (fm_of_term ps vs t1, fm_of_term ps vs t2) | fm_of_term ps vs (Const (\<^const_name>\HOL.disj\, _) $ t1 $ t2) = Proc.Or (fm_of_term ps vs t1, fm_of_term ps vs t2) | fm_of_term ps vs (Const (\<^const_name>\HOL.implies\, _) $ t1 $ t2) = Proc.Imp (fm_of_term ps vs t1, fm_of_term ps vs t2) | fm_of_term ps vs (\<^term>\(=) :: bool => _ \ $ t1 $ t2) = Proc.Iff (fm_of_term ps vs t1, fm_of_term ps vs t2) | fm_of_term ps vs (Const (\<^const_name>\Not\, _) $ t') = Proc.NOT (fm_of_term ps vs t') | fm_of_term ps vs (Const (\<^const_name>\Ex\, _) $ Abs abs) = Proc.E (uncurry (fm_of_term ps) (descend vs abs)) | fm_of_term ps vs (Const (\<^const_name>\All\, _) $ Abs abs) = Proc.A (uncurry (fm_of_term ps) (descend vs abs)) | fm_of_term ps vs (\<^term>\(=) :: int => _\ $ t1 $ t2) = Proc.Eq (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | fm_of_term ps vs (Const (\<^const_name>\Orderings.less_eq\, _) $ t1 $ t2) = Proc.Le (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | fm_of_term ps vs (Const (\<^const_name>\Orderings.less\, _) $ t1 $ t2) = Proc.Lt (Proc.Sub (num_of_term vs t1, num_of_term vs t2)) | fm_of_term ps vs (Const (\<^const_name>\Rings.dvd\, _) $ t1 $ t2) = (case perhaps_number t1 of SOME n => Proc.Dvd (Proc.Int_of_integer n, num_of_term vs t2) | NONE => raise COOPER "reification: unsupported dvd") | fm_of_term ps vs t = let val n = find_index (fn t' => t aconv t') ps in if n > 0 then Proc.Closed (Proc.nat_of_integer n) else raise COOPER "reification: unknown term" end; fun term_of_num vs (Proc.C i) = HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i) | term_of_num vs (Proc.Bound n) = Free (nth vs (Proc.integer_of_nat n)) | term_of_num vs (Proc.Neg t') = \<^term>\uminus :: int => _\ $ term_of_num vs t' | term_of_num vs (Proc.Add (t1, t2)) = \<^term>\(+) :: int => _\ $ term_of_num vs t1 $ term_of_num vs t2 | term_of_num vs (Proc.Sub (t1, t2)) = \<^term>\(-) :: int => _\ $ term_of_num vs t1 $ term_of_num vs t2 | term_of_num vs (Proc.Mul (i, t2)) = \<^term>\(*) :: int => _\ $ HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i) $ term_of_num vs t2 | term_of_num vs (Proc.CN (n, i, t')) = term_of_num vs (Proc.Add (Proc.Mul (i, Proc.Bound n), t')); fun term_of_fm ps vs Proc.T = \<^term>\True\ | term_of_fm ps vs Proc.F = \<^term>\False\ | term_of_fm ps vs (Proc.And (t1, t2)) = HOLogic.conj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | term_of_fm ps vs (Proc.Or (t1, t2)) = HOLogic.disj $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | term_of_fm ps vs (Proc.Imp (t1, t2)) = HOLogic.imp $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | term_of_fm ps vs (Proc.Iff (t1, t2)) = \<^term>\(=) :: bool => _\ $ term_of_fm ps vs t1 $ term_of_fm ps vs t2 | term_of_fm ps vs (Proc.NOT t') = HOLogic.Not $ term_of_fm ps vs t' | term_of_fm ps vs (Proc.Eq t') = \<^term>\(=) :: int => _ \ $ term_of_num vs t'$ \<^term>\0::int\ | term_of_fm ps vs (Proc.NEq t') = term_of_fm ps vs (Proc.NOT (Proc.Eq t')) | term_of_fm ps vs (Proc.Lt t') = \<^term>\(<) :: int => _ \ $ term_of_num vs t' $ \<^term>\0::int\ | term_of_fm ps vs (Proc.Le t') = \<^term>\(<=) :: int => _ \ $ term_of_num vs t' $ \<^term>\0::int\ | term_of_fm ps vs (Proc.Gt t') = \<^term>\(<) :: int => _ \ $ \<^term>\0::int\ $ term_of_num vs t' | term_of_fm ps vs (Proc.Ge t') = \<^term>\(<=) :: int => _ \ $ \<^term>\0::int\ $ term_of_num vs t' | term_of_fm ps vs (Proc.Dvd (i, t')) = \<^term>\(dvd) :: int => _ \ $ HOLogic.mk_number HOLogic.intT (Proc.integer_of_int i) $ term_of_num vs t' | term_of_fm ps vs (Proc.NDvd (i, t')) = term_of_fm ps vs (Proc.NOT (Proc.Dvd (i, t'))) | term_of_fm ps vs (Proc.Closed n) = nth ps (Proc.integer_of_nat n) | term_of_fm ps vs (Proc.NClosed n) = term_of_fm ps vs (Proc.NOT (Proc.Closed n)); fun procedure t = let val vs = Term.add_frees t []; val ps = add_bools t []; in (term_of_fm ps vs o Proc.pa o fm_of_term ps vs) t end; end; val (_, oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\cooper\, (fn (ctxt, t) => (Thm.cterm_of ctxt o Logic.mk_equals o apply2 HOLogic.mk_Trueprop) (t, procedure t))))); val comp_ss = simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms semiring_norm}); fun strip_objimp ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.implies\, _) $ _ $ _ => let val (A, B) = Thm.dest_binop ct in A :: strip_objimp B end | _ => [ct]); fun strip_objall ct = case Thm.term_of ct of Const (\<^const_name>\All\, _) $ Abs (xn,_,_) => let val (a,(v,t')) = (apsnd (Thm.dest_abs (SOME xn)) o Thm.dest_comb) ct in apfst (cons (a,v)) (strip_objall t') end | _ => ([],ct); local val all_maxscope_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps map (fn th => th RS sym) @{thms "all_simps"}) in fun thin_prems_tac ctxt P = simp_tac (put_simpset all_maxscope_ss ctxt) THEN' CSUBGOAL (fn (p', i) => let val (qvs, p) = strip_objall (Thm.dest_arg p') val (ps, c) = split_last (strip_objimp p) val qs = filter P ps val q = if P c then c else \<^cterm>\False\ val ng = fold_rev (fn (a,v) => fn t => Thm.apply a (Thm.lambda v t)) qvs (fold_rev (fn p => fn q => Thm.apply (Thm.apply \<^cterm>\HOL.implies\ p) q) qs q) val g = Thm.apply (Thm.apply \<^cterm>\(==>)\ (Thm.apply \<^cterm>\Trueprop\ ng)) p' val ntac = (case qs of [] => q aconvc \<^cterm>\False\ | _ => false) in if ntac then no_tac else (case \<^try>\Goal.prove_internal ctxt [] g (K (blast_tac (put_claset HOL_cs ctxt) 1))\ of NONE => no_tac | SOME r => resolve_tac ctxt [r] i) end) end; local fun isnum t = case t of Const(\<^const_name>\Groups.zero\,_) => true | Const(\<^const_name>\Groups.one\,_) => true | \<^term>\Suc\$s => isnum s | \<^term>\nat\$s => isnum s | \<^term>\int\$s => isnum s | Const(\<^const_name>\Groups.uminus\,_)$s => isnum s | Const(\<^const_name>\Groups.plus\,_)$l$r => isnum l andalso isnum r | Const(\<^const_name>\Groups.times\,_)$l$r => isnum l andalso isnum r | Const(\<^const_name>\Groups.minus\,_)$l$r => isnum l andalso isnum r | Const(\<^const_name>\Power.power\,_)$l$r => isnum l andalso isnum r | Const(\<^const_name>\Rings.modulo\,_)$l$r => isnum l andalso isnum r | Const(\<^const_name>\Rings.divide\,_)$l$r => isnum l andalso isnum r | _ => is_number t orelse can HOLogic.dest_nat t fun ty cts t = if not (member (op =) [HOLogic.intT, HOLogic.natT, HOLogic.boolT] (Thm.typ_of_cterm t)) then false else case Thm.term_of t of c$l$r => if member (op =) [\<^term>\(*)::int => _\, \<^term>\(*)::nat => _\] c then not (isnum l orelse isnum r) else not (member (op aconv) cts c) | c$_ => not (member (op aconv) cts c) | c => not (member (op aconv) cts c) val term_constants = let fun h acc t = case t of Const _ => insert (op aconv) t acc | a$b => h (h acc a) b | Abs (_,_,t) => h acc t | _ => acc in h [] end; in fun is_relevant ctxt ct = subset (op aconv) (term_constants (Thm.term_of ct), snd (get ctxt)) andalso forall (fn Free (_, T) => member (op =) [\<^typ>\int\, \<^typ>\nat\] T) (Misc_Legacy.term_frees (Thm.term_of ct)) andalso forall (fn Var (_, T) => member (op =) [\<^typ>\int\, \<^typ>\nat\] T) (Misc_Legacy.term_vars (Thm.term_of ct)); fun int_nat_terms ctxt ct = let val cts = snd (get ctxt) fun h acc t = if ty cts t then insert (op aconvc) t acc else case Thm.term_of t of _$_ => h (h acc (Thm.dest_arg t)) (Thm.dest_fun t) | Abs(_,_,_) => Thm.dest_abs NONE t ||> h acc |> uncurry (remove (op aconvc)) | _ => acc in h [] ct end end; fun generalize_tac ctxt f = CSUBGOAL (fn (p, _) => PRIMITIVE (fn st => let fun all x t = Thm.apply (Thm.cterm_of ctxt (Logic.all_const (Thm.typ_of_cterm x))) (Thm.lambda x t) val ts = sort Thm.fast_term_ord (f p) val p' = fold_rev all ts p in Thm.implies_intr p' (Thm.implies_elim st (fold Thm.forall_elim ts (Thm.assume p'))) end)); local val ss1 = simpset_of (put_simpset comp_ss \<^context> addsimps @{thms simp_thms} @ [@{thm "nat_numeral"} RS sym, @{thm int_dvd_int_iff [symmetric]}, @{thm "of_nat_add"}, @{thm "of_nat_mult"}] @ map (fn r => r RS sym) [@{thm "int_int_eq"}, @{thm "zle_int"}, @{thm "of_nat_less_iff" [where ?'a = int]}] |> Splitter.add_split @{thm "zdiff_int_split"}) val ss2 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps [@{thm "nat_0_le"}, @{thm "of_nat_numeral"}, @{thm "all_nat"}, @{thm "ex_nat"}, @{thm "zero_le_numeral"}, @{thm "le_numeral_extra"(3)}, @{thm "of_nat_0"}, @{thm "of_nat_1"}, @{thm "Suc_eq_plus1"}] |> fold Simplifier.add_cong [@{thm "conj_le_cong"}, @{thm "imp_le_cong"}]) val div_mod_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms simp_thms mod_eq_0_iff_dvd mod_add_left_eq mod_add_right_eq mod_add_eq div_add1_eq [symmetric] div_add1_eq [symmetric] mod_self mod_by_0 div_by_0 div_0 mod_0 div_by_1 mod_by_1 div_by_Suc_0 mod_by_Suc_0 Suc_eq_plus1 ac_simps} addsimprocs [\<^simproc>\cancel_div_mod_nat\, \<^simproc>\cancel_div_mod_int\]) val splits_ss = simpset_of (put_simpset comp_ss \<^context> addsimps [@{thm minus_div_mult_eq_mod [symmetric]}] |> fold Splitter.add_split [@{thm "split_zdiv"}, @{thm "split_zmod"}, @{thm "split_div'"}, @{thm "split_min"}, @{thm "split_max"}, @{thm "abs_split"}]) in fun nat_to_int_tac ctxt = simp_tac (put_simpset ss1 ctxt) THEN_ALL_NEW simp_tac (put_simpset ss2 ctxt) THEN_ALL_NEW simp_tac (put_simpset comp_ss ctxt); fun div_mod_tac ctxt = simp_tac (put_simpset div_mod_ss ctxt); fun splits_tac ctxt = simp_tac (put_simpset splits_ss ctxt); end; fun core_tac ctxt = CSUBGOAL (fn (p, i) => let val cpth = if Config.get ctxt quick_and_dirty then oracle (ctxt, Envir.beta_norm (Envir.eta_long [] (Thm.term_of (Thm.dest_arg p)))) else Conv.arg_conv (conv ctxt) p val p' = Thm.rhs_of cpth val th = Thm.implies_intr p' (Thm.equal_elim (Thm.symmetric cpth) (Thm.assume p')) in resolve_tac ctxt [th] i end handle COOPER _ => no_tac); fun finish_tac ctxt q = SUBGOAL (fn (_, i) => (if q then I else TRY) (resolve_tac ctxt [TrueI] i)); fun tac elim add_ths del_ths = Subgoal.FOCUS_PARAMS (fn {context = ctxt, ...} => let val simpset_ctxt = put_simpset (fst (get ctxt)) ctxt delsimps del_ths addsimps add_ths in Method.insert_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\arith\)) THEN_ALL_NEW Object_Logic.full_atomize_tac ctxt THEN_ALL_NEW CONVERSION Thm.eta_long_conversion THEN_ALL_NEW simp_tac simpset_ctxt THEN_ALL_NEW (TRY o generalize_tac ctxt (int_nat_terms ctxt)) THEN_ALL_NEW Object_Logic.full_atomize_tac ctxt THEN_ALL_NEW (thin_prems_tac ctxt (is_relevant ctxt)) THEN_ALL_NEW Object_Logic.full_atomize_tac ctxt THEN_ALL_NEW div_mod_tac ctxt THEN_ALL_NEW splits_tac ctxt THEN_ALL_NEW simp_tac simpset_ctxt THEN_ALL_NEW CONVERSION Thm.eta_long_conversion THEN_ALL_NEW nat_to_int_tac ctxt THEN_ALL_NEW core_tac ctxt THEN_ALL_NEW finish_tac ctxt elim end 1); (* attribute syntax *) local fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); val constsN = "consts"; val any_keyword = keyword constsN val thms = Scan.repeats (Scan.unless any_keyword Attrib.multi_thm); val terms = thms >> map (Thm.term_of o Drule.dest_term); fun optional scan = Scan.optional scan []; in val _ = Theory.setup (Attrib.setup \<^binding>\presburger\ ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del || optional (keyword constsN |-- terms) >> add) "data for Cooper's algorithm" #> Arith_Data.add_tactic "Presburger arithmetic" (tac true [] [])); end; end; diff --git a/src/HOL/Tools/Quickcheck/random_generators.ML b/src/HOL/Tools/Quickcheck/random_generators.ML --- a/src/HOL/Tools/Quickcheck/random_generators.ML +++ b/src/HOL/Tools/Quickcheck/random_generators.ML @@ -1,490 +1,490 @@ (* Title: HOL/Tools/Quickcheck/random_generators.ML Author: Florian Haftmann, TU Muenchen Random generators for various types. *) signature RANDOM_GENERATORS = sig type seed = Random_Engine.seed val random_fun: typ -> typ -> ('a -> 'a -> bool) -> ('a -> term) -> (seed -> ('b * (unit -> term)) * seed) -> (seed -> seed * seed) -> seed -> (('a -> 'b) * (unit -> term)) * seed val compile_generator_expr: Proof.context -> (term * term list) list -> bool -> int list -> (bool * term list) option * Quickcheck.report option val put_counterexample: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> (bool * term list) option * seed) -> Proof.context -> Proof.context val put_counterexample_report: (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed) -> Proof.context -> Proof.context val instantiate_random_datatype : Old_Datatype_Aux.config -> Old_Datatype_Aux.descr -> (string * sort) list -> string list -> string -> string list * string list -> typ list * typ list -> theory -> theory end; structure Random_Generators : RANDOM_GENERATORS = struct (** abstract syntax **) fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\) val size = \<^term>\i::natural\; val size_pred = \<^term>\(i::natural) - 1\; val size' = \<^term>\j::natural\; val seed = \<^term>\s::Random.seed\; val resultT = \<^typ>\(bool \ term list) option\; (** typ "'a \ 'b" **) type seed = Random_Engine.seed; fun random_fun T1 T2 eq term_of random random_split seed = let val fun_upd = Const (\<^const_name>\fun_upd\, (T1 --> T2) --> T1 --> T2 --> T1 --> T2); val ((_, t2), seed') = random seed; val (seed'', seed''') = random_split seed'; val state = Unsynchronized.ref (seed'', [], fn () => Abs ("x", T1, t2 ())); fun random_fun' x = let val (seed, fun_map, f_t) = ! state; in case AList.lookup (uncurry eq) fun_map x of SOME y => y | NONE => let val t1 = term_of x; val ((y, t2), seed') = random seed; val fun_map' = (x, y) :: fun_map; val f_t' = fn () => fun_upd $ f_t () $ t1 $ t2 (); val _ = state := (seed', fun_map', f_t'); in y end end; fun term_fun' () = #3 (! state) (); in ((random_fun', term_fun'), seed''') end; (** datatypes **) (* definitional scheme for random instances on datatypes *) local val eq = Thm.cprop_of @{thm random_aux_rec} |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_arg; val lhs = eq |> Thm.dest_arg1; val pt_random_aux = lhs |> Thm.dest_fun; val pt_rhs = eq |> Thm.dest_arg |> Thm.dest_fun; val a_v = pt_random_aux |> Thm.ctyp_of_cterm |> Thm.dest_ctyp1 |> Thm.typ_of |> dest_TVar; val rew_thms = map mk_meta_eq [@{thm natural_zero_minus_one}, @{thm Suc_natural_minus_one}, @{thm select_weight_cons_zero}, @{thm beyond_zero}]; val rew_ts = map (Logic.dest_equals o Thm.prop_of) rew_thms; val rew_ss = simpset_of (put_simpset HOL_ss \<^context> addsimps rew_thms); in fun random_aux_primrec eq lthy = let val thy = Proof_Context.theory_of lthy; val ((t_random_aux as Free (random_aux, T)) $ (t_k as Free (v, _)), proto_t_rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq; val Type (_, [_, iT]) = T; val icT = Thm.ctyp_of lthy iT; - val inst = Thm.instantiate_cterm ([(a_v, icT)], []); + val inst = Thm.instantiate_cterm (TVars.make [(a_v, icT)], Vars.empty); fun subst_v t' = map_aterms (fn t as Free (w, _) => if v = w then t' else t | t => t); val t_rhs = lambda t_k proto_t_rhs; val eqs0 = [subst_v \<^term>\0::natural\ eq, subst_v (\<^const>\Code_Numeral.Suc\ $ t_k) eq]; val eqs1 = map (Pattern.rewrite_term thy rew_ts []) eqs0; val ((_, (_, eqs2)), lthy') = lthy |> BNF_LFP_Compat.primrec_simple [((Binding.concealed (Binding.name random_aux), T), NoSyn)] eqs1; val cT_random_aux = inst pt_random_aux |> Thm.term_of |> dest_Var; val cT_rhs = inst pt_rhs |> Thm.term_of |> dest_Var; val rule = @{thm random_aux_rec} |> Drule.instantiate_normalize - ([(a_v, icT)], - [(cT_random_aux, Thm.cterm_of lthy' t_random_aux), + (TVars.make [(a_v, icT)], + Vars.make [(cT_random_aux, Thm.cterm_of lthy' t_random_aux), (cT_rhs, Thm.cterm_of lthy' t_rhs)]); fun tac ctxt = ALLGOALS (resolve_tac ctxt [rule]) THEN ALLGOALS (simp_tac (put_simpset rew_ss ctxt)) THEN (ALLGOALS (Proof_Context.fact_tac ctxt eqs2)); val simp = Goal.prove_sorry lthy' [v] [] eq (tac o #context); in (simp, lthy') end; end; fun random_aux_primrec_multi auxname [eq] lthy = lthy |> random_aux_primrec eq |>> single | random_aux_primrec_multi auxname (eqs as _ :: _ :: _) lthy = let val thy = Proof_Context.theory_of lthy; val (lhss, rhss) = map_split (HOLogic.dest_eq o HOLogic.dest_Trueprop) eqs; val (vs, (arg as Free (v, _)) :: _) = map_split (fn (t1 $ t2) => (t1, t2)) lhss; val Ts = map fastype_of lhss; val tupleT = foldr1 HOLogic.mk_prodT Ts; val aux_lhs = Free ("mutual_" ^ auxname, fastype_of arg --> tupleT) $ arg; val aux_eq = (HOLogic.mk_Trueprop o HOLogic.mk_eq) (aux_lhs, foldr1 HOLogic.mk_prod rhss); fun mk_proj t [T] = [t] | mk_proj t (Ts as T :: (Ts' as _ :: _)) = Const (\<^const_name>\fst\, foldr1 HOLogic.mk_prodT Ts --> T) $ t :: mk_proj (Const (\<^const_name>\snd\, foldr1 HOLogic.mk_prodT Ts --> foldr1 HOLogic.mk_prodT Ts') $ t) Ts'; val projs = mk_proj (aux_lhs) Ts; val proj_eqs = map2 (fn v => fn proj => (v, lambda arg proj)) vs projs; val proj_defs = map2 (fn Free (name, _) => fn (_, rhs) => ((Binding.concealed (Binding.name name), NoSyn), (apfst Binding.concealed Binding.empty_atts, rhs))) vs proj_eqs; val aux_eq' = Pattern.rewrite_term thy proj_eqs [] aux_eq; fun prove_eqs aux_simp proj_defs lthy = let val proj_simps = map (snd o snd) proj_defs; fun tac { context = ctxt, prems = _ } = ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps proj_simps)) THEN ALLGOALS (EqSubst.eqsubst_tac ctxt [0] [aux_simp]) THEN ALLGOALS (simp_tac (put_simpset HOL_ss ctxt addsimps @{thms fst_conv snd_conv})); in (map (fn prop => Goal.prove_sorry lthy [v] [] prop tac) eqs, lthy) end; in lthy |> random_aux_primrec aux_eq' ||>> fold_map Local_Theory.define proj_defs |-> uncurry prove_eqs end; fun random_aux_specification prfx name eqs lthy = let val vs = fold Term.add_free_names ((snd o strip_comb o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o hd) eqs) []; fun mk_proto_eq eq = let val (head $ t $ u, rhs) = (HOLogic.dest_eq o HOLogic.dest_Trueprop) eq; in ((HOLogic.mk_Trueprop o HOLogic.mk_eq) (head, lambda t (lambda u rhs))) end; val proto_eqs = map mk_proto_eq eqs; fun prove_simps proto_simps lthy = let val ext_simps = map (fn thm => fun_cong OF [fun_cong OF [thm]]) proto_simps; val tac = ALLGOALS (Proof_Context.fact_tac lthy ext_simps); in (map (fn prop => Goal.prove_sorry lthy vs [] prop (K tac)) eqs, lthy) end; val b = Binding.concealed (Binding.qualify true prfx (Binding.qualify true name (Binding.name "simps"))); in lthy |> random_aux_primrec_multi (name ^ prfx) proto_eqs |-> prove_simps |-> (fn simps => Local_Theory.note ((b, @{attributes [simp, nitpick_simp]}), simps)) |-> (fn (_, thms) => Code.declare_default_eqns (map (rpair true) thms)) end (* constructing random instances on datatypes *) val random_auxN = "random_aux"; fun mk_random_aux_eqs thy descr vs (names, auxnames) (Ts, Us) = let val mk_const = curry (Sign.mk_const thy); val random_auxsN = map (prefix (random_auxN ^ "_")) (names @ auxnames); val rTs = Ts @ Us; fun random_resultT T = \<^typ>\Random.seed\ --> HOLogic.mk_prodT (termifyT T,\<^typ>\Random.seed\); fun sizeT T = \<^typ>\natural\ --> \<^typ>\natural\ --> T; val random_auxT = sizeT o random_resultT; val random_auxs = map2 (fn s => fn rT => Free (s, random_auxT rT)) random_auxsN rTs; fun mk_random_call T = (NONE, (HOLogic.mk_random T size', T)); fun mk_random_aux_call fTs (k, _) (tyco, Ts) = let val T = Type (tyco, Ts); fun mk_random_fun_lift [] t = t | mk_random_fun_lift (fT :: fTs) t = mk_const \<^const_name>\random_fun_lift\ [fTs ---> T, fT] $ mk_random_fun_lift fTs t; val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size'); val size = Option.map snd (Old_Datatype_Aux.find_shortest_path descr k) |> the_default 0; in (SOME size, (t, fTs ---> T)) end; val tss = Old_Datatype_Aux.interpret_construction descr vs { atyp = mk_random_call, dtyp = mk_random_aux_call }; fun mk_consexpr simpleT (c, xs) = let val (ks, simple_tTs) = split_list xs; val T = termifyT simpleT; val tTs = (map o apsnd) termifyT simple_tTs; val is_rec = exists is_some ks; val k = fold (fn NONE => I | SOME k => Integer.max k) ks 0; val vs = Name.invent_names Name.context "x" (map snd simple_tTs); val tc = HOLogic.mk_return T \<^typ>\Random.seed\ (HOLogic.mk_valtermify_app c vs simpleT); val t = HOLogic.mk_ST (map2 (fn (t, _) => fn (v, T') => ((t, \<^typ>\Random.seed\), SOME ((v, termifyT T')))) tTs vs) tc \<^typ>\Random.seed\ (SOME T, \<^typ>\Random.seed\); val tk = if is_rec then if k = 0 then size else \<^term>\Quickcheck_Random.beyond :: natural \ natural \ natural\ $ HOLogic.mk_number \<^typ>\natural\ k $ size else \<^term>\1::natural\ in (is_rec, HOLogic.mk_prod (tk, t)) end; fun sort_rec xs = map_filter (fn (true, t) => SOME t | _ => NONE) xs @ map_filter (fn (false, t) => SOME t | _ => NONE) xs; val gen_exprss = tss |> (map o apfst) Type |> map (fn (T, cs) => (T, (sort_rec o map (mk_consexpr T)) cs)); fun mk_select (rT, xs) = mk_const \<^const_name>\Quickcheck_Random.collapse\ [\<^typ>\Random.seed\, termifyT rT] $ (mk_const \<^const_name>\Random.select_weight\ [random_resultT rT] $ HOLogic.mk_list (HOLogic.mk_prodT (\<^typ>\natural\, random_resultT rT)) xs) $ seed; val auxs_lhss = map (fn t => t $ size $ size' $ seed) random_auxs; val auxs_rhss = map mk_select gen_exprss; in (random_auxs, auxs_lhss ~~ auxs_rhss) end; fun instantiate_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = let val _ = Old_Datatype_Aux.message config "Creating quickcheck generators ..."; val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; fun mk_size_arg k = case Old_Datatype_Aux.find_shortest_path descr k of SOME (_, l) => if l = 0 then size else \<^term>\max :: natural \ natural \ natural\ $ HOLogic.mk_number \<^typ>\natural\ l $ size | NONE => size; val (random_auxs, auxs_eqs) = (apsnd o map) mk_prop_eq (mk_random_aux_eqs thy descr vs (names, auxnames) (Ts, Us)); val random_defs = map_index (fn (k, T) => mk_prop_eq (HOLogic.mk_random T size, nth random_auxs k $ mk_size_arg k $ size)) Ts; in thy |> Class.instantiation (tycos, vs, \<^sort>\random\) |> random_aux_specification prfx random_auxN auxs_eqs |> `(fn lthy => map (Syntax.check_term lthy) random_defs) |-> (fn random_defs' => fold_map (fn random_def => Specification.definition NONE [] [] ((Binding.concealed Binding.empty, []), random_def)) random_defs') |> snd |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) end; (** building and compiling generator expressions **) structure Data = Proof_Data ( type T = (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> (bool * term list) option * seed) * (unit -> Code_Numeral.natural -> bool -> Code_Numeral.natural -> seed -> ((bool * term list) option * (bool list * bool)) * seed); val empty: T = (fn () => raise Fail "counterexample", fn () => raise Fail "counterexample_report"); fun init _ = empty; ); val get_counterexample = #1 o Data.get; val get_counterexample_report = #2 o Data.get; val put_counterexample = Data.map o @{apply 2(1)} o K; val put_counterexample_report = Data.map o @{apply 2(2)} o K; val target = "Quickcheck"; fun mk_generator_expr ctxt (t, _) = let val thy = Proof_Context.theory_of ctxt val prop = fold_rev absfree (Term.add_frees t []) t val Ts = (map snd o fst o strip_abs) prop val bound_max = length Ts - 1; val bounds = map_index (fn (i, ty) => (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; val result = list_comb (prop, map (fn (i, _, _, _) => Bound i) bounds); val terms = HOLogic.mk_list \<^typ>\term\ (map (fn (_, i, _, _) => Bound i $ \<^term>\()\) bounds); val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt val genuine_only = Free (genuine_only_name, \<^typ>\bool\) val none_t = Const (\<^const_name>\None\, resultT) val check = Quickcheck_Common.mk_safe_if genuine_only none_t (result, none_t, fn genuine => \<^term>\Some :: bool \ term list => (bool \ term list) option\ $ HOLogic.mk_prod (Quickcheck_Common.reflect_bool genuine, terms)) val return = HOLogic.pair_const resultT \<^typ>\Random.seed\; fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\scomp\, liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun mk_case_prod T = Sign.mk_const thy (\<^const_name>\case_prod\, [T, \<^typ>\unit \ term\, liftT resultT \<^typ>\Random.seed\]); fun mk_scomp_split T t t' = mk_scomp (mk_termtyp T) resultT \<^typ>\Random.seed\ t (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\unit => term\, t'))); fun mk_bindclause (_, _, i, T) = mk_scomp_split T (Sign.mk_const thy (\<^const_name>\Quickcheck_Random.random\, [T]) $ Bound i); in lambda genuine_only (Abs ("n", \<^typ>\natural\, fold_rev mk_bindclause bounds (return $ check true))) end; fun mk_reporting_generator_expr ctxt (t, _) = let val thy = Proof_Context.theory_of ctxt val resultT = \<^typ>\(bool \ term list) option \ (bool list \ bool)\ val prop = fold_rev absfree (Term.add_frees t []) t val Ts = (map snd o fst o strip_abs) prop val bound_max = length Ts - 1 val bounds = map_index (fn (i, ty) => (2 * (bound_max - i) + 1, 2 * (bound_max - i), 2 * i, ty)) Ts; val prop' = betapplys (prop, map (fn (i, _, _, _) => Bound i) bounds); val terms = HOLogic.mk_list \<^typ>\term\ (map (fn (_, i, _, _) => Bound i $ \<^term>\()\) bounds) val (assms, concl) = Quickcheck_Common.strip_imp prop' val return = HOLogic.pair_const resultT \<^typ>\Random.seed\; fun mk_assms_report i = HOLogic.mk_prod (\<^term>\None :: (bool \ term list) option\, HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate i \<^term>\True\ @ replicate (length assms - i) \<^term>\False\), \<^term>\False\)) fun mk_concl_report b = HOLogic.mk_prod (HOLogic.mk_list HOLogic.boolT (replicate (length assms) \<^term>\True\), Quickcheck_Common.reflect_bool b) val ([genuine_only_name], _) = Variable.variant_fixes ["genuine_only"] ctxt val genuine_only = Free (genuine_only_name, \<^typ>\bool\) val none_t = HOLogic.mk_prod (\<^term>\None :: (bool \ term list) option\, mk_concl_report true) val concl_check = Quickcheck_Common.mk_safe_if genuine_only none_t (concl, none_t, fn genuine => HOLogic.mk_prod (\<^term>\Some :: bool \ term list => (bool \ term list) option\ $ HOLogic.mk_prod (Quickcheck_Common.reflect_bool genuine, terms), mk_concl_report false)) val check = fold_rev (fn (i, assm) => fn t => Quickcheck_Common.mk_safe_if genuine_only (mk_assms_report i) (HOLogic.mk_not assm, mk_assms_report i, t)) (map_index I assms) concl_check fun liftT T sT = sT --> HOLogic.mk_prodT (T, sT); fun mk_termtyp T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); fun mk_scomp T1 T2 sT f g = Const (\<^const_name>\scomp\, liftT T1 sT --> (T1 --> liftT T2 sT) --> liftT T2 sT) $ f $ g; fun mk_case_prod T = Sign.mk_const thy (\<^const_name>\case_prod\, [T, \<^typ>\unit \ term\, liftT resultT \<^typ>\Random.seed\]); fun mk_scomp_split T t t' = mk_scomp (mk_termtyp T) resultT \<^typ>\Random.seed\ t (mk_case_prod T $ Abs ("", T, Abs ("", \<^typ>\unit \ term\, t'))); fun mk_bindclause (_, _, i, T) = mk_scomp_split T (Sign.mk_const thy (\<^const_name>\Quickcheck_Random.random\, [T]) $ Bound i); in lambda genuine_only (Abs ("n", \<^typ>\natural\, fold_rev mk_bindclause bounds (return $ check true))) end val mk_parametric_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr ((mk_generator_expr, absdummy \<^typ>\bool\ (absdummy \<^typ>\natural\ \<^term>\Pair None :: Random.seed \ (bool \ term list) option \ Random.seed\)), \<^typ>\bool \ natural \ Random.seed \ (bool \ term list) option \ Random.seed\) val mk_parametric_reporting_generator_expr = Quickcheck_Common.gen_mk_parametric_generator_expr ((mk_reporting_generator_expr, absdummy \<^typ>\bool\ (absdummy \<^typ>\natural\ \<^term>\Pair (None, ([], False)) :: Random.seed \ ((bool \ term list) option \ (bool list \ bool)) \ Random.seed\)), \<^typ>\bool \ natural \ Random.seed \ ((bool \ term list) option \ (bool list \ bool)) \ Random.seed\) (* single quickcheck report *) datatype single_report = Run of bool list * bool | MatchExc fun collect_single_report single_report (Quickcheck.Report {iterations = iterations, raised_match_errors = raised_match_errors, satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests}) = case single_report of MatchExc => Quickcheck.Report {iterations = iterations + 1, raised_match_errors = raised_match_errors + 1, satisfied_assms = satisfied_assms, positive_concl_tests = positive_concl_tests} | Run (assms, concl) => Quickcheck.Report {iterations = iterations + 1, raised_match_errors = raised_match_errors, satisfied_assms = map2 (fn b => fn s => if b then s + 1 else s) assms (if null satisfied_assms then replicate (length assms) 0 else satisfied_assms), positive_concl_tests = if concl then positive_concl_tests + 1 else positive_concl_tests} val empty_report = Quickcheck.Report { iterations = 0, raised_match_errors = 0, satisfied_assms = [], positive_concl_tests = 0 } fun compile_generator_expr_raw ctxt ts = let val iterations = Config.get ctxt Quickcheck.iterations in if Config.get ctxt Quickcheck.report then let val t' = mk_parametric_reporting_generator_expr ctxt ts; val compile = Code_Runtime.dynamic_value_strict (get_counterexample_report, put_counterexample_report, "Random_Generators.put_counterexample_report") ctxt (SOME target) (fn proc => fn g => fn c => fn b => fn s => g c b s #>> (apfst o Option.map o apsnd o map) proc) t' []; fun single_tester c b s = compile c b s |> Random_Engine.run fun iterate_and_collect _ _ 0 report = (NONE, report) | iterate_and_collect genuine_only (card, size) j report = let val (test_result, single_report) = apsnd Run (single_tester card genuine_only size) val report = collect_single_report single_report report in case test_result of NONE => iterate_and_collect genuine_only (card, size) (j - 1) report | SOME q => (SOME q, report) end in fn genuine_only => fn [card, size] => apsnd SOME (iterate_and_collect genuine_only (card, size) iterations empty_report) end else let val t' = mk_parametric_generator_expr ctxt ts; val compile = Code_Runtime.dynamic_value_strict (get_counterexample, put_counterexample, "Random_Generators.put_counterexample") ctxt (SOME target) (fn proc => fn g => fn c => fn b => fn s => g c b s #>> (Option.map o apsnd o map) proc) t' []; fun single_tester c b s = compile c b s |> Random_Engine.run fun iterate _ _ 0 = NONE | iterate genuine_only (card, size) j = case single_tester card genuine_only size of NONE => iterate genuine_only (card, size) (j - 1) | SOME q => SOME q in fn genuine_only => fn [card, size] => (rpair NONE (iterate genuine_only (card, size) iterations)) end end; fun compile_generator_expr ctxt ts = let val compiled = compile_generator_expr_raw ctxt ts in fn genuine_only => fn [card, size] => compiled genuine_only [Code_Numeral.natural_of_integer card, Code_Numeral.natural_of_integer size] end; val size_types = [\<^type_name>\Enum.finite_1\, \<^type_name>\Enum.finite_2\, \<^type_name>\Enum.finite_3\, \<^type_name>\Enum.finite_4\, \<^type_name>\Enum.finite_5\]; fun size_matters_for _ Ts = not (forall (fn Type (tyco, []) => member (op =) size_types tyco | _ => false) Ts); val test_goals = Quickcheck_Common.generator_test_goal_terms ("random", (size_matters_for, compile_generator_expr)); (** setup **) val active = Attrib.setup_config_bool \<^binding>\quickcheck_random_active\ (K false); val _ = Theory.setup (Quickcheck_Common.datatype_interpretation \<^plugin>\quickcheck_random\ (\<^sort>\random\, instantiate_random_datatype) #> Context.theory_map (Quickcheck.add_tester ("random", (active, test_goals)))); end; diff --git a/src/HOL/Tools/Quotient/quotient_tacs.ML b/src/HOL/Tools/Quotient/quotient_tacs.ML --- a/src/HOL/Tools/Quotient/quotient_tacs.ML +++ b/src/HOL/Tools/Quotient/quotient_tacs.ML @@ -1,764 +1,765 @@ (* Title: HOL/Tools/Quotient/quotient_tacs.ML Author: Cezary Kaliszyk and Christian Urban Tactics for solving goal arising from lifting theorems to quotient types. *) signature QUOTIENT_TACS = sig val regularize_tac: Proof.context -> int -> tactic val injection_tac: Proof.context -> int -> tactic val all_injection_tac: Proof.context -> int -> tactic val clean_tac: Proof.context -> int -> tactic val descend_procedure_tac: Proof.context -> thm list -> int -> tactic val descend_tac: Proof.context -> thm list -> int -> tactic val partiality_descend_procedure_tac: Proof.context -> thm list -> int -> tactic val partiality_descend_tac: Proof.context -> thm list -> int -> tactic val lift_procedure_tac: Proof.context -> thm list -> thm -> int -> tactic val lift_tac: Proof.context -> thm list -> thm list -> int -> tactic val lifted: Proof.context -> typ list -> thm list -> thm -> thm val lifted_attrib: attribute end; structure Quotient_Tacs: QUOTIENT_TACS = struct (** various helper fuctions **) (* Since HOL_basic_ss is too "big" for us, we *) (* need to set up our own minimal simpset. *) fun mk_minimal_simpset ctxt = empty_simpset ctxt |> Simplifier.set_subgoaler asm_simp_tac |> Simplifier.set_mksimps (mksimps []) (* composition of two theorems, used in maps *) fun OF1 thm1 thm2 = thm2 RS thm1 fun atomize_thm ctxt thm = let val thm' = Thm.legacy_freezeT (Thm.forall_intr_vars thm) (* FIXME/TODO: is this proper Isar-technology? no! *) val thm'' = Object_Logic.atomize ctxt (Thm.cprop_of thm') in @{thm equal_elim_rule1} OF [thm'', thm'] end (*** Regularize Tactic ***) (** solvers for equivp and quotient assumptions **) fun equiv_tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\))) val equiv_solver = mk_solver "Equivalence goal solver" equiv_tac fun quotient_tac ctxt = (REPEAT_ALL_NEW (FIRST' [resolve_tac ctxt @{thms identity_quotient3}, resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_thm\))])) val quotient_solver = mk_solver "Quotient goal solver" quotient_tac fun solve_quotient_assm ctxt thm = case Seq.pull (quotient_tac ctxt 1 thm) of SOME (t, _) => t | _ => error "Solve_quotient_assm failed. Possibly a quotient theorem is missing." fun get_match_inst thy pat trm = let val univ = Unify.matchers (Context.Theory thy) [(pat, trm)] val SOME (env, _) = Seq.pull univ (* raises Bind, if no unifier *) (* FIXME fragile *) val tenv = Vartab.dest (Envir.term_env env) val tyenv = Vartab.dest (Envir.type_env env) fun prep_ty (x, (S, ty)) = ((x, S), Thm.global_ctyp_of thy ty) fun prep_trm (x, (T, t)) = ((x, T), Thm.global_cterm_of thy t) in - (map prep_ty tyenv, map prep_trm tenv) + (TVars.make (map prep_ty tyenv), Vars.make (map prep_trm tenv)) end (* Calculates the instantiations for the lemmas: ball_reg_eqv_range and bex_reg_eqv_range Since the left-hand-side contains a non-pattern '?P (f ?x)' we rely on unification/instantiation to check whether the theorem applies and return NONE if it doesn't. *) fun calculate_inst ctxt ball_bex_thm redex R1 R2 = let val thy = Proof_Context.theory_of ctxt fun get_lhs thm = fst (Logic.dest_equals (Thm.concl_of thm)) val ty_inst = map (SOME o Thm.ctyp_of ctxt) [domain_type (fastype_of R2)] val trm_inst = map (SOME o Thm.cterm_of ctxt) [R2, R1] in (case try (Thm.instantiate' ty_inst trm_inst) ball_bex_thm of NONE => NONE | SOME thm' => (case try (get_match_inst thy (get_lhs thm')) (Thm.term_of redex) of NONE => NONE | SOME inst2 => try (Drule.instantiate_normalize inst2) thm')) end fun ball_bex_range_simproc ctxt redex = (case Thm.term_of redex of (Const (\<^const_name>\Ball\, _) $ (Const (\<^const_name>\Respects\, _) $ (Const (\<^const_name>\rel_fun\, _) $ R1 $ R2)) $ _) => calculate_inst ctxt @{thm ball_reg_eqv_range[THEN eq_reflection]} redex R1 R2 | (Const (\<^const_name>\Bex\, _) $ (Const (\<^const_name>\Respects\, _) $ (Const (\<^const_name>\rel_fun\, _) $ R1 $ R2)) $ _) => calculate_inst ctxt @{thm bex_reg_eqv_range[THEN eq_reflection]} redex R1 R2 | _ => NONE) (* Regularize works as follows: 0. preliminary simplification step according to ball_reg_eqv bex_reg_eqv babs_reg_eqv ball_reg_eqv_range bex_reg_eqv_range 1. eliminating simple Ball/Bex instances (ball_reg_right bex_reg_left) 2. monos 3. commutation rules for ball and bex (ball_all_comm bex_ex_comm) 4. then rel-equalities, which need to be instantiated with 'eq_imp_rel' to avoid loops 5. then simplification like 0 finally jump back to 1 *) fun reflp_get ctxt = map_filter (fn th => if Thm.no_prems th then SOME (OF1 @{thm equivp_reflp} th) else NONE handle THM _ => NONE) (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\)) val eq_imp_rel = @{lemma "equivp R \ a = b \ R a b" by (simp add: equivp_reflp)} fun eq_imp_rel_get ctxt = map (OF1 eq_imp_rel) (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_equiv\)) val regularize_simproc = Simplifier.make_simproc \<^context> "regularize" {lhss = [\<^term>\Ball (Respects (R1 ===> R2)) P\, \<^term>\Bex (Respects (R1 ===> R2)) P\], proc = K ball_bex_range_simproc}; fun regularize_tac ctxt = let val simpset = mk_minimal_simpset ctxt addsimps @{thms ball_reg_eqv bex_reg_eqv babs_reg_eqv babs_simp} addsimprocs [regularize_simproc] addSolver equiv_solver addSolver quotient_solver val eq_eqvs = eq_imp_rel_get ctxt in simp_tac simpset THEN' TRY o REPEAT_ALL_NEW (CHANGED o FIRST' [resolve_tac ctxt @{thms ball_reg_right bex_reg_left bex1_bexeq_reg}, resolve_tac ctxt (Inductive.get_monos ctxt), resolve_tac ctxt @{thms ball_all_comm bex_ex_comm}, resolve_tac ctxt eq_eqvs, simp_tac simpset]) end (*** Injection Tactic ***) (* Looks for Quot_True assumptions, and in case its parameter is an application, it returns the function and the argument. *) fun find_qt_asm asms = let fun find_fun trm = (case trm of (Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\Quot_True\, _) $ _)) => true | _ => false) in (case find_first find_fun asms of SOME (_ $ (_ $ (f $ a))) => SOME (f, a) | _ => NONE) end fun quot_true_simple_conv ctxt fnctn ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\Quot_True\, _) $ x) => let val fx = fnctn x; val cx = Thm.cterm_of ctxt x; val cfx = Thm.cterm_of ctxt fx; val cxt = Thm.ctyp_of ctxt (fastype_of x); val cfxt = Thm.ctyp_of ctxt (fastype_of fx); val thm = Thm.instantiate' [SOME cxt, SOME cfxt] [SOME cx, SOME cfx] @{thm QT_imp} in Conv.rewr_conv thm ctrm end) fun quot_true_conv ctxt fnctn ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\Quot_True\, _) $ _) => quot_true_simple_conv ctxt fnctn ctrm | _ $ _ => Conv.comb_conv (quot_true_conv ctxt fnctn) ctrm | Abs _ => Conv.abs_conv (fn (_, ctxt) => quot_true_conv ctxt fnctn) ctxt ctrm | _ => Conv.all_conv ctrm) fun quot_true_tac ctxt fnctn = CONVERSION ((Conv.params_conv ~1 (fn ctxt => (Conv.prems_conv ~1 (quot_true_conv ctxt fnctn)))) ctxt) fun dest_comb (f $ a) = (f, a) fun dest_bcomb ((_ $ l) $ r) = (l, r) fun unlam t = (case t of Abs a => snd (Term.dest_abs a) | _ => unlam (Abs("", domain_type (fastype_of t), (incr_boundvars 1 t) $ (Bound 0)))) val bare_concl = HOLogic.dest_Trueprop o Logic.strip_assums_concl (* We apply apply_rsp only in case if the type needs lifting. This is the case if the type of the data in the Quot_True assumption is different from the corresponding type in the goal. *) val apply_rsp_tac = Subgoal.FOCUS (fn {concl, asms, context = ctxt,...} => let val bare_concl = HOLogic.dest_Trueprop (Thm.term_of concl) val qt_asm = find_qt_asm (map Thm.term_of asms) in case (bare_concl, qt_asm) of (R2 $ (f $ x) $ (g $ y), SOME (qt_fun, qt_arg)) => if fastype_of qt_fun = fastype_of f then no_tac else let val ty_x = fastype_of x val ty_b = fastype_of qt_arg val ty_f = range_type (fastype_of f) val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_x, ty_b, ty_f] val t_inst = map (SOME o Thm.cterm_of ctxt) [R2, f, g, x, y]; val inst_thm = Thm.instantiate' ty_inst ([NONE, NONE, NONE] @ t_inst) @{thm apply_rspQ3} in (resolve_tac ctxt [inst_thm] THEN' SOLVED' (quotient_tac ctxt)) 1 end | _ => no_tac end) (* Instantiates and applies 'equals_rsp'. Since the theorem is complex we rely on instantiation to tell us if it applies *) fun equals_rsp_tac R ctxt = case try (Thm.cterm_of ctxt) R of (* There can be loose bounds in R *) (* FIXME fragile *) SOME ctm => let val ty = domain_type (fastype_of R) in case try (Thm.instantiate' [SOME (Thm.ctyp_of ctxt ty)] [SOME (Thm.cterm_of ctxt R)]) @{thm equals_rsp} of SOME thm => resolve_tac ctxt [thm] THEN' quotient_tac ctxt | NONE => K no_tac end | _ => K no_tac fun rep_abs_rsp_tac ctxt = SUBGOAL (fn (goal, i) => (case try bare_concl goal of SOME (rel $ _ $ (rep $ (abs $ _))) => (let val (ty_a, ty_b) = dest_funT (fastype_of abs); val ty_inst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b]; in case try (map (SOME o Thm.cterm_of ctxt)) [rel, abs, rep] of SOME t_inst => (case try (Thm.instantiate' ty_inst t_inst) @{thm rep_abs_rsp} of SOME inst_thm => (resolve_tac ctxt [inst_thm] THEN' quotient_tac ctxt) i | NONE => no_tac) | NONE => no_tac end handle TERM _ => no_tac) | _ => no_tac)) (* Injection means to prove that the regularized theorem implies the abs/rep injected one. The deterministic part: - remove lambdas from both sides - prove Ball/Bex/Babs equalities using ball_rsp, bex_rsp, babs_rsp - prove Ball/Bex relations using rel_funI - reflexivity of equality - prove equality of relations using equals_rsp - use user-supplied RSP theorems - solve 'relation of relations' goals using quot_rel_rsp - remove rep_abs from the right side (Lambdas under respects may have left us some assumptions) Then in order: - split applications of lifted type (apply_rsp) - split applications of non-lifted type (cong_tac) - apply extentionality - assumption - reflexivity of the relation *) fun injection_match_tac ctxt = SUBGOAL (fn (goal, i) => (case bare_concl goal of (* (R1 ===> R2) (%x...) (%x...) ----> [|R1 x y|] ==> R2 (...x) (...y) *) (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Abs _) $ (Abs _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam (* (op =) (Ball...) (Ball...) ----> (op =) (...) (...) *) | (Const (\<^const_name>\HOL.eq\,_) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _)) => resolve_tac ctxt @{thms ball_rsp} THEN' dresolve_tac ctxt @{thms QT_all} (* (R1 ===> op =) (Ball...) (Ball...) ----> [|R1 x y|] ==> (Ball...x) = (Ball...y) *) | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Ball\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam (* (op =) (Bex...) (Bex...) ----> (op =) (...) (...) *) | Const (\<^const_name>\HOL.eq\,_) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms bex_rsp} THEN' dresolve_tac ctxt @{thms QT_ex} (* (R1 ===> op =) (Bex...) (Bex...) ----> [|R1 x y|] ==> (Bex...x) = (Bex...y) *) | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Bex\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) => resolve_tac ctxt @{thms rel_funI} THEN' quot_true_tac ctxt unlam | (Const (\<^const_name>\rel_fun\, _) $ _ $ _) $ (Const(\<^const_name>\Bex1_rel\,_) $ _) $ (Const(\<^const_name>\Bex1_rel\,_) $ _) => resolve_tac ctxt @{thms bex1_rel_rsp} THEN' quotient_tac ctxt | (_ $ (Const(\<^const_name>\Babs\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _) $ (Const(\<^const_name>\Babs\,_) $ (Const (\<^const_name>\Respects\, _) $ _) $ _)) => resolve_tac ctxt @{thms babs_rsp} THEN' quotient_tac ctxt | Const (\<^const_name>\HOL.eq\,_) $ (R $ _ $ _) $ (_ $ _ $ _) => (resolve_tac ctxt @{thms refl} ORELSE' (equals_rsp_tac R ctxt THEN' RANGE [ quot_true_tac ctxt (fst o dest_bcomb), quot_true_tac ctxt (snd o dest_bcomb)])) (* reflexivity of operators arising from Cong_tac *) | Const (\<^const_name>\HOL.eq\,_) $ _ $ _ => resolve_tac ctxt @{thms refl} (* respectfulness of constants; in particular of a simple relation *) | _ $ (Const _) $ (Const _) (* rel_fun, list_rel, etc but not equality *) => resolve_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\quot_respect\)) THEN_ALL_NEW quotient_tac ctxt (* R (...) (Rep (Abs ...)) ----> R (...) (...) *) (* observe map_fun *) | _ $ _ $ _ => (resolve_tac ctxt @{thms quot_rel_rsp} THEN_ALL_NEW quotient_tac ctxt) ORELSE' rep_abs_rsp_tac ctxt | _ => K no_tac) i) fun injection_step_tac ctxt rel_refl = FIRST' [ injection_match_tac ctxt, (* R (t $ ...) (t' $ ...) ----> apply_rsp provided type of t needs lifting *) apply_rsp_tac ctxt THEN' RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], (* (op =) (t $ ...) (t' $ ...) ----> Cong provided type of t does not need lifting *) (* merge with previous tactic *) Cong_Tac.cong_tac ctxt @{thm cong} THEN' RANGE [quot_true_tac ctxt (fst o dest_comb), quot_true_tac ctxt (snd o dest_comb)], (* resolving with R x y assumptions *) assume_tac ctxt, (* (op =) (%x...) (%y...) ----> (op =) (...) (...) *) resolve_tac ctxt @{thms ext} THEN' quot_true_tac ctxt unlam, (* reflexivity of the basic relations *) (* R ... ... *) resolve_tac ctxt rel_refl] fun injection_tac ctxt = let val rel_refl = reflp_get ctxt in injection_step_tac ctxt rel_refl end fun all_injection_tac ctxt = REPEAT_ALL_NEW (injection_tac ctxt) (*** Cleaning of the Theorem ***) (* expands all map_funs, except in front of the (bound) variables listed in xs *) fun map_fun_simple_conv xs ctrm = (case Thm.term_of ctrm of ((Const (\<^const_name>\map_fun\, _) $ _ $ _) $ h $ _) => if member (op=) xs h then Conv.all_conv ctrm else Conv.rewr_conv @{thm map_fun_apply [THEN eq_reflection]} ctrm | _ => Conv.all_conv ctrm) fun map_fun_conv xs ctxt ctrm = (case Thm.term_of ctrm of _ $ _ => (Conv.comb_conv (map_fun_conv xs ctxt) then_conv map_fun_simple_conv xs) ctrm | Abs _ => Conv.abs_conv (fn (x, ctxt) => map_fun_conv (Thm.term_of x :: xs) ctxt) ctxt ctrm | _ => Conv.all_conv ctrm) fun map_fun_tac ctxt = CONVERSION (map_fun_conv [] ctxt) (* custom matching functions *) fun mk_abs u i t = if incr_boundvars i u aconv t then Bound i else case t of t1 $ t2 => mk_abs u i t1 $ mk_abs u i t2 | Abs (s, T, t') => Abs (s, T, mk_abs u (i + 1) t') | Bound j => if i = j then error "make_inst" else t | _ => t fun make_inst lhs t = let val _ $ (Abs (_, _, (_ $ ((f as Var (_, Type ("fun", [T, _]))) $ u)))) = lhs; val _ $ (Abs (_, _, (_ $ g))) = t; in (f, Abs ("x", T, mk_abs u 0 g)) end fun make_inst_id lhs t = let val _ $ (Abs (_, _, (f as Var (_, Type ("fun", [T, _]))) $ u)) = lhs; val _ $ (Abs (_, _, g)) = t; in (f, Abs ("x", T, mk_abs u 0 g)) end (* Simplifies a redex using the 'lambda_prs' theorem. First instantiates the types and known subterms. Then solves the quotient assumptions to get Rep2 and Abs1 Finally instantiates the function f using make_inst If Rep2 is an identity then the pattern is simpler and make_inst_id is used *) fun lambda_prs_simple_conv ctxt ctrm = (case Thm.term_of ctrm of (Const (\<^const_name>\map_fun\, _) $ r1 $ a2) $ (Abs _) => let val (ty_b, ty_a) = dest_funT (fastype_of r1) val (ty_c, ty_d) = dest_funT (fastype_of a2) val tyinst = map (SOME o Thm.ctyp_of ctxt) [ty_a, ty_b, ty_c, ty_d] val tinst = [NONE, NONE, SOME (Thm.cterm_of ctxt r1), NONE, SOME (Thm.cterm_of ctxt a2)] val thm1 = Thm.instantiate' tyinst tinst @{thm lambda_prs[THEN eq_reflection]} val thm2 = solve_quotient_assm ctxt (solve_quotient_assm ctxt thm1) val thm3 = rewrite_rule ctxt @{thms id_apply[THEN eq_reflection]} thm2 val (insp, inst) = if ty_c = ty_d then make_inst_id (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) else make_inst (Thm.term_of (Thm.lhs_of thm3)) (Thm.term_of ctrm) val thm4 = - Drule.instantiate_normalize ([], [(dest_Var insp, Thm.cterm_of ctxt inst)]) thm3 + Drule.instantiate_normalize + (TVars.empty, Vars.make [(dest_Var insp, Thm.cterm_of ctxt inst)]) thm3 in Conv.rewr_conv thm4 ctrm end | _ => Conv.all_conv ctrm) fun lambda_prs_conv ctxt = Conv.top_conv lambda_prs_simple_conv ctxt fun lambda_prs_tac ctxt = CONVERSION (lambda_prs_conv ctxt) (* Cleaning consists of: 1. unfolding of ---> in front of everything, except bound variables (this prevents lambda_prs from becoming stuck) 2. simplification with lambda_prs 3. simplification with: - Quotient_abs_rep Quotient_rel_rep babs_prs all_prs ex_prs ex1_prs - id_simps and preservation lemmas and - symmetric versions of the definitions (that is definitions of quotient constants are folded) 4. test for refl *) fun clean_tac ctxt = let val thy = Proof_Context.theory_of ctxt val defs = map (Thm.symmetric o #def) (Quotient_Info.dest_quotconsts_global thy) val prs = rev (Named_Theorems.get ctxt \<^named_theorems>\quot_preserve\) val ids = rev (Named_Theorems.get ctxt \<^named_theorems>\id_simps\) val thms = @{thms Quotient3_abs_rep Quotient3_rel_rep babs_prs all_prs ex_prs ex1_prs} @ ids @ prs @ defs val simpset = (mk_minimal_simpset ctxt) addsimps thms addSolver quotient_solver in EVERY' [ map_fun_tac ctxt, lambda_prs_tac ctxt, simp_tac simpset, TRY o resolve_tac ctxt [refl]] end (* Tactic for Generalising Free Variables in a Goal *) fun inst_spec ctrm = Thm.instantiate' [SOME (Thm.ctyp_of_cterm ctrm)] [NONE, SOME ctrm] @{thm spec} fun inst_spec_tac ctxt ctrms = EVERY' (map (dresolve_tac ctxt o single o inst_spec) ctrms) fun all_list xs trm = fold (fn (x, T) => fn t' => HOLogic.mk_all (x, T, t')) xs trm fun apply_under_Trueprop f = HOLogic.dest_Trueprop #> f #> HOLogic.mk_Trueprop fun gen_frees_tac ctxt = SUBGOAL (fn (concl, i) => let val vrs = Term.add_frees concl [] val cvrs = map (Thm.cterm_of ctxt o Free) vrs val concl' = apply_under_Trueprop (all_list vrs) concl val goal = Logic.mk_implies (concl', concl) val rule = Goal.prove ctxt [] [] goal (K (EVERY1 [inst_spec_tac ctxt (rev cvrs), assume_tac ctxt])) in resolve_tac ctxt [rule] i end) (** The General Shape of the Lifting Procedure **) (* - A is the original raw theorem - B is the regularized theorem - C is the rep/abs injected version of B - D is the lifted theorem - 1st prem is the regularization step - 2nd prem is the rep/abs injection step - 3rd prem is the cleaning part the Quot_True premise in 2nd records the lifted theorem *) val procedure_thm = @{lemma "\A; A \ B; Quot_True D \ B = C; C = D\ \ D" by (simp add: Quot_True_def)} (* in case of partial equivalence relations, this form of the procedure theorem results in solvable proof obligations *) val partiality_procedure_thm = @{lemma "[|B; Quot_True D ==> B = C; C = D|] ==> D" by (simp add: Quot_True_def)} fun lift_match_error ctxt msg rtrm qtrm = let val rtrm_str = Syntax.string_of_term ctxt rtrm val qtrm_str = Syntax.string_of_term ctxt qtrm val msg = cat_lines [enclose "[" "]" msg, "The quotient theorem", qtrm_str, "", "does not match with original theorem", rtrm_str] in error msg end fun procedure_inst ctxt rtrm qtrm = let val rtrm' = HOLogic.dest_Trueprop rtrm val qtrm' = HOLogic.dest_Trueprop qtrm val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Thm.instantiate' [] [SOME (Thm.cterm_of ctxt rtrm'), SOME (Thm.cterm_of ctxt reg_goal), NONE, SOME (Thm.cterm_of ctxt inj_goal)] procedure_thm end (* Since we use Ball and Bex during the lifting and descending, we cannot deal with lemmas containing them, unless we unfold them by default. *) val default_unfolds = @{thms Ball_def Bex_def} (** descending as tactic **) fun descend_procedure_tac ctxt simps = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal val rule = procedure_inst ctxt rtrm goal in resolve_tac ctxt [rule] i end) end fun descend_tac ctxt simps = let val mk_tac_raw = descend_procedure_tac ctxt simps THEN' RANGE [Object_Logic.rulify_tac ctxt THEN' (K all_tac), regularize_tac ctxt, all_injection_tac ctxt, clean_tac ctxt] in Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw end (** descending for partial equivalence relations **) fun partiality_procedure_inst ctxt rtrm qtrm = let val rtrm' = HOLogic.dest_Trueprop rtrm val qtrm' = HOLogic.dest_Trueprop qtrm val reg_goal = Quotient_Term.regularize_trm_chk ctxt (rtrm', qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm val inj_goal = Quotient_Term.inj_repabs_trm_chk ctxt (reg_goal, qtrm') handle Quotient_Term.LIFT_MATCH msg => lift_match_error ctxt msg rtrm qtrm in Thm.instantiate' [] [SOME (Thm.cterm_of ctxt reg_goal), NONE, SOME (Thm.cterm_of ctxt inj_goal)] partiality_procedure_thm end fun partiality_descend_procedure_tac ctxt simps = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) val rtrm = Quotient_Term.derive_rtrm ctxt qtys goal val rule = partiality_procedure_inst ctxt rtrm goal in resolve_tac ctxt [rule] i end) end fun partiality_descend_tac ctxt simps = let val mk_tac_raw = partiality_descend_procedure_tac ctxt simps THEN' RANGE [Object_Logic.rulify_tac ctxt THEN' (K all_tac), all_injection_tac ctxt, clean_tac ctxt] in Goal.conjunction_tac THEN_ALL_NEW mk_tac_raw end (** lifting as a tactic **) (* the tactic leaves three subgoals to be proved *) fun lift_procedure_tac ctxt simps rthm = let val simpset = (mk_minimal_simpset ctxt) addsimps (simps @ default_unfolds) in full_simp_tac simpset THEN' Object_Logic.full_atomize_tac ctxt THEN' gen_frees_tac ctxt THEN' SUBGOAL (fn (goal, i) => let (* full_atomize_tac contracts eta redexes, so we do it also in the original theorem *) val rthm' = rthm |> full_simplify simpset |> Drule.eta_contraction_rule |> Thm.forall_intr_frees |> atomize_thm ctxt val rule = procedure_inst ctxt (Thm.prop_of rthm') goal in (resolve_tac ctxt [rule] THEN' resolve_tac ctxt [rthm']) i end) end fun lift_single_tac ctxt simps rthm = lift_procedure_tac ctxt simps rthm THEN' RANGE [ regularize_tac ctxt, all_injection_tac ctxt, clean_tac ctxt ] fun lift_tac ctxt simps rthms = Goal.conjunction_tac THEN' RANGE (map (lift_single_tac ctxt simps) rthms) (* automated lifting with pre-simplification of the theorems; for internal usage *) fun lifted ctxt qtys simps rthm = let val ((_, [rthm']), ctxt') = Variable.import true [rthm] ctxt val goal = Quotient_Term.derive_qtrm ctxt' qtys (Thm.prop_of rthm') in Goal.prove ctxt' [] [] goal (K (HEADGOAL (lift_single_tac ctxt' simps rthm'))) |> singleton (Proof_Context.export ctxt' ctxt) end (* lifting as an attribute *) val lifted_attrib = Thm.rule_attribute [] (fn context => let val ctxt = Context.proof_of context val qtys = map #qtyp (Quotient_Info.dest_quotients ctxt) in lifted ctxt qtys [] end) end; (* structure *) diff --git a/src/HOL/Tools/SMT/smt_normalize.ML b/src/HOL/Tools/SMT/smt_normalize.ML --- a/src/HOL/Tools/SMT/smt_normalize.ML +++ b/src/HOL/Tools/SMT/smt_normalize.ML @@ -1,535 +1,535 @@ (* Title: HOL/Tools/SMT/smt_normalize.ML Author: Sascha Boehme, TU Muenchen Normalization steps on theorems required by SMT solvers. *) signature SMT_NORMALIZE = sig val drop_fact_warning: Proof.context -> thm -> unit val atomize_conv: Proof.context -> conv val special_quant_table: (string * thm) list val case_bool_entry: string * thm val abs_min_max_table: (string * thm) list type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list val add_extra_norm: SMT_Util.class * extra_norm -> Context.generic -> Context.generic val normalize: Proof.context -> thm list -> (int * thm) list end; structure SMT_Normalize: SMT_NORMALIZE = struct fun drop_fact_warning ctxt = SMT_Config.verbose_msg ctxt (prefix "Warning: dropping assumption: " o Thm.string_of_thm ctxt) (* general theorem normalizations *) (** instantiate elimination rules **) local val (cpfalse, cfalse) = `SMT_Util.mk_cprop (Thm.cterm_of \<^context> \<^const>\False\) fun inst f ct thm = let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm)) - in Thm.instantiate ([], [(dest_Var (Thm.term_of cv), ct)]) thm end + in Thm.instantiate (TVars.empty, Vars.make [(dest_Var (Thm.term_of cv), ct)]) thm end in fun instantiate_elim thm = (case Thm.concl_of thm of \<^const>\Trueprop\ $ Var (_, \<^typ>\bool\) => inst Thm.dest_arg cfalse thm | Var _ => inst I cpfalse thm | _ => thm) end (** normalize definitions **) fun norm_def thm = (case Thm.prop_of thm of \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ _ $ Abs _) => norm_def (thm RS @{thm fun_cong}) | Const (\<^const_name>\Pure.eq\, _) $ _ $ Abs _ => norm_def (HOLogic.mk_obj_eq thm) | _ => thm) (** atomization **) fun atomize_conv ctxt ct = (case Thm.term_of ct of \<^const>\Pure.imp\ $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_imp} | Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => Conv.binop_conv (atomize_conv ctxt) then_conv Conv.rewr_conv @{thm atomize_eq} | Const (\<^const_name>\Pure.all\, _) $ Abs _ => Conv.binder_conv (atomize_conv o snd) ctxt then_conv Conv.rewr_conv @{thm atomize_all} | _ => Conv.all_conv) ct handle CTERM _ => Conv.all_conv ct val setup_atomize = fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\Pure.imp\, \<^const_name>\Pure.eq\, \<^const_name>\Pure.all\, \<^const_name>\Trueprop\] (** unfold special quantifiers **) val special_quant_table = [ (\<^const_name>\Ex1\, @{thm Ex1_def_raw}), (\<^const_name>\Ball\, @{thm Ball_def_raw}), (\<^const_name>\Bex\, @{thm Bex_def_raw})] local fun special_quant (Const (n, _)) = AList.lookup (op =) special_quant_table n | special_quant _ = NONE fun special_quant_conv _ ct = (case special_quant (Thm.term_of ct) of SOME thm => Conv.rewr_conv thm | NONE => Conv.all_conv) ct in fun unfold_special_quants_conv ctxt = SMT_Util.if_exists_conv (is_some o special_quant) (Conv.top_conv special_quant_conv ctxt) val setup_unfolded_quants = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) special_quant_table end (** trigger inference **) local (*** check trigger syntax ***) fun dest_trigger (Const (\<^const_name>\pat\, _) $ _) = SOME true | dest_trigger (Const (\<^const_name>\nopat\, _) $ _) = SOME false | dest_trigger _ = NONE fun eq_list [] = false | eq_list (b :: bs) = forall (equal b) bs fun proper_trigger t = t |> these o try SMT_Util.dest_symb_list |> map (map_filter dest_trigger o these o try SMT_Util.dest_symb_list) |> (fn [] => false | bss => forall eq_list bss) fun proper_quant inside f t = (case t of Const (\<^const_name>\All\, _) $ Abs (_, _, u) => proper_quant true f u | Const (\<^const_name>\Ex\, _) $ Abs (_, _, u) => proper_quant true f u | \<^const>\trigger\ $ p $ u => (if inside then f p else false) andalso proper_quant false f u | Abs (_, _, u) => proper_quant false f u | u1 $ u2 => proper_quant false f u1 andalso proper_quant false f u2 | _ => true) fun check_trigger_error ctxt t = error ("SMT triggers must only occur under quantifier and multipatterns " ^ "must have the same kind: " ^ Syntax.string_of_term ctxt t) fun check_trigger_conv ctxt ct = if proper_quant false proper_trigger (SMT_Util.term_of ct) then Conv.all_conv ct else check_trigger_error ctxt (Thm.term_of ct) (*** infer simple triggers ***) fun dest_cond_eq ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Thm.dest_binop ct | \<^const>\HOL.implies\ $ _ $ _ => dest_cond_eq (Thm.dest_arg ct) | _ => raise CTERM ("no equation", [ct])) fun get_constrs thy (Type (n, _)) = these (BNF_LFP_Compat.get_constrs thy n) | get_constrs _ _ = [] fun is_constr thy (n, T) = let fun match (m, U) = m = n andalso Sign.typ_instance thy (T, U) in can (the o find_first match o get_constrs thy o Term.body_type) T end fun is_constr_pat thy t = (case Term.strip_comb t of (Free _, []) => true | (Const c, ts) => is_constr thy c andalso forall (is_constr_pat thy) ts | _ => false) fun is_simp_lhs ctxt t = (case Term.strip_comb t of (Const c, ts as _ :: _) => not (SMT_Builtin.is_builtin_fun_ext ctxt c ts) andalso forall (is_constr_pat (Proof_Context.theory_of ctxt)) ts | _ => false) fun has_all_vars vs t = subset (op aconv) (vs, map Free (Term.add_frees t [])) fun minimal_pats vs ct = if has_all_vars vs (Thm.term_of ct) then (case Thm.term_of ct of _ $ _ => (case apply2 (minimal_pats vs) (Thm.dest_comb ct) of ([], []) => [[ct]] | (ctss, ctss') => union (eq_set (op aconvc)) ctss ctss') | _ => []) else [] fun proper_mpat _ _ _ [] = false | proper_mpat thy gen u cts = let val tps = (op ~~) (`gen (map Thm.term_of cts)) fun some_match u = tps |> exists (fn (t', t) => Pattern.matches thy (t', u) andalso not (t aconv u)) in not (Term.exists_subterm some_match u) end val pat = SMT_Util.mk_const_pat \<^theory> \<^const_name>\pat\ Thm.dest_ctyp0 fun mk_pat ct = Thm.apply (SMT_Util.instT' ct pat) ct fun mk_clist T = apply2 (Thm.cterm_of \<^context>) (SMT_Util.symb_cons_const T, SMT_Util.symb_nil_const T) fun mk_list (ccons, cnil) f cts = fold_rev (Thm.mk_binop ccons o f) cts cnil val mk_pat_list = mk_list (mk_clist \<^typ>\pattern\) val mk_mpat_list = mk_list (mk_clist \<^typ>\pattern symb_list\) fun mk_trigger ctss = mk_mpat_list (mk_pat_list mk_pat) ctss val trigger_eq = mk_meta_eq @{lemma "p = trigger t p" by (simp add: trigger_def)} fun insert_trigger_conv [] ct = Conv.all_conv ct | insert_trigger_conv ctss ct = let val (ctr, cp) = Thm.dest_binop (Thm.rhs_of trigger_eq) ||> rpair ct val inst = map (apfst (dest_Var o Thm.term_of)) [cp, (ctr, mk_trigger ctss)] - in Thm.instantiate ([], inst) trigger_eq end + in Thm.instantiate (TVars.empty, Vars.make inst) trigger_eq end fun infer_trigger_eq_conv outer_ctxt (ctxt, cvs) ct = let val (lhs, rhs) = dest_cond_eq ct val vs = map Thm.term_of cvs val thy = Proof_Context.theory_of ctxt fun get_mpats ct = if is_simp_lhs ctxt (Thm.term_of ct) then minimal_pats vs ct else [] val gen = Variable.export_terms ctxt outer_ctxt val filter_mpats = filter (proper_mpat thy gen (Thm.term_of rhs)) in insert_trigger_conv (filter_mpats (get_mpats lhs)) ct end fun has_trigger (\<^const>\trigger\ $ _ $ _) = true | has_trigger _ = false fun try_trigger_conv cv ct = if SMT_Util.under_quant has_trigger (SMT_Util.term_of ct) then Conv.all_conv ct else Conv.try_conv cv ct fun infer_trigger_conv ctxt = if Config.get ctxt SMT_Config.infer_triggers then try_trigger_conv (SMT_Util.under_quant_conv (infer_trigger_eq_conv ctxt) ctxt) else Conv.all_conv in fun trigger_conv ctxt = SMT_Util.prop_conv (check_trigger_conv ctxt then_conv infer_trigger_conv ctxt) val setup_trigger = fold SMT_Builtin.add_builtin_fun_ext'' [\<^const_name>\pat\, \<^const_name>\nopat\, \<^const_name>\trigger\] end (** combined general normalizations **) fun gen_normalize1_conv ctxt = atomize_conv ctxt then_conv unfold_special_quants_conv ctxt then_conv Thm.beta_conversion true then_conv trigger_conv ctxt fun gen_normalize1 ctxt = instantiate_elim #> norm_def #> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion) #> Thm.forall_intr_vars #> Conv.fconv_rule (gen_normalize1_conv ctxt) #> (* Z3 4.3.1 silently normalizes "P --> Q --> R" to "P & Q --> R" *) Raw_Simplifier.rewrite_rule ctxt @{thms HOL.imp_conjL[symmetric, THEN eq_reflection]} fun gen_norm1_safe ctxt (i, thm) = (case try (gen_normalize1 ctxt) thm of SOME thm' => SOME (i, thm') | NONE => (drop_fact_warning ctxt thm; NONE)) fun gen_normalize ctxt iwthms = map_filter (gen_norm1_safe ctxt) iwthms (* unfolding of definitions and theory-specific rewritings *) fun expand_head_conv cv ct = (case Thm.term_of ct of _ $ _ => Conv.fun_conv (expand_head_conv cv) then_conv Conv.try_conv (Thm.beta_conversion false) | _ => cv) ct (** rewrite bool case expressions as if expressions **) val case_bool_entry = (\<^const_name>\bool.case_bool\, @{thm case_bool_if}) local fun is_case_bool (Const (\<^const_name>\bool.case_bool\, _)) = true | is_case_bool _ = false fun unfold_conv _ = SMT_Util.if_true_conv (is_case_bool o Term.head_of) (expand_head_conv (Conv.rewr_conv @{thm case_bool_if})) in fun rewrite_case_bool_conv ctxt = SMT_Util.if_exists_conv is_case_bool (Conv.top_conv unfold_conv ctxt) val setup_case_bool = SMT_Builtin.add_builtin_fun_ext'' \<^const_name>\bool.case_bool\ end (** unfold abs, min and max **) val abs_min_max_table = [ (\<^const_name>\min\, @{thm min_def_raw}), (\<^const_name>\max\, @{thm max_def_raw}), (\<^const_name>\abs\, @{thm abs_if_raw})] local fun abs_min_max ctxt (Const (n, Type (\<^type_name>\fun\, [T, _]))) = (case AList.lookup (op =) abs_min_max_table n of NONE => NONE | SOME thm => if SMT_Builtin.is_builtin_typ_ext ctxt T then SOME thm else NONE) | abs_min_max _ _ = NONE fun unfold_amm_conv ctxt ct = (case abs_min_max ctxt (Term.head_of (Thm.term_of ct)) of SOME thm => expand_head_conv (Conv.rewr_conv thm) | NONE => Conv.all_conv) ct in fun unfold_abs_min_max_conv ctxt = SMT_Util.if_exists_conv (is_some o abs_min_max ctxt) (Conv.top_conv unfold_amm_conv ctxt) val setup_abs_min_max = fold (SMT_Builtin.add_builtin_fun_ext'' o fst) abs_min_max_table end (** embedding of standard natural number operations into integer operations **) local val simple_nat_ops = [ @{const HOL.eq (nat)}, @{const less (nat)}, @{const less_eq (nat)}, \<^const>\Suc\, @{const plus (nat)}, @{const minus (nat)}] val nat_consts = simple_nat_ops @ [@{const numeral (nat)}, @{const zero_class.zero (nat)}, @{const one_class.one (nat)}] @ [@{const times (nat)}, @{const divide (nat)}, @{const modulo (nat)}] val is_nat_const = member (op aconv) nat_consts val nat_int_thm = Thm.symmetric (mk_meta_eq @{thm nat_int}) val nat_int_comp_thms = map mk_meta_eq @{thms nat_int_comparison} val int_ops_thms = map mk_meta_eq @{thms int_ops} val int_if_thm = mk_meta_eq @{thm int_if} fun if_conv cv1 cv2 = Conv.combination_conv (Conv.combination_conv (Conv.arg_conv cv1) cv2) cv2 fun int_ops_conv cv ctxt ct = (case Thm.term_of ct of @{const of_nat (int)} $ (Const (\<^const_name>\If\, _) $ _ $ _ $ _) => Conv.rewr_conv int_if_thm then_conv if_conv (cv ctxt) (int_ops_conv cv ctxt) | @{const of_nat (int)} $ _ => (Conv.rewrs_conv int_ops_thms then_conv Conv.top_sweep_conv (int_ops_conv cv) ctxt) else_conv Conv.arg_conv (Conv.sub_conv cv ctxt) | _ => Conv.no_conv) ct val unfold_nat_let_conv = Conv.rewr_conv @{lemma "Let (n::nat) f \ f n" by (rule Let_def)} val drop_nat_int_conv = Conv.rewr_conv (Thm.symmetric nat_int_thm) fun nat_to_int_conv ctxt ct = ( Conv.top_conv (K (Conv.try_conv unfold_nat_let_conv)) ctxt then_conv Conv.top_sweep_conv nat_conv ctxt then_conv Conv.top_conv (K (Conv.try_conv drop_nat_int_conv)) ctxt) ct and nat_conv ctxt ct = ( Conv.rewrs_conv (nat_int_thm :: nat_int_comp_thms) then_conv Conv.top_sweep_conv (int_ops_conv nat_to_int_conv) ctxt) ct fun add_int_of_nat vs ct cu (q, cts) = (case Thm.term_of ct of @{const of_nat(int)} => if Term.exists_subterm (member (op aconv) vs) (Thm.term_of cu) then (true, cts) else (q, insert (op aconvc) cu cts) | _ => (q, cts)) fun add_apps f vs ct = (case Thm.term_of ct of _ $ _ => let val (cu1, cu2) = Thm.dest_comb ct in f vs cu1 cu2 #> add_apps f vs cu1 #> add_apps f vs cu2 end | Abs _ => let val (cv, cu) = Thm.dest_abs NONE ct in add_apps f (Thm.term_of cv :: vs) cu end | _ => I) val int_thm = @{lemma "(0::int) <= int (n::nat)" by simp} val nat_int_thms = @{lemma "\n::nat. (0::int) <= int n" "\n::nat. nat (int n) = n" "\i::int. int (nat i) = (if 0 <= i then i else 0)" by simp_all} val var = Term.dest_Var (Thm.term_of (funpow 3 Thm.dest_arg (Thm.cprop_of int_thm))) in fun nat_as_int_conv ctxt = SMT_Util.if_exists_conv is_nat_const (nat_to_int_conv ctxt) fun add_int_of_nat_constraints thms = let val (q, cts) = fold (add_apps add_int_of_nat [] o Thm.cprop_of) thms (false, []) in if q then (thms, nat_int_thms) - else (thms, map (fn ct => Thm.instantiate ([], [(var, ct)]) int_thm) cts) + else (thms, map (fn ct => Thm.instantiate (TVars.empty, Vars.make [(var, ct)]) int_thm) cts) end val setup_nat_as_int = SMT_Builtin.add_builtin_typ_ext (\<^typ>\nat\, fn ctxt => K (Config.get ctxt SMT_Config.nat_as_int)) #> fold (SMT_Builtin.add_builtin_fun_ext' o Term.dest_Const) simple_nat_ops end (** normalize numerals **) local (* rewrite Numeral1 into 1 rewrite - 0 into 0 *) fun is_irregular_number (Const (\<^const_name>\numeral\, _) $ Const (\<^const_name>\num.One\, _)) = true | is_irregular_number (Const (\<^const_name>\uminus\, _) $ Const (\<^const_name>\Groups.zero\, _)) = true | is_irregular_number _ = false fun is_strange_number ctxt t = is_irregular_number t andalso SMT_Builtin.is_builtin_num ctxt t val proper_num_ss = simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms Num.numeral_One minus_zero}) fun norm_num_conv ctxt = SMT_Util.if_conv (is_strange_number ctxt) (Simplifier.rewrite (put_simpset proper_num_ss ctxt)) Conv.no_conv in fun normalize_numerals_conv ctxt = SMT_Util.if_exists_conv (is_strange_number ctxt) (Conv.top_sweep_conv norm_num_conv ctxt) end (** combined unfoldings and rewritings **) fun burrow_ids f ithms = let val (is, thms) = split_list ithms val (thms', extra_thms) = f thms in (is ~~ thms') @ map (pair ~1) extra_thms end fun unfold_conv ctxt = rewrite_case_bool_conv ctxt then_conv unfold_abs_min_max_conv ctxt then_conv (if Config.get ctxt SMT_Config.nat_as_int then nat_as_int_conv ctxt else Conv.all_conv) then_conv Thm.beta_conversion true fun unfold_polymorph ctxt = map (apsnd (Conv.fconv_rule (unfold_conv ctxt))) fun unfold_monomorph ctxt = map (apsnd (Conv.fconv_rule (normalize_numerals_conv ctxt))) #> Config.get ctxt SMT_Config.nat_as_int ? burrow_ids add_int_of_nat_constraints (* overall normalization *) type extra_norm = Proof.context -> thm list * thm list -> thm list * thm list structure Extra_Norms = Generic_Data ( type T = extra_norm SMT_Util.dict val empty = [] val extend = I fun merge data = SMT_Util.dict_merge fst data ) fun add_extra_norm (cs, norm) = Extra_Norms.map (SMT_Util.dict_update (cs, norm)) fun apply_extra_norms ctxt ithms = let val cs = SMT_Config.solver_class_of ctxt val es = SMT_Util.dict_lookup (Extra_Norms.get (Context.Proof ctxt)) cs in burrow_ids (fold (fn e => e ctxt) es o rpair []) ithms end local val ignored = member (op =) [\<^const_name>\All\, \<^const_name>\Ex\, \<^const_name>\Let\, \<^const_name>\If\, \<^const_name>\HOL.eq\] val schematic_consts_of = let fun collect (\<^const>\trigger\ $ p $ t) = collect_trigger p #> collect t | collect (t $ u) = collect t #> collect u | collect (Abs (_, _, t)) = collect t | collect (t as Const (n, _)) = if not (ignored n) then Monomorph.add_schematic_consts_of t else I | collect _ = I and collect_trigger t = let val dest = these o try SMT_Util.dest_symb_list in fold (fold collect_pat o dest) (dest t) end and collect_pat (Const (\<^const_name>\pat\, _) $ t) = collect t | collect_pat (Const (\<^const_name>\nopat\, _) $ t) = collect t | collect_pat _ = I in (fn t => collect t Symtab.empty) end in fun monomorph ctxt xthms = let val (xs, thms) = split_list xthms in map (pair 1) thms |> Monomorph.monomorph schematic_consts_of ctxt |> maps (uncurry (map o pair)) o map2 pair xs o map (map snd) end end fun normalize ctxt wthms = wthms |> map_index I |> gen_normalize ctxt |> unfold_polymorph ctxt |> monomorph ctxt |> unfold_monomorph ctxt |> apply_extra_norms ctxt val _ = Theory.setup (Context.theory_map ( setup_atomize #> setup_unfolded_quants #> setup_trigger #> setup_case_bool #> setup_abs_min_max #> setup_nat_as_int)) end; diff --git a/src/HOL/Tools/SMT/smt_replay.ML b/src/HOL/Tools/SMT/smt_replay.ML --- a/src/HOL/Tools/SMT/smt_replay.ML +++ b/src/HOL/Tools/SMT/smt_replay.ML @@ -1,296 +1,297 @@ (* Title: HOL/Tools/SMT/smt_replay.ML Author: Sascha Boehme, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Mathias Fleury, MPII Shared library for parsing and replay. *) signature SMT_REPLAY = sig (*theorem nets*) val thm_net_of: ('a -> thm) -> 'a list -> 'a Net.net val net_instances: (int * thm) Net.net -> cterm -> (int * thm) list (*proof combinators*) val under_assumption: (thm -> thm) -> cterm -> thm val discharge: thm -> thm -> thm (*a faster COMP*) type compose_data = cterm list * (cterm -> cterm list) * thm val precompose: (cterm -> cterm list) -> thm -> compose_data val precompose2: (cterm -> cterm * cterm) -> thm -> compose_data val compose: compose_data -> thm -> thm (*simpset*) val add_simproc: Simplifier.simproc -> Context.generic -> Context.generic val make_simpset: Proof.context -> thm list -> simpset (*assertion*) val add_asserted: ('a * ('b * thm) -> 'c -> 'c) -> 'c -> ('d -> 'a * 'e * term * 'b) -> ('e -> bool) -> Proof.context -> thm list -> (int * thm) list -> 'd list -> Proof.context -> ((int * ('a * thm)) list * thm list) * (Proof.context * 'c) (*statistics*) val pretty_statistics: string -> int -> int list Symtab.table -> Pretty.T val intermediate_statistics: Proof.context -> Timing.start -> int -> int -> unit (*theorem transformation*) val varify: Proof.context -> thm -> thm val params_of: term -> (string * typ) list (*spy*) val spying: bool -> Proof.context -> (unit -> string) -> string -> unit val print_stats: (string * int list) list -> string end; structure SMT_Replay : SMT_REPLAY = struct (* theorem nets *) fun thm_net_of f xthms = let fun insert xthm = Net.insert_term (K false) (Thm.prop_of (f xthm), xthm) in fold insert xthms Net.empty end fun maybe_instantiate ct thm = try Thm.first_order_match (Thm.cprop_of thm, ct) |> Option.map (fn inst => Thm.instantiate inst thm) local fun instances_from_net match f net ct = let val lookup = if match then Net.match_term else Net.unify_term val xthms = lookup net (Thm.term_of ct) fun select ct = map_filter (f (maybe_instantiate ct)) xthms fun select' ct = let val thm = Thm.trivial ct in map_filter (f (try (fn rule => rule COMP thm))) xthms end in (case select ct of [] => select' ct | xthms' => xthms') end in fun net_instances net = instances_from_net false (fn f => fn (i, thm) => Option.map (pair i) (f thm)) net end (* proof combinators *) fun under_assumption f ct = let val ct' = SMT_Util.mk_cprop ct in Thm.implies_intr ct' (f (Thm.assume ct')) end fun discharge p pq = Thm.implies_elim pq p (* a faster COMP *) type compose_data = cterm list * (cterm -> cterm list) * thm fun list2 (x, y) = [x, y] fun precompose f rule : compose_data = (f (Thm.cprem_of rule 1), f, rule) fun precompose2 f rule : compose_data = precompose (list2 o f) rule fun compose (cvs, f, rule) thm = discharge thm - (Thm.instantiate ([], map (dest_Var o Thm.term_of) cvs ~~ f (Thm.cprop_of thm)) rule) + (Thm.instantiate + (TVars.empty, Vars.make (map (dest_Var o Thm.term_of) cvs ~~ f (Thm.cprop_of thm))) rule) (* simpset *) local val antisym_le1 = mk_meta_eq @{thm order_class.antisym_conv} val antisym_le2 = mk_meta_eq @{thm order_class.antisym_conv2} val antisym_less1 = mk_meta_eq @{thm order_class.antisym_conv1} val antisym_less2 = mk_meta_eq @{thm linorder_class.antisym_conv3} fun eq_prop t thm = HOLogic.mk_Trueprop t aconv Thm.prop_of thm fun dest_binop ((c as Const _) $ t $ u) = (c, t, u) | dest_binop t = raise TERM ("dest_binop", [t]) fun prove_antisym_le ctxt ct = let val (le, r, s) = dest_binop (Thm.term_of ct) val less = Const (\<^const_name>\less\, Term.fastype_of le) val prems = Simplifier.prems_of ctxt in (case find_first (eq_prop (le $ s $ r)) prems of NONE => find_first (eq_prop (HOLogic.mk_not (less $ r $ s))) prems |> Option.map (fn thm => thm RS antisym_less1) | SOME thm => SOME (thm RS antisym_le1)) end handle THM _ => NONE fun prove_antisym_less ctxt ct = let val (less, r, s) = dest_binop (HOLogic.dest_not (Thm.term_of ct)) val le = Const (\<^const_name>\less_eq\, Term.fastype_of less) val prems = Simplifier.prems_of ctxt in (case find_first (eq_prop (le $ r $ s)) prems of NONE => find_first (eq_prop (HOLogic.mk_not (less $ s $ r))) prems |> Option.map (fn thm => thm RS antisym_less2) | SOME thm => SOME (thm RS antisym_le2)) end handle THM _ => NONE val basic_simpset = simpset_of (put_simpset HOL_ss \<^context> addsimps @{thms field_simps times_divide_eq_right times_divide_eq_left arith_special arith_simps rel_simps array_rules z3div_def z3mod_def NO_MATCH_def} addsimprocs [\<^simproc>\numeral_divmod\, Simplifier.make_simproc \<^context> "fast_int_arith" {lhss = [\<^term>\(m::int) < n\, \<^term>\(m::int) \ n\, \<^term>\(m::int) = n\], proc = K Lin_Arith.simproc}, Simplifier.make_simproc \<^context> "antisym_le" {lhss = [\<^term>\(x::'a::order) \ y\], proc = K prove_antisym_le}, Simplifier.make_simproc \<^context> "antisym_less" {lhss = [\<^term>\\ (x::'a::linorder) < y\], proc = K prove_antisym_less}]) structure Simpset = Generic_Data ( type T = simpset val empty = basic_simpset val extend = I val merge = Simplifier.merge_ss ) in fun add_simproc simproc context = Simpset.map (simpset_map (Context.proof_of context) (fn ctxt => ctxt addsimprocs [simproc])) context fun make_simpset ctxt rules = simpset_of (put_simpset (Simpset.get (Context.Proof ctxt)) ctxt addsimps rules) end local val remove_trigger = mk_meta_eq @{thm trigger_def} val remove_fun_app = mk_meta_eq @{thm fun_app_def} fun rewrite_conv _ [] = Conv.all_conv | rewrite_conv ctxt eqs = Simplifier.full_rewrite (empty_simpset ctxt addsimps eqs) val rewrite_true_rule = @{lemma "True \ \ False" by simp} val prep_rules = [@{thm Let_def}, remove_trigger, remove_fun_app, rewrite_true_rule] fun rewrite _ [] = I | rewrite ctxt eqs = Conv.fconv_rule (rewrite_conv ctxt eqs) fun lookup_assm assms_net ct = net_instances assms_net ct |> map (fn ithm as (_, thm) => (ithm, Thm.cprop_of thm aconvc ct)) in fun add_asserted tab_update tab_empty p_extract cond outer_ctxt rewrite_rules assms steps ctxt0 = let val eqs = map (rewrite ctxt0 [rewrite_true_rule]) rewrite_rules val eqs' = union Thm.eq_thm eqs prep_rules val assms_net = assms |> map (apsnd (rewrite ctxt0 eqs')) |> map (apsnd (Conv.fconv_rule Thm.eta_conversion)) |> thm_net_of snd fun revert_conv ctxt = rewrite_conv ctxt eqs' then_conv Thm.eta_conversion fun assume thm ctxt = let val ct = Thm.cprem_of thm 1 val (thm', ctxt') = yield_singleton Assumption.add_assumes ct ctxt in (thm' RS thm, ctxt') end fun add1 id fixes thm1 ((i, th), exact) ((iidths, thms), (ctxt, ptab)) = let val (thm, ctxt') = if exact then (Thm.implies_elim thm1 th, ctxt) else assume thm1 ctxt val thms' = if exact then thms else th :: thms in (((i, (id, th)) :: iidths, thms'), (ctxt', tab_update (id, (fixes, thm)) ptab)) end fun add step (cx as ((iidths, thms), (ctxt, ptab))) = let val (id, rule, concl, fixes) = p_extract step in if (*Z3_Proof.is_assumption rule andalso rule <> Z3_Proof.Hypothesis*) cond rule then let val ct = Thm.cterm_of ctxt concl val thm1 = Thm.trivial ct |> Conv.fconv_rule (Conv.arg1_conv (revert_conv outer_ctxt)) val thm2 = singleton (Variable.export ctxt outer_ctxt) thm1 in (case lookup_assm assms_net (Thm.cprem_of thm2 1) of [] => let val (thm, ctxt') = assume thm1 ctxt in ((iidths, thms), (ctxt', tab_update (id, (fixes, thm)) ptab)) end | ithms => fold (add1 id fixes thm1) ithms cx) end else cx end in fold add steps (([], []), (ctxt0, tab_empty)) end end fun params_of t = Term.strip_qnt_vars \<^const_name>\Pure.all\ t fun varify ctxt thm = let val maxidx = Thm.maxidx_of thm + 1 val vs = params_of (Thm.prop_of thm) val vars = map_index (fn (i, (n, T)) => Var ((n, i + maxidx), T)) vs in Drule.forall_elim_list (map (Thm.cterm_of ctxt) vars) thm end fun intermediate_statistics ctxt start total = SMT_Config.statistics_msg ctxt (fn current => "Reconstructed " ^ string_of_int current ^ " of " ^ string_of_int total ^ " steps in " ^ string_of_int (Time.toMilliseconds (#elapsed (Timing.result start))) ^ " ms") fun pretty_statistics solver total stats = let val stats = Symtab.map (K (map (fn i => curry Int.div i 1000000))) stats fun mean_of is = let val len = length is val mid = len div 2 in if len mod 2 = 0 then (nth is (mid - 1) + nth is mid) div 2 else nth is mid end fun pretty_item name p = Pretty.item (Pretty.separate ":" [Pretty.str name, p]) fun pretty (name, milliseconds) = (Pretty.block (Pretty.str (name ^": ") :: Pretty.separate "," [ Pretty.str (string_of_int (length milliseconds) ^ " occurrences") , Pretty.str (string_of_int (mean_of milliseconds) ^ " ms mean time"), Pretty.str (string_of_int (fold Integer.max milliseconds 0) ^ " ms maximum time"), Pretty.str (string_of_int (fold Integer.add milliseconds 0) ^ " ms total time")])) in Pretty.big_list (solver ^ " proof reconstruction statistics:") ( pretty_item "total time" (Pretty.str (string_of_int total ^ " ms")) :: map pretty (Symtab.dest stats)) end fun timestamp_format time = Date.fmt "%Y-%m-%d %H:%M:%S." (Date.fromTimeLocal time) ^ (StringCvt.padLeft #"0" 3 (string_of_int (Time.toMilliseconds time - 1000 * Time.toSeconds time))) fun print_stats stats = let fun print_list xs = fold (fn x => fn msg => msg ^ string_of_int x ^ ",") xs "" in fold (fn (x,y) => fn msg => msg ^ x ^ ": " ^ print_list y ^ "\n") stats "" end fun spying false _ _ _ = () | spying true ctxt f filename = let val message = f () val thy = Context.theory_long_name ((Context.theory_of o Context.Proof) ctxt) val spying_version = "1" in File.append (Path.explode ("$ISABELLE_HOME_USER/" ^ filename)) (spying_version ^ "; " ^ thy ^ "; " ^ (timestamp_format (Time.now ())) ^ ";\n" ^ message ^ "\n") end end; diff --git a/src/HOL/Tools/SMT/smt_replay_methods.ML b/src/HOL/Tools/SMT/smt_replay_methods.ML --- a/src/HOL/Tools/SMT/smt_replay_methods.ML +++ b/src/HOL/Tools/SMT/smt_replay_methods.ML @@ -1,518 +1,518 @@ (* Title: HOL/Tools/SMT/smt_replay_methods.ML Author: Sascha Boehme, TU Muenchen Author: Jasmin Blanchette, TU Muenchen Author: Mathias Fleury, MPII Proof methods for replaying SMT proofs. *) signature SMT_REPLAY_METHODS = sig (*Printing*) val pretty_goal: Proof.context -> string -> string -> thm list -> term -> Pretty.T val trace_goal: Proof.context -> string -> thm list -> term -> unit val trace: Proof.context -> (unit -> string) -> unit (*Replaying*) val replay_error: Proof.context -> string -> string -> thm list -> term -> 'a val replay_rule_error: string -> Proof.context -> string -> thm list -> term -> 'a (*theory lemma methods*) type th_lemma_method = Proof.context -> thm list -> term -> thm val add_th_lemma_method: string * th_lemma_method -> Context.generic -> Context.generic val get_th_lemma_method: Proof.context -> th_lemma_method Symtab.table val discharge: int -> thm list -> thm -> thm val match_instantiate: Proof.context -> term -> thm -> thm val prove: Proof.context -> term -> (Proof.context -> int -> tactic) -> thm val prove_arith_rewrite: ((term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> Proof.context -> term -> thm (*abstraction*) type abs_context = int * term Termtab.table type 'a abstracter = term -> abs_context -> 'a * abs_context val add_arith_abstracter: (term abstracter -> term option abstracter) -> Context.generic -> Context.generic val abstract_lit: term -> abs_context -> term * abs_context val abstract_conj: term -> abs_context -> term * abs_context val abstract_disj: term -> abs_context -> term * abs_context val abstract_not: (term -> abs_context -> term * abs_context) -> term -> abs_context -> term * abs_context val abstract_unit: term -> abs_context -> term * abs_context val abstract_bool: term -> abs_context -> term * abs_context (* also abstracts over equivalences and conjunction and implication*) val abstract_bool_shallow: term -> abs_context -> term * abs_context (* abstracts down to equivalence symbols *) val abstract_bool_shallow_equivalence: term -> abs_context -> term * abs_context val abstract_prop: term -> abs_context -> term * abs_context val abstract_term: term -> abs_context -> term * abs_context val abstract_eq: (term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> term -> int * term Termtab.table -> term * (int * term Termtab.table) val abstract_neq: (term -> int * term Termtab.table -> term * (int * term Termtab.table)) -> term -> int * term Termtab.table -> term * (int * term Termtab.table) val abstract_arith: Proof.context -> term -> abs_context -> term * abs_context (* also abstracts over if-then-else *) val abstract_arith_shallow: Proof.context -> term -> abs_context -> term * abs_context val prove_abstract: Proof.context -> thm list -> term -> (Proof.context -> thm list -> int -> tactic) -> (abs_context -> (term list * term) * abs_context) -> thm val prove_abstract': Proof.context -> term -> (Proof.context -> thm list -> int -> tactic) -> (abs_context -> term * abs_context) -> thm val try_provers: string -> Proof.context -> string -> (string * (term -> 'a)) list -> thm list -> term -> 'a (*shared tactics*) val cong_unfolding_trivial: Proof.context -> thm list -> term -> thm val cong_basic: Proof.context -> thm list -> term -> thm val cong_full: Proof.context -> thm list -> term -> thm val cong_unfolding_first: Proof.context -> thm list -> term -> thm val arith_th_lemma: Proof.context -> thm list -> term -> thm val arith_th_lemma_wo: Proof.context -> thm list -> term -> thm val arith_th_lemma_wo_shallow: Proof.context -> thm list -> term -> thm val arith_th_lemma_tac_with_wo: Proof.context -> thm list -> int -> tactic val dest_thm: thm -> term val prop_tac: Proof.context -> thm list -> int -> tactic val certify_prop: Proof.context -> term -> cterm val dest_prop: term -> term end; structure SMT_Replay_Methods: SMT_REPLAY_METHODS = struct (* utility functions *) fun trace ctxt f = SMT_Config.trace_msg ctxt f () fun pretty_thm ctxt thm = Syntax.pretty_term ctxt (Thm.concl_of thm) fun pretty_goal ctxt msg rule thms t = let val full_msg = msg ^ ": " ^ quote rule val assms = if null thms then [] else [Pretty.big_list "assumptions:" (map (pretty_thm ctxt) thms)] val concl = Pretty.big_list "proposition:" [Syntax.pretty_term ctxt t] in Pretty.big_list full_msg (assms @ [concl]) end fun replay_error ctxt msg rule thms t = error (Pretty.string_of (pretty_goal ctxt msg rule thms t)) fun replay_rule_error name ctxt = replay_error ctxt ("Failed to replay" ^ name ^ " proof step") fun trace_goal ctxt rule thms t = trace ctxt (fn () => Pretty.string_of (pretty_goal ctxt "Goal" rule thms t)) fun as_prop (t as Const (\<^const_name>\Trueprop\, _) $ _) = t | as_prop t = HOLogic.mk_Trueprop t fun dest_prop (Const (\<^const_name>\Trueprop\, _) $ t) = t | dest_prop t = t fun dest_thm thm = dest_prop (Thm.concl_of thm) (* plug-ins *) type abs_context = int * term Termtab.table type 'a abstracter = term -> abs_context -> 'a * abs_context type th_lemma_method = Proof.context -> thm list -> term -> thm fun id_ord ((id1, _), (id2, _)) = int_ord (id1, id2) structure Plugins = Generic_Data ( type T = (int * (term abstracter -> term option abstracter)) list * th_lemma_method Symtab.table val empty = ([], Symtab.empty) val extend = I fun merge ((abss1, ths1), (abss2, ths2)) = ( Ord_List.merge id_ord (abss1, abss2), Symtab.merge (K true) (ths1, ths2)) ) fun add_arith_abstracter abs = Plugins.map (apfst (Ord_List.insert id_ord (serial (), abs))) fun get_arith_abstracters ctxt = map snd (fst (Plugins.get (Context.Proof ctxt))) fun add_th_lemma_method method = Plugins.map (apsnd (Symtab.update_new method)) fun get_th_lemma_method ctxt = snd (Plugins.get (Context.Proof ctxt)) fun match ctxt pat t = (Vartab.empty, Vartab.empty) |> Pattern.first_order_match (Proof_Context.theory_of ctxt) (pat, t) fun gen_certify_inst sel cert ctxt thm t = let val inst = match ctxt (dest_thm thm) (dest_prop t) fun cert_inst (ix, (a, b)) = ((ix, a), cert b) in Vartab.fold (cons o cert_inst) (sel inst) [] end fun match_instantiateT ctxt t thm = if Term.exists_type (Term.exists_subtype Term.is_TVar) (dest_thm thm) then - Thm.instantiate (gen_certify_inst fst (Thm.ctyp_of ctxt) ctxt thm t, []) thm + Thm.instantiate (TVars.make (gen_certify_inst fst (Thm.ctyp_of ctxt) ctxt thm t), Vars.empty) thm else thm fun match_instantiate ctxt t thm = let val thm' = match_instantiateT ctxt t thm in - Thm.instantiate ([], gen_certify_inst snd (Thm.cterm_of ctxt) ctxt thm' t) thm' + Thm.instantiate (TVars.empty, Vars.make (gen_certify_inst snd (Thm.cterm_of ctxt) ctxt thm' t)) thm' end fun discharge _ [] thm = thm | discharge i (rule :: rules) thm = discharge (i + Thm.nprems_of rule) rules (rule RSN (i, thm)) fun by_tac ctxt thms ns ts t tac = Goal.prove ctxt [] (map as_prop ts) (as_prop t) (fn {context, prems} => HEADGOAL (tac context prems)) |> Drule.generalize (Names.empty, Names.make_set ns) |> discharge 1 thms fun prove ctxt t tac = by_tac ctxt [] [] [] t (K o tac) (* abstraction *) fun prove_abstract ctxt thms t tac f = let val ((prems, concl), (_, ts)) = f (1, Termtab.empty) val ns = Termtab.fold (fn (_, v) => cons (fst (Term.dest_Free v))) ts [] in by_tac ctxt [] ns prems concl tac |> match_instantiate ctxt t |> discharge 1 thms end fun prove_abstract' ctxt t tac f = prove_abstract ctxt [] t tac (f #>> pair []) fun lookup_term (_, terms) t = Termtab.lookup terms t fun abstract_sub t f cx = (case lookup_term cx t of SOME v => (v, cx) | NONE => f cx) fun mk_fresh_free t (i, terms) = let val v = Free ("t" ^ string_of_int i, fastype_of t) in (v, (i + 1, Termtab.update (t, v) terms)) end fun apply_abstracters _ [] _ cx = (NONE, cx) | apply_abstracters abs (abstracter :: abstracters) t cx = (case abstracter abs t cx of (NONE, _) => apply_abstracters abs abstracters t cx | x as (SOME _, _) => x) fun abstract_term (t as _ $ _) = abstract_sub t (mk_fresh_free t) | abstract_term (t as Abs _) = abstract_sub t (mk_fresh_free t) | abstract_term t = pair t fun abstract_bin abs f t t1 t2 = abstract_sub t (abs t1 ##>> abs t2 #>> f) fun abstract_ter abs f t t1 t2 t3 = abstract_sub t (abs t1 ##>> abs t2 ##>> abs t3 #>> (Scan.triple1 #> f)) fun abstract_lit (\<^const>\HOL.Not\ $ t) = abstract_term t #>> HOLogic.mk_not | abstract_lit t = abstract_term t fun abstract_not abs (t as \<^const>\HOL.Not\ $ t1) = abstract_sub t (abs t1 #>> HOLogic.mk_not) | abstract_not _ t = abstract_lit t fun abstract_conj (t as \<^const>\HOL.conj\ $ t1 $ t2) = abstract_bin abstract_conj HOLogic.mk_conj t t1 t2 | abstract_conj t = abstract_lit t fun abstract_disj (t as \<^const>\HOL.disj\ $ t1 $ t2) = abstract_bin abstract_disj HOLogic.mk_disj t t1 t2 | abstract_disj t = abstract_lit t fun abstract_prop (t as (c as @{const If (bool)}) $ t1 $ t2 $ t3) = abstract_ter abstract_prop (fn (t1, t2, t3) => c $ t1 $ t2 $ t3) t t1 t2 t3 | abstract_prop (t as \<^const>\HOL.disj\ $ t1 $ t2) = abstract_bin abstract_prop HOLogic.mk_disj t t1 t2 | abstract_prop (t as \<^const>\HOL.conj\ $ t1 $ t2) = abstract_bin abstract_prop HOLogic.mk_conj t t1 t2 | abstract_prop (t as \<^const>\HOL.implies\ $ t1 $ t2) = abstract_bin abstract_prop HOLogic.mk_imp t t1 t2 | abstract_prop (t as \<^term>\HOL.eq :: bool => _\ $ t1 $ t2) = abstract_bin abstract_prop HOLogic.mk_eq t t1 t2 | abstract_prop t = abstract_not abstract_prop t fun abstract_arith ctxt u = let fun abs (t as (Const (\<^const_name>\HOL.The\, _) $ Abs (_, _, _))) = abstract_sub t (abstract_term t) | abs (t as (c as Const _) $ Abs (s, T, t')) = abstract_sub t (abs t' #>> (fn u' => c $ Abs (s, T, u'))) | abs (t as (c as Const (\<^const_name>\If\, _)) $ t1 $ t2 $ t3) = abstract_ter abs (fn (t1, t2, t3) => c $ t1 $ t2 $ t3) t t1 t2 t3 | abs (t as \<^const>\HOL.Not\ $ t1) = abstract_sub t (abs t1 #>> HOLogic.mk_not) | abs (t as \<^const>\HOL.disj\ $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> HOLogic.mk_disj) | abs (t as (c as Const (\<^const_name>\uminus_class.uminus\, _)) $ t1) = abstract_sub t (abs t1 #>> (fn u => c $ u)) | abs (t as (c as Const (\<^const_name>\plus_class.plus\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\minus_class.minus\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\times_class.times\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\z3div\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\z3mod\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\HOL.eq\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\ord_class.less\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\ord_class.less_eq\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs t = abstract_sub t (fn cx => if can HOLogic.dest_number t then (t, cx) else (case apply_abstracters abs (get_arith_abstracters ctxt) t cx of (SOME u, cx') => (u, cx') | (NONE, _) => abstract_term t cx)) in abs u end fun abstract_unit (t as (\<^const>\HOL.Not\ $ (\<^const>\HOL.disj\ $ t1 $ t2))) = abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_not o HOLogic.mk_disj) | abstract_unit (t as (\<^const>\HOL.disj\ $ t1 $ t2)) = abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_disj) | abstract_unit (t as (Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2)) = if fastype_of t1 = \<^typ>\bool\ then abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_eq) else abstract_lit t | abstract_unit (t as (\<^const>\HOL.Not\ $ (Const(\<^const_name>\HOL.eq\, _) $ t1 $ t2))) = if fastype_of t1 = \<^typ>\bool\ then abstract_sub t (abstract_unit t1 ##>> abstract_unit t2 #>> HOLogic.mk_eq #>> HOLogic.mk_not) else abstract_lit t | abstract_unit (t as (\<^const>\HOL.Not\ $ t1)) = abstract_sub t (abstract_unit t1 #>> HOLogic.mk_not) | abstract_unit t = abstract_lit t fun abstract_bool (t as (@{const HOL.disj} $ t1 $ t2)) = abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> HOLogic.mk_disj) | abstract_bool (t as (@{const HOL.conj} $ t1 $ t2)) = abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> HOLogic.mk_conj) | abstract_bool (t as (Const(@{const_name HOL.eq}, _) $ t1 $ t2)) = if fastype_of t1 = @{typ bool} then abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> HOLogic.mk_eq) else abstract_lit t | abstract_bool (t as (@{const HOL.Not} $ t1)) = abstract_sub t (abstract_bool t1 #>> HOLogic.mk_not) | abstract_bool (t as (@{const HOL.implies} $ t1 $ t2)) = abstract_sub t (abstract_bool t1 ##>> abstract_bool t2 #>> HOLogic.mk_imp) | abstract_bool t = abstract_lit t fun abstract_bool_shallow (t as (@{const HOL.disj} $ t1 $ t2)) = abstract_sub t (abstract_bool_shallow t1 ##>> abstract_bool_shallow t2 #>> HOLogic.mk_disj) | abstract_bool_shallow (t as (@{const HOL.Not} $ t1)) = abstract_sub t (abstract_bool_shallow t1 #>> HOLogic.mk_not) | abstract_bool_shallow t = abstract_term t fun abstract_bool_shallow_equivalence (t as (@{const HOL.disj} $ t1 $ t2)) = abstract_sub t (abstract_bool_shallow_equivalence t1 ##>> abstract_bool_shallow_equivalence t2 #>> HOLogic.mk_disj) | abstract_bool_shallow_equivalence (t as (Const(@{const_name HOL.eq}, _) $ t1 $ t2)) = if fastype_of t1 = @{typ bool} then abstract_sub t (abstract_lit t1 ##>> abstract_lit t2 #>> HOLogic.mk_eq) else abstract_lit t | abstract_bool_shallow_equivalence (t as (@{const HOL.Not} $ t1)) = abstract_sub t (abstract_bool_shallow_equivalence t1 #>> HOLogic.mk_not) | abstract_bool_shallow_equivalence t = abstract_lit t fun abstract_arith_shallow ctxt u = let fun abs (t as (Const (\<^const_name>\HOL.The\, _) $ Abs (_, _, _))) = abstract_sub t (abstract_term t) | abs (t as (c as Const _) $ Abs (s, T, t')) = abstract_sub t (abs t' #>> (fn u' => c $ Abs (s, T, u'))) | abs (t as (Const (\<^const_name>\If\, _)) $ _ $ _ $ _) = abstract_sub t (abstract_term t) | abs (t as \<^const>\HOL.Not\ $ t1) = abstract_sub t (abs t1 #>> HOLogic.mk_not) | abs (t as \<^const>\HOL.disj\ $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> HOLogic.mk_disj) | abs (t as (c as Const (\<^const_name>\uminus_class.uminus\, _)) $ t1) = abstract_sub t (abs t1 #>> (fn u => c $ u)) | abs (t as (c as Const (\<^const_name>\plus_class.plus\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\minus_class.minus\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\times_class.times\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (Const (\<^const_name>\z3div\, _)) $ _ $ _) = abstract_sub t (abstract_term t) | abs (t as (Const (\<^const_name>\z3mod\, _)) $ _ $ _) = abstract_sub t (abstract_term t) | abs (t as (c as Const (\<^const_name>\HOL.eq\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\ord_class.less\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs (t as (c as Const (\<^const_name>\ord_class.less_eq\, _)) $ t1 $ t2) = abstract_sub t (abs t1 ##>> abs t2 #>> (fn (u1, u2) => c $ u1 $ u2)) | abs t = abstract_sub t (fn cx => if can HOLogic.dest_number t then (t, cx) else (case apply_abstracters abs (get_arith_abstracters ctxt) t cx of (SOME u, cx') => (u, cx') | (NONE, _) => abstract_term t cx)) in abs u end (* theory lemmas *) fun try_provers prover_name ctxt rule [] thms t = replay_rule_error prover_name ctxt rule thms t | try_provers prover_name ctxt rule ((name, prover) :: named_provers) thms t = (case (trace ctxt (K ("Trying prover " ^ quote name)); try prover t) of SOME thm => thm | NONE => try_provers prover_name ctxt rule named_provers thms t) (*theory-lemma*) fun arith_th_lemma_tac ctxt prems = Method.insert_tac ctxt prems THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms z3div_def z3mod_def}) THEN' Arith_Data.arith_tac ctxt fun arith_th_lemma ctxt thms t = prove_abstract ctxt thms t arith_th_lemma_tac ( fold_map (abstract_arith ctxt o dest_thm) thms ##>> abstract_arith ctxt (dest_prop t)) fun arith_th_lemma_tac_with_wo ctxt prems = Method.insert_tac ctxt prems THEN' SELECT_GOAL (Local_Defs.unfold0_tac ctxt @{thms z3div_def z3mod_def int_distrib}) THEN' Simplifier.asm_full_simp_tac (empty_simpset ctxt addsimprocs [(*@{simproc linordered_ring_strict_eq_cancel_factor_poly}, @{simproc linordered_ring_strict_le_cancel_factor_poly}, @{simproc linordered_ring_strict_less_cancel_factor_poly}, @{simproc nat_eq_cancel_factor_poly}, @{simproc nat_le_cancel_factor_poly}, @{simproc nat_less_cancel_factor_poly}*)]) THEN' (fn i => TRY (Arith_Data.arith_tac ctxt i)) fun arith_th_lemma_wo ctxt thms t = prove_abstract ctxt thms t arith_th_lemma_tac_with_wo ( fold_map (abstract_arith ctxt o dest_thm) thms ##>> abstract_arith ctxt (dest_prop t)) fun arith_th_lemma_wo_shallow ctxt thms t = prove_abstract ctxt thms t arith_th_lemma_tac_with_wo ( fold_map (abstract_arith_shallow ctxt o dest_thm) thms ##>> abstract_arith_shallow ctxt (dest_prop t)) val _ = Theory.setup (Context.theory_map (add_th_lemma_method ("arith", arith_th_lemma))) (* congruence *) fun certify_prop ctxt t = Thm.cterm_of ctxt (as_prop t) fun ctac ctxt prems i st = st |> ( resolve_tac ctxt (@{thm refl} :: prems) i ORELSE (cong_tac ctxt i THEN ctac ctxt prems (i + 1) THEN ctac ctxt prems i)) fun cong_basic ctxt thms t = let val st = Thm.trivial (certify_prop ctxt t) in (case Seq.pull (ctac ctxt thms 1 st) of SOME (thm, _) => thm | NONE => raise THM ("cong", 0, thms @ [st])) end val cong_dest_rules = @{lemma "(\ P \ Q) \ (P \ \ Q) \ P = Q" "(P \ \ Q) \ (\ P \ Q) \ P = Q" by fast+} fun cong_full_core_tac ctxt = eresolve_tac ctxt @{thms subst} THEN' resolve_tac ctxt @{thms refl} ORELSE' Classical.fast_tac ctxt fun cong_full ctxt thms t = prove ctxt t (fn ctxt' => Method.insert_tac ctxt thms THEN' (cong_full_core_tac ctxt' ORELSE' dresolve_tac ctxt cong_dest_rules THEN' cong_full_core_tac ctxt')) local val reorder_for_simp = try (fn thm => let val t = Thm.prop_of (@{thm eq_reflection} OF [thm]) val thm = (case Logic.dest_equals t of (t1, t2) => if t1 aconv t2 then raise TERM("identical terms", [t1, t2]) else if Term.size_of_term t1 > Term.size_of_term t2 then @{thm eq_reflection} OF [thm] else @{thm eq_reflection} OF [@{thm sym} OF [thm]]) handle TERM("dest_equals", _) => @{thm eq_reflection} OF [thm] in thm end) in fun cong_unfolding_trivial ctxt thms t = prove ctxt t (fn _ => EVERY' (map (fn thm => K (unfold_tac ctxt [thm])) ((map_filter reorder_for_simp thms))) THEN' (fn i => TRY (resolve_tac ctxt @{thms refl} i))) fun cong_unfolding_first ctxt thms t = let val ctxt = ctxt |> empty_simpset |> put_simpset HOL_basic_ss |> (fn ctxt => ctxt addsimps @{thms not_not eq_commute}) in prove ctxt t (fn _ => EVERY' (map (fn thm => K (unfold_tac ctxt [thm])) ((map_filter reorder_for_simp thms))) THEN' Method.insert_tac ctxt thms THEN' (full_simp_tac ctxt) THEN' K (ALLGOALS (K (Clasimp.auto_tac ctxt)))) end end fun arith_rewrite_tac ctxt _ = let val backup_tac = Arith_Data.arith_tac ctxt ORELSE' Clasimp.force_tac ctxt in (TRY o Simplifier.simp_tac ctxt) THEN_ALL_NEW backup_tac ORELSE' backup_tac end fun abstract_eq f (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = f t1 ##>> f t2 #>> HOLogic.mk_eq | abstract_eq _ t = abstract_term t fun abstract_neq f (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2) = f t1 ##>> f t2 #>> HOLogic.mk_eq | abstract_neq f (Const (\<^const_name>\HOL.Not\, _) $ (Const (\<^const_name>\HOL.eq\, _) $ t1 $ t2)) = f t1 ##>> f t2 #>> HOLogic.mk_eq #>> curry (op $) HOLogic.Not | abstract_neq f (Const (\<^const_name>\HOL.disj\, _) $ t1 $ t2) = f t1 ##>> f t2 #>> HOLogic.mk_disj | abstract_neq _ t = abstract_term t fun prove_arith_rewrite abstracter ctxt t = prove_abstract' ctxt t arith_rewrite_tac ( abstracter (abstract_arith ctxt) (dest_prop t)) fun prop_tac ctxt prems = Method.insert_tac ctxt prems THEN' SUBGOAL (fn (prop, i) => if Term.size_of_term prop > 100 then SAT.satx_tac ctxt i else (Classical.fast_tac ctxt ORELSE' Clasimp.force_tac ctxt) i) end; diff --git a/src/HOL/Tools/SMT/smt_util.ML b/src/HOL/Tools/SMT/smt_util.ML --- a/src/HOL/Tools/SMT/smt_util.ML +++ b/src/HOL/Tools/SMT/smt_util.ML @@ -1,242 +1,243 @@ (* Title: HOL/Tools/SMT/smt_util.ML Author: Sascha Boehme, TU Muenchen General utility functions. *) signature SMT_UTIL = sig (*basic combinators*) val repeat: ('a -> 'a option) -> 'a -> 'a val repeat_yield: ('a -> 'b -> ('a * 'b) option) -> 'a -> 'b -> 'a * 'b datatype order = First_Order | Higher_Order (*class dictionaries*) type class = string list val basicC: class val string_of_class: class -> string type 'a dict = (class * 'a) Ord_List.T val dict_map_default: class * 'a -> ('a -> 'a) -> 'a dict -> 'a dict val dict_update: class * 'a -> 'a dict -> 'a dict val dict_merge: ('a * 'a -> 'a) -> 'a dict * 'a dict -> 'a dict val dict_lookup: 'a dict -> class -> 'a list val dict_get: 'a dict -> class -> 'a option (*types*) val dest_funT: int -> typ -> typ list * typ (*terms*) val dest_conj: term -> term * term val dest_disj: term -> term * term val under_quant: (term -> 'a) -> term -> 'a val is_number: term -> bool (*symbolic lists*) val symb_nil_const: typ -> term val symb_cons_const: typ -> term val mk_symb_list: typ -> term list -> term val dest_symb_list: term -> term list (*patterns and instantiations*) val mk_const_pat: theory -> string -> (ctyp -> 'a) -> 'a * cterm val instTs: ctyp list -> ctyp list * cterm -> cterm val instT: ctyp -> ctyp * cterm -> cterm val instT': cterm -> ctyp * cterm -> cterm (*certified terms*) val dest_cabs: cterm -> Proof.context -> cterm * Proof.context val dest_all_cabs: cterm -> Proof.context -> cterm * Proof.context val dest_cbinder: cterm -> Proof.context -> cterm * Proof.context val dest_all_cbinders: cterm -> Proof.context -> cterm * Proof.context val mk_cprop: cterm -> cterm val dest_cprop: cterm -> cterm val mk_cequals: cterm -> cterm -> cterm val term_of: cterm -> term val prop_of: thm -> term (*conversions*) val if_conv: (term -> bool) -> conv -> conv -> conv val if_true_conv: (term -> bool) -> conv -> conv val if_exists_conv: (term -> bool) -> conv -> conv val binders_conv: (Proof.context -> conv) -> Proof.context -> conv val under_quant_conv: (Proof.context * cterm list -> conv) -> Proof.context -> conv val prop_conv: conv -> conv end; structure SMT_Util: SMT_UTIL = struct (* basic combinators *) fun repeat f = let fun rep x = (case f x of SOME y => rep y | NONE => x) in rep end fun repeat_yield f = let fun rep x y = (case f x y of SOME (x', y') => rep x' y' | NONE => (x, y)) in rep end (* order *) datatype order = First_Order | Higher_Order (* class dictionaries *) type class = string list val basicC = [] fun string_of_class [] = "basic" | string_of_class cs = "basic." ^ space_implode "." cs type 'a dict = (class * 'a) Ord_List.T fun class_ord ((cs1, _), (cs2, _)) = rev_order (list_ord fast_string_ord (cs1, cs2)) fun dict_insert (cs, x) d = if AList.defined (op =) d cs then d else Ord_List.insert class_ord (cs, x) d fun dict_map_default (cs, x) f = dict_insert (cs, x) #> AList.map_entry (op =) cs f fun dict_update (e as (_, x)) = dict_map_default e (K x) fun dict_merge val_merge = sort class_ord o AList.join (op =) (K val_merge) fun dict_lookup d cs = let fun match (cs', x) = if is_prefix (op =) cs' cs then SOME x else NONE in map_filter match d end fun dict_get d cs = (case AList.lookup (op =) d cs of NONE => (case cs of [] => NONE | _ => dict_get d (take (length cs - 1) cs)) | SOME x => SOME x) (* types *) val dest_funT = let fun dest Ts 0 T = (rev Ts, T) | dest Ts i (Type ("fun", [T, U])) = dest (T::Ts) (i-1) U | dest _ _ T = raise TYPE ("not a function type", [T], []) in dest [] end (* terms *) fun dest_conj (\<^const>\HOL.conj\ $ t $ u) = (t, u) | dest_conj t = raise TERM ("not a conjunction", [t]) fun dest_disj (\<^const>\HOL.disj\ $ t $ u) = (t, u) | dest_disj t = raise TERM ("not a disjunction", [t]) fun under_quant f t = (case t of Const (\<^const_name>\All\, _) $ Abs (_, _, u) => under_quant f u | Const (\<^const_name>\Ex\, _) $ Abs (_, _, u) => under_quant f u | _ => f t) val is_number = let fun is_num env (Const (\<^const_name>\Let\, _) $ t $ Abs (_, _, u)) = is_num (t :: env) u | is_num env (Bound i) = i < length env andalso is_num env (nth env i) | is_num _ t = can HOLogic.dest_number t in is_num [] end (* symbolic lists *) fun symb_listT T = Type (\<^type_name>\symb_list\, [T]) fun symb_nil_const T = Const (\<^const_name>\Symb_Nil\, symb_listT T) fun symb_cons_const T = let val listT = symb_listT T in Const (\<^const_name>\Symb_Cons\, T --> listT --> listT) end fun mk_symb_list T ts = fold_rev (fn t => fn u => symb_cons_const T $ t $ u) ts (symb_nil_const T) fun dest_symb_list (Const (\<^const_name>\Symb_Nil\, _)) = [] | dest_symb_list (Const (\<^const_name>\Symb_Cons\, _) $ t $ u) = t :: dest_symb_list u (* patterns and instantiations *) fun mk_const_pat thy name destT = let val cpat = Thm.global_cterm_of thy (Const (name, Sign.the_const_type thy name)) in (destT (Thm.ctyp_of_cterm cpat), cpat) end -fun instTs cUs (cTs, ct) = Thm.instantiate_cterm (map (dest_TVar o Thm.typ_of) cTs ~~ cUs, []) ct +fun instTs cUs (cTs, ct) = + Thm.instantiate_cterm (TVars.make (map (dest_TVar o Thm.typ_of) cTs ~~ cUs), Vars.empty) ct fun instT cU (cT, ct) = instTs [cU] ([cT], ct) fun instT' ct = instT (Thm.ctyp_of_cterm ct) (* certified terms *) fun dest_cabs ct ctxt = (case Thm.term_of ct of Abs _ => let val (n, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt in (snd (Thm.dest_abs (SOME n) ct), ctxt') end | _ => raise CTERM ("no abstraction", [ct])) val dest_all_cabs = repeat_yield (try o dest_cabs) fun dest_cbinder ct ctxt = (case Thm.term_of ct of Const _ $ Abs _ => dest_cabs (Thm.dest_arg ct) ctxt | _ => raise CTERM ("not a binder", [ct])) val dest_all_cbinders = repeat_yield (try o dest_cbinder) val mk_cprop = Thm.apply (Thm.cterm_of \<^context> \<^const>\Trueprop\) fun dest_cprop ct = (case Thm.term_of ct of \<^const>\Trueprop\ $ _ => Thm.dest_arg ct | _ => raise CTERM ("not a property", [ct])) val equals = mk_const_pat \<^theory> \<^const_name>\Pure.eq\ Thm.dest_ctyp0 fun mk_cequals ct cu = Thm.mk_binop (instT' ct equals) ct cu val dest_prop = (fn \<^const>\Trueprop\ $ t => t | t => t) fun term_of ct = dest_prop (Thm.term_of ct) fun prop_of thm = dest_prop (Thm.prop_of thm) (* conversions *) fun if_conv pred cv1 cv2 ct = if pred (Thm.term_of ct) then cv1 ct else cv2 ct fun if_true_conv pred cv = if_conv pred cv Conv.all_conv fun if_exists_conv pred = if_true_conv (Term.exists_subterm pred) fun binders_conv cv ctxt = Conv.binder_conv (binders_conv cv o snd) ctxt else_conv cv ctxt fun under_quant_conv cv ctxt = let fun quant_conv inside ctxt cvs ct = (case Thm.term_of ct of Const (\<^const_name>\All\, _) $ Abs _ => Conv.binder_conv (under_conv cvs) ctxt | Const (\<^const_name>\Ex\, _) $ Abs _ => Conv.binder_conv (under_conv cvs) ctxt | _ => if inside then cv (ctxt, cvs) else Conv.all_conv) ct and under_conv cvs (cv, ctxt) = quant_conv true ctxt (cv :: cvs) in quant_conv false ctxt [] end fun prop_conv cv ct = (case Thm.term_of ct of \<^const>\Trueprop\ $ _ => Conv.arg_conv cv ct | _ => raise CTERM ("not a property", [ct])) end; diff --git a/src/HOL/Tools/Transfer/transfer.ML b/src/HOL/Tools/Transfer/transfer.ML --- a/src/HOL/Tools/Transfer/transfer.ML +++ b/src/HOL/Tools/Transfer/transfer.ML @@ -1,932 +1,932 @@ (* Title: HOL/Tools/Transfer/transfer.ML Author: Brian Huffman, TU Muenchen Author: Ondrej Kuncar, TU Muenchen Generic theorem transfer method. *) signature TRANSFER = sig type pred_data val mk_pred_data: thm -> thm -> thm list -> pred_data val rel_eq_onp: pred_data -> thm val pred_def: pred_data -> thm val pred_simps: pred_data -> thm list val update_pred_simps: thm list -> pred_data -> pred_data val bottom_rewr_conv: thm list -> conv val top_rewr_conv: thm list -> conv val top_sweep_rewr_conv: thm list -> conv val prep_conv: conv val fold_relator_eqs_conv: Proof.context -> conv val unfold_relator_eqs_conv: Proof.context -> conv val get_transfer_raw: Proof.context -> thm list val get_relator_eq: Proof.context -> thm list val retrieve_relator_eq: Proof.context -> term -> thm list val get_sym_relator_eq: Proof.context -> thm list val get_relator_eq_raw: Proof.context -> thm list val get_relator_domain: Proof.context -> thm list val morph_pred_data: morphism -> pred_data -> pred_data val lookup_pred_data: Proof.context -> string -> pred_data option val update_pred_data: string -> pred_data -> Context.generic -> Context.generic val is_compound_lhs: Proof.context -> term -> bool val is_compound_rhs: Proof.context -> term -> bool val transfer_add: attribute val transfer_del: attribute val transfer_raw_add: thm -> Context.generic -> Context.generic val transfer_raw_del: thm -> Context.generic -> Context.generic val transferred_attribute: thm list -> attribute val untransferred_attribute: thm list -> attribute val prep_transfer_domain_thm: Proof.context -> thm -> thm val transfer_domain_add: attribute val transfer_domain_del: attribute val transfer_rule_of_term: Proof.context -> bool -> term -> thm val transfer_rule_of_lhs: Proof.context -> term -> thm val eq_tac: Proof.context -> int -> tactic val transfer_start_tac: bool -> Proof.context -> int -> tactic val transfer_prover_start_tac: Proof.context -> int -> tactic val transfer_step_tac: Proof.context -> int -> tactic val transfer_end_tac: Proof.context -> int -> tactic val transfer_prover_end_tac: Proof.context -> int -> tactic val transfer_tac: bool -> Proof.context -> int -> tactic val transfer_prover_tac: Proof.context -> int -> tactic val gen_frees_tac: (string * typ) list -> Proof.context -> int -> tactic end structure Transfer : TRANSFER = struct fun bottom_rewr_conv rewrs = Conv.bottom_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context> fun top_rewr_conv rewrs = Conv.top_conv (K (Conv.try_conv (Conv.rewrs_conv rewrs))) \<^context> fun top_sweep_rewr_conv rewrs = Conv.top_sweep_conv (K (Conv.rewrs_conv rewrs)) \<^context> (** Theory Data **) val compound_xhs_empty_net = Item_Net.init (Thm.eq_thm_prop o apply2 snd) (single o fst); val rewr_rules = Item_Net.init Thm.eq_thm_prop (single o fst o HOLogic.dest_eq o HOLogic.dest_Trueprop o Thm.concl_of); datatype pred_data = PRED_DATA of {pred_def: thm, rel_eq_onp: thm, pred_simps: thm list} fun mk_pred_data pred_def rel_eq_onp pred_simps = PRED_DATA {pred_def = pred_def, rel_eq_onp = rel_eq_onp, pred_simps = pred_simps} fun map_pred_data' f1 f2 f3 (PRED_DATA {pred_def, rel_eq_onp, pred_simps}) = PRED_DATA {pred_def = f1 pred_def, rel_eq_onp = f2 rel_eq_onp, pred_simps = f3 pred_simps} fun rep_pred_data (PRED_DATA p) = p val rel_eq_onp = #rel_eq_onp o rep_pred_data val pred_def = #pred_def o rep_pred_data val pred_simps = #pred_simps o rep_pred_data fun update_pred_simps new_pred_data = map_pred_data' I I (K new_pred_data) structure Data = Generic_Data ( type T = { transfer_raw : thm Item_Net.T, known_frees : (string * typ) list, compound_lhs : (term * thm) Item_Net.T, compound_rhs : (term * thm) Item_Net.T, relator_eq : thm Item_Net.T, relator_eq_raw : thm Item_Net.T, relator_domain : thm Item_Net.T, pred_data : pred_data Symtab.table } val empty = { transfer_raw = Thm.item_net_intro, known_frees = [], compound_lhs = compound_xhs_empty_net, compound_rhs = compound_xhs_empty_net, relator_eq = rewr_rules, relator_eq_raw = Thm.item_net, relator_domain = Thm.item_net, pred_data = Symtab.empty } val extend = I fun merge ( { transfer_raw = t1, known_frees = k1, compound_lhs = l1, compound_rhs = c1, relator_eq = r1, relator_eq_raw = rw1, relator_domain = rd1, pred_data = pd1 }, { transfer_raw = t2, known_frees = k2, compound_lhs = l2, compound_rhs = c2, relator_eq = r2, relator_eq_raw = rw2, relator_domain = rd2, pred_data = pd2 } ) = { transfer_raw = Item_Net.merge (t1, t2), known_frees = Library.merge (op =) (k1, k2), compound_lhs = Item_Net.merge (l1, l2), compound_rhs = Item_Net.merge (c1, c2), relator_eq = Item_Net.merge (r1, r2), relator_eq_raw = Item_Net.merge (rw1, rw2), relator_domain = Item_Net.merge (rd1, rd2), pred_data = Symtab.merge (K true) (pd1, pd2) } ) fun get_net_content f ctxt = Item_Net.content (f (Data.get (Context.Proof ctxt))) |> map (Thm.transfer' ctxt) val get_transfer_raw = get_net_content #transfer_raw val get_known_frees = #known_frees o Data.get o Context.Proof fun is_compound f ctxt t = Item_Net.retrieve (f (Data.get (Context.Proof ctxt))) t |> exists (fn (pat, _) => Pattern.matches (Proof_Context.theory_of ctxt) (pat, t)); val is_compound_lhs = is_compound #compound_lhs val is_compound_rhs = is_compound #compound_rhs val get_relator_eq = get_net_content #relator_eq #> map safe_mk_meta_eq fun retrieve_relator_eq ctxt t = Item_Net.retrieve (#relator_eq (Data.get (Context.Proof ctxt))) t |> map (Thm.transfer' ctxt) val get_sym_relator_eq = get_net_content #relator_eq #> map (safe_mk_meta_eq #> Thm.symmetric) val get_relator_eq_raw = get_net_content #relator_eq_raw val get_relator_domain = get_net_content #relator_domain val get_pred_data = #pred_data o Data.get o Context.Proof fun map_data f1 f2 f3 f4 f5 f6 f7 f8 { transfer_raw, known_frees, compound_lhs, compound_rhs, relator_eq, relator_eq_raw, relator_domain, pred_data } = { transfer_raw = f1 transfer_raw, known_frees = f2 known_frees, compound_lhs = f3 compound_lhs, compound_rhs = f4 compound_rhs, relator_eq = f5 relator_eq, relator_eq_raw = f6 relator_eq_raw, relator_domain = f7 relator_domain, pred_data = f8 pred_data } fun map_transfer_raw f = map_data f I I I I I I I fun map_known_frees f = map_data I f I I I I I I fun map_compound_lhs f = map_data I I f I I I I I fun map_compound_rhs f = map_data I I I f I I I I fun map_relator_eq f = map_data I I I I f I I I fun map_relator_eq_raw f = map_data I I I I I f I I fun map_relator_domain f = map_data I I I I I I f I fun map_pred_data f = map_data I I I I I I I f val add_transfer_thm = Thm.trim_context #> (fn thm => Data.map (map_transfer_raw (Item_Net.update thm) o map_compound_lhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ (lhs as (_ $ _)) $ _ => Item_Net.update (lhs, thm) | _ => I) o map_compound_rhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ _ $ (rhs as (_ $ _)) => Item_Net.update (rhs, thm) | _ => I) o map_known_frees (Term.add_frees (Thm.concl_of thm)))) fun del_transfer_thm thm = Data.map (map_transfer_raw (Item_Net.remove thm) o map_compound_lhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ (lhs as (_ $ _)) $ _ => Item_Net.remove (lhs, thm) | _ => I) o map_compound_rhs (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Rel\, _) $ _ $ _ $ (rhs as (_ $ _)) => Item_Net.remove (rhs, thm) | _ => I)) fun transfer_raw_add thm ctxt = add_transfer_thm thm ctxt fun transfer_raw_del thm ctxt = del_transfer_thm thm ctxt (** Conversions **) fun transfer_rel_conv conv = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.fun2_conv (Conv.arg_conv conv))) val Rel_rule = Thm.symmetric @{thm Rel_def} fun Rel_conv ct = let val (cT, cT') = Thm.dest_funT (Thm.ctyp_of_cterm ct) val (cU, _) = Thm.dest_funT cT' in Thm.instantiate' [SOME cT, SOME cU] [SOME ct] Rel_rule end (* Conversion to preprocess a transfer rule *) fun safe_Rel_conv ct = Conv.try_conv (HOLogic.Trueprop_conv (Conv.fun_conv (Conv.fun_conv Rel_conv))) ct fun prep_conv ct = ( Conv.implies_conv safe_Rel_conv prep_conv else_conv safe_Rel_conv else_conv Conv.all_conv) ct fun fold_relator_eqs_conv ctxt ct = (bottom_rewr_conv (get_relator_eq ctxt)) ct; fun unfold_relator_eqs_conv ctxt ct = (top_rewr_conv (get_sym_relator_eq ctxt)) ct; (** Replacing explicit equalities with is_equality premises **) fun mk_is_equality t = Const (\<^const_name>\is_equality\, Term.fastype_of t --> HOLogic.boolT) $ t fun gen_abstract_equalities ctxt (dest : term -> term * (term -> term)) thm = let val prop = Thm.prop_of thm val (t, mk_prop') = dest prop (* Only consider "(=)" at non-base types *) fun is_eq (Const (\<^const_name>\HOL.eq\, Type ("fun", [T, _]))) = (case T of Type (_, []) => false | _ => true) | is_eq _ = false val add_eqs = Term.fold_aterms (fn t => if is_eq t then insert (op =) t else I) val eq_consts = rev (add_eqs t []) val eqTs = map (snd o dest_Const) eq_consts val used = Term.add_free_names prop [] val names = map (K "") eqTs |> Name.variant_list used val frees = map Free (names ~~ eqTs) val prems = map (HOLogic.mk_Trueprop o mk_is_equality) frees val prop1 = mk_prop' (Term.subst_atomic (eq_consts ~~ frees) t) val prop2 = fold Logic.all frees (Logic.list_implies (prems, prop1)) val cprop = Thm.cterm_of ctxt prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false @{thms is_equality_lemma} cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm in forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2})) end handle TERM _ => thm fun abstract_equalities_transfer ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (rel, fn rel' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel' $ x $ y))) end val contracted_eq_thm = Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm handle CTERM _ => thm in gen_abstract_equalities ctxt dest contracted_eq_thm end fun abstract_equalities_relator_eq ctxt rel_eq_thm = gen_abstract_equalities ctxt (fn x => (x, I)) (rel_eq_thm RS @{thm is_equality_def [THEN iffD2]}) fun abstract_equalities_domain ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((eq, dom), y) = apfst Term.dest_comb (Term.dest_comb concl) in (dom, fn dom' => Logic.list_implies (prems, HOLogic.mk_Trueprop (eq $ dom' $ y))) end fun transfer_rel_conv conv = Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv conv))) val contracted_eq_thm = Conv.fconv_rule (transfer_rel_conv (fold_relator_eqs_conv ctxt)) thm in gen_abstract_equalities ctxt dest contracted_eq_thm end (** Replacing explicit Domainp predicates with Domainp assumptions **) fun mk_Domainp_assm (T, R) = HOLogic.mk_eq ((Const (\<^const_name>\Domainp\, Term.fastype_of T --> Term.fastype_of R) $ T), R) fun fold_Domainp f (t as Const (\<^const_name>\Domainp\,_) $ (Var (_,_))) = f t | fold_Domainp f (t $ u) = fold_Domainp f t #> fold_Domainp f u | fold_Domainp f (Abs (_, _, t)) = fold_Domainp f t | fold_Domainp _ _ = I fun subst_terms tab t = let val t' = Termtab.lookup tab t in case t' of SOME t' => t' | NONE => (case t of u $ v => (subst_terms tab u) $ (subst_terms tab v) | Abs (a, T, t) => Abs (a, T, subst_terms tab t) | t => t) end fun gen_abstract_domains ctxt (dest : term -> term * (term -> term)) thm = let val prop = Thm.prop_of thm val (t, mk_prop') = dest prop val Domainp_tms = rev (fold_Domainp (fn t => insert op= t) t []) val Domainp_Ts = map (snd o dest_funT o snd o dest_Const o fst o dest_comb) Domainp_tms val used = Term.add_free_names t [] val rels = map (snd o dest_comb) Domainp_tms val rel_names = map (fst o fst o dest_Var) rels val names = map (fn name => ("D" ^ name)) rel_names |> Name.variant_list used val frees = map Free (names ~~ Domainp_Ts) val prems = map (HOLogic.mk_Trueprop o mk_Domainp_assm) (rels ~~ frees); val t' = subst_terms (fold Termtab.update (Domainp_tms ~~ frees) Termtab.empty) t val prop1 = fold Logic.all frees (Logic.list_implies (prems, mk_prop' t')) val prop2 = Logic.list_rename_params (rev names) prop1 val cprop = Thm.cterm_of ctxt prop2 val equal_thm = Raw_Simplifier.rewrite ctxt false @{thms Domainp_lemma} cprop fun forall_elim thm = Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm; in forall_elim (thm COMP (equal_thm COMP @{thm equal_elim_rule2})) end handle TERM _ => thm fun abstract_domains_transfer ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (x, fn x' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x' $ y))) end in gen_abstract_domains ctxt dest thm end fun abstract_domains_relator_domain ctxt thm = let fun dest prop = let val prems = Logic.strip_imp_prems prop val concl = HOLogic.dest_Trueprop (Logic.strip_imp_concl prop) val ((rel, x), y) = apfst Term.dest_comb (Term.dest_comb concl) in (y, fn y' => Logic.list_implies (prems, HOLogic.mk_Trueprop (rel $ x $ y'))) end in gen_abstract_domains ctxt dest thm end fun detect_transfer_rules thm = let fun is_transfer_rule tm = case (HOLogic.dest_Trueprop tm) of (Const (\<^const_name>\HOL.eq\, _)) $ ((Const (\<^const_name>\Domainp\, _)) $ _) $ _ => false | _ $ _ $ _ => true | _ => false fun safe_transfer_rule_conv ctm = if is_transfer_rule (Thm.term_of ctm) then safe_Rel_conv ctm else Conv.all_conv ctm in Conv.fconv_rule (Conv.prems_conv ~1 safe_transfer_rule_conv) thm end (** Adding transfer domain rules **) fun prep_transfer_domain_thm ctxt = abstract_equalities_domain ctxt o detect_transfer_rules fun add_transfer_domain_thm thm ctxt = (add_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt fun del_transfer_domain_thm thm ctxt = (del_transfer_thm o prep_transfer_domain_thm (Context.proof_of ctxt)) thm ctxt (** Transfer proof method **) val post_simps = @{thms transfer_forall_eq [symmetric] transfer_implies_eq [symmetric] transfer_bforall_unfold} fun gen_frees_tac keepers ctxt = SUBGOAL (fn (t, i) => let val keepers = keepers @ get_known_frees ctxt val vs = rev (Term.add_frees t []) val vs' = filter_out (member (op =) keepers) vs in Induct.arbitrary_tac ctxt 0 vs' i end) fun mk_relT (T, U) = T --> U --> HOLogic.boolT fun mk_Rel t = let val T = fastype_of t in Const (\<^const_name>\Transfer.Rel\, T --> T) $ t end fun transfer_rule_of_terms (prj : typ * typ -> typ) ctxt tab t u = let (* precondition: prj(T,U) must consist of only TFrees and type "fun" *) fun rel (T as Type ("fun", [T1, T2])) (U as Type ("fun", [U1, U2])) = let val r1 = rel T1 U1 val r2 = rel T2 U2 val rT = fastype_of r1 --> fastype_of r2 --> mk_relT (T, U) in Const (\<^const_name>\rel_fun\, rT) $ r1 $ r2 end | rel T U = let val (a, _) = dest_TFree (prj (T, U)) in Free (the (AList.lookup (op =) tab a), mk_relT (T, U)) end fun zip _ thms (Bound i) (Bound _) = (nth thms i, []) | zip ctxt thms (Abs (x, T, t)) (Abs (y, U, u)) = let val ([x', y'], ctxt') = Variable.variant_fixes [x, y] ctxt val prop = mk_Rel (rel T U) $ Free (x', T) $ Free (y', U) val cprop = Thm.cterm_of ctxt' (HOLogic.mk_Trueprop prop) val thm0 = Thm.assume cprop val (thm1, hyps) = zip ctxt' (thm0 :: thms) t u val ((r1, x), y) = apfst Thm.dest_comb (Thm.dest_comb (Thm.dest_arg cprop)) val r2 = Thm.dest_fun2 (Thm.dest_arg (Thm.cprop_of thm1)) val (a1, (b1, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r1)) val (a2, (b2, _)) = apsnd Thm.dest_funT (Thm.dest_funT (Thm.ctyp_of_cterm r2)) val tinsts = [SOME a1, SOME b1, SOME a2, SOME b2] val insts = [SOME (Thm.dest_arg r1), SOME (Thm.dest_arg r2)] val rule = Thm.instantiate' tinsts insts @{thm Rel_abs} val thm2 = Thm.forall_intr x (Thm.forall_intr y (Thm.implies_intr cprop thm1)) in (thm2 COMP rule, hyps) end | zip ctxt thms (f $ t) (g $ u) = let val (thm1, hyps1) = zip ctxt thms f g val (thm2, hyps2) = zip ctxt thms t u in (thm2 RS (thm1 RS @{thm Rel_app}), hyps1 @ hyps2) end | zip _ _ t u = let val T = fastype_of t val U = fastype_of u val prop = mk_Rel (rel T U) $ t $ u val cprop = Thm.cterm_of ctxt (HOLogic.mk_Trueprop prop) in (Thm.assume cprop, [cprop]) end val r = mk_Rel (rel (fastype_of t) (fastype_of u)) val goal = HOLogic.mk_Trueprop (r $ t $ u) val rename = Thm.trivial (Thm.cterm_of ctxt goal) val (thm, hyps) = zip ctxt [] t u in Drule.implies_intr_list hyps (thm RS rename) end (* create a lambda term of the same shape as the given term *) fun skeleton is_atom = let fun dummy ctxt = let val (c, ctxt') = yield_singleton Variable.variant_fixes "a" ctxt in (Free (c, dummyT), ctxt') end fun skel (Bound i) ctxt = (Bound i, ctxt) | skel (Abs (x, _, t)) ctxt = let val (t', ctxt) = skel t ctxt in (Abs (x, dummyT, t'), ctxt) end | skel (tu as t $ u) ctxt = if is_atom tu andalso not (Term.is_open tu) then dummy ctxt else let val (t', ctxt) = skel t ctxt val (u', ctxt) = skel u ctxt in (t' $ u', ctxt) end | skel _ ctxt = dummy ctxt in fn ctxt => fn t => fst (skel t ctxt) |> Syntax.check_term ctxt (* FIXME avoid syntax operation!? *) end (** Monotonicity analysis **) (* TODO: Put extensible table in theory data *) val monotab = Symtab.make [(\<^const_name>\transfer_implies\, [~1, 1]), (\<^const_name>\transfer_forall\, [1])(*, (@{const_name implies}, [~1, 1]), (@{const_name All}, [1])*)] (* Function bool_insts determines the set of boolean-relation variables that can be instantiated to implies, rev_implies, or iff. Invariants: bool_insts p (t, u) requires that u :: _ => _ => ... => bool, and t is a skeleton of u *) fun bool_insts p (t, u) = let fun strip2 (t1 $ t2, u1 $ u2, tus) = strip2 (t1, u1, (t2, u2) :: tus) | strip2 x = x fun or3 ((a, b, c), (x, y, z)) = (a orelse x, b orelse y, c orelse z) fun go Ts p (Abs (_, T, t), Abs (_, _, u)) tab = go (T :: Ts) p (t, u) tab | go Ts p (t, u) tab = let val (a, _) = dest_TFree (Term.body_type (Term.fastype_of1 (Ts, t))) val (_, tf, tus) = strip2 (t, u, []) val ps_opt = case tf of Const (c, _) => Symtab.lookup monotab c | _ => NONE val tab1 = case ps_opt of SOME ps => let val ps' = map (fn x => p * x) (take (length tus) ps) in fold I (map2 (go Ts) ps' tus) tab end | NONE => tab val tab2 = Symtab.make [(a, (p >= 0, p <= 0, is_none ps_opt))] in Symtab.join (K or3) (tab1, tab2) end val tab = go [] p (t, u) Symtab.empty fun f (a, (true, false, false)) = SOME (a, \<^const>\implies\) | f (a, (false, true, false)) = SOME (a, \<^const>\rev_implies\) | f (a, (true, true, _)) = SOME (a, HOLogic.eq_const HOLogic.boolT) | f _ = NONE in map_filter f (Symtab.dest tab) end fun transfer_rule_of_term ctxt equiv t = let val s = skeleton (is_compound_rhs ctxt) ctxt t val frees = map fst (Term.add_frees s []) val tfrees = map fst (Term.add_tfrees s []) fun prep a = "R" ^ Library.unprefix "'" a val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt val tab = tfrees ~~ rnames fun prep a = the (AList.lookup (op =) tab a) val thm = transfer_rule_of_terms fst ctxt' tab s t val binsts = bool_insts (if equiv then 0 else 1) (s, t) val idx = Thm.maxidx_of thm + 1 fun tinst (a, _) = (((a, idx), \<^sort>\type\), \<^ctyp>\bool\) fun inst (a, t) = ((Name.clean_index (prep a, idx), \<^typ>\bool \ bool \ bool\), Thm.cterm_of ctxt' t) in thm |> Thm.generalize (Names.make_set tfrees, Names.make_set (rnames @ frees)) idx - |> Thm.instantiate (map tinst binsts, map inst binsts) + |> Thm.instantiate (TVars.make (map tinst binsts), Vars.make (map inst binsts)) end fun transfer_rule_of_lhs ctxt t = let val s = skeleton (is_compound_lhs ctxt) ctxt t val frees = map fst (Term.add_frees s []) val tfrees = map fst (Term.add_tfrees s []) fun prep a = "R" ^ Library.unprefix "'" a val (rnames, ctxt') = Variable.variant_fixes (map prep tfrees) ctxt val tab = tfrees ~~ rnames fun prep a = the (AList.lookup (op =) tab a) val thm = transfer_rule_of_terms snd ctxt' tab t s val binsts = bool_insts 1 (s, t) val idx = Thm.maxidx_of thm + 1 fun tinst (a, _) = (((a, idx), \<^sort>\type\), \<^ctyp>\bool\) fun inst (a, t) = ((Name.clean_index (prep a, idx), \<^typ>\bool \ bool \ bool\), Thm.cterm_of ctxt' t) in thm |> Thm.generalize (Names.make_set tfrees, Names.make_set (rnames @ frees)) idx - |> Thm.instantiate (map tinst binsts, map inst binsts) + |> Thm.instantiate (TVars.make (map tinst binsts), Vars.make (map inst binsts)) end fun eq_rules_tac ctxt eq_rules = TRY o REPEAT_ALL_NEW (resolve_tac ctxt eq_rules) THEN_ALL_NEW resolve_tac ctxt @{thms is_equality_eq} fun eq_tac ctxt = eq_rules_tac ctxt (get_relator_eq_raw ctxt) fun transfer_step_tac ctxt = REPEAT_ALL_NEW (resolve_tac ctxt (get_transfer_raw ctxt)) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt (get_relator_eq_raw ctxt)) fun transfer_start_tac equiv ctxt i = let val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} val start_rule = if equiv then @{thm transfer_start} else @{thm transfer_start'} val err_msg = "Transfer failed to convert goal to an object-logic formula" fun main_tac (t, i) = resolve_tac ctxt [start_rule] i THEN (resolve_tac ctxt [transfer_rule_of_term ctxt equiv (HOLogic.dest_Trueprop t)]) (i + 1) handle TERM (_, ts) => raise TERM (err_msg, ts) in EVERY [rewrite_goal_tac ctxt pre_simps i THEN SUBGOAL main_tac i] end; fun transfer_prover_start_tac ctxt = SUBGOAL (fn (t, i) => let val rhs = (snd o Term.dest_comb o HOLogic.dest_Trueprop) t val rule1 = transfer_rule_of_term ctxt false rhs val expand_eqs_in_rel_conv = transfer_rel_conv (unfold_relator_eqs_conv ctxt) in EVERY [CONVERSION prep_conv i, CONVERSION expand_eqs_in_rel_conv i, resolve_tac ctxt @{thms transfer_prover_start} i, resolve_tac ctxt [rule1] (i + 1)] end); fun transfer_end_tac ctxt i = let val post_simps = @{thms transfer_forall_eq [symmetric] transfer_implies_eq [symmetric] transfer_bforall_unfold} in EVERY [rewrite_goal_tac ctxt post_simps i, Goal.norm_hhf_tac ctxt i] end; fun transfer_prover_end_tac ctxt i = resolve_tac ctxt @{thms refl} i local infix 1 THEN_ALL_BUT_FIRST_NEW fun (tac1 THEN_ALL_BUT_FIRST_NEW tac2) i st = st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 (i + 1) (i + Thm.nprems_of st' - Thm.nprems_of st) st')); in fun transfer_tac equiv ctxt i = let val rules = get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt (* allow unsolved subgoals only for standard transfer method, not for transfer' *) val end_tac = if equiv then K all_tac else K no_tac fun transfer_search_tac i = (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt eq_rules)) ORELSE' end_tac) i in (transfer_start_tac equiv ctxt THEN_ALL_BUT_FIRST_NEW transfer_search_tac THEN' transfer_end_tac ctxt) i end fun transfer_prover_tac ctxt i = let val rules = get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt fun transfer_prover_search_tac i = (REPEAT_ALL_NEW (resolve_tac ctxt rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt eq_rules)) i in (transfer_prover_start_tac ctxt THEN_ALL_BUT_FIRST_NEW transfer_prover_search_tac THEN' transfer_prover_end_tac ctxt) i end end; (** Transfer attribute **) fun transferred ctxt extra_rules thm = let val rules = extra_rules @ get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} val thm1 = Thm.forall_intr_vars thm val instT = rev (Term.add_tvars (Thm.full_prop_of thm1) []) |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S)))) val thm2 = thm1 - |> Thm.instantiate (instT, []) + |> Thm.instantiate (TVars.make instT, Vars.empty) |> Raw_Simplifier.rewrite_rule ctxt pre_simps val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt val rule = transfer_rule_of_lhs ctxt' (HOLogic.dest_Trueprop (Thm.concl_of thm2)) in Goal.prove_internal ctxt' [] (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\bool\)))) (fn _ => resolve_tac ctxt' [thm2 RS @{thm transfer_start'}, thm2 RS @{thm transfer_start}] 1 THEN (resolve_tac ctxt' [rule] THEN_ALL_NEW (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1 handle TERM (_, ts) => raise TERM ("Transfer failed to convert goal to an object-logic formula", ts)) |> Raw_Simplifier.rewrite_rule ctxt' post_simps |> Simplifier.norm_hhf ctxt' |> Drule.generalize (Names.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Names.empty) |> Drule.zero_var_indexes end (* handle THM _ => thm *) fun untransferred ctxt extra_rules thm = let val rules = extra_rules @ get_transfer_raw ctxt val eq_rules = get_relator_eq_raw ctxt val pre_simps = @{thms transfer_forall_eq transfer_implies_eq} val thm1 = Thm.forall_intr_vars thm val instT = rev (Term.add_tvars (Thm.full_prop_of thm1) []) |> map (fn v as ((a, _), S) => (v, Thm.ctyp_of ctxt (TFree (a, S)))) val thm2 = thm1 - |> Thm.instantiate (instT, []) + |> Thm.instantiate (TVars.make instT, Vars.empty) |> Raw_Simplifier.rewrite_rule ctxt pre_simps val ctxt' = Variable.declare_names (Thm.full_prop_of thm2) ctxt val t = HOLogic.dest_Trueprop (Thm.concl_of thm2) val rule = transfer_rule_of_term ctxt' true t in Goal.prove_internal ctxt' [] (Thm.cterm_of ctxt' (HOLogic.mk_Trueprop (Var (("P", 0), \<^typ>\bool\)))) (fn _ => resolve_tac ctxt' [thm2 RS @{thm untransfer_start}] 1 THEN (resolve_tac ctxt' [rule] THEN_ALL_NEW (SOLVED' (REPEAT_ALL_NEW (resolve_tac ctxt' rules) THEN_ALL_NEW (DETERM o eq_rules_tac ctxt' eq_rules)))) 1 handle TERM (_, ts) => raise TERM ("Transfer failed to convert goal to an object-logic formula", ts)) |> Raw_Simplifier.rewrite_rule ctxt' post_simps |> Simplifier.norm_hhf ctxt' |> Drule.generalize (Names.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Names.empty) |> Drule.zero_var_indexes end (** Methods and attributes **) val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)) val fixing = Scan.optional (Scan.lift (Args.$$$ "fixing" -- Args.colon) |-- Scan.repeat free) [] val reverse_prems = fn _ => PRIMITIVE (fn thm => fold_rev (fn i => Thm.permute_prems i 1) (0 upto Thm.nprems_of thm - 1) thm); fun transfer_start_method equiv : (Proof.context -> Proof.method) context_parser = fixing >> (fn vs => fn ctxt => SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_start_tac equiv ctxt THEN' reverse_prems)); fun transfer_method equiv : (Proof.context -> Proof.method) context_parser = fixing >> (fn vs => fn ctxt => SIMPLE_METHOD' (gen_frees_tac vs ctxt THEN' transfer_tac equiv ctxt)) val transfer_prover_start_method : (Proof.context -> Proof.method) context_parser = Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_start_tac ctxt THEN' reverse_prems)) val transfer_prover_method : (Proof.context -> Proof.method) context_parser = Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_tac ctxt)) (* Attribute for transfer rules *) fun prep_rule ctxt = abstract_domains_transfer ctxt o abstract_equalities_transfer ctxt o Conv.fconv_rule prep_conv val transfer_add = Thm.declaration_attribute (fn thm => fn ctxt => (add_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt) val transfer_del = Thm.declaration_attribute (fn thm => fn ctxt => (del_transfer_thm o prep_rule (Context.proof_of ctxt)) thm ctxt) val transfer_attribute = Attrib.add_del transfer_add transfer_del (* Attributes for transfer domain rules *) val transfer_domain_add = Thm.declaration_attribute add_transfer_domain_thm val transfer_domain_del = Thm.declaration_attribute del_transfer_domain_thm val transfer_domain_attribute = Attrib.add_del transfer_domain_add transfer_domain_del (* Attributes for transferred rules *) fun transferred_attribute thms = Thm.rule_attribute thms (fn context => transferred (Context.proof_of context) thms) fun untransferred_attribute thms = Thm.rule_attribute thms (fn context => untransferred (Context.proof_of context) thms) val transferred_attribute_parser = Attrib.thms >> transferred_attribute val untransferred_attribute_parser = Attrib.thms >> untransferred_attribute fun morph_pred_data phi = let val morph_thm = Morphism.thm phi in map_pred_data' morph_thm morph_thm (map morph_thm) end fun lookup_pred_data ctxt type_name = Symtab.lookup (get_pred_data ctxt) type_name |> Option.map (morph_pred_data (Morphism.transfer_morphism' ctxt)) fun update_pred_data type_name qinfo ctxt = Data.map (map_pred_data (Symtab.update (type_name, qinfo))) ctxt (* Theory setup *) val _ = Theory.setup let val name = \<^binding>\relator_eq\ fun add_thm thm context = context |> Data.map (map_relator_eq (Item_Net.update (Thm.trim_context thm))) |> Data.map (map_relator_eq_raw (Item_Net.update (Thm.trim_context (abstract_equalities_relator_eq (Context.proof_of context) thm)))) fun del_thm thm context = context |> Data.map (map_relator_eq (Item_Net.remove thm)) |> Data.map (map_relator_eq_raw (Item_Net.remove (abstract_equalities_relator_eq (Context.proof_of context) thm))) val add = Thm.declaration_attribute add_thm val del = Thm.declaration_attribute del_thm val text = "declaration of relator equality rule (used by transfer method)" val content = Item_Net.content o #relator_eq o Data.get in Attrib.setup name (Attrib.add_del add del) text #> Global_Theory.add_thms_dynamic (name, content) end val _ = Theory.setup let val name = \<^binding>\relator_domain\ fun add_thm thm context = let val thm' = thm |> abstract_domains_relator_domain (Context.proof_of context) |> Thm.trim_context in context |> Data.map (map_relator_domain (Item_Net.update thm')) |> add_transfer_domain_thm thm' end fun del_thm thm context = let val thm' = abstract_domains_relator_domain (Context.proof_of context) thm in context |> Data.map (map_relator_domain (Item_Net.remove thm')) |> del_transfer_domain_thm thm' end val add = Thm.declaration_attribute add_thm val del = Thm.declaration_attribute del_thm val text = "declaration of relator domain rule (used by transfer method)" val content = Item_Net.content o #relator_domain o Data.get in Attrib.setup name (Attrib.add_del add del) text #> Global_Theory.add_thms_dynamic (name, content) end val _ = Theory.setup (Attrib.setup \<^binding>\transfer_rule\ transfer_attribute "transfer rule for transfer method" #> Global_Theory.add_thms_dynamic (\<^binding>\transfer_raw\, Item_Net.content o #transfer_raw o Data.get) #> Attrib.setup \<^binding>\transfer_domain_rule\ transfer_domain_attribute "transfer domain rule for transfer method" #> Attrib.setup \<^binding>\transferred\ transferred_attribute_parser "raw theorem transferred to abstract theorem using transfer rules" #> Attrib.setup \<^binding>\untransferred\ untransferred_attribute_parser "abstract theorem transferred to raw theorem using transfer rules" #> Global_Theory.add_thms_dynamic (\<^binding>\relator_eq_raw\, Item_Net.content o #relator_eq_raw o Data.get) #> Method.setup \<^binding>\transfer_start\ (transfer_start_method true) "firtst step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_start'\ (transfer_start_method false) "firtst step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_prover_start\ transfer_prover_start_method "firtst step in the transfer_prover algorithm (for debugging transfer_prover)" #> Method.setup \<^binding>\transfer_step\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_step_tac ctxt))) "step in the search for transfer rules (for debugging transfer and transfer_prover)" #> Method.setup \<^binding>\transfer_end\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_end_tac ctxt))) "last step in the transfer algorithm (for debugging transfer)" #> Method.setup \<^binding>\transfer_prover_end\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (transfer_prover_end_tac ctxt))) "last step in the transfer_prover algorithm (for debugging transfer_prover)" #> Method.setup \<^binding>\transfer\ (transfer_method true) "generic theorem transfer method" #> Method.setup \<^binding>\transfer'\ (transfer_method false) "generic theorem transfer method" #> Method.setup \<^binding>\transfer_prover\ transfer_prover_method "for proving transfer rules") end diff --git a/src/HOL/Tools/coinduction.ML b/src/HOL/Tools/coinduction.ML --- a/src/HOL/Tools/coinduction.ML +++ b/src/HOL/Tools/coinduction.ML @@ -1,170 +1,170 @@ (* Title: HOL/Tools/coinduction.ML Author: Johannes Hölzl, TU Muenchen Author: Dmitriy Traytel, TU Muenchen Copyright 2013 Coinduction method that avoids some boilerplate compared to coinduct. *) signature COINDUCTION = sig val coinduction_context_tactic: term list -> thm list option -> thm list -> int -> context_tactic val coinduction_tac: Proof.context -> term list -> thm list option -> thm list -> int -> tactic end; structure Coinduction : COINDUCTION = struct fun filter_in_out _ [] = ([], []) | filter_in_out P (x :: xs) = let val (ins, outs) = filter_in_out P xs; in if P x then (x :: ins, outs) else (ins, x :: outs) end; fun ALLGOALS_SKIP skip tac st = let fun doall n = if n = skip then all_tac else tac n THEN doall (n - 1) in doall (Thm.nprems_of st) st end; fun THEN_ALL_NEW_SKIP skip tac1 tac2 i st = st |> (tac1 i THEN (fn st' => Seq.INTERVAL tac2 (i + skip) (i + Thm.nprems_of st' - Thm.nprems_of st) st')); fun DELETE_PREMS_AFTER skip tac i st = let val n = nth (Thm.prems_of st) (i - 1) |> Logic.strip_assums_hyp |> length; in (THEN_ALL_NEW_SKIP skip tac (REPEAT_DETERM_N n o eresolve0_tac [thin_rl])) i st end; fun coinduction_context_tactic raw_vars opt_raw_thms prems = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => let val lhs_of_eq = HOLogic.dest_Trueprop #> HOLogic.dest_eq #> fst; fun find_coinduct t = Induct.find_coinductP ctxt t @ (try (Induct.find_coinductT ctxt o fastype_of o lhs_of_eq) t |> the_default []); val raw_thms = (case opt_raw_thms of SOME raw_thms => raw_thms | NONE => goal |> Logic.strip_assums_concl |> find_coinduct); val thy = Proof_Context.theory_of ctxt; val concl = Variable.focus NONE goal ctxt |> fst |> snd |> Logic.strip_imp_concl; val raw_thm = (case find_first (fn thm => Pattern.matches thy (Thm.concl_of thm, concl)) raw_thms of SOME thm => thm | NONE => error "No matching coinduction rule found") val skip = Integer.max 1 (Rule_Cases.get_consumes raw_thm) - 1; val cases = Rule_Cases.get raw_thm |> fst; in ((Object_Logic.rulify_tac ctxt THEN' Method.insert_tac ctxt prems THEN' Object_Logic.atomize_prems_tac ctxt THEN' DELETE_PREMS_AFTER skip (Subgoal.FOCUS (fn {concl, context = ctxt, params, prems, ...} => let val vars = raw_vars @ map (Thm.term_of o snd) params; val names_ctxt = ctxt |> fold Variable.declare_names vars |> fold Variable.declare_thm (raw_thm :: prems); val thm_concl = Thm.cprop_of raw_thm |> strip_imp_concl; val (instT, inst) = Thm.match (thm_concl, concl); - val rhoTs = map (fn (v, T) => (TVar v, Thm.typ_of T)) instT; - val rhots = map (fn (v, t) => (Var v, Thm.term_of t)) inst; + val rhoTs = map (fn (v, T) => (TVar v, Thm.typ_of T)) (TVars.dest instT); + val rhots = map (fn (v, t) => (Var v, Thm.term_of t)) (Vars.dest inst); val xs = hd (Thm.prems_of raw_thm) |> HOLogic.dest_Trueprop |> strip_comb |> snd |> map (subst_atomic_types rhoTs); val raw_eqs = map (fn x => (x, AList.lookup op aconv rhots x |> the)) xs; val ((names, ctxt), Ts) = map_split (apfst fst o dest_Var o fst) raw_eqs |>> (fn names => Variable.variant_fixes names names_ctxt) ; val eqs = @{map 3} (fn name => fn T => fn (_, rhs) => HOLogic.mk_eq (Free (name, T), rhs)) names Ts raw_eqs; val phi = eqs @ map (HOLogic.dest_Trueprop o Thm.prop_of) prems |> try (Library.foldr1 HOLogic.mk_conj) |> the_default \<^term>\True\ |> Ctr_Sugar_Util.list_exists_free vars |> Term.map_abs_vars (Variable.revert_fixed ctxt) |> fold_rev Term.absfree (names ~~ Ts) |> Thm.cterm_of ctxt; val thm = infer_instantiate' ctxt [SOME phi] raw_thm; val e = length eqs; val p = length prems; in HEADGOAL (EVERY' [resolve_tac ctxt [thm], EVERY' (map (fn var => resolve_tac ctxt [infer_instantiate' ctxt [NONE, SOME (Thm.cterm_of ctxt var)] exI]) vars), if p = 0 then Ctr_Sugar_Util.CONJ_WRAP' (K (resolve_tac ctxt [refl])) eqs else REPEAT_DETERM_N e o (resolve_tac ctxt [conjI] THEN' resolve_tac ctxt [refl]) THEN' Ctr_Sugar_Util.CONJ_WRAP' (resolve_tac ctxt o single) prems, K (ALLGOALS_SKIP skip (REPEAT_DETERM_N (length vars) o (eresolve_tac ctxt [exE] THEN' rotate_tac ~1) THEN' DELETE_PREMS_AFTER 0 (Subgoal.FOCUS (fn {prems, params, context = ctxt, ...} => (case prems of [] => all_tac | inv :: case_prems => let val (init, last) = funpow_yield (p + e - 1) (HOLogic.conj_elim ctxt) inv; val inv_thms = init @ [last]; val eqs = take e inv_thms; fun is_local_var t = member (fn (t, (_, t')) => t aconv (Thm.term_of t')) params t; val (eqs, assms') = filter_in_out (is_local_var o lhs_of_eq o Thm.prop_of) eqs; val assms = assms' @ drop e inv_thms in HEADGOAL (Method.insert_tac ctxt (assms @ case_prems)) THEN Ctr_Sugar_General_Tactics.unfold_thms_tac ctxt eqs end)) ctxt)))]) end) ctxt) THEN' K (prune_params_tac ctxt)) i) st |> Seq.maps (fn st' => CONTEXT_CASES (Rule_Cases.make_common ctxt (Thm.prop_of st') cases) all_tac (ctxt, st')) end); fun coinduction_tac ctxt x1 x2 x3 x4 = NO_CONTEXT_TACTIC ctxt (coinduction_context_tactic x1 x2 x3 x4); local val ruleN = "rule" val arbitraryN = "arbitrary" fun named_rule k arg get = Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|-- (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name => (case get (Context.proof_of context) name of SOME x => x | NONE => error ("No rule for " ^ k ^ " " ^ quote name)))))); fun rule get_type get_pred = named_rule Induct.typeN (Args.type_name {proper = false, strict = false}) get_type || named_rule Induct.predN (Args.const {proper = false, strict = false}) get_pred || named_rule Induct.setN (Args.const {proper = false, strict = false}) get_pred || Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms; val coinduct_rules = rule Induct.lookup_coinductT Induct.lookup_coinductP; fun unless_more_args scan = Scan.unless (Scan.lift ((Args.$$$ arbitraryN || Args.$$$ Induct.typeN || Args.$$$ Induct.predN || Args.$$$ Induct.setN || Args.$$$ ruleN) -- Args.colon)) scan; val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |-- Scan.repeat1 (unless_more_args Args.term)) []; in val _ = Theory.setup (Method.setup \<^binding>\coinduction\ (arbitrary -- Scan.option coinduct_rules >> (fn (arbitrary, opt_rules) => fn _ => fn facts => Seq.DETERM (coinduction_context_tactic arbitrary opt_rules facts 1))) "coinduction on types or predicates/sets"); end; end; diff --git a/src/HOL/Tools/hologic.ML b/src/HOL/Tools/hologic.ML --- a/src/HOL/Tools/hologic.ML +++ b/src/HOL/Tools/hologic.ML @@ -1,721 +1,721 @@ (* Title: HOL/Tools/hologic.ML Author: Lawrence C Paulson and Markus Wenzel Abstract syntax operations for HOL. *) signature HOLOGIC = sig val id_const: typ -> term val mk_comp: term * term -> term val boolN: string val boolT: typ val mk_obj_eq: thm -> thm val Trueprop: term val mk_Trueprop: term -> term val dest_Trueprop: term -> term val Trueprop_conv: conv -> conv val mk_induct_forall: typ -> term val mk_setT: typ -> typ val dest_setT: typ -> typ val Collect_const: typ -> term val mk_Collect: string * typ * term -> term val mk_mem: term * term -> term val dest_mem: term -> term * term val mk_set: typ -> term list -> term val dest_set: term -> term list val mk_UNIV: typ -> term val conj_intr: Proof.context -> thm -> thm -> thm val conj_elim: Proof.context -> thm -> thm * thm val conj_elims: Proof.context -> thm -> thm list val conj: term val disj: term val imp: term val Not: term val mk_conj: term * term -> term val mk_disj: term * term -> term val mk_imp: term * term -> term val mk_not: term -> term val dest_conj: term -> term list val conjuncts: term -> term list val dest_disj: term -> term list val disjuncts: term -> term list val dest_imp: term -> term * term val dest_not: term -> term val conj_conv: conv -> conv -> conv val eq_const: typ -> term val mk_eq: term * term -> term val dest_eq: term -> term * term val eq_conv: conv -> conv -> conv val all_const: typ -> term val mk_all: string * typ * term -> term val list_all: (string * typ) list * term -> term val exists_const: typ -> term val mk_exists: string * typ * term -> term val choice_const: typ -> term val class_equal: string val mk_binop: string -> term * term -> term val mk_binrel: string -> term * term -> term val dest_bin: string -> typ -> term -> term * term val unitT: typ val is_unitT: typ -> bool val unit: term val is_unit: term -> bool val mk_prodT: typ * typ -> typ val dest_prodT: typ -> typ * typ val pair_const: typ -> typ -> term val mk_prod: term * term -> term val dest_prod: term -> term * term val mk_fst: term -> term val mk_snd: term -> term val case_prod_const: typ * typ * typ -> term val mk_case_prod: term -> term val flatten_tupleT: typ -> typ list val tupled_lambda: term -> term -> term val mk_tupleT: typ list -> typ val strip_tupleT: typ -> typ list val mk_tuple: term list -> term val strip_tuple: term -> term list val mk_ptupleT: int list list -> typ list -> typ val strip_ptupleT: int list list -> typ -> typ list val flat_tupleT_paths: typ -> int list list val mk_ptuple: int list list -> typ -> term list -> term val strip_ptuple: int list list -> term -> term list val flat_tuple_paths: term -> int list list val mk_ptupleabs: int list list -> typ -> typ -> term -> term val strip_ptupleabs: term -> term * typ list * int list list val natT: typ val zero: term val is_zero: term -> bool val mk_Suc: term -> term val dest_Suc: term -> term val Suc_zero: term val mk_nat: int -> term val dest_nat: term -> int val class_size: string val size_const: typ -> term val intT: typ val one_const: term val bit0_const: term val bit1_const: term val mk_numeral: int -> term val dest_numeral: term -> int val numeral_const: typ -> term val add_numerals: term -> (term * typ) list -> (term * typ) list val mk_number: typ -> int -> term val dest_number: term -> typ * int val code_integerT: typ val code_naturalT: typ val realT: typ val charT: typ val mk_char: int -> term val dest_char: term -> int val listT: typ -> typ val nil_const: typ -> term val cons_const: typ -> term val mk_list: typ -> term list -> term val dest_list: term -> term list val stringT: typ val mk_string: string -> term val dest_string: term -> string val literalT: typ val mk_literal: string -> term val dest_literal: term -> string val mk_typerep: typ -> term val termT: typ val term_of_const: typ -> term val mk_term_of: typ -> term -> term val reflect_term: term -> term val mk_valtermify_app: string -> (string * typ) list -> typ -> term val mk_return: typ -> typ -> term -> term val mk_ST: ((term * typ) * (string * typ) option) list -> term -> typ -> typ option * typ -> term val mk_random: typ -> term -> term end; structure HOLogic: HOLOGIC = struct (* functions *) fun id_const T = Const ("Fun.id", T --> T); fun mk_comp (f, g) = let val fT = fastype_of f; val gT = fastype_of g; val compT = fT --> gT --> domain_type gT --> range_type fT; in Const ("Fun.comp", compT) $ f $ g end; (* bool and set *) val boolN = "HOL.bool"; val boolT = Type (boolN, []); fun mk_induct_forall T = Const ("HOL.induct_forall", (T --> boolT) --> boolT); fun mk_setT T = Type ("Set.set", [T]); fun dest_setT (Type ("Set.set", [T])) = T | dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []); fun mk_set T ts = let val sT = mk_setT T; val empty = Const ("Orderings.bot_class.bot", sT); fun insert t u = Const ("Set.insert", T --> sT --> sT) $ t $ u; in fold_rev insert ts empty end; fun mk_UNIV T = Const ("Orderings.top_class.top", mk_setT T); fun dest_set (Const ("Orderings.bot_class.bot", _)) = [] | dest_set (Const ("Set.insert", _) $ t $ u) = t :: dest_set u | dest_set t = raise TERM ("dest_set", [t]); fun Collect_const T = Const ("Set.Collect", (T --> boolT) --> mk_setT T); fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T) t; fun mk_mem (x, A) = let val setT = fastype_of A in Const ("Set.member", dest_setT setT --> setT --> boolT) $ x $ A end; fun dest_mem (Const ("Set.member", _) $ x $ A) = (x, A) | dest_mem t = raise TERM ("dest_mem", [t]); (* logic *) fun mk_obj_eq th = th RS @{thm meta_eq_to_obj_eq}; val Trueprop = Const (\<^const_name>\Trueprop\, boolT --> propT); fun mk_Trueprop P = Trueprop $ P; fun dest_Trueprop (Const (\<^const_name>\Trueprop\, _) $ P) = P | dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); fun Trueprop_conv cv ct = (case Thm.term_of ct of Const (\<^const_name>\Trueprop\, _) $ _ => Conv.arg_conv cv ct | _ => raise CTERM ("Trueprop_conv", [ct])) fun conj_intr ctxt thP thQ = let val (P, Q) = apply2 (Object_Logic.dest_judgment ctxt o Thm.cprop_of) (thP, thQ) handle CTERM (msg, _) => raise THM (msg, 0, [thP, thQ]); - val inst = Thm.instantiate ([], [((("P", 0), boolT), P), ((("Q", 0), boolT), Q)]); - in Drule.implies_elim_list (inst @{thm conjI}) [thP, thQ] end; + val inst = (TVars.empty, Vars.make [((("P", 0), boolT), P), ((("Q", 0), boolT), Q)]); + in Drule.implies_elim_list (Thm.instantiate inst @{thm conjI}) [thP, thQ] end; fun conj_elim ctxt thPQ = let val (P, Q) = Thm.dest_binop (Object_Logic.dest_judgment ctxt (Thm.cprop_of thPQ)) handle CTERM (msg, _) => raise THM (msg, 0, [thPQ]); - val inst = Thm.instantiate ([], [((("P", 0), boolT), P), ((("Q", 0), boolT), Q)]); - val thP = Thm.implies_elim (inst @{thm conjunct1}) thPQ; - val thQ = Thm.implies_elim (inst @{thm conjunct2}) thPQ; + val inst = (TVars.empty, Vars.make [((("P", 0), boolT), P), ((("Q", 0), boolT), Q)]); + val thP = Thm.implies_elim (Thm.instantiate inst @{thm conjunct1}) thPQ; + val thQ = Thm.implies_elim (Thm.instantiate inst @{thm conjunct2}) thPQ; in (thP, thQ) end; fun conj_elims ctxt th = let val (th1, th2) = conj_elim ctxt th in conj_elims ctxt th1 @ conj_elims ctxt th2 end handle THM _ => [th]; val conj = \<^term>\HOL.conj\ and disj = \<^term>\HOL.disj\ and imp = \<^term>\implies\ and Not = \<^term>\Not\; fun mk_conj (t1, t2) = conj $ t1 $ t2 and mk_disj (t1, t2) = disj $ t1 $ t2 and mk_imp (t1, t2) = imp $ t1 $ t2 and mk_not t = Not $ t; fun dest_conj (Const ("HOL.conj", _) $ t $ t') = t :: dest_conj t' | dest_conj t = [t]; (*Like dest_conj, but flattens conjunctions however nested*) fun conjuncts_aux (Const ("HOL.conj", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs) | conjuncts_aux t conjs = t::conjs; fun conjuncts t = conjuncts_aux t []; fun dest_disj (Const ("HOL.disj", _) $ t $ t') = t :: dest_disj t' | dest_disj t = [t]; (*Like dest_disj, but flattens disjunctions however nested*) fun disjuncts_aux (Const ("HOL.disj", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs) | disjuncts_aux t disjs = t::disjs; fun disjuncts t = disjuncts_aux t []; fun dest_imp (Const ("HOL.implies", _) $ A $ B) = (A, B) | dest_imp t = raise TERM ("dest_imp", [t]); fun dest_not (Const ("HOL.Not", _) $ t) = t | dest_not t = raise TERM ("dest_not", [t]); fun conj_conv cv1 cv2 ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.conj\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct | _ => raise CTERM ("conj_conv", [ct])); fun eq_const T = Const (\<^const_name>\HOL.eq\, T --> T --> boolT); fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; fun dest_eq (Const (\<^const_name>\HOL.eq\, _) $ lhs $ rhs) = (lhs, rhs) | dest_eq t = raise TERM ("dest_eq", [t]) fun eq_conv cv1 cv2 ct = (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct | _ => raise CTERM ("eq_conv", [ct])); fun all_const T = Const ("HOL.All", (T --> boolT) --> boolT); fun mk_all (x, T, P) = all_const T $ absfree (x, T) P; fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t; fun exists_const T = Const ("HOL.Ex", (T --> boolT) --> boolT); fun mk_exists (x, T, P) = exists_const T $ absfree (x, T) P; fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T); val class_equal = "HOL.equal"; (* binary operations and relations *) fun mk_binop c (t, u) = let val T = fastype_of t in Const (c, T --> T --> T) $ t $ u end; fun mk_binrel c (t, u) = let val T = fastype_of t in Const (c, T --> T --> boolT) $ t $ u end; (*destruct the application of a binary operator. The dummyT case is a crude way of handling polymorphic operators.*) fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) = if c = c' andalso (T=T' orelse T=dummyT) then (t, u) else raise TERM ("dest_bin " ^ c, [tm]) | dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]); (* unit *) val unitT = Type ("Product_Type.unit", []); fun is_unitT (Type ("Product_Type.unit", [])) = true | is_unitT _ = false; val unit = Const ("Product_Type.Unity", unitT); fun is_unit (Const ("Product_Type.Unity", _)) = true | is_unit _ = false; (* prod *) fun mk_prodT (T1, T2) = Type ("Product_Type.prod", [T1, T2]); fun dest_prodT (Type ("Product_Type.prod", [T1, T2])) = (T1, T2) | dest_prodT T = raise TYPE ("dest_prodT", [T], []); fun pair_const T1 T2 = Const ("Product_Type.Pair", T1 --> T2 --> mk_prodT (T1, T2)); fun mk_prod (t1, t2) = let val T1 = fastype_of t1 and T2 = fastype_of t2 in pair_const T1 T2 $ t1 $ t2 end; fun dest_prod (Const ("Product_Type.Pair", _) $ t1 $ t2) = (t1, t2) | dest_prod t = raise TERM ("dest_prod", [t]); fun mk_fst p = let val pT = fastype_of p in Const ("Product_Type.prod.fst", pT --> fst (dest_prodT pT)) $ p end; fun mk_snd p = let val pT = fastype_of p in Const ("Product_Type.prod.snd", pT --> snd (dest_prodT pT)) $ p end; fun case_prod_const (A, B, C) = Const ("Product_Type.prod.case_prod", (A --> B --> C) --> mk_prodT (A, B) --> C); fun mk_case_prod t = (case Term.fastype_of t of T as (Type ("fun", [A, Type ("fun", [B, C])])) => Const ("Product_Type.prod.case_prod", T --> mk_prodT (A, B) --> C) $ t | _ => raise TERM ("mk_case_prod: bad body type", [t])); (*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*) fun flatten_tupleT (Type ("Product_Type.prod", [T1, T2])) = flatten_tupleT T1 @ flatten_tupleT T2 | flatten_tupleT T = [T]; (*abstraction over nested tuples*) fun tupled_lambda (x as Free _) b = lambda x b | tupled_lambda (x as Var _) b = lambda x b | tupled_lambda (Const ("Product_Type.Pair", _) $ u $ v) b = mk_case_prod (tupled_lambda u (tupled_lambda v b)) | tupled_lambda (Const ("Product_Type.Unity", _)) b = Abs ("x", unitT, b) | tupled_lambda t _ = raise TERM ("tupled_lambda: bad tuple", [t]); (* tuples with right-fold structure *) fun mk_tupleT [] = unitT | mk_tupleT Ts = foldr1 mk_prodT Ts; fun strip_tupleT (Type ("Product_Type.unit", [])) = [] | strip_tupleT (Type ("Product_Type.prod", [T1, T2])) = T1 :: strip_tupleT T2 | strip_tupleT T = [T]; fun mk_tuple [] = unit | mk_tuple ts = foldr1 mk_prod ts; fun strip_tuple (Const ("Product_Type.Unity", _)) = [] | strip_tuple (Const ("Product_Type.Pair", _) $ t1 $ t2) = t1 :: strip_tuple t2 | strip_tuple t = [t]; (* tuples with specific arities an "arity" of a tuple is a list of lists of integers, denoting paths to subterms that are pairs *) fun ptuple_err s = raise TERM (s ^ ": inconsistent use of nested products", []); fun mk_ptupleT ps = let fun mk p Ts = if member (op =) ps p then let val (T, Ts') = mk (1::p) Ts; val (U, Ts'') = mk (2::p) Ts' in (mk_prodT (T, U), Ts'') end else (hd Ts, tl Ts) in fst o mk [] end; fun strip_ptupleT ps = let fun factors p T = if member (op =) ps p then (case T of Type ("Product_Type.prod", [T1, T2]) => factors (1::p) T1 @ factors (2::p) T2 | _ => ptuple_err "strip_ptupleT") else [T] in factors [] end; val flat_tupleT_paths = let fun factors p (Type ("Product_Type.prod", [T1, T2])) = p :: factors (1::p) T1 @ factors (2::p) T2 | factors p _ = [] in factors [] end; fun mk_ptuple ps = let fun mk p T ts = if member (op =) ps p then (case T of Type ("Product_Type.prod", [T1, T2]) => let val (t, ts') = mk (1::p) T1 ts; val (u, ts'') = mk (2::p) T2 ts' in (pair_const T1 T2 $ t $ u, ts'') end | _ => ptuple_err "mk_ptuple") else (hd ts, tl ts) in fst oo mk [] end; fun strip_ptuple ps = let fun dest p t = if member (op =) ps p then (case t of Const ("Product_Type.Pair", _) $ t $ u => dest (1::p) t @ dest (2::p) u | _ => ptuple_err "strip_ptuple") else [t] in dest [] end; val flat_tuple_paths = let fun factors p (Const ("Product_Type.Pair", _) $ t $ u) = p :: factors (1::p) t @ factors (2::p) u | factors p _ = [] in factors [] end; (*In mk_ptupleabs ps S T u, term u expects separate arguments for the factors of S, with result type T. The call creates a new term expecting one argument of type S.*) fun mk_ptupleabs ps T T3 u = let fun ap ((p, T) :: pTs) = if member (op =) ps p then (case T of Type ("Product_Type.prod", [T1, T2]) => case_prod_const (T1, T2, map snd pTs ---> T3) $ ap ((1::p, T1) :: (2::p, T2) :: pTs) | _ => ptuple_err "mk_ptupleabs") else Abs ("x", T, ap pTs) | ap [] = let val k = length ps in list_comb (incr_boundvars (k + 1) u, map Bound (k downto 0)) end in ap [([], T)] end; val strip_ptupleabs = let fun strip [] qs Ts t = (t, rev Ts, qs) | strip (p :: ps) qs Ts (Const ("Product_Type.prod.case_prod", _) $ t) = strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t | strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t | strip (p :: ps) qs Ts t = strip ps qs (hd (binder_types (fastype_of1 (Ts, t))) :: Ts) (incr_boundvars 1 t $ Bound 0) in strip [[]] [] [] end; (* nat *) val natT = Type ("Nat.nat", []); val zero = Const ("Groups.zero_class.zero", natT); fun is_zero (Const ("Groups.zero_class.zero", _)) = true | is_zero _ = false; fun mk_Suc t = Const ("Nat.Suc", natT --> natT) $ t; fun dest_Suc (Const ("Nat.Suc", _) $ t) = t | dest_Suc t = raise TERM ("dest_Suc", [t]); val Suc_zero = mk_Suc zero; fun mk_nat n = let fun mk 0 = zero | mk n = mk_Suc (mk (n - 1)); in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end; fun dest_nat (Const ("Groups.zero_class.zero", _)) = 0 | dest_nat (Const ("Nat.Suc", _) $ t) = dest_nat t + 1 | dest_nat t = raise TERM ("dest_nat", [t]); val class_size = "Nat.size"; fun size_const T = Const ("Nat.size_class.size", T --> natT); (* binary numerals and int *) val numT = Type ("Num.num", []); val intT = Type ("Int.int", []); val one_const = Const ("Num.num.One", numT) and bit0_const = Const ("Num.num.Bit0", numT --> numT) and bit1_const = Const ("Num.num.Bit1", numT --> numT); fun mk_numeral i = let fun mk 1 = one_const | mk i = let val (q, r) = Integer.div_mod i 2 in (if r = 0 then bit0_const else bit1_const) $ mk q end in if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, []) end fun dest_numeral (Const ("Num.num.One", _)) = 1 | dest_numeral (Const ("Num.num.Bit0", _) $ bs) = 2 * dest_numeral bs | dest_numeral (Const ("Num.num.Bit1", _) $ bs) = 2 * dest_numeral bs + 1 | dest_numeral t = raise TERM ("dest_num", [t]); fun numeral_const T = Const ("Num.numeral_class.numeral", numT --> T); fun add_numerals (Const ("Num.numeral_class.numeral", Type (_, [_, T])) $ t) = cons (t, T) | add_numerals (t $ u) = add_numerals t #> add_numerals u | add_numerals (Abs (_, _, t)) = add_numerals t | add_numerals _ = I; fun mk_number T 0 = Const ("Groups.zero_class.zero", T) | mk_number T 1 = Const ("Groups.one_class.one", T) | mk_number T i = if i > 0 then numeral_const T $ mk_numeral i else Const ("Groups.uminus_class.uminus", T --> T) $ mk_number T (~ i); fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0) | dest_number (Const ("Groups.one_class.one", T)) = (T, 1) | dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) = (T, dest_numeral t) | dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", _)) $ t) = apsnd (op ~) (dest_number t) | dest_number t = raise TERM ("dest_number", [t]); (* code target numerals *) val code_integerT = Type ("Code_Numeral.integer", []); val code_naturalT = Type ("Code_Numeral.natural", []); (* real *) val realT = Type ("Real.real", []); (* list *) fun listT T = Type ("List.list", [T]); fun nil_const T = Const ("List.list.Nil", listT T); fun cons_const T = let val lT = listT T in Const ("List.list.Cons", T --> lT --> lT) end; fun mk_list T ts = let val lT = listT T; val Nil = Const ("List.list.Nil", lT); fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u; in fold_rev Cons ts Nil end; fun dest_list (Const ("List.list.Nil", _)) = [] | dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u | dest_list t = raise TERM ("dest_list", [t]); (* booleans as bits *) fun mk_bit b = if b = 0 then \<^term>\False\ else if b = 1 then \<^term>\True\ else raise TERM ("mk_bit", []); fun mk_bits len = map mk_bit o Integer.radicify 2 len; fun dest_bit \<^term>\False\ = 0 | dest_bit \<^term>\True\ = 1 | dest_bit _ = raise TERM ("dest_bit", []); val dest_bits = Integer.eval_radix 2 o map dest_bit; (* char *) val charT = Type ("String.char", []); val Char_const = Const ("String.char.Char", replicate 8 boolT ---> charT); fun mk_char i = if 0 <= i andalso i <= 255 then list_comb (Char_const, mk_bits 8 i) else raise TERM ("mk_char", []) fun dest_char (Const ("String.char.Char", _) $ b0 $ b1 $ b2 $ b3 $ b4 $ b5 $ b6 $ b7) = dest_bits [b0, b1, b2, b3, b4, b5, b6, b7] | dest_char t = raise TERM ("dest_char", [t]); (* string *) val stringT = listT charT; val mk_string = mk_list charT o map (mk_char o ord) o raw_explode; val dest_string = implode o map (chr o dest_char) o dest_list; (* literal *) val literalT = Type ("String.literal", []); val Literal_const = Const ("String.Literal", replicate 7 boolT ---> literalT --> literalT); fun mk_literal s = let fun mk [] = Const ("Groups.zero_class.zero", literalT) | mk (c :: cs) = list_comb (Literal_const, mk_bits 7 c) $ mk cs; val cs = map ord (raw_explode s); in if forall (fn c => c < 128) cs then mk cs else raise TERM ("mk_literal", []) end; val dest_literal = let fun dest (Const ("Groups.zero_class.zero", Type ("String.literal", []))) = [] | dest (Const ("String.Literal", _) $ b0 $ b1 $ b2 $ b3 $ b4 $ b5 $ b6 $ t) = chr (dest_bits [b0, b1, b2, b3, b4, b5, b6]) :: dest t | dest t = raise TERM ("dest_literal", [t]); in implode o dest end; (* typerep and term *) val typerepT = Type ("Typerep.typerep", []); fun mk_typerep (Type (tyco, Ts)) = Const ("Typerep.typerep.Typerep", literalT --> listT typerepT --> typerepT) $ mk_literal tyco $ mk_list typerepT (map mk_typerep Ts) | mk_typerep (T as TFree _) = Const ("Typerep.typerep_class.typerep", Term.itselfT T --> typerepT) $ Logic.mk_type T; val termT = Type ("Code_Evaluation.term", []); fun term_of_const T = Const ("Code_Evaluation.term_of_class.term_of", T --> termT); fun mk_term_of T t = term_of_const T $ t; fun reflect_term (Const (c, T)) = Const ("Code_Evaluation.Const", literalT --> typerepT --> termT) $ mk_literal c $ mk_typerep T | reflect_term (t1 $ t2) = Const ("Code_Evaluation.App", termT --> termT --> termT) $ reflect_term t1 $ reflect_term t2 | reflect_term (Abs (v, _, t)) = Abs (v, termT, reflect_term t) | reflect_term t = t; fun mk_valtermify_app c vs T = let fun termifyT T = mk_prodT (T, unitT --> termT); fun valapp T T' = Const ("Code_Evaluation.valapp", termifyT (T --> T') --> termifyT T --> termifyT T'); fun mk_fTs [] _ = [] | mk_fTs (_ :: Ts) T = (Ts ---> T) :: mk_fTs Ts T; val Ts = map snd vs; val t = Const (c, Ts ---> T); val tt = mk_prod (t, Abs ("u", unitT, reflect_term t)); fun app (fT, (v, T)) t = valapp T fT $ t $ Free (v, termifyT T); in fold app (mk_fTs Ts T ~~ vs) tt end; (* open state monads *) fun mk_return T U x = pair_const T U $ x; fun mk_ST clauses t U (someT, V) = let val R = case someT of SOME T => mk_prodT (T, V) | NONE => V fun mk_clause ((t, U), SOME (v, T)) (t', U') = (Const ("Product_Type.scomp", (U --> mk_prodT (T, U')) --> (T --> U' --> R) --> U --> R) $ t $ lambda (Free (v, T)) t', U) | mk_clause ((t, U), NONE) (t', U') = (Const ("Product_Type.fcomp", (U --> U') --> (U' --> R) --> U --> R) $ t $ t', U) in fold_rev mk_clause clauses (t, U) |> fst end; (* random seeds *) val random_seedT = mk_prodT (code_naturalT, code_naturalT); fun mk_random T t = Const ("Quickcheck_Random.random_class.random", code_naturalT --> random_seedT --> mk_prodT (mk_prodT (T, unitT --> termT), random_seedT)) $ t; end; diff --git a/src/HOL/Tools/inductive_set.ML b/src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML +++ b/src/HOL/Tools/inductive_set.ML @@ -1,561 +1,561 @@ (* Title: HOL/Tools/inductive_set.ML Author: Stefan Berghofer, TU Muenchen Wrapper for defining inductive sets using package for inductive predicates, including infrastructure for converting between predicates and sets. *) signature INDUCTIVE_SET = sig val to_set_att: thm list -> attribute val to_pred_att: thm list -> attribute val to_pred : thm list -> Context.generic -> thm -> thm val pred_set_conv_att: attribute val add_inductive: Inductive.flags -> ((binding * typ) * mixfix) list -> (string * typ) list -> (Attrib.binding * term) list -> thm list -> local_theory -> Inductive.result * local_theory val add_inductive_cmd: bool -> bool -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> Specification.multi_specs_cmd -> (Facts.ref * Token.src list) list -> local_theory -> Inductive.result * local_theory val mono_add: attribute val mono_del: attribute end; structure Inductive_Set: INDUCTIVE_SET = struct (***********************************************************************************) (* simplifies (%x y. (x, y) : S & P x y) to (%x y. (x, y) : S Int {(x, y). P x y}) *) (* and (%x y. (x, y) : S | P x y) to (%x y. (x, y) : S Un {(x, y). P x y}) *) (* used for converting "strong" (co)induction rules *) (***********************************************************************************) val anyt = Free ("t", TFree ("'t", [])); fun strong_ind_simproc tab = Simplifier.make_simproc \<^context> "strong_ind" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let fun close p t f = let val vs = Term.add_vars t [] in Thm.instantiate' [] (rev (map (SOME o Thm.cterm_of ctxt o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; fun mkop \<^const_name>\HOL.conj\ T x = SOME (Const (\<^const_name>\Lattices.inf\, T --> T --> T), x) | mkop \<^const_name>\HOL.disj\ T x = SOME (Const (\<^const_name>\Lattices.sup\, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T in HOLogic.Collect_const U $ HOLogic.mk_ptupleabs (HOLogic.flat_tuple_paths p) U HOLogic.boolT t end; fun decomp (Const (s, _) $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S) $ u) = mkop s T (m, p, S, mk_collect p T (head_of u)) | decomp (Const (s, _) $ u $ ((m as Const (\<^const_name>\Set.member\, Type (_, [_, Type (_, [T, _])]))) $ p $ S)) = mkop s T (m, p, mk_collect p T (head_of u), S) | decomp _ = NONE; val simp = full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv}) 1; fun mk_rew t = (case strip_abs_vars t of [] => NONE | xs => (case decomp (strip_abs_body t) of NONE => NONE | SOME (bop, (m, p, S, S')) => SOME (close (Goal.prove ctxt [] []) (Logic.mk_equals (t, fold_rev Term.abs xs (m $ p $ (bop $ S $ S')))) (K (EVERY [resolve_tac ctxt [eq_reflection] 1, REPEAT (resolve_tac ctxt @{thms ext} 1), resolve_tac ctxt @{thms iffI} 1, EVERY [eresolve_tac ctxt @{thms conjE} 1, resolve_tac ctxt @{thms IntI} 1, simp, simp, eresolve_tac ctxt @{thms IntE} 1, resolve_tac ctxt @{thms conjI} 1, simp, simp] ORELSE EVERY [eresolve_tac ctxt @{thms disjE} 1, resolve_tac ctxt @{thms UnI1} 1, simp, resolve_tac ctxt @{thms UnI2} 1, simp, eresolve_tac ctxt @{thms UnE} 1, resolve_tac ctxt @{thms disjI1} 1, simp, resolve_tac ctxt @{thms disjI2} 1, simp]]))) handle ERROR _ => NONE)) in (case strip_comb (Thm.term_of ct) of (h as Const (name, _), ts) => if Symtab.defined tab name then let val rews = map mk_rew ts in if forall is_none rews then NONE else SOME (fold (fn th1 => fn th2 => Thm.combination th2 th1) (map2 (fn SOME r => K r | NONE => Thm.reflexive o Thm.cterm_of ctxt) rews ts) (Thm.reflexive (Thm.cterm_of ctxt h))) end else NONE | _ => NONE) end}; (* only eta contract terms occurring as arguments of functions satisfying p *) fun eta_contract p = let fun eta b (Abs (a, T, body)) = (case eta b body of body' as (f $ Bound 0) => if Term.is_dependent f orelse not b then Abs (a, T, body') else incr_boundvars ~1 f | body' => Abs (a, T, body')) | eta b (t $ u) = eta b t $ eta (p (head_of t)) u | eta b t = t in eta false end; fun eta_contract_thm ctxt p = Conv.fconv_rule (Conv.then_conv (Thm.beta_conversion true, fn ct => Thm.transitive (Thm.eta_conversion ct) (Thm.symmetric (Thm.eta_conversion (Thm.cterm_of ctxt (eta_contract p (Thm.term_of ct))))))); (***********************************************************) (* rules for converting between predicate and set notation *) (* *) (* rules for converting predicates to sets have the form *) (* P (%x y. (x, y) : s) = (%x y. (x, y) : S s) *) (* *) (* rules for converting sets to predicates have the form *) (* S {(x, y). p x y} = {(x, y). P p x y} *) (* *) (* where s and p are parameters *) (***********************************************************) structure Data = Generic_Data ( type T = {(* rules for converting predicates to sets *) to_set_simps: thm list, (* rules for converting sets to predicates *) to_pred_simps: thm list, (* arities of functions of type t set => ... => u set *) set_arities: (typ * (int list list option list * int list list option)) list Symtab.table, (* arities of functions of type (t => ... => bool) => u => ... => bool *) pred_arities: (typ * (int list list option list * int list list option)) list Symtab.table}; val empty = {to_set_simps = [], to_pred_simps = [], set_arities = Symtab.empty, pred_arities = Symtab.empty}; val extend = I; fun merge ({to_set_simps = to_set_simps1, to_pred_simps = to_pred_simps1, set_arities = set_arities1, pred_arities = pred_arities1}, {to_set_simps = to_set_simps2, to_pred_simps = to_pred_simps2, set_arities = set_arities2, pred_arities = pred_arities2}) : T = {to_set_simps = Thm.merge_thms (to_set_simps1, to_set_simps2), to_pred_simps = Thm.merge_thms (to_pred_simps1, to_pred_simps2), set_arities = Symtab.merge_list (op =) (set_arities1, set_arities2), pred_arities = Symtab.merge_list (op =) (pred_arities1, pred_arities2)}; ); fun name_type_of (Free p) = SOME p | name_type_of (Const p) = SOME p | name_type_of _ = NONE; fun map_type f (Free (s, T)) = Free (s, f T) | map_type f (Var (ixn, T)) = Var (ixn, f T) | map_type f _ = error "map_type"; fun find_most_specific is_inst f eq xs T = find_first (fn U => is_inst (T, f U) andalso forall (fn U' => eq (f U, f U') orelse not (is_inst (T, f U') andalso is_inst (f U', f U))) xs) xs; fun lookup_arity thy arities (s, T) = case Symtab.lookup arities s of NONE => NONE | SOME xs => find_most_specific (Sign.typ_instance thy) fst (op =) xs T; fun lookup_rule thy f rules = find_most_specific (swap #> Pattern.matches thy) (f #> fst) (op aconv) rules; fun infer_arities thy arities (optf, t) fs = case strip_comb t of (Abs (_, _, u), []) => infer_arities thy arities (NONE, u) fs | (Abs _, _) => infer_arities thy arities (NONE, Envir.beta_norm t) fs | (u, ts) => (case Option.map (lookup_arity thy arities) (name_type_of u) of SOME (SOME (_, (arity, _))) => (fold (infer_arities thy arities) (arity ~~ List.take (ts, length arity)) fs handle General.Subscript => error "infer_arities: bad term") | _ => fold (infer_arities thy arities) (map (pair NONE) ts) (case optf of NONE => fs | SOME f => AList.update op = (u, the_default f (Option.map (fn g => inter (op =) g f) (AList.lookup op = fs u))) fs)); (**************************************************************) (* derive the to_pred equation from the to_set equation *) (* *) (* 1. instantiate each set parameter with {(x, y). p x y} *) (* 2. apply %P. {(x, y). P x y} to both sides of the equation *) (* 3. simplify *) (**************************************************************) fun mk_to_pred_inst ctxt fs = map (fn (x, ps) => let val (Ts, T) = strip_type (fastype_of x); val U = HOLogic.dest_setT T; val x' = map_type (K (Ts @ HOLogic.strip_ptupleT ps U ---> HOLogic.boolT)) x; in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs ps U HOLogic.boolT (list_comb (x', map Bound (length Ts - 1 downto 0)))))) end) fs; fun mk_to_pred_eq ctxt p fs optfs' T thm = let val insts = mk_to_pred_inst ctxt fs; - val thm' = Thm.instantiate ([], insts) thm; + val thm' = Thm.instantiate (TVars.empty, Vars.make insts) thm; val thm'' = (case optfs' of NONE => thm' RS sym | SOME fs' => let val U = HOLogic.dest_setT (body_type T); val Ts = HOLogic.strip_ptupleT fs' U; val arg_cong' = Thm.incr_indexes (Thm.maxidx_of thm + 1) arg_cong; val (Var (arg_cong_f, _), _) = arg_cong' |> Thm.concl_of |> dest_comb |> snd |> strip_comb |> snd |> hd |> dest_comb; in thm' RS (infer_instantiate ctxt [(arg_cong_f, Thm.cterm_of ctxt (Abs ("P", Ts ---> HOLogic.boolT, HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs' U HOLogic.boolT (Bound 0))))] arg_cong' RS sym) end) in Simplifier.simplify (put_simpset HOL_basic_ss ctxt addsimps @{thms mem_Collect_eq case_prod_conv} addsimprocs [\<^simproc>\Collect_mem\]) thm'' |> zero_var_indexes |> eta_contract_thm ctxt (equal p) end; (**** declare rules for converting predicates to sets ****) exception Malformed of string; fun add context thm (tab as {to_set_simps, to_pred_simps, set_arities, pred_arities}) = (case Thm.prop_of thm of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ lhs $ rhs) => (case body_type T of \<^typ>\bool\ => let val thy = Context.theory_of context; val ctxt = Context.proof_of context; fun factors_of t fs = case strip_abs_body t of Const (\<^const_name>\Set.member\, _) $ u $ S => if is_Free S orelse is_Var S then let val ps = HOLogic.flat_tuple_paths u in (SOME ps, (S, ps) :: fs) end else (NONE, fs) | _ => (NONE, fs); val (h, ts) = strip_comb lhs val (pfs, fs) = fold_map factors_of ts []; val ((h', ts'), fs') = (case rhs of Abs _ => (case strip_abs_body rhs of Const (\<^const_name>\Set.member\, _) $ u $ S => (strip_comb S, SOME (HOLogic.flat_tuple_paths u)) | _ => raise Malformed "member symbol on right-hand side expected") | _ => (strip_comb rhs, NONE)) in case (name_type_of h, name_type_of h') of (SOME (s, T), SOME (s', T')) => if exists (fn (U, _) => Sign.typ_instance thy (T', U) andalso Sign.typ_instance thy (U, T')) (Symtab.lookup_list set_arities s') then (if Context_Position.is_really_visible ctxt then warning ("Ignoring conversion rule for operator " ^ s') else (); tab) else {to_set_simps = Thm.trim_context thm :: to_set_simps, to_pred_simps = Thm.trim_context (mk_to_pred_eq ctxt h fs fs' T' thm) :: to_pred_simps, set_arities = Symtab.insert_list op = (s', (T', (map (AList.lookup op = fs) ts', fs'))) set_arities, pred_arities = Symtab.insert_list op = (s, (T, (pfs, fs'))) pred_arities} | _ => raise Malformed "set / predicate constant expected" end | _ => raise Malformed "equation between predicates expected") | _ => raise Malformed "equation expected") handle Malformed msg => let val ctxt = Context.proof_of context val _ = if Context_Position.is_really_visible ctxt then warning ("Ignoring malformed set / predicate conversion rule: " ^ msg ^ "\n" ^ Thm.string_of_thm ctxt thm) else (); in tab end; val pred_set_conv_att = Thm.declaration_attribute (fn thm => fn ctxt => Data.map (add ctxt thm) ctxt); (**** convert theorem in set notation to predicate notation ****) fun is_pred tab t = case Option.map (Symtab.lookup tab o fst) (name_type_of t) of SOME (SOME _) => true | _ => false; fun to_pred_simproc rules = let val rules' = map mk_meta_eq rules in Simplifier.make_simproc \<^context> "to_pred" {lhss = [anyt], proc = fn _ => fn ctxt => fn ct => lookup_rule (Proof_Context.theory_of ctxt) (Thm.prop_of #> Logic.dest_equals) rules' (Thm.term_of ct)} end; fun to_pred_proc thy rules t = case lookup_rule thy I rules t of NONE => NONE | SOME (lhs, rhs) => SOME (Envir.subst_term (Pattern.match thy (lhs, t) (Vartab.empty, Vartab.empty)) rhs); fun to_pred thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_pred_simps, set_arities, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy set_arities (NONE, Thm.prop_of thm) []); (* instantiate each set parameter with {(x, y). p x y} *) val insts = mk_to_pred_inst ctxt fs in thm |> - Thm.instantiate ([], insts) |> + Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimprocs [to_pred_simproc (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: map (Thm.transfer thy) to_pred_simps)]) |> eta_contract_thm ctxt (is_pred pred_arities) |> Rule_Cases.save thm end; val to_pred_att = Thm.rule_attribute [] o to_pred; (**** convert theorem in predicate notation to set notation ****) fun to_set thms context thm = let val thy = Context.theory_of context; val ctxt = Context.proof_of context; val {to_set_simps, pred_arities, ...} = fold (add context) thms (Data.get context); val fs = filter (is_Var o fst) (infer_arities thy pred_arities (NONE, Thm.prop_of thm) []); (* instantiate each predicate parameter with %x y. (x, y) : s *) val insts = map (fn (x, ps) => let val Ts = binder_types (fastype_of x); val l = length Ts; val k = length ps; val (Rs, Us) = chop (l - k - 1) Ts; val T = HOLogic.mk_ptupleT ps Us; val x' = map_type (K (Rs ---> HOLogic.mk_setT T)) x in (dest_Var x, Thm.cterm_of ctxt (fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple ps T (map Bound (k downto 0)), list_comb (x', map Bound (l - 1 downto k + 1)))))) end) fs; in thm |> - Thm.instantiate ([], insts) |> + Thm.instantiate (TVars.empty, Vars.make insts) |> Simplifier.full_simplify (put_simpset HOL_basic_ss ctxt addsimps to_set_simps addsimprocs [strong_ind_simproc pred_arities, \<^simproc>\Collect_mem\]) |> Rule_Cases.save thm end; val to_set_att = Thm.rule_attribute [] o to_set; (**** definition of inductive sets ****) fun add_ind_set_def {quiet_mode, verbose, alt_name, coind, no_elim, no_ind, skip_mono} cs intros monos params cnames_syn lthy = let val thy = Proof_Context.theory_of lthy; val {set_arities, pred_arities, to_pred_simps, ...} = Data.get (Context.Proof lthy); fun infer (Abs (_, _, t)) = infer t | infer (Const (\<^const_name>\Set.member\, _) $ t $ u) = infer_arities thy set_arities (SOME (HOLogic.flat_tuple_paths t), u) | infer (t $ u) = infer t #> infer u | infer _ = I; val new_arities = filter_out (fn (x as Free (_, T), _) => member (op =) params x andalso length (binder_types T) > 0 | _ => false) (fold (snd #> infer) intros []); val params' = map (fn x => (case AList.lookup op = new_arities x of SOME fs => let val T = HOLogic.dest_setT (fastype_of x); val Ts = HOLogic.strip_ptupleT fs T; val x' = map_type (K (Ts ---> HOLogic.boolT)) x in (x, (x', (HOLogic.Collect_const T $ HOLogic.mk_ptupleabs fs T HOLogic.boolT x', fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs T (map Bound (length fs downto 0)), x))))) end | NONE => (x, (x, (x, x))))) params; val (params1, (params2, params3)) = params' |> map snd |> split_list ||> split_list; val paramTs = map fastype_of params; (* equations for converting sets to predicates *) val ((cs', cs_info), eqns) = cs |> map (fn c as Free (s, T) => let val fs = the_default [] (AList.lookup op = new_arities c); val (Us, U) = strip_type T |> apsnd HOLogic.dest_setT; val _ = Us = paramTs orelse error (Pretty.string_of (Pretty.chunks [Pretty.str "Argument types", Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) Us)), Pretty.str ("of " ^ s ^ " do not agree with types"), Pretty.block (Pretty.commas (map (Syntax.pretty_typ lthy) paramTs)), Pretty.str "of declared parameters"])); val Ts = HOLogic.strip_ptupleT fs U; val c' = Free (s ^ "p", map fastype_of params1 @ Ts ---> HOLogic.boolT) in ((c', (fs, U, Ts)), (list_comb (c, params2), HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (c', params1)))) end) |> split_list |>> split_list; val eqns' = eqns @ map (Thm.prop_of #> HOLogic.dest_Trueprop #> HOLogic.dest_eq) (@{thm mem_Collect_eq} :: @{thm case_prod_conv} :: to_pred_simps); (* predicate version of the introduction rules *) val intros' = map (fn (name_atts, t) => (name_atts, t |> map_aterms (fn u => (case AList.lookup op = params' u of SOME (_, (u', _)) => u' | NONE => u)) |> Pattern.rewrite_term thy [] [to_pred_proc thy eqns'] |> eta_contract (member op = cs' orf is_pred pred_arities))) intros; val cnames_syn' = map (fn (b, _) => (Binding.suffix_name "p" b, NoSyn)) cnames_syn; val monos' = map (to_pred [] (Context.Proof lthy)) monos; val ({preds, intrs, elims, raw_induct, eqs, ...}, lthy1) = Inductive.add_ind_def {quiet_mode = quiet_mode, verbose = verbose, alt_name = Binding.empty, coind = coind, no_elim = no_elim, no_ind = no_ind, skip_mono = skip_mono} cs' intros' monos' params1 cnames_syn' lthy; (* define inductive sets using previously defined predicates *) val (defs, lthy2) = lthy1 |> fold_map Local_Theory.define (map (fn (((b, mx), (fs, U, _)), p) => ((b, mx), ((Thm.def_binding b, []), fold_rev lambda params (HOLogic.Collect_const U $ HOLogic.mk_ptupleabs fs U HOLogic.boolT (list_comb (p, params3)))))) (cnames_syn ~~ cs_info ~~ preds)); (* prove theorems for converting predicate to set notation *) val lthy3 = fold (fn (((p, c as Free (s, _)), (fs, U, Ts)), (_, (_, def))) => fn lthy => let val conv_thm = Goal.prove lthy (map (fst o dest_Free) params) [] (HOLogic.mk_Trueprop (HOLogic.mk_eq (list_comb (p, params3), fold_rev (Term.abs o pair "x") Ts (HOLogic.mk_mem (HOLogic.mk_ptuple fs U (map Bound (length fs downto 0)), list_comb (c, params)))))) (K (REPEAT (resolve_tac lthy @{thms ext} 1) THEN simp_tac (put_simpset HOL_basic_ss lthy addsimps [def, @{thm mem_Collect_eq}, @{thm case_prod_conv}]) 1)) in lthy |> Local_Theory.note ((Binding.name (s ^ "p_" ^ s ^ "_eq"), [Attrib.internal (K pred_set_conv_att)]), [conv_thm]) |> snd end) (preds ~~ cs ~~ cs_info ~~ defs) lthy2; (* convert theorems to set notation *) val rec_name = if Binding.is_empty alt_name then Binding.conglomerate (map #1 cnames_syn) else alt_name; val cnames = map (Local_Theory.full_name lthy3 o #1) cnames_syn; (* FIXME *) val spec_name = Binding.conglomerate (map #1 cnames_syn); val (intr_names, intr_atts) = split_list (map fst intros); val raw_induct' = to_set [] (Context.Proof lthy3) raw_induct; val (intrs', elims', eqs', induct, inducts, lthy4) = Inductive.declare_rules rec_name coind no_ind spec_name cnames (map fst defs) (map (to_set [] (Context.Proof lthy3)) intrs) intr_names intr_atts (map (fn th => (to_set [] (Context.Proof lthy3) th, map (fst o fst) (fst (Rule_Cases.get th)), Rule_Cases.get_constraints th)) elims) (map (to_set [] (Context.Proof lthy3)) eqs) raw_induct' lthy3; in ({intrs = intrs', elims = elims', induct = induct, inducts = inducts, raw_induct = raw_induct', preds = map fst defs, eqs = eqs'}, lthy4) end; val add_inductive = Inductive.gen_add_inductive add_ind_set_def; val add_inductive_cmd = Inductive.gen_add_inductive_cmd add_ind_set_def; fun mono_att att = Thm.declaration_attribute (fn thm => fn context => Thm.attribute_declaration att (to_pred [] context thm) context); val mono_add = mono_att Inductive.mono_add; val mono_del = mono_att Inductive.mono_del; (** package setup **) (* attributes *) val _ = Theory.setup (Attrib.setup \<^binding>\pred_set_conv\ (Scan.succeed pred_set_conv_att) "declare rules for converting between predicate and set notation" #> Attrib.setup \<^binding>\to_set\ (Attrib.thms >> to_set_att) "convert rule to set notation" #> Attrib.setup \<^binding>\to_pred\ (Attrib.thms >> to_pred_att) "convert rule to predicate notation" #> Attrib.setup \<^binding>\mono_set\ (Attrib.add_del mono_add mono_del) "declare of monotonicity rule for set operators"); (* commands *) val ind_set_decl = Inductive.gen_ind_decl add_ind_set_def; val _ = Outer_Syntax.local_theory \<^command_keyword>\inductive_set\ "define inductive sets" (ind_set_decl false); val _ = Outer_Syntax.local_theory \<^command_keyword>\coinductive_set\ "define coinductive sets" (ind_set_decl true); end; diff --git a/src/HOL/Tools/lin_arith.ML b/src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML +++ b/src/HOL/Tools/lin_arith.ML @@ -1,976 +1,977 @@ (* Title: HOL/Tools/lin_arith.ML Author: Tjark Weber and Tobias Nipkow, TU Muenchen HOL setup for linear arithmetic (see Provers/Arith/fast_lin_arith.ML). *) signature LIN_ARITH = sig val pre_tac: Proof.context -> int -> tactic val simple_tac: Proof.context -> int -> tactic val tac: Proof.context -> int -> tactic val simproc: Proof.context -> cterm -> thm option val add_inj_thms: thm list -> Context.generic -> Context.generic val add_lessD: thm -> Context.generic -> Context.generic val add_simps: thm list -> Context.generic -> Context.generic val add_simprocs: simproc list -> Context.generic -> Context.generic val add_inj_const: string * typ -> Context.generic -> Context.generic val add_discrete_type: string -> Context.generic -> Context.generic val set_number_of: (Proof.context -> typ -> int -> cterm) -> Context.generic -> Context.generic val global_setup: theory -> theory val init_arith_data: Context.generic -> Context.generic val split_limit: int Config.T val neq_limit: int Config.T val trace: bool Config.T end; structure Lin_Arith: LIN_ARITH = struct (* Parameters data for general linear arithmetic functor *) structure LA_Logic: LIN_ARITH_LOGIC = struct val ccontr = @{thm ccontr}; val conjI = conjI; val notI = notI; val sym = sym; val trueI = TrueI; val not_lessD = @{thm linorder_not_less} RS iffD1; val not_leD = @{thm linorder_not_le} RS iffD1; fun mk_Eq thm = thm RS @{thm Eq_FalseI} handle THM _ => thm RS @{thm Eq_TrueI}; val mk_Trueprop = HOLogic.mk_Trueprop; fun atomize thm = case Thm.prop_of thm of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\HOL.conj\, _) $ _ $ _) => atomize (thm RS conjunct1) @ atomize (thm RS conjunct2) | _ => [thm]; fun neg_prop ((TP as Const(\<^const_name>\Trueprop\, _)) $ (Const (\<^const_name>\Not\, _) $ t)) = TP $ t | neg_prop ((TP as Const(\<^const_name>\Trueprop\, _)) $ t) = TP $ (HOLogic.Not $t) | neg_prop t = raise TERM ("neg_prop", [t]); fun is_False thm = let val _ $ t = Thm.prop_of thm in t = \<^term>\False\ end; fun is_nat t = (fastype_of1 t = HOLogic.natT); fun mk_nat_thm thy t = - let val ct = Thm.global_cterm_of thy t - in Drule.instantiate_normalize ([], [((("n", 0), HOLogic.natT), ct)]) @{thm le0} end; + let val ct = Thm.global_cterm_of thy t in + Drule.instantiate_normalize (TVars.empty, Vars.make [((("n", 0), HOLogic.natT), ct)]) @{thm le0} + end; end; (* arith context data *) structure Lin_Arith_Data = Generic_Data ( type T = {splits: thm list, inj_consts: (string * typ) list, discrete: string list}; val empty = {splits = [], inj_consts = [], discrete = []}; val extend = I; fun merge ({splits = splits1, inj_consts = inj_consts1, discrete = discrete1}, {splits = splits2, inj_consts = inj_consts2, discrete = discrete2}) : T = {splits = Thm.merge_thms (splits1, splits2), inj_consts = Library.merge (op =) (inj_consts1, inj_consts2), discrete = Library.merge (op =) (discrete1, discrete2)}; ); val get_arith_data = Lin_Arith_Data.get o Context.Proof; fun get_splits ctxt = #splits (get_arith_data ctxt) |> map (Thm.transfer' ctxt); fun add_split thm = Lin_Arith_Data.map (fn {splits, inj_consts, discrete} => {splits = update Thm.eq_thm_prop (Thm.trim_context thm) splits, inj_consts = inj_consts, discrete = discrete}); fun add_discrete_type d = Lin_Arith_Data.map (fn {splits, inj_consts, discrete} => {splits = splits, inj_consts = inj_consts, discrete = update (op =) d discrete}); fun add_inj_const c = Lin_Arith_Data.map (fn {splits, inj_consts, discrete} => {splits = splits, inj_consts = update (op =) c inj_consts, discrete = discrete}); val split_limit = Attrib.setup_config_int \<^binding>\linarith_split_limit\ (K 9); val neq_limit = Attrib.setup_config_int \<^binding>\linarith_neq_limit\ (K 9); val trace = Attrib.setup_config_bool \<^binding>\linarith_trace\ (K false); structure LA_Data: LIN_ARITH_DATA = struct val neq_limit = neq_limit; val trace = trace; (* Decomposition of terms *) (*internal representation of linear (in-)equations*) type decomp = ((term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat * bool); fun nT (Type ("fun", [N, _])) = (N = HOLogic.natT) | nT _ = false; fun add_atom (t : term) (m : Rat.rat) (p : (term * Rat.rat) list, i : Rat.rat) : (term * Rat.rat) list * Rat.rat = case AList.lookup Envir.aeconv p t of NONE => ((t, m) :: p, i) | SOME n => (AList.update Envir.aeconv (t, Rat.add n m) p, i); (* decompose nested multiplications, bracketing them to the right and combining all their coefficients inj_consts: list of constants to be ignored when encountered (e.g. arithmetic type conversions that preserve value) m: multiplicity associated with the entire product returns either (SOME term, associated multiplicity) or (NONE, constant) *) fun of_field_sort thy U = Sign.of_sort thy (U, \<^sort>\inverse\); fun demult thy (inj_consts : (string * typ) list) : term * Rat.rat -> term option * Rat.rat = let fun demult ((mC as Const (\<^const_name>\Groups.times\, _)) $ s $ t, m) = (case s of Const (\<^const_name>\Groups.times\, _) $ s1 $ s2 => (* bracketing to the right: '(s1 * s2) * t' becomes 's1 * (s2 * t)' *) demult (mC $ s1 $ (mC $ s2 $ t), m) | _ => (* product 's * t', where either factor can be 'NONE' *) (case demult (s, m) of (SOME s', m') => (case demult (t, m') of (SOME t', m'') => (SOME (mC $ s' $ t'), m'') | (NONE, m'') => (SOME s', m'')) | (NONE, m') => demult (t, m'))) | demult (atom as (mC as Const (\<^const_name>\Rings.divide\, T)) $ s $ t, m) = (* FIXME: Shouldn't we simplify nested quotients, e.g. '(s/t)/u' could become 's/(t*u)', and '(s*t)/u' could become 's*(t/u)' ? Note that if we choose to do so here, the simpset used by arith must be able to perform the same simplifications. *) (* quotient 's / t', where the denominator t can be NONE *) (* Note: will raise Div iff m' is @0 *) if of_field_sort thy (domain_type T) then let val (os',m') = demult (s, m); val (ot',p) = demult (t, @1) in (case (os',ot') of (SOME s', SOME t') => SOME (mC $ s' $ t') | (SOME s', NONE) => SOME s' | (NONE, SOME t') => SOME (mC $ Const (\<^const_name>\Groups.one\, domain_type (snd (dest_Const mC))) $ t') | (NONE, NONE) => NONE, Rat.mult m' (Rat.inv p)) end else (SOME atom, m) (* terms that evaluate to numeric constants *) | demult (Const (\<^const_name>\Groups.uminus\, _) $ t, m) = demult (t, ~ m) | demult (Const (\<^const_name>\Groups.zero\, _), _) = (NONE, @0) | demult (Const (\<^const_name>\Groups.one\, _), m) = (NONE, m) (*Warning: in rare cases (neg_)numeral encloses a non-numeral, in which case dest_numeral raises TERM; hence all the handles below. Same for Suc-terms that turn out not to be numerals - although the simplifier should eliminate those anyway ...*) | demult (t as Const ("Num.numeral_class.numeral", _) (*DYNAMIC BINDING!*) $ n, m) = ((NONE, Rat.mult m (Rat.of_int (HOLogic.dest_numeral n))) handle TERM _ => (SOME t, m)) | demult (t as Const (\<^const_name>\Suc\, _) $ _, m) = ((NONE, Rat.mult m (Rat.of_int (HOLogic.dest_nat t))) handle TERM _ => (SOME t, m)) (* injection constants are ignored *) | demult (t as Const f $ x, m) = if member (op =) inj_consts f then demult (x, m) else (SOME t, m) (* everything else is considered atomic *) | demult (atom, m) = (SOME atom, m) in demult end; fun decomp0 thy (inj_consts : (string * typ) list) (rel : string, lhs : term, rhs : term) : ((term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat) option = let (* Turns a term 'all' and associated multiplicity 'm' into a list 'p' of summands and associated multiplicities, plus a constant 'i' (with implicit multiplicity 1) *) fun poly (Const (\<^const_name>\Groups.plus\, _) $ s $ t, m : Rat.rat, pi : (term * Rat.rat) list * Rat.rat) = poly (s, m, poly (t, m, pi)) | poly (all as Const (\<^const_name>\Groups.minus\, T) $ s $ t, m, pi) = if nT T then add_atom all m pi else poly (s, m, poly (t, ~ m, pi)) | poly (all as Const (\<^const_name>\Groups.uminus\, T) $ t, m, pi) = if nT T then add_atom all m pi else poly (t, ~ m, pi) | poly (Const (\<^const_name>\Groups.zero\, _), _, pi) = pi | poly (Const (\<^const_name>\Groups.one\, _), m, (p, i)) = (p, Rat.add i m) | poly (all as Const ("Num.numeral_class.numeral", _) (*DYNAMIC BINDING!*) $ t, m, pi as (p, i)) = (let val k = HOLogic.dest_numeral t in (p, Rat.add i (Rat.mult m (Rat.of_int k))) end handle TERM _ => add_atom all m pi) | poly (Const (\<^const_name>\Suc\, _) $ t, m, (p, i)) = poly (t, m, (p, Rat.add i m)) | poly (all as Const (\<^const_name>\Groups.times\, _) $ _ $ _, m, pi as (p, i)) = (case demult thy inj_consts (all, m) of (NONE, m') => (p, Rat.add i m') | (SOME u, m') => add_atom u m' pi) | poly (all as Const (\<^const_name>\Rings.divide\, T) $ _ $ _, m, pi as (p, i)) = if of_field_sort thy (domain_type T) then (case demult thy inj_consts (all, m) of (NONE, m') => (p, Rat.add i m') | (SOME u, m') => add_atom u m' pi) else add_atom all m pi | poly (all as Const f $ x, m, pi) = if member (op =) inj_consts f then poly (x, m, pi) else add_atom all m pi | poly (all, m, pi) = add_atom all m pi val (p, i) = poly (lhs, @1, ([], @0)) val (q, j) = poly (rhs, @1, ([], @0)) in case rel of \<^const_name>\Orderings.less\ => SOME (p, i, "<", q, j) | \<^const_name>\Orderings.less_eq\ => SOME (p, i, "<=", q, j) | \<^const_name>\HOL.eq\ => SOME (p, i, "=", q, j) | _ => NONE end handle General.Div => NONE; fun of_lin_arith_sort thy U = Sign.of_sort thy (U, \<^sort>\Rings.linordered_idom\); fun allows_lin_arith thy (discrete : string list) (U as Type (D, [])) : bool * bool = if of_lin_arith_sort thy U then (true, member (op =) discrete D) else if member (op =) discrete D then (true, true) else (false, false) | allows_lin_arith sg _ U = (of_lin_arith_sort sg U, false); fun decomp_typecheck thy (discrete, inj_consts) (T : typ, xxx) : decomp option = case T of Type ("fun", [U, _]) => (case allows_lin_arith thy discrete U of (true, d) => (case decomp0 thy inj_consts xxx of NONE => NONE | SOME (p, i, rel, q, j) => SOME (p, i, rel, q, j, d)) | (false, _) => NONE) | _ => NONE; fun negate (SOME (x, i, rel, y, j, d)) = SOME (x, i, "~" ^ rel, y, j, d) | negate NONE = NONE; fun decomp_negation thy data ((Const (\<^const_name>\Trueprop\, _)) $ (Const (rel, T) $ lhs $ rhs)) : decomp option = decomp_typecheck thy data (T, (rel, lhs, rhs)) | decomp_negation thy data ((Const (\<^const_name>\Trueprop\, _)) $ (Const (\<^const_name>\Not\, _) $ (Const (rel, T) $ lhs $ rhs))) = negate (decomp_typecheck thy data (T, (rel, lhs, rhs))) | decomp_negation _ _ _ = NONE; fun decomp ctxt : term -> decomp option = let val thy = Proof_Context.theory_of ctxt val {discrete, inj_consts, ...} = get_arith_data ctxt in decomp_negation thy (discrete, inj_consts) end; fun domain_is_nat (_ $ (Const (_, T) $ _ $ _)) = nT T | domain_is_nat (_ $ (Const (\<^const_name>\Not\, _) $ (Const (_, T) $ _ $ _))) = nT T | domain_is_nat _ = false; (* Abstraction of terms *) (* Abstract terms contain only arithmetic operators and relations. When constructing an abstract term for an arbitrary term, non-arithmetic sub-terms are replaced by fresh variables which are declared in the context. Constructing an abstract term from an arbitrary term follows the strategy of decomp. *) fun apply t u = t $ u fun with2 f c t u cx = f t cx ||>> f u |>> (fn (t, u) => c $ t $ u) fun abstract_atom (t as Free _) cx = (t, cx) | abstract_atom (t as Const _) cx = (t, cx) | abstract_atom t (cx as (terms, ctxt)) = (case AList.lookup Envir.aeconv terms t of SOME u => (u, cx) | NONE => let val (n, ctxt') = yield_singleton Variable.variant_fixes "" ctxt val u = Free (n, fastype_of t) in (u, ((t, u) :: terms, ctxt')) end) fun abstract_num t cx = if can HOLogic.dest_number t then (t, cx) else abstract_atom t cx fun is_field_sort (_, ctxt) T = of_field_sort (Proof_Context.theory_of ctxt) (domain_type T) fun is_inj_const (_, ctxt) f = let val {inj_consts, ...} = get_arith_data ctxt in member (op =) inj_consts f end fun abstract_arith ((c as Const (\<^const_name>\Groups.plus\, _)) $ u1 $ u2) cx = with2 abstract_arith c u1 u2 cx | abstract_arith (t as (c as Const (\<^const_name>\Groups.minus\, T)) $ u1 $ u2) cx = if nT T then abstract_atom t cx else with2 abstract_arith c u1 u2 cx | abstract_arith (t as (c as Const (\<^const_name>\Groups.uminus\, T)) $ u) cx = if nT T then abstract_atom t cx else abstract_arith u cx |>> apply c | abstract_arith ((c as Const (\<^const_name>\Suc\, _)) $ u) cx = abstract_arith u cx |>> apply c | abstract_arith ((c as Const (\<^const_name>\Groups.times\, _)) $ u1 $ u2) cx = with2 abstract_arith c u1 u2 cx | abstract_arith (t as (c as Const (\<^const_name>\Rings.divide\, T)) $ u1 $ u2) cx = if is_field_sort cx T then with2 abstract_arith c u1 u2 cx else abstract_atom t cx | abstract_arith (t as (c as Const f) $ u) cx = if is_inj_const cx f then abstract_arith u cx |>> apply c else abstract_num t cx | abstract_arith t cx = abstract_num t cx fun is_lin_arith_rel \<^const_name>\Orderings.less\ = true | is_lin_arith_rel \<^const_name>\Orderings.less_eq\ = true | is_lin_arith_rel \<^const_name>\HOL.eq\ = true | is_lin_arith_rel _ = false fun is_lin_arith_type (_, ctxt) T = let val {discrete, ...} = get_arith_data ctxt in fst (allows_lin_arith (Proof_Context.theory_of ctxt) discrete T) end fun abstract_rel (t as (r as Const (rel, Type ("fun", [U, _]))) $ lhs $ rhs) cx = if is_lin_arith_rel rel andalso is_lin_arith_type cx U then with2 abstract_arith r lhs rhs cx else abstract_atom t cx | abstract_rel t cx = abstract_atom t cx fun abstract_neg ((c as Const (\<^const_name>\Not\, _)) $ t) cx = abstract_rel t cx |>> apply c | abstract_neg t cx = abstract_rel t cx fun abstract ((c as Const (\<^const_name>\Trueprop\, _)) $ t) cx = abstract_neg t cx |>> apply c | abstract t cx = abstract_atom t cx (*---------------------------------------------------------------------------*) (* the following code performs splitting of certain constants (e.g., min, *) (* max) in a linear arithmetic problem; similar to what split_tac later does *) (* to the proof state *) (*---------------------------------------------------------------------------*) (* checks if splitting with 'thm' is implemented *) fun is_split_thm ctxt thm = (case Thm.concl_of thm of _ $ (_ $ (_ $ lhs) $ _) => (* Trueprop $ ((op =) $ (?P $ lhs) $ rhs) *) (case head_of lhs of Const (a, _) => member (op =) [\<^const_name>\Orderings.max\, \<^const_name>\Orderings.min\, \<^const_name>\Groups.abs\, \<^const_name>\Groups.minus\, "Int.nat" (*DYNAMIC BINDING!*), \<^const_name>\Rings.modulo\, \<^const_name>\Rings.divide\] a | _ => (if Context_Position.is_visible ctxt then warning ("Lin. Arith.: wrong format for split rule " ^ Thm.string_of_thm ctxt thm) else (); false)) | _ => (if Context_Position.is_visible ctxt then warning ("Lin. Arith.: wrong format for split rule " ^ Thm.string_of_thm ctxt thm) else (); false)); (* substitute new for occurrences of old in a term, incrementing bound *) (* variables as needed when substituting inside an abstraction *) fun subst_term ([] : (term * term) list) (t : term) = t | subst_term pairs t = (case AList.lookup Envir.aeconv pairs t of SOME new => new | NONE => (case t of Abs (a, T, body) => let val pairs' = map (apply2 (incr_boundvars 1)) pairs in Abs (a, T, subst_term pairs' body) end | t1 $ t2 => subst_term pairs t1 $ subst_term pairs t2 | _ => t)); (* approximates the effect of one application of split_tac (followed by NNF *) (* normalization) on the subgoal represented by '(Ts, terms)'; returns a *) (* list of new subgoals (each again represented by a typ list for bound *) (* variables and a term list for premises), or NONE if split_tac would fail *) (* on the subgoal *) (* FIXME: currently only the effect of certain split theorems is reproduced *) (* (which is why we need 'is_split_thm'). A more canonical *) (* implementation should analyze the right-hand side of the split *) (* theorem that can be applied, and modify the subgoal accordingly. *) (* Or even better, the splitter should be extended to provide *) (* splitting on terms as well as splitting on theorems (where the *) (* former can have a faster implementation as it does not need to be *) (* proof-producing). *) fun split_once_items ctxt (Ts : typ list, terms : term list) : (typ list * term list) list option = let val thy = Proof_Context.theory_of ctxt (* takes a list [t1, ..., tn] to the term *) (* tn' --> ... --> t1' --> False , *) (* where ti' = HOLogic.dest_Trueprop ti *) fun REPEAT_DETERM_etac_rev_mp tms = fold (curry HOLogic.mk_imp) (map HOLogic.dest_Trueprop tms) \<^term>\False\ val split_thms = filter (is_split_thm ctxt) (get_splits ctxt) val cmap = Splitter.cmap_of_split_thms split_thms val goal_tm = REPEAT_DETERM_etac_rev_mp terms val splits = Splitter.split_posns cmap thy Ts goal_tm val split_limit = Config.get ctxt split_limit in if length splits > split_limit then ( tracing ("linarith_split_limit exceeded (current value is " ^ string_of_int split_limit ^ ")"); NONE ) else case splits of [] => (* split_tac would fail: no possible split *) NONE | (_, _::_, _, _, _) :: _ => (* disallow a split that involves non-locally bound variables (except *) (* when bound by outermost meta-quantifiers) *) NONE | (_, [], _, split_type, split_term) :: _ => (* ignore all but the first possible split *) (case strip_comb split_term of (* ?P (max ?i ?j) = ((?i <= ?j --> ?P ?j) & (~ ?i <= ?j --> ?P ?i)) *) (Const (\<^const_name>\Orderings.max\, _), [t1, t2]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(split_term, t2)]) rev_terms val t1_leq_t2 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 val not_t1_leq_t2 = HOLogic.Not $ t1_leq_t2 val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_leq_t2) :: terms2 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop not_t1_leq_t2) :: terms1 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (min ?i ?j) = ((?i <= ?j --> ?P ?i) & (~ ?i <= ?j --> ?P ?j)) *) | (Const (\<^const_name>\Orderings.min\, _), [t1, t2]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(split_term, t2)]) rev_terms val t1_leq_t2 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 val not_t1_leq_t2 = HOLogic.Not $ t1_leq_t2 val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_leq_t2) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop not_t1_leq_t2) :: terms2 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (abs ?a) = ((0 <= ?a --> ?P ?a) & (?a < 0 --> ?P (- ?a))) *) | (Const (\<^const_name>\Groups.abs\, _), [t1]) => let val rev_terms = rev terms val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(split_term, Const (\<^const_name>\Groups.uminus\, split_type --> split_type) $ t1)]) rev_terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val zero_leq_t1 = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ t1 val t1_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t1 $ zero val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop zero_leq_t1) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_lt_zero) :: terms2 @ [not_false] in SOME [(Ts, subgoal1), (Ts, subgoal2)] end (* ?P (?a - ?b) = ((?a < ?b --> ?P 0) & (ALL d. ?a = ?b + d --> ?P d)) *) | (Const (\<^const_name>\Groups.minus\, _), [t1, t2]) => let (* "d" in the above theorem becomes a new bound variable after NNF *) (* transformation, therefore some adjustment of indices is necessary *) val rev_terms = rev terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val d = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms val terms2 = map (subst_term [(incr_boundvars 1 split_term, d)]) (map (incr_boundvars 1) rev_terms) val t1' = incr_boundvars 1 t1 val t2' = incr_boundvars 1 t2 val t1_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t1 $ t2 val t1_eq_t2_plus_d = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ t2' $ d) val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_lt_t2) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_eq_t2_plus_d) :: terms2 @ [not_false] in SOME [(Ts, subgoal1), (split_type :: Ts, subgoal2)] end (* ?P (nat ?i) = ((ALL n. ?i = of_nat n --> ?P n) & (?i < 0 --> ?P 0)) *) | (Const ("Int.nat", _), (*DYNAMIC BINDING!*) [t1]) => let val rev_terms = rev terms val zero_int = Const (\<^const_name>\Groups.zero\, HOLogic.intT) val zero_nat = Const (\<^const_name>\Groups.zero\, HOLogic.natT) val n = Bound 0 val terms1 = map (subst_term [(incr_boundvars 1 split_term, n)]) (map (incr_boundvars 1) rev_terms) val terms2 = map (subst_term [(split_term, zero_nat)]) rev_terms val t1' = incr_boundvars 1 t1 val t1_eq_nat_n = Const (\<^const_name>\HOL.eq\, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\of_nat\, HOLogic.natT --> HOLogic.intT) $ n) val t1_lt_zero = Const (\<^const_name>\Orderings.less\, HOLogic.intT --> HOLogic.intT --> HOLogic.boolT) $ t1 $ zero_int val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t1_eq_nat_n) :: terms1 @ [not_false] val subgoal2 = (HOLogic.mk_Trueprop t1_lt_zero) :: terms2 @ [not_false] in SOME [(HOLogic.natT :: Ts, subgoal1), (Ts, subgoal2)] end (* ?P ((?n::nat) mod (numeral ?k)) = ((numeral ?k = 0 --> ?P ?n) & (~ (numeral ?k = 0) --> (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P j))) *) | (Const (\<^const_name>\Rings.modulo\, Type ("fun", [\<^typ>\nat\, _])), [t1, t2]) => let val rev_terms = rev terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2 = map (subst_term [(incr_boundvars 2 split_term, j)]) (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero val t2_neq_zero = HOLogic.mk_not (Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero) val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [t2_neq_zero, j_lt_t2, t1_eq_t2_times_i_plus_j]) @ terms2 @ [not_false] in SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)] end (* ?P ((?n::nat) div (numeral ?k)) = ((numeral ?k = 0 --> ?P 0) & (~ (numeral ?k = 0) --> (ALL i j. j < numeral ?k --> ?n = numeral ?k * i + j --> ?P i))) *) | (Const (\<^const_name>\Rings.divide\, Type ("fun", [\<^typ>\nat\, _])), [t1, t2]) => let val rev_terms = rev terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms val terms2 = map (subst_term [(incr_boundvars 2 split_term, i)]) (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero val t2_neq_zero = HOLogic.mk_not (Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero) val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [t2_neq_zero, j_lt_t2, t1_eq_t2_times_i_plus_j]) @ terms2 @ [not_false] in SOME [(Ts, subgoal1), (split_type :: split_type :: Ts, subgoal2)] end (* ?P ((?n::int) mod (numeral ?k)) = ((numeral ?k = 0 --> ?P ?n) & (0 < numeral ?k --> (ALL i j. 0 <= j & j < numeral ?k & ?n = numeral ?k * i + j --> ?P j)) & (numeral ?k < 0 --> (ALL i j. numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P j))) *) | (Const (\<^const_name>\Rings.modulo\, Type ("fun", [Type ("Int.int", []), _])), (*DYNAMIC BINDING!*) [t1, t2]) => let val rev_terms = rev terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, t1)]) rev_terms val terms2_3 = map (subst_term [(incr_boundvars 2 split_term, j)]) (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero val zero_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ zero $ t2' val t2_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero val zero_leq_j = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ j val j_leq_zero = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ j $ zero val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' val t2_lt_j = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ t2' $ j val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j]) @ hd terms2_3 :: (if tl terms2_3 = [] then [not_false] else []) @ (map HOLogic.mk_Trueprop [j_lt_t2, t1_eq_t2_times_i_plus_j]) @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false]) val subgoal3 = (map HOLogic.mk_Trueprop [t2_lt_zero, t2_lt_j]) @ hd terms2_3 :: (if tl terms2_3 = [] then [not_false] else []) @ (map HOLogic.mk_Trueprop [j_leq_zero, t1_eq_t2_times_i_plus_j]) @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false]) val Ts' = split_type :: split_type :: Ts in SOME [(Ts, subgoal1), (Ts', subgoal2), (Ts', subgoal3)] end (* ?P ((?n::int) div (numeral ?k)) = ((numeral ?k = 0 --> ?P 0) & (0 < numeral ?k --> (ALL i j. 0 <= j & j < numeral ?k & ?n = numeral ?k * i + j --> ?P i)) & (numeral ?k < 0 --> (ALL i j. numeral ?k < j & j <= 0 & ?n = numeral ?k * i + j --> ?P i))) *) | (Const (\<^const_name>\Rings.divide\, Type ("fun", [Type ("Int.int", []), _])), (*DYNAMIC BINDING!*) [t1, t2]) => let val rev_terms = rev terms val zero = Const (\<^const_name>\Groups.zero\, split_type) val i = Bound 1 val j = Bound 0 val terms1 = map (subst_term [(split_term, zero)]) rev_terms val terms2_3 = map (subst_term [(incr_boundvars 2 split_term, i)]) (map (incr_boundvars 2) rev_terms) val t1' = incr_boundvars 2 t1 val t2' = incr_boundvars 2 t2 val t2_eq_zero = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t2 $ zero val zero_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ zero $ t2' val t2_lt_zero = Const (\<^const_name>\Orderings.less\, split_type --> split_type --> HOLogic.boolT) $ t2' $ zero val zero_leq_j = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ zero $ j val j_leq_zero = Const (\<^const_name>\Orderings.less_eq\, split_type --> split_type --> HOLogic.boolT) $ j $ zero val j_lt_t2 = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ j $ t2' val t2_lt_j = Const (\<^const_name>\Orderings.less\, split_type --> split_type--> HOLogic.boolT) $ t2' $ j val t1_eq_t2_times_i_plus_j = Const (\<^const_name>\HOL.eq\, split_type --> split_type --> HOLogic.boolT) $ t1' $ (Const (\<^const_name>\Groups.plus\, split_type --> split_type --> split_type) $ (Const (\<^const_name>\Groups.times\, split_type --> split_type --> split_type) $ t2' $ i) $ j) val not_false = HOLogic.mk_Trueprop (HOLogic.Not $ \<^term>\False\) val subgoal1 = (HOLogic.mk_Trueprop t2_eq_zero) :: terms1 @ [not_false] val subgoal2 = (map HOLogic.mk_Trueprop [zero_lt_t2, zero_leq_j]) @ hd terms2_3 :: (if tl terms2_3 = [] then [not_false] else []) @ (map HOLogic.mk_Trueprop [j_lt_t2, t1_eq_t2_times_i_plus_j]) @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false]) val subgoal3 = (map HOLogic.mk_Trueprop [t2_lt_zero, t2_lt_j]) @ hd terms2_3 :: (if tl terms2_3 = [] then [not_false] else []) @ (map HOLogic.mk_Trueprop [j_leq_zero, t1_eq_t2_times_i_plus_j]) @ (if tl terms2_3 = [] then [] else tl terms2_3 @ [not_false]) val Ts' = split_type :: split_type :: Ts in SOME [(Ts, subgoal1), (Ts', subgoal2), (Ts', subgoal3)] end (* this will only happen if a split theorem can be applied for which no *) (* code exists above -- in which case either the split theorem should be *) (* implemented above, or 'is_split_thm' should be modified to filter it *) (* out *) | (t, ts) => (if Context_Position.is_visible ctxt then warning ("Lin. Arith.: split rule for " ^ Syntax.string_of_term ctxt t ^ " (with " ^ string_of_int (length ts) ^ " argument(s)) not implemented; proof reconstruction is likely to fail") else (); NONE)) end; (* split_once_items *) (* remove terms that do not satisfy 'p'; change the order of the remaining *) (* terms in the same way as filter_prems_tac does *) fun filter_prems_tac_items (p : term -> bool) (terms : term list) : term list = let fun filter_prems t (left, right) = if p t then (left, right @ [t]) else (left @ right, []) val (left, right) = fold filter_prems terms ([], []) in right @ left end; (* return true iff TRY (etac notE) THEN eq_assume_tac would succeed on a *) (* subgoal that has 'terms' as premises *) fun negated_term_occurs_positively (terms : term list) : bool = exists (fn (Trueprop $ (Const (\<^const_name>\Not\, _) $ t)) => member Envir.aeconv terms (Trueprop $ t) | _ => false) terms; fun pre_decomp ctxt (Ts : typ list, terms : term list) : (typ list * term list) list = let (* repeatedly split (including newly emerging subgoals) until no further *) (* splitting is possible *) fun split_loop ([] : (typ list * term list) list) = ([] : (typ list * term list) list) | split_loop (subgoal::subgoals) = (case split_once_items ctxt subgoal of SOME new_subgoals => split_loop (new_subgoals @ subgoals) | NONE => subgoal :: split_loop subgoals) fun is_relevant t = is_some (decomp ctxt t) (* filter_prems_tac is_relevant: *) val relevant_terms = filter_prems_tac_items is_relevant terms (* split_tac, NNF normalization: *) val split_goals = split_loop [(Ts, relevant_terms)] (* necessary because split_once_tac may normalize terms: *) val beta_eta_norm = map (apsnd (map (Envir.eta_contract o Envir.beta_norm))) split_goals (* TRY (etac notE) THEN eq_assume_tac: *) val result = filter_out (negated_term_occurs_positively o snd) beta_eta_norm in result end; (* takes the i-th subgoal [| A1; ...; An |] ==> B to *) (* An --> ... --> A1 --> B, performs splitting with the given 'split_thms' *) (* (resulting in a different subgoal P), takes P to ~P ==> False, *) (* performs NNF-normalization of ~P, and eliminates conjunctions, *) (* disjunctions and existential quantifiers from the premises, possibly (in *) (* the case of disjunctions) resulting in several new subgoals, each of the *) (* general form [| Q1; ...; Qm |] ==> False. Fails if more than *) (* !split_limit splits are possible. *) local fun nnf_simpset ctxt = (empty_simpset ctxt |> Simplifier.set_mkeqTrue mk_eq_True |> Simplifier.set_mksimps (mksimps mksimps_pairs)) addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj}, @{thm de_Morgan_conj}, not_all, not_ex, not_not] fun prem_nnf_tac ctxt = full_simp_tac (nnf_simpset ctxt) in fun split_once_tac ctxt split_thms = let val thy = Proof_Context.theory_of ctxt val cond_split_tac = SUBGOAL (fn (subgoal, i) => let val Ts = rev (map snd (Logic.strip_params subgoal)) val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl subgoal) val cmap = Splitter.cmap_of_split_thms split_thms val splits = Splitter.split_posns cmap thy Ts concl in if null splits orelse length splits > Config.get ctxt split_limit then no_tac else if null (#2 (hd splits)) then split_tac ctxt split_thms i else (* disallow a split that involves non-locally bound variables *) (* (except when bound by outermost meta-quantifiers) *) no_tac end) in EVERY' [ REPEAT_DETERM o eresolve_tac ctxt [rev_mp], cond_split_tac, resolve_tac ctxt @{thms ccontr}, prem_nnf_tac ctxt, TRY o REPEAT_ALL_NEW (DETERM o (eresolve_tac ctxt [conjE, exE] ORELSE' eresolve_tac ctxt [disjE])) ] end; end; (* local *) (* remove irrelevant premises, then split the i-th subgoal (and all new *) (* subgoals) by using 'split_once_tac' repeatedly. Beta-eta-normalize new *) (* subgoals and finally attempt to solve them by finding an immediate *) (* contradiction (i.e., a term and its negation) in their premises. *) fun pre_tac ctxt i = let val split_thms = filter (is_split_thm ctxt) (get_splits ctxt) fun is_relevant t = is_some (decomp ctxt t) in DETERM ( TRY (filter_prems_tac ctxt is_relevant i) THEN ( (TRY o REPEAT_ALL_NEW (split_once_tac ctxt split_thms)) THEN_ALL_NEW (CONVERSION Drule.beta_eta_conversion THEN' (TRY o (eresolve_tac ctxt [notE] THEN' eq_assume_tac))) ) i ) end; end; (* LA_Data *) val pre_tac = LA_Data.pre_tac; structure Fast_Arith = Fast_Lin_Arith(structure LA_Logic = LA_Logic and LA_Data = LA_Data); val add_inj_thms = Fast_Arith.add_inj_thms; val add_lessD = Fast_Arith.add_lessD; val add_simps = Fast_Arith.add_simps; val add_simprocs = Fast_Arith.add_simprocs; val set_number_of = Fast_Arith.set_number_of; val simple_tac = Fast_Arith.lin_arith_tac; (* reduce contradictory <= to False. Most of the work is done by the cancel tactics. *) val init_arith_data = Fast_Arith.map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, number_of, ...} => {add_mono_thms = map Thm.trim_context @{thms add_mono_thms_linordered_semiring add_mono_thms_linordered_field} @ add_mono_thms, mult_mono_thms = map Thm.trim_context (@{thms mult_strict_left_mono mult_left_mono} @ [@{lemma "a = b ==> c * a = c * b" by (rule arg_cong)}]) @ mult_mono_thms, inj_thms = inj_thms, lessD = lessD, neqE = map Thm.trim_context @{thms linorder_neqE_nat linorder_neqE_linordered_idom} @ neqE, simpset = put_simpset HOL_basic_ss \<^context> |> Simplifier.add_cong @{thm if_weak_cong} |> simpset_of, number_of = number_of}); (* FIXME !?? *) fun add_arith_facts ctxt = Simplifier.add_prems (rev (Named_Theorems.get ctxt \<^named_theorems>\arith\)) ctxt; val simproc = add_arith_facts #> Fast_Arith.lin_arith_simproc; (* generic refutation procedure *) (* parameters: test: term -> bool tests if a term is at all relevant to the refutation proof; if not, then it can be discarded. Can improve performance, esp. if disjunctions can be discarded (no case distinction needed!). prep_tac: int -> tactic A preparation tactic to be applied to the goal once all relevant premises have been moved to the conclusion. ref_tac: int -> tactic the actual refutation tactic. Should be able to deal with goals [| A1; ...; An |] ==> False where the Ai are atomic, i.e. no top-level &, | or EX *) local fun nnf_simpset ctxt = (empty_simpset ctxt |> Simplifier.set_mkeqTrue mk_eq_True |> Simplifier.set_mksimps (mksimps mksimps_pairs)) addsimps [@{thm imp_conv_disj}, @{thm iff_conv_conj_imp}, @{thm de_Morgan_disj}, @{thm de_Morgan_conj}, @{thm not_all}, @{thm not_ex}, @{thm not_not}]; fun prem_nnf_tac ctxt = full_simp_tac (nnf_simpset ctxt); in fun refute_tac ctxt test prep_tac ref_tac = let val refute_prems_tac = REPEAT_DETERM (eresolve_tac ctxt [@{thm conjE}, @{thm exE}] 1 ORELSE filter_prems_tac ctxt test 1 ORELSE eresolve_tac ctxt @{thms disjE} 1) THEN (DETERM (eresolve_tac ctxt @{thms notE} 1 THEN eq_assume_tac 1) ORELSE ref_tac 1); in EVERY'[TRY o filter_prems_tac ctxt test, REPEAT_DETERM o eresolve_tac ctxt @{thms rev_mp}, prep_tac, resolve_tac ctxt @{thms ccontr}, prem_nnf_tac ctxt, SELECT_GOAL (DEPTH_SOLVE refute_prems_tac)] end; end; (* arith proof method *) local fun raw_tac ctxt = (* FIXME: K true should be replaced by a sensible test (perhaps "is_some o decomp sg"? -- but note that the test is applied to terms already before they are split/normalized) to speed things up in case there are lots of irrelevant terms involved; elimination of min/max can be optimized: (max m n + k <= r) = (m+k <= r & n+k <= r) (l <= min m n + k) = (l <= m+k & l <= n+k) *) refute_tac ctxt (K true) (* Splitting is also done inside simple_tac, but not completely -- *) (* split_tac may use split theorems that have not been implemented in *) (* simple_tac (cf. pre_decomp and split_once_items above), and *) (* split_limit may trigger. *) (* Therefore splitting outside of simple_tac may allow us to prove *) (* some goals that simple_tac alone would fail on. *) (REPEAT_DETERM o split_tac ctxt (get_splits ctxt)) (Fast_Arith.lin_arith_tac ctxt); in fun tac ctxt = FIRST' [simple_tac ctxt, Object_Logic.full_atomize_tac ctxt THEN' (REPEAT_DETERM o resolve_tac ctxt [impI]) THEN' raw_tac ctxt]; end; (* context setup *) val global_setup = map_theory_simpset (fn ctxt => ctxt addSolver (mk_solver "lin_arith" (add_arith_facts #> Fast_Arith.prems_lin_arith_tac))) #> Attrib.setup \<^binding>\arith_split\ (Scan.succeed (Thm.declaration_attribute add_split)) "declaration of split rules for arithmetic procedure" #> Method.setup \<^binding>\linarith\ (Scan.succeed (fn ctxt => METHOD (fn facts => HEADGOAL (Method.insert_tac ctxt (rev (Named_Theorems.get ctxt \<^named_theorems>\arith\) @ facts) THEN' tac ctxt)))) "linear arithmetic" #> Arith_Data.add_tactic "linear arithmetic" tac; end; diff --git a/src/HOL/Tools/monomorph.ML b/src/HOL/Tools/monomorph.ML --- a/src/HOL/Tools/monomorph.ML +++ b/src/HOL/Tools/monomorph.ML @@ -1,323 +1,323 @@ (* Title: HOL/Tools/monomorph.ML Author: Sascha Boehme, TU Muenchen Monomorphization of theorems, i.e., computation of ground instances for theorems with type variables. This procedure is incomplete in general, but works well for most practical problems. Monomorphization is essentially an enumeration of substitutions that map schematic types to ground types. Applying these substitutions to theorems with type variables results in monomorphized ground instances. The enumeration is driven by schematic constants (constants occurring with type variables) and all ground instances of such constants (occurrences without type variables). The enumeration is organized in rounds in which all substitutions for the schematic constants are computed that are induced by the ground instances. Any new ground instance may induce further substitutions in a subsequent round. To prevent nontermination, there is an upper limit of rounds involved and of the number of monomorphized ground instances computed. Theorems to be monomorphized must be tagged with a number indicating the initial round in which they participate first. The initial round is ignored for theorems without type variables. For any other theorem, the initial round must be greater than zero. Returned monomorphized theorems carry a number showing from which monomorphization round they emerged. *) signature MONOMORPH = sig (* utility functions *) val typ_has_tvars: typ -> bool val all_schematic_consts_of: term -> typ list Symtab.table val add_schematic_consts_of: term -> typ list Symtab.table -> typ list Symtab.table (* configuration options *) val max_rounds: int Config.T val max_new_instances: int Config.T val max_thm_instances: int Config.T val max_new_const_instances_per_round: int Config.T val max_duplicated_instances: int Config.T (* monomorphization *) val monomorph: (term -> typ list Symtab.table) -> Proof.context -> (int * thm) list -> (int * thm) list list end structure Monomorph: MONOMORPH = struct (* utility functions *) val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false) fun add_schematic_const (c as (_, T)) = if typ_has_tvars T then Symtab.insert_list (op =) c else I fun add_schematic_consts_of t = Term.fold_aterms (fn Const c => add_schematic_const c | _ => I) t fun all_schematic_consts_of t = add_schematic_consts_of t Symtab.empty fun clear_grounds grounds = Symtab.map (K (K [])) grounds (* configuration options *) val max_rounds = Attrib.setup_config_int \<^binding>\monomorph_max_rounds\ (K 5) val max_new_instances = Attrib.setup_config_int \<^binding>\monomorph_max_new_instances\ (K 500) val max_thm_instances = Attrib.setup_config_int \<^binding>\monomorph_max_thm_instances\ (K 20) val max_new_const_instances_per_round = Attrib.setup_config_int \<^binding>\monomorph_max_new_const_instances_per_round\ (K 5) val max_duplicated_instances = Attrib.setup_config_int \<^binding>\monomorph_max_duplicated_instances\ (K 16000) fun limit_rounds ctxt f = let val max = Config.get ctxt max_rounds fun round i x = if i > max then x else round (i + 1) (f ctxt i x) in round 1 end (* theorem information and related functions *) datatype thm_info = Ground of thm | Ignored | Schematic of { id: int, theorem: thm, tvars: (indexname * sort) list, schematics: (string * typ) list, initial_round: int} fun fold_grounds f = fold (fn Ground thm => f thm | _ => I) fun fold_schematic f thm_info = (case thm_info of Schematic {id, theorem, tvars, schematics, initial_round} => f id theorem tvars schematics initial_round | _ => I) fun fold_schematics pred f = let fun apply id thm tvars schematics initial_round x = if pred initial_round then f id thm tvars schematics x else x in fold (fold_schematic apply) end (* collecting data *) (* Theorems with type variables that cannot be instantiated should be ignored. A type variable has only a chance to be instantiated if it occurs in the type of one of the schematic constants. *) fun groundable all_tvars schematics = let val tvars' = Symtab.fold (fold Term.add_tvarsT o snd) schematics [] in forall (member (op =) tvars') all_tvars end fun prepare schematic_consts_of rthms = let fun prep (initial_round, thm) ((id, idx), consts) = if Term.exists_type typ_has_tvars (Thm.prop_of thm) then let (* increase indices to avoid clashes of type variables *) val thm' = Thm.incr_indexes idx thm val idx' = Thm.maxidx_of thm' + 1 val tvars = Term.add_tvars (Thm.prop_of thm') [] val schematics = schematic_consts_of (Thm.prop_of thm') val schematics' = Symtab.fold (fn (n, Ts) => fold (cons o pair n) Ts) schematics [] (* collect the names of all constants that need to be instantiated *) val consts' = consts |> Symtab.fold (fn (n, _) => Symtab.update (n, [])) schematics val thm_info = if not (groundable tvars schematics) then Ignored else Schematic { id = id, theorem = thm', tvars = tvars, schematics = schematics', initial_round = initial_round} in (thm_info, ((id + 1, idx'), consts')) end else (Ground thm, ((id + 1, idx + Thm.maxidx_of thm + 1), consts)) in fold_map prep rthms ((0, 0), Symtab.empty) ||> snd end (* collecting instances *) fun instantiate ctxt subst = let fun cert (ix, (S, T)) = ((ix, S), Thm.ctyp_of ctxt T) fun cert' subst = Vartab.fold (cons o cert) subst [] - in Thm.instantiate (cert' subst, []) end + in Thm.instantiate (TVars.make (cert' subst), Vars.empty) end fun add_new_grounds used_grounds new_grounds thm = let fun mem tab (n, T) = member (op =) (Symtab.lookup_list tab n) T fun add (Const (c as (n, _))) = if mem used_grounds c orelse mem new_grounds c then I else if not (Symtab.defined used_grounds n) then I else Symtab.insert_list (op =) c | add _ = I in Term.fold_aterms add (Thm.prop_of thm) end fun add_insts max_instances max_duplicated_instances max_thm_insts ctxt round used_grounds new_grounds id thm tvars schematics cx = let exception ENOUGH of typ list Symtab.table * (int * int * ((int * (sort * typ) Vartab.table) * thm) list Inttab.table) val thy = Proof_Context.theory_of ctxt fun add subst (cx as (next_grounds, (hits, misses, insts))) = if hits >= max_instances orelse misses >= max_duplicated_instances then raise ENOUGH cx else let val thm' = instantiate ctxt subst thm val rthm = ((round, subst), thm') val rthms = Inttab.lookup_list insts id val n_insts' = if member (eq_snd Thm.eq_thm) rthms rthm then (hits, misses + 1, insts) else let val (hits', misses') = if length rthms >= max_thm_insts then (hits, misses + 1) else (hits + 1, misses) in (hits', misses', Inttab.cons_list (id, rthm) insts) end val next_grounds' = add_new_grounds used_grounds new_grounds thm' next_grounds in (next_grounds', n_insts') end fun with_grounds (n, T) f subst (n', Us) = let fun matching U = (* one-step refinement of the given substitution *) (case try (Sign.typ_match thy (T, U)) subst of NONE => I | SOME subst' => f subst') in if n = n' then fold matching Us else I end fun with_matching_ground c subst f = (* Try new grounds before already used grounds. Otherwise only substitutions already seen in previous rounds get enumerated. *) Symtab.fold (with_grounds c (f true) subst) new_grounds #> Symtab.fold (with_grounds c (f false) subst) used_grounds fun is_complete subst = (* Check if a substitution is defined for all TVars of the theorem, which guarantees that the instantiation with this substitution results in a ground theorem since all matchings that led to this substitution are with ground types only. *) subset (op =) (tvars, Vartab.fold (cons o apsnd fst) subst []) fun for_schematics _ [] _ = I | for_schematics used_new (c :: cs) subst = with_matching_ground c subst (fn new => fn subst' => if is_complete subst' then if used_new orelse new then add subst' else I else for_schematics (used_new orelse new) cs subst') #> for_schematics used_new cs subst in (* Enumerate all substitutions that lead to a ground instance of the theorem not seen before. A necessary condition for such a new ground instance is the usage of at least one ground from the new_grounds table. The approach used here is to match all schematics of the theorem with all relevant grounds. *) for_schematics false schematics Vartab.empty cx handle ENOUGH cx' => cx' end fun is_new round initial_round = (round = initial_round) fun is_active round initial_round = (round > initial_round) fun find_instances max_instances max_duplicated_instances max_thm_insts max_new_grounds thm_infos ctxt round (known_grounds, new_grounds0, insts) = let val new_grounds = Symtab.map (fn _ => fn grounds => if length grounds <= max_new_grounds then grounds else take max_new_grounds (sort Term_Ord.typ_ord grounds)) new_grounds0 val add_new = add_insts max_instances max_duplicated_instances max_thm_insts ctxt round fun consider_all pred f (cx as (_, (hits, misses, _))) = if hits >= max_instances orelse misses >= max_duplicated_instances then cx else fold_schematics pred f thm_infos cx val known_grounds' = Symtab.merge_list (op =) (known_grounds, new_grounds) val empty_grounds = clear_grounds known_grounds' val (new_grounds', insts') = (Symtab.empty, insts) |> consider_all (is_active round) (add_new known_grounds new_grounds) |> consider_all (is_new round) (add_new empty_grounds known_grounds') in (known_grounds', new_grounds', insts') end fun add_ground_types thm = let fun add (n, T) = Symtab.map_entry n (insert (op =) T) in Term.fold_aterms (fn Const c => add c | _ => I) (Thm.prop_of thm) end fun collect_instances ctxt max_thm_insts max_new_grounds thm_infos consts = let val known_grounds = fold_grounds add_ground_types thm_infos consts val empty_grounds = clear_grounds known_grounds val max_instances = Config.get ctxt max_new_instances |> fold (fn Schematic _ => Integer.add 1 | _ => I) thm_infos val max_duplicated_instances = Config.get ctxt max_duplicated_instances val (_, _, (_, _, insts)) = limit_rounds ctxt (find_instances max_instances max_duplicated_instances max_thm_insts max_new_grounds thm_infos) (empty_grounds, known_grounds, (0, 0, Inttab.empty)) in insts end (* monomorphization *) fun size_of_subst subst = Vartab.fold (Integer.add o size_of_typ o snd o snd) subst 0 fun subst_ord subst = int_ord (apply2 size_of_subst subst) fun instantiated_thms _ _ (Ground thm) = [(0, thm)] | instantiated_thms _ _ Ignored = [] | instantiated_thms max_thm_insts insts (Schematic {id, ...}) = Inttab.lookup_list insts id |> (fn rthms => if length rthms <= max_thm_insts then rthms else take max_thm_insts (sort (prod_ord int_ord subst_ord o apply2 fst) rthms)) |> map (apfst fst) fun monomorph schematic_consts_of ctxt rthms = let val max_thm_insts = Config.get ctxt max_thm_instances val max_new_grounds = Config.get ctxt max_new_const_instances_per_round val (thm_infos, consts) = prepare schematic_consts_of rthms val insts = if Symtab.is_empty consts then Inttab.empty else collect_instances ctxt max_thm_insts max_new_grounds thm_infos consts in map (instantiated_thms max_thm_insts insts) thm_infos end end diff --git a/src/HOL/Tools/numeral.ML b/src/HOL/Tools/numeral.ML --- a/src/HOL/Tools/numeral.ML +++ b/src/HOL/Tools/numeral.ML @@ -1,116 +1,116 @@ (* Title: HOL/Tools/numeral.ML Author: Makarius Logical and syntactic operations on numerals (see also HOL/Tools/hologic.ML). *) signature NUMERAL = sig val mk_cnumber: ctyp -> int -> cterm val mk_number_syntax: int -> term val dest_num_syntax: term -> int val add_code: string -> (int -> int) -> (Code_Printer.literals -> int -> string) -> string -> theory -> theory end; structure Numeral: NUMERAL = struct (* numeral *) fun dest_num_syntax (Const (\<^const_syntax>\Num.Bit0\, _) $ t) = 2 * dest_num_syntax t | dest_num_syntax (Const (\<^const_syntax>\Num.Bit1\, _) $ t) = 2 * dest_num_syntax t + 1 | dest_num_syntax (Const (\<^const_syntax>\Num.One\, _)) = 1; fun mk_num_syntax n = if n > 0 then (case Integer.quot_rem n 2 of (0, 1) => Syntax.const \<^const_syntax>\One\ | (n, 0) => Syntax.const \<^const_syntax>\Bit0\ $ mk_num_syntax n | (n, 1) => Syntax.const \<^const_syntax>\Bit1\ $ mk_num_syntax n) else raise Match fun mk_cbit 0 = \<^cterm>\Num.Bit0\ | mk_cbit 1 = \<^cterm>\Num.Bit1\ | mk_cbit _ = raise CTERM ("mk_cbit", []); fun mk_cnumeral i = let fun mk 1 = \<^cterm>\Num.One\ | mk i = let val (q, r) = Integer.div_mod i 2 in Thm.apply (mk_cbit r) (mk q) end in if i > 0 then mk i else raise CTERM ("mk_cnumeral: negative input", []) end (* number *) local val cterm_of = Thm.cterm_of \<^context>; fun tvar S = (("'a", 0), S); val zero_tvar = tvar \<^sort>\zero\; val zero = cterm_of (Const (\<^const_name>\zero_class.zero\, TVar zero_tvar)); val one_tvar = tvar \<^sort>\one\; val one = cterm_of (Const (\<^const_name>\one_class.one\, TVar one_tvar)); val numeral_tvar = tvar \<^sort>\numeral\; val numeral = cterm_of (Const (\<^const_name>\numeral\, \<^typ>\num\ --> TVar numeral_tvar)); val uminus_tvar = tvar \<^sort>\uminus\; val uminus = cterm_of (Const (\<^const_name>\uminus\, TVar uminus_tvar --> TVar uminus_tvar)); -fun instT T v = Thm.instantiate_cterm ([(v, T)], []); +fun instT T v = Thm.instantiate_cterm (TVars.make [(v, T)], Vars.empty); in fun mk_cnumber T 0 = instT T zero_tvar zero | mk_cnumber T 1 = instT T one_tvar one | mk_cnumber T i = if i > 0 then Thm.apply (instT T numeral_tvar numeral) (mk_cnumeral i) else Thm.apply (instT T uminus_tvar uminus) (Thm.apply (instT T numeral_tvar numeral) (mk_cnumeral (~ i))); end; fun mk_number_syntax n = if n = 0 then Syntax.const \<^const_syntax>\Groups.zero\ else if n = 1 then Syntax.const \<^const_syntax>\Groups.one\ else Syntax.const \<^const_syntax>\numeral\ $ mk_num_syntax n; (* code generator *) local open Basic_Code_Thingol in fun dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\Num.One\, ... }) = SOME 1 | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\Num.Bit0\, ... } `$ t) = (case dest_num_code t of SOME n => SOME (2 * n) | _ => NONE) | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\Num.Bit1\, ... } `$ t) = (case dest_num_code t of SOME n => SOME (2 * n + 1) | _ => NONE) | dest_num_code _ = NONE; fun add_code number_of preproc print target thy = let fun pretty literals _ thm _ _ [(t, _)] = case dest_num_code t of SOME n => (Code_Printer.str o print literals o preproc) n | NONE => Code_Printer.eqn_error thy thm "Illegal numeral expression: illegal term"; in thy |> Code_Target.set_printings (Code_Symbol.Constant (number_of, [(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))])) end; end; (*local*) end; diff --git a/src/HOL/Tools/numeral_simprocs.ML b/src/HOL/Tools/numeral_simprocs.ML --- a/src/HOL/Tools/numeral_simprocs.ML +++ b/src/HOL/Tools/numeral_simprocs.ML @@ -1,741 +1,741 @@ (* Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 2000 University of Cambridge Simprocs for the (integer) numerals. *) (*To quote from Provers/Arith/cancel_numeral_factor.ML: Cancels common coefficients in balanced expressions: u*#m ~~ u'*#m' == #n*u ~~ #n'*u' where ~~ is an appropriate balancing operation (e.g. =, <=, <, div, /) and d = gcd(m,m') and n=m/d and n'=m'/d. *) signature NUMERAL_SIMPROCS = sig val trans_tac: Proof.context -> thm option -> tactic val assoc_fold: Proof.context -> cterm -> thm option val combine_numerals: Proof.context -> cterm -> thm option val eq_cancel_numerals: Proof.context -> cterm -> thm option val less_cancel_numerals: Proof.context -> cterm -> thm option val le_cancel_numerals: Proof.context -> cterm -> thm option val eq_cancel_factor: Proof.context -> cterm -> thm option val le_cancel_factor: Proof.context -> cterm -> thm option val less_cancel_factor: Proof.context -> cterm -> thm option val div_cancel_factor: Proof.context -> cterm -> thm option val mod_cancel_factor: Proof.context -> cterm -> thm option val dvd_cancel_factor: Proof.context -> cterm -> thm option val divide_cancel_factor: Proof.context -> cterm -> thm option val eq_cancel_numeral_factor: Proof.context -> cterm -> thm option val less_cancel_numeral_factor: Proof.context -> cterm -> thm option val le_cancel_numeral_factor: Proof.context -> cterm -> thm option val div_cancel_numeral_factor: Proof.context -> cterm -> thm option val divide_cancel_numeral_factor: Proof.context -> cterm -> thm option val field_combine_numerals: Proof.context -> cterm -> thm option val field_divide_cancel_numeral_factor: simproc val num_ss: simpset val field_comp_conv: Proof.context -> conv end; structure Numeral_Simprocs : NUMERAL_SIMPROCS = struct fun trans_tac _ NONE = all_tac | trans_tac ctxt (SOME th) = ALLGOALS (resolve_tac ctxt [th RS trans]); val mk_number = Arith_Data.mk_number; val mk_sum = Arith_Data.mk_sum; val long_mk_sum = Arith_Data.long_mk_sum; val dest_sum = Arith_Data.dest_sum; val mk_times = HOLogic.mk_binop \<^const_name>\Groups.times\; fun one_of T = Const(\<^const_name>\Groups.one\, T); (* build product with trailing 1 rather than Numeral 1 in order to avoid the unnecessary restriction to type class number_ring which is not required for cancellation of common factors in divisions. UPDATE: this reasoning no longer applies (number_ring is gone) *) fun mk_prod T = let val one = one_of T fun mk [] = one | mk [t] = t | mk (t :: ts) = if t = one then mk ts else mk_times (t, mk ts) in mk end; (*This version ALWAYS includes a trailing one*) fun long_mk_prod T [] = one_of T | long_mk_prod T (t :: ts) = mk_times (t, mk_prod T ts); val dest_times = HOLogic.dest_bin \<^const_name>\Groups.times\ dummyT; fun dest_prod t = let val (t,u) = dest_times t in dest_prod t @ dest_prod u end handle TERM _ => [t]; fun find_first_numeral past (t::terms) = ((snd (HOLogic.dest_number t), rev past @ terms) handle TERM _ => find_first_numeral (t::past) terms) | find_first_numeral past [] = raise TERM("find_first_numeral", []); (*DON'T do the obvious simplifications; that would create special cases*) fun mk_coeff (k, t) = mk_times (mk_number (Term.fastype_of t) k, t); (*Express t as a product of (possibly) a numeral with other sorted terms*) fun dest_coeff sign (Const (\<^const_name>\Groups.uminus\, _) $ t) = dest_coeff (~sign) t | dest_coeff sign t = let val ts = sort Term_Ord.term_ord (dest_prod t) val (n, ts') = find_first_numeral [] ts handle TERM _ => (1, ts) in (sign*n, mk_prod (Term.fastype_of t) ts') end; (*Find first coefficient-term THAT MATCHES u*) fun find_first_coeff past u [] = raise TERM("find_first_coeff", []) | find_first_coeff past u (t::terms) = let val (n,u') = dest_coeff 1 t in if u aconv u' then (n, rev past @ terms) else find_first_coeff (t::past) u terms end handle TERM _ => find_first_coeff (t::past) u terms; (*Fractions as pairs of ints. Can't use Rat.rat because the representation needs to preserve negative values in the denominator.*) fun mk_frac (p, q) = if q = 0 then raise Div else (p, q); (*Don't reduce fractions; sums must be proved by rule add_frac_eq. Fractions are reduced later by the cancel_numeral_factor simproc.*) fun add_frac ((p1, q1), (p2, q2)) = (p1 * q2 + p2 * q1, q1 * q2); val mk_divide = HOLogic.mk_binop \<^const_name>\Rings.divide\; (*Build term (p / q) * t*) fun mk_fcoeff ((p, q), t) = let val T = Term.fastype_of t in mk_times (mk_divide (mk_number T p, mk_number T q), t) end; (*Express t as a product of a fraction with other sorted terms*) fun dest_fcoeff sign (Const (\<^const_name>\Groups.uminus\, _) $ t) = dest_fcoeff (~sign) t | dest_fcoeff sign (Const (\<^const_name>\Rings.divide\, _) $ t $ u) = let val (p, t') = dest_coeff sign t val (q, u') = dest_coeff 1 u in (mk_frac (p, q), mk_divide (t', u')) end | dest_fcoeff sign t = let val (p, t') = dest_coeff sign t val T = Term.fastype_of t in (mk_frac (p, 1), mk_divide (t', one_of T)) end; (** New term ordering so that AC-rewriting brings numerals to the front **) (*Order integers by absolute value and then by sign. The standard integer ordering is not well-founded.*) fun num_ord (i,j) = (case int_ord (abs i, abs j) of EQUAL => int_ord (Int.sign i, Int.sign j) | ord => ord); (*This resembles Term_Ord.term_ord, but it puts binary numerals before other non-atomic terms.*) local open Term in fun numterm_ord (t, u) = case (try HOLogic.dest_number t, try HOLogic.dest_number u) of (SOME (_, i), SOME (_, j)) => num_ord (i, j) | (SOME _, NONE) => LESS | (NONE, SOME _) => GREATER | _ => ( case (t, u) of (Abs (_, T, t), Abs(_, U, u)) => (prod_ord numterm_ord Term_Ord.typ_ord ((t, T), (u, U))) | _ => ( case int_ord (size_of_term t, size_of_term u) of EQUAL => let val (f, ts) = strip_comb t and (g, us) = strip_comb u in (prod_ord Term_Ord.hd_ord numterms_ord ((f, ts), (g, us))) end | ord => ord)) and numterms_ord (ts, us) = list_ord numterm_ord (ts, us) end; val num_ss = simpset_of (put_simpset HOL_basic_ss \<^context> |> Simplifier.set_term_ord numterm_ord); (*Maps 1 to Numeral1 so that arithmetic isn't complicated by the abstract 1.*) val numeral_syms = [@{thm numeral_One} RS sym]; (*Simplify 0+n, n+0, Numeral1*n, n*Numeral1, 1*x, x*1, x/1 *) val add_0s = @{thms add_0_left add_0_right}; val mult_1s = @{thms mult_1s divide_numeral_1 mult_1_left mult_1_right mult_minus1 mult_minus1_right div_by_1}; (* For post-simplification of the rhs of simproc-generated rules *) val post_simps = [@{thm numeral_One}, @{thm add_0_left}, @{thm add_0_right}, @{thm mult_zero_left}, @{thm mult_zero_right}, @{thm mult_1_left}, @{thm mult_1_right}, @{thm mult_minus1}, @{thm mult_minus1_right}] val field_post_simps = post_simps @ [@{thm div_0}, @{thm div_by_1}] (*Simplify inverse Numeral1*) val inverse_1s = [@{thm inverse_numeral_1}]; (*To perform binary arithmetic. The "left" rewriting handles patterns created by the Numeral_Simprocs, such as 3 * (5 * x). *) val simps = [@{thm numeral_One} RS sym] @ @{thms add_numeral_left} @ @{thms add_neg_numeral_left} @ @{thms mult_numeral_left} @ @{thms arith_simps} @ @{thms rel_simps}; (*Binary arithmetic BUT NOT ADDITION since it may collapse adjacent terms during re-arrangement*) val non_add_simps = subtract Thm.eq_thm (@{thms add_numeral_left} @ @{thms add_neg_numeral_left} @ @{thms numeral_plus_numeral} @ @{thms add_neg_numeral_simps}) simps; (*To evaluate binary negations of coefficients*) val minus_simps = [@{thm minus_zero}, @{thm minus_minus}]; (*To let us treat subtraction as addition*) val diff_simps = [@{thm diff_conv_add_uminus}, @{thm minus_add_distrib}, @{thm minus_minus}]; (*To let us treat division as multiplication*) val divide_simps = [@{thm divide_inverse}, @{thm inverse_mult_distrib}, @{thm inverse_inverse_eq}]; (*to extract again any uncancelled minuses*) val minus_from_mult_simps = [@{thm minus_minus}, @{thm mult_minus_left}, @{thm mult_minus_right}]; (*combine unary minus with numeric literals, however nested within a product*) val mult_minus_simps = [@{thm mult.assoc}, @{thm minus_mult_right}, @{thm minus_mult_commute}, @{thm numeral_times_minus_swap}]; val norm_ss1 = simpset_of (put_simpset num_ss \<^context> addsimps numeral_syms @ add_0s @ mult_1s @ diff_simps @ minus_simps @ @{thms ac_simps}) val norm_ss2 = simpset_of (put_simpset num_ss \<^context> addsimps non_add_simps @ mult_minus_simps) val norm_ss3 = simpset_of (put_simpset num_ss \<^context> addsimps minus_from_mult_simps @ @{thms ac_simps} @ @{thms ac_simps minus_mult_commute}) structure CancelNumeralsCommon = struct val mk_sum = mk_sum val dest_sum = dest_sum val mk_coeff = mk_coeff val dest_coeff = dest_coeff 1 val find_first_coeff = find_first_coeff [] val trans_tac = trans_tac fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) val numeral_simp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps add_0s @ simps) fun numeral_simp_tac ctxt = ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps val prove_conv = Arith_Data.prove_conv end; structure EqCancelNumerals = CancelNumeralsFun (open CancelNumeralsCommon val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ dummyT val bal_add1 = @{thm eq_add_iff1} RS trans val bal_add2 = @{thm eq_add_iff2} RS trans ); structure LessCancelNumerals = CancelNumeralsFun (open CancelNumeralsCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ dummyT val bal_add1 = @{thm less_add_iff1} RS trans val bal_add2 = @{thm less_add_iff2} RS trans ); structure LeCancelNumerals = CancelNumeralsFun (open CancelNumeralsCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ dummyT val bal_add1 = @{thm le_add_iff1} RS trans val bal_add2 = @{thm le_add_iff2} RS trans ); val eq_cancel_numerals = EqCancelNumerals.proc val less_cancel_numerals = LessCancelNumerals.proc val le_cancel_numerals = LeCancelNumerals.proc structure CombineNumeralsData = struct type coeff = int val iszero = (fn x => x = 0) val add = op + val mk_sum = long_mk_sum (*to work for e.g. 2*x + 3*x *) val dest_sum = dest_sum val mk_coeff = mk_coeff val dest_coeff = dest_coeff 1 val left_distrib = @{thm combine_common_factor} RS trans val prove_conv = Arith_Data.prove_conv_nohyps val trans_tac = trans_tac fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) val numeral_simp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps add_0s @ simps) fun numeral_simp_tac ctxt = ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq post_simps end; structure CombineNumerals = CombineNumeralsFun(CombineNumeralsData); (*Version for fields, where coefficients can be fractions*) structure FieldCombineNumeralsData = struct type coeff = int * int val iszero = (fn (p, _) => p = 0) val add = add_frac val mk_sum = long_mk_sum val dest_sum = dest_sum val mk_coeff = mk_fcoeff val dest_coeff = dest_fcoeff 1 val left_distrib = @{thm combine_common_factor} RS trans val prove_conv = Arith_Data.prove_conv_nohyps val trans_tac = trans_tac val norm_ss1a = simpset_of (put_simpset norm_ss1 \<^context> addsimps inverse_1s @ divide_simps) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss1a ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) val numeral_simp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps add_0s @ simps @ [@{thm add_frac_eq}, @{thm not_False_eq_True}]) fun numeral_simp_tac ctxt = ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq field_post_simps end; structure FieldCombineNumerals = CombineNumeralsFun(FieldCombineNumeralsData); val combine_numerals = CombineNumerals.proc val field_combine_numerals = FieldCombineNumerals.proc (** Constant folding for multiplication in semirings **) (*We do not need folding for addition: combine_numerals does the same thing*) structure Semiring_Times_Assoc_Data : ASSOC_FOLD_DATA = struct val assoc_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms ac_simps minus_mult_commute}) val eq_reflection = eq_reflection val is_numeral = can HOLogic.dest_number end; structure Semiring_Times_Assoc = Assoc_Fold (Semiring_Times_Assoc_Data); fun assoc_fold ctxt ct = Semiring_Times_Assoc.proc ctxt (Thm.term_of ct) structure CancelNumeralFactorCommon = struct val mk_coeff = mk_coeff val dest_coeff = dest_coeff 1 val trans_tac = trans_tac val norm_ss1 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps minus_from_mult_simps @ mult_1s) val norm_ss2 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps simps @ mult_minus_simps) val norm_ss3 = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms ac_simps minus_mult_commute numeral_times_minus_swap}) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss1 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss2 ctxt)) THEN ALLGOALS (simp_tac (put_simpset norm_ss3 ctxt)) (* simp_thms are necessary because some of the cancellation rules below (e.g. mult_less_cancel_left) introduce various logical connectives *) val numeral_simp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps simps @ @{thms simp_thms}) fun numeral_simp_tac ctxt = ALLGOALS (simp_tac (put_simpset numeral_simp_ss ctxt)) val simplify_meta_eq = Arith_Data.simplify_meta_eq ([@{thm Nat.add_0}, @{thm Nat.add_0_right}] @ post_simps) val prove_conv = Arith_Data.prove_conv end (*Version for semiring_div*) structure DivCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val mk_bal = HOLogic.mk_binop \<^const_name>\Rings.divide\ val dest_bal = HOLogic.dest_bin \<^const_name>\Rings.divide\ dummyT val cancel = @{thm div_mult_mult1} RS trans val neg_exchanges = false ) (*Version for fields*) structure DivideCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val mk_bal = HOLogic.mk_binop \<^const_name>\Rings.divide\ val dest_bal = HOLogic.dest_bin \<^const_name>\Rings.divide\ dummyT val cancel = @{thm mult_divide_mult_cancel_left} RS trans val neg_exchanges = false ) structure EqCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ dummyT val cancel = @{thm mult_cancel_left} RS trans val neg_exchanges = false ) structure LessCancelNumeralFactor = CancelNumeralFactorFun (open CancelNumeralFactorCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ dummyT val cancel = @{thm mult_less_cancel_left} RS trans val neg_exchanges = true ) structure LeCancelNumeralFactor = CancelNumeralFactorFun ( open CancelNumeralFactorCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ dummyT val cancel = @{thm mult_le_cancel_left} RS trans val neg_exchanges = true ) val eq_cancel_numeral_factor = EqCancelNumeralFactor.proc val less_cancel_numeral_factor = LessCancelNumeralFactor.proc val le_cancel_numeral_factor = LeCancelNumeralFactor.proc val div_cancel_numeral_factor = DivCancelNumeralFactor.proc val divide_cancel_numeral_factor = DivideCancelNumeralFactor.proc val field_divide_cancel_numeral_factor = Simplifier.make_simproc \<^context> "field_divide_cancel_numeral_factor" {lhss = [\<^term>\((l::'a::field) * m) / n\, \<^term>\(l::'a::field) / (m * n)\, \<^term>\((numeral v)::'a::field) / (numeral w)\, \<^term>\((numeral v)::'a::field) / (- numeral w)\, \<^term>\((- numeral v)::'a::field) / (numeral w)\, \<^term>\((- numeral v)::'a::field) / (- numeral w)\], proc = K DivideCancelNumeralFactor.proc} val field_cancel_numeral_factors = [Simplifier.make_simproc \<^context> "field_eq_cancel_numeral_factor" {lhss = [\<^term>\(l::'a::field) * m = n\, \<^term>\(l::'a::field) = m * n\], proc = K EqCancelNumeralFactor.proc}, field_divide_cancel_numeral_factor] (** Declarations for ExtractCommonTerm **) (*Find first term that matches u*) fun find_first_t past u [] = raise TERM ("find_first_t", []) | find_first_t past u (t::terms) = if u aconv t then (rev past @ terms) else find_first_t (t::past) u terms handle TERM _ => find_first_t (t::past) u terms; (** Final simplification for the CancelFactor simprocs **) val simplify_one = Arith_Data.simplify_meta_eq [@{thm mult_1_left}, @{thm mult_1_right}, @{thm div_by_1}, @{thm numeral_One}]; fun cancel_simplify_meta_eq ctxt cancel_th th = simplify_one ctxt (([th, cancel_th]) MRS trans); local val Tp_Eq = Thm.reflexive (Thm.cterm_of \<^theory_context>\HOL\ HOLogic.Trueprop) fun Eq_True_elim Eq = Thm.equal_elim (Thm.combination Tp_Eq (Thm.symmetric Eq)) @{thm TrueI} in fun sign_conv pos_th neg_th ctxt t = let val T = fastype_of t; val zero = Const(\<^const_name>\Groups.zero\, T); val less = Const(\<^const_name>\Orderings.less\, [T,T] ---> HOLogic.boolT); val pos = less $ zero $ t and neg = less $ t $ zero fun prove p = SOME (Eq_True_elim (Simplifier.asm_rewrite ctxt (Thm.cterm_of ctxt p))) handle THM _ => NONE in case prove pos of SOME th => SOME(th RS pos_th) | NONE => (case prove neg of SOME th => SOME(th RS neg_th) | NONE => NONE) end; end structure CancelFactorCommon = struct val mk_sum = long_mk_prod val dest_sum = dest_prod val mk_coeff = mk_coeff val dest_coeff = dest_coeff val find_first = find_first_t [] val trans_tac = trans_tac val norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps mult_1s @ @{thms ac_simps minus_mult_commute}) fun norm_tac ctxt = ALLGOALS (simp_tac (put_simpset norm_ss ctxt)) val simplify_meta_eq = cancel_simplify_meta_eq fun mk_eq (a, b) = HOLogic.mk_Trueprop (HOLogic.mk_eq (a, b)) end; (*mult_cancel_left requires a ring with no zero divisors.*) structure EqCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_eq val dest_bal = HOLogic.dest_bin \<^const_name>\HOL.eq\ dummyT fun simp_conv _ _ = SOME @{thm mult_cancel_left} ); (*for ordered rings*) structure LeCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less_eq\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less_eq\ dummyT val simp_conv = sign_conv @{thm mult_le_cancel_left_pos} @{thm mult_le_cancel_left_neg} ); (*for ordered rings*) structure LessCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Orderings.less\ val dest_bal = HOLogic.dest_bin \<^const_name>\Orderings.less\ dummyT val simp_conv = sign_conv @{thm mult_less_cancel_left_pos} @{thm mult_less_cancel_left_neg} ); (*for semirings with division*) structure DivCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binop \<^const_name>\Rings.divide\ val dest_bal = HOLogic.dest_bin \<^const_name>\Rings.divide\ dummyT fun simp_conv _ _ = SOME @{thm div_mult_mult1_if} ); structure ModCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binop \<^const_name>\modulo\ val dest_bal = HOLogic.dest_bin \<^const_name>\modulo\ dummyT fun simp_conv _ _ = SOME @{thm mod_mult_mult1} ); (*for idoms*) structure DvdCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binrel \<^const_name>\Rings.dvd\ val dest_bal = HOLogic.dest_bin \<^const_name>\Rings.dvd\ dummyT fun simp_conv _ _ = SOME @{thm dvd_mult_cancel_left} ); (*Version for all fields, including unordered ones (type complex).*) structure DivideCancelFactor = ExtractCommonTermFun (open CancelFactorCommon val mk_bal = HOLogic.mk_binop \<^const_name>\Rings.divide\ val dest_bal = HOLogic.dest_bin \<^const_name>\Rings.divide\ dummyT fun simp_conv _ _ = SOME @{thm mult_divide_mult_cancel_left_if} ); fun eq_cancel_factor ctxt ct = EqCancelFactor.proc ctxt (Thm.term_of ct) fun le_cancel_factor ctxt ct = LeCancelFactor.proc ctxt (Thm.term_of ct) fun less_cancel_factor ctxt ct = LessCancelFactor.proc ctxt (Thm.term_of ct) fun div_cancel_factor ctxt ct = DivCancelFactor.proc ctxt (Thm.term_of ct) fun mod_cancel_factor ctxt ct = ModCancelFactor.proc ctxt (Thm.term_of ct) fun dvd_cancel_factor ctxt ct = DvdCancelFactor.proc ctxt (Thm.term_of ct) fun divide_cancel_factor ctxt ct = DivideCancelFactor.proc ctxt (Thm.term_of ct) local val cterm_of = Thm.cterm_of \<^context>; fun tvar S = (("'a", 0), S); val zero_tvar = tvar \<^sort>\zero\; val zero = cterm_of (Const (\<^const_name>\zero_class.zero\, TVar zero_tvar)); val type_tvar = tvar \<^sort>\type\; val geq = cterm_of (Const (\<^const_name>\HOL.eq\, TVar type_tvar --> TVar type_tvar --> \<^typ>\bool\)); val add_frac_eq = mk_meta_eq @{thm "add_frac_eq"} val add_frac_num = mk_meta_eq @{thm "add_frac_num"} val add_num_frac = mk_meta_eq @{thm "add_num_frac"} fun prove_nz ctxt T t = let - val z = Thm.instantiate_cterm ([(zero_tvar, T)], []) zero - val eq = Thm.instantiate_cterm ([(type_tvar, T)], []) geq + val z = Thm.instantiate_cterm (TVars.make [(zero_tvar, T)], Vars.empty) zero + val eq = Thm.instantiate_cterm (TVars.make [(type_tvar, T)], Vars.empty) geq val th = Simplifier.rewrite (ctxt addsimps @{thms simp_thms}) (Thm.apply \<^cterm>\Trueprop\ (Thm.apply \<^cterm>\Not\ (Thm.apply (Thm.apply eq t) z))) in Thm.equal_elim (Thm.symmetric th) TrueI end fun proc ctxt ct = let val ((x,y),(w,z)) = (Thm.dest_binop #> (fn (a,b) => (Thm.dest_binop a, Thm.dest_binop b))) ct val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z,w] val T = Thm.ctyp_of_cterm x val [y_nz, z_nz] = map (prove_nz ctxt T) [y, z] val th = Thm.instantiate' [SOME T] (map SOME [y,z,x,w]) add_frac_eq in SOME (Thm.implies_elim (Thm.implies_elim th y_nz) z_nz) end handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE fun proc2 ctxt ct = let val (l,r) = Thm.dest_binop ct val T = Thm.ctyp_of_cterm l in (case (Thm.term_of l, Thm.term_of r) of (Const(\<^const_name>\Rings.divide\,_)$_$_, _) => let val (x,y) = Thm.dest_binop l val z = r val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z] val ynz = prove_nz ctxt T y in SOME (Thm.implies_elim (Thm.instantiate' [SOME T] (map SOME [y,x,z]) add_frac_num) ynz) end | (_, Const (\<^const_name>\Rings.divide\,_)$_$_) => let val (x,y) = Thm.dest_binop r val z = l val _ = map (HOLogic.dest_number o Thm.term_of) [x,y,z] val ynz = prove_nz ctxt T y in SOME (Thm.implies_elim (Thm.instantiate' [SOME T] (map SOME [y,z,x]) add_num_frac) ynz) end | _ => NONE) end handle CTERM _ => NONE | TERM _ => NONE | THM _ => NONE fun is_number (Const(\<^const_name>\Rings.divide\,_)$a$b) = is_number a andalso is_number b | is_number t = can HOLogic.dest_number t val is_number = is_number o Thm.term_of fun proc3 ctxt ct = (case Thm.term_of ct of Const(\<^const_name>\Orderings.less\,_)$(Const(\<^const_name>\Rings.divide\,_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_less_eq"} in SOME (mk_meta_eq th) end | Const(\<^const_name>\Orderings.less_eq\,_)$(Const(\<^const_name>\Rings.divide\,_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_le_eq"} in SOME (mk_meta_eq th) end | Const(\<^const_name>\HOL.eq\,_)$(Const(\<^const_name>\Rings.divide\,_)$_$_)$_ => let val ((a,b),c) = Thm.dest_binop ct |>> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "divide_eq_eq"} in SOME (mk_meta_eq th) end | Const(\<^const_name>\Orderings.less\,_)$_$(Const(\<^const_name>\Rings.divide\,_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "less_divide_eq"} in SOME (mk_meta_eq th) end | Const(\<^const_name>\Orderings.less_eq\,_)$_$(Const(\<^const_name>\Rings.divide\,_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "le_divide_eq"} in SOME (mk_meta_eq th) end | Const(\<^const_name>\HOL.eq\,_)$_$(Const(\<^const_name>\Rings.divide\,_)$_$_) => let val (a,(b,c)) = Thm.dest_binop ct ||> Thm.dest_binop val _ = map is_number [a,b,c] val T = Thm.ctyp_of_cterm c val th = Thm.instantiate' [SOME T] (map SOME [a,b,c]) @{thm "eq_divide_eq"} in SOME (mk_meta_eq th) end | _ => NONE) handle TERM _ => NONE | CTERM _ => NONE | THM _ => NONE val add_frac_frac_simproc = Simplifier.make_simproc \<^context> "add_frac_frac_simproc" {lhss = [\<^term>\(x::'a::field) / y + (w::'a::field) / z\], proc = K proc} val add_frac_num_simproc = Simplifier.make_simproc \<^context> "add_frac_num_simproc" {lhss = [\<^term>\(x::'a::field) / y + z\, \<^term>\z + (x::'a::field) / y\], proc = K proc2} val ord_frac_simproc = Simplifier.make_simproc \<^context> "ord_frac_simproc" {lhss = [\<^term>\(a::'a::{field,ord}) / b < c\, \<^term>\(a::'a::{field,ord}) / b \ c\, \<^term>\c < (a::'a::{field,ord}) / b\, \<^term>\c \ (a::'a::{field,ord}) / b\, \<^term>\c = (a::'a::{field,ord}) / b\, \<^term>\(a::'a::{field, ord}) / b = c\], proc = K proc3} val ths = [@{thm "mult_numeral_1"}, @{thm "mult_numeral_1_right"}, @{thm "divide_numeral_1"}, @{thm "div_by_0"}, @{thm div_0}, @{thm "divide_divide_eq_left"}, @{thm "times_divide_eq_left"}, @{thm "times_divide_eq_right"}, @{thm "times_divide_times_eq"}, @{thm "divide_divide_eq_right"}, @{thm diff_conv_add_uminus}, @{thm "minus_divide_left"}, @{thm "add_divide_distrib"} RS sym, @{thm Fields.field_divide_inverse} RS sym, @{thm inverse_divide}, Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Conv.rewr_conv (mk_meta_eq @{thm mult.commute})))) (@{thm Fields.field_divide_inverse} RS sym)] val field_comp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms "semiring_norm"} addsimps ths addsimps @{thms simp_thms} addsimprocs field_cancel_numeral_factors addsimprocs [add_frac_frac_simproc, add_frac_num_simproc, ord_frac_simproc] |> Simplifier.add_cong @{thm "if_weak_cong"}) in fun field_comp_conv ctxt = Simplifier.rewrite (put_simpset field_comp_ss ctxt) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One}]) end end; diff --git a/src/HOL/Tools/record.ML b/src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML +++ b/src/HOL/Tools/record.ML @@ -1,2425 +1,2426 @@ (* Title: HOL/Tools/record.ML Author: Wolfgang Naraschewski, TU Muenchen Author: Markus Wenzel, TU Muenchen Author: Norbert Schirmer, TU Muenchen Author: Thomas Sewell, NICTA Extensible records with structural subtyping. *) signature RECORD = sig val type_abbr: bool Config.T val type_as_fields: bool Config.T val timing: bool Config.T type info = {args: (string * sort) list, parent: (typ list * string) option, fields: (string * typ) list, extension: (string * typ list), ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm, select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list, fold_congs: thm list, unfold_congs: thm list, splits: thm list, defs: thm list, surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm, cases: thm, simps: thm list, iffs: thm list} val get_info: theory -> string -> info option val the_info: theory -> string -> info val get_hierarchy: theory -> (string * typ list) -> (string * ((string * sort) * typ) list) list val add_record: {overloaded: bool} -> (string * sort) list * binding -> (typ list * string) option -> (binding * typ * mixfix) list -> theory -> theory val last_extT: typ -> (string * typ list) option val dest_recTs: typ -> (string * typ list) list val get_extT_fields: theory -> typ -> (string * typ) list * (string * typ) val get_recT_fields: theory -> typ -> (string * typ) list * (string * typ) val get_parent: theory -> string -> (typ list * string) option val get_extension: theory -> string -> (string * typ list) option val get_extinjects: theory -> thm list val get_simpset: theory -> simpset val simproc: simproc val eq_simproc: simproc val upd_simproc: simproc val split_simproc: (term -> int) -> simproc val ex_sel_eq_simproc: simproc val split_tac: Proof.context -> int -> tactic val split_simp_tac: Proof.context -> thm list -> (term -> int) -> int -> tactic val split_wrapper: string * (Proof.context -> wrapper) val pretty_recT: Proof.context -> typ -> Pretty.T val string_of_record: Proof.context -> string -> string val codegen: bool Config.T val updateN: string val ext_typeN: string val extN: string end; signature ISO_TUPLE_SUPPORT = sig val add_iso_tuple_type: {overloaded: bool} -> binding * (string * sort) list -> typ * typ -> theory -> (term * term) * theory val mk_cons_tuple: term * term -> term val dest_cons_tuple: term -> term * term val iso_tuple_intros_tac: Proof.context -> int -> tactic end; structure Iso_Tuple_Support: ISO_TUPLE_SUPPORT = struct val isoN = "_Tuple_Iso"; val iso_tuple_intro = @{thm isomorphic_tuple_intro}; val iso_tuple_intros = Tactic.build_net @{thms isomorphic_tuple.intros}; val tuple_iso_tuple = (\<^const_name>\Record.tuple_iso_tuple\, @{thm tuple_iso_tuple}); structure Iso_Tuple_Thms = Theory_Data ( type T = thm Symtab.table; val empty = Symtab.make [tuple_iso_tuple]; val extend = I; fun merge data = Symtab.merge Thm.eq_thm_prop data; (* FIXME handle Symtab.DUP ?? *) ); fun get_typedef_info tyco vs (({rep_type, Abs_name, ...}, {Rep_inject, Abs_inverse, ... }) : Typedef.info) thy = let val exists_thm = UNIV_I |> Thm.instantiate' [SOME (Thm.global_ctyp_of thy (Logic.varifyT_global rep_type))] []; val proj_constr = Abs_inverse OF [exists_thm]; val absT = Type (tyco, map TFree vs); in thy |> pair (tyco, ((Rep_inject, proj_constr), Const (Abs_name, rep_type --> absT), absT)) end fun do_typedef overloaded raw_tyco repT raw_vs thy = let val ctxt = Proof_Context.init_global thy |> Variable.declare_typ repT; val vs = map (Proof_Context.check_tfree ctxt) raw_vs; in thy |> Named_Target.theory_map_result (apsnd o Typedef.transform_info) (Typedef.add_typedef overloaded (raw_tyco, vs, NoSyn) (HOLogic.mk_UNIV repT) NONE (fn ctxt' => resolve_tac ctxt' [UNIV_witness] 1)) |-> (fn (tyco, info) => get_typedef_info tyco vs info) end; fun mk_cons_tuple (left, right) = let val (leftT, rightT) = (fastype_of left, fastype_of right); val prodT = HOLogic.mk_prodT (leftT, rightT); val isomT = Type (\<^type_name>\tuple_isomorphism\, [prodT, leftT, rightT]); in Const (\<^const_name>\Record.iso_tuple_cons\, isomT --> leftT --> rightT --> prodT) $ Const (fst tuple_iso_tuple, isomT) $ left $ right end; fun dest_cons_tuple (Const (\<^const_name>\Record.iso_tuple_cons\, _) $ Const _ $ t $ u) = (t, u) | dest_cons_tuple t = raise TERM ("dest_cons_tuple", [t]); fun add_iso_tuple_type overloaded (b, alphas) (leftT, rightT) thy = let val repT = HOLogic.mk_prodT (leftT, rightT); val ((_, ((rep_inject, abs_inverse), absC, absT)), typ_thy) = thy |> do_typedef overloaded b repT alphas ||> Sign.add_path (Binding.name_of b); (*FIXME proper prefixing instead*) val typ_ctxt = Proof_Context.init_global typ_thy; (*construct a type and body for the isomorphism constant by instantiating the theorem to which the definition will be applied*) val intro_inst = rep_inject RS infer_instantiate typ_ctxt [(("abst", 0), Thm.cterm_of typ_ctxt absC)] iso_tuple_intro; val (_, body) = Logic.dest_equals (List.last (Thm.prems_of intro_inst)); val isomT = fastype_of body; val isom_binding = Binding.suffix_name isoN b; val isom_name = Sign.full_name typ_thy isom_binding; val isom = Const (isom_name, isomT); val ([isom_def], cdef_thy) = typ_thy |> Sign.declare_const_global ((isom_binding, isomT), NoSyn) |> snd |> Global_Theory.add_defs false [((Binding.concealed (Thm.def_binding isom_binding), Logic.mk_equals (isom, body)), [])]; val iso_tuple = isom_def RS (abs_inverse RS (rep_inject RS iso_tuple_intro)); val cons = Const (\<^const_name>\Record.iso_tuple_cons\, isomT --> leftT --> rightT --> absT); val thm_thy = cdef_thy |> Iso_Tuple_Thms.map (Symtab.insert Thm.eq_thm_prop (isom_name, iso_tuple)) |> Sign.restore_naming thy in ((isom, cons $ isom), thm_thy) end; fun iso_tuple_intros_tac ctxt = resolve_from_net_tac ctxt iso_tuple_intros THEN' CSUBGOAL (fn (cgoal, i) => let val goal = Thm.term_of cgoal; val isthms = Iso_Tuple_Thms.get (Proof_Context.theory_of ctxt); fun err s t = raise TERM ("iso_tuple_intros_tac: " ^ s, [t]); val goal' = Envir.beta_eta_contract goal; val is = (case goal' of Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\isomorphic_tuple\, _) $ Const is) => is | _ => err "unexpected goal format" goal'); val isthm = (case Symtab.lookup isthms (#1 is) of SOME isthm => isthm | NONE => err "no thm found for constant" (Const is)); in resolve_tac ctxt [isthm] i end); end; structure Record: RECORD = struct val surject_assistI = @{thm iso_tuple_surjective_proof_assistI}; val surject_assist_idE = @{thm iso_tuple_surjective_proof_assist_idE}; val updacc_accessor_eqE = @{thm update_accessor_accessor_eqE}; val updacc_updator_eqE = @{thm update_accessor_updator_eqE}; val updacc_eq_idI = @{thm iso_tuple_update_accessor_eq_assist_idI}; val updacc_eq_triv = @{thm iso_tuple_update_accessor_eq_assist_triv}; val updacc_foldE = @{thm update_accessor_congruence_foldE}; val updacc_unfoldE = @{thm update_accessor_congruence_unfoldE}; val updacc_noopE = @{thm update_accessor_noopE}; val updacc_noop_compE = @{thm update_accessor_noop_compE}; val updacc_cong_idI = @{thm update_accessor_cong_assist_idI}; val updacc_cong_triv = @{thm update_accessor_cong_assist_triv}; val updacc_cong_from_eq = @{thm iso_tuple_update_accessor_cong_from_eq}; val codegen = Attrib.setup_config_bool \<^binding>\record_codegen\ (K true); (** name components **) val rN = "r"; val wN = "w"; val moreN = "more"; val schemeN = "_scheme"; val ext_typeN = "_ext"; val inner_typeN = "_inner"; val extN ="_ext"; val updateN = "_update"; val makeN = "make"; val fields_selN = "fields"; val extendN = "extend"; val truncateN = "truncate"; (*** utilities ***) fun varifyT idx = map_type_tfree (fn (a, S) => TVar ((a, idx), S)); (* timing *) val timing = Attrib.setup_config_bool \<^binding>\record_timing\ (K false); fun timeit_msg ctxt s x = if Config.get ctxt timing then (warning s; timeit x) else x (); fun timing_msg ctxt s = if Config.get ctxt timing then warning s else (); (* syntax *) val Trueprop = HOLogic.mk_Trueprop; infix 0 :== ===; infixr 0 ==>; val op :== = Misc_Legacy.mk_defpair; val op === = Trueprop o HOLogic.mk_eq; val op ==> = Logic.mk_implies; (* constructor *) fun mk_ext (name, T) ts = let val Ts = map fastype_of ts in list_comb (Const (suffix extN name, Ts ---> T), ts) end; (* selector *) fun mk_selC sT (c, T) = (c, sT --> T); fun mk_sel s (c, T) = let val sT = fastype_of s in Const (mk_selC sT (c, T)) $ s end; (* updates *) fun mk_updC sfx sT (c, T) = (suffix sfx c, (T --> T) --> sT --> sT); fun mk_upd' sfx c v sT = let val vT = domain_type (fastype_of v); in Const (mk_updC sfx sT (c, vT)) $ v end; fun mk_upd sfx c v s = mk_upd' sfx c v (fastype_of s) $ s; (* types *) fun dest_recT (typ as Type (c_ext_type, Ts as (_ :: _))) = (case try (unsuffix ext_typeN) c_ext_type of NONE => raise TYPE ("Record.dest_recT", [typ], []) | SOME c => ((c, Ts), List.last Ts)) | dest_recT typ = raise TYPE ("Record.dest_recT", [typ], []); val is_recT = can dest_recT; fun dest_recTs T = let val ((c, Ts), U) = dest_recT T in (c, Ts) :: dest_recTs U end handle TYPE _ => []; fun last_extT T = let val ((c, Ts), U) = dest_recT T in (case last_extT U of NONE => SOME (c, Ts) | SOME l => SOME l) end handle TYPE _ => NONE; fun rec_id i T = let val rTs = dest_recTs T; val rTs' = if i < 0 then rTs else take i rTs; in implode (map #1 rTs') end; (*** extend theory by record definition ***) (** record info **) (* type info and parent_info *) type info = {args: (string * sort) list, parent: (typ list * string) option, fields: (string * typ) list, extension: (string * typ list), ext_induct: thm, ext_inject: thm, ext_surjective: thm, ext_split: thm, ext_def: thm, select_convs: thm list, update_convs: thm list, select_defs: thm list, update_defs: thm list, fold_congs: thm list, (* potentially used in L4.verified *) unfold_congs: thm list, (* potentially used in L4.verified *) splits: thm list, defs: thm list, surjective: thm, equality: thm, induct_scheme: thm, induct: thm, cases_scheme: thm, cases: thm, simps: thm list, iffs: thm list}; fun make_info args parent fields extension ext_induct ext_inject ext_surjective ext_split ext_def select_convs update_convs select_defs update_defs fold_congs unfold_congs splits defs surjective equality induct_scheme induct cases_scheme cases simps iffs : info = {args = args, parent = parent, fields = fields, extension = extension, ext_induct = ext_induct, ext_inject = ext_inject, ext_surjective = ext_surjective, ext_split = ext_split, ext_def = ext_def, select_convs = select_convs, update_convs = update_convs, select_defs = select_defs, update_defs = update_defs, fold_congs = fold_congs, unfold_congs = unfold_congs, splits = splits, defs = defs, surjective = surjective, equality = equality, induct_scheme = induct_scheme, induct = induct, cases_scheme = cases_scheme, cases = cases, simps = simps, iffs = iffs}; type parent_info = {name: string, fields: (string * typ) list, extension: (string * typ list), induct_scheme: thm, ext_def: thm}; fun make_parent_info name fields extension ext_def induct_scheme : parent_info = {name = name, fields = fields, extension = extension, ext_def = ext_def, induct_scheme = induct_scheme}; (* theory data *) type data = {records: info Symtab.table, sel_upd: {selectors: (int * bool) Symtab.table, updates: string Symtab.table, simpset: simpset, defset: simpset}, equalities: thm Symtab.table, extinjects: thm list, extsplit: thm Symtab.table, (*maps extension name to split rule*) splits: (thm * thm * thm * thm) Symtab.table, (*!!, ALL, EX - split-equalities, induct rule*) extfields: (string * typ) list Symtab.table, (*maps extension to its fields*) fieldext: (string * typ list) Symtab.table}; (*maps field to its extension*) fun make_data records sel_upd equalities extinjects extsplit splits extfields fieldext = {records = records, sel_upd = sel_upd, equalities = equalities, extinjects=extinjects, extsplit = extsplit, splits = splits, extfields = extfields, fieldext = fieldext }: data; structure Data = Theory_Data ( type T = data; val empty = make_data Symtab.empty {selectors = Symtab.empty, updates = Symtab.empty, simpset = HOL_basic_ss, defset = HOL_basic_ss} Symtab.empty [] Symtab.empty Symtab.empty Symtab.empty Symtab.empty; val extend = I; fun merge ({records = recs1, sel_upd = {selectors = sels1, updates = upds1, simpset = ss1, defset = ds1}, equalities = equalities1, extinjects = extinjects1, extsplit = extsplit1, splits = splits1, extfields = extfields1, fieldext = fieldext1}, {records = recs2, sel_upd = {selectors = sels2, updates = upds2, simpset = ss2, defset = ds2}, equalities = equalities2, extinjects = extinjects2, extsplit = extsplit2, splits = splits2, extfields = extfields2, fieldext = fieldext2}) = make_data (Symtab.merge (K true) (recs1, recs2)) {selectors = Symtab.merge (K true) (sels1, sels2), updates = Symtab.merge (K true) (upds1, upds2), simpset = Simplifier.merge_ss (ss1, ss2), defset = Simplifier.merge_ss (ds1, ds2)} (Symtab.merge Thm.eq_thm_prop (equalities1, equalities2)) (Thm.merge_thms (extinjects1, extinjects2)) (Symtab.merge Thm.eq_thm_prop (extsplit1, extsplit2)) (Symtab.merge (fn ((a, b, c, d), (w, x, y, z)) => Thm.eq_thm (a, w) andalso Thm.eq_thm (b, x) andalso Thm.eq_thm (c, y) andalso Thm.eq_thm (d, z)) (splits1, splits2)) (Symtab.merge (K true) (extfields1, extfields2)) (Symtab.merge (K true) (fieldext1, fieldext2)); ); (* access 'records' *) val get_info = Symtab.lookup o #records o Data.get; fun the_info thy name = (case get_info thy name of SOME info => info | NONE => error ("Unknown record type " ^ quote name)); fun put_record name info = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data (Symtab.update (name, info) records) sel_upd equalities extinjects extsplit splits extfields fieldext); (* access 'sel_upd' *) val get_sel_upd = #sel_upd o Data.get; val is_selector = Symtab.defined o #selectors o get_sel_upd; val get_updates = Symtab.lookup o #updates o get_sel_upd; val get_simpset = #simpset o get_sel_upd; val get_sel_upd_defs = #defset o get_sel_upd; fun get_update_details u thy = let val sel_upd = get_sel_upd thy in (case Symtab.lookup (#updates sel_upd) u of SOME s => let val SOME (dep, ismore) = Symtab.lookup (#selectors sel_upd) s in SOME (s, dep, ismore) end | NONE => NONE) end; fun put_sel_upd names more depth simps defs thy = let val ctxt0 = Proof_Context.init_global thy; val all = names @ [more]; val sels = map (rpair (depth, false)) names @ [(more, (depth, true))]; val upds = map (suffix updateN) all ~~ all; val {records, sel_upd = {selectors, updates, simpset, defset}, equalities, extinjects, extsplit, splits, extfields, fieldext} = Data.get thy; val data = make_data records {selectors = fold Symtab.update_new sels selectors, updates = fold Symtab.update_new upds updates, simpset = simpset_map ctxt0 (fn ctxt => ctxt addsimps simps) simpset, defset = simpset_map ctxt0 (fn ctxt => ctxt addsimps defs) defset} equalities extinjects extsplit splits extfields fieldext; in Data.put data thy end; (* access 'equalities' *) fun add_equalities name thm = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data records sel_upd (Symtab.update_new (name, thm) equalities) extinjects extsplit splits extfields fieldext); val get_equalities = Symtab.lookup o #equalities o Data.get; (* access 'extinjects' *) fun add_extinjects thm = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data records sel_upd equalities (insert Thm.eq_thm_prop thm extinjects) extsplit splits extfields fieldext); val get_extinjects = rev o #extinjects o Data.get; (* access 'extsplit' *) fun add_extsplit name thm = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data records sel_upd equalities extinjects (Symtab.update_new (name, thm) extsplit) splits extfields fieldext); (* access 'splits' *) fun add_splits name thmP = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data records sel_upd equalities extinjects extsplit (Symtab.update_new (name, thmP) splits) extfields fieldext); val get_splits = Symtab.lookup o #splits o Data.get; (* parent/extension of named record *) val get_parent = (Option.join o Option.map #parent) oo (Symtab.lookup o #records o Data.get); val get_extension = Option.map #extension oo (Symtab.lookup o #records o Data.get); (* access 'extfields' *) fun add_extfields name fields = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => make_data records sel_upd equalities extinjects extsplit splits (Symtab.update_new (name, fields) extfields) fieldext); val get_extfields = Symtab.lookup o #extfields o Data.get; fun get_extT_fields thy T = let val ((name, Ts), moreT) = dest_recT T; val recname = let val (nm :: _ :: rst) = rev (Long_Name.explode name) (* FIXME !? *) in Long_Name.implode (rev (nm :: rst)) end; val varifyT = varifyT (maxidx_of_typs (moreT :: Ts) + 1); val {records, extfields, ...} = Data.get thy; val (fields, (more, _)) = split_last (Symtab.lookup_list extfields name); val args = map varifyT (snd (#extension (the (Symtab.lookup records recname)))); val subst = fold (Sign.typ_match thy) (#1 (split_last args) ~~ #1 (split_last Ts)) Vartab.empty; val fields' = map (apsnd (Envir.norm_type subst o varifyT)) fields; in (fields', (more, moreT)) end; fun get_recT_fields thy T = let val (root_fields, (root_more, root_moreT)) = get_extT_fields thy T; val (rest_fields, rest_more) = if is_recT root_moreT then get_recT_fields thy root_moreT else ([], (root_more, root_moreT)); in (root_fields @ rest_fields, rest_more) end; (* access 'fieldext' *) fun add_fieldext extname_types fields = Data.map (fn {records, sel_upd, equalities, extinjects, extsplit, splits, extfields, fieldext} => let val fieldext' = fold (fn field => Symtab.update_new (field, extname_types)) fields fieldext; in make_data records sel_upd equalities extinjects extsplit splits extfields fieldext' end); val get_fieldext = Symtab.lookup o #fieldext o Data.get; (* parent records *) local fun add_parents _ NONE = I | add_parents thy (SOME (types, name)) = let fun err msg = error (msg ^ " parent record " ^ quote name); val {args, parent, ...} = (case get_info thy name of SOME info => info | NONE => err "Unknown"); val _ = if length types <> length args then err "Bad number of arguments for" else (); fun bad_inst ((x, S), T) = if Sign.of_sort thy (T, S) then NONE else SOME x val bads = map_filter bad_inst (args ~~ types); val _ = null bads orelse err ("Ill-sorted instantiation of " ^ commas bads ^ " in"); val inst = args ~~ types; val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); val parent' = Option.map (apfst (map subst)) parent; in cons (name, inst) #> add_parents thy parent' end; in fun get_hierarchy thy (name, types) = add_parents thy (SOME (types, name)) []; fun get_parent_info thy parent = add_parents thy parent [] |> map (fn (name, inst) => let val subst = Term.map_type_tfree (the o AList.lookup (op =) inst); val {fields, extension, induct_scheme, ext_def, ...} = the_info thy name; val fields' = map (apsnd subst) fields; val extension' = apsnd (map subst) extension; in make_parent_info name fields' extension' ext_def induct_scheme end); end; (** concrete syntax for records **) (* parse translations *) local fun split_args (field :: fields) ((name, arg) :: fargs) = if can (unsuffix name) field then let val (args, rest) = split_args fields fargs in (arg :: args, rest) end else raise Fail ("expecting field " ^ quote field ^ " but got " ^ quote name) | split_args [] (fargs as (_ :: _)) = ([], fargs) | split_args (_ :: _) [] = raise Fail "expecting more fields" | split_args _ _ = ([], []); fun field_type_tr ((Const (\<^syntax_const>\_field_type\, _) $ Const (name, _) $ arg)) = (name, arg) | field_type_tr t = raise TERM ("field_type_tr", [t]); fun field_types_tr (Const (\<^syntax_const>\_field_types\, _) $ t $ u) = field_type_tr t :: field_types_tr u | field_types_tr t = [field_type_tr t]; fun record_field_types_tr more ctxt t = let val thy = Proof_Context.theory_of ctxt; fun err msg = raise TERM ("Error in record-type input: " ^ msg, [t]); fun mk_ext (fargs as (name, _) :: _) = (case get_fieldext thy (Proof_Context.intern_const ctxt name) of SOME (ext, alphas) => (case get_extfields thy ext of SOME fields => let val (fields', _) = split_last fields; val types = map snd fields'; val (args, rest) = split_args (map fst fields') fargs handle Fail msg => err msg; val argtypes = Syntax.check_typs ctxt (map Syntax_Phases.decode_typ args); val varifyT = varifyT (fold Term.maxidx_typ argtypes ~1 + 1); val vartypes = map varifyT types; val subst = Type.raw_matches (vartypes, argtypes) Vartab.empty handle Type.TYPE_MATCH => err "type is no proper record (extension)"; val alphas' = map (Syntax_Phases.term_of_typ ctxt o Envir.norm_type subst o varifyT) (#1 (split_last alphas)); val more' = mk_ext rest; in list_comb (Syntax.const (Lexicon.mark_type (suffix ext_typeN ext)), alphas' @ [more']) end | NONE => err ("no fields defined for " ^ quote ext)) | NONE => err (quote name ^ " is no proper field")) | mk_ext [] = more; in mk_ext (field_types_tr t) end; fun record_type_tr ctxt [t] = record_field_types_tr (Syntax.const \<^type_syntax>\unit\) ctxt t | record_type_tr _ ts = raise TERM ("record_type_tr", ts); fun record_type_scheme_tr ctxt [t, more] = record_field_types_tr more ctxt t | record_type_scheme_tr _ ts = raise TERM ("record_type_scheme_tr", ts); fun field_tr ((Const (\<^syntax_const>\_field\, _) $ Const (name, _) $ arg)) = (name, arg) | field_tr t = raise TERM ("field_tr", [t]); fun fields_tr (Const (\<^syntax_const>\_fields\, _) $ t $ u) = field_tr t :: fields_tr u | fields_tr t = [field_tr t]; fun record_fields_tr more ctxt t = let val thy = Proof_Context.theory_of ctxt; fun err msg = raise TERM ("Error in record input: " ^ msg, [t]); fun mk_ext (fargs as (name, _) :: _) = (case get_fieldext thy (Proof_Context.intern_const ctxt name) of SOME (ext, _) => (case get_extfields thy ext of SOME fields => let val (args, rest) = split_args (map fst (fst (split_last fields))) fargs handle Fail msg => err msg; val more' = mk_ext rest; in list_comb (Syntax.const (Lexicon.mark_const (ext ^ extN)), args @ [more']) end | NONE => err ("no fields defined for " ^ quote ext)) | NONE => err (quote name ^ " is no proper field")) | mk_ext [] = more; in mk_ext (fields_tr t) end; fun record_tr ctxt [t] = record_fields_tr (Syntax.const \<^const_syntax>\Unity\) ctxt t | record_tr _ ts = raise TERM ("record_tr", ts); fun record_scheme_tr ctxt [t, more] = record_fields_tr more ctxt t | record_scheme_tr _ ts = raise TERM ("record_scheme_tr", ts); fun field_update_tr (Const (\<^syntax_const>\_field_update\, _) $ Const (name, _) $ arg) = Syntax.const (suffix updateN name) $ Abs (Name.uu_, dummyT, arg) | field_update_tr t = raise TERM ("field_update_tr", [t]); fun field_updates_tr (Const (\<^syntax_const>\_field_updates\, _) $ t $ u) = field_update_tr t :: field_updates_tr u | field_updates_tr t = [field_update_tr t]; fun record_update_tr [t, u] = fold (curry op $) (field_updates_tr u) t | record_update_tr ts = raise TERM ("record_update_tr", ts); in val _ = Theory.setup (Sign.parse_translation [(\<^syntax_const>\_record_update\, K record_update_tr), (\<^syntax_const>\_record\, record_tr), (\<^syntax_const>\_record_scheme\, record_scheme_tr), (\<^syntax_const>\_record_type\, record_type_tr), (\<^syntax_const>\_record_type_scheme\, record_type_scheme_tr)]); end; (* print translations *) val type_abbr = Attrib.setup_config_bool \<^binding>\record_type_abbr\ (K true); val type_as_fields = Attrib.setup_config_bool \<^binding>\record_type_as_fields\ (K true); local (* FIXME early extern (!??) *) (* FIXME Syntax.free (??) *) fun field_type_tr' (c, t) = Syntax.const \<^syntax_const>\_field_type\ $ Syntax.const c $ t; fun field_types_tr' (t, u) = Syntax.const \<^syntax_const>\_field_types\ $ t $ u; fun record_type_tr' ctxt t = let val thy = Proof_Context.theory_of ctxt; val T = Syntax_Phases.decode_typ t; val varifyT = varifyT (Term.maxidx_of_typ T + 1); fun strip_fields T = (case T of Type (ext, args as _ :: _) => (case try (unsuffix ext_typeN) ext of SOME ext' => (case get_extfields thy ext' of SOME (fields as (x, _) :: _) => (case get_fieldext thy x of SOME (_, alphas) => (let val (f :: fs, _) = split_last fields; val fields' = apfst (Proof_Context.extern_const ctxt) f :: map (apfst Long_Name.base_name) fs; val (args', more) = split_last args; val alphavars = map varifyT (#1 (split_last alphas)); val subst = Type.raw_matches (alphavars, args') Vartab.empty; val fields'' = (map o apsnd) (Envir.norm_type subst o varifyT) fields'; in fields'' @ strip_fields more end handle Type.TYPE_MATCH => [("", T)]) | _ => [("", T)]) | _ => [("", T)]) | _ => [("", T)]) | _ => [("", T)]); val (fields, (_, moreT)) = split_last (strip_fields T); val _ = null fields andalso raise Match; val u = foldr1 field_types_tr' (map (field_type_tr' o apsnd (Syntax_Phases.term_of_typ ctxt)) fields); in if not (Config.get ctxt type_as_fields) orelse null fields then raise Match else if moreT = HOLogic.unitT then Syntax.const \<^syntax_const>\_record_type\ $ u else Syntax.const \<^syntax_const>\_record_type_scheme\ $ u $ Syntax_Phases.term_of_typ ctxt moreT end; (*try to reconstruct the record name type abbreviation from the (nested) extension types*) fun record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt tm = let val T = Syntax_Phases.decode_typ tm; val varifyT = varifyT (maxidx_of_typ T + 1); fun mk_type_abbr subst name args = let val abbrT = Type (name, map (varifyT o TFree) args) in Syntax_Phases.term_of_typ ctxt (Envir.norm_type subst abbrT) end; fun match rT T = Type.raw_match (varifyT rT, T) Vartab.empty; in if Config.get ctxt type_abbr then (case last_extT T of SOME (name, _) => if name = last_ext then let val subst = match schemeT T in if HOLogic.is_unitT (Envir.norm_type subst (varifyT (TFree zeta))) then mk_type_abbr subst abbr alphas else mk_type_abbr subst (suffix schemeN abbr) (alphas @ [zeta]) end handle Type.TYPE_MATCH => record_type_tr' ctxt tm else raise Match (*give print translation of specialised record a chance*) | _ => raise Match) else record_type_tr' ctxt tm end; in fun record_ext_type_tr' name = let val ext_type_name = Lexicon.mark_type (suffix ext_typeN name); fun tr' ctxt ts = record_type_tr' ctxt (list_comb (Syntax.const ext_type_name, ts)); in (ext_type_name, tr') end; fun record_ext_type_abbr_tr' abbr alphas zeta last_ext schemeT name = let val ext_type_name = Lexicon.mark_type (suffix ext_typeN name); fun tr' ctxt ts = record_type_abbr_tr' abbr alphas zeta last_ext schemeT ctxt (list_comb (Syntax.const ext_type_name, ts)); in (ext_type_name, tr') end; end; local (* FIXME Syntax.free (??) *) fun field_tr' (c, t) = Syntax.const \<^syntax_const>\_field\ $ Syntax.const c $ t; fun fields_tr' (t, u) = Syntax.const \<^syntax_const>\_fields\ $ t $ u; fun record_tr' ctxt t = let val thy = Proof_Context.theory_of ctxt; fun strip_fields t = (case strip_comb t of (Const (ext, _), args as (_ :: _)) => (case try (Lexicon.unmark_const o unsuffix extN) ext of SOME ext' => (case get_extfields thy ext' of SOME fields => (let val (f :: fs, _) = split_last (map fst fields); val fields' = Proof_Context.extern_const ctxt f :: map Long_Name.base_name fs; val (args', more) = split_last args; in (fields' ~~ args') @ strip_fields more end handle ListPair.UnequalLengths => [("", t)]) | NONE => [("", t)]) | NONE => [("", t)]) | _ => [("", t)]); val (fields, (_, more)) = split_last (strip_fields t); val _ = null fields andalso raise Match; val u = foldr1 fields_tr' (map field_tr' fields); in (case more of Const (\<^const_syntax>\Unity\, _) => Syntax.const \<^syntax_const>\_record\ $ u | _ => Syntax.const \<^syntax_const>\_record_scheme\ $ u $ more) end; in fun record_ext_tr' name = let val ext_name = Lexicon.mark_const (name ^ extN); fun tr' ctxt ts = record_tr' ctxt (list_comb (Syntax.const ext_name, ts)); in (ext_name, tr') end; end; local fun dest_update ctxt c = (case try Lexicon.unmark_const c of SOME d => try (unsuffix updateN) (Proof_Context.extern_const ctxt d) | NONE => NONE); fun field_updates_tr' ctxt (tm as Const (c, _) $ k $ u) = (case dest_update ctxt c of SOME name => (case try Syntax_Trans.const_abs_tr' k of SOME t => apfst (cons (Syntax.const \<^syntax_const>\_field_update\ $ Syntax.free name $ t)) (field_updates_tr' ctxt u) | NONE => ([], tm)) | NONE => ([], tm)) | field_updates_tr' _ tm = ([], tm); fun record_update_tr' ctxt tm = (case field_updates_tr' ctxt tm of ([], _) => raise Match | (ts, u) => Syntax.const \<^syntax_const>\_record_update\ $ u $ foldr1 (fn (v, w) => Syntax.const \<^syntax_const>\_field_updates\ $ v $ w) (rev ts)); in fun field_update_tr' name = let val update_name = Lexicon.mark_const (name ^ updateN); fun tr' ctxt [t, u] = record_update_tr' ctxt (Syntax.const update_name $ t $ u) | tr' _ _ = raise Match; in (update_name, tr') end; end; (** record simprocs **) fun is_sel_upd_pair thy (Const (s, _)) (Const (u, t')) = (case get_updates thy u of SOME u_name => u_name = s | NONE => raise TERM ("is_sel_upd_pair: not update", [Const (u, t')])); fun mk_comp_id f = let val T = range_type (fastype_of f) in HOLogic.mk_comp (Const (\<^const_name>\Fun.id\, T --> T), f) end; fun get_upd_funs (upd $ _ $ t) = upd :: get_upd_funs t | get_upd_funs _ = []; fun get_accupd_simps ctxt term defset = let val thy = Proof_Context.theory_of ctxt; val (acc, [body]) = strip_comb term; val upd_funs = sort_distinct Term_Ord.fast_term_ord (get_upd_funs body); fun get_simp upd = let (* FIXME fresh "f" (!?) *) val T = domain_type (fastype_of upd); val lhs = HOLogic.mk_comp (acc, upd $ Free ("f", T)); val rhs = if is_sel_upd_pair thy acc upd then HOLogic.mk_comp (Free ("f", T), acc) else mk_comp_id acc; val prop = lhs === rhs; val othm = Goal.prove ctxt [] [] prop (fn {context = ctxt', ...} => simp_tac (put_simpset defset ctxt') 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN TRY (simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms id_apply id_o o_id}) 1)); val dest = if is_sel_upd_pair thy acc upd then @{thm o_eq_dest} else @{thm o_eq_id_dest}; in Drule.export_without_context (othm RS dest) end; in map get_simp upd_funs end; fun get_updupd_simp ctxt defset u u' comp = let (* FIXME fresh "f" (!?) *) val f = Free ("f", domain_type (fastype_of u)); val f' = Free ("f'", domain_type (fastype_of u')); val lhs = HOLogic.mk_comp (u $ f, u' $ f'); val rhs = if comp then u $ HOLogic.mk_comp (f, f') else HOLogic.mk_comp (u' $ f', u $ f); val prop = lhs === rhs; val othm = Goal.prove ctxt [] [] prop (fn {context = ctxt', ...} => simp_tac (put_simpset defset ctxt') 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN TRY (simp_tac (put_simpset HOL_ss ctxt' addsimps @{thms id_apply}) 1)); val dest = if comp then @{thm o_eq_dest_lhs} else @{thm o_eq_dest}; in Drule.export_without_context (othm RS dest) end; fun get_updupd_simps ctxt term defset = let val upd_funs = get_upd_funs term; val cname = fst o dest_Const; fun getswap u u' = get_updupd_simp ctxt defset u u' (cname u = cname u'); fun build_swaps_to_eq _ [] swaps = swaps | build_swaps_to_eq upd (u :: us) swaps = let val key = (cname u, cname upd); val newswaps = if Symreltab.defined swaps key then swaps else Symreltab.insert (K true) (key, getswap u upd) swaps; in if cname u = cname upd then newswaps else build_swaps_to_eq upd us newswaps end; fun swaps_needed [] _ _ swaps = map snd (Symreltab.dest swaps) | swaps_needed (u :: us) prev seen swaps = if Symtab.defined seen (cname u) then swaps_needed us prev seen (build_swaps_to_eq u prev swaps) else swaps_needed us (u :: prev) (Symtab.insert (K true) (cname u, ()) seen) swaps; in swaps_needed upd_funs [] Symtab.empty Symreltab.empty end; fun prove_unfold_defs thy ex_simps ex_simprs prop = let val ctxt = Proof_Context.init_global thy; val defset = get_sel_upd_defs thy; val prop' = Envir.beta_eta_contract prop; val (lhs, _) = Logic.dest_equals (Logic.strip_assums_concl prop'); val (_, args) = strip_comb lhs; val simps = (if length args = 1 then get_accupd_simps else get_updupd_simps) ctxt lhs defset; in Goal.prove ctxt [] [] prop' (fn {context = ctxt', ...} => simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (simps @ @{thms K_record_comp})) 1 THEN TRY (simp_tac (put_simpset HOL_basic_ss ctxt' addsimps ex_simps addsimprocs ex_simprs) 1)) end; local fun eq (s1: string) (s2: string) = (s1 = s2); fun has_field extfields f T = exists (fn (eN, _) => exists (eq f o fst) (Symtab.lookup_list extfields eN)) (dest_recTs T); fun K_skeleton n (T as Type (_, [_, kT])) (b as Bound i) (Abs (x, xT, t)) = if null (loose_bnos t) then ((n, kT), (Abs (x, xT, Bound (i + 1)))) else ((n, T), b) | K_skeleton n T b _ = ((n, T), b); in (* simproc *) (* Simplify selections of an record update: (1) S (S_update k r) = k (S r) (2) S (X_update k r) = S r The simproc skips multiple updates at once, eg: S (X_update x (Y_update y (S_update k r))) = k (S r) But be careful in (2) because of the extensibility of records. - If S is a more-selector we have to make sure that the update on component X does not affect the selected subrecord. - If X is a more-selector we have to make sure that S is not in the updated subrecord. *) val simproc = Simplifier.make_simproc \<^context> "record" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; val t = Thm.term_of ct; in (case t of (sel as Const (s, Type (_, [_, rangeS]))) $ ((upd as Const (u, Type (_, [_, Type (_, [rT, _])]))) $ k $ r) => if is_selector thy s andalso is_some (get_updates thy u) then let val {sel_upd = {updates, ...}, extfields, ...} = Data.get thy; fun mk_eq_terms ((upd as Const (u, Type(_, [kT, _]))) $ k $ r) = (case Symtab.lookup updates u of NONE => NONE | SOME u_name => if u_name = s then (case mk_eq_terms r of NONE => let val rv = ("r", rT); val rb = Bound 0; val (kv, kb) = K_skeleton "k" kT (Bound 1) k; in SOME (upd $ kb $ rb, kb $ (sel $ rb), [kv, rv]) end | SOME (trm, trm', vars) => let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k; in SOME (upd $ kb $ trm, kb $ trm', kv :: vars) end) else if has_field extfields u_name rangeS orelse has_field extfields s (domain_type kT) then NONE else (case mk_eq_terms r of SOME (trm, trm', vars) => let val (kv, kb) = K_skeleton "k" kT (Bound (length vars)) k in SOME (upd $ kb $ trm, trm', kv :: vars) end | NONE => let val rv = ("r", rT); val rb = Bound 0; val (kv, kb) = K_skeleton "k" kT (Bound 1) k; in SOME (upd $ kb $ rb, sel $ rb, [kv, rv]) end)) | mk_eq_terms _ = NONE; in (case mk_eq_terms (upd $ k $ r) of SOME (trm, trm', vars) => SOME (prove_unfold_defs thy [] [] (Logic.list_all (vars, Logic.mk_equals (sel $ trm, trm')))) | NONE => NONE) end else NONE | _ => NONE) end}; fun get_upd_acc_cong_thm upd acc thy ss = let val ctxt = Proof_Context.init_global thy; val prop = infer_instantiate ctxt [(("upd", 0), Thm.cterm_of ctxt upd), (("ac", 0), Thm.cterm_of ctxt acc)] updacc_cong_triv |> Thm.concl_of; in Goal.prove ctxt [] [] prop (fn {context = ctxt', ...} => simp_tac (put_simpset ss ctxt') 1 THEN REPEAT_DETERM (Iso_Tuple_Support.iso_tuple_intros_tac ctxt' 1) THEN TRY (resolve_tac ctxt' [updacc_cong_idI] 1)) end; (* upd_simproc *) (*Simplify multiple updates: (1) "N_update y (M_update g (N_update x (M_update f r))) = (N_update (y o x) (M_update (g o f) r))" (2) "r(|M:= M r|) = r" In both cases "more" updates complicate matters: for this reason we omit considering further updates if doing so would introduce both a more update and an update to a field within it.*) val upd_simproc = Simplifier.make_simproc \<^context> "record_upd" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; val t = Thm.term_of ct; (*We can use more-updators with other updators as long as none of the other updators go deeper than any more updator. min here is the depth of the deepest other updator, max the depth of the shallowest more updator.*) fun include_depth (dep, true) (min, max) = if min <= dep then SOME (min, if dep <= max orelse max = ~1 then dep else max) else NONE | include_depth (dep, false) (min, max) = if dep <= max orelse max = ~1 then SOME (if min <= dep then dep else min, max) else NONE; fun getupdseq (term as (upd as Const (u, _)) $ f $ tm) min max = (case get_update_details u thy of SOME (s, dep, ismore) => (case include_depth (dep, ismore) (min, max) of SOME (min', max') => let val (us, bs, _) = getupdseq tm min' max' in ((upd, s, f) :: us, bs, fastype_of term) end | NONE => ([], term, HOLogic.unitT)) | NONE => ([], term, HOLogic.unitT)) | getupdseq term _ _ = ([], term, HOLogic.unitT); val (upds, base, baseT) = getupdseq t 0 ~1; fun is_upd_noop s (Abs (n, T, Const (s', T') $ tm')) tm = if s = s' andalso null (loose_bnos tm') andalso subst_bound (HOLogic.unit, tm') = tm then (true, Abs (n, T, Const (s', T') $ Bound 1)) else (false, HOLogic.unit) | is_upd_noop _ _ _ = (false, HOLogic.unit); fun get_noop_simps (upd as Const _) (Abs (_, _, (acc as Const _) $ _)) = let val ss = get_sel_upd_defs thy; val uathm = get_upd_acc_cong_thm upd acc thy ss; in [Drule.export_without_context (uathm RS updacc_noopE), Drule.export_without_context (uathm RS updacc_noop_compE)] end; (*If f is constant then (f o g) = f. We know that K_skeleton only returns constant abstractions thus when we see an abstraction we can discard inner updates.*) fun add_upd (f as Abs _) _ = [f] | add_upd f fs = (f :: fs); (*mk_updterm returns (orig-term-skeleton, simplified-skeleton, variables, duplicate-updates, simp-flag, noop-simps) where duplicate-updates is a table used to pass upward the list of update functions which can be composed into an update above them, simp-flag indicates whether any simplification was achieved, and noop-simps are used for eliminating case (2) defined above*) fun mk_updterm ((upd as Const (u, T), s, f) :: upds) above term = let val (lhs, rhs, vars, dups, simp, noops) = mk_updterm upds (Symtab.update (u, ()) above) term; val (fvar, skelf) = K_skeleton (Long_Name.base_name s) (domain_type T) (Bound (length vars)) f; val (isnoop, skelf') = is_upd_noop s f term; val funT = domain_type T; fun mk_comp_local (f, f') = Const (\<^const_name>\Fun.comp\, funT --> funT --> funT) $ f $ f'; in if isnoop then (upd $ skelf' $ lhs, rhs, vars, Symtab.update (u, []) dups, true, if Symtab.defined noops u then noops else Symtab.update (u, get_noop_simps upd skelf') noops) else if Symtab.defined above u then (upd $ skelf $ lhs, rhs, fvar :: vars, Symtab.map_default (u, []) (add_upd skelf) dups, true, noops) else (case Symtab.lookup dups u of SOME fs => (upd $ skelf $ lhs, upd $ foldr1 mk_comp_local (add_upd skelf fs) $ rhs, fvar :: vars, dups, true, noops) | NONE => (upd $ skelf $ lhs, upd $ skelf $ rhs, fvar :: vars, dups, simp, noops)) end | mk_updterm [] _ _ = (Bound 0, Bound 0, [("r", baseT)], Symtab.empty, false, Symtab.empty) | mk_updterm us _ _ = raise TERM ("mk_updterm match", map (fn (x, _, _) => x) us); val (lhs, rhs, vars, _, simp, noops) = mk_updterm upds Symtab.empty base; val noops' = maps snd (Symtab.dest noops); in if simp then SOME (prove_unfold_defs thy noops' [simproc] (Logic.list_all (vars, Logic.mk_equals (lhs, rhs)))) else NONE end}; end; (* eq_simproc *) (*Look up the most specific record-equality. Note on efficiency: Testing equality of records boils down to the test of equality of all components. Therefore the complexity is: #components * complexity for single component. Especially if a record has a lot of components it may be better to split up the record first and do simplification on that (split_simp_tac). e.g. r(|lots of updates|) = x eq_simproc split_simp_tac Complexity: #components * #updates #updates *) val eq_simproc = Simplifier.make_simproc \<^context> "record_eq" {lhss = [\<^term>\r = s\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (\<^const_name>\HOL.eq\, Type (_, [T, _])) $ _ $ _ => (case rec_id ~1 T of "" => NONE | name => (case get_equalities (Proof_Context.theory_of ctxt) name of NONE => NONE | SOME thm => SOME (thm RS @{thm Eq_TrueI}))) | _ => NONE)}; (* split_simproc *) (*Split quantified occurrences of records, for which P holds. P can peek on the subterm starting at the quantified occurrence of the record (including the quantifier): P t = 0: do not split P t = ~1: completely split P t > 0: split up to given bound of record extensions.*) fun split_simproc P = Simplifier.make_simproc \<^context> "record_split" {lhss = [\<^term>\x::'a::{}\], proc = fn _ => fn ctxt => fn ct => (case Thm.term_of ct of Const (quantifier, Type (_, [Type (_, [T, _]), _])) $ _ => if quantifier = \<^const_name>\Pure.all\ orelse quantifier = \<^const_name>\All\ orelse quantifier = \<^const_name>\Ex\ then (case rec_id ~1 T of "" => NONE | _ => let val split = P (Thm.term_of ct) in if split <> 0 then (case get_splits (Proof_Context.theory_of ctxt) (rec_id split T) of NONE => NONE | SOME (all_thm, All_thm, Ex_thm, _) => SOME (case quantifier of \<^const_name>\Pure.all\ => all_thm | \<^const_name>\All\ => All_thm RS @{thm eq_reflection} | \<^const_name>\Ex\ => Ex_thm RS @{thm eq_reflection} | _ => raise Fail "split_simproc")) else NONE end) else NONE | _ => NONE)}; val ex_sel_eq_simproc = Simplifier.make_simproc \<^context> "ex_sel_eq" {lhss = [\<^term>\Ex t\], proc = fn _ => fn ctxt => fn ct => let val thy = Proof_Context.theory_of ctxt; val t = Thm.term_of ct; fun mkeq (lr, Teq, (sel, Tsel), x) i = if is_selector thy sel then let val x' = if not (Term.is_dependent x) then Free ("x" ^ string_of_int i, range_type Tsel) else raise TERM ("", [x]); val sel' = Const (sel, Tsel) $ Bound 0; val (l, r) = if lr then (sel', x') else (x', sel'); in Const (\<^const_name>\HOL.eq\, Teq) $ l $ r end else raise TERM ("", [Const (sel, Tsel)]); fun dest_sel_eq (Const (\<^const_name>\HOL.eq\, Teq) $ (Const (sel, Tsel) $ Bound 0) $ X) = (true, Teq, (sel, Tsel), X) | dest_sel_eq (Const (\<^const_name>\HOL.eq\, Teq) $ X $ (Const (sel, Tsel) $ Bound 0)) = (false, Teq, (sel, Tsel), X) | dest_sel_eq _ = raise TERM ("", []); in (case t of Const (\<^const_name>\Ex\, Tex) $ Abs (s, T, t) => (let val eq = mkeq (dest_sel_eq t) 0; val prop = Logic.list_all ([("r", T)], Logic.mk_equals (Const (\<^const_name>\Ex\, Tex) $ Abs (s, T, eq), \<^term>\True\)); in SOME (Goal.prove_sorry_global thy [] [] prop (fn _ => simp_tac (put_simpset (get_simpset thy) ctxt addsimps @{thms simp_thms} addsimprocs [split_simproc (K ~1)]) 1)) end handle TERM _ => NONE) | _ => NONE) end}; (* split_simp_tac *) (*Split (and simplify) all records in the goal for which P holds. For quantified occurrences of a record P can peek on the whole subterm (including the quantifier); for free variables P can only peek on the variable itself. P t = 0: do not split P t = ~1: completely split P t > 0: split up to given bound of record extensions.*) fun split_simp_tac ctxt thms P = CSUBGOAL (fn (cgoal, i) => let val thy = Proof_Context.theory_of ctxt; val goal = Thm.term_of cgoal; val frees = filter (is_recT o #2) (Term.add_frees goal []); val has_rec = exists_Const (fn (s, Type (_, [Type (_, [T, _]), _])) => (s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\ orelse s = \<^const_name>\Ex\) andalso is_recT T | _ => false); fun mk_split_free_tac free induct_thm i = let val _ $ (_ $ Var (r, _)) = Thm.concl_of induct_thm; val thm = infer_instantiate ctxt [(r, Thm.cterm_of ctxt free)] induct_thm; in simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_atomize}) i THEN resolve_tac ctxt [thm] i THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms induct_rulify}) i end; val split_frees_tacs = frees |> map_filter (fn (x, T) => (case rec_id ~1 T of "" => NONE | _ => let val free = Free (x, T); val split = P free; in if split <> 0 then (case get_splits thy (rec_id split T) of NONE => NONE | SOME (_, _, _, induct_thm) => SOME (mk_split_free_tac free induct_thm i)) else NONE end)); val simprocs = if has_rec goal then [split_simproc P] else []; val thms' = @{thms o_apply K_record_comp} @ thms; in EVERY split_frees_tacs THEN full_simp_tac (put_simpset (get_simpset thy) ctxt addsimps thms' addsimprocs simprocs) i end); (* split_tac *) (*Split all records in the goal, which are quantified by !! or ALL.*) fun split_tac ctxt = CSUBGOAL (fn (cgoal, i) => let val goal = Thm.term_of cgoal; val has_rec = exists_Const (fn (s, Type (_, [Type (_, [T, _]), _])) => (s = \<^const_name>\Pure.all\ orelse s = \<^const_name>\All\) andalso is_recT T | _ => false); fun is_all (Const (\<^const_name>\Pure.all\, _) $ _) = ~1 | is_all (Const (\<^const_name>\All\, _) $ _) = ~1 | is_all _ = 0; in if has_rec goal then full_simp_tac (put_simpset HOL_basic_ss ctxt addsimprocs [split_simproc is_all]) i else no_tac end); val _ = Theory.setup (map_theory_simpset (fn ctxt => ctxt addsimprocs [simproc, upd_simproc, eq_simproc])); (* wrapper *) val split_name = "record_split_tac"; val split_wrapper = (split_name, fn ctxt => fn tac => split_tac ctxt ORELSE' tac); (** theory extender interface **) (* attributes *) fun case_names_fields x = Rule_Cases.case_names ["fields"] x; fun induct_type_global name = [case_names_fields, Induct.induct_type name]; fun cases_type_global name = [case_names_fields, Induct.cases_type name]; (* tactics *) (*Do case analysis / induction according to rule on last parameter of ith subgoal (or on s if there are no parameters). Instatiation of record variable (and predicate) in rule is calculated to avoid problems with higher order unification.*) fun try_param_tac ctxt s rule = CSUBGOAL (fn (cgoal, i) => let val g = Thm.term_of cgoal; val params = Logic.strip_params g; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl g); val rule' = Thm.lift_rule cgoal rule; val (P, ys) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_assums_concl (Thm.prop_of rule'))); (*ca indicates if rule is a case analysis or induction rule*) val (x, ca) = (case rev (drop (length params) ys) of [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (hd (rev (Logic.strip_assums_hyp (hd (Thm.prems_of rule')))))))), true) | [x] => (head_of x, false)); val rule'' = infer_instantiate ctxt (map (apsnd (Thm.cterm_of ctxt)) (case rev params of [] => (case AList.lookup (op =) (Term.add_frees g []) s of NONE => error "try_param_tac: no such variable" | SOME T => [(#1 (dest_Var P), if ca then concl else lambda (Free (s, T)) concl), (#1 (dest_Var x), Free (s, T))]) | (_, T) :: _ => [(#1 (dest_Var P), fold_rev Term.abs params (if ca then concl else incr_boundvars 1 (Abs (s, T, concl)))), (#1 (dest_Var x), fold_rev Term.abs params (Bound 0))])) rule'; in compose_tac ctxt (false, rule'', Thm.nprems_of rule) i end); fun extension_definition overloaded name fields alphas zeta moreT more vars thy = let val ctxt = Proof_Context.init_global thy; val base_name = Long_Name.base_name name; val fieldTs = map snd fields; val fields_moreTs = fieldTs @ [moreT]; val alphas_zeta = alphas @ [zeta]; val ext_binding = Binding.name (suffix extN base_name); val ext_name = suffix extN name; val ext_tyco = suffix ext_typeN name; val extT = Type (ext_tyco, map TFree alphas_zeta); val ext_type = fields_moreTs ---> extT; (* the tree of new types that will back the record extension *) val mktreeV = Balanced_Tree.make Iso_Tuple_Support.mk_cons_tuple; fun mk_iso_tuple (left, right) (thy, i) = let val suff = if i = 0 then ext_typeN else inner_typeN ^ string_of_int i; val ((_, cons), thy') = thy |> Iso_Tuple_Support.add_iso_tuple_type overloaded (Binding.suffix_name suff (Binding.name base_name), alphas_zeta) (fastype_of left, fastype_of right); in (cons $ left $ right, (thy', i + 1)) end; (*trying to create a 1-element iso_tuple will fail, and is pointless anyway*) fun mk_even_iso_tuple [arg] = pair arg | mk_even_iso_tuple args = mk_iso_tuple (Iso_Tuple_Support.dest_cons_tuple (mktreeV args)); fun build_meta_tree_type i thy vars more = let val len = length vars in if len < 1 then raise TYPE ("meta_tree_type args too short", [], vars) else if len > 16 then let fun group16 [] = [] | group16 xs = take 16 xs :: group16 (drop 16 xs); val vars' = group16 vars; val (composites, (thy', i')) = fold_map mk_even_iso_tuple vars' (thy, i); in build_meta_tree_type i' thy' composites more end else let val (term, (thy', _)) = mk_iso_tuple (mktreeV vars, more) (thy, 0) in (term, thy') end end; val _ = timing_msg ctxt "record extension preparing definitions"; (* 1st stage part 1: introduce the tree of new types *) val (ext_body, typ_thy) = timeit_msg ctxt "record extension nested type def:" (fn () => build_meta_tree_type 1 thy vars more); (* prepare declarations and definitions *) (* 1st stage part 2: define the ext constant *) fun mk_ext args = list_comb (Const (ext_name, ext_type), args); val ext_spec = Logic.mk_equals (mk_ext (vars @ [more]), ext_body); val ([ext_def], defs_thy) = timeit_msg ctxt "record extension constructor def:" (fn () => typ_thy |> Sign.declare_const_global ((ext_binding, ext_type), NoSyn) |> snd |> Global_Theory.add_defs false [((Thm.def_binding ext_binding, ext_spec), [])]); val defs_ctxt = Proof_Context.init_global defs_thy; (* prepare propositions *) val _ = timing_msg ctxt "record extension preparing propositions"; val vars_more = vars @ [more]; val variants = map (fn Free (x, _) => x) vars_more; val ext = mk_ext vars_more; val s = Free (rN, extT); val P = Free (singleton (Name.variant_list variants) "P", extT --> HOLogic.boolT); val inject_prop = (* FIXME local x x' *) let val vars_more' = map (fn (Free (x, T)) => Free (x ^ "'", T)) vars_more in HOLogic.mk_conj (HOLogic.eq_const extT $ mk_ext vars_more $ mk_ext vars_more', \<^term>\True\) === foldr1 HOLogic.mk_conj (map HOLogic.mk_eq (vars_more ~~ vars_more') @ [\<^term>\True\]) end; val induct_prop = (fold_rev Logic.all vars_more (Trueprop (P $ ext)), Trueprop (P $ s)); val split_meta_prop = (* FIXME local P *) let val P = Free (singleton (Name.variant_list variants) "P", extT --> propT) in Logic.mk_equals (Logic.all s (P $ s), fold_rev Logic.all vars_more (P $ ext)) end; val inject = timeit_msg ctxt "record extension inject proof:" (fn () => simplify (put_simpset HOL_ss defs_ctxt) (Goal.prove_sorry_global defs_thy [] [] inject_prop (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_basic_ss ctxt addsimps [ext_def]) 1 THEN REPEAT_DETERM (resolve_tac ctxt @{thms refl_conj_eq} 1 ORELSE Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE resolve_tac ctxt [refl] 1)))); (*We need a surjection property r = (| f = f r, g = g r ... |) to prove other theorems. We haven't given names to the accessors f, g etc yet however, so we generate an ext structure with free variables as all arguments and allow the introduction tactic to operate on it as far as it can. We then use Drule.export_without_context to convert the free variables into unifiable variables and unify them with (roughly) the definition of the accessor.*) val surject = timeit_msg ctxt "record extension surjective proof:" (fn () => let val start = infer_instantiate defs_ctxt [(("y", 0), Thm.cterm_of defs_ctxt ext)] surject_assist_idE; val tactic1 = simp_tac (put_simpset HOL_basic_ss defs_ctxt addsimps [ext_def]) 1 THEN REPEAT_ALL_NEW (Iso_Tuple_Support.iso_tuple_intros_tac defs_ctxt) 1; val tactic2 = REPEAT (resolve_tac defs_ctxt [surject_assistI] 1 THEN resolve_tac defs_ctxt [refl] 1); val [halfway] = Seq.list_of (tactic1 start); val [surject] = Seq.list_of (tactic2 (Drule.export_without_context halfway)); in surject end); val split_meta = timeit_msg ctxt "record extension split_meta proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_meta_prop (fn {context = ctxt, ...} => EVERY1 [resolve_tac ctxt @{thms equal_intr_rule}, Goal.norm_hhf_tac ctxt, eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt, resolve_tac ctxt [@{thm prop_subst} OF [surject]], REPEAT o eresolve_tac ctxt @{thms meta_allE}, assume_tac ctxt])); val induct = timeit_msg ctxt "record extension induct proof:" (fn () => let val (assm, concl) = induct_prop in Goal.prove_sorry_global defs_thy [] [assm] concl (fn {context = ctxt, prems, ...} => cut_tac (split_meta RS Drule.equal_elim_rule2) 1 THEN resolve_tac ctxt prems 2 THEN asm_simp_tac (put_simpset HOL_ss ctxt) 1) end); val ([(_, [induct']), (_, [inject']), (_, [surjective']), (_, [split_meta'])], thm_thy) = defs_thy |> Global_Theory.note_thmss "" [((Binding.name "ext_induct", []), [([induct], [])]), ((Binding.name "ext_inject", []), [([inject], [])]), ((Binding.name "ext_surjective", []), [([surject], [])]), ((Binding.name "ext_split", []), [([split_meta], [])])]; in (((ext_name, ext_type), (ext_tyco, alphas_zeta), extT, induct', inject', surjective', split_meta', ext_def), thm_thy) end; fun chunks [] [] = [] | chunks [] xs = [xs] | chunks (l :: ls) xs = take l xs :: chunks ls (drop l xs); fun chop_last [] = error "chop_last: list should not be empty" | chop_last [x] = ([], x) | chop_last (x :: xs) = let val (tl, l) = chop_last xs in (x :: tl, l) end; fun subst_last _ [] = error "subst_last: list should not be empty" | subst_last s [_] = [s] | subst_last s (x :: xs) = x :: subst_last s xs; (* mk_recordT *) (*build up the record type from the current extension tpye extT and a list of parent extensions, starting with the root of the record hierarchy*) fun mk_recordT extT = fold_rev (fn (parent, Ts) => fn T => Type (parent, subst_last T Ts)) extT; (* code generation *) fun mk_random_eq tyco vs extN Ts = let (* FIXME local i etc. *) val size = \<^term>\i::natural\; fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); val T = Type (tyco, map TFree vs); val Tm = termifyT T; val params = Name.invent_names Name.context "x" Ts; val lhs = HOLogic.mk_random T size; val tc = HOLogic.mk_return Tm \<^typ>\Random.seed\ (HOLogic.mk_valtermify_app extN params T); val rhs = HOLogic.mk_ST (map (fn (v, T') => ((HOLogic.mk_random T' size, \<^typ>\Random.seed\), SOME (v, termifyT T'))) params) tc \<^typ>\Random.seed\ (SOME Tm, \<^typ>\Random.seed\); in (lhs, rhs) end fun mk_full_exhaustive_eq tyco vs extN Ts = let (* FIXME local i etc. *) val size = \<^term>\i::natural\; fun termifyT T = HOLogic.mk_prodT (T, \<^typ>\unit \ term\); val T = Type (tyco, map TFree vs); val test_function = Free ("f", termifyT T --> \<^typ>\(bool \ term list) option\); val params = Name.invent_names Name.context "x" Ts; fun full_exhaustiveT T = (termifyT T --> \<^typ>\(bool \ Code_Evaluation.term list) option\) --> \<^typ>\natural\ --> \<^typ>\(bool \ Code_Evaluation.term list) option\; fun mk_full_exhaustive T = Const (\<^const_name>\Quickcheck_Exhaustive.full_exhaustive_class.full_exhaustive\, full_exhaustiveT T); val lhs = mk_full_exhaustive T $ test_function $ size; val tc = test_function $ (HOLogic.mk_valtermify_app extN params T); val rhs = fold_rev (fn (v, T) => fn cont => mk_full_exhaustive T $ (lambda (Free (v, termifyT T)) cont) $ size) params tc; in (lhs, rhs) end; fun instantiate_sort_record (sort, mk_eq) tyco vs extN Ts thy = let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (mk_eq tyco vs extN Ts)); in thy |> Class.instantiation ([tyco], vs, sort) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition NONE [] [] ((Binding.concealed Binding.empty, []), eq)) |> snd |> Class.prove_instantiation_exit (fn ctxt => Class.intro_classes_tac ctxt []) end; fun ensure_sort_record (sort, mk_eq) ext_tyco vs extN Ts thy = let val algebra = Sign.classes_of thy; val has_inst = Sorts.has_instance algebra ext_tyco sort; in if has_inst then thy else (case Quickcheck_Common.perhaps_constrain thy (map (rpair sort) Ts) vs of SOME constrain => instantiate_sort_record (sort, mk_eq) ext_tyco (map constrain vs) extN ((map o map_atyps) (fn TFree v => TFree (constrain v)) Ts) thy | NONE => thy) end; fun add_code ext_tyco vs extT ext simps inject thy = if Config.get_global thy codegen then let val eq = HOLogic.mk_Trueprop (HOLogic.mk_eq (Const (\<^const_name>\HOL.equal\, extT --> extT --> HOLogic.boolT), Const (\<^const_name>\HOL.eq\, extT --> extT --> HOLogic.boolT))); fun tac ctxt eq_def = Class.intro_classes_tac ctxt [] THEN rewrite_goals_tac ctxt [Simpdata.mk_eq eq_def] THEN ALLGOALS (resolve_tac ctxt @{thms refl}); fun mk_eq ctxt eq_def = rewrite_rule ctxt [Axclass.unoverload ctxt (Thm.symmetric (Simpdata.mk_eq eq_def))] inject; fun mk_eq_refl ctxt = @{thm equal_refl} |> Thm.instantiate - ([((("'a", 0), \<^sort>\equal\), Thm.ctyp_of ctxt (Logic.varifyT_global extT))], []) + (TVars.make [((("'a", 0), \<^sort>\equal\), Thm.ctyp_of ctxt (Logic.varifyT_global extT))], + Vars.empty) |> Axclass.unoverload ctxt; val ensure_random_record = ensure_sort_record (\<^sort>\random\, mk_random_eq); val ensure_exhaustive_record = ensure_sort_record (\<^sort>\full_exhaustive\, mk_full_exhaustive_eq); fun add_code eq_def thy = let val ctxt = Proof_Context.init_global thy; in thy |> Code.declare_default_eqns_global [(mk_eq (Proof_Context.init_global thy) eq_def, true), (mk_eq_refl (Proof_Context.init_global thy), false)] end; in thy |> Code.declare_datatype_global [ext] |> Code.declare_default_eqns_global (map (rpair true) simps) |> Class.instantiation ([ext_tyco], vs, [HOLogic.class_equal]) |> `(fn lthy => Syntax.check_term lthy eq) |-> (fn eq => Specification.definition NONE [] [] (Binding.empty_atts, eq)) |-> (fn (_, (_, eq_def)) => Class.prove_instantiation_exit_result Morphism.thm tac eq_def) |-> add_code |> ensure_random_record ext_tyco vs (fst ext) (binder_types (snd ext)) |> ensure_exhaustive_record ext_tyco vs (fst ext) (binder_types (snd ext)) end else thy; fun add_ctr_sugar ctr exhaust inject sel_thms = Ctr_Sugar.default_register_ctr_sugar_global (K true) {kind = Ctr_Sugar.Record, T = body_type (fastype_of ctr), ctrs = [ctr], casex = Term.dummy, discs = [], selss = [], exhaust = exhaust, nchotomy = Drule.dummy_thm, injects = [inject], distincts = [], case_thms = [], case_cong = Drule.dummy_thm, case_cong_weak = Drule.dummy_thm, case_distribs = [], split = Drule.dummy_thm, split_asm = Drule.dummy_thm, disc_defs = [], disc_thmss = [], discIs = [], disc_eq_cases = [], sel_defs = [], sel_thmss = [sel_thms], distinct_discsss = [], exhaust_discs = [], exhaust_sels = [], collapses = [], expands = [], split_sels = [], split_sel_asms = [], case_eq_ifs = []}; fun lhs_of_equation (Const (\<^const_name>\Pure.eq\, _) $ t $ _) = t | lhs_of_equation (\<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ t $ _)) = t; fun add_spec_rule rule = let val head = head_of (lhs_of_equation (Thm.prop_of rule)) in Spec_Rules.add_global Binding.empty Spec_Rules.equational [head] [rule] end; (* definition *) fun definition overloaded (alphas, binding) parent (parents: parent_info list) raw_fields thy0 = let val ctxt0 = Proof_Context.init_global thy0; val prefix = Binding.name_of binding; val name = Sign.full_name thy0 binding; val full = Sign.full_name_path thy0 prefix; val bfields = map (fn (x, T, _) => (x, T)) raw_fields; val field_syntax = map #3 raw_fields; val parent_fields = maps #fields parents; val parent_chunks = map (length o #fields) parents; val parent_names = map fst parent_fields; val parent_types = map snd parent_fields; val parent_fields_len = length parent_fields; val parent_variants = Name.variant_list [moreN, rN, rN ^ "'", wN] (map Long_Name.base_name parent_names); val parent_vars = map2 (curry Free) parent_variants parent_types; val parent_len = length parents; val fields = map (apfst full) bfields; val names = map fst fields; val types = map snd fields; val alphas_fields = fold Term.add_tfreesT types []; val alphas_ext = inter (op =) alphas_fields alphas; val len = length fields; val variants = Name.variant_list (moreN :: rN :: (rN ^ "'") :: wN :: parent_variants) (map (Binding.name_of o fst) bfields); val vars = map2 (curry Free) variants types; val named_vars = names ~~ vars; val idxms = 0 upto len; val all_fields = parent_fields @ fields; val all_types = parent_types @ types; val all_variants = parent_variants @ variants; val all_vars = parent_vars @ vars; val all_named_vars = (parent_names ~~ parent_vars) @ named_vars; val zeta = (singleton (Name.variant_list (map #1 alphas)) "'z", \<^sort>\type\); val moreT = TFree zeta; val more = Free (moreN, moreT); val full_moreN = full (Binding.name moreN); val bfields_more = bfields @ [(Binding.name moreN, moreT)]; val fields_more = fields @ [(full_moreN, moreT)]; val named_vars_more = named_vars @ [(full_moreN, more)]; val all_vars_more = all_vars @ [more]; val all_named_vars_more = all_named_vars @ [(full_moreN, more)]; (* 1st stage: ext_thy *) val extension_name = full binding; val ((ext, (ext_tyco, vs), extT, ext_induct, ext_inject, ext_surjective, ext_split, ext_def), ext_thy) = thy0 |> Sign.qualified_path false binding |> extension_definition overloaded extension_name fields alphas_ext zeta moreT more vars; val ext_ctxt = Proof_Context.init_global ext_thy; val _ = timing_msg ext_ctxt "record preparing definitions"; val Type extension_scheme = extT; val extension_name = unsuffix ext_typeN (fst extension_scheme); val extension = let val (n, Ts) = extension_scheme in (n, subst_last HOLogic.unitT Ts) end; val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extension_name]; val extension_id = implode extension_names; fun rec_schemeT n = mk_recordT (map #extension (drop n parents)) extT; val rec_schemeT0 = rec_schemeT 0; fun recT n = let val (c, Ts) = extension in mk_recordT (map #extension (drop n parents)) (Type (c, subst_last HOLogic.unitT Ts)) end; val recT0 = recT 0; fun mk_rec args n = let val (args', more) = chop_last args; fun mk_ext' ((name, T), args) more = mk_ext (name, T) (args @ [more]); fun build Ts = fold_rev mk_ext' (drop n ((extension_names ~~ Ts) ~~ chunks parent_chunks args')) more; in if more = HOLogic.unit then build (map_range recT (parent_len + 1)) else build (map_range rec_schemeT (parent_len + 1)) end; val r_rec0 = mk_rec all_vars_more 0; val r_rec_unit0 = mk_rec (all_vars @ [HOLogic.unit]) 0; fun r n = Free (rN, rec_schemeT n); val r0 = r 0; fun r_unit n = Free (rN, recT n); val r_unit0 = r_unit 0; (* print translations *) val record_ext_type_abbr_tr's = let val trname = hd extension_names; val last_ext = unsuffix ext_typeN (fst extension); in [record_ext_type_abbr_tr' name alphas zeta last_ext rec_schemeT0 trname] end; val record_ext_type_tr's = let (*avoid conflict with record_type_abbr_tr's*) val trnames = if parent_len > 0 then [extension_name] else []; in map record_ext_type_tr' trnames end; val print_translation = map field_update_tr' (full_moreN :: names) @ [record_ext_tr' extension_name] @ record_ext_type_tr's @ record_ext_type_abbr_tr's; (* prepare declarations *) val sel_decls = map (mk_selC rec_schemeT0 o apfst Binding.name_of) bfields_more; val upd_decls = map (mk_updC updateN rec_schemeT0 o apfst Binding.name_of) bfields_more; val make_decl = (makeN, all_types ---> recT0); val fields_decl = (fields_selN, types ---> Type extension); val extend_decl = (extendN, recT0 --> moreT --> rec_schemeT0); val truncate_decl = (truncateN, rec_schemeT0 --> recT0); (* prepare definitions *) val ext_defs = ext_def :: map #ext_def parents; (*Theorems from the iso_tuple intros. By unfolding ext_defs from r_rec0 we create a tree of constructor calls (many of them Pair, but others as well). The introduction rules for update_accessor_eq_assist can unify two different ways on these constructors. If we take the complete result sequence of running a the introduction tactic, we get one theorem for each upd/acc pair, from which we can derive the bodies of our selector and updator and their convs.*) val (accessor_thms, updator_thms, upd_acc_cong_assists) = timeit_msg ext_ctxt "record getting tree access/updates:" (fn () => let val r_rec0_Vars = let (*pick variable indices of 1 to avoid possible variable collisions with existing variables in updacc_eq_triv*) fun to_Var (Free (c, T)) = Var ((c, 1), T); in mk_rec (map to_Var all_vars_more) 0 end; val init_thm = infer_instantiate ext_ctxt [(("v", 0), Thm.cterm_of ext_ctxt r_rec0), (("v'", 0), Thm.cterm_of ext_ctxt r_rec0_Vars)] updacc_eq_triv; val terminal = resolve_tac ext_ctxt [updacc_eq_idI] 1 THEN resolve_tac ext_ctxt [refl] 1; val tactic = simp_tac (put_simpset HOL_basic_ss ext_ctxt addsimps ext_defs) 1 THEN REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac ext_ctxt 1 ORELSE terminal); val updaccs = Seq.list_of (tactic init_thm); in (updaccs RL [updacc_accessor_eqE], updaccs RL [updacc_updator_eqE], updaccs RL [updacc_cong_from_eq]) end); fun lastN xs = drop parent_fields_len xs; (*selectors*) fun mk_sel_spec ((c, T), thm) = let val (acc $ arg, _) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm))); val _ = if arg aconv r_rec0 then () else raise TERM ("mk_sel_spec: different arg", [arg]); in Const (mk_selC rec_schemeT0 (c, T)) :== acc end; val sel_specs = map mk_sel_spec (fields_more ~~ lastN accessor_thms); (*updates*) fun mk_upd_spec ((c, T), thm) = let val (upd $ _ $ arg, _) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Envir.beta_eta_contract (Thm.concl_of thm))); val _ = if arg aconv r_rec0 then () else raise TERM ("mk_sel_spec: different arg", [arg]); in Const (mk_updC updateN rec_schemeT0 (c, T)) :== upd end; val upd_specs = map mk_upd_spec (fields_more ~~ lastN updator_thms); (*derived operations*) val make_spec = list_comb (Const (full (Binding.name makeN), all_types ---> recT0), all_vars) :== mk_rec (all_vars @ [HOLogic.unit]) 0; val fields_spec = list_comb (Const (full (Binding.name fields_selN), types ---> Type extension), vars) :== mk_rec (all_vars @ [HOLogic.unit]) parent_len; val extend_spec = Const (full (Binding.name extendN), recT0 --> moreT --> rec_schemeT0) $ r_unit0 $ more :== mk_rec ((map (mk_sel r_unit0) all_fields) @ [more]) 0; val truncate_spec = Const (full (Binding.name truncateN), rec_schemeT0 --> recT0) $ r0 :== mk_rec ((map (mk_sel r0) all_fields) @ [HOLogic.unit]) 0; (* 2st stage: defs_thy *) val (((sel_defs, upd_defs), derived_defs), defs_thy) = timeit_msg ext_ctxt "record trfuns/tyabbrs/selectors/updates/make/fields/extend/truncate defs:" (fn () => ext_thy |> Sign.print_translation print_translation |> Sign.restore_naming thy0 |> Typedecl.abbrev_global (binding, map #1 alphas, NoSyn) recT0 |> snd |> Typedecl.abbrev_global (Binding.suffix_name schemeN binding, map #1 (alphas @ [zeta]), NoSyn) rec_schemeT0 |> snd |> Sign.qualified_path false binding |> fold (fn ((x, T), mx) => snd o Sign.declare_const_global ((Binding.name x, T), mx)) (sel_decls ~~ (field_syntax @ [NoSyn])) |> fold (fn (x, T) => snd o Sign.declare_const_global ((Binding.name x, T), NoSyn)) (upd_decls @ [make_decl, fields_decl, extend_decl, truncate_decl]) |> (Global_Theory.add_defs false o map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) sel_specs ||>> (Global_Theory.add_defs false o map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) upd_specs ||>> (Global_Theory.add_defs false o map (Thm.no_attributes o apfst (Binding.concealed o Binding.name))) [make_spec, fields_spec, extend_spec, truncate_spec]); val defs_ctxt = Proof_Context.init_global defs_thy; (* prepare propositions *) val _ = timing_msg defs_ctxt "record preparing propositions"; val P = Free (singleton (Name.variant_list all_variants) "P", rec_schemeT0 --> HOLogic.boolT); val C = Free (singleton (Name.variant_list all_variants) "C", HOLogic.boolT); val P_unit = Free (singleton (Name.variant_list all_variants) "P", recT0 --> HOLogic.boolT); (*selectors*) val sel_conv_props = map (fn (c, x as Free (_, T)) => mk_sel r_rec0 (c, T) === x) named_vars_more; (*updates*) fun mk_upd_prop i (c, T) = let val x' = Free (singleton (Name.variant_list all_variants) (Long_Name.base_name c ^ "'"), T --> T); val n = parent_fields_len + i; val args' = nth_map n (K (x' $ nth all_vars_more n)) all_vars_more; in mk_upd updateN c x' r_rec0 === mk_rec args' 0 end; val upd_conv_props = map2 mk_upd_prop idxms fields_more; (*induct*) val induct_scheme_prop = fold_rev Logic.all all_vars_more (Trueprop (P $ r_rec0)) ==> Trueprop (P $ r0); val induct_prop = (fold_rev Logic.all all_vars (Trueprop (P_unit $ r_rec_unit0)), Trueprop (P_unit $ r_unit0)); (*surjective*) val surjective_prop = let val args = map (fn (c, Free (_, T)) => mk_sel r0 (c, T)) all_named_vars_more in r0 === mk_rec args 0 end; (*cases*) val cases_scheme_prop = (fold_rev Logic.all all_vars_more ((r0 === r_rec0) ==> Trueprop C), Trueprop C); val cases_prop = fold_rev Logic.all all_vars ((r_unit0 === r_rec_unit0) ==> Trueprop C) ==> Trueprop C; (*split*) val split_meta_prop = let val P = Free (singleton (Name.variant_list all_variants) "P", rec_schemeT0 --> propT); in Logic.mk_equals (Logic.all r0 (P $ r0), fold_rev Logic.all all_vars_more (P $ r_rec0)) end; val split_object_prop = let val ALL = fold_rev (fn (v, T) => fn t => HOLogic.mk_all (v, T, t)) in ALL [dest_Free r0] (P $ r0) === ALL (map dest_Free all_vars_more) (P $ r_rec0) end; val split_ex_prop = let val EX = fold_rev (fn (v, T) => fn t => HOLogic.mk_exists (v, T, t)) in EX [dest_Free r0] (P $ r0) === EX (map dest_Free all_vars_more) (P $ r_rec0) end; (*equality*) val equality_prop = let val s' = Free (rN ^ "'", rec_schemeT0); fun mk_sel_eq (c, Free (_, T)) = mk_sel r0 (c, T) === mk_sel s' (c, T); val seleqs = map mk_sel_eq all_named_vars_more; in Logic.all r0 (Logic.all s' (Logic.list_implies (seleqs, r0 === s'))) end; (* 3rd stage: thms_thy *) val record_ss = get_simpset defs_thy; val sel_upd_ctxt = put_simpset record_ss defs_ctxt addsimps (sel_defs @ accessor_thms @ upd_defs @ updator_thms); val (sel_convs, upd_convs) = timeit_msg defs_ctxt "record sel_convs/upd_convs proof:" (fn () => grouped 10 Par_List.map (fn prop => Goal.prove_sorry_global defs_thy [] [] prop (fn _ => ALLGOALS (asm_full_simp_tac sel_upd_ctxt))) (sel_conv_props @ upd_conv_props)) |> chop (length sel_conv_props); val (fold_congs, unfold_congs) = timeit_msg defs_ctxt "record upd fold/unfold congs:" (fn () => let val symdefs = map Thm.symmetric (sel_defs @ upd_defs); val fold_ctxt = put_simpset HOL_basic_ss defs_ctxt addsimps symdefs; val ua_congs = map (Drule.export_without_context o simplify fold_ctxt) upd_acc_cong_assists; in (ua_congs RL [updacc_foldE], ua_congs RL [updacc_unfoldE]) end); val parent_induct = Option.map #induct_scheme (try List.last parents); val induct_scheme = timeit_msg defs_ctxt "record induct_scheme proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] induct_scheme_prop (fn {context = ctxt, ...} => EVERY [case parent_induct of NONE => all_tac | SOME ind => try_param_tac ctxt rN ind 1, try_param_tac ctxt rN ext_induct 1, asm_simp_tac (put_simpset HOL_basic_ss ctxt) 1])); val induct = timeit_msg defs_ctxt "record induct proof:" (fn () => Goal.prove_sorry_global defs_thy [] [#1 induct_prop] (#2 induct_prop) (fn {context = ctxt, prems, ...} => try_param_tac ctxt rN induct_scheme 1 THEN try_param_tac ctxt "more" @{thm unit.induct} 1 THEN resolve_tac ctxt prems 1)); val surjective = timeit_msg defs_ctxt "record surjective proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] surjective_prop (fn {context = ctxt, ...} => EVERY [resolve_tac ctxt [surject_assist_idE] 1, simp_tac (put_simpset HOL_basic_ss ctxt addsimps ext_defs) 1, REPEAT (Iso_Tuple_Support.iso_tuple_intros_tac ctxt 1 ORELSE (resolve_tac ctxt [surject_assistI] 1 THEN simp_tac (put_simpset (get_sel_upd_defs defs_thy) ctxt addsimps (sel_defs @ @{thms o_assoc id_apply id_o o_id})) 1))])); val cases_scheme = timeit_msg defs_ctxt "record cases_scheme proof:" (fn () => Goal.prove_sorry_global defs_thy [] [#1 cases_scheme_prop] (#2 cases_scheme_prop) (fn {context = ctxt, prems, ...} => resolve_tac ctxt prems 1 THEN resolve_tac ctxt [surjective] 1)); val cases = timeit_msg defs_ctxt "record cases proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] cases_prop (fn {context = ctxt, ...} => try_param_tac ctxt rN cases_scheme 1 THEN ALLGOALS (asm_full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms unit_all_eq1})))); val split_meta = timeit_msg defs_ctxt "record split_meta proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_meta_prop (fn {context = ctxt', ...} => EVERY1 [resolve_tac ctxt' @{thms equal_intr_rule}, Goal.norm_hhf_tac ctxt', eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt', resolve_tac ctxt' [@{thm prop_subst} OF [surjective]], REPEAT o eresolve_tac ctxt' @{thms meta_allE}, assume_tac ctxt'])); val split_object = timeit_msg defs_ctxt "record split_object proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_object_prop (fn {context = ctxt, ...} => resolve_tac ctxt [@{lemma "Trueprop A \ Trueprop B \ A = B" by (rule iffI) unfold}] 1 THEN rewrite_goals_tac ctxt @{thms atomize_all [symmetric]} THEN resolve_tac ctxt [split_meta] 1)); val split_ex = timeit_msg defs_ctxt "record split_ex proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] split_ex_prop (fn {context = ctxt, ...} => simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{lemma "\x. P x \ \ (\x. \ P x)" by simp} :: @{thms not_not Not_eq_iff})) 1 THEN resolve_tac ctxt [split_object] 1)); val equality = timeit_msg defs_ctxt "record equality proof:" (fn () => Goal.prove_sorry_global defs_thy [] [] equality_prop (fn {context = ctxt, ...} => asm_full_simp_tac (put_simpset record_ss ctxt addsimps (split_meta :: sel_convs)) 1)); val ([(_, sel_convs'), (_, upd_convs'), (_, sel_defs'), (_, upd_defs'), (_, fold_congs'), (_, unfold_congs'), (_, splits' as [split_meta', split_object', split_ex']), (_, derived_defs'), (_, [surjective']), (_, [equality']), (_, [induct_scheme']), (_, [induct']), (_, [cases_scheme']), (_, [cases'])], thms_thy) = defs_thy |> Code.declare_default_eqns_global (map (rpair true) derived_defs) |> Global_Theory.note_thmss "" [((Binding.name "select_convs", []), [(sel_convs, [])]), ((Binding.name "update_convs", []), [(upd_convs, [])]), ((Binding.name "select_defs", []), [(sel_defs, [])]), ((Binding.name "update_defs", []), [(upd_defs, [])]), ((Binding.name "fold_congs", []), [(fold_congs, [])]), ((Binding.name "unfold_congs", []), [(unfold_congs, [])]), ((Binding.name "splits", []), [([split_meta, split_object, split_ex], [])]), ((Binding.name "defs", []), [(derived_defs, [])]), ((Binding.name "surjective", []), [([surjective], [])]), ((Binding.name "equality", []), [([equality], [])]), ((Binding.name "induct_scheme", induct_type_global (suffix schemeN name)), [([induct_scheme], [])]), ((Binding.name "induct", induct_type_global name), [([induct], [])]), ((Binding.name "cases_scheme", cases_type_global (suffix schemeN name)), [([cases_scheme], [])]), ((Binding.name "cases", cases_type_global name), [([cases], [])])]; val sel_upd_simps = sel_convs' @ upd_convs'; val sel_upd_defs = sel_defs' @ upd_defs'; val depth = parent_len + 1; val ([(_, simps'), (_, iffs')], thms_thy') = thms_thy |> Global_Theory.note_thmss "" [((Binding.name "simps", [Simplifier.simp_add]), [(sel_upd_simps, [])]), ((Binding.name "iffs", [iff_add]), [([ext_inject], [])])]; val info = make_info alphas parent fields extension ext_induct ext_inject ext_surjective ext_split ext_def sel_convs' upd_convs' sel_defs' upd_defs' fold_congs' unfold_congs' splits' derived_defs' surjective' equality' induct_scheme' induct' cases_scheme' cases' simps' iffs'; val final_thy = thms_thy' |> put_record name info |> put_sel_upd names full_moreN depth sel_upd_simps sel_upd_defs |> add_equalities extension_id equality' |> add_extinjects ext_inject |> add_extsplit extension_name ext_split |> add_splits extension_id (split_meta', split_object', split_ex', induct_scheme') |> add_extfields extension_name (fields @ [(full_moreN, moreT)]) |> add_fieldext (extension_name, snd extension) names |> add_code ext_tyco vs extT ext simps' ext_inject |> add_ctr_sugar (Const ext) cases_scheme' ext_inject sel_convs' |> fold add_spec_rule (sel_convs' @ upd_convs' @ derived_defs') |> Sign.restore_naming thy0; in final_thy end; (* add_record *) local fun read_parent NONE ctxt = (NONE, ctxt) | read_parent (SOME raw_T) ctxt = (case Proof_Context.read_typ_abbrev ctxt raw_T of Type (name, Ts) => (SOME (Ts, name), fold Variable.declare_typ Ts ctxt) | T => error ("Bad parent record specification: " ^ Syntax.string_of_typ ctxt T)); fun read_fields raw_fields ctxt = let val Ts = Syntax.read_typs ctxt (map (fn (_, raw_T, _) => raw_T) raw_fields); val fields = map2 (fn (x, _, mx) => fn T => (x, T, mx)) raw_fields Ts; val ctxt' = fold Variable.declare_typ Ts ctxt; in (fields, ctxt') end; in fun add_record overloaded (params, binding) raw_parent raw_fields thy = let val ctxt = Proof_Context.init_global thy; fun cert_typ T = Type.no_tvars (Proof_Context.cert_typ ctxt T) handle TYPE (msg, _, _) => error msg; (* specification *) val parent = Option.map (apfst (map cert_typ)) raw_parent handle ERROR msg => cat_error msg ("The error(s) above occurred in parent record specification"); val parent_args = (case parent of SOME (Ts, _) => Ts | NONE => []); val parents = get_parent_info thy parent; val bfields = raw_fields |> map (fn (x, raw_T, mx) => (x, cert_typ raw_T, mx) handle ERROR msg => cat_error msg ("The error(s) above occurred in record field " ^ Binding.print x)); (* errors *) val name = Sign.full_name thy binding; val err_dup_record = if is_none (get_info thy name) then [] else ["Duplicate definition of record " ^ quote name]; val spec_frees = fold Term.add_tfreesT (parent_args @ map #2 bfields) []; val err_extra_frees = (case subtract (op =) params spec_frees of [] => [] | extras => ["Extra free type variable(s) " ^ commas (map (Syntax.string_of_typ ctxt o TFree) extras)]); val err_no_fields = if null bfields then ["No fields present"] else []; val err_dup_fields = (case duplicates Binding.eq_name (map #1 bfields) of [] => [] | dups => ["Duplicate field(s) " ^ commas (map Binding.print dups)]); val err_bad_fields = if forall (not_equal moreN o Binding.name_of o #1) bfields then [] else ["Illegal field name " ^ quote moreN]; val errs = err_dup_record @ err_extra_frees @ err_no_fields @ err_dup_fields @ err_bad_fields; val _ = if null errs then () else error (cat_lines errs); in thy |> definition overloaded (params, binding) parent parents bfields end handle ERROR msg => cat_error msg ("Failed to define record " ^ Binding.print binding); fun add_record_cmd overloaded (raw_params, binding) raw_parent raw_fields thy = let val ctxt = Proof_Context.init_global thy; val params = map (apsnd (Typedecl.read_constraint ctxt)) raw_params; val ctxt1 = fold (Variable.declare_typ o TFree) params ctxt; val (parent, ctxt2) = read_parent raw_parent ctxt1; val (fields, ctxt3) = read_fields raw_fields ctxt2; val params' = map (Proof_Context.check_tfree ctxt3) params; in thy |> add_record overloaded (params', binding) parent fields end; end; (* printing *) local fun the_parent_recT (Type (parent, [Type (_, [unit as Type (_,[])])])) = Type (parent, [unit]) | the_parent_recT (Type (extT, [T])) = Type (extT, [the_parent_recT T]) | the_parent_recT T = raise TYPE ("Not a unit record scheme with parent: ", [T], []) in fun pretty_recT ctxt typ = let val thy = Proof_Context.theory_of ctxt val (fs, (_, moreT)) = get_recT_fields thy typ val _ = if moreT = HOLogic.unitT then () else raise TYPE ("Not a unit record scheme: ", [typ], []) val parent = if length (dest_recTs typ) >= 2 then SOME (the_parent_recT typ) else NONE val pfs = case parent of SOME p => fst (get_recT_fields thy p) | NONE => [] val fs' = drop (length pfs) fs fun pretty_field (name, typ) = Pretty.block [ Syntax.pretty_term ctxt (Const (name, typ)), Pretty.brk 1, Pretty.str "::", Pretty.brk 1, Syntax.pretty_typ ctxt typ ] in Pretty.block (Library.separate (Pretty.brk 1) ([Pretty.keyword1 "record", Syntax.pretty_typ ctxt typ, Pretty.str "="] @ (case parent of SOME p => [Syntax.pretty_typ ctxt p, Pretty.str "+"] | NONE => [])) @ Pretty.fbrk :: Pretty.fbreaks (map pretty_field fs')) end end fun string_of_record ctxt s = let val T = Syntax.read_typ ctxt s in Pretty.string_of (pretty_recT ctxt T) handle TYPE _ => error ("Unknown record: " ^ Syntax.string_of_typ ctxt T) end val print_record = let fun print_item string_of (modes, arg) = Toplevel.keep (fn state => Print_Mode.with_modes modes (fn () => Output.writeln (string_of state arg)) ()); in print_item (string_of_record o Toplevel.context_of) end (* outer syntax *) val _ = Outer_Syntax.command \<^command_keyword>\record\ "define extensible record" (Parse_Spec.overloaded -- (Parse.type_args_constrained -- Parse.binding) -- (\<^keyword>\=\ |-- Scan.option (Parse.typ --| \<^keyword>\+\) -- Scan.repeat1 Parse.const_binding) >> (fn ((overloaded, x), (y, z)) => Toplevel.theory (add_record_cmd {overloaded = overloaded} x y z))); val opt_modes = Scan.optional (\<^keyword>\(\ |-- Parse.!!! (Scan.repeat1 Parse.name --| \<^keyword>\)\)) [] val _ = Outer_Syntax.command \<^command_keyword>\print_record\ "print record definiton" (opt_modes -- Parse.typ >> print_record); end diff --git a/src/HOL/Tools/reification.ML b/src/HOL/Tools/reification.ML --- a/src/HOL/Tools/reification.ML +++ b/src/HOL/Tools/reification.ML @@ -1,298 +1,299 @@ (* Title: HOL/Tools/reification.ML Author: Amine Chaieb, TU Muenchen A trial for automatical reification. *) signature REIFICATION = sig val conv: Proof.context -> thm list -> conv val tac: Proof.context -> thm list -> term option -> int -> tactic val lift_conv: Proof.context -> conv -> term option -> int -> tactic val dereify: Proof.context -> thm list -> conv end; structure Reification : REIFICATION = struct fun dest_listT (Type (\<^type_name>\list\, [T])) = T; val FWD = curry (op OF); fun rewrite_with ctxt eqs = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps eqs); val pure_subst = @{lemma "x == y ==> PROP P y ==> PROP P x" by simp} fun lift_conv ctxt conv some_t = Subgoal.FOCUS (fn {context = ctxt', concl, ...} => let val ct = (case some_t of NONE => Thm.dest_arg concl | SOME t => Thm.cterm_of ctxt' t) val thm = conv ct; in if Thm.is_reflexive thm then no_tac else ALLGOALS (resolve_tac ctxt [pure_subst OF [thm]]) end) ctxt; (* Make a congruence rule out of a defining equation for the interpretation th is one defining equation of f, i.e. th is "f (Cp ?t1 ... ?tn) = P(f ?t1, .., f ?tn)" Cp is a constructor pattern and P is a pattern The result is: [|?A1 = f ?t1 ; .. ; ?An= f ?tn |] ==> P (?A1, .., ?An) = f (Cp ?t1 .. ?tn) + the a list of names of the A1 .. An, Those are fresh in the ctxt *) fun mk_congeq ctxt fs th = let val Const (fN, _) = th |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst; val ((_, [th']), ctxt') = Variable.import true [th] ctxt; val (lhs, rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of th')); fun add_fterms (t as t1 $ t2) = if exists (fn f => Term.could_unify (t |> strip_comb |> fst, f)) fs then insert (op aconv) t else add_fterms t1 #> add_fterms t2 | add_fterms (t as Abs _) = if exists_Const (fn (c, _) => c = fN) t then K [t] else K [] | add_fterms _ = I; val fterms = add_fterms rhs []; val (xs, ctxt'') = Variable.variant_fixes (replicate (length fterms) "x") ctxt'; val tys = map fastype_of fterms; val vs = map Free (xs ~~ tys); val env = fterms ~~ vs; (*FIXME*) fun replace_fterms (t as t1 $ t2) = (case AList.lookup (op aconv) env t of SOME v => v | NONE => replace_fterms t1 $ replace_fterms t2) | replace_fterms t = (case AList.lookup (op aconv) env t of SOME v => v | NONE => t); fun mk_def (Abs (x, xT, t), v) = HOLogic.mk_Trueprop (HOLogic.all_const xT $ Abs (x, xT, HOLogic.mk_eq (v $ Bound 0, t))) | mk_def (t, v) = HOLogic.mk_Trueprop (HOLogic.mk_eq (v, t)); fun tryext x = (x RS @{lemma "(\x. f x = g x) \ f = g" by blast} handle THM _ => x); val cong = (Goal.prove ctxt'' [] (map mk_def env) (HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs, replace_fterms rhs))) (fn {context, prems, ...} => Local_Defs.unfold0_tac context (map tryext prems) THEN resolve_tac ctxt'' [th'] 1)) RS sym; val (cong' :: vars') = Variable.export ctxt'' ctxt (cong :: map (Drule.mk_term o Thm.cterm_of ctxt'') vs); val vs' = map (fst o fst o Term.dest_Var o Thm.term_of o Drule.dest_term) vars'; in (vs', cong') end; (* congs is a list of pairs (P,th) where th is a theorem for [| f p1 = A1; ...; f pn = An|] ==> f (C p1 .. pn) = P *) fun rearrange congs = let fun P (_, th) = let val \<^term>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ l $ _) = Thm.concl_of th in can dest_Var l end; val (yes, no) = List.partition P congs; in no @ yes end; fun dereify ctxt eqs = rewrite_with ctxt (eqs @ @{thms nth_Cons_0 nth_Cons_Suc}); fun index_of t bds = let val tt = HOLogic.listT (fastype_of t); in (case AList.lookup Type.could_unify bds tt of NONE => error "index_of: type not found in environements!" | SOME (tbs, tats) => let val i = find_index (fn t' => t' = t) tats; val j = find_index (fn t' => t' = t) tbs; in if j = ~1 then if i = ~1 then (length tbs + length tats, AList.update Type.could_unify (tt, (tbs, tats @ [t])) bds) else (i, bds) else (j, bds) end) end; (* Generic decomp for reification : matches the actual term with the rhs of one cong rule. The result of the matching guides the proof synthesis: The matches of the introduced Variables A1 .. An are processed recursively The rest is instantiated in the cong rule,i.e. no reification is needed *) (* da is the decomposition for atoms, ie. it returns ([],g) where g returns the right instance f (AtC n) = t , where AtC is the Atoms constructor and n is the number of the atom corresponding to t *) fun decomp_reify da cgns (ct, ctxt) bds = let val thy = Proof_Context.theory_of ctxt; fun tryabsdecomp (ct, ctxt) bds = (case Thm.term_of ct of Abs (_, xT, ta) => let val ([raw_xn], ctxt') = Variable.variant_fixes ["x"] ctxt; val (xn, ta) = Syntax_Trans.variant_abs (raw_xn, xT, ta); (* FIXME !? *) val x = Free (xn, xT); val cx = Thm.cterm_of ctxt' x; val cta = Thm.cterm_of ctxt' ta; val bds = (case AList.lookup Type.could_unify bds (HOLogic.listT xT) of NONE => error "tryabsdecomp: Type not found in the Environement" | SOME (bsT, atsT) => AList.update Type.could_unify (HOLogic.listT xT, (x :: bsT, atsT)) bds); in (([(cta, ctxt')], fn ([th], bds) => (hd (Variable.export ctxt' ctxt [(Thm.forall_intr cx th) COMP allI]), let val (bsT, asT) = the (AList.lookup Type.could_unify bds (HOLogic.listT xT)); in AList.update Type.could_unify (HOLogic.listT xT, (tl bsT, asT)) bds end)), bds) end | _ => da (ct, ctxt) bds) in (case cgns of [] => tryabsdecomp (ct, ctxt) bds | ((vns, cong) :: congs) => (let val (tyenv, tmenv) = Pattern.match thy ((fst o HOLogic.dest_eq o HOLogic.dest_Trueprop) (Thm.concl_of cong), Thm.term_of ct) (Vartab.empty, Vartab.empty); val (fnvs, invs) = List.partition (fn ((vn, _),_) => member (op =) vns vn) (Vartab.dest tmenv); val (fts, its) = (map (snd o snd) fnvs, map (fn ((vn, vi), (tT, t)) => (((vn, vi), tT), Thm.cterm_of ctxt t)) invs); val ctyenv = map (fn ((vn, vi), (s, ty)) => (((vn, vi), s), Thm.ctyp_of ctxt ty)) (Vartab.dest tyenv); in ((map (Thm.cterm_of ctxt) fts ~~ replicate (length fts) ctxt, - apfst (FWD (Drule.instantiate_normalize (ctyenv, its) cong))), bds) + apfst (FWD (Drule.instantiate_normalize (TVars.make ctyenv, Vars.make its) cong))), bds) end handle Pattern.MATCH => decomp_reify da congs (ct, ctxt) bds)) end; fun get_nths (t as (Const (\<^const_name>\List.nth\, _) $ vs $ n)) = AList.update (op aconv) (t, (vs, n)) | get_nths (t1 $ t2) = get_nths t1 #> get_nths t2 | get_nths (Abs (_, _, t')) = get_nths t' | get_nths _ = I; fun tryeqs [] (ct, ctxt) bds = error "Cannot find the atoms equation" | tryeqs (eq :: eqs) (ct, ctxt) bds = (( let val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd; val nths = get_nths rhs []; val (vss, _) = fold_rev (fn (_, (vs, n)) => fn (vss, ns) => (insert (op aconv) vs vss, insert (op aconv) n ns)) nths ([], []); val (vsns, ctxt') = Variable.variant_fixes (replicate (length vss) "vs") ctxt; val (xns, ctxt'') = Variable.variant_fixes (replicate (length nths) "x") ctxt'; val thy = Proof_Context.theory_of ctxt''; val vsns_map = vss ~~ vsns; val xns_map = fst (split_list nths) ~~ xns; val subst = map (fn (nt, xn) => (nt, Var ((xn, 0), fastype_of nt))) xns_map; val rhs_P = subst_free subst rhs; val (tyenv, tmenv) = Pattern.match thy (rhs_P, Thm.term_of ct) (Vartab.empty, Vartab.empty); val sbst = Envir.subst_term (tyenv, tmenv); val sbsT = Envir.subst_type tyenv; val subst_ty = map (fn (n, (s, t)) => ((n, s), Thm.ctyp_of ctxt'' t)) (Vartab.dest tyenv) val tml = Vartab.dest tmenv; val (subst_ns, bds) = fold_map (fn (Const _ $ _ $ n, Var (xn0, _)) => fn bds => let val name = snd (the (AList.lookup (op =) tml xn0)); val (idx, bds) = index_of name bds; in (apply2 (Thm.cterm_of ctxt'') (n, idx |> HOLogic.mk_nat), bds) end) subst bds; val subst_vs = let fun h (Const _ $ (vs as Var (_, lT)) $ _, Var (_, T)) = let val cns = sbst (Const (\<^const_name>\List.Cons\, T --> lT --> lT)); val lT' = sbsT lT; val (bsT, _) = the (AList.lookup Type.could_unify bds lT); val vsn = the (AList.lookup (op =) vsns_map vs); val vs' = fold_rev (fn x => fn xs => cns $ x $xs) bsT (Free (vsn, lT')); in apply2 (Thm.cterm_of ctxt'') (vs, vs') end; in map h subst end; val cts = map (fn ((vn, vi), (tT, t)) => apply2 (Thm.cterm_of ctxt'') (Var ((vn, vi), tT), t)) (fold (AList.delete (fn (((a : string), _), (b, _)) => a = b)) (map (fn n => (n, 0)) xns) tml); val substt = let - val ih = Drule.cterm_rule (Thm.instantiate (subst_ty, [])); + val ih = Drule.cterm_rule (Thm.instantiate (TVars.make subst_ty, Vars.empty)); in map (apply2 ih) (subst_ns @ subst_vs @ cts) end; val th = - (Drule.instantiate_normalize (subst_ty, map (apfst (dest_Var o Thm.term_of)) substt) eq) + (Drule.instantiate_normalize + (TVars.make subst_ty, Vars.make (map (apfst (dest_Var o Thm.term_of)) substt)) eq) RS sym; in (hd (Variable.export ctxt'' ctxt [th]), bds) end) handle Pattern.MATCH => tryeqs eqs (ct, ctxt) bds); (* looks for the atoms equation and instantiates it with the right number *) fun mk_decompatom eqs (ct, ctxt) bds = (([], fn (_, bds) => let val tT = fastype_of (Thm.term_of ct); fun isat eq = let val rhs = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> snd; in exists_Const (fn (n, ty) => n = \<^const_name>\List.nth\ andalso AList.defined Type.could_unify bds (domain_type ty)) rhs andalso Type.could_unify (fastype_of rhs, tT) end; in tryeqs (filter isat eqs) (ct, ctxt) bds end), bds); (* Generic reification procedure: *) (* creates all needed cong rules and then just uses the theorem synthesis *) fun mk_congs ctxt eqs = let val fs = fold_rev (fn eq => insert (op =) (eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb |> fst)) eqs []; val tys = fold_rev (fn f => fold (insert (op =)) (f |> fastype_of |> binder_types |> tl)) fs []; val (vs, ctxt') = Variable.variant_fixes (replicate (length tys) "vs") ctxt; val subst = the o AList.lookup (op =) (map2 (fn T => fn v => (T, Thm.cterm_of ctxt' (Free (v, T)))) tys vs); fun prep_eq eq = let val (_, _ :: vs) = eq |> Thm.prop_of |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst |> strip_comb; val subst = map_filter (fn Var v => SOME (v, subst (#2 v)) | _ => NONE) vs; - in Thm.instantiate ([], subst) eq end; + in Thm.instantiate (TVars.empty, Vars.make subst) eq end; val (ps, congs) = map_split (mk_congeq ctxt' fs o prep_eq) eqs; val bds = AList.make (K ([], [])) tys; in (ps ~~ Variable.export ctxt' ctxt congs, bds) end fun conv ctxt eqs ct = let val (congs, bds) = mk_congs ctxt eqs; val congs = rearrange congs; val (th, bds') = apfst mk_eq (divide_and_conquer' (decomp_reify (mk_decompatom eqs) congs) (ct, ctxt) bds); fun is_list_var (Var (_, t)) = can dest_listT t | is_list_var _ = false; val vars = th |> Thm.prop_of |> Logic.dest_equals |> snd |> strip_comb |> snd |> filter is_list_var; val vs = map (fn Var (v as (_, T)) => (v, the (AList.lookup Type.could_unify bds' T) |> snd |> HOLogic.mk_list (dest_listT T))) vars; val th' = - Drule.instantiate_normalize ([], map (apsnd (Thm.cterm_of ctxt)) vs) th; + Drule.instantiate_normalize (TVars.empty, Vars.make (map (apsnd (Thm.cterm_of ctxt)) vs)) th; val th'' = Thm.symmetric (dereify ctxt [] (Thm.lhs_of th')); in Thm.transitive th'' th' end; fun tac ctxt eqs = lift_conv ctxt (conv ctxt eqs); end; diff --git a/src/HOL/Tools/sat.ML b/src/HOL/Tools/sat.ML --- a/src/HOL/Tools/sat.ML +++ b/src/HOL/Tools/sat.ML @@ -1,470 +1,471 @@ (* Title: HOL/Tools/sat.ML Author: Stephan Merz and Alwen Tiu, QSL Team, LORIA (http://qsl.loria.fr) Author: Tjark Weber, TU Muenchen Proof reconstruction from SAT solvers. Description: This file defines several tactics to invoke a proof-producing SAT solver on a propositional goal in clausal form. We use a sequent presentation of clauses to speed up resolution proof reconstruction. We call such clauses "raw clauses", which are of the form [x1, ..., xn, P] |- False (note the use of |- instead of ==>, i.e. of Isabelle's (meta-)hyps here), where each xi is a literal (see also comments in cnf.ML). This does not work for goals containing schematic variables! The tactic produces a clause representation of the given goal in DIMACS format and invokes a SAT solver, which should return a proof consisting of a sequence of resolution steps, indicating the two input clauses, and resulting in new clauses, leading to the empty clause (i.e. "False"). The tactic replays this proof in Isabelle and thus solves the overall goal. There are three SAT tactics available. They differ in the CNF transformation used. "sat_tac" uses naive CNF transformation to transform the theorem to be proved before giving it to the SAT solver. The naive transformation in the worst case can lead to an exponential blow up in formula size. Another tactic, "satx_tac", uses "definitional CNF transformation" which attempts to produce a formula of linear size increase compared to the input formula, at the cost of possibly introducing new variables. See cnf.ML for more comments on the CNF transformation. "rawsat_tac" should be used with caution: no CNF transformation is performed, and the tactic's behavior is undefined if the subgoal is not already given as [| C1; ...; Cn |] ==> False, where each Ci is a disjunction. The SAT solver to be used can be set via the "solver" reference. See sat_solvers.ML for possible values, and etc/settings for required (solver- dependent) configuration settings. To replay SAT proofs in Isabelle, you must of course use a proof-producing SAT solver in the first place. Proofs are replayed only if "quick_and_dirty" is false. If "quick_and_dirty" is true, the theorem (in case the SAT solver claims its negation to be unsatisfiable) is proved via an oracle. *) signature SAT = sig val trace: bool Config.T (* print trace messages *) val solver: string Config.T (* name of SAT solver to be used *) val counter: int Unsynchronized.ref (* output: number of resolution steps during last proof replay *) val rawsat_thm: Proof.context -> cterm list -> thm val rawsat_tac: Proof.context -> int -> tactic val sat_tac: Proof.context -> int -> tactic val satx_tac: Proof.context -> int -> tactic end structure SAT : SAT = struct val trace = Attrib.setup_config_bool \<^binding>\sat_trace\ (K false); fun cond_tracing ctxt msg = if Config.get ctxt trace then tracing (msg ()) else (); val solver = Attrib.setup_config_string \<^binding>\sat_solver\ (K "cdclite"); (*see HOL/Tools/sat_solver.ML for possible values*) val counter = Unsynchronized.ref 0; val resolution_thm = @{lemma "(P \ False) \ (\ P \ False) \ False" by (rule case_split)} (* ------------------------------------------------------------------------- *) (* lit_ord: an order on integers that considers their absolute values only, *) (* thereby treating integers that represent the same atom (positively *) (* or negatively) as equal *) (* ------------------------------------------------------------------------- *) fun lit_ord (i, j) = int_ord (abs i, abs j); (* ------------------------------------------------------------------------- *) (* CLAUSE: during proof reconstruction, three kinds of clauses are *) (* distinguished: *) (* 1. NO_CLAUSE: clause not proved (yet) *) (* 2. ORIG_CLAUSE: a clause as it occurs in the original problem *) (* 3. RAW_CLAUSE: a raw clause, with additional precomputed information *) (* (a mapping from int's to its literals) for faster proof *) (* reconstruction *) (* ------------------------------------------------------------------------- *) datatype CLAUSE = NO_CLAUSE | ORIG_CLAUSE of thm | RAW_CLAUSE of thm * (int * cterm) list; (* ------------------------------------------------------------------------- *) (* resolve_raw_clauses: given a non-empty list of raw clauses, we fold *) (* resolution over the list (starting with its head), i.e. with two raw *) (* clauses *) (* [P, x1, ..., a, ..., xn] |- False *) (* and *) (* [Q, y1, ..., a', ..., ym] |- False *) (* (where a and a' are dual to each other), we convert the first clause *) (* to *) (* [P, x1, ..., xn] |- a ==> False , *) (* the second clause to *) (* [Q, y1, ..., ym] |- a' ==> False *) (* and then perform resolution with *) (* [| ?P ==> False; ~?P ==> False |] ==> False *) (* to produce *) (* [P, Q, x1, ..., xn, y1, ..., ym] |- False *) (* Each clause is accompanied with an association list mapping integers *) (* (positive for positive literals, negative for negative literals, and *) (* the same absolute value for dual literals) to the actual literals as *) (* cterms. *) (* ------------------------------------------------------------------------- *) fun resolve_raw_clauses _ [] = raise THM ("Proof reconstruction failed (empty list of resolvents)!", 0, []) | resolve_raw_clauses ctxt (c::cs) = let (* merges two sorted lists wrt. 'lit_ord', suppressing duplicates *) fun merge xs [] = xs | merge [] ys = ys | merge (x :: xs) (y :: ys) = (case lit_ord (apply2 fst (x, y)) of LESS => x :: merge xs (y :: ys) | EQUAL => x :: merge xs ys | GREATER => y :: merge (x :: xs) ys) (* find out which two hyps are used in the resolution *) fun find_res_hyps ([], _) _ = raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) | find_res_hyps (_, []) _ = raise THM ("Proof reconstruction failed (no literal for resolution)!", 0, []) | find_res_hyps (h1 :: hyps1, h2 :: hyps2) acc = (case lit_ord (apply2 fst (h1, h2)) of LESS => find_res_hyps (hyps1, h2 :: hyps2) (h1 :: acc) | EQUAL => let val (i1, chyp1) = h1 val (i2, chyp2) = h2 in if i1 = ~ i2 then (i1 < 0, chyp1, chyp2, rev acc @ merge hyps1 hyps2) else (* i1 = i2 *) find_res_hyps (hyps1, hyps2) (h1 :: acc) end | GREATER => find_res_hyps (h1 :: hyps1, hyps2) (h2 :: acc)) fun resolution (c1, hyps1) (c2, hyps2) = let val _ = cond_tracing ctxt (fn () => "Resolving clause: " ^ Thm.string_of_thm ctxt c1 ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c1) ^ ")\nwith clause: " ^ Thm.string_of_thm ctxt c2 ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c2) ^ ")") (* the two literals used for resolution *) val (hyp1_is_neg, hyp1, hyp2, new_hyps) = find_res_hyps (hyps1, hyps2) [] val c1' = Thm.implies_intr hyp1 c1 (* Gamma1 |- hyp1 ==> False *) val c2' = Thm.implies_intr hyp2 c2 (* Gamma2 |- hyp2 ==> False *) val res_thm = (* |- (lit ==> False) ==> (~lit ==> False) ==> False *) let val cLit = snd (Thm.dest_comb (if hyp1_is_neg then hyp2 else hyp1)) (* strip Trueprop *) in - Thm.instantiate ([], [((("P", 0), HOLogic.boolT), cLit)]) resolution_thm + Thm.instantiate (TVars.empty, Vars.make [((("P", 0), HOLogic.boolT), cLit)]) + resolution_thm end val _ = cond_tracing ctxt (fn () => "Resolution theorem: " ^ Thm.string_of_thm ctxt res_thm) (* Gamma1, Gamma2 |- False *) val c_new = Thm.implies_elim (Thm.implies_elim res_thm (if hyp1_is_neg then c2' else c1')) (if hyp1_is_neg then c1' else c2') val _ = cond_tracing ctxt (fn () => "Resulting clause: " ^ Thm.string_of_thm ctxt c_new ^ " (hyps: " ^ ML_Syntax.print_list (Syntax.string_of_term ctxt) (Thm.hyps_of c_new) ^ ")") val _ = Unsynchronized.inc counter in (c_new, new_hyps) end in fold resolution cs c end; (* ------------------------------------------------------------------------- *) (* replay_proof: replays the resolution proof returned by the SAT solver; *) (* cf. SAT_Solver.proof for details of the proof format. Updates the *) (* 'clauses' array with derived clauses, and returns the derived clause *) (* at index 'empty_id' (which should just be "False" if proof *) (* reconstruction was successful, with the used clauses as hyps). *) (* 'atom_table' must contain an injective mapping from all atoms that *) (* occur (as part of a literal) in 'clauses' to positive integers. *) (* ------------------------------------------------------------------------- *) fun replay_proof ctxt atom_table clauses (clause_table, empty_id) = let fun index_of_literal chyp = (case (HOLogic.dest_Trueprop o Thm.term_of) chyp of (Const (\<^const_name>\Not\, _) $ atom) => SOME (~ (the (Termtab.lookup atom_table atom))) | atom => SOME (the (Termtab.lookup atom_table atom))) handle TERM _ => NONE; (* 'chyp' is not a literal *) fun prove_clause id = (case Array.sub (clauses, id) of RAW_CLAUSE clause => clause | ORIG_CLAUSE thm => (* convert the original clause *) let val _ = cond_tracing ctxt (fn () => "Using original clause #" ^ string_of_int id) val raw = CNF.clause2raw_thm ctxt thm val hyps = sort (lit_ord o apply2 fst) (map_filter (fn chyp => Option.map (rpair chyp) (index_of_literal chyp)) (Thm.chyps_of raw)) val clause = (raw, hyps) val _ = Array.update (clauses, id, RAW_CLAUSE clause) in clause end | NO_CLAUSE => (* prove the clause, using information from 'clause_table' *) let val _ = cond_tracing ctxt (fn () => "Proving clause #" ^ string_of_int id ^ " ...") val ids = the (Inttab.lookup clause_table id) val clause = resolve_raw_clauses ctxt (map prove_clause ids) val _ = Array.update (clauses, id, RAW_CLAUSE clause) val _ = cond_tracing ctxt (fn () => "Replay chain successful; clause stored at #" ^ string_of_int id) in clause end) val _ = counter := 0 val empty_clause = fst (prove_clause empty_id) val _ = cond_tracing ctxt (fn () => "Proof reconstruction successful; " ^ string_of_int (!counter) ^ " resolution step(s) total.") in empty_clause end; (* ------------------------------------------------------------------------- *) (* string_of_prop_formula: return a human-readable string representation of *) (* a 'prop_formula' (just for tracing) *) (* ------------------------------------------------------------------------- *) fun string_of_prop_formula Prop_Logic.True = "True" | string_of_prop_formula Prop_Logic.False = "False" | string_of_prop_formula (Prop_Logic.BoolVar i) = "x" ^ string_of_int i | string_of_prop_formula (Prop_Logic.Not fm) = "~" ^ string_of_prop_formula fm | string_of_prop_formula (Prop_Logic.Or (fm1, fm2)) = "(" ^ string_of_prop_formula fm1 ^ " v " ^ string_of_prop_formula fm2 ^ ")" | string_of_prop_formula (Prop_Logic.And (fm1, fm2)) = "(" ^ string_of_prop_formula fm1 ^ " & " ^ string_of_prop_formula fm2 ^ ")"; (* ------------------------------------------------------------------------- *) (* rawsat_thm: run external SAT solver with the given clauses. Reconstructs *) (* a proof from the resulting proof trace of the SAT solver. The *) (* theorem returned is just "False" (with some of the given clauses as *) (* hyps). *) (* ------------------------------------------------------------------------- *) fun rawsat_thm ctxt clauses = let (* remove premises that equal "True" *) val clauses' = filter (fn clause => (not_equal \<^term>\True\ o HOLogic.dest_Trueprop o Thm.term_of) clause handle TERM ("dest_Trueprop", _) => true) clauses (* remove non-clausal premises -- of course this shouldn't actually *) (* remove anything as long as 'rawsat_tac' is only called after the *) (* premises have been converted to clauses *) val clauses'' = filter (fn clause => ((CNF.is_clause o HOLogic.dest_Trueprop o Thm.term_of) clause handle TERM ("dest_Trueprop", _) => false) orelse (if Context_Position.is_visible ctxt then warning ("Ignoring non-clausal premise " ^ Syntax.string_of_term ctxt (Thm.term_of clause)) else (); false)) clauses' (* remove trivial clauses -- this is necessary because zChaff removes *) (* trivial clauses during preprocessing, and otherwise our clause *) (* numbering would be off *) val nontrivial_clauses = filter (not o CNF.clause_is_trivial o HOLogic.dest_Trueprop o Thm.term_of) clauses'' (* sort clauses according to the term order -- an optimization, *) (* useful because forming the union of hypotheses, as done by *) (* Conjunction.intr_balanced and fold Thm.weaken below, is quadratic for *) (* terms sorted in descending order, while only linear for terms *) (* sorted in ascending order *) val sorted_clauses = sort Thm.fast_term_ord nontrivial_clauses val _ = cond_tracing ctxt (fn () => "Sorted non-trivial clauses:\n" ^ cat_lines (map (Syntax.string_of_term ctxt o Thm.term_of) sorted_clauses)) (* translate clauses from HOL terms to Prop_Logic.prop_formula *) val (fms, atom_table) = fold_map (Prop_Logic.prop_formula_of_term o HOLogic.dest_Trueprop o Thm.term_of) sorted_clauses Termtab.empty val _ = cond_tracing ctxt (fn () => "Invoking SAT solver on clauses:\n" ^ cat_lines (map string_of_prop_formula fms)) val fm = Prop_Logic.all fms fun make_quick_and_dirty_thm () = let val _ = cond_tracing ctxt (fn () => "quick_and_dirty is set: proof reconstruction skipped, using oracle instead") val False_thm = Skip_Proof.make_thm_cterm \<^cprop>\False\ in (* 'fold Thm.weaken (rev sorted_clauses)' is linear, while 'fold *) (* Thm.weaken sorted_clauses' would be quadratic, since we sorted *) (* clauses in ascending order (which is linear for *) (* 'Conjunction.intr_balanced', used below) *) fold Thm.weaken (rev sorted_clauses) False_thm end in case let val the_solver = Config.get ctxt solver val _ = cond_tracing ctxt (fn () => "Invoking solver " ^ the_solver) in SAT_Solver.invoke_solver the_solver fm end of SAT_Solver.UNSATISFIABLE (SOME (clause_table, empty_id)) => (cond_tracing ctxt (fn () => "Proof trace from SAT solver:\n" ^ "clauses: " ^ ML_Syntax.print_list (ML_Syntax.print_pair ML_Syntax.print_int (ML_Syntax.print_list ML_Syntax.print_int)) (Inttab.dest clause_table) ^ "\n" ^ "empty clause: " ^ string_of_int empty_id); if Config.get ctxt quick_and_dirty then make_quick_and_dirty_thm () else let (* optimization: convert the given clauses to "[c_1 && ... && c_n] |- c_i"; *) (* this avoids accumulation of hypotheses during resolution *) (* [c_1, ..., c_n] |- c_1 && ... && c_n *) val clauses_thm = Conjunction.intr_balanced (map Thm.assume sorted_clauses) (* [c_1 && ... && c_n] |- c_1 && ... && c_n *) val cnf_cterm = Thm.cprop_of clauses_thm val cnf_thm = Thm.assume cnf_cterm (* [[c_1 && ... && c_n] |- c_1, ..., [c_1 && ... && c_n] |- c_n] *) val cnf_clauses = Conjunction.elim_balanced (length sorted_clauses) cnf_thm (* initialize the clause array with the given clauses *) val max_idx = fst (the (Inttab.max clause_table)) val clause_arr = Array.array (max_idx + 1, NO_CLAUSE) val _ = fold (fn thm => fn idx => (Array.update (clause_arr, idx, ORIG_CLAUSE thm); idx+1)) cnf_clauses 0 (* replay the proof to derive the empty clause *) (* [c_1 && ... && c_n] |- False *) val raw_thm = replay_proof ctxt atom_table clause_arr (clause_table, empty_id) in (* [c_1, ..., c_n] |- False *) Thm.implies_elim (Thm.implies_intr cnf_cterm raw_thm) clauses_thm end) | SAT_Solver.UNSATISFIABLE NONE => if Config.get ctxt quick_and_dirty then (if Context_Position.is_visible ctxt then warning "SAT solver claims the formula to be unsatisfiable, but did not provide a proof" else (); make_quick_and_dirty_thm ()) else raise THM ("SAT solver claims the formula to be unsatisfiable, but did not provide a proof", 0, []) | SAT_Solver.SATISFIABLE assignment => let val msg = "SAT solver found a countermodel:\n" ^ (commas o map (fn (term, idx) => Syntax.string_of_term_global \<^theory> term ^ ": " ^ (case assignment idx of NONE => "arbitrary" | SOME true => "true" | SOME false => "false"))) (Termtab.dest atom_table) in raise THM (msg, 0, []) end | SAT_Solver.UNKNOWN => raise THM ("SAT solver failed to decide the formula", 0, []) end; (* ------------------------------------------------------------------------- *) (* Tactics *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) (* rawsat_tac: solves the i-th subgoal of the proof state; this subgoal *) (* should be of the form *) (* [| c1; c2; ...; ck |] ==> False *) (* where each cj is a non-empty clause (i.e. a disjunction of literals) *) (* or "True" *) (* ------------------------------------------------------------------------- *) fun rawsat_tac ctxt i = Subgoal.FOCUS (fn {context = ctxt', prems, ...} => resolve_tac ctxt' [rawsat_thm ctxt' (map Thm.cprop_of prems)] 1) ctxt i; (* ------------------------------------------------------------------------- *) (* pre_cnf_tac: converts the i-th subgoal *) (* [| A1 ; ... ; An |] ==> B *) (* to *) (* [| A1; ... ; An ; ~B |] ==> False *) (* (handling meta-logical connectives in B properly before negating), *) (* then replaces meta-logical connectives in the premises (i.e. "==>", *) (* "!!" and "==") by connectives of the HOL object-logic (i.e. by *) (* "-->", "!", and "="), then performs beta-eta-normalization on the *) (* subgoal *) (* ------------------------------------------------------------------------- *) fun pre_cnf_tac ctxt = resolve_tac ctxt @{thms ccontr} THEN' Object_Logic.atomize_prems_tac ctxt THEN' CONVERSION Drule.beta_eta_conversion; (* ------------------------------------------------------------------------- *) (* cnfsat_tac: checks if the empty clause "False" occurs among the premises; *) (* if not, eliminates conjunctions (i.e. each clause of the CNF formula *) (* becomes a separate premise), then applies 'rawsat_tac' to solve the *) (* subgoal *) (* ------------------------------------------------------------------------- *) fun cnfsat_tac ctxt i = (eresolve_tac ctxt [FalseE] i) ORELSE (REPEAT_DETERM (eresolve_tac ctxt [conjE] i) THEN rawsat_tac ctxt i); (* ------------------------------------------------------------------------- *) (* cnfxsat_tac: checks if the empty clause "False" occurs among the *) (* premises; if not, eliminates conjunctions (i.e. each clause of the *) (* CNF formula becomes a separate premise) and existential quantifiers, *) (* then applies 'rawsat_tac' to solve the subgoal *) (* ------------------------------------------------------------------------- *) fun cnfxsat_tac ctxt i = (eresolve_tac ctxt [FalseE] i) ORELSE (REPEAT_DETERM (eresolve_tac ctxt [conjE] i ORELSE eresolve_tac ctxt [exE] i) THEN rawsat_tac ctxt i); (* ------------------------------------------------------------------------- *) (* sat_tac: tactic for calling an external SAT solver, taking as input an *) (* arbitrary formula. The input is translated to CNF, possibly causing *) (* an exponential blowup. *) (* ------------------------------------------------------------------------- *) fun sat_tac ctxt i = pre_cnf_tac ctxt i THEN CNF.cnf_rewrite_tac ctxt i THEN cnfsat_tac ctxt i; (* ------------------------------------------------------------------------- *) (* satx_tac: tactic for calling an external SAT solver, taking as input an *) (* arbitrary formula. The input is translated to CNF, possibly *) (* introducing new literals. *) (* ------------------------------------------------------------------------- *) fun satx_tac ctxt i = pre_cnf_tac ctxt i THEN CNF.cnfx_rewrite_tac ctxt i THEN cnfxsat_tac ctxt i; end; diff --git a/src/HOL/Tools/semiring_normalizer.ML b/src/HOL/Tools/semiring_normalizer.ML --- a/src/HOL/Tools/semiring_normalizer.ML +++ b/src/HOL/Tools/semiring_normalizer.ML @@ -1,884 +1,886 @@ (* Title: HOL/Tools/semiring_normalizer.ML Author: Amine Chaieb, TU Muenchen Normalization of expressions in semirings. *) signature SEMIRING_NORMALIZER = sig type entry val match: Proof.context -> cterm -> entry option val the_semiring: Proof.context -> thm -> cterm list * thm list val the_ring: Proof.context -> thm -> cterm list * thm list val the_field: Proof.context -> thm -> cterm list * thm list val the_idom: Proof.context -> thm -> thm list val the_ideal: Proof.context -> thm -> thm list val declare: thm -> {semiring: term list * thm list, ring: term list * thm list, field: term list * thm list, idom: thm list, ideal: thm list} -> local_theory -> local_theory val semiring_normalize_conv: Proof.context -> conv val semiring_normalize_ord_conv: Proof.context -> cterm ord -> conv val semiring_normalize_wrapper: Proof.context -> entry -> conv val semiring_normalize_ord_wrapper: Proof.context -> entry -> cterm ord -> conv val semiring_normalizers_conv: cterm list -> cterm list * thm list -> cterm list * thm list -> cterm list * thm list -> (cterm -> bool) * conv * conv * conv -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} val semiring_normalizers_ord_wrapper: Proof.context -> entry -> cterm ord -> {add: Proof.context -> conv, mul: Proof.context -> conv, neg: Proof.context -> conv, main: Proof.context -> conv, pow: Proof.context -> conv, sub: Proof.context -> conv} end structure Semiring_Normalizer: SEMIRING_NORMALIZER = struct (** data **) type entry = {vars: cterm list, semiring: cterm list * thm list, ring: cterm list * thm list, field: cterm list * thm list, idom: thm list, ideal: thm list} * {is_const: cterm -> bool, dest_const: cterm -> Rat.rat, mk_const: ctyp -> Rat.rat -> cterm, conv: Proof.context -> cterm -> thm}; structure Data = Generic_Data ( type T = (thm * entry) list; val empty = []; val extend = I; fun merge data = AList.merge Thm.eq_thm (K true) data; ); fun the_rules ctxt = fst o the o AList.lookup Thm.eq_thm (Data.get (Context.Proof ctxt)) val the_semiring = #semiring oo the_rules val the_ring = #ring oo the_rules val the_field = #field oo the_rules val the_idom = #idom oo the_rules val the_ideal = #ideal oo the_rules fun match ctxt tm = let fun match_inst ({vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal}, fns) pat = let fun h instT = let - val substT = Thm.instantiate (instT, []); + val substT = Thm.instantiate (instT, Vars.empty); val substT_cterm = Drule.cterm_rule substT; val vars' = map substT_cterm vars; val semiring' = (map substT_cterm sr_ops, map substT sr_rules); val ring' = (map substT_cterm r_ops, map substT r_rules); val field' = (map substT_cterm f_ops, map substT f_rules); val idom' = map substT idom; val ideal' = map substT ideal; val result = ({vars = vars', semiring = semiring', ring = ring', field = field', idom = idom', ideal = ideal'}, fns); in SOME result end in (case try Thm.match (pat, tm) of NONE => NONE | SOME (instT, _) => h instT) end; fun match_struct (_, entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) = get_first (match_inst entry) (sr_ops @ r_ops @ f_ops); in get_first match_struct (Data.get (Context.Proof ctxt)) end; (* extra-logical functions *) val semiring_norm_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms semiring_norm}); val semiring_funs = {is_const = can HOLogic.dest_number o Thm.term_of, dest_const = (fn ct => Rat.of_int (snd (HOLogic.dest_number (Thm.term_of ct) handle TERM _ => error "ring_dest_const"))), mk_const = (fn cT => fn x => Numeral.mk_cnumber cT (case Rat.dest x of (i, 1) => i | _ => error "int_of_rat: bad int")), conv = (fn ctxt => Simplifier.rewrite (put_simpset semiring_norm_ss ctxt) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms numeral_One}))}; val divide_const = Thm.cterm_of \<^context> (Logic.varify_global \<^term>\(/)\); val [divide_tvar] = Term.add_tvars (Thm.term_of divide_const) []; val field_funs = let fun numeral_is_const ct = case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b => can HOLogic.dest_number a andalso can HOLogic.dest_number b | Const (\<^const_name>\Fields.inverse\,_)$t => can HOLogic.dest_number t | t => can HOLogic.dest_number t fun dest_const ct = ((case Thm.term_of ct of Const (\<^const_name>\Rings.divide\,_) $ a $ b=> Rat.make (snd (HOLogic.dest_number a), snd (HOLogic.dest_number b)) | Const (\<^const_name>\Fields.inverse\,_)$t => Rat.inv (Rat.of_int (snd (HOLogic.dest_number t))) | t => Rat.of_int (snd (HOLogic.dest_number t))) handle TERM _ => error "ring_dest_const") fun mk_const cT x = let val (a, b) = Rat.dest x in if b = 1 then Numeral.mk_cnumber cT a else Thm.apply - (Thm.apply (Thm.instantiate_cterm ([(divide_tvar, cT)], []) divide_const) - (Numeral.mk_cnumber cT a)) + (Thm.apply + (Thm.instantiate_cterm (TVars.make [(divide_tvar, cT)], Vars.empty) divide_const) + (Numeral.mk_cnumber cT a)) (Numeral.mk_cnumber cT b) end in {is_const = numeral_is_const, dest_const = dest_const, mk_const = mk_const, conv = Numeral_Simprocs.field_comp_conv} end; (* logical content *) val semiringN = "semiring"; val ringN = "ring"; val fieldN = "field"; val idomN = "idom"; fun declare raw_key {semiring = raw_semiring0, ring = raw_ring0, field = raw_field0, idom = raw_idom, ideal = raw_ideal} lthy = let val ctxt' = fold Proof_Context.augment (fst raw_semiring0 @ fst raw_ring0 @ fst raw_field0) lthy; val prepare_ops = apfst (Variable.export_terms ctxt' lthy #> map (Thm.cterm_of lthy)); val raw_semiring = prepare_ops raw_semiring0; val raw_ring = prepare_ops raw_ring0; val raw_field = prepare_ops raw_field0; in lthy |> Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => fn context => let val ctxt = Context.proof_of context; val key = Morphism.thm phi raw_key; fun transform_ops_rules (ops, rules) = (map (Morphism.cterm phi) ops, Morphism.fact phi rules); val (sr_ops, sr_rules) = transform_ops_rules raw_semiring; val (r_ops, r_rules) = transform_ops_rules raw_ring; val (f_ops, f_rules) = transform_ops_rules raw_field; val idom = Morphism.fact phi raw_idom; val ideal = Morphism.fact phi raw_ideal; fun check kind name xs n = null xs orelse length xs = n orelse error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name); val check_ops = check "operations"; val check_rules = check "rules"; val _ = check_ops semiringN sr_ops 5 andalso check_rules semiringN sr_rules 36 andalso check_ops ringN r_ops 2 andalso check_rules ringN r_rules 2 andalso check_ops fieldN f_ops 2 andalso check_rules fieldN f_rules 2 andalso check_rules idomN idom 2; val mk_meta = Local_Defs.meta_rewrite_rule ctxt; val sr_rules' = map mk_meta sr_rules; val r_rules' = map mk_meta r_rules; val f_rules' = map mk_meta f_rules; fun rule i = nth sr_rules' (i - 1); val (cx, cy) = Thm.dest_binop (hd sr_ops); val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg; val ((clx, crx), (cly, cry)) = rule 13 |> Thm.rhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val ((ca, cb), (cc, cd)) = rule 20 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_binop; val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg; val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> apply2 Thm.dest_arg; val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry]; val semiring = (sr_ops, sr_rules'); val ring = (r_ops, r_rules'); val field = (f_ops, f_rules'); val ideal' = map (Thm.symmetric o mk_meta) ideal in context |> Data.map (AList.update Thm.eq_thm (key, ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal'}, (if null f_ops then semiring_funs else field_funs)))) end) end; (** auxiliary **) fun is_comb ct = (case Thm.term_of ct of _ $ _ => true | _ => false); val concl = Thm.cprop_of #> Thm.dest_arg; fun is_binop ct ct' = (case Thm.term_of ct' of c $ _ $ _ => Thm.term_of ct aconv c | _ => false); fun dest_binop ct ct' = if is_binop ct ct' then Thm.dest_binop ct' else raise CTERM ("dest_binop: bad binop", [ct, ct']) -fun inst_thm inst = Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) inst); +fun inst_thm inst = + Thm.instantiate (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) inst)); val dest_number = Thm.term_of #> HOLogic.dest_number #> snd; val is_number = can dest_number; fun numeral01_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One}]); fun zero1_numeral_conv ctxt = Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [@{thm numeral_One} RS sym]); fun zerone_conv ctxt cv = zero1_numeral_conv ctxt then_conv cv then_conv numeral01_conv ctxt; val nat_add_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps @{thms arith_simps} @ @{thms diff_nat_numeral} @ @{thms rel_simps} @ @{thms if_False if_True Nat.add_0 add_Suc add_numeral_left Suc_eq_plus1} @ map (fn th => th RS sym) @{thms numerals}); fun nat_add_conv ctxt = zerone_conv ctxt (Simplifier.rewrite (put_simpset nat_add_ss ctxt)); val zeron_tm = \<^cterm>\0::nat\; val onen_tm = \<^cterm>\1::nat\; val true_tm = \<^cterm>\True\; (** normalizing conversions **) (* core conversion *) fun semiring_normalizers_conv vars (sr_ops, sr_rules) (r_ops, r_rules) (f_ops, f_rules) (is_semiring_constant, semiring_add_conv, semiring_mul_conv, semiring_pow_conv) = let val [pthm_02, pthm_03, pthm_04, pthm_05, pthm_07, pthm_08, pthm_09, pthm_10, pthm_11, pthm_12, pthm_13, pthm_14, pthm_15, pthm_16, pthm_17, pthm_18, pthm_19, pthm_21, pthm_22, pthm_23, pthm_24, pthm_25, pthm_26, pthm_27, pthm_28, pthm_29, pthm_30, pthm_31, pthm_32, pthm_33, pthm_34, pthm_35, pthm_36, pthm_37, pthm_38, _] = sr_rules; val [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry] = vars; val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; val [add_tm, mul_tm, pow_tm] = map (Thm.dest_fun o Thm.dest_fun) [add_pat, mul_pat, pow_pat]; val dest_add = dest_binop add_tm val dest_mul = dest_binop mul_tm fun dest_pow tm = let val (l,r) = dest_binop pow_tm tm in if is_number r then (l,r) else raise CTERM ("dest_pow",[tm]) end; val is_add = is_binop add_tm val is_mul = is_binop mul_tm val (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, cx', cy') = (case (r_ops, r_rules) of ([sub_pat, neg_pat], [neg_mul, sub_add]) => let val sub_tm = Thm.dest_fun (Thm.dest_fun sub_pat) val neg_tm = Thm.dest_fun neg_pat val dest_sub = dest_binop sub_tm in (neg_mul, sub_add, sub_tm, neg_tm, dest_sub, neg_mul |> concl |> Thm.dest_arg, sub_add |> concl |> Thm.dest_arg |> Thm.dest_arg) end | _ => (TrueI, TrueI, true_tm, true_tm, (fn t => (t,t)), true_tm, true_tm)); val (divide_inverse, divide_tm, inverse_tm) = (case (f_ops, f_rules) of ([divide_pat, inverse_pat], [div_inv, _]) => let val div_tm = funpow 2 Thm.dest_fun divide_pat val inv_tm = Thm.dest_fun inverse_pat in (div_inv, div_tm, inv_tm) end | _ => (TrueI, true_tm, true_tm)); in fn variable_ord => let (* Conversion for "x^n * x^m", with either x^n = x and/or x^m = x possible. *) (* Also deals with "const * const", but both terms must involve powers of *) (* the same variable, or both be constants, or behaviour may be incorrect. *) fun powvar_mul_conv ctxt tm = let val (l,r) = dest_mul tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_mul_conv tm else ((let val (lx,ln) = dest_pow l in ((let val (_, rn) = dest_pow r val th1 = inst_thm [(cx,lx),(cp,ln),(cq,rn)] pthm_29 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => (let val th1 = inst_thm [(cx,lx),(cq,ln)] pthm_31 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end)) end) handle CTERM _ => ((let val (rx,rn) = dest_pow r val th1 = inst_thm [(cx,rx),(cq,rn)] pthm_30 val (tm1,tm2) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule tm1 (nat_add_conv ctxt tm2)) end) handle CTERM _ => inst_thm [(cx,l)] pthm_32 )) end; (* Remove "1 * m" from a monomial, and just leave m. *) fun monomial_deone th = (let val (l,r) = dest_mul(concl th) in if l aconvc one_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_13) else th end) handle CTERM _ => th; (* Conversion for "(monomial)^n", where n is a numeral. *) fun monomial_pow_conv ctxt = let fun monomial_pow tm bod ntm = if not(is_comb bod) then Thm.reflexive tm else if is_semiring_constant bod then semiring_pow_conv tm else let val (lopr,r) = Thm.dest_comb bod in if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = inst_thm [(cx,l),(cp,r),(cq,ntm)] pthm_34 val (l,r) = Thm.dest_comb(concl th1) in Thm.transitive th1 (Drule.arg_cong_rule l (nat_add_conv ctxt r)) end else if opr aconvc mul_tm then let val th1 = inst_thm [(cx,l),(cy,r),(cq,ntm)] pthm_33 val (xy,z) = Thm.dest_comb(concl th1) val (x,y) = Thm.dest_comb xy val thl = monomial_pow y l ntm val thr = monomial_pow z r ntm in Thm.transitive th1 (Thm.combination (Drule.arg_cong_rule x thl) thr) end else Thm.reflexive tm end end in fn tm => let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if not (opr aconvc pow_tm) orelse not(is_number r) then raise CTERM ("monomial_pow_conv", [tm]) else if r aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if r aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else monomial_deone(monomial_pow tm l r) end end; (* Multiplication of canonical monomials. *) fun monomial_mul_conv ctxt = let fun powvar tm = if is_semiring_constant tm then one_tm else ((let val (lopr,r) = Thm.dest_comb tm val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then l else raise CTERM ("monomial_mul_conv",[tm]) end) handle CTERM _ => tm) (* FIXME !? *) fun vorder x y = if x aconvc y then 0 else if x aconvc one_tm then ~1 else if y aconvc one_tm then 1 else if is_less (variable_ord (x, y)) then ~1 else 1 fun monomial_mul tm l r = ((let val (lx,ly) = dest_mul l val vl = powvar lx in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] pthm_15 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 val th3 = Thm.transitive th1 th2 val (tm5,tm6) = Thm.dest_comb(concl th3) val (tm7,tm8) = Thm.dest_comb tm6 val th4 = monomial_mul tm6 (Thm.dest_arg tm7) tm8 in Thm.transitive th3 (Drule.arg_cong_rule tm5 th4) end else let val th0 = if ord < 0 then pthm_16 else pthm_17 val th1 = inst_thm [(clx,lx),(cly,ly),(crx,rx),(cry,ry)] th0 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_18 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2 in Thm.transitive th1 th2 end else if ord < 0 then let val th1 = inst_thm [(clx,lx),(cly,ly),(crx,r)] pthm_19 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else inst_thm [(ca,l),(cb,r)] pthm_09 end)) end) handle CTERM _ => (let val vl = powvar l in ((let val (rx,ry) = dest_mul r val vr = powvar rx val ord = vorder vl vr in if ord = 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_21 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 in Thm.transitive th1 (Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (powvar_mul_conv ctxt tm4)) tm2) end else if ord > 0 then let val th1 = inst_thm [(clx,l),(crx,rx),(cry,ry)] pthm_22 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm2 in Thm.transitive th1 (Drule.arg_cong_rule tm1 (monomial_mul tm2 (Thm.dest_arg tm3) tm4)) end else Thm.reflexive tm end) handle CTERM _ => (let val vr = powvar r val ord = vorder vl vr in if ord = 0 then powvar_mul_conv ctxt tm else if ord > 0 then inst_thm [(ca,l),(cb,r)] pthm_09 else Thm.reflexive tm end)) end)) in fn tm => let val (l,r) = dest_mul tm in monomial_deone(monomial_mul tm l r) end end; (* Multiplication by monomial of a polynomial. *) fun polynomial_monomial_mul_conv ctxt = let fun pmm_conv tm = let val (l,r) = dest_mul tm in ((let val (y,z) = dest_add r val th1 = inst_thm [(cx,l),(cy,y),(cz,z)] pthm_37 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Thm.combination (Drule.arg_cong_rule tm3 (monomial_mul_conv ctxt tm4)) (pmm_conv tm2) in Thm.transitive th1 th2 end) handle CTERM _ => monomial_mul_conv ctxt tm) end in pmm_conv end; (* Addition of two monomials identical except for constant multiples. *) fun monomial_add_conv tm = let val (l,r) = dest_add tm in if is_semiring_constant l andalso is_semiring_constant r then semiring_add_conv tm else let val th1 = if is_mul l andalso is_semiring_constant(Thm.dest_arg1 l) then if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(ca,Thm.dest_arg1 l),(cm,Thm.dest_arg r), (cb,Thm.dest_arg1 r)] pthm_02 else inst_thm [(ca,Thm.dest_arg1 l),(cm,r)] pthm_03 else if is_mul r andalso is_semiring_constant(Thm.dest_arg1 r) then inst_thm [(cm,l),(ca,Thm.dest_arg1 r)] pthm_04 else inst_thm [(cm,r)] pthm_05 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (semiring_add_conv tm4) val th3 = Thm.transitive th1 (Drule.fun_cong_rule th2 tm2) val tm5 = concl th3 in if (Thm.dest_arg1 tm5) aconvc zero_tm then Thm.transitive th3 (inst_thm [(ca,Thm.dest_arg tm5)] pthm_11) else monomial_deone th3 end end; (* Ordering on monomials. *) fun striplist dest = let fun strip x acc = ((let val (l,r) = dest x in strip l (strip r acc) end) handle CTERM _ => x::acc) (* FIXME !? *) in fn x => strip x [] end; fun powervars tm = let val ptms = striplist dest_mul tm in if is_semiring_constant (hd ptms) then tl ptms else ptms end; val num_0 = 0; val num_1 = 1; fun dest_varpow tm = ((let val (x,n) = dest_pow tm in (x,dest_number n) end) handle CTERM _ => (tm,(if is_semiring_constant tm then num_0 else num_1))); val morder = let fun lexorder ls = case ls of ([],[]) => 0 | (_ ,[]) => ~1 | ([], _) => 1 | (((x1,n1)::vs1),((x2,n2)::vs2)) => (case variable_ord (x1, x2) of LESS => 1 | GREATER => ~1 | EQUAL => if n1 < n2 then ~1 else if n2 < n1 then 1 else lexorder (vs1, vs2)) in fn tm1 => fn tm2 => let val vdegs1 = map dest_varpow (powervars tm1) val vdegs2 = map dest_varpow (powervars tm2) val deg1 = fold (Integer.add o snd) vdegs1 num_0 val deg2 = fold (Integer.add o snd) vdegs2 num_0 in if deg1 < deg2 then ~1 else if deg1 > deg2 then 1 else lexorder (vdegs1, vdegs2) end end; (* Addition of two polynomials. *) fun polynomial_add_conv ctxt = let fun dezero_rule th = let val tm = concl th in if not(is_add tm) then th else let val (lopr,r) = Thm.dest_comb tm val l = Thm.dest_arg lopr in if l aconvc zero_tm then Thm.transitive th (inst_thm [(ca,r)] pthm_07) else if r aconvc zero_tm then Thm.transitive th (inst_thm [(ca,l)] pthm_08) else th end end fun padd tm = let val (l,r) = dest_add tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_07 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_08 else if is_add l then let val (a,b) = dest_add l in if is_add r then let val (c,d) = dest_add r val ord = morder a c in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,c),(cd,d)] pthm_23 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (monomial_add_conv tm4) in dezero_rule (Thm.transitive th1 (Thm.combination th2 (padd tm2))) end else (* ord <> 0*) let val th1 = if ord > 0 then inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 else inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else (* not (is_add r)*) let val ord = morder a r in if ord = 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_26 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else (* ord <> 0*) if ord > 0 then let val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_24 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end else (* not (is_add l)*) if is_add r then let val (c,d) = dest_add r val ord = morder l c in if ord = 0 then let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_28 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.fun_cong_rule (Drule.arg_cong_rule tm3 (monomial_add_conv tm4)) tm2 in dezero_rule (Thm.transitive th1 th2) end else if ord > 0 then Thm.reflexive tm else let val th1 = inst_thm [(ca,l),(cc,c),(cd,d)] pthm_25 val (tm1,tm2) = Thm.dest_comb(concl th1) in dezero_rule (Thm.transitive th1 (Drule.arg_cong_rule tm1 (padd tm2))) end end else let val ord = morder l r in if ord = 0 then monomial_add_conv tm else if ord > 0 then dezero_rule(Thm.reflexive tm) else dezero_rule (inst_thm [(ca,l),(cc,r)] pthm_27) end end in padd end; (* Multiplication of two polynomials. *) fun polynomial_mul_conv ctxt = let fun pmul tm = let val (l,r) = dest_mul tm in if not(is_add l) then polynomial_monomial_mul_conv ctxt tm else if not(is_add r) then let val th1 = inst_thm [(ca,l),(cb,r)] pthm_09 in Thm.transitive th1 (polynomial_monomial_mul_conv ctxt (concl th1)) end else let val (a,b) = dest_add l val th1 = inst_thm [(ca,a),(cb,b),(cc,r)] pthm_10 val (tm1,tm2) = Thm.dest_comb(concl th1) val (tm3,tm4) = Thm.dest_comb tm1 val th2 = Drule.arg_cong_rule tm3 (polynomial_monomial_mul_conv ctxt tm4) val th3 = Thm.transitive th1 (Thm.combination th2 (pmul tm2)) in Thm.transitive th3 (polynomial_add_conv ctxt (concl th3)) end end in fn tm => let val (l,r) = dest_mul tm in if l aconvc zero_tm then inst_thm [(ca,r)] pthm_11 else if r aconvc zero_tm then inst_thm [(ca,l)] pthm_12 else if l aconvc one_tm then inst_thm [(ca,r)] pthm_13 else if r aconvc one_tm then inst_thm [(ca,l)] pthm_14 else pmul tm end end; (* Power of polynomial (optimized for the monomial and trivial cases). *) fun num_conv ctxt n = nat_add_conv ctxt (Thm.apply \<^cterm>\Suc\ (Numeral.mk_cnumber \<^ctyp>\nat\ (dest_number n - 1))) |> Thm.symmetric; fun polynomial_pow_conv ctxt = let fun ppow tm = let val (l,n) = dest_pow tm in if n aconvc zeron_tm then inst_thm [(cx,l)] pthm_35 else if n aconvc onen_tm then inst_thm [(cx,l)] pthm_36 else let val th1 = num_conv ctxt n val th2 = inst_thm [(cx,l),(cq,Thm.dest_arg (concl th1))] pthm_38 val (tm1,tm2) = Thm.dest_comb(concl th2) val th3 = Thm.transitive th2 (Drule.arg_cong_rule tm1 (ppow tm2)) val th4 = Thm.transitive (Drule.arg_cong_rule (Thm.dest_fun tm) th1) th3 in Thm.transitive th4 (polynomial_mul_conv ctxt (concl th4)) end end in fn tm => if is_add(Thm.dest_arg1 tm) then ppow tm else monomial_pow_conv ctxt tm end; (* Negation. *) fun polynomial_neg_conv ctxt tm = let val (l,r) = Thm.dest_comb tm in if not (l aconvc neg_tm) then raise CTERM ("polynomial_neg_conv",[tm]) else let val th1 = inst_thm [(cx', r)] neg_mul val th2 = Thm.transitive th1 (Conv.arg1_conv semiring_mul_conv (concl th1)) in Thm.transitive th2 (polynomial_monomial_mul_conv ctxt (concl th2)) end end; (* Subtraction. *) fun polynomial_sub_conv ctxt tm = let val (l,r) = dest_sub tm val th1 = inst_thm [(cx', l), (cy', r)] sub_add val (tm1,tm2) = Thm.dest_comb(concl th1) val th2 = Drule.arg_cong_rule tm1 (polynomial_neg_conv ctxt tm2) in Thm.transitive th1 (Thm.transitive th2 (polynomial_add_conv ctxt (concl th2))) end; (* Conversion from HOL term. *) fun polynomial_conv ctxt tm = if is_semiring_constant tm then semiring_add_conv tm else if not(is_comb tm) then Thm.reflexive tm else let val (lopr,r) = Thm.dest_comb tm in if lopr aconvc neg_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (polynomial_neg_conv ctxt (concl th1)) end else if lopr aconvc inverse_tm then let val th1 = Drule.arg_cong_rule lopr (polynomial_conv ctxt r) in Thm.transitive th1 (semiring_mul_conv (concl th1)) end else if not(is_comb lopr) then Thm.reflexive tm else let val (opr,l) = Thm.dest_comb lopr in if opr aconvc pow_tm andalso is_number r then let val th1 = Drule.fun_cong_rule (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) r in Thm.transitive th1 (polynomial_pow_conv ctxt (concl th1)) end else if opr aconvc divide_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val th2 = (Conv.rewr_conv divide_inverse then_conv polynomial_mul_conv ctxt) (Thm.rhs_of th1) in Thm.transitive th1 th2 end else if opr aconvc add_tm orelse opr aconvc mul_tm orelse opr aconvc sub_tm then let val th1 = Thm.combination (Drule.arg_cong_rule opr (polynomial_conv ctxt l)) (polynomial_conv ctxt r) val f = if opr aconvc add_tm then polynomial_add_conv ctxt else if opr aconvc mul_tm then polynomial_mul_conv ctxt else polynomial_sub_conv ctxt in Thm.transitive th1 (f (concl th1)) end else Thm.reflexive tm end end; in {main = polynomial_conv, add = polynomial_add_conv, mul = polynomial_mul_conv, pow = polynomial_pow_conv, neg = polynomial_neg_conv, sub = polynomial_sub_conv} end end; val nat_exp_ss = simpset_of (put_simpset HOL_basic_ss \<^context> addsimps (@{thms eval_nat_numeral} @ @{thms diff_nat_numeral} @ @{thms arith_simps} @ @{thms rel_simps}) addsimps [@{thm Let_def}, @{thm if_False}, @{thm if_True}, @{thm Nat.add_0}, @{thm add_Suc}]); (* various normalizing conversions *) fun semiring_normalizers_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = let val pow_conv = Conv.arg_conv (Simplifier.rewrite (put_simpset nat_exp_ss ctxt)) then_conv Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [nth (snd semiring) 31, nth (snd semiring) 34]) then_conv conv ctxt val dat = (is_const, conv ctxt, conv ctxt, pow_conv) in semiring_normalizers_conv vars semiring ring field dat term_ord end; fun semiring_normalize_ord_wrapper ctxt ({vars, semiring, ring, field, idom, ideal}, {conv, dest_const, mk_const, is_const}) term_ord = #main (semiring_normalizers_ord_wrapper ctxt ({vars = vars, semiring = semiring, ring = ring, field = field, idom = idom, ideal = ideal}, {conv = conv, dest_const = dest_const, mk_const = mk_const, is_const = is_const}) term_ord) ctxt; fun semiring_normalize_wrapper ctxt data = semiring_normalize_ord_wrapper ctxt data Thm.term_ord; fun semiring_normalize_ord_conv ctxt ord tm = (case match ctxt tm of NONE => Thm.reflexive tm | SOME res => semiring_normalize_ord_wrapper ctxt res ord tm); fun semiring_normalize_conv ctxt = semiring_normalize_ord_conv ctxt Thm.term_ord; end; diff --git a/src/HOL/Tools/split_rule.ML b/src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML +++ b/src/HOL/Tools/split_rule.ML @@ -1,123 +1,124 @@ (* Title: HOL/Tools/split_rule.ML Author: Stefan Berghofer, David von Oheimb, and Markus Wenzel, TU Muenchen Some tools for managing tupled arguments and abstractions in rules. *) signature SPLIT_RULE = sig val split_rule_var: Proof.context -> term -> thm -> thm val split_rule: Proof.context -> thm -> thm val complete_split_rule: Proof.context -> thm -> thm end; structure Split_Rule: SPLIT_RULE = struct (** split rules **) fun internal_case_prod_const (Ta, Tb, Tc) = Const (\<^const_name>\Product_Type.internal_case_prod\, [[Ta, Tb] ---> Tc, HOLogic.mk_prodT (Ta, Tb)] ---> Tc); fun eval_internal_split ctxt = hol_simplify ctxt @{thms internal_case_prod_def} o hol_simplify ctxt @{thms internal_case_prod_conv}; fun remove_internal_split ctxt = eval_internal_split ctxt o split_all ctxt; (*In ap_split S T u, term u expects separate arguments for the factors of S, with result type T. The call creates a new term expecting one argument of type S.*) fun ap_split (Type (\<^type_name>\Product_Type.prod\, [T1, T2])) T3 u = internal_case_prod_const (T1, T2, T3) $ Abs ("v", T1, ap_split T2 T3 ((ap_split T1 (HOLogic.flatten_tupleT T2 ---> T3) (incr_boundvars 1 u)) $ Bound 0)) | ap_split _ T3 u = u; (*Curries any Var of function type in the rule*) fun split_rule_var' ctxt (t as Var (v, Type ("fun", [T1, T2]))) rl = let val T' = HOLogic.flatten_tupleT T1 ---> T2; val newt = ap_split T1 T2 (Var (v, T')); - in Thm.instantiate ([], [(dest_Var t, Thm.cterm_of ctxt newt)]) rl end + in Thm.instantiate (TVars.empty, Vars.make [(dest_Var t, Thm.cterm_of ctxt newt)]) rl end | split_rule_var' _ _ rl = rl; (* complete splitting of partially split rules *) fun ap_split' (T::Ts) U u = Abs ("v", T, ap_split' Ts U (ap_split T (maps HOLogic.flatten_tupleT Ts ---> U) (incr_boundvars 1 u) $ Bound 0)) | ap_split' _ _ u = u; fun complete_split_rule_var ctxt (t as Var (v, T), ts) (rl, vs) = let val (Us', U') = strip_type T; val Us = take (length ts) Us'; val U = drop (length ts) Us' ---> U'; val T' = maps HOLogic.flatten_tupleT Us ---> U; fun mk_tuple (v as Var ((a, _), T)) (xs, insts) = let val Ts = HOLogic.flatten_tupleT T; val ys = Name.variant_list xs (replicate (length Ts) a); in (xs @ ys, (dest_Var v, Thm.cterm_of ctxt (HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths T) T (map (Var o apfst (rpair 0)) (ys ~~ Ts)))) :: insts) end | mk_tuple _ x = x; val newt = ap_split' Us U (Var (v, T')); val (vs', insts) = fold mk_tuple ts (vs, []); in - (Drule.instantiate_normalize ([], ((v, T), Thm.cterm_of ctxt newt) :: insts) rl, vs') + (Drule.instantiate_normalize + (TVars.empty, Vars.make (((v, T), Thm.cterm_of ctxt newt) :: insts)) rl, vs') end | complete_split_rule_var _ _ x = x; fun collect_vars (Abs (_, _, t)) = collect_vars t | collect_vars t = (case strip_comb t of (v as Var _, ts) => cons (v, ts) | (_, ts) => fold collect_vars ts); fun split_rule_var ctxt = (Drule.export_without_context o remove_internal_split ctxt) oo split_rule_var' ctxt; (*curries ALL function variables occurring in a rule's conclusion*) fun split_rule ctxt rl = fold_rev (split_rule_var' ctxt) (Misc_Legacy.term_vars (Thm.concl_of rl)) rl |> remove_internal_split ctxt |> Drule.export_without_context; (*curries ALL function variables*) fun complete_split_rule ctxt rl = let val prop = Thm.prop_of rl; val xs = Term.fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I) prop []; val vars = collect_vars prop []; in fst (fold_rev (complete_split_rule_var ctxt) vars (rl, xs)) |> remove_internal_split ctxt |> Drule.export_without_context |> Rule_Cases.save rl end; (* attribute syntax *) val _ = Theory.setup (Attrib.setup \<^binding>\split_format\ (Scan.lift (Args.parens (Args.$$$ "complete") >> K (Thm.rule_attribute [] (complete_split_rule o Context.proof_of)))) "split pair-typed subterms in premises, or function arguments" #> Attrib.setup \<^binding>\split_rule\ (Scan.succeed (Thm.rule_attribute [] (split_rule o Context.proof_of))) "curries ALL function variables occurring in a rule's conclusion"); end; diff --git a/src/HOL/Types_To_Sets/unoverload_type.ML b/src/HOL/Types_To_Sets/unoverload_type.ML --- a/src/HOL/Types_To_Sets/unoverload_type.ML +++ b/src/HOL/Types_To_Sets/unoverload_type.ML @@ -1,62 +1,62 @@ (* Title: HOL/Types_To_Sets/unoverload_type.ML Author: Fabian Immler, TU München Internalize sorts and unoverload parameters of a type variable. *) signature UNOVERLOAD_TYPE = sig val unoverload_type: Context.generic -> indexname list -> thm -> thm val unoverload_type_attr: indexname list -> attribute end; structure Unoverload_Type : UNOVERLOAD_TYPE = struct fun params_of_class thy class = try (Axclass.get_info thy #> #params) class |> these fun params_of_super_classes thy class = class::Sorts.super_classes (Sign.classes_of thy) class |> maps (params_of_class thy) fun params_of_sort thy sort = maps (params_of_super_classes thy) sort fun subst_TFree n' ty' ty = map_type_tfree (fn x as (n, _) => if n = n' then ty' else TFree x) ty fun unoverload_single_type context tvar thm = let val tvars = Term.add_tvars (Thm.prop_of thm) [] val thy = Context.theory_of context in case find_first (fn (y, _) => tvar = y) tvars of NONE => raise TYPE ("unoverload_type: no such type variable in theorem", [TVar (tvar,[])], []) | SOME (x as (_, sort)) => let val (_, thm') = Internalize_Sort.internalize_sort (Thm.global_ctyp_of thy (TVar x)) thm val prop' = fold (fn _ => Term.dest_comb #> #2) (replicate (Thm.nprems_of thm' - Thm.nprems_of thm) ()) (Thm.prop_of thm') val (tyenv, _) = Pattern.first_order_match thy ((prop', Thm.prop_of thm)) (Vartab.empty, Vartab.empty) val tyenv_list = tyenv |> Vartab.dest |> map (fn (x, (s, TVar (x', _))) => ((x, s), Thm.ctyp_of (Context.proof_of context) (TVar(x', s)))) - val thm'' = Drule.instantiate_normalize (tyenv_list, []) thm' + val thm'' = Drule.instantiate_normalize (TVars.make tyenv_list, Vars.empty) thm' val varify_const = apsnd (subst_TFree "'a" (TVar (tvar, @{sort type}))) #> Const #> Thm.global_cterm_of thy val consts = params_of_sort thy sort |> map varify_const in fold Unoverloading.unoverload consts thm'' |> Raw_Simplifier.norm_hhf (Context.proof_of context) end end fun unoverload_type context xs = fold (unoverload_single_type context) xs fun unoverload_type_attr xs = Thm.rule_attribute [] (fn context => unoverload_type context xs) val _ = Context.>> (Context.map_theory (Attrib.setup \<^binding>\unoverload_type\ (Scan.lift (Scan.repeat Args.var) >> unoverload_type_attr) "internalize and unoverload type class parameters")) end \ No newline at end of file diff --git a/src/Provers/Arith/fast_lin_arith.ML b/src/Provers/Arith/fast_lin_arith.ML --- a/src/Provers/Arith/fast_lin_arith.ML +++ b/src/Provers/Arith/fast_lin_arith.ML @@ -1,821 +1,821 @@ (* Title: Provers/Arith/fast_lin_arith.ML Author: Tobias Nipkow and Tjark Weber and Sascha Boehme A generic linear arithmetic package. Only take premises and conclusions into account that are already (negated) (in)equations. lin_arith_simproc tries to prove or disprove the term. *) (*** Data needed for setting up the linear arithmetic package ***) signature LIN_ARITH_LOGIC = sig val conjI : thm (* P ==> Q ==> P & Q *) val ccontr : thm (* (~ P ==> False) ==> P *) val notI : thm (* (P ==> False) ==> ~ P *) val not_lessD : thm (* ~(m < n) ==> n <= m *) val not_leD : thm (* ~(m <= n) ==> n < m *) val sym : thm (* x = y ==> y = x *) val trueI : thm (* True *) val mk_Eq : thm -> thm val atomize : thm -> thm list val mk_Trueprop : term -> term val neg_prop : term -> term val is_False : thm -> bool val is_nat : typ list * term -> bool val mk_nat_thm : theory -> term -> thm end; (* mk_Eq(~in) = `in == False' mk_Eq(in) = `in == True' where `in' is an (in)equality. neg_prop(t) = neg if t is wrapped up in Trueprop and neg is the (logically) negated version of t (again wrapped up in Trueprop), where the negation of a negative term is the term itself (no double negation!); raises TERM ("neg_prop", [t]) if t is not of the form 'Trueprop $ _' is_nat(parameter-types,t) = t:nat mk_nat_thm(t) = "0 <= t" *) signature LIN_ARITH_DATA = sig (*internal representation of linear (in-)equations:*) type decomp = (term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat * bool val decomp: Proof.context -> term -> decomp option val domain_is_nat: term -> bool (*abstraction for proof replay*) val abstract_arith: term -> (term * term) list * Proof.context -> term * ((term * term) list * Proof.context) val abstract: term -> (term * term) list * Proof.context -> term * ((term * term) list * Proof.context) (*preprocessing, performed on a representation of subgoals as list of premises:*) val pre_decomp: Proof.context -> typ list * term list -> (typ list * term list) list (*preprocessing, performed on the goal -- must do the same as 'pre_decomp':*) val pre_tac: Proof.context -> int -> tactic (*the limit on the number of ~= allowed; because each ~= is split into two cases, this can lead to an explosion*) val neq_limit: int Config.T val trace: bool Config.T end; (* decomp(`x Rel y') should yield (p,i,Rel,q,j,d) where Rel is one of "<", "~<", "<=", "~<=" and "=" and p (q, respectively) is the decomposition of the sum term x (y, respectively) into a list of summand * multiplicity pairs and a constant summand and d indicates if the domain is discrete. domain_is_nat(`x Rel y') t should yield true iff x is of type "nat". The relationship between pre_decomp and pre_tac is somewhat tricky. The internal representation of a subgoal and the corresponding theorem must be modified by pre_decomp (pre_tac, resp.) in a corresponding way. See the comment for split_items below. (This is even necessary for eta- and beta-equivalent modifications, as some of the lin. arith. code is not insensitive to them.) Simpset must reduce contradictory <= to False. It should also cancel common summands to keep <= reduced; otherwise <= can grow to massive proportions. *) signature FAST_LIN_ARITH = sig val prems_lin_arith_tac: Proof.context -> int -> tactic val lin_arith_tac: Proof.context -> int -> tactic val lin_arith_simproc: Proof.context -> cterm -> thm option val map_data: ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list, lessD: thm list, neqE: thm list, simpset: simpset, number_of: (Proof.context -> typ -> int -> cterm) option} -> {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list, lessD: thm list, neqE: thm list, simpset: simpset, number_of: (Proof.context -> typ -> int -> cterm) option}) -> Context.generic -> Context.generic val add_inj_thms: thm list -> Context.generic -> Context.generic val add_lessD: thm -> Context.generic -> Context.generic val add_simps: thm list -> Context.generic -> Context.generic val add_simprocs: simproc list -> Context.generic -> Context.generic val set_number_of: (Proof.context -> typ -> int -> cterm) -> Context.generic -> Context.generic end; functor Fast_Lin_Arith (structure LA_Logic: LIN_ARITH_LOGIC and LA_Data: LIN_ARITH_DATA): FAST_LIN_ARITH = struct (** theory data **) structure Data = Generic_Data ( type T = {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list, lessD: thm list, neqE: thm list, simpset: simpset, number_of: (Proof.context -> typ -> int -> cterm) option}; val empty : T = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [], lessD = [], neqE = [], simpset = empty_ss, number_of = NONE}; val extend = I; fun merge ({add_mono_thms = add_mono_thms1, mult_mono_thms = mult_mono_thms1, inj_thms = inj_thms1, lessD = lessD1, neqE = neqE1, simpset = simpset1, number_of = number_of1}, {add_mono_thms = add_mono_thms2, mult_mono_thms = mult_mono_thms2, inj_thms = inj_thms2, lessD = lessD2, neqE = neqE2, simpset = simpset2, number_of = number_of2}) : T = {add_mono_thms = Thm.merge_thms (add_mono_thms1, add_mono_thms2), mult_mono_thms = Thm.merge_thms (mult_mono_thms1, mult_mono_thms2), inj_thms = Thm.merge_thms (inj_thms1, inj_thms2), lessD = Thm.merge_thms (lessD1, lessD2), neqE = Thm.merge_thms (neqE1, neqE2), simpset = merge_ss (simpset1, simpset2), number_of = merge_options (number_of1, number_of2)}; ); val map_data = Data.map; val get_data = Data.get o Context.Proof; fun get_neqE ctxt = map (Thm.transfer' ctxt) (#neqE (get_data ctxt)); fun map_inj_thms f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} = {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = f inj_thms, lessD = lessD, neqE = neqE, simpset = simpset, number_of = number_of}; fun map_lessD f {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} = {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, lessD = f lessD, neqE = neqE, simpset = simpset, number_of = number_of}; fun map_simpset f context = map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, number_of} => {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, lessD = lessD, neqE = neqE, simpset = simpset_map (Context.proof_of context) f simpset, number_of = number_of}) context; fun add_inj_thms thms = map_data (map_inj_thms (append (map Thm.trim_context thms))); fun add_lessD thm = map_data (map_lessD (fn thms => thms @ [Thm.trim_context thm])); fun add_simps thms = map_simpset (fn ctxt => ctxt addsimps thms); fun add_simprocs procs = map_simpset (fn ctxt => ctxt addsimprocs procs); fun set_number_of f = map_data (fn {add_mono_thms, mult_mono_thms, inj_thms, lessD, neqE, simpset, ...} => {add_mono_thms = add_mono_thms, mult_mono_thms = mult_mono_thms, inj_thms = inj_thms, lessD = lessD, neqE = neqE, simpset = simpset, number_of = SOME f}); fun number_of ctxt = (case get_data ctxt of {number_of = SOME f, ...} => f ctxt | _ => fn _ => fn _ => raise CTERM ("number_of", [])); (*** A fast decision procedure ***) (*** Code ported from HOL Light ***) (* possible optimizations: use (var,coeff) rep or vector rep tp save space; treat non-negative atoms separately rather than adding 0 <= atom *) datatype lineq_type = Eq | Le | Lt; datatype injust = Asm of int | Nat of int (* index of atom *) | LessD of injust | NotLessD of injust | NotLeD of injust | NotLeDD of injust | Multiplied of int * injust | Added of injust * injust; datatype lineq = Lineq of int * lineq_type * int list * injust; (* ------------------------------------------------------------------------- *) (* Finding a (counter) example from the trace of a failed elimination *) (* ------------------------------------------------------------------------- *) (* Examples are represented as rational numbers, *) (* Dont blame John Harrison for this code - it is entirely mine. TN *) exception NoEx; (* Coding: (i,true,cs) means i <= cs and (i,false,cs) means i < cs. In general, true means the bound is included, false means it is excluded. Need to know if it is a lower or upper bound for unambiguous interpretation! *) (* ------------------------------------------------------------------------- *) (* End of counterexample finder. The actual decision procedure starts here. *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) (* Calculate new (in)equality type after addition. *) (* ------------------------------------------------------------------------- *) fun find_add_type(Eq,x) = x | find_add_type(x,Eq) = x | find_add_type(_,Lt) = Lt | find_add_type(Lt,_) = Lt | find_add_type(Le,Le) = Le; (* ------------------------------------------------------------------------- *) (* Multiply out an (in)equation. *) (* ------------------------------------------------------------------------- *) fun multiply_ineq n (i as Lineq(k,ty,l,just)) = if n = 1 then i else if n = 0 andalso ty = Lt then raise Fail "multiply_ineq" else if n < 0 andalso (ty=Le orelse ty=Lt) then raise Fail "multiply_ineq" else Lineq (n * k, ty, map (Integer.mult n) l, Multiplied (n, just)); (* ------------------------------------------------------------------------- *) (* Add together (in)equations. *) (* ------------------------------------------------------------------------- *) fun add_ineq (Lineq (k1,ty1,l1,just1)) (Lineq (k2,ty2,l2,just2)) = let val l = map2 Integer.add l1 l2 in Lineq(k1+k2,find_add_type(ty1,ty2),l,Added(just1,just2)) end; (* ------------------------------------------------------------------------- *) (* Elimination of variable between a single pair of (in)equations. *) (* If they're both inequalities, 1st coefficient must be +ve, 2nd -ve. *) (* ------------------------------------------------------------------------- *) fun elim_var v (i1 as Lineq(_,ty1,l1,_)) (i2 as Lineq(_,ty2,l2,_)) = let val c1 = nth l1 v and c2 = nth l2 v val m = Integer.lcm c1 c2 val m1 = m div (abs c1) and m2 = m div (abs c2) val (n1,n2) = if (c1 >= 0) = (c2 >= 0) then if ty1 = Eq then (~m1,m2) else if ty2 = Eq then (m1,~m2) else raise Fail "elim_var" else (m1,m2) val (p1,p2) = if ty1=Eq andalso ty2=Eq andalso (n1 = ~1 orelse n2 = ~1) then (~n1,~n2) else (n1,n2) in add_ineq (multiply_ineq p1 i1) (multiply_ineq p2 i2) end; (* ------------------------------------------------------------------------- *) (* The main refutation-finding code. *) (* ------------------------------------------------------------------------- *) fun is_trivial (Lineq(_,_,l,_)) = forall (fn i => i=0) l; fun is_contradictory (Lineq(k,ty,_,_)) = case ty of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0; fun calc_blowup l = let val (p,n) = List.partition (curry (op <) 0) (filter (curry (op <>) 0) l) in length p * length n end; (* ------------------------------------------------------------------------- *) (* Main elimination code: *) (* *) (* (1) Looks for immediate solutions (false assertions with no variables). *) (* *) (* (2) If there are any equations, picks a variable with the lowest absolute *) (* coefficient in any of them, and uses it to eliminate. *) (* *) (* (3) Otherwise, chooses a variable in the inequality to minimize the *) (* blowup (number of consequences generated) and eliminates it. *) (* ------------------------------------------------------------------------- *) fun extract_first p = let fun extract xs (y::ys) = if p y then (y, xs @ ys) else extract (y::xs) ys | extract _ [] = raise List.Empty in extract [] end; fun print_ineqs ctxt ineqs = if Config.get ctxt LA_Data.trace then tracing(cat_lines(""::map (fn Lineq(c,t,l,_) => string_of_int c ^ (case t of Eq => " = " | Lt=> " < " | Le => " <= ") ^ commas(map string_of_int l)) ineqs)) else (); type history = (int * lineq list) list; datatype result = Success of injust | Failure of history; fun elim ctxt (ineqs, hist) = let val _ = print_ineqs ctxt ineqs val (triv, nontriv) = List.partition is_trivial ineqs in if not (null triv) then case find_first is_contradictory triv of NONE => elim ctxt (nontriv, hist) | SOME(Lineq(_,_,_,j)) => Success j else if null nontriv then Failure hist else let val (eqs, noneqs) = List.partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in if not (null eqs) then let val c = fold (fn Lineq(_,_,l,_) => fn cs => union (op =) l cs) eqs [] |> filter (fn i => i <> 0) |> sort (int_ord o apply2 abs) |> hd val (eq as Lineq(_,_,ceq,_),othereqs) = extract_first (fn Lineq(_,_,l,_) => member (op =) l c) eqs val v = find_index (fn v => v = c) ceq val (ioth,roth) = List.partition (fn (Lineq(_,_,l,_)) => nth l v = 0) (othereqs @ noneqs) val others = map (elim_var v eq) roth @ ioth in elim ctxt (others,(v,nontriv)::hist) end else let val lists = map (fn (Lineq(_,_,l,_)) => l) noneqs val numlist = 0 upto (length (hd lists) - 1) val coeffs = map (fn i => map (fn xs => nth xs i) lists) numlist val blows = map calc_blowup coeffs val iblows = blows ~~ numlist val nziblows = filter_out (fn (i, _) => i = 0) iblows in if null nziblows then Failure((~1,nontriv)::hist) else let val (_,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows) val (no,yes) = List.partition (fn (Lineq(_,_,l,_)) => nth l v = 0) ineqs val (pos,neg) = List.partition(fn (Lineq(_,_,l,_)) => nth l v > 0) yes in elim ctxt (no @ map_product (elim_var v) pos neg, (v,nontriv)::hist) end end end end; (* ------------------------------------------------------------------------- *) (* Translate back a proof. *) (* ------------------------------------------------------------------------- *) fun trace_thm ctxt msgs th = (if Config.get ctxt LA_Data.trace then tracing (cat_lines (msgs @ [Thm.string_of_thm ctxt th])) else (); th); fun trace_term ctxt msgs t = (if Config.get ctxt LA_Data.trace then tracing (cat_lines (msgs @ [Syntax.string_of_term ctxt t])) else (); t); fun trace_msg ctxt msg = if Config.get ctxt LA_Data.trace then tracing msg else (); val union_term = union Envir.aeconv; fun add_atoms (lhs, _, _, rhs, _, _) = union_term (map fst lhs) o union_term (map fst rhs); fun atoms_of ds = fold add_atoms ds []; (* Simplification may detect a contradiction 'prematurely' due to type information: n+1 <= 0 is simplified to False and does not need to be crossed with 0 <= n. *) local exception FalseE of thm * (int * cterm) list * Proof.context in fun mkthm ctxt asms (just: injust) = let val thy = Proof_Context.theory_of ctxt; val {add_mono_thms = add_mono_thms0, mult_mono_thms = mult_mono_thms0, inj_thms = inj_thms0, lessD = lessD0, simpset, ...} = get_data ctxt; val add_mono_thms = map (Thm.transfer thy) add_mono_thms0; val mult_mono_thms = map (Thm.transfer thy) mult_mono_thms0; val inj_thms = map (Thm.transfer thy) inj_thms0; val lessD = map (Thm.transfer thy) lessD0; val number_of = number_of ctxt; val simpset_ctxt = put_simpset simpset ctxt; fun only_concl f thm = if Thm.no_prems thm then f (Thm.concl_of thm) else NONE; val atoms = atoms_of (map_filter (only_concl (LA_Data.decomp ctxt)) asms); fun use_first rules thm = get_first (fn th => SOME (thm RS th) handle THM _ => NONE) rules fun add2 thm1 thm2 = use_first add_mono_thms (thm1 RS (thm2 RS LA_Logic.conjI)); fun try_add thms thm = get_first (fn th => add2 th thm) thms; fun add_thms thm1 thm2 = (case add2 thm1 thm2 of NONE => (case try_add ([thm1] RL inj_thms) thm2 of NONE => (the (try_add ([thm2] RL inj_thms) thm1) handle Option.Option => (trace_thm ctxt [] thm1; trace_thm ctxt [] thm2; raise Fail "Linear arithmetic: failed to add thms")) | SOME thm => thm) | SOME thm => thm); fun mult_by_add n thm = let fun mul i th = if i = 1 then th else mul (i - 1) (add_thms thm th) in mul n thm end; val rewr = Simplifier.rewrite simpset_ctxt; val rewrite_concl = Conv.fconv_rule (Conv.concl_conv ~1 (Conv.arg_conv (Conv.binop_conv rewr))); fun discharge_prem thm = if Thm.nprems_of thm = 0 then thm else let val cv = Conv.arg1_conv (Conv.arg_conv rewr) in Thm.implies_elim (Conv.fconv_rule cv thm) LA_Logic.trueI end fun mult n thm = (case use_first mult_mono_thms thm of NONE => mult_by_add n thm | SOME mth => let val cv = mth |> Thm.cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg1 val T = Thm.typ_of_cterm cv in mth - |> Thm.instantiate ([], [(dest_Var (Thm.term_of cv), number_of T n)]) + |> Thm.instantiate (TVars.empty, Vars.make [(dest_Var (Thm.term_of cv), number_of T n)]) |> rewrite_concl |> discharge_prem handle CTERM _ => mult_by_add n thm | THM _ => mult_by_add n thm end); fun mult_thm n thm = if n = ~1 then thm RS LA_Logic.sym else if n < 0 then mult (~n) thm RS LA_Logic.sym else mult n thm; fun simp thm (cx as (_, hyps, ctxt')) = let val thm' = trace_thm ctxt ["Simplified:"] (full_simplify simpset_ctxt thm) in if LA_Logic.is_False thm' then raise FalseE (thm', hyps, ctxt') else (thm', cx) end; fun abs_thm i (cx as (terms, hyps, ctxt)) = (case AList.lookup (op =) hyps i of SOME ct => (Thm.assume ct, cx) | NONE => let val thm = nth asms i val (t, (terms', ctxt')) = LA_Data.abstract (Thm.prop_of thm) (terms, ctxt) val ct = Thm.cterm_of ctxt' t in (Thm.assume ct, (terms', (i, ct) :: hyps, ctxt')) end); fun nat_thm t (terms, hyps, ctxt) = let val (t', (terms', ctxt')) = LA_Data.abstract_arith t (terms, ctxt) in (LA_Logic.mk_nat_thm thy t', (terms', hyps, ctxt')) end; fun step0 msg (thm, cx) = (trace_thm ctxt [msg] thm, cx) fun step1 msg j f cx = mk j cx |>> f |>> trace_thm ctxt [msg] and step2 msg j1 j2 f cx = mk j1 cx ||>> mk j2 |>> f |>> trace_thm ctxt [msg] and mk (Asm i) cx = step0 ("Asm " ^ string_of_int i) (abs_thm i cx) | mk (Nat i) cx = step0 ("Nat " ^ string_of_int i) (nat_thm (nth atoms i) cx) | mk (LessD j) cx = step1 "L" j (fn thm => hd ([thm] RL lessD)) cx | mk (NotLeD j) cx = step1 "NLe" j (fn thm => thm RS LA_Logic.not_leD) cx | mk (NotLeDD j) cx = step1 "NLeD" j (fn thm => hd ([thm RS LA_Logic.not_leD] RL lessD)) cx | mk (NotLessD j) cx = step1 "NL" j (fn thm => thm RS LA_Logic.not_lessD) cx | mk (Added (j1, j2)) cx = step2 "+" j1 j2 (uncurry add_thms) cx |-> simp | mk (Multiplied (n, j)) cx = (trace_msg ctxt ("*" ^ string_of_int n); step1 "*" j (mult_thm n) cx) fun finish ctxt' hyps thm = thm |> fold_rev (Thm.implies_intr o snd) hyps |> singleton (Variable.export ctxt' ctxt) |> fold (fn (i, _) => fn thm => nth asms i RS thm) hyps in let val _ = trace_msg ctxt "mkthm"; val (thm, (_, hyps, ctxt')) = mk just ([], [], ctxt); val _ = trace_thm ctxt ["Final thm:"] thm; val fls = simplify simpset_ctxt thm; val _ = trace_thm ctxt ["After simplification:"] fls; val _ = if LA_Logic.is_False fls then () else (tracing (cat_lines (["Assumptions:"] @ map (Thm.string_of_thm ctxt) asms @ [""] @ ["Proved:", Thm.string_of_thm ctxt fls, ""])); warning "Linear arithmetic should have refuted the assumptions.\n\ \Please inform Tobias Nipkow.") in finish ctxt' hyps fls end handle FalseE (thm, hyps, ctxt') => trace_thm ctxt ["False reached early:"] (finish ctxt' hyps thm) end; end; fun coeff poly atom = AList.lookup Envir.aeconv poly atom |> the_default 0; fun integ(rlhs,r,rel,rrhs,s,d) = let val (rn,rd) = Rat.dest r and (sn,sd) = Rat.dest s val m = Integer.lcms(map (snd o Rat.dest) (r :: s :: map snd rlhs @ map snd rrhs)) fun mult(t,r) = let val (i,j) = Rat.dest r in (t,i * (m div j)) end in (m,(map mult rlhs, rn*(m div rd), rel, map mult rrhs, sn*(m div sd), d)) end fun mklineq atoms = fn (item, k) => let val (m, (lhs,i,rel,rhs,j,discrete)) = integ item val lhsa = map (coeff lhs) atoms and rhsa = map (coeff rhs) atoms val diff = map2 (curry (op -)) rhsa lhsa val c = i-j val just = Asm k fun lineq(c,le,cs,j) = Lineq(c,le,cs, if m=1 then j else Multiplied(m,j)) in case rel of "<=" => lineq(c,Le,diff,just) | "~<=" => if discrete then lineq(1-c,Le,map (op ~) diff,NotLeDD(just)) else lineq(~c,Lt,map (op ~) diff,NotLeD(just)) | "<" => if discrete then lineq(c+1,Le,diff,LessD(just)) else lineq(c,Lt,diff,just) | "~<" => lineq(~c,Le,map (op~) diff,NotLessD(just)) | "=" => lineq(c,Eq,diff,just) | _ => raise Fail ("mklineq" ^ rel) end; (* ------------------------------------------------------------------------- *) (* Print (counter) example *) (* ------------------------------------------------------------------------- *) (* ------------------------------------------------------------------------- *) fun mknat (pTs : typ list) (ixs : int list) (atom : term, i : int) : lineq option = if LA_Logic.is_nat (pTs, atom) then let val l = map (fn j => if j=i then 1 else 0) ixs in SOME (Lineq (0, Le, l, Nat i)) end else NONE; (* This code is tricky. It takes a list of premises in the order they occur in the subgoal. Numerical premises are coded as SOME(tuple), non-numerical ones as NONE. Going through the premises, each numeric one is converted into a Lineq. The tricky bit is to convert ~= which is split into two cases < and >. Thus split_items returns a list of equation systems. This may blow up if there are many ~=, but in practice it does not seem to happen. The really tricky bit is to arrange the order of the cases such that they coincide with the order in which the cases are in the end generated by the tactic that applies the generated refutation thms (see function 'refute_tac'). For variables n of type nat, a constraint 0 <= n is added. *) (* FIXME: To optimize, the splitting of cases and the search for refutations *) (* could be intertwined: separate the first (fully split) case, *) (* refute it, continue with splitting and refuting. Terminate with *) (* failure as soon as a case could not be refuted; i.e. delay further *) (* splitting until after a refutation for other cases has been found. *) fun split_items ctxt do_pre split_neq (Ts, terms) : (typ list * (LA_Data.decomp * int) list) list = let (* splits inequalities '~=' into '<' and '>'; this corresponds to *) (* 'REPEAT_DETERM (eresolve_tac neqE i)' at the theorem/tactic *) (* level *) (* FIXME: this is currently sensitive to the order of theorems in *) (* neqE: The theorem for type "nat" must come first. A *) (* better (i.e. less likely to break when neqE changes) *) (* implementation should *test* which theorem from neqE *) (* can be applied, and split the premise accordingly. *) fun elim_neq (ineqs : (LA_Data.decomp option * bool) list) : (LA_Data.decomp option * bool) list list = let fun elim_neq' _ ([] : (LA_Data.decomp option * bool) list) : (LA_Data.decomp option * bool) list list = [[]] | elim_neq' nat_only ((NONE, is_nat) :: ineqs) = map (cons (NONE, is_nat)) (elim_neq' nat_only ineqs) | elim_neq' nat_only ((ineq as (SOME (l, i, rel, r, j, d), is_nat)) :: ineqs) = if rel = "~=" andalso (not nat_only orelse is_nat) then (* [| ?l ~= ?r; ?l < ?r ==> ?R; ?r < ?l ==> ?R |] ==> ?R *) elim_neq' nat_only (ineqs @ [(SOME (l, i, "<", r, j, d), is_nat)]) @ elim_neq' nat_only (ineqs @ [(SOME (r, j, "<", l, i, d), is_nat)]) else map (cons ineq) (elim_neq' nat_only ineqs) in ineqs |> elim_neq' true |> maps (elim_neq' false) end fun ignore_neq (NONE, bool) = (NONE, bool) | ignore_neq (ineq as SOME (_, _, rel, _, _, _), bool) = if rel = "~=" then (NONE, bool) else (ineq, bool) fun number_hyps _ [] = [] | number_hyps n (NONE::xs) = number_hyps (n+1) xs | number_hyps n ((SOME x)::xs) = (x, n) :: number_hyps (n+1) xs val result = (Ts, terms) |> (* user-defined preprocessing of the subgoal *) (if do_pre then LA_Data.pre_decomp ctxt else Library.single) |> tap (fn subgoals => trace_msg ctxt ("Preprocessing yields " ^ string_of_int (length subgoals) ^ " subgoal(s) total.")) |> (* produce the internal encoding of (in-)equalities *) map (apsnd (map (fn t => (LA_Data.decomp ctxt t, LA_Data.domain_is_nat t)))) |> (* splitting of inequalities *) map (apsnd (if split_neq then elim_neq else Library.single o map ignore_neq)) |> maps (fn (Ts, subgoals) => map (pair Ts o map fst) subgoals) |> (* numbering of hypotheses, ignoring irrelevant ones *) map (apsnd (number_hyps 0)) in trace_msg ctxt ("Splitting of inequalities yields " ^ string_of_int (length result) ^ " subgoal(s) total."); result end; fun refutes ctxt : (typ list * (LA_Data.decomp * int) list) list -> injust list -> injust list option = let fun refute ((Ts, initems : (LA_Data.decomp * int) list) :: initemss) (js: injust list) = let val atoms = atoms_of (map fst initems) val n = length atoms val mkleq = mklineq atoms val ixs = 0 upto (n - 1) val iatoms = atoms ~~ ixs val natlineqs = map_filter (mknat Ts ixs) iatoms val ineqs = map mkleq initems @ natlineqs in (case elim ctxt (ineqs, []) of Success j => (trace_msg ctxt ("Contradiction! (" ^ string_of_int (length js + 1) ^ ")"); refute initemss (js @ [j])) | Failure _ => NONE) end | refute [] js = SOME js in refute end; fun refute ctxt params do_pre split_neq terms : injust list option = refutes ctxt (split_items ctxt do_pre split_neq (map snd params, terms)) []; fun count P xs = length (filter P xs); fun prove ctxt params do_pre Hs concl : bool * injust list option = let val _ = trace_msg ctxt "prove:" (* append the negated conclusion to 'Hs' -- this corresponds to *) (* 'DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i)' at the *) (* theorem/tactic level *) val Hs' = Hs @ [LA_Logic.neg_prop concl] fun is_neq NONE = false | is_neq (SOME (_,_,r,_,_,_)) = (r = "~=") val neq_limit = Config.get ctxt LA_Data.neq_limit val split_neq = count is_neq (map (LA_Data.decomp ctxt) Hs') <= neq_limit in if split_neq then () else trace_msg ctxt ("neq_limit exceeded (current value is " ^ string_of_int neq_limit ^ "), ignoring all inequalities"); (split_neq, refute ctxt params do_pre split_neq Hs') end handle TERM ("neg_prop", _) => (* since no meta-logic negation is available, we can only fail if *) (* the conclusion is not of the form 'Trueprop $ _' (simply *) (* dropping the conclusion doesn't work either, because even *) (* 'False' does not imply arbitrary 'concl::prop') *) (trace_msg ctxt "prove failed (cannot negate conclusion)."; (false, NONE)); fun refute_tac ctxt (i, split_neq, justs) = fn state => let val _ = trace_thm ctxt ["refute_tac (on subgoal " ^ string_of_int i ^ ", with " ^ string_of_int (length justs) ^ " justification(s)):"] state val neqE = get_neqE ctxt; fun just1 j = (* eliminate inequalities *) (if split_neq then REPEAT_DETERM (eresolve_tac ctxt neqE i) else all_tac) THEN PRIMITIVE (trace_thm ctxt ["State after neqE:"]) THEN (* use theorems generated from the actual justifications *) Subgoal.FOCUS (fn {prems, ...} => resolve_tac ctxt [mkthm ctxt prems j] 1) ctxt i in (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *) DETERM (resolve_tac ctxt [LA_Logic.notI, LA_Logic.ccontr] i) THEN (* user-defined preprocessing of the subgoal *) DETERM (LA_Data.pre_tac ctxt i) THEN PRIMITIVE (trace_thm ctxt ["State after pre_tac:"]) THEN (* prove every resulting subgoal, using its justification *) EVERY (map just1 justs) end state; (* Fast but very incomplete decider. Only premises and conclusions that are already (negated) (in)equations are taken into account. *) fun simpset_lin_arith_tac ctxt = SUBGOAL (fn (A, i) => let val params = rev (Logic.strip_params A) val Hs = Logic.strip_assums_hyp A val concl = Logic.strip_assums_concl A val _ = trace_term ctxt ["Trying to refute subgoal " ^ string_of_int i] A in case prove ctxt params true Hs concl of (_, NONE) => (trace_msg ctxt "Refutation failed."; no_tac) | (split_neq, SOME js) => (trace_msg ctxt "Refutation succeeded."; refute_tac ctxt (i, split_neq, js)) end); fun prems_lin_arith_tac ctxt = Method.insert_tac ctxt (Simplifier.prems_of ctxt) THEN' simpset_lin_arith_tac ctxt; fun lin_arith_tac ctxt = simpset_lin_arith_tac (empty_simpset ctxt); (** Forward proof from theorems **) (* More tricky code. Needs to arrange the proofs of the multiple cases (due to splits of ~= premises) such that it coincides with the order of the cases generated by function split_items. *) datatype splittree = Tip of thm list | Spl of thm * cterm * splittree * cterm * splittree; (* "(ct1 ==> ?R) ==> (ct2 ==> ?R) ==> ?R" is taken to (ct1, ct2) *) fun extract (imp : cterm) : cterm * cterm = let val (Il, r) = Thm.dest_comb imp val (_, imp1) = Thm.dest_comb Il val (Ict1, _) = Thm.dest_comb imp1 val (_, ct1) = Thm.dest_comb Ict1 val (Ir, _) = Thm.dest_comb r val (_, Ict2r) = Thm.dest_comb Ir val (Ict2, _) = Thm.dest_comb Ict2r val (_, ct2) = Thm.dest_comb Ict2 in (ct1, ct2) end; fun splitasms ctxt (asms : thm list) : splittree = let val neqE = get_neqE ctxt fun elim_neq [] (asms', []) = Tip (rev asms') | elim_neq [] (asms', asms) = Tip (rev asms' @ asms) | elim_neq (_ :: neqs) (asms', []) = elim_neq neqs ([],rev asms') | elim_neq (neqs as (neq :: _)) (asms', asm::asms) = (case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) [neq] of SOME spl => let val (ct1, ct2) = extract (Thm.cprop_of spl) val thm1 = Thm.assume ct1 val thm2 = Thm.assume ct2 in Spl (spl, ct1, elim_neq neqs (asms', asms@[thm1]), ct2, elim_neq neqs (asms', asms@[thm2])) end | NONE => elim_neq neqs (asm::asms', asms)) in elim_neq neqE ([], asms) end; fun fwdproof ctxt (Tip asms : splittree) (j::js : injust list) = (mkthm ctxt asms j, js) | fwdproof ctxt (Spl (thm, ct1, tree1, ct2, tree2)) js = let val (thm1, js1) = fwdproof ctxt tree1 js val (thm2, js2) = fwdproof ctxt tree2 js1 val thm1' = Thm.implies_intr ct1 thm1 val thm2' = Thm.implies_intr ct2 thm2 in (thm2' COMP (thm1' COMP thm), js2) end; (* FIXME needs handle THM _ => NONE ? *) fun prover ctxt thms Tconcl (js : injust list) split_neq pos : thm option = let val nTconcl = LA_Logic.neg_prop Tconcl val cnTconcl = Thm.cterm_of ctxt nTconcl val nTconclthm = Thm.assume cnTconcl val tree = (if split_neq then splitasms ctxt else Tip) (thms @ [nTconclthm]) val (Falsethm, _) = fwdproof ctxt tree js val contr = if pos then LA_Logic.ccontr else LA_Logic.notI val concl = Thm.implies_intr cnTconcl Falsethm COMP contr in SOME (trace_thm ctxt ["Proved by lin. arith. prover:"] (LA_Logic.mk_Eq concl)) end (*in case concl contains ?-var, which makes assume fail:*) (* FIXME Variable.import_terms *) handle THM _ => NONE; (* PRE: concl is not negated! This assumption is OK because 1. lin_arith_simproc tries both to prove and disprove concl and 2. lin_arith_simproc is applied by the Simplifier which dives into terms and will thus try the non-negated concl anyway. *) fun lin_arith_simproc ctxt concl = let val thms = maps LA_Logic.atomize (Simplifier.prems_of ctxt) val Hs = map Thm.prop_of thms val Tconcl = LA_Logic.mk_Trueprop (Thm.term_of concl) in case prove ctxt [] false Hs Tconcl of (* concl provable? *) (split_neq, SOME js) => prover ctxt thms Tconcl js split_neq true | (_, NONE) => let val nTconcl = LA_Logic.neg_prop Tconcl in case prove ctxt [] false Hs nTconcl of (* ~concl provable? *) (split_neq, SOME js) => prover ctxt thms nTconcl js split_neq false | (_, NONE) => NONE end end; end; diff --git a/src/Provers/hypsubst.ML b/src/Provers/hypsubst.ML --- a/src/Provers/hypsubst.ML +++ b/src/Provers/hypsubst.ML @@ -1,306 +1,306 @@ (* Title: Provers/hypsubst.ML Authors: Martin D Coen, Tobias Nipkow and Lawrence C Paulson Copyright 1995 University of Cambridge Basic equational reasoning: hyp_subst_tac and methods "hypsubst", "simplesubst". Tactic to substitute using (at least) the assumption x=t in the rest of the subgoal, and to delete (at least) that assumption. Original version due to Martin Coen. This version uses the simplifier, and requires it to be already present. Test data: Goal "!!x.[| Q(x,y,z); y=x; a=x; z=y; P(y) |] ==> P(z)"; Goal "!!x.[| Q(x,y,z); z=f(x); x=z |] ==> P(z)"; Goal "!!y. [| ?x=y; P(?x) |] ==> y = a"; Goal "!!z. [| ?x=y; P(?x) |] ==> y = a"; Goal "!!x a. [| x = f(b); g(a) = b |] ==> P(x)"; by (bound_hyp_subst_tac 1); by (hyp_subst_tac 1); Here hyp_subst_tac goes wrong; harder still to prove P(f(f(a))) & P(f(a)) Goal "P(a) --> (EX y. a=y --> P(f(a)))"; Goal "!!x. [| Q(x,h1); P(a,h2); R(x,y,h3); R(y,z,h4); x=f(y); \ \ P(x,h5); P(y,h6); K(x,h7) |] ==> Q(x,c)"; by (blast_hyp_subst_tac true 1); *) signature HYPSUBST_DATA = sig val dest_Trueprop : term -> term val dest_eq : term -> term * term val dest_imp : term -> term * term val eq_reflection : thm (* a=b ==> a==b *) val rev_eq_reflection: thm (* a==b ==> a=b *) val imp_intr : thm (* (P ==> Q) ==> P-->Q *) val rev_mp : thm (* [| P; P-->Q |] ==> Q *) val subst : thm (* [| a=b; P(a) |] ==> P(b) *) val sym : thm (* a=b ==> b=a *) val thin_refl : thm (* [|x=x; PROP W|] ==> PROP W *) end; signature HYPSUBST = sig val bound_hyp_subst_tac : Proof.context -> int -> tactic val hyp_subst_tac_thin : bool -> Proof.context -> int -> tactic val hyp_subst_thin : bool Config.T val hyp_subst_tac : Proof.context -> int -> tactic val blast_hyp_subst_tac : Proof.context -> bool -> int -> tactic val stac : Proof.context -> thm -> int -> tactic end; functor Hypsubst(Data: HYPSUBST_DATA): HYPSUBST = struct exception EQ_VAR; (*Simplifier turns Bound variables to special Free variables: change it back (any Bound variable will do)*) fun inspect_contract t = (case Envir.eta_contract t of Free (a, T) => if Name.is_bound a then Bound 0 else Free (a, T) | t' => t'); val has_vars = Term.exists_subterm Term.is_Var; val has_tvars = Term.exists_type (Term.exists_subtype Term.is_TVar); (*If novars then we forbid Vars in the equality. If bnd then we only look for Bound variables to eliminate. When can we safely delete the equality? Not if it equates two constants; consider 0=1. Not if it resembles x=t[x], since substitution does not eliminate x. Not if it resembles ?x=0; consider ?x=0 ==> ?x=1 or even ?x=0 ==> P Not if it involves a variable free in the premises, but we can't check for this -- hence bnd and bound_hyp_subst_tac Prefer to eliminate Bound variables if possible. Result: true = use as is, false = reorient first also returns var to substitute, relevant if it is Free *) fun inspect_pair bnd novars (t, u) = if novars andalso (has_tvars t orelse has_tvars u) then raise Match (*variables in the type!*) else (case apply2 inspect_contract (t, u) of (Bound i, _) => if loose_bvar1 (u, i) orelse novars andalso has_vars u then raise Match else (true, Bound i) (*eliminates t*) | (_, Bound i) => if loose_bvar1 (t, i) orelse novars andalso has_vars t then raise Match else (false, Bound i) (*eliminates u*) | (t' as Free _, _) => if bnd orelse Logic.occs (t', u) orelse novars andalso has_vars u then raise Match else (true, t') (*eliminates t*) | (_, u' as Free _) => if bnd orelse Logic.occs (u', t) orelse novars andalso has_vars t then raise Match else (false, u') (*eliminates u*) | _ => raise Match); (*Locates a substitutable variable on the left (resp. right) of an equality assumption. Returns the number of intervening assumptions. *) fun eq_var bnd novars check_frees t = let fun check_free ts (orient, Free f) = if not check_frees orelse not orient orelse exists (curry Logic.occs (Free f)) ts then (orient, true) else raise Match | check_free ts (orient, _) = (orient, false) fun eq_var_aux k (Const(\<^const_name>\Pure.all\,_) $ Abs(_,_,t)) hs = eq_var_aux k t hs | eq_var_aux k (Const(\<^const_name>\Pure.imp\,_) $ A $ B) hs = ((k, check_free (B :: hs) (inspect_pair bnd novars (Data.dest_eq (Data.dest_Trueprop A)))) handle TERM _ => eq_var_aux (k+1) B (A :: hs) | Match => eq_var_aux (k+1) B (A :: hs)) | eq_var_aux k _ _ = raise EQ_VAR in eq_var_aux 0 t [] end; fun thin_free_eq_tac ctxt = SUBGOAL (fn (t, i) => let val (k, _) = eq_var false false false t val ok = (eq_var false false true t |> fst) > k handle EQ_VAR => true in if ok then EVERY [rotate_tac k i, eresolve_tac ctxt [thin_rl] i, rotate_tac (~k) i] else no_tac end handle EQ_VAR => no_tac) (*For the simpset. Adds ALL suitable equalities, even if not first! No vars are allowed here, as simpsets are built from meta-assumptions*) fun mk_eqs bnd th = [ if inspect_pair bnd false (Data.dest_eq (Data.dest_Trueprop (Thm.prop_of th))) |> fst then th RS Data.eq_reflection else Thm.symmetric(th RS Data.eq_reflection) (*reorient*) ] handle TERM _ => [] | Match => []; (*Select a suitable equality assumption; substitute throughout the subgoal If bnd is true, then it replaces Bound variables only. *) fun gen_hyp_subst_tac ctxt bnd = SUBGOAL (fn (Bi, i) => let val (k, (orient, is_free)) = eq_var bnd true true Bi val hyp_subst_ctxt = empty_simpset ctxt |> Simplifier.set_mksimps (K (mk_eqs bnd)) in EVERY [rotate_tac k i, asm_lr_simp_tac hyp_subst_ctxt i, if not is_free then eresolve_tac ctxt [thin_rl] i else if orient then eresolve_tac ctxt [Data.rev_mp] i else eresolve_tac ctxt [Data.sym RS Data.rev_mp] i, rotate_tac (~k) i, if is_free then resolve_tac ctxt [Data.imp_intr] i else all_tac] end handle THM _ => no_tac | EQ_VAR => no_tac) val ssubst = Drule.zero_var_indexes (Data.sym RS Data.subst); fun inst_subst_tac ctxt b rl = CSUBGOAL (fn (cBi, i) => case try (Logic.strip_assums_hyp #> hd #> Data.dest_Trueprop #> Data.dest_eq #> apply2 Envir.eta_contract) (Thm.term_of cBi) of SOME (t, t') => let val Bi = Thm.term_of cBi; val ps = Logic.strip_params Bi; val U = Term.fastype_of1 (rev (map snd ps), t); val Q = Data.dest_Trueprop (Logic.strip_assums_concl Bi); val rl' = Thm.lift_rule cBi rl; val (ixn, T) = dest_Var (Term.head_of (Data.dest_Trueprop (Logic.strip_assums_concl (Thm.prop_of rl')))); val (v1, v2) = Data.dest_eq (Data.dest_Trueprop (Logic.strip_assums_concl (hd (Thm.prems_of rl')))); val (Ts, V) = split_last (Term.binder_types T); val u = fold_rev Term.abs (ps @ [("x", U)]) (case (if b then t else t') of Bound j => subst_bounds (map Bound ((1 upto j) @ 0 :: (j + 2 upto length ps)), Q) | t => Term.abstract_over (t, Term.incr_boundvars 1 Q)); val (instT, _) = Thm.match (apply2 (Thm.cterm_of ctxt o Logic.mk_type) (V, U)); in compose_tac ctxt (true, Drule.instantiate_normalize (instT, - map (apsnd (Thm.cterm_of ctxt)) + Vars.make (map (apsnd (Thm.cterm_of ctxt)) [((ixn, Ts ---> U --> body_type T), u), ((fst (dest_Var (head_of v1)), Ts ---> U), fold_rev Term.abs ps t), - ((fst (dest_Var (head_of v2)), Ts ---> U), fold_rev Term.abs ps t')]) rl', + ((fst (dest_Var (head_of v2)), Ts ---> U), fold_rev Term.abs ps t')])) rl', Thm.nprems_of rl) i end | NONE => no_tac); fun imp_intr_tac ctxt = resolve_tac ctxt [Data.imp_intr]; fun rev_dup_elim ctxt th = (th RSN (2, revcut_rl)) |> Thm.assumption (SOME ctxt) 2 |> Seq.hd; fun dup_subst ctxt = rev_dup_elim ctxt ssubst (* FIXME: "etac Data.rev_mp i" will not behave as expected if goal has *) (* premises containing meta-implications or quantifiers *) (*Old version of the tactic above -- slower but the only way to handle equalities containing Vars.*) fun vars_gen_hyp_subst_tac ctxt bnd = SUBGOAL(fn (Bi,i) => let val n = length(Logic.strip_assums_hyp Bi) - 1 val (k, (orient, is_free)) = eq_var bnd false true Bi val rl = if is_free then dup_subst ctxt else ssubst val rl = if orient then rl else Data.sym RS rl in DETERM (EVERY [REPEAT_DETERM_N k (eresolve_tac ctxt [Data.rev_mp] i), rotate_tac 1 i, REPEAT_DETERM_N (n-k) (eresolve_tac ctxt [Data.rev_mp] i), inst_subst_tac ctxt orient rl i, REPEAT_DETERM_N n (imp_intr_tac ctxt i THEN rotate_tac ~1 i)]) end handle THM _ => no_tac | EQ_VAR => no_tac); (*Substitutes for Free or Bound variables, discarding equalities on Bound variables and on Free variables if thin=true*) fun hyp_subst_tac_thin thin ctxt = REPEAT_DETERM1 o FIRST' [ematch_tac ctxt [Data.thin_refl], gen_hyp_subst_tac ctxt false, vars_gen_hyp_subst_tac ctxt false, if thin then thin_free_eq_tac ctxt else K no_tac]; val hyp_subst_thin = Attrib.setup_config_bool \<^binding>\hypsubst_thin\ (K false); fun hyp_subst_tac ctxt = hyp_subst_tac_thin (Config.get ctxt hyp_subst_thin) ctxt; (*Substitutes for Bound variables only -- this is always safe*) fun bound_hyp_subst_tac ctxt = REPEAT_DETERM1 o (gen_hyp_subst_tac ctxt true ORELSE' vars_gen_hyp_subst_tac ctxt true); (** Version for Blast_tac. Hyps that are affected by the substitution are moved to the front. Defect: even trivial changes are noticed, such as substitutions in the arguments of a function Var. **) (*final re-reversal of the changed assumptions*) fun reverse_n_tac _ 0 i = all_tac | reverse_n_tac _ 1 i = rotate_tac ~1 i | reverse_n_tac ctxt n i = REPEAT_DETERM_N n (rotate_tac ~1 i THEN eresolve_tac ctxt [Data.rev_mp] i) THEN REPEAT_DETERM_N n (imp_intr_tac ctxt i THEN rotate_tac ~1 i); (*Use imp_intr, comparing the old hyps with the new ones as they come out.*) fun all_imp_intr_tac ctxt hyps i = let fun imptac (r, []) st = reverse_n_tac ctxt r i st | imptac (r, hyp::hyps) st = let val (hyp', _) = Thm.term_of (Thm.cprem_of st i) |> Logic.strip_assums_concl |> Data.dest_Trueprop |> Data.dest_imp; val (r', tac) = if Envir.aeconv (hyp, hyp') then (r, imp_intr_tac ctxt i THEN rotate_tac ~1 i) else (*leave affected hyps at end*) (r + 1, imp_intr_tac ctxt i); in (case Seq.pull (tac st) of NONE => Seq.single st | SOME (st', _) => imptac (r', hyps) st') end in imptac (0, rev hyps) end; fun blast_hyp_subst_tac ctxt trace = SUBGOAL(fn (Bi, i) => let val (k, (symopt, _)) = eq_var false false false Bi val hyps0 = map Data.dest_Trueprop (Logic.strip_assums_hyp Bi) (*omit selected equality, returning other hyps*) val hyps = List.take(hyps0, k) @ List.drop(hyps0, k+1) val n = length hyps in if trace then tracing "Substituting an equality" else (); DETERM (EVERY [REPEAT_DETERM_N k (eresolve_tac ctxt [Data.rev_mp] i), rotate_tac 1 i, REPEAT_DETERM_N (n-k) (eresolve_tac ctxt [Data.rev_mp] i), inst_subst_tac ctxt symopt (if symopt then ssubst else Data.subst) i, all_imp_intr_tac ctxt hyps i]) end handle THM _ => no_tac | EQ_VAR => no_tac); (*apply an equality or definition ONCE; fails unless the substitution has an effect*) fun stac ctxt th = let val th' = th RS Data.rev_eq_reflection handle THM _ => th in CHANGED_GOAL (resolve_tac ctxt [th' RS ssubst]) end; (* method setup *) val _ = Theory.setup (Method.setup \<^binding>\hypsubst\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o hyp_subst_tac ctxt))) "substitution using an assumption (improper)" #> Method.setup \<^binding>\hypsubst_thin\ (Scan.succeed (fn ctxt => SIMPLE_METHOD' (CHANGED_PROP o hyp_subst_tac_thin true ctxt))) "substitution using an assumption, eliminating assumptions" #> Method.setup \<^binding>\simplesubst\ (Attrib.thm >> (fn th => fn ctxt => SIMPLE_METHOD' (stac ctxt th))) "simple substitution"); end; diff --git a/src/Pure/Isar/class_declaration.ML b/src/Pure/Isar/class_declaration.ML --- a/src/Pure/Isar/class_declaration.ML +++ b/src/Pure/Isar/class_declaration.ML @@ -1,395 +1,398 @@ (* Title: Pure/Isar/class_declaration.ML Author: Florian Haftmann, TU Muenchen Declaring classes and subclass relations. *) signature CLASS_DECLARATION = sig val class: binding -> Bundle.name list -> class list -> Element.context_i list -> theory -> string * local_theory val class_cmd: binding -> (xstring * Position.T) list -> xstring list -> Element.context list -> theory -> string * local_theory val prove_subclass: tactic -> class -> local_theory -> local_theory val subclass: class -> local_theory -> Proof.state val subclass_cmd: xstring -> local_theory -> Proof.state end; structure Class_Declaration: CLASS_DECLARATION = struct (** class definitions **) local (* calculating class-related rules including canonical interpretation *) fun calculate thy class sups base_sort param_map assm_axiom = let val thy_ctxt = Proof_Context.init_global thy; val certT = Thm.trim_context_ctyp o Thm.global_ctyp_of thy; val cert = Thm.trim_context_cterm o Thm.global_cterm_of thy; (* instantiation of canonical interpretation *) val a_tfree = (Name.aT, base_sort); val a_type = TFree a_tfree; val param_map_const = (map o apsnd) Const param_map; - val param_map_inst = param_map |> map (fn (x, (c, T)) => - let val T' = map_atyps (K a_type) T in ((x, T'), cert (Const (c, T'))) end); + val param_map_inst = + Frees.build (param_map |> fold (fn (x, (c, T)) => + let val T' = map_atyps (K a_type) T + in Frees.add ((x, T'), cert (Const (c, T'))) end)); val const_morph = - Element.instantiate_normalize_morphism ([], param_map_inst); + Element.instantiate_normalize_morphism (TFrees.empty, param_map_inst); val typ_morph = - Element.instantiate_normalize_morphism ([(a_tfree, certT (Term.aT [class]))], []) + Element.instantiate_normalize_morphism + (TFrees.make [(a_tfree, certT (Term.aT [class]))], Frees.empty) val (([raw_props], _, [(_, raw_inst_morph)], _, export_morph), _) = thy_ctxt |> Expression.cert_goal_expression ([(class, (("", false), (Expression.Named param_map_const, [])))], []); val (props, inst_morph) = if null param_map then (raw_props |> map (Morphism.term typ_morph), raw_inst_morph $> typ_morph) else (raw_props, raw_inst_morph); (*FIXME proper handling in locale.ML / expression.ML would be desirable*) (* witness for canonical interpretation *) val some_prop = try the_single props; val some_witn = Option.map (fn prop => let val sup_axioms = map_filter (fst o Class.rules thy) sups; val loc_intro_tac = (case Locale.intros_of thy class of (_, NONE) => all_tac | (_, SOME intro) => ALLGOALS (resolve_tac thy_ctxt [intro])); val tac = loc_intro_tac THEN ALLGOALS (Proof_Context.fact_tac thy_ctxt (sup_axioms @ the_list assm_axiom)); in Element.prove_witness thy_ctxt prop tac end) some_prop; val some_axiom = Option.map (Element.conclude_witness thy_ctxt) some_witn; (* canonical interpretation *) val base_morph = inst_morph $> Morphism.binding_morphism "class_binding" (Binding.prefix false (Class.class_prefix class)) $> Element.satisfy_morphism (the_list some_witn); val eq_morph = Element.eq_morphism thy (Class.these_defs thy sups); (* assm_intro *) fun prove_assm_intro thm = let val ((_, [thm']), _) = Variable.import true [thm] thy_ctxt; val const_eq_morph = (case eq_morph of SOME eq_morph => const_morph $> eq_morph | NONE => const_morph); val thm'' = Morphism.thm const_eq_morph thm'; in Goal.prove_sorry_global thy [] [] (Thm.prop_of thm'') (fn {context = ctxt, ...} => ALLGOALS (Proof_Context.fact_tac ctxt [thm''])) end; val some_assm_intro = Option.map prove_assm_intro (fst (Locale.intros_of thy class)); (* of_class *) val of_class_prop_concl = Logic.mk_of_class (a_type, class); val of_class_prop = (case some_prop of NONE => of_class_prop_concl | SOME prop => Logic.mk_implies (Morphism.term const_morph ((map_types o map_atyps) (K a_type) prop), of_class_prop_concl)); val sup_of_classes = map (snd o Class.rules thy) sups; val loc_axiom_intros = map Drule.export_without_context_open (Locale.axioms_of thy class); val axclass_intro = #intro (Axclass.get_info thy class); val base_sort_trivs = Thm.of_sort (Thm.global_ctyp_of thy a_type, base_sort); fun tac ctxt = REPEAT (SOMEGOAL (match_tac ctxt (axclass_intro :: sup_of_classes @ loc_axiom_intros @ base_sort_trivs) ORELSE' assume_tac ctxt)); val of_class = Goal.prove_sorry_global thy [] [] of_class_prop (tac o #context); in (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) end; (* reading and processing class specifications *) fun prep_class_elems prep_decl ctxt sups raw_elems = let (* user space type system: only permits 'a type variable, improves towards 'a *) val thy = Proof_Context.theory_of ctxt; val algebra = Sign.classes_of thy; val inter_sort = curry (Sorts.inter_sort algebra); val proto_base_sort = if null sups then Sign.defaultS thy else fold inter_sort (map (Class.base_sort thy) sups) []; val is_param = member (op =) (map fst (Class.these_params thy sups)); val base_constraints = (map o apsnd) (map_type_tfree (K (TVar ((Name.aT, 0), proto_base_sort))) o fst o fst o snd) (Class.these_operations thy sups); val singleton_fixate = burrow_types (fn Ts => let val tfrees = fold Term.add_tfreesT Ts []; val inferred_sort = (fold o fold_atyps) (fn TVar (_, S) => inter_sort S | _ => I) Ts []; val fixate_sort = (case tfrees of [] => inferred_sort | [(a, S)] => if a <> Name.aT then error ("No type variable other than " ^ Name.aT ^ " allowed in class specification") else if Sorts.sort_le algebra (S, inferred_sort) then S else error ("Type inference imposes additional sort constraint " ^ Syntax.string_of_sort_global thy inferred_sort ^ " of type parameter " ^ Name.aT ^ " of sort " ^ Syntax.string_of_sort_global thy S) | _ => error "Multiple type variables in class specification"); val fixateT = Term.aT fixate_sort; in (map o map_atyps) (fn T as TVar (xi, _) => if Type_Infer.is_param xi then fixateT else T | T => T) Ts end); fun unify_params ts = let val param_Ts = (fold o fold_aterms) (fn Free (v, T) => if is_param v then fold_atyps (insert (op =)) T else I | _ => I) ts []; val param_namesT = map_filter (try (fst o dest_TVar)) param_Ts; val param_T = if null param_namesT then NONE else SOME (case get_first (try dest_TFree) param_Ts of SOME v_sort => TFree v_sort | NONE => TVar (hd param_namesT, proto_base_sort)); in case param_T of NONE => ts | SOME T => map (subst_TVars (map (rpair T) param_namesT)) ts end; (* preprocessing elements, retrieving base sort from type-checked elements *) val raw_supexpr = (map (fn sup => (sup, (("", false), (Expression.Positional [], [])))) sups, []); val init_class_body = fold (Proof_Context.add_const_constraint o apsnd SOME) base_constraints #> Class.redeclare_operations thy sups #> Context.proof_map (Syntax_Phases.term_check 0 "singleton_fixate" (K singleton_fixate)); val ((raw_supparams, _, raw_inferred_elems, _), _) = ctxt |> Context.proof_map (Syntax_Phases.term_check 0 "unify_params" (K unify_params)) |> prep_decl raw_supexpr init_class_body raw_elems; fun filter_element (Element.Fixes []) = NONE | filter_element (e as Element.Fixes _) = SOME e | filter_element (Element.Constrains []) = NONE | filter_element (e as Element.Constrains _) = SOME e | filter_element (Element.Assumes []) = NONE | filter_element (e as Element.Assumes _) = SOME e | filter_element (Element.Defines _) = error ("\"defines\" element not allowed in class specification.") | filter_element (Element.Notes _) = error ("\"notes\" element not allowed in class specification.") | filter_element (Element.Lazy_Notes _) = error ("\"notes\" element not allowed in class specification."); val inferred_elems = map_filter filter_element raw_inferred_elems; fun fold_element_types f (Element.Fixes fxs) = fold (fn (_, SOME T, _) => f T) fxs | fold_element_types f (Element.Constrains cnstrs) = fold (f o snd) cnstrs | fold_element_types f (Element.Assumes assms) = fold (fold (fn (t, ts) => fold_types f t #> (fold o fold_types) f ts) o snd) assms; val base_sort = if null inferred_elems then proto_base_sort else (case (fold o fold_element_types) Term.add_tfreesT inferred_elems [] of [] => error "No type variable in class specification" | [(_, sort)] => sort | _ => error "Multiple type variables in class specification"); val supparams = map (fn ((c, T), _) => (c, map_atyps (K (Term.aT base_sort)) T)) raw_supparams; val supparam_names = map fst supparams; fun mk_param ((c, _), _) = Free (c, (the o AList.lookup (op =) supparams) c); val supexpr = (map (fn sup => (sup, (("", false), (Expression.Positional (map (SOME o mk_param) (Locale.params_of thy sup)), [])))) sups, map (fn (c, T) => (Binding.name c, SOME T, NoSyn)) supparams); in (base_sort, supparam_names, supexpr, inferred_elems) end; val cert_class_elems = prep_class_elems Expression.cert_declaration; val read_class_elems = prep_class_elems Expression.cert_read_declaration; fun prep_class_spec prep_class prep_include prep_class_elems ctxt raw_supclasses raw_includes raw_elems = let val thy = Proof_Context.theory_of ctxt; (* prepare import *) val inter_sort = curry (Sorts.inter_sort (Sign.classes_of thy)); val sups = Sign.minimize_sort thy (map (prep_class ctxt) raw_supclasses); val _ = (case filter_out (Class.is_class thy) sups of [] => () | no_classes => error ("No (proper) classes: " ^ commas_quote no_classes)); val raw_supparams = (map o apsnd) (snd o snd) (Class.these_params thy sups); val raw_supparam_names = map fst raw_supparams; val _ = if has_duplicates (op =) raw_supparam_names then error ("Duplicate parameter(s) in superclasses: " ^ (commas_quote (duplicates (op =) raw_supparam_names))) else (); (* infer types and base sort *) val includes = map (prep_include ctxt) raw_includes; val includes_ctxt = Bundle.includes includes ctxt; val (base_sort, supparam_names, supexpr, inferred_elems) = prep_class_elems includes_ctxt sups raw_elems; val sup_sort = inter_sort base_sort sups; (* process elements as class specification *) val class_ctxt = Class.begin sups base_sort includes_ctxt; val ((_, _, syntax_elems, _), _) = class_ctxt |> Expression.cert_declaration supexpr I inferred_elems; fun check_vars e vs = if null vs then error ("No type variable in part of specification element " ^ Pretty.string_of (Pretty.chunks (Element.pretty_ctxt class_ctxt e))) else (); fun check_element (e as Element.Fixes fxs) = List.app (fn (_, SOME T, _) => check_vars e (Term.add_tfreesT T [])) fxs | check_element (e as Element.Assumes assms) = List.app (fn (_, ts_pss) => List.app (fn (t, _) => check_vars e (Term.add_tfrees t [])) ts_pss) assms | check_element _ = (); val _ = List.app check_element syntax_elems; fun fork_syn (Element.Fixes xs) = fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs #>> Element.Fixes | fork_syn x = pair x; val (elems, global_syntax) = fold_map fork_syn syntax_elems []; in (((sups, supparam_names), (sup_sort, base_sort, supexpr)), (includes, elems, global_syntax)) end; val cert_class_spec = prep_class_spec (K I) (K I) cert_class_elems; val read_class_spec = prep_class_spec Proof_Context.read_class Bundle.check read_class_elems; (* class establishment *) fun add_consts class base_sort sups supparam_names global_syntax thy = let (*FIXME simplify*) val supconsts = supparam_names |> AList.make (snd o the o AList.lookup (op =) (Class.these_params thy sups)) |> (map o apsnd o apsnd o map_atyps) (K (Term.aT [class])); val all_params = Locale.params_of thy class; val raw_params = (snd o chop (length supparam_names)) all_params; fun add_const ((raw_c, raw_ty), _) thy = let val b = Binding.name raw_c; val c = Sign.full_name thy b; val ty = map_atyps (K (Term.aT base_sort)) raw_ty; val ty0 = Type.strip_sorts ty; val ty' = map_atyps (K (Term.aT [class])) ty0; val syn = (the_default NoSyn o AList.lookup Binding.eq_name global_syntax) b; in thy |> Sign.declare_const_global ((b, ty0), syn) |> snd |> pair ((Variable.check_name b, ty), (c, ty')) end; in thy |> Sign.add_path (Class.class_prefix class) |> fold_map add_const raw_params ||> Sign.restore_naming thy |-> (fn params => pair (supconsts @ (map o apfst) fst params, params)) end; fun adjungate_axclass bname class base_sort sups supsort supparam_names global_syntax thy = let (*FIXME simplify*) fun globalize param_map = map_aterms (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty) | t => t); val raw_pred = Locale.intros_of thy class |> fst |> Option.map (Logic.unvarify_global o Logic.strip_imp_concl o Thm.prop_of); fun get_axiom thy = (case #axioms (Axclass.get_info thy class) of [] => NONE | [thm] => SOME thm); in thy |> add_consts class base_sort sups supparam_names global_syntax |-> (fn (param_map, params) => Axclass.define_class (bname, supsort) (map (fst o snd) params) [(Binding.empty_atts, Option.map (globalize param_map) raw_pred |> the_list)] #> snd #> `get_axiom #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params #> pair (param_map, params, assm_axiom))) end; fun gen_class prep_class_spec b raw_includes raw_supclasses raw_elems thy = let val class = Sign.full_name thy b; val prefix = Binding.qualify true "class"; val ctxt = Proof_Context.init_global thy; val (((sups, supparam_names), (supsort, base_sort, supexpr)), (includes, elems, global_syntax)) = prep_class_spec ctxt raw_supclasses raw_includes raw_elems; in thy |> Expression.add_locale b (prefix b) includes supexpr elems |> snd |> Local_Theory.exit_global |> adjungate_axclass b class base_sort sups supsort supparam_names global_syntax |-> (fn (param_map, params, assm_axiom) => `(fn thy => calculate thy class sups base_sort param_map assm_axiom) #-> (fn (base_morph, eq_morph, export_morph, some_axiom, some_assm_intro, of_class) => Context.theory_map (Locale.add_registration {inst = (class, base_morph), mixin = Option.map (rpair true) eq_morph, export = export_morph}) #> Class.register class sups params base_sort base_morph export_morph some_axiom some_assm_intro of_class #> Global_Theory.store_thm (prefix (Binding.qualified_name (class ^ ".of_class.intro")), of_class))) |> snd |> Named_Target.init includes class |> pair class end; in val class = gen_class cert_class_spec; val class_cmd = gen_class read_class_spec; end; (*local*) (** subclass relations **) local fun gen_subclass prep_class do_proof raw_sup lthy = let val thy = Proof_Context.theory_of lthy; val proto_sup = prep_class thy raw_sup; val proto_sub = (case Named_Target.class_of lthy of SOME class => class | NONE => error "Not in a class target"); val (sub, sup) = Axclass.cert_classrel thy (proto_sub, proto_sup); val expr = ([(sup, (("", false), (Expression.Positional [], [])))], []); val (([props], _, deps, _, export), goal_ctxt) = Expression.cert_goal_expression expr lthy; val some_prop = try the_single props; val some_dep_morph = try the_single (map snd deps); fun after_qed some_wit = Class.register_subclass (sub, sup) some_dep_morph some_wit export; in do_proof after_qed some_prop goal_ctxt end; fun user_proof after_qed some_prop = Element.witness_proof (after_qed o try the_single o the_single) [the_list some_prop]; fun tactic_proof tac after_qed some_prop ctxt = after_qed (Option.map (fn prop => Element.prove_witness ctxt prop tac) some_prop) ctxt; in fun prove_subclass tac = gen_subclass (K I) (tactic_proof tac); fun subclass x = gen_subclass (K I) user_proof x; fun subclass_cmd x = gen_subclass (Proof_Context.read_class o Proof_Context.init_global) user_proof x; end; (*local*) end; diff --git a/src/Pure/Isar/code.ML b/src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML +++ b/src/Pure/Isar/code.ML @@ -1,1584 +1,1587 @@ (* Title: Pure/Isar/code.ML Author: Florian Haftmann, TU Muenchen Abstract executable ingredients of theory. Management of data dependent on executable ingredients as synchronized cache; purged on any change of underlying executable ingredients. *) signature CODE = sig (*constants*) val check_const: theory -> term -> string val read_const: theory -> string -> string val string_of_const: theory -> string -> string val args_number: theory -> string -> int (*constructor sets*) val constrset_of_consts: theory -> (string * typ) list -> string * ((string * sort) list * (string * ((string * sort) list * typ list)) list) (*code equations and certificates*) val assert_eqn: theory -> thm * bool -> thm * bool val assert_abs_eqn: theory -> string option -> thm -> thm * (string * string) type cert val constrain_cert: theory -> sort list -> cert -> cert val conclude_cert: cert -> cert val typargs_deps_of_cert: theory -> cert -> (string * sort) list * (string * typ list) list val equations_of_cert: theory -> cert -> ((string * sort) list * typ) * (((term * string option) list * (term * string option)) * (thm option * bool)) list option val pretty_cert: theory -> cert -> Pretty.T list (*executable code*) type constructors type abs_type val type_interpretation: (string -> theory -> theory) -> theory -> theory val datatype_interpretation: (string * constructors -> theory -> theory) -> theory -> theory val abstype_interpretation: (string * abs_type -> theory -> theory) -> theory -> theory val declare_datatype_global: (string * typ) list -> theory -> theory val declare_datatype_cmd: string list -> theory -> theory val declare_abstype: thm -> local_theory -> local_theory val declare_abstype_global: thm -> theory -> theory val declare_default_eqns: (thm * bool) list -> local_theory -> local_theory val declare_default_eqns_global: (thm * bool) list -> theory -> theory val declare_eqns: (thm * bool) list -> local_theory -> local_theory val declare_eqns_global: (thm * bool) list -> theory -> theory val add_eqn_global: thm * bool -> theory -> theory val del_eqn_global: thm -> theory -> theory val declare_abstract_eqn: thm -> local_theory -> local_theory val declare_abstract_eqn_global: thm -> theory -> theory val declare_aborting_global: string -> theory -> theory val declare_unimplemented_global: string -> theory -> theory val declare_case_global: thm -> theory -> theory val declare_undefined_global: string -> theory -> theory val get_type: theory -> string -> constructors * bool val get_type_of_constr_or_abstr: theory -> string -> (string * bool) option val is_constr: theory -> string -> bool val is_abstr: theory -> string -> bool val get_cert: Proof.context -> ((thm * bool) list -> (thm * bool) list option) list -> string -> cert type case_schema val get_case_schema: theory -> string -> case_schema option val get_case_cong: theory -> string -> thm option val is_undefined: theory -> string -> bool val print_codesetup: theory -> unit end; signature CODE_DATA_ARGS = sig type T val empty: T end; signature CODE_DATA = sig type T val change: theory option -> (T -> T) -> T val change_yield: theory option -> (T -> 'a * T) -> 'a * T end; signature PRIVATE_CODE = sig include CODE val declare_data: Any.T -> serial val change_yield_data: serial * ('a -> Any.T) * (Any.T -> 'a) -> theory -> ('a -> 'b * 'a) -> 'b * 'a end; structure Code : PRIVATE_CODE = struct (** auxiliary **) (* printing *) fun string_of_typ thy = Syntax.string_of_typ (Config.put show_sorts true (Syntax.init_pretty_global thy)); fun string_of_const thy c = let val ctxt = Proof_Context.init_global thy in case Axclass.inst_of_param thy c of SOME (c, tyco) => Proof_Context.extern_const ctxt c ^ " " ^ enclose "[" "]" (Proof_Context.extern_type ctxt tyco) | NONE => Proof_Context.extern_const ctxt c end; (* constants *) fun const_typ thy = Type.strip_sorts o Sign.the_const_type thy; fun args_number thy = length o binder_types o const_typ thy; fun devarify ty = let val tys = fold_atyps (fn TVar vi_sort => AList.update (op =) vi_sort) ty []; val vs = Name.invent Name.context Name.aT (length tys); val mapping = map2 (fn v => fn (vi, sort) => (vi, TFree (v, sort))) vs tys; in Term.typ_subst_TVars mapping ty end; fun typscheme thy (c, ty) = (map dest_TFree (Sign.const_typargs thy (c, ty)), Type.strip_sorts ty); fun typscheme_equiv (ty1, ty2) = Type.raw_instance (devarify ty1, ty2) andalso Type.raw_instance (devarify ty2, ty1); fun check_bare_const thy t = case try dest_Const t of SOME c_ty => c_ty | NONE => error ("Not a constant: " ^ Syntax.string_of_term_global thy t); fun check_unoverload thy (c, ty) = let val c' = Axclass.unoverload_const thy (c, ty); val ty_decl = const_typ thy c'; in if typscheme_equiv (ty_decl, Logic.varifyT_global ty) then c' else error ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_const thy = check_unoverload thy o check_bare_const thy; fun read_bare_const thy = check_bare_const thy o Syntax.read_term_global thy; fun read_const thy = check_unoverload thy o read_bare_const thy; (** executable specifications **) (* types *) datatype type_spec = Constructors of { constructors: (string * ((string * sort) list * typ list)) list, case_combinators: string list} | Abstractor of { abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string, more_abstract_functions: string list}; fun concrete_constructors_of (Constructors {constructors, ...}) = constructors | concrete_constructors_of _ = []; fun constructors_of (Constructors {constructors, ...}) = (constructors, false) | constructors_of (Abstractor {abstractor = (co, (vs, ty)), ...}) = ([(co, (vs, [ty]))], true); fun case_combinators_of (Constructors {case_combinators, ...}) = case_combinators | case_combinators_of (Abstractor _) = []; fun add_case_combinator c (vs, Constructors {constructors, case_combinators}) = (vs, Constructors {constructors = constructors, case_combinators = insert (op =) c case_combinators}); fun projection_of (Constructors _) = NONE | projection_of (Abstractor {projection, ...}) = SOME projection; fun abstract_functions_of (Constructors _) = [] | abstract_functions_of (Abstractor {more_abstract_functions, projection, ...}) = projection :: more_abstract_functions; fun add_abstract_function c (vs, Abstractor {abs_rep, abstractor, projection, more_abstract_functions}) = (vs, Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = insert (op =) c more_abstract_functions}); fun join_same_types' (Constructors {constructors, case_combinators = case_combinators1}, Constructors {case_combinators = case_combinators2, ...}) = Constructors {constructors = constructors, case_combinators = merge (op =) (case_combinators1, case_combinators2)} | join_same_types' (Abstractor {abs_rep, abstractor, projection, more_abstract_functions = more_abstract_functions1}, Abstractor {more_abstract_functions = more_abstract_functions2, ...}) = Abstractor {abs_rep = abs_rep, abstractor = abstractor, projection = projection, more_abstract_functions = merge (op =) (more_abstract_functions1, more_abstract_functions2)}; fun join_same_types ((vs, spec1), (_, spec2)) = (vs, join_same_types' (spec1, spec2)); (* functions *) datatype fun_spec = Eqns of bool * (thm * bool) list | Proj of term * (string * string) | Abstr of thm * (string * string); val unimplemented = Eqns (true, []); fun is_unimplemented (Eqns (true, [])) = true | is_unimplemented _ = false; fun is_default (Eqns (true, _)) = true | is_default _ = false; val aborting = Eqns (false, []); fun associated_abstype (Proj (_, tyco_abs)) = SOME tyco_abs | associated_abstype (Abstr (_, tyco_abs)) = SOME tyco_abs | associated_abstype _ = NONE; (* cases *) type case_schema = int * (int * string option list); datatype case_spec = No_Case | Case of {schema: case_schema, tycos: string list, cong: thm} | Undefined; fun associated_datatypes (Case {tycos, schema = (_, (_, raw_cos)), ...}) = (tycos, map_filter I raw_cos) | associated_datatypes _ = ([], []); (** background theory data store **) (* historized declaration data *) structure History = struct type 'a T = { entry: 'a, suppressed: bool, (*incompatible entries are merely suppressed after theory merge but sustain*) history: serial list (*explicit trace of declaration history supports non-monotonic declarations*) } Symtab.table; fun some_entry (SOME {suppressed = false, entry, ...}) = SOME entry | some_entry _ = NONE; fun lookup table = Symtab.lookup table #> some_entry; fun register key entry table = if is_some (Symtab.lookup table key) then Symtab.map_entry key (fn {history, ...} => {entry = entry, suppressed = false, history = serial () :: history}) table else Symtab.update (key, {entry = entry, suppressed = false, history = [serial ()]}) table; fun modify_entry key f = Symtab.map_entry key (fn {entry, suppressed, history} => {entry = f entry, suppressed = suppressed, history = history}); fun all table = Symtab.dest table |> map_filter (fn (key, {entry, suppressed = false, ...}) => SOME (key, entry) | _ => NONE); local fun merge_history join_same ({entry = entry1, history = history1, ...}, {entry = entry2, history = history2, ...}) = let val history = merge (op =) (history1, history2); val entry = if hd history1 = hd history2 then join_same (entry1, entry2) else if hd history = hd history1 then entry1 else entry2; in {entry = entry, suppressed = false, history = history} end; in fun join join_same tables = Symtab.join (K (merge_history join_same)) tables; fun suppress key = Symtab.map_entry key (fn {entry, history, ...} => {entry = entry, suppressed = true, history = history}); fun suppress_except f = Symtab.map (fn key => fn {entry, suppressed, history} => {entry = entry, suppressed = suppressed orelse (not o f) (key, entry), history = history}); end; end; datatype specs = Specs of { types: ((string * sort) list * type_spec) History.T, pending_eqns: (thm * bool) list Symtab.table, functions: fun_spec History.T, cases: case_spec History.T }; fun types_of (Specs {types, ...}) = types; fun pending_eqns_of (Specs {pending_eqns, ...}) = pending_eqns; fun functions_of (Specs {functions, ...}) = functions; fun cases_of (Specs {cases, ...}) = cases; fun make_specs (types, ((pending_eqns, functions), cases)) = Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}; val empty_specs = make_specs (Symtab.empty, ((Symtab.empty, Symtab.empty), Symtab.empty)); fun map_specs f (Specs {types = types, pending_eqns = pending_eqns, functions = functions, cases = cases}) = make_specs (f (types, ((pending_eqns, functions), cases))); fun merge_specs (Specs {types = types1, pending_eqns = _, functions = functions1, cases = cases1}, Specs {types = types2, pending_eqns = _, functions = functions2, cases = cases2}) = let val types = History.join join_same_types (types1, types2); val all_types = map (snd o snd) (History.all types); fun check_abstype (c, fun_spec) = case associated_abstype fun_spec of NONE => true | SOME (tyco, abs) => (case History.lookup types tyco of NONE => false | SOME (_, Constructors _) => false | SOME (_, Abstractor {abstractor = (abs', _), projection, more_abstract_functions, ...}) => abs = abs' andalso (c = projection orelse member (op =) more_abstract_functions c)); fun check_datatypes (_, case_spec) = let val (tycos, required_constructors) = associated_datatypes case_spec; val allowed_constructors = tycos |> maps (these o Option.map (concrete_constructors_of o snd) o History.lookup types) |> map fst; in subset (op =) (required_constructors, allowed_constructors) end; val all_constructors = maps (fst o constructors_of) all_types; val functions = History.join fst (functions1, functions2) |> fold (History.suppress o fst) all_constructors |> History.suppress_except check_abstype; val cases = History.join fst (cases1, cases2) |> History.suppress_except check_datatypes; in make_specs (types, ((Symtab.empty, functions), cases)) end; val map_types = map_specs o apfst; val map_pending_eqns = map_specs o apsnd o apfst o apfst; val map_functions = map_specs o apsnd o apfst o apsnd; val map_cases = map_specs o apsnd o apsnd; (* data slots dependent on executable code *) (*private copy avoids potential conflict of table exceptions*) structure Datatab = Table(type key = int val ord = int_ord); local type kind = {empty: Any.T}; val kinds = Synchronized.var "Code_Data" (Datatab.empty: kind Datatab.table); fun invoke f k = (case Datatab.lookup (Synchronized.value kinds) k of SOME kind => f kind | NONE => raise Fail "Invalid code data identifier"); in fun declare_data empty = let val k = serial (); val kind = {empty = empty}; val _ = Synchronized.change kinds (Datatab.update (k, kind)); in k end; fun invoke_init k = invoke (fn kind => #empty kind) k; end; (*local*) (* global theory store *) local type data = Any.T Datatab.table; fun make_dataref thy = (Context.theory_long_name thy, Synchronized.var "code data" (NONE : (data * Context.theory_id) option)); structure Code_Data = Theory_Data ( type T = specs * (string * (data * Context.theory_id) option Synchronized.var); val empty = (empty_specs, make_dataref (Context.the_global_context ())); val extend = I; fun merge ((specs1, dataref), (specs2, _)) = (merge_specs (specs1, specs2), dataref); ); fun init_dataref thy = if #1 (#2 (Code_Data.get thy)) = Context.theory_long_name thy then NONE else SOME ((Code_Data.map o apsnd) (fn _ => make_dataref thy) thy) in val _ = Theory.setup (Theory.at_begin init_dataref); (* access to executable specifications *) val specs_of : theory -> specs = fst o Code_Data.get; fun modify_specs f thy = Code_Data.map (fn (specs, _) => (f specs, make_dataref thy)) thy; (* access to data dependent on executable specifications *) fun change_yield_data (kind, mk, dest) theory f = let val dataref = #2 (#2 (Code_Data.get theory)); val (datatab, thy_id) = case Synchronized.value dataref of SOME (datatab, thy_id) => if Context.eq_thy_id (Context.theory_id theory, thy_id) then (datatab, thy_id) else (Datatab.empty, Context.theory_id theory) | NONE => (Datatab.empty, Context.theory_id theory) val data = case Datatab.lookup datatab kind of SOME data => data | NONE => invoke_init kind; val result as (_, data') = f (dest data); val _ = Synchronized.change dataref ((K o SOME) (Datatab.update (kind, mk data') datatab, thy_id)); in result end; end; (*local*) (* pending function equations *) (* Ideally, *all* equations implementing a functions would be treated as *one* atomic declaration; unfortunately, we cannot implement this: the too-well-established declaration interface are Isar attributes which operate on *one* single theorem. Hence we treat such Isar declarations as "pending" and historize them as proper declarations at the end of each theory. *) fun modify_pending_eqns c f specs = let val existing_eqns = case History.lookup (functions_of specs) c of SOME (Eqns (false, eqns)) => eqns | _ => []; in specs |> map_pending_eqns (Symtab.map_default (c, existing_eqns) f) end; fun register_fun_spec c spec = map_pending_eqns (Symtab.delete_safe c) #> map_functions (History.register c spec); fun lookup_fun_spec specs c = case Symtab.lookup (pending_eqns_of specs) c of SOME eqns => Eqns (false, eqns) | NONE => (case History.lookup (functions_of specs) c of SOME spec => spec | NONE => unimplemented); fun lookup_proper_fun_spec specs c = let val spec = lookup_fun_spec specs c in if is_unimplemented spec then NONE else SOME spec end; fun all_fun_specs specs = map_filter (fn c => Option.map (pair c) (lookup_proper_fun_spec specs c)) (union (op =) ((Symtab.keys o pending_eqns_of) specs) ((Symtab.keys o functions_of) specs)); fun historize_pending_fun_specs thy = let val pending_eqns = (pending_eqns_of o specs_of) thy; in if Symtab.is_empty pending_eqns then NONE else thy |> modify_specs (map_functions (Symtab.fold (fn (c, eqs) => History.register c (Eqns (false, eqs))) pending_eqns) #> map_pending_eqns (K Symtab.empty)) |> SOME end; val _ = Theory.setup (Theory.at_end historize_pending_fun_specs); (** foundation **) (* types *) fun no_constr thy s (c, ty) = error ("Not a datatype constructor:\n" ^ string_of_const thy c ^ " :: " ^ string_of_typ thy ty ^ "\n" ^ enclose "(" ")" s); fun analyze_constructor thy (c, ty) = let val _ = Thm.global_cterm_of thy (Const (c, ty)); val ty_decl = devarify (const_typ thy c); fun last_typ c_ty ty = let val tfrees = Term.add_tfreesT ty []; val (tyco, vs) = (apsnd o map) dest_TFree (dest_Type (body_type ty)) handle TYPE _ => no_constr thy "bad type" c_ty val _ = if tyco = "fun" then no_constr thy "bad type" c_ty else (); val _ = if has_duplicates (eq_fst (op =)) vs then no_constr thy "duplicate type variables in datatype" c_ty else (); val _ = if length tfrees <> length vs then no_constr thy "type variables missing in datatype" c_ty else (); in (tyco, vs) end; val (tyco, _) = last_typ (c, ty) ty_decl; val (_, vs) = last_typ (c, ty) ty; in ((tyco, map snd vs), (c, (map fst vs, ty))) end; fun constrset_of_consts thy consts = let val _ = map (fn (c, _) => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) consts; val raw_constructors = map (analyze_constructor thy) consts; val tyco = case distinct (op =) (map (fst o fst) raw_constructors) of [tyco] => tyco | [] => error "Empty constructor set" | tycos => error ("Different type constructors in constructor set: " ^ commas_quote tycos) val vs = Name.invent Name.context Name.aT (Sign.arity_number thy tyco); fun inst vs' (c, (vs, ty)) = let val the_v = the o AList.lookup (op =) (vs ~~ vs'); val ty' = map_type_tfree (fn (v, _) => TFree (the_v v, [])) ty; val (vs'', ty'') = typscheme thy (c, ty'); in (c, (vs'', binder_types ty'')) end; val constructors = map (inst vs o snd) raw_constructors; in (tyco, (map (rpair []) vs, constructors)) end; fun lookup_vs_type_spec thy = History.lookup ((types_of o specs_of) thy); type constructors = (string * sort) list * (string * ((string * sort) list * typ list)) list; fun get_type thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, type_spec) => apfst (pair vs) (constructors_of type_spec) | NONE => Sign.arity_number thy tyco |> Name.invent Name.context Name.aT |> map (rpair []) |> rpair [] |> rpair false; type abs_type = (string * sort) list * {abs_rep: thm, abstractor: string * ((string * sort) list * typ), projection: string}; fun get_abstype_spec thy tyco = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor {abs_rep, abstractor, projection, ...}) => (vs, {abs_rep = abs_rep, abstractor = abstractor, projection = projection}) | _ => error ("Not an abstract type: " ^ tyco); fun get_type_of_constr_or_abstr thy c = case (body_type o const_typ thy) c of Type (tyco, _) => let val ((_, cos), abstract) = get_type thy tyco in if member (op =) (map fst cos) c then SOME (tyco, abstract) else NONE end | _ => NONE; fun is_constr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, false) => true | _ => false; fun is_abstr thy c = case get_type_of_constr_or_abstr thy c of SOME (_, true) => true | _ => false; (* bare code equations *) (* convention for variables: ?x ?'a for free-floating theorems (e.g. in the data store) ?x 'a for certificates x 'a for final representation of equations *) exception BAD_THM of string; fun bad_thm msg = raise BAD_THM msg; datatype strictness = Silent | Liberal | Strict fun handle_strictness thm_of f strictness thy x = SOME (f x) handle BAD_THM msg => case strictness of Silent => NONE | Liberal => (warning (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); NONE) | Strict => error (msg ^ ", in theorem:\n" ^ Thm.string_of_thm_global thy (thm_of x)); fun is_linear thm = let val (_, args) = (strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of) thm in not (has_duplicates (op =) ((fold o fold_aterms) (fn Var (v, _) => cons v | _ => I) args [])) end; fun check_decl_ty thy (c, ty) = let val ty_decl = const_typ thy c; in if typscheme_equiv (ty_decl, ty) then () else bad_thm ("Type\n" ^ string_of_typ thy ty ^ "\nof constant " ^ quote c ^ "\nis too specific compared to declared type\n" ^ string_of_typ thy ty_decl) end; fun check_eqn thy {allow_nonlinear, allow_consts, allow_pats} thm (lhs, rhs) = let fun vars_of t = fold_aterms (fn Var (v, _) => insert (op =) v | Free _ => bad_thm "Illegal free variable" | _ => I) t []; fun tvars_of t = fold_term_types (fn _ => fold_atyps (fn TVar (v, _) => insert (op =) v | TFree _ => bad_thm "Illegal free type variable")) t []; val lhs_vs = vars_of lhs; val rhs_vs = vars_of rhs; val lhs_tvs = tvars_of lhs; val rhs_tvs = tvars_of rhs; val _ = if null (subtract (op =) lhs_vs rhs_vs) then () else bad_thm "Free variables on right hand side of equation"; val _ = if null (subtract (op =) lhs_tvs rhs_tvs) then () else bad_thm "Free type variables on right hand side of equation"; val (head, args) = strip_comb lhs; val (c, ty) = case head of Const (c_ty as (_, ty)) => (Axclass.unoverload_const thy c_ty, ty) | _ => bad_thm "Equation not headed by constant"; fun check _ (Abs _) = bad_thm "Abstraction on left hand side of equation" | check 0 (Var _) = () | check _ (Var _) = bad_thm "Variable with application on left hand side of equation" | check n (t1 $ t2) = (check (n+1) t1; check 0 t2) | check n (Const (c_ty as (c, ty))) = if allow_pats then let val c' = Axclass.unoverload_const thy c_ty in if n = (length o binder_types) ty then if allow_consts orelse is_constr thy c' then () else bad_thm (quote c ^ " is not a constructor, on left hand side of equation") else bad_thm ("Partially applied constant " ^ quote c ^ " on left hand side of equation") end else bad_thm ("Pattern not allowed here, but constant " ^ quote c ^ " encountered on left hand side of equation") val _ = map (check 0) args; val _ = if allow_nonlinear orelse is_linear thm then () else bad_thm "Duplicate variables on left hand side of equation"; val _ = if (is_none o Axclass.class_of_param thy) c then () else bad_thm "Overloaded constant as head in equation"; val _ = if not (is_constr thy c) then () else bad_thm "Constructor as head in equation"; val _ = if not (is_abstr thy c) then () else bad_thm "Abstractor as head in equation"; val _ = check_decl_ty thy (c, ty); val _ = case strip_type ty of (Type (tyco, _) :: _, _) => (case lookup_vs_type_spec thy tyco of SOME (_, type_spec) => (case projection_of type_spec of SOME proj => if c = proj then bad_thm "Projection as head in equation" else () | _ => ()) | _ => ()) | _ => (); in () end; local fun raw_assert_eqn thy check_patterns (thm, proper) = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val _ = check_eqn thy {allow_nonlinear = not proper, allow_consts = not (proper andalso check_patterns), allow_pats = true} thm (lhs, rhs); in (thm, proper) end; fun raw_assert_abs_eqn thy some_tyco thm = let val (full_lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val (proj_t, lhs) = dest_comb full_lhs handle TERM _ => bad_thm "Not an abstract equation"; val (proj, ty) = dest_Const proj_t handle TERM _ => bad_thm "Not an abstract equation"; val (tyco, Ts) = (dest_Type o domain_type) ty handle TERM _ => bad_thm "Not an abstract equation" | TYPE _ => bad_thm "Not an abstract equation"; val _ = case some_tyco of SOME tyco' => if tyco = tyco' then () else bad_thm ("Abstract type mismatch:" ^ quote tyco ^ " vs. " ^ quote tyco') | NONE => (); val (vs, proj', (abs', _)) = case lookup_vs_type_spec thy tyco of SOME (vs, Abstractor spec) => (vs, #projection spec, #abstractor spec) | _ => bad_thm ("Not an abstract type: " ^ tyco); val _ = if proj = proj' then () else bad_thm ("Projection mismatch: " ^ quote proj ^ " vs. " ^ quote proj'); val _ = check_eqn thy {allow_nonlinear = false, allow_consts = false, allow_pats = false} thm (lhs, rhs); val _ = if ListPair.all (fn (T, (_, sort)) => Sign.of_sort thy (T, sort)) (Ts, vs) then () else error ("Type arguments do not satisfy sort constraints of abstype certificate."); in (thm, (tyco, abs')) end; in fun generic_assert_eqn strictness thy check_patterns eqn = handle_strictness fst (raw_assert_eqn thy check_patterns) strictness thy eqn; fun generic_assert_abs_eqn strictness thy check_patterns thm = handle_strictness I (raw_assert_abs_eqn thy check_patterns) strictness thy thm; end; fun assert_eqn thy = the o generic_assert_eqn Strict thy true; fun assert_abs_eqn thy some_tyco = the o generic_assert_abs_eqn Strict thy some_tyco; val head_eqn = dest_Const o fst o strip_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun const_typ_eqn thy thm = let val (c, ty) = head_eqn thm; val c' = Axclass.unoverload_const thy (c, ty); (*permissive wrt. to overloaded constants!*) in (c', ty) end; fun const_eqn thy = fst o const_typ_eqn thy; fun const_abs_eqn thy = Axclass.unoverload_const thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.plain_prop_of; fun mk_proj tyco vs ty abs rep = let val ty_abs = Type (tyco, map TFree vs); val xarg = Var (("x", 0), ty); in Logic.mk_equals (Const (rep, ty_abs --> ty) $ (Const (abs, ty --> ty_abs) $ xarg), xarg) end; (* technical transformations of code equations *) fun meta_rewrite thy = Local_Defs.meta_rewrite_rule (Proof_Context.init_global thy); fun expand_eta thy k thm = let val (lhs, rhs) = (Logic.dest_equals o Thm.plain_prop_of) thm; val (_, args) = strip_comb lhs; val l = if k = ~1 then (length o fst o strip_abs) rhs else Int.max (0, k - length args); val (raw_vars, _) = Term.strip_abs_eta l rhs; val vars = burrow_fst (Name.variant_list (map (fst o fst) (Term.add_vars lhs []))) raw_vars; fun expand (v, ty) thm = Drule.fun_cong_rule thm (Thm.global_cterm_of thy (Var ((v, 0), ty))); in thm |> fold expand vars |> Conv.fconv_rule Drule.beta_eta_conversion end; fun same_arity thy thms = let val num_args_of = length o snd o strip_comb o fst o Logic.dest_equals; val k = fold (Integer.max o num_args_of o Thm.prop_of) thms 0; in map (expand_eta thy k) thms end; fun mk_desymbolization pre post mk vs = let val names = map (pre o fst o fst) vs |> map (Name.desymbolize (SOME false)) |> Name.variant_list [] |> map post; in map_filter (fn (((v, i), x), v') => if v = v' andalso i = 0 then NONE else SOME (((v, i), x), mk ((v', 0), x))) (vs ~~ names) end; fun desymbolize_tvars thy thms = let val tvs = fold (Term.add_tvars o Thm.prop_of) thms []; val instT = mk_desymbolization (unprefix "'") (prefix "'") (Thm.global_ctyp_of thy o TVar) tvs; - in map (Thm.instantiate (instT, [])) thms end; + in map (Thm.instantiate (TVars.make instT, Vars.empty)) thms end; fun desymbolize_vars thy thm = let val vs = Term.add_vars (Thm.prop_of thm) []; val inst = mk_desymbolization I I (Thm.global_cterm_of thy o Var) vs; - in Thm.instantiate ([], inst) thm end; + in Thm.instantiate (TVars.empty, Vars.make inst) thm end; fun canonize_thms thy = desymbolize_tvars thy #> same_arity thy #> map (desymbolize_vars thy); (* preparation and classification of code equations *) fun prep_eqn strictness thy = apfst (meta_rewrite thy) #> generic_assert_eqn strictness thy false #> Option.map (fn eqn => (const_eqn thy (fst eqn), eqn)); fun prep_eqns strictness thy = map_filter (prep_eqn strictness thy) #> AList.group (op =); fun prep_abs_eqn strictness thy = meta_rewrite thy #> generic_assert_abs_eqn strictness thy NONE #> Option.map (fn abs_eqn => (const_abs_eqn thy (fst abs_eqn), abs_eqn)); fun prep_maybe_abs_eqn thy raw_thm = let val thm = meta_rewrite thy raw_thm; val some_abs_thm = generic_assert_abs_eqn Silent thy NONE thm; in case some_abs_thm of SOME (thm, tyco) => SOME (const_abs_eqn thy thm, ((thm, true), SOME tyco)) | NONE => generic_assert_eqn Liberal thy false (thm, false) |> Option.map (fn (thm, _) => (const_eqn thy thm, ((thm, is_linear thm), NONE))) end; (* abstype certificates *) local fun raw_abstype_cert thy proto_thm = let val thm = (Axclass.unoverload (Proof_Context.init_global thy) o meta_rewrite thy) proto_thm; val (lhs, rhs) = Logic.dest_equals (Thm.plain_prop_of thm) handle TERM _ => bad_thm "Not an equation" | THM _ => bad_thm "Not a proper equation"; val ((abs, raw_ty), ((rep, rep_ty), param)) = (apsnd (apfst dest_Const o dest_comb) o apfst dest_Const o dest_comb) lhs handle TERM _ => bad_thm "Not an abstype certificate"; val _ = apply2 (fn c => if (is_some o Axclass.class_of_param thy) c then error ("Is a class parameter: " ^ string_of_const thy c) else ()) (abs, rep); val _ = check_decl_ty thy (abs, raw_ty); val _ = check_decl_ty thy (rep, rep_ty); val _ = if length (binder_types raw_ty) = 1 then () else bad_thm "Bad type for abstract constructor"; val _ = (fst o dest_Var) param handle TERM _ => bad_thm "Not an abstype certificate"; val _ = if param = rhs then () else bad_thm "Not an abstype certificate"; val ((tyco, sorts), (abs, (vs, ty'))) = analyze_constructor thy (abs, devarify raw_ty); val ty = domain_type ty'; val (vs', _) = typscheme thy (abs, ty'); in (tyco, (vs ~~ sorts, ((abs, (vs', ty)), (rep, thm)))) end; in fun check_abstype_cert strictness thy proto_thm = handle_strictness I (raw_abstype_cert thy) strictness thy proto_thm; end; (* code equation certificates *) fun build_head thy (c, ty) = Thm.global_cterm_of thy (Logic.mk_equals (Free ("HEAD", ty), Const (c, ty))); fun get_head thy cert_thm = let val [head] = Thm.chyps_of cert_thm; val (_, Const (c, ty)) = (Logic.dest_equals o Thm.term_of) head; in (typscheme thy (c, ty), head) end; fun typscheme_projection thy = typscheme thy o dest_Const o fst o dest_comb o fst o Logic.dest_equals; fun typscheme_abs thy = typscheme thy o dest_Const o fst o strip_comb o snd o dest_comb o fst o Logic.dest_equals o Thm.prop_of; fun constrain_thm thy vs sorts thm = let val mapping = map2 (fn (v, sort) => fn sort' => (v, Sorts.inter_sort (Sign.classes_of thy) (sort, sort'))) vs sorts; - val inst = map2 (fn (v, sort) => fn (_, sort') => - (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping; + val instT = + TVars.build + (fold2 (fn (v, sort) => fn (_, sort') => + TVars.add (((v, 0), sort), Thm.global_ctyp_of thy (TFree (v, sort')))) vs mapping); val subst = (Term.map_types o map_type_tfree) (fn (v, _) => TFree (v, the (AList.lookup (op =) mapping v))); in thm |> Thm.varifyT_global - |> Thm.instantiate (inst, []) + |> Thm.instantiate (instT, Vars.empty) |> pair subst end; fun concretify_abs thy tyco abs_thm = let val (_, {abstractor = (c_abs, _), abs_rep, ...}) = get_abstype_spec thy tyco; val lhs = (fst o Logic.dest_equals o Thm.prop_of) abs_thm val ty = fastype_of lhs; val ty_abs = (fastype_of o snd o dest_comb) lhs; val abs = Thm.global_cterm_of thy (Const (c_abs, ty --> ty_abs)); val raw_concrete_thm = Drule.transitive_thm OF [Thm.symmetric abs_rep, Thm.combination (Thm.reflexive abs) abs_thm]; in (c_abs, (Thm.varifyT_global o zero_var_indexes) raw_concrete_thm) end; fun add_rhss_of_eqn thy t = let val (args, rhs) = (apfst (snd o strip_comb) o Logic.dest_equals) t; fun add_const (Const (c, ty)) = insert (op =) (c, Sign.const_typargs thy (c, ty)) | add_const _ = I val add_consts = fold_aterms add_const in add_consts rhs o fold add_consts args end; val dest_eqn = apfst (snd o strip_comb) o Logic.dest_equals o Logic.unvarify_global; abstype cert = Nothing of thm | Equations of thm * bool list | Projection of term * string | Abstract of thm * string with fun dummy_thm ctxt c = let val thy = Proof_Context.theory_of ctxt; val raw_ty = devarify (const_typ thy c); val (vs, _) = typscheme thy (c, raw_ty); val sortargs = case Axclass.class_of_param thy c of SOME class => [[class]] | NONE => (case get_type_of_constr_or_abstr thy c of SOME (tyco, _) => (map snd o fst o the) (AList.lookup (op =) ((snd o fst o get_type thy) tyco) c) | NONE => replicate (length vs) []); val the_sort = the o AList.lookup (op =) (map fst vs ~~ sortargs); val ty = map_type_tfree (fn (v, _) => TFree (v, the_sort v)) raw_ty val chead = build_head thy (c, ty); in Thm.weaken chead Drule.dummy_thm end; fun nothing_cert ctxt c = Nothing (dummy_thm ctxt c); fun cert_of_eqns ctxt c [] = Equations (dummy_thm ctxt c, []) | cert_of_eqns ctxt c raw_eqns = let val thy = Proof_Context.theory_of ctxt; val eqns = burrow_fst (canonize_thms thy) raw_eqns; val _ = map (assert_eqn thy) eqns; val (thms, propers) = split_list eqns; val _ = map (fn thm => if c = const_eqn thy thm then () else error ("Wrong head of code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy thm)) thms; val tvars_of = build_rev o Term.add_tvarsT; val vss = map (tvars_of o snd o head_eqn) thms; fun inter_sorts vs = fold (curry (Sorts.inter_sort (Sign.classes_of thy)) o snd) vs []; val sorts = map_transpose inter_sorts vss; val vts = Name.invent_names Name.context Name.aT sorts; - val thms' = - map2 (fn vs => Thm.instantiate (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts, [])) vss thms; + fun instantiate vs = + Thm.instantiate (TVars.make (vs ~~ map (Thm.ctyp_of ctxt o TFree) vts), Vars.empty); + val thms' = map2 instantiate vss thms; val head_thm = Thm.symmetric (Thm.assume (build_head thy (head_eqn (hd thms')))); fun head_conv ct = if can Thm.dest_comb ct then Conv.fun_conv head_conv ct else Conv.rewr_conv head_thm ct; val rewrite_head = Conv.fconv_rule (Conv.arg1_conv head_conv); val cert_thm = Conjunction.intr_balanced (map rewrite_head thms'); in Equations (cert_thm, propers) end; fun cert_of_proj ctxt proj tyco = let val thy = Proof_Context.theory_of ctxt val (vs, {abstractor = (abs, (_, ty)), projection = proj', ...}) = get_abstype_spec thy tyco; val _ = if proj = proj' then () else error ("Wrong head of projection,\nexpected constant " ^ string_of_const thy proj); in Projection (mk_proj tyco vs ty abs proj, tyco) end; fun cert_of_abs ctxt tyco c raw_abs_thm = let val thy = Proof_Context.theory_of ctxt; val abs_thm = singleton (canonize_thms thy) raw_abs_thm; val _ = assert_abs_eqn thy (SOME tyco) abs_thm; val _ = if c = const_abs_eqn thy abs_thm then () else error ("Wrong head of abstract code equation,\nexpected constant " ^ string_of_const thy c ^ "\n" ^ Thm.string_of_thm_global thy abs_thm); in Abstract (Thm.legacy_freezeT abs_thm, tyco) end; fun constrain_cert_thm thy sorts cert_thm = let val ((vs, _), head) = get_head thy cert_thm; val (subst, cert_thm') = cert_thm |> Thm.implies_intr head |> constrain_thm thy vs sorts; val head' = Thm.term_of head |> subst |> Thm.global_cterm_of thy; val cert_thm'' = cert_thm' |> Thm.elim_implies (Thm.assume head'); in cert_thm'' end; fun constrain_cert thy sorts (Nothing cert_thm) = Nothing (constrain_cert_thm thy sorts cert_thm) | constrain_cert thy sorts (Equations (cert_thm, propers)) = Equations (constrain_cert_thm thy sorts cert_thm, propers) | constrain_cert _ _ (cert as Projection _) = cert | constrain_cert thy sorts (Abstract (abs_thm, tyco)) = Abstract (snd (constrain_thm thy (fst (typscheme_abs thy abs_thm)) sorts abs_thm), tyco); fun conclude_cert (Nothing cert_thm) = Nothing (Thm.close_derivation \<^here> cert_thm) | conclude_cert (Equations (cert_thm, propers)) = Equations (Thm.close_derivation \<^here> cert_thm, propers) | conclude_cert (cert as Projection _) = cert | conclude_cert (Abstract (abs_thm, tyco)) = Abstract (Thm.close_derivation \<^here> abs_thm, tyco); fun typscheme_of_cert thy (Nothing cert_thm) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Equations (cert_thm, _)) = fst (get_head thy cert_thm) | typscheme_of_cert thy (Projection (proj, _)) = typscheme_projection thy proj | typscheme_of_cert thy (Abstract (abs_thm, _)) = typscheme_abs thy abs_thm; fun typargs_deps_of_cert thy (Nothing cert_thm) = let val vs = (fst o fst) (get_head thy cert_thm); in (vs, []) end | typargs_deps_of_cert thy (Equations (cert_thm, propers)) = let val vs = (fst o fst) (get_head thy cert_thm); val equations = if null propers then [] else Thm.prop_of cert_thm |> Logic.dest_conjunction_balanced (length propers); in (vs, fold (add_rhss_of_eqn thy) equations []) end | typargs_deps_of_cert thy (Projection (t, _)) = (fst (typscheme_projection thy t), add_rhss_of_eqn thy t []) | typargs_deps_of_cert thy (Abstract (abs_thm, tyco)) = let val vs = fst (typscheme_abs thy abs_thm); val (_, concrete_thm) = concretify_abs thy tyco abs_thm; in (vs, add_rhss_of_eqn thy (Logic.unvarify_types_global (Thm.prop_of concrete_thm)) []) end; fun equations_of_cert thy (cert as Nothing _) = (typscheme_of_cert thy cert, NONE) | equations_of_cert thy (cert as Equations (cert_thm, propers)) = let val tyscm = typscheme_of_cert thy cert; val thms = if null propers then [] else cert_thm |> Local_Defs.expand [snd (get_head thy cert_thm)] |> Thm.varifyT_global |> Conjunction.elim_balanced (length propers); fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, NONE)); in (tyscm, SOME (map (abstractions o dest_eqn o Thm.prop_of) thms ~~ (map SOME thms ~~ propers))) end | equations_of_cert thy (Projection (t, tyco)) = let val (_, {abstractor = (abs, _), ...}) = get_abstype_spec thy tyco; val tyscm = typscheme_projection thy t; val t' = Logic.varify_types_global t; fun abstractions (args, rhs) = (map (rpair (SOME abs)) args, (rhs, NONE)); in (tyscm, SOME [((abstractions o dest_eqn) t', (NONE, true))]) end | equations_of_cert thy (Abstract (abs_thm, tyco)) = let val tyscm = typscheme_abs thy abs_thm; val (abs, concrete_thm) = concretify_abs thy tyco abs_thm; fun abstractions (args, rhs) = (map (rpair NONE) args, (rhs, (SOME abs))); in (tyscm, SOME [((abstractions o dest_eqn o Thm.prop_of) concrete_thm, (SOME (Thm.varifyT_global abs_thm), true))]) end; fun pretty_cert _ (Nothing _) = [] | pretty_cert thy (cert as Equations _) = (map_filter (Option.map (Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy)) o fst o snd) o these o snd o equations_of_cert thy) cert | pretty_cert thy (Projection (t, _)) = [Syntax.pretty_term_global thy (Logic.varify_types_global t)] | pretty_cert thy (Abstract (abs_thm, _)) = [(Thm.pretty_thm_global thy o Axclass.overload (Proof_Context.init_global thy) o Thm.varifyT_global) abs_thm]; end; (* code certificate access with preprocessing *) fun eqn_conv conv ct = let fun lhs_conv ct = if can Thm.dest_comb ct then Conv.combination_conv lhs_conv conv ct else Conv.all_conv ct; in Conv.combination_conv (Conv.arg_conv lhs_conv) conv ct end; fun rewrite_eqn conv ctxt = singleton (Variable.trade (K (map (Conv.fconv_rule (conv (Simplifier.rewrite ctxt))))) ctxt) fun preprocess conv ctxt = Thm.transfer' ctxt #> rewrite_eqn conv ctxt #> Axclass.unoverload ctxt; fun cert_of_eqns_preprocess ctxt functrans c = let fun trace_eqns s eqns = (Pretty.writeln o Pretty.chunks) (Pretty.str s :: map (Thm.pretty_thm ctxt o fst) eqns); val tracing = if Config.get ctxt simp_trace then trace_eqns else (K o K) (); in tap (tracing "before function transformation") #> (perhaps o perhaps_loop o perhaps_apply) functrans #> tap (tracing "after function transformation") #> (map o apfst) (preprocess eqn_conv ctxt) #> cert_of_eqns ctxt c end; fun get_cert ctxt functrans c = case lookup_proper_fun_spec (specs_of (Proof_Context.theory_of ctxt)) c of NONE => nothing_cert ctxt c | SOME (Eqns (_, eqns)) => eqns |> cert_of_eqns_preprocess ctxt functrans c | SOME (Proj (_, (tyco, _))) => cert_of_proj ctxt c tyco | SOME (Abstr (abs_thm, (tyco, _))) => abs_thm |> preprocess Conv.arg_conv ctxt |> cert_of_abs ctxt tyco c; (* case certificates *) local fun raw_case_cert thm = let val ((head, raw_case_expr), cases) = (apfst Logic.dest_equals o apsnd Logic.dest_conjunctions o Logic.dest_implies o Thm.plain_prop_of) thm; val _ = case head of Free _ => () | Var _ => () | _ => raise TERM ("case_cert", []); val ([(case_var, _)], case_expr) = Term.strip_abs_eta 1 raw_case_expr; val (Const (case_const, _), raw_params) = strip_comb case_expr; val n = find_index (fn Free (v, _) => v = case_var | _ => false) raw_params; val _ = if n = ~1 then raise TERM ("case_cert", []) else (); val params = map (fst o dest_Var) (nth_drop n raw_params); fun dest_case t = let val (head' $ t_co, rhs) = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val (Const (co, _), args) = strip_comb t_co; val (Var (param, _), args') = strip_comb rhs; val _ = if args' = args then () else raise TERM ("case_cert", []); in (param, co) end; fun analyze_cases cases = let val co_list = fold (AList.update (op =) o dest_case) cases []; in map (AList.lookup (op =) co_list) params end; fun analyze_let t = let val (head' $ arg, Var (param', _) $ arg') = Logic.dest_equals t; val _ = if head' = head then () else raise TERM ("case_cert", []); val _ = if arg' = arg then () else raise TERM ("case_cert", []); val _ = if [param'] = params then () else raise TERM ("case_cert", []); in [] end; fun analyze (cases as [let_case]) = (analyze_cases cases handle Bind => analyze_let let_case) | analyze cases = analyze_cases cases; in (case_const, (n, analyze cases)) end; in fun case_cert thm = raw_case_cert thm handle Bind => error "bad case certificate" | TERM _ => error "bad case certificate"; end; fun lookup_case_spec thy = History.lookup ((cases_of o specs_of) thy); fun get_case_schema thy c = case lookup_case_spec thy c of SOME (Case {schema, ...}) => SOME schema | _ => NONE; fun get_case_cong thy c = case lookup_case_spec thy c of SOME (Case {cong, ...}) => SOME cong | _ => NONE; fun is_undefined thy c = case lookup_case_spec thy c of SOME Undefined => true | _ => false; (* diagnostic *) fun print_codesetup thy = let val ctxt = Proof_Context.init_global thy; val specs = specs_of thy; fun pretty_equations const thms = (Pretty.block o Pretty.fbreaks) (Pretty.str (string_of_const thy const) :: map (Thm.pretty_thm_item ctxt) thms); fun pretty_function (const, Eqns (_, eqns)) = pretty_equations const (map fst eqns) | pretty_function (const, Proj (proj, _)) = Pretty.block [Pretty.str (string_of_const thy const), Pretty.fbrk, Syntax.pretty_term ctxt proj] | pretty_function (const, Abstr (thm, _)) = pretty_equations const [thm]; fun pretty_typ (tyco, vs) = Pretty.str (string_of_typ thy (Type (tyco, map TFree vs))); fun pretty_type_spec (typ, (cos, abstract)) = if null cos then pretty_typ typ else (Pretty.block o Pretty.breaks) ( pretty_typ typ :: Pretty.str "=" :: (if abstract then [Pretty.str "(abstract)"] else []) @ separate (Pretty.str "|") (map (fn (c, (_, [])) => Pretty.str (string_of_const thy c) | (c, (_, tys)) => (Pretty.block o Pretty.breaks) (Pretty.str (string_of_const thy c) :: Pretty.str "of" :: map (Pretty.quote o Syntax.pretty_typ_global thy) tys)) cos) ); fun pretty_case_param NONE = "" | pretty_case_param (SOME c) = string_of_const thy c fun pretty_case (const, Case {schema = (_, (_, [])), ...}) = Pretty.str (string_of_const thy const) | pretty_case (const, Case {schema = (_, (_, cos)), ...}) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str "with", (Pretty.block o Pretty.commas o map (Pretty.str o pretty_case_param)) cos] | pretty_case (const, Undefined) = (Pretty.block o Pretty.breaks) [ Pretty.str (string_of_const thy const), Pretty.str ""]; val functions = all_fun_specs specs |> sort (string_ord o apply2 fst); val types = History.all (types_of specs) |> map (fn (tyco, (vs, spec)) => ((tyco, vs), constructors_of spec)) |> sort (string_ord o apply2 (fst o fst)); val cases = History.all (cases_of specs) |> filter (fn (_, No_Case) => false | _ => true) |> sort (string_ord o apply2 fst); in Pretty.writeln_chunks [ Pretty.block ( Pretty.str "types:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_type_spec) types ), Pretty.block ( Pretty.str "functions:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_function) functions ), Pretty.block ( Pretty.str "cases:" :: Pretty.fbrk :: (Pretty.fbreaks o map pretty_case) cases ) ] end; (** declaration of executable ingredients **) (* plugins for dependent applications *) structure Codetype_Plugin = Plugin(type T = string); val codetype_plugin = Plugin_Name.declare_setup \<^binding>\codetype\; fun type_interpretation f = Codetype_Plugin.interpretation codetype_plugin (fn tyco => Local_Theory.background_theory (fn thy => thy |> Sign.root_path |> Sign.add_path (Long_Name.qualifier tyco) |> f tyco |> Sign.restore_naming thy)); fun datatype_interpretation f = type_interpretation (fn tyco => fn thy => case get_type thy tyco of (spec, false) => f (tyco, spec) thy | (_, true) => thy ); fun abstype_interpretation f = type_interpretation (fn tyco => fn thy => case try (get_abstype_spec thy) tyco of SOME spec => f (tyco, spec) thy | NONE => thy ); fun register_tyco_for_plugin tyco = Named_Target.theory_map (Codetype_Plugin.data_default tyco); (* abstract code declarations *) local fun generic_code_declaration strictness lift_phi f x = Local_Theory.declaration {syntax = false, pervasive = false} (fn phi => Context.mapping (f strictness (lift_phi phi x)) I); in fun silent_code_declaration lift_phi = generic_code_declaration Silent lift_phi; fun code_declaration lift_phi = generic_code_declaration Liberal lift_phi; end; (* types *) fun invalidate_constructors_of (_, type_spec) = fold (fn (c, _) => History.register c unimplemented) (fst (constructors_of type_spec)); fun invalidate_abstract_functions_of (_, type_spec) = fold (fn c => History.register c unimplemented) (abstract_functions_of type_spec); fun invalidate_case_combinators_of (_, type_spec) = fold (fn c => History.register c No_Case) (case_combinators_of type_spec); fun register_type (tyco, vs_typ_spec) specs = let val olds = the_list (History.lookup (types_of specs) tyco); in specs |> map_functions (fold invalidate_abstract_functions_of olds #> invalidate_constructors_of vs_typ_spec) |> map_cases (fold invalidate_case_combinators_of olds) |> map_types (History.register tyco vs_typ_spec) end; fun declare_datatype_global proto_constrs thy = let fun unoverload_const_typ (c, ty) = (Axclass.unoverload_const thy (c, ty), ty); val constrs = map unoverload_const_typ proto_constrs; val (tyco, (vs, cos)) = constrset_of_consts thy constrs; in thy |> modify_specs (register_type (tyco, (vs, Constructors {constructors = cos, case_combinators = []}))) |> register_tyco_for_plugin tyco end; fun declare_datatype_cmd raw_constrs thy = declare_datatype_global (map (read_bare_const thy) raw_constrs) thy; fun generic_declare_abstype strictness proto_thm thy = case check_abstype_cert strictness thy proto_thm of SOME (tyco, (vs, (abstractor as (abs, (_, ty)), (proj, abs_rep)))) => thy |> modify_specs (register_type (tyco, (vs, Abstractor {abstractor = abstractor, projection = proj, abs_rep = abs_rep, more_abstract_functions = []})) #> register_fun_spec proj (Proj (Logic.varify_types_global (mk_proj tyco vs ty abs proj), (tyco, abs)))) |> register_tyco_for_plugin tyco | NONE => thy; val declare_abstype_global = generic_declare_abstype Strict; val declare_abstype = code_declaration Morphism.thm generic_declare_abstype; (* functions *) (* strictness wrt. shape of theorem propositions: * default equations: silent * using declarations and attributes: warnings (after morphism application!) * using global declarations (... -> thy -> thy): strict * internal processing after storage: strict *) local fun subsumptive_add thy verbose (thm, proper) eqns = let val args_of = drop_prefix is_Var o rev o snd o strip_comb o Term.map_types Type.strip_sorts o fst o Logic.dest_equals o Thm.plain_prop_of o Thm.transfer thy; val args = args_of thm; val incr_idx = Logic.incr_indexes ([], [], Thm.maxidx_of thm + 1); fun matches_args args' = let val k = length args' - length args in if k >= 0 then Pattern.matchess thy (args, (map incr_idx o drop k) args') else false end; fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then (if verbose then warning ("Code generator: dropping subsumed code equation\n" ^ Thm.string_of_thm_global thy thm') else (); true) else false; in (thm |> Thm.close_derivation \<^here> |> Thm.trim_context, proper) :: filter_out drop eqns end; fun add_eqn_for (c, eqn) thy = thy |> modify_specs (modify_pending_eqns c (subsumptive_add thy true eqn)); fun add_eqns_for default (c, proto_eqns) thy = thy |> modify_specs (fn specs => if is_default (lookup_fun_spec specs c) orelse not default then let val eqns = [] |> fold_rev (subsumptive_add thy (not default)) proto_eqns; in specs |> register_fun_spec c (Eqns (default, eqns)) end else specs); fun add_abstract_for (c, (thm, tyco_abs as (tyco, _))) = modify_specs (register_fun_spec c (Abstr (Thm.close_derivation \<^here> thm, tyco_abs)) #> map_types (History.modify_entry tyco (add_abstract_function c))) in fun generic_declare_eqns default strictness raw_eqns thy = fold (add_eqns_for default) (prep_eqns strictness thy raw_eqns) thy; fun generic_add_eqn strictness raw_eqn thy = fold add_eqn_for (the_list (prep_eqn strictness thy raw_eqn)) thy; fun generic_declare_abstract_eqn strictness raw_abs_eqn thy = fold add_abstract_for (the_list (prep_abs_eqn strictness thy raw_abs_eqn)) thy; fun add_maybe_abs_eqn_liberal thm thy = case prep_maybe_abs_eqn thy thm of SOME (c, (eqn, NONE)) => add_eqn_for (c, eqn) thy | SOME (c, ((thm, _), SOME tyco)) => add_abstract_for (c, (thm, tyco)) thy | NONE => thy; end; val declare_default_eqns_global = generic_declare_eqns true Silent; val declare_default_eqns = silent_code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns true); val declare_eqns_global = generic_declare_eqns false Strict; val declare_eqns = code_declaration (map o apfst o Morphism.thm) (generic_declare_eqns false); val add_eqn_global = generic_add_eqn Strict; fun del_eqn_global thm thy = case prep_eqn Liberal thy (thm, false) of SOME (c, (thm, _)) => modify_specs (modify_pending_eqns c (filter_out (fn (thm', _) => Thm.eq_thm_prop (thm, thm')))) thy | NONE => thy; val declare_abstract_eqn_global = generic_declare_abstract_eqn Strict; val declare_abstract_eqn = code_declaration Morphism.thm generic_declare_abstract_eqn; fun declare_aborting_global c = modify_specs (register_fun_spec c aborting); fun declare_unimplemented_global c = modify_specs (register_fun_spec c unimplemented); (* cases *) fun case_cong thy case_const (num_args, (pos, _)) = let val ([x, y], ctxt) = fold_map Name.variant ["A", "A'"] Name.context; val (zs, _) = fold_map Name.variant (replicate (num_args - 1) "") ctxt; val (ws, vs) = chop pos zs; val T = devarify (const_typ thy case_const); val Ts = binder_types T; val T_cong = nth Ts pos; fun mk_prem z = Free (z, T_cong); fun mk_concl z = list_comb (Const (case_const, T), map2 (curry Free) (ws @ z :: vs) Ts); val (prem, concl) = apply2 Logic.mk_equals (apply2 mk_prem (x, y), apply2 mk_concl (x, y)); in Goal.prove_sorry_global thy (x :: y :: zs) [prem] concl (fn {context = ctxt', prems} => Simplifier.rewrite_goals_tac ctxt' prems THEN ALLGOALS (Proof_Context.fact_tac ctxt' [Drule.reflexive_thm])) end; fun declare_case_global thm thy = let val (case_const, (k, cos)) = case_cert thm; fun get_type_of_constr c = case get_type_of_constr_or_abstr thy c of SOME (c, false) => SOME c | _ => NONE; val cos_with_tycos = (map_filter o Option.map) (fn c => (c, get_type_of_constr c)) cos; val _ = case map_filter (fn (c, NONE) => SOME c | _ => NONE) cos_with_tycos of [] => () | cs => error ("Non-constructor(s) in case certificate: " ^ commas_quote cs); val tycos = distinct (op =) (map_filter snd cos_with_tycos); val schema = (1 + Int.max (1, length cos), (k, cos)); val cong = case_cong thy case_const schema; in thy |> modify_specs (map_cases (History.register case_const (Case {schema = schema, tycos = tycos, cong = cong})) #> map_types (fold (fn tyco => History.modify_entry tyco (add_case_combinator case_const)) tycos)) end; fun declare_undefined_global c = (modify_specs o map_cases) (History.register c Undefined); (* attributes *) fun code_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); fun code_thm_attribute g f = Scan.lift (g |-- Scan.succeed (code_attribute f)); fun code_const_attribute g f = Scan.lift (g -- Args.colon) |-- Scan.repeat1 Args.term >> (fn ts => code_attribute (K (fold (fn t => fn thy => f ((check_const thy o Logic.unvarify_types_global) t) thy) ts))); val _ = Theory.setup (let val code_attribute_parser = code_thm_attribute (Args.$$$ "equation") (fn thm => generic_add_eqn Liberal (thm, true)) || code_thm_attribute (Args.$$$ "nbe") (fn thm => generic_add_eqn Liberal (thm, false)) || code_thm_attribute (Args.$$$ "abstract") (generic_declare_abstract_eqn Liberal) || code_thm_attribute (Args.$$$ "abstype") (generic_declare_abstype Liberal) || code_thm_attribute Args.del del_eqn_global || code_const_attribute (Args.$$$ "abort") declare_aborting_global || code_const_attribute (Args.$$$ "drop") declare_unimplemented_global || Scan.succeed (code_attribute add_maybe_abs_eqn_liberal); in Attrib.setup \<^binding>\code\ code_attribute_parser "declare theorems for code generation" end); end; (*struct*) (* type-safe interfaces for data dependent on executable code *) functor Code_Data(Data: CODE_DATA_ARGS): CODE_DATA = struct type T = Data.T; exception Data of T; fun dest (Data x) = x val kind = Code.declare_data (Data Data.empty); val data_op = (kind, Data, dest); fun change_yield (SOME thy) f = Code.change_yield_data data_op thy f | change_yield NONE f = f Data.empty fun change some_thy f = snd (change_yield some_thy (pair () o f)); end; structure Code : CODE = struct open Code; end; diff --git a/src/Pure/Isar/element.ML b/src/Pure/Isar/element.ML --- a/src/Pure/Isar/element.ML +++ b/src/Pure/Isar/element.ML @@ -1,473 +1,473 @@ (* Title: Pure/Isar/element.ML Author: Makarius Explicit data structures for some Isar language elements, with derived logical operations. *) signature ELEMENT = sig type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list) type obtains = (string, string) obtain list type obtains_i = (typ, term) obtain list datatype ('typ, 'term) stmt = Shows of (Attrib.binding * ('term * 'term list) list) list | Obtains of ('typ, 'term) obtain list type statement = (string, string) stmt type statement_i = (typ, term) stmt datatype ('typ, 'term, 'fact) ctxt = Fixes of (binding * 'typ option * mixfix) list | Constrains of (string * 'typ) list | Assumes of (Attrib.binding * ('term * 'term list) list) list | Defines of (Attrib.binding * ('term * 'term list)) list | Notes of string * (Attrib.binding * ('fact * Token.src list) list) list | Lazy_Notes of string * (binding * 'fact lazy) type context = (string, string, Facts.ref) ctxt type context_i = (typ, term, thm list) ctxt val map_ctxt: {binding: binding -> binding, typ: 'typ -> 'a, term: 'term -> 'b, pattern: 'term -> 'b, fact: 'fact -> 'c, attrib: Token.src -> Token.src} -> ('typ, 'term, 'fact) ctxt -> ('a, 'b, 'c) ctxt val map_ctxt_attrib: (Token.src -> Token.src) -> ('typ, 'term, 'fact) ctxt -> ('typ, 'term, 'fact) ctxt val transform_ctxt: morphism -> context_i -> context_i val pretty_stmt: Proof.context -> statement_i -> Pretty.T list val pretty_ctxt: Proof.context -> context_i -> Pretty.T list val pretty_ctxt_no_attribs: Proof.context -> context_i -> Pretty.T list val pretty_statement: Proof.context -> string -> thm -> Pretty.T type witness val prove_witness: Proof.context -> term -> tactic -> witness val witness_proof: (witness list list -> Proof.context -> Proof.context) -> term list list -> Proof.context -> Proof.state val witness_proof_eqs: (witness list list -> thm list -> Proof.context -> Proof.context) -> term list list -> term list -> Proof.context -> Proof.state val witness_local_proof: (witness list list -> Proof.state -> Proof.state) -> string -> term list list -> Proof.context -> Proof.state -> Proof.state val witness_local_proof_eqs: (witness list list -> thm list -> Proof.state -> Proof.state) -> string -> term list list -> term list -> Proof.context -> Proof.state -> Proof.state val transform_witness: morphism -> witness -> witness val conclude_witness: Proof.context -> witness -> thm val pretty_witness: Proof.context -> witness -> Pretty.T - val instantiate_normalize_morphism: - ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> morphism + val instantiate_normalize_morphism: ctyp TFrees.table * cterm Frees.table -> morphism val satisfy_morphism: witness list -> morphism val eq_term_morphism: theory -> term list -> morphism option val eq_morphism: theory -> thm list -> morphism option val init: context_i -> Context.generic -> Context.generic val activate_i: context_i -> Proof.context -> context_i * Proof.context val activate: (typ, term, Facts.ref) ctxt -> Proof.context -> context_i * Proof.context end; structure Element: ELEMENT = struct (** language elements **) (* statement *) type ('typ, 'term) obtain = binding * ((binding * 'typ option * mixfix) list * 'term list); type obtains = (string, string) obtain list; type obtains_i = (typ, term) obtain list; datatype ('typ, 'term) stmt = Shows of (Attrib.binding * ('term * 'term list) list) list | Obtains of ('typ, 'term) obtain list; type statement = (string, string) stmt; type statement_i = (typ, term) stmt; (* context *) datatype ('typ, 'term, 'fact) ctxt = Fixes of (binding * 'typ option * mixfix) list | Constrains of (string * 'typ) list | Assumes of (Attrib.binding * ('term * 'term list) list) list | Defines of (Attrib.binding * ('term * 'term list)) list | Notes of string * (Attrib.binding * ('fact * Token.src list) list) list | Lazy_Notes of string * (binding * 'fact lazy); type context = (string, string, Facts.ref) ctxt; type context_i = (typ, term, thm list) ctxt; fun map_ctxt {binding, typ, term, pattern, fact, attrib} = fn Fixes fixes => Fixes (fixes |> map (fn (x, T, mx) => (binding x, Option.map typ T, mx))) | Constrains xs => Constrains (xs |> map (fn (x, T) => (Variable.check_name (binding (Binding.name x)), typ T))) | Assumes asms => Assumes (asms |> map (fn ((a, atts), propps) => ((binding a, map attrib atts), propps |> map (fn (t, ps) => (term t, map pattern ps))))) | Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => ((binding a, map attrib atts), (term t, map pattern ps)))) | Notes (kind, facts) => Notes (kind, facts |> map (fn ((a, atts), bs) => ((binding a, map attrib atts), bs |> map (fn (ths, btts) => (fact ths, map attrib btts))))) | Lazy_Notes (kind, (a, ths)) => Lazy_Notes (kind, (binding a, Lazy.map fact ths)); fun map_ctxt_attrib attrib = map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = I, attrib = attrib}; fun transform_ctxt phi = map_ctxt {binding = Morphism.binding phi, typ = Morphism.typ phi, term = Morphism.term phi, pattern = Morphism.term phi, fact = Morphism.fact phi, attrib = map (Token.transform phi)}; (** pretty printing **) fun pretty_items _ _ [] = [] | pretty_items keyword sep (x :: ys) = Pretty.block [Pretty.keyword2 keyword, Pretty.brk 1, x] :: map (fn y => Pretty.block [Pretty.str " ", Pretty.keyword2 sep, Pretty.brk 1, y]) ys; (* pretty_stmt *) fun pretty_stmt ctxt = let val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; val prt_term = Pretty.quote o Syntax.pretty_term ctxt; val prt_terms = separate (Pretty.keyword2 "and") o map prt_term; val prt_binding = Attrib.pretty_binding ctxt; val prt_name = Proof_Context.pretty_name ctxt; fun prt_show (a, ts) = Pretty.block (Pretty.breaks (prt_binding a ":" @ prt_terms (map fst ts))); fun prt_var (x, SOME T, _) = Pretty.block [prt_name (Binding.name_of x), Pretty.str " ::", Pretty.brk 1, prt_typ T] | prt_var (x, NONE, _) = prt_name (Binding.name_of x); val prt_vars = separate (Pretty.keyword2 "and") o map prt_var; fun prt_obtain (_, ([], props)) = Pretty.block (Pretty.breaks (prt_terms props)) | prt_obtain (_, (vars, props)) = Pretty.block (Pretty.breaks (prt_vars vars @ [Pretty.keyword2 "where"] @ prt_terms props)); in fn Shows shows => pretty_items "shows" "and" (map prt_show shows) | Obtains obtains => pretty_items "obtains" "|" (map prt_obtain obtains) end; (* pretty_ctxt *) fun gen_pretty_ctxt show_attribs ctxt = let val prt_typ = Pretty.quote o Syntax.pretty_typ ctxt; val prt_term = Pretty.quote o Syntax.pretty_term ctxt; val prt_thm = Pretty.cartouche o Thm.pretty_thm ctxt; val prt_name = Proof_Context.pretty_name ctxt; fun prt_binding (b, atts) = Attrib.pretty_binding ctxt (b, if show_attribs then atts else []); fun prt_fact (ths, atts) = if not show_attribs orelse null atts then map prt_thm ths else Pretty.enclose "(" ")" (Pretty.breaks (map prt_thm ths)) :: Attrib.pretty_attribs ctxt atts; fun prt_mixfix NoSyn = [] | prt_mixfix mx = [Pretty.brk 2, Mixfix.pretty_mixfix mx]; fun prt_fix (x, SOME T, mx) = Pretty.block (prt_name (Binding.name_of x) :: Pretty.str " ::" :: Pretty.brk 1 :: prt_typ T :: prt_mixfix mx) | prt_fix (x, NONE, mx) = Pretty.block (prt_name (Binding.name_of x) :: prt_mixfix mx); fun prt_constrain (x, T) = prt_fix (Binding.name x, SOME T, NoSyn); fun prt_asm (a, ts) = Pretty.block (Pretty.breaks (prt_binding a ":" @ map (prt_term o fst) ts)); fun prt_def (a, (t, _)) = Pretty.block (Pretty.breaks (prt_binding a ":" @ [prt_term t])); fun prt_note (a, ths) = Pretty.block (Pretty.breaks (flat (prt_binding a " =" :: map prt_fact ths))); fun notes_kind "" = "notes" | notes_kind kind = "notes " ^ kind; in fn Fixes fixes => pretty_items "fixes" "and" (map prt_fix fixes) | Constrains xs => pretty_items "constrains" "and" (map prt_constrain xs) | Assumes asms => pretty_items "assumes" "and" (map prt_asm asms) | Defines defs => pretty_items "defines" "and" (map prt_def defs) | Notes (kind, facts) => pretty_items (notes_kind kind) "and" (map prt_note facts) | Lazy_Notes (kind, (a, ths)) => pretty_items (notes_kind kind) "and" [prt_note ((a, []), [(Lazy.force ths, [])])] end; val pretty_ctxt = gen_pretty_ctxt true; val pretty_ctxt_no_attribs = gen_pretty_ctxt false; (* pretty_statement *) local fun standard_elim ctxt th = (case Object_Logic.elim_concl ctxt th of SOME C => let val thesis = Var ((Auto_Bind.thesisN, Thm.maxidx_of th + 1), fastype_of C); - val th' = Thm.instantiate ([], [(Term.dest_Var C, Thm.cterm_of ctxt thesis)]) th; + val insts = (TVars.empty, Vars.make [(Term.dest_Var C, Thm.cterm_of ctxt thesis)]); + val th' = Thm.instantiate insts th; in (th', true) end | NONE => (th, false)); fun thm_name ctxt kind th prts = let val head = if Thm.has_name_hint th then Pretty.block [Pretty.keyword1 kind, Pretty.brk 1, Proof_Context.pretty_name ctxt (Long_Name.base_name (Thm.get_name_hint th)), Pretty.str ":"] else Pretty.keyword1 kind in Pretty.block (Pretty.fbreaks (head :: prts)) end; fun obtain prop ctxt = let val ((ps, prop'), ctxt') = Variable.focus NONE prop ctxt; fun fix (x, T) = (Binding.name (Variable.revert_fixed ctxt' x), SOME T, NoSyn); val xs = map (fix o #2) ps; val As = Logic.strip_imp_prems prop'; in ((Binding.empty, (xs, As)), ctxt') end; in fun pretty_statement ctxt kind raw_th = let val (th, is_elim) = standard_elim ctxt (Raw_Simplifier.norm_hhf ctxt raw_th); val ((_, [th']), ctxt') = Variable.import true [th] (Variable.set_body true ctxt); val prop = Thm.prop_of th'; val (prems, concl) = Logic.strip_horn prop; val concl_term = Object_Logic.drop_judgment ctxt concl; val (assumes, cases) = chop_suffix (fn prem => is_elim andalso concl aconv Logic.strip_assums_concl prem) prems; val is_thesis = if null cases then K false else fn v => v aconv concl_term; val fixes = rev (fold_aterms (fn v as Free (x, T) => if Variable.is_newly_fixed ctxt' ctxt x andalso not (is_thesis v) then insert (op =) (Variable.revert_fixed ctxt' x, T) else I | _ => I) prop []); in pretty_ctxt ctxt' (Fixes (map (fn (x, T) => (Binding.name x, SOME T, NoSyn)) fixes)) @ pretty_ctxt ctxt' (Assumes (map (fn t => (Binding.empty_atts, [(t, [])])) assumes)) @ (if null cases then pretty_stmt ctxt' (Shows [(Binding.empty_atts, [(concl, [])])]) else let val (clauses, ctxt'') = fold_map obtain cases ctxt' in pretty_stmt ctxt'' (Obtains clauses) end) end |> thm_name ctxt kind raw_th; end; (** logical operations **) (* witnesses -- hypotheses as protected facts *) datatype witness = Witness of term * thm; val mark_witness = Logic.protect; fun witness_prop (Witness (t, _)) = t; fun witness_hyps (Witness (_, th)) = Thm.hyps_of th; fun map_witness f (Witness witn) = Witness (f witn); fun transform_witness phi = map_witness (fn (t, th) => (Morphism.term phi t, Morphism.thm phi th)); fun prove_witness ctxt t tac = Witness (t, Goal.prove ctxt [] [] (mark_witness t) (fn _ => resolve_tac ctxt [Drule.protectI] 1 THEN tac) |> Thm.close_derivation \<^here>); local val refine_witness = Proof.refine_singleton (Method.Basic (fn ctxt => CONTEXT_TACTIC o K (ALLGOALS (CONJUNCTS (ALLGOALS (CONJUNCTS (TRYALL (resolve_tac ctxt [Drule.protectI])))))))); fun gen_witness_proof proof after_qed wit_propss eq_props = let val propss = (map o map) (fn prop => (mark_witness prop, [])) wit_propss @ [map (rpair []) eq_props]; fun after_qed' thmss = let val (wits, eqs) = split_last ((map o map) (Thm.close_derivation \<^here>) thmss); in after_qed ((map2 o map2) (curry Witness) wit_propss wits) eqs end; in proof after_qed' propss #> refine_witness end; fun proof_local cmd goal_ctxt after_qed propp = let fun after_qed' (result_ctxt, results) state' = after_qed (burrow (Proof_Context.export result_ctxt (Proof.context_of state')) results) state'; in Proof.map_context (K goal_ctxt) #> Proof.internal_goal (K (K ())) (Proof_Context.get_mode goal_ctxt) true cmd NONE after_qed' [] [] (map (pair Binding.empty_atts) propp) #> snd end; in fun witness_proof after_qed wit_propss = gen_witness_proof (Proof.theorem NONE) (fn wits => fn _ => after_qed wits) wit_propss []; val witness_proof_eqs = gen_witness_proof (Proof.theorem NONE); fun witness_local_proof after_qed cmd wit_propss goal_ctxt = gen_witness_proof (proof_local cmd goal_ctxt) (fn wits => fn _ => after_qed wits) wit_propss []; fun witness_local_proof_eqs after_qed cmd wit_propss eq_props goal_ctxt = gen_witness_proof (proof_local cmd goal_ctxt) after_qed wit_propss eq_props; end; fun conclude_witness ctxt (Witness (_, th)) = Goal.conclude th |> Raw_Simplifier.norm_hhf_protect ctxt |> Thm.close_derivation \<^here>; fun pretty_witness ctxt witn = let val prt_term = Pretty.quote o Syntax.pretty_term ctxt in Pretty.block (prt_term (witness_prop witn) :: (if Config.get ctxt show_hyps then [Pretty.brk 2, Pretty.list "[" "]" (map prt_term (witness_hyps witn))] else [])) end; (* instantiate frees, with beta normalization *) fun instantiate_normalize_morphism insts = Morphism.instantiate_frees_morphism insts $> Morphism.term_morphism "beta_norm" Envir.beta_norm $> Morphism.thm_morphism "beta_conversion" (Conv.fconv_rule (Thm.beta_conversion true)); (* satisfy hypotheses *) local val norm_term = Envir.beta_eta_contract; val norm_conv = Drule.beta_eta_conversion; val norm_cterm = Thm.rhs_of o norm_conv; fun find_witness witns hyp = (case find_first (fn Witness (t, _) => hyp aconv t) witns of NONE => let val hyp' = norm_term hyp in find_first (fn Witness (t, _) => hyp' aconv norm_term t) witns end | some => some); fun compose_witness (Witness (_, th)) r = let val th' = Goal.conclude th; val A = Thm.cprem_of r 1; in Thm.implies_elim (Conv.gconv_rule norm_conv 1 r) (Conv.fconv_rule norm_conv (Thm.instantiate (Thm.match (apply2 norm_cterm (Thm.cprop_of th', A))) th')) end; in fun satisfy_thm witns thm = (Thm.chyps_of thm, thm) |-> fold (fn hyp => (case find_witness witns (Thm.term_of hyp) of NONE => I | SOME w => Thm.implies_intr hyp #> compose_witness w)); val satisfy_morphism = Morphism.thm_morphism "Element.satisfy" o satisfy_thm; end; (* rewriting with equalities *) (* for activating declarations only *) fun eq_term_morphism _ [] = NONE | eq_term_morphism thy props = let fun decomp_simp prop = let val ctxt = Proof_Context.init_global thy; val _ = Logic.count_prems prop = 0 orelse error ("Bad conditional rewrite rule " ^ Syntax.string_of_term ctxt prop); val lhsrhs = Logic.dest_equals prop handle TERM _ => error ("Rewrite rule not a meta-equality " ^ Syntax.string_of_term ctxt prop); in lhsrhs end; val phi = Morphism.morphism "Element.eq_term_morphism" {binding = [], typ = [], term = [Pattern.rewrite_term thy (map decomp_simp props) []], fact = [fn _ => error "Illegal application of Element.eq_term_morphism"]}; in SOME phi end; fun eq_morphism _ [] = NONE | eq_morphism thy thms = let (* FIXME proper context!? *) fun rewrite th = rewrite_rule (Proof_Context.init_global (Thm.theory_of_thm th)) thms th; val phi = Morphism.morphism "Element.eq_morphism" {binding = [], typ = [], term = [Raw_Simplifier.rewrite_term thy thms []], fact = [map rewrite]}; in SOME phi end; (** activate in context **) (* init *) fun init (Fixes fixes) = Context.map_proof (Proof_Context.add_fixes fixes #> #2) | init (Constrains _) = I | init (Assumes asms) = Context.map_proof (fn ctxt => let val asms' = Attrib.map_specs (map (Attrib.attribute ctxt)) asms; val (_, ctxt') = ctxt |> fold Proof_Context.augment (maps (map #1 o #2) asms') |> Proof_Context.add_assms Assumption.assume_export asms'; in ctxt' end) | init (Defines defs) = Context.map_proof (fn ctxt => let val defs' = Attrib.map_specs (map (Attrib.attribute ctxt)) defs; val asms = defs' |> map (fn (b, (t, ps)) => let val (_, t') = Local_Defs.cert_def ctxt (K []) t (* FIXME adapt ps? *) in (t', (b, [(t', ps)])) end); val (_, ctxt') = ctxt |> fold Proof_Context.augment (map #1 asms) |> Proof_Context.add_assms Local_Defs.def_export (map #2 asms); in ctxt' end) | init (Notes (kind, facts)) = Attrib.generic_notes kind facts #> #2 | init (Lazy_Notes (kind, ths)) = Attrib.lazy_notes kind ths; (* activate *) fun activate_i elem ctxt = let val elem' = (case (map_ctxt_attrib o map) Token.init_assignable elem of Defines defs => Defines (defs |> map (fn ((a, atts), (t, ps)) => ((Thm.def_binding_optional (Binding.name (#1 (#1 (Local_Defs.cert_def ctxt (K []) t)))) a, atts), (t, ps)))) | e => e); val ctxt' = Context.proof_map (init elem') ctxt; in ((map_ctxt_attrib o map) Token.closure elem', ctxt') end; fun activate raw_elem ctxt = let val elem = raw_elem |> map_ctxt {binding = I, typ = I, term = I, pattern = I, fact = Proof_Context.get_fact ctxt, attrib = Attrib.check_src ctxt} in activate_i elem ctxt end; end; diff --git a/src/Pure/Isar/expression.ML b/src/Pure/Isar/expression.ML --- a/src/Pure/Isar/expression.ML +++ b/src/Pure/Isar/expression.ML @@ -1,879 +1,880 @@ (* Title: Pure/Isar/expression.ML Author: Clemens Ballarin, TU Muenchen Locale expressions and user interface layer of locales. *) signature EXPRESSION = sig (* Locale expressions *) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list type 'term rewrites = (Attrib.binding * 'term) list type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list type expression_i = (string, term) expr * (binding * typ option * mixfix) list type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list (* Processing of context statements *) val cert_statement: Element.context_i list -> Element.statement_i -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context val read_statement: Element.context list -> Element.statement -> Proof.context -> (Attrib.binding * (term * term list) list) list * Proof.context (* Declaring locales *) val cert_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context_i list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val cert_read_declaration: expression_i -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) (*FIXME*) val read_declaration: expression -> (Proof.context -> Proof.context) -> Element.context list -> Proof.context -> (((string * typ) * mixfix) list * (string * morphism) list * Element.context_i list * Proof.context) * ((string * typ) list * Proof.context) val add_locale: binding -> binding -> Bundle.name list -> expression_i -> Element.context_i list -> theory -> string * local_theory val add_locale_cmd: binding -> binding -> (xstring * Position.T) list -> expression -> Element.context list -> theory -> string * local_theory (* Processing of locale expressions *) val cert_goal_expression: expression_i -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context val read_goal_expression: expression -> Proof.context -> (term list list * term list list * (string * morphism) list * (Attrib.binding * term) list list * morphism) * Proof.context end; structure Expression : EXPRESSION = struct datatype ctxt = datatype Element.ctxt; (*** Expressions ***) datatype 'term map = Positional of 'term option list | Named of (string * 'term) list; type 'term rewrites = (Attrib.binding * 'term) list; type ('name, 'term) expr = ('name * ((string * bool) * ('term map * 'term rewrites))) list; type expression_i = (string, term) expr * (binding * typ option * mixfix) list; type expression = (xstring * Position.T, string) expr * (binding * string option * mixfix) list; (** Internalise locale names in expr **) fun check_expr thy instances = map (apfst (Locale.check thy)) instances; (** Parameters of expression **) (*Sanity check of instantiations and extraction of implicit parameters. The latter only occurs iff strict = false. Positional instantiations are extended to match full length of parameter list of instantiated locale.*) fun parameters_of thy strict (expr, fixed) = let val ctxt = Proof_Context.init_global thy; fun reject_dups message xs = (case duplicates (op =) xs of [] => () | dups => error (message ^ commas dups)); fun parm_eq ((p1, mx1), (p2, mx2)) = p1 = p2 andalso (Mixfix.equal (mx1, mx2) orelse error ("Conflicting syntax for parameter " ^ quote p1 ^ " in expression" ^ Position.here_list [Mixfix.pos_of mx1, Mixfix.pos_of mx2])); fun params_loc loc = Locale.params_of thy loc |> map (apfst #1); fun params_inst (loc, (prfx, (Positional insts, eqns))) = let val ps = params_loc loc; val d = length ps - length insts; val insts' = if d < 0 then error ("More arguments than parameters in instantiation of locale " ^ quote (Locale.markup_name ctxt loc)) else insts @ replicate d NONE; val ps' = (ps ~~ insts') |> map_filter (fn (p, NONE) => SOME p | (_, SOME _) => NONE); in (ps', (loc, (prfx, (Positional insts', eqns)))) end | params_inst (loc, (prfx, (Named insts, eqns))) = let val _ = reject_dups "Duplicate instantiation of the following parameter(s): " (map fst insts); val ps' = (insts, params_loc loc) |-> fold (fn (p, _) => fn ps => if AList.defined (op =) ps p then AList.delete (op =) p ps else error (quote p ^ " not a parameter of instantiated expression")); in (ps', (loc, (prfx, (Named insts, eqns)))) end; fun params_expr is = let val (is', ps') = fold_map (fn i => fn ps => let val (ps', i') = params_inst i; val ps'' = distinct parm_eq (ps @ ps'); in (i', ps'') end) is [] in (ps', is') end; val (implicit, expr') = params_expr expr; val implicit' = map #1 implicit; val fixed' = map (Variable.check_name o #1) fixed; val _ = reject_dups "Duplicate fixed parameter(s): " fixed'; val implicit'' = if strict then [] else let val _ = reject_dups "Parameter(s) declared simultaneously in expression and for clause: " (implicit' @ fixed'); in map (fn (x, mx) => (Binding.name x, NONE, mx)) implicit end; in (expr', implicit'' @ fixed) end; (** Read instantiation **) (* Parse positional or named instantiation *) local fun prep_inst prep_term ctxt parms (Positional insts) = (insts ~~ parms) |> map (fn (NONE, p) => Free (p, dummyT) | (SOME t, _) => prep_term ctxt t) | prep_inst prep_term ctxt parms (Named insts) = parms |> map (fn p => (case AList.lookup (op =) insts p of SOME t => prep_term ctxt t | NONE => Free (p, dummyT))); in fun parse_inst x = prep_inst Syntax.parse_term x; fun make_inst x = prep_inst (K I) x; end; (* Instantiation morphism *) fun inst_morphism params ((prfx, mandatory), insts') ctxt = let (* parameters *) val parm_types = map #2 params; val type_parms = fold Term.add_tfreesT parm_types []; (* type inference *) val parm_types' = map (Type_Infer.paramify_vars o Logic.varifyT_global) parm_types; val type_parms' = fold Term.add_tvarsT parm_types' []; val checked = (map (Logic.mk_type o TVar) type_parms' @ map2 Type.constraint parm_types' insts') |> Syntax.check_terms (Config.put Type_Infer.object_logic false ctxt) val (type_parms'', insts'') = chop (length type_parms') checked; (* context *) val ctxt' = fold Proof_Context.augment checked ctxt; val certT = Thm.trim_context_ctyp o Thm.ctyp_of ctxt'; val cert = Thm.trim_context_cterm o Thm.cterm_of ctxt'; (* instantiation *) val instT = - (type_parms ~~ map Logic.dest_type type_parms'') - |> map_filter (fn (v, T) => if TFree v = T then NONE else SOME (v, T)); + TFrees.build + (fold2 (fn v => fn T => not (TFree v = T) ? TFrees.add (v, T)) + type_parms (map Logic.dest_type type_parms'')); val cert_inst = - ((map #1 params ~~ - map (Term_Subst.instantiateT_frees (TFrees.make instT)) parm_types) ~~ insts'') - |> map_filter (fn (v, t) => if Free v = t then NONE else SOME (v, cert t)); + Frees.build + (fold2 (fn v => fn t => not (Free v = t) ? Frees.add (v, cert t)) + (map #1 params ~~ map (Term_Subst.instantiateT_frees instT) parm_types) insts''); in - (Element.instantiate_normalize_morphism (map (apsnd certT) instT, cert_inst) $> + (Element.instantiate_normalize_morphism (TFrees.map (K certT) instT, cert_inst) $> Morphism.binding_morphism "Expression.inst" (Binding.prefix mandatory prfx), ctxt') end; (*** Locale processing ***) (** Parsing **) fun parse_elem prep_typ prep_term ctxt = Element.map_ctxt {binding = I, typ = prep_typ ctxt, term = prep_term (Proof_Context.set_mode Proof_Context.mode_schematic ctxt), pattern = prep_term (Proof_Context.set_mode Proof_Context.mode_pattern ctxt), fact = I, attrib = I}; fun prepare_stmt prep_prop prep_obtains ctxt stmt = (case stmt of Element.Shows raw_shows => raw_shows |> (map o apsnd o map) (fn (t, ps) => (prep_prop (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) t, map (prep_prop (Proof_Context.set_mode Proof_Context.mode_pattern ctxt)) ps)) | Element.Obtains raw_obtains => let val ((_, thesis), thesis_ctxt) = Obtain.obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; in map (fn (b, t) => ((b, []), [(t, [])])) obtains end); (** Simultaneous type inference: instantiations + elements + statement **) local fun mk_type T = (Logic.mk_type T, []); fun mk_term t = (t, []); fun mk_propp (p, pats) = (Type.constraint propT p, pats); fun dest_type (T, []) = Logic.dest_type T; fun dest_term (t, []) = t; fun dest_propp (p, pats) = (p, pats); fun extract_inst (_, (_, ts)) = map mk_term ts; fun restore_inst ((l, (p, _)), cs) = (l, (p, map dest_term cs)); fun extract_eqns es = map (mk_term o snd) es; fun restore_eqns (es, cs) = map2 (fn (b, _) => fn c => (b, dest_term c)) es cs; fun extract_elem (Fixes fixes) = map (#2 #> the_list #> map mk_type) fixes | extract_elem (Constrains csts) = map (#2 #> single #> map mk_type) csts | extract_elem (Assumes asms) = map (#2 #> map mk_propp) asms | extract_elem (Defines defs) = map (fn (_, (t, ps)) => [mk_propp (t, ps)]) defs | extract_elem (Notes _) = [] | extract_elem (Lazy_Notes _) = []; fun restore_elem (Fixes fixes, css) = (fixes ~~ css) |> map (fn ((x, _, mx), cs) => (x, cs |> map dest_type |> try hd, mx)) |> Fixes | restore_elem (Constrains csts, css) = (csts ~~ css) |> map (fn ((x, _), cs) => (x, cs |> map dest_type |> hd)) |> Constrains | restore_elem (Assumes asms, css) = (asms ~~ css) |> map (fn ((b, _), cs) => (b, map dest_propp cs)) |> Assumes | restore_elem (Defines defs, css) = (defs ~~ css) |> map (fn ((b, _), [c]) => (b, dest_propp c)) |> Defines | restore_elem (elem as Notes _, _) = elem | restore_elem (elem as Lazy_Notes _, _) = elem; fun prep (_, pats) (ctxt, t :: ts) = let val ctxt' = Proof_Context.augment t ctxt in ((t, Syntax.check_props (Proof_Context.set_mode Proof_Context.mode_pattern ctxt') pats), (ctxt', ts)) end; fun check cs ctxt = let val (cs', (ctxt', _)) = fold_map prep cs (ctxt, Syntax.check_terms (Proof_Context.set_mode Proof_Context.mode_schematic ctxt) (map fst cs)); in (cs', ctxt') end; in fun check_autofix insts eqnss elems concl ctxt = let val inst_cs = map extract_inst insts; val eqns_cs = map extract_eqns eqnss; val elem_css = map extract_elem elems; val concl_cs = (map o map) mk_propp (map snd concl); (* Type inference *) val (inst_cs' :: eqns_cs' :: css', ctxt') = (fold_burrow o fold_burrow) check (inst_cs :: eqns_cs :: elem_css @ [concl_cs]) ctxt; val (elem_css', [concl_cs']) = chop (length elem_css) css'; in ((map restore_inst (insts ~~ inst_cs'), map restore_eqns (eqnss ~~ eqns_cs'), map restore_elem (elems ~~ elem_css'), map fst concl ~~ concl_cs'), ctxt') end; end; (** Prepare locale elements **) fun declare_elem prep_var (Fixes fixes) ctxt = let val (vars, _) = fold_map prep_var fixes ctxt in ctxt |> Proof_Context.add_fixes vars |> snd end | declare_elem prep_var (Constrains csts) ctxt = ctxt |> fold_map (fn (x, T) => prep_var (Binding.name x, SOME T, NoSyn)) csts |> snd | declare_elem _ (Assumes _) ctxt = ctxt | declare_elem _ (Defines _) ctxt = ctxt | declare_elem _ (Notes _) ctxt = ctxt | declare_elem _ (Lazy_Notes _) ctxt = ctxt; (** Finish locale elements **) fun finish_inst ctxt (loc, (prfx, inst)) = let val thy = Proof_Context.theory_of ctxt; val (morph, _) = inst_morphism (map #1 (Locale.params_of thy loc)) (prfx, inst) ctxt; in (loc, morph) end; fun finish_fixes (parms: (string * typ) list) = map (fn (binding, _, mx) => let val x = Binding.name_of binding in (binding, AList.lookup (op =) parms x, mx) end); local fun closeup _ _ false elem = elem | closeup (outer_ctxt, ctxt) parms true elem = let (* FIXME consider closing in syntactic phase -- before type checking *) fun close_frees t = let val rev_frees = Term.fold_aterms (fn Free (x, T) => if Variable.is_fixed outer_ctxt x orelse AList.defined (op =) parms x then I else insert (op =) (x, T) | _ => I) t []; in fold (Logic.all o Free) rev_frees t end; fun no_binds [] = [] | no_binds _ = error "Illegal term bindings in context element"; in (case elem of Assumes asms => Assumes (asms |> map (fn (a, propps) => (a, map (fn (t, ps) => (close_frees t, no_binds ps)) propps))) | Defines defs => Defines (defs |> map (fn ((name, atts), (t, ps)) => let val ((c, _), t') = Local_Defs.cert_def ctxt (K []) (close_frees t) in ((Thm.def_binding_optional (Binding.name c) name, atts), (t', no_binds ps)) end)) | e => e) end; in fun finish_elem _ parms _ (Fixes fixes) = Fixes (finish_fixes parms fixes) | finish_elem _ _ _ (Constrains _) = Constrains [] | finish_elem ctxts parms do_close (Assumes asms) = closeup ctxts parms do_close (Assumes asms) | finish_elem ctxts parms do_close (Defines defs) = closeup ctxts parms do_close (Defines defs) | finish_elem _ _ _ (elem as Notes _) = elem | finish_elem _ _ _ (elem as Lazy_Notes _) = elem; end; (** Process full context statement: instantiations + elements + statement **) (* Interleave incremental parsing and type inference over entire parsed stretch. *) local fun abs_def ctxt = Thm.cterm_of ctxt #> Assumption.assume ctxt #> Local_Defs.abs_def_rule ctxt #> Thm.prop_of; fun prep_full_context_statement parse_typ parse_prop prep_obtains prep_var_elem prep_inst prep_eqns prep_attr prep_var_inst prep_expr {strict, do_close, fixed_frees} raw_import init_body raw_elems raw_stmt ctxt1 = let val thy = Proof_Context.theory_of ctxt1; val (raw_insts, fixed) = parameters_of thy strict (apfst (prep_expr thy) raw_import); fun prep_insts_cumulative (loc, (prfx, (inst, eqns))) (i, insts, eqnss, ctxt) = let val params = map #1 (Locale.params_of thy loc); val inst' = prep_inst ctxt (map #1 params) inst; val parm_types' = params |> map (#2 #> Logic.varifyT_global #> Term.map_type_tvar (fn ((x, _), S) => TVar ((x, i), S)) #> Type_Infer.paramify_vars); val inst'' = map2 Type.constraint parm_types' inst'; val insts' = insts @ [(loc, (prfx, inst''))]; val ((insts'', _, _, _), ctxt2) = check_autofix insts' [] [] [] ctxt; val inst''' = insts'' |> List.last |> snd |> snd; val (inst_morph, _) = inst_morphism params (prfx, inst''') ctxt; val ctxt' = Locale.activate_declarations (loc, inst_morph) ctxt2 handle ERROR msg => if null eqns then error msg else (Locale.tracing ctxt1 (msg ^ "\nFalling back to reading rewrites clause before activation."); ctxt2); val attrss = map (apsnd (map (prep_attr ctxt)) o fst) eqns; val eqns' = (prep_eqns ctxt' o map snd) eqns; val eqnss' = [attrss ~~ eqns']; val ((_, [eqns''], _, _), _) = check_autofix insts'' eqnss' [] [] ctxt'; val rewrite_morph = eqns' |> map (abs_def ctxt') |> Variable.export_terms ctxt' ctxt |> Element.eq_term_morphism (Proof_Context.theory_of ctxt) |> the_default Morphism.identity; val ctxt'' = Locale.activate_declarations (loc, inst_morph $> rewrite_morph) ctxt; val eqnss' = eqnss @ [attrss ~~ Variable.export_terms ctxt' ctxt eqns']; in (i + 1, insts', eqnss', ctxt'') end; fun prep_elem raw_elem ctxt = let val ctxt' = ctxt |> Context_Position.set_visible false |> declare_elem prep_var_elem raw_elem |> Context_Position.restore_visible ctxt; val elems' = parse_elem parse_typ parse_prop ctxt' raw_elem; in (elems', ctxt') end; val fors = fold_map prep_var_inst fixed ctxt1 |> fst; val ctxt2 = ctxt1 |> Proof_Context.add_fixes fors |> snd; val (_, insts', eqnss', ctxt3) = fold prep_insts_cumulative raw_insts (0, [], [], ctxt2); fun prep_stmt elems ctxt = check_autofix insts' [] elems (prepare_stmt parse_prop prep_obtains ctxt raw_stmt) ctxt; val _ = if fixed_frees then () else (case fold (fold (Variable.add_frees ctxt3) o snd o snd) insts' [] of [] => () | frees => error ("Illegal free variables in expression: " ^ commas_quote (map (Syntax.string_of_term ctxt3 o Free) (rev frees)))); val ((insts, _, elems', concl), ctxt4) = ctxt3 |> init_body |> fold_map prep_elem raw_elems |-> prep_stmt; (* parameters from expression and elements *) val xs = maps (fn Fixes fixes => map (Variable.check_name o #1) fixes | _ => []) (Fixes fors :: elems'); val (parms, ctxt5) = fold_map Proof_Context.inferred_param xs ctxt4; val fors' = finish_fixes parms fors; val fixed = map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fors'; val deps = map (finish_inst ctxt5) insts; val elems'' = map (finish_elem (ctxt1, ctxt5) parms do_close) elems'; in ((fixed, deps, eqnss', elems'', concl), (parms, ctxt5)) end; in fun cert_full_context_statement x = prep_full_context_statement (K I) (K I) Obtain.cert_obtains Proof_Context.cert_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun cert_read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var make_inst Syntax.check_props (K I) Proof_Context.cert_var (K I) x; fun read_full_context_statement x = prep_full_context_statement Syntax.parse_typ Syntax.parse_prop Obtain.parse_obtains Proof_Context.read_var parse_inst Syntax.read_props Attrib.check_src Proof_Context.read_var check_expr x; end; (* Context statement: elements + statement *) local fun prep_statement prep activate raw_elems raw_stmt ctxt = let val ((_, _, _, elems, concl), _) = prep {strict = true, do_close = false, fixed_frees = true} ([], []) I raw_elems raw_stmt ctxt; val ctxt' = ctxt |> Proof_Context.set_stmt true |> fold_map activate elems |> #2 |> Proof_Context.restore_stmt ctxt; in (concl, ctxt') end; in fun cert_statement x = prep_statement cert_full_context_statement Element.activate_i x; fun read_statement x = prep_statement read_full_context_statement Element.activate x; end; (* Locale declaration: import + elements *) fun fix_params params = Proof_Context.add_fixes (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) params) #> snd; local fun prep_declaration prep activate raw_import init_body raw_elems ctxt = let val ((fixed, deps, eqnss, elems, _), (parms, ctxt0)) = prep {strict = false, do_close = true, fixed_frees = false} raw_import init_body raw_elems (Element.Shows []) ctxt; val _ = null (flat eqnss) orelse error "Illegal rewrites clause(s) in declaration of locale"; (* Declare parameters and imported facts *) val ctxt' = ctxt |> fix_params fixed |> fold (Context.proof_map o Locale.activate_facts NONE) deps; val (elems', ctxt'') = ctxt' |> Proof_Context.set_stmt true |> fold_map activate elems ||> Proof_Context.restore_stmt ctxt'; in ((fixed, deps, elems', ctxt''), (parms, ctxt0)) end; in fun cert_declaration x = prep_declaration cert_full_context_statement Element.activate_i x; fun cert_read_declaration x = prep_declaration cert_read_full_context_statement Element.activate x; fun read_declaration x = prep_declaration read_full_context_statement Element.activate x; end; (* Locale expression to set up a goal *) local fun props_of thy (name, morph) = let val (asm, defs) = Locale.specification_of thy name in map (Morphism.term morph) (the_list asm @ defs) end; fun prep_goal_expression prep expression ctxt = let val thy = Proof_Context.theory_of ctxt; val ((fixed, deps, eqnss, _, _), _) = prep {strict = true, do_close = true, fixed_frees = true} expression I [] (Element.Shows []) ctxt; (* proof obligations *) val propss = map (props_of thy) deps; val eq_propss = (map o map) snd eqnss; val goal_ctxt = ctxt |> fix_params fixed |> (fold o fold) Proof_Context.augment (propss @ eq_propss); val export = Proof_Context.export_morphism goal_ctxt ctxt; val exp_fact = Drule.zero_var_indexes_list o map Thm.strip_shyps o Morphism.fact export; val exp_term = Term_Subst.zero_var_indexes o Morphism.term export; val exp_typ = Logic.type_map exp_term; val export' = Morphism.morphism "Expression.prep_goal" {binding = [], typ = [exp_typ], term = [exp_term], fact = [exp_fact]}; in ((propss, eq_propss, deps, eqnss, export'), goal_ctxt) end; in fun cert_goal_expression x = prep_goal_expression cert_full_context_statement x; fun read_goal_expression x = prep_goal_expression read_full_context_statement x; end; (*** Locale declarations ***) (* extract specification text *) val norm_term = Envir.beta_norm oo Term.subst_atomic; fun bind_def ctxt eq (env, eqs) = let val _ = Local_Defs.cert_def ctxt (K []) eq; val ((y, T), b) = Local_Defs.abs_def eq; val b' = norm_term env b; fun err msg = error (msg ^ ": " ^ quote y); in (case filter (fn (Free (y', _), _) => y = y' | _ => false) env of [] => ((Free (y, T), b') :: env, eq :: eqs) | dups => if forall (fn (_, b'') => b' aconv b'') dups then (env, eqs) else err "Attempt to redefine variable") end; (* text has the following structure: (((exts, exts'), (ints, ints')), (xs, env, defs)) where exts: external assumptions (terms in assumes elements) exts': dito, normalised wrt. env ints: internal assumptions (terms in assumptions from insts) ints': dito, normalised wrt. env xs: the free variables in exts' and ints' and rhss of definitions, this includes parameters except defined parameters env: list of term pairs encoding substitutions, where the first term is a free variable; substitutions represent defines elements and the rhs is normalised wrt. the previous env defs: the equations from the defines elements *) fun eval_text _ _ (Fixes _) text = text | eval_text _ _ (Constrains _) text = text | eval_text _ is_ext (Assumes asms) (((exts, exts'), (ints, ints')), (env, defs)) = let val ts = maps (map #1 o #2) asms; val ts' = map (norm_term env) ts; val spec' = if is_ext then ((exts @ ts, exts' @ ts'), (ints, ints')) else ((exts, exts'), (ints @ ts, ints' @ ts')); in (spec', (env, defs)) end | eval_text ctxt _ (Defines defs) (spec, binds) = (spec, fold (bind_def ctxt o #1 o #2) defs binds) | eval_text _ _ (Notes _) text = text | eval_text _ _ (Lazy_Notes _) text = text; fun eval_inst ctxt (loc, morph) text = let val thy = Proof_Context.theory_of ctxt; val (asm, defs) = Locale.specification_of thy loc; val asm' = Option.map (Morphism.term morph) asm; val defs' = map (Morphism.term morph) defs; val text' = text |> (if is_some asm then eval_text ctxt false (Assumes [(Binding.empty_atts, [(the asm', [])])]) else I) |> (if not (null defs) then eval_text ctxt false (Defines (map (fn def => (Binding.empty_atts, (def, []))) defs')) else I) (* FIXME clone from locale.ML *) in text' end; fun eval_elem ctxt elem text = eval_text ctxt true elem text; fun eval ctxt deps elems = let val text' = fold (eval_inst ctxt) deps ((([], []), ([], [])), ([], [])); val ((spec, (_, defs))) = fold (eval_elem ctxt) elems text'; in (spec, defs) end; (* axiomsN: name of theorem set with destruct rules for locale predicates, also name suffix of delta predicates and assumptions. *) val axiomsN = "axioms"; local (* introN: name of theorems for introduction rules of locale and delta predicates *) val introN = "intro"; fun atomize_spec ctxt ts = let val t = Logic.mk_conjunction_balanced ts; val body = Object_Logic.atomize_term ctxt t; val bodyT = Term.fastype_of body; in if bodyT = propT then (t, propT, Thm.reflexive (Thm.cterm_of ctxt t)) else (body, bodyT, Object_Logic.atomize ctxt (Thm.cterm_of ctxt t)) end; (* achieve plain syntax for locale predicates (without "PROP") *) fun aprop_tr' n c = let val c' = Lexicon.mark_const c; fun tr' (_: Proof.context) T args = if T <> dummyT andalso length args = n then Syntax.const "_aprop" $ Term.list_comb (Syntax.const c', args) else raise Match; in (c', tr') end; (* define one predicate including its intro rule and axioms - binding: predicate name - parms: locale parameters - defs: thms representing substitutions from defines elements - ts: terms representing locale assumptions (not normalised wrt. defs) - norm_ts: terms representing locale assumptions (normalised wrt. defs) - thy: the theory *) fun def_pred binding parms defs ts norm_ts thy = let val name = Sign.full_name thy binding; val thy_ctxt = Proof_Context.init_global thy; val (body, bodyT, body_eq) = atomize_spec thy_ctxt norm_ts; val env = Names.build (Names.add_free_names body); val xs = filter (Names.defined env o #1) parms; val Ts = map #2 xs; val type_tfrees = TFrees.build (fold TFrees.add_tfreesT Ts); val extra_tfrees = TFrees.build (TFrees.add_tfrees_unless (TFrees.defined type_tfrees) body) |> TFrees.keys |> map TFree; val predT = map Term.itselfT extra_tfrees ---> Ts ---> bodyT; val args = map Logic.mk_type extra_tfrees @ map Free xs; val head = Term.list_comb (Const (name, predT), args); val statement = Object_Logic.ensure_propT thy_ctxt head; val ([pred_def], defs_thy) = thy |> bodyT = propT ? Sign.typed_print_translation [aprop_tr' (length args) name] |> Sign.declare_const_global ((binding, predT), NoSyn) |> snd |> Global_Theory.add_defs false [((Thm.def_binding binding, Logic.mk_equals (head, body)), [])]; val defs_ctxt = Proof_Context.init_global defs_thy |> Variable.declare_term head; val intro = Goal.prove_global defs_thy [] norm_ts statement (fn {context = ctxt, ...} => rewrite_goals_tac ctxt [pred_def] THEN compose_tac defs_ctxt (false, body_eq RS Drule.equal_elim_rule1, 1) 1 THEN compose_tac defs_ctxt (false, Conjunction.intr_balanced (map (Thm.assume o Thm.cterm_of defs_ctxt) norm_ts), 0) 1); val conjuncts = (Drule.equal_elim_rule2 OF [body_eq, rewrite_rule defs_ctxt [pred_def] (Thm.assume (Thm.cterm_of defs_ctxt statement))]) |> Conjunction.elim_balanced (length ts); val (_, axioms_ctxt) = defs_ctxt |> Assumption.add_assumes (maps Thm.chyps_of (defs @ conjuncts)); val axioms = ts ~~ conjuncts |> map (fn (t, ax) => Element.prove_witness axioms_ctxt t (rewrite_goals_tac axioms_ctxt defs THEN compose_tac axioms_ctxt (false, ax, 0) 1)); in ((statement, intro, axioms), defs_thy) end; in (* main predicate definition function *) fun define_preds binding parms (((exts, exts'), (ints, ints')), defs) thy = let val ctxt = Proof_Context.init_global thy; val defs' = map (Thm.cterm_of ctxt #> Assumption.assume ctxt #> Drule.abs_def) defs; val (a_pred, a_intro, a_axioms, thy'') = if null exts then (NONE, NONE, [], thy) else let val abinding = if null ints then binding else Binding.suffix_name ("_" ^ axiomsN) binding; val ((statement, intro, axioms), thy') = thy |> def_pred abinding parms defs' exts exts'; val ((_, [intro']), thy'') = thy' |> Sign.qualified_path true abinding |> Global_Theory.note_thms "" ((Binding.name introN, []), [([intro], [Locale.unfold_add])]) ||> Sign.restore_naming thy'; in (SOME statement, SOME intro', axioms, thy'') end; val (b_pred, b_intro, b_axioms, thy'''') = if null ints then (NONE, NONE, [], thy'') else let val ((statement, intro, axioms), thy''') = thy'' |> def_pred binding parms defs' (ints @ the_list a_pred) (ints' @ the_list a_pred); val ctxt''' = Proof_Context.init_global thy'''; val ([(_, [intro']), _], thy'''') = thy''' |> Sign.qualified_path true binding |> Global_Theory.note_thmss "" [((Binding.name introN, []), [([intro], [Locale.intro_add])]), ((Binding.name axiomsN, []), [(map (Drule.export_without_context o Element.conclude_witness ctxt''') axioms, [])])] ||> Sign.restore_naming thy'''; in (SOME statement, SOME intro', axioms, thy'''') end; in ((a_pred, a_intro, a_axioms), (b_pred, b_intro, b_axioms), thy'''') end; end; local fun assumes_to_notes (Assumes asms) axms = fold_map (fn (a, spec) => fn axs => let val (ps, qs) = chop (length spec) axs in ((a, [(ps, [])]), qs) end) asms axms |> apfst (curry Notes "") | assumes_to_notes e axms = (e, axms); fun defines_to_notes ctxt (Defines defs) = Notes ("", map (fn (a, (def, _)) => (a, [([Assumption.assume ctxt (Thm.cterm_of ctxt def)], [(Attrib.internal o K) Locale.witness_add])])) defs) | defines_to_notes _ e = e; val is_hyp = fn Assumes _ => true | Defines _ => true | _ => false; fun gen_add_locale prep_include prep_decl binding raw_predicate_binding raw_includes raw_import raw_body thy = let val name = Sign.full_name thy binding; val _ = Locale.defined thy name andalso error ("Duplicate definition of locale " ^ quote name); val ctxt = Proof_Context.init_global thy; val includes = map (prep_include ctxt) raw_includes; val ((fixed, deps, body_elems, _), (parms, ctxt')) = ctxt |> Bundle.includes includes |> prep_decl raw_import I raw_body; val text as (((_, exts'), _), defs) = eval ctxt' deps body_elems; val type_tfrees = TFrees.build (fold (TFrees.add_tfreesT o #2) parms); val extra_tfrees = TFrees.build (fold (TFrees.add_tfrees_unless (TFrees.defined type_tfrees)) exts') |> TFrees.keys; val _ = if null extra_tfrees then () else warning ("Additional type variable(s) in locale specification " ^ Binding.print binding ^ ": " ^ commas (map (Syntax.string_of_typ ctxt' o TFree) extra_tfrees)); val predicate_binding = if Binding.is_empty raw_predicate_binding then binding else raw_predicate_binding; val ((a_statement, a_intro, a_axioms), (b_statement, b_intro, b_axioms), thy') = define_preds predicate_binding parms text thy; val pred_ctxt = Proof_Context.init_global thy'; val a_satisfy = Element.satisfy_morphism a_axioms; val b_satisfy = Element.satisfy_morphism b_axioms; val params = fixed @ maps (fn Fixes fixes => map (fn (b, SOME T, mx) => ((Binding.name_of b, T), mx)) fixes | _ => []) body_elems; val asm = if is_some b_statement then b_statement else a_statement; val hyp_spec = filter is_hyp body_elems; val notes = if is_some asm then [("", [((Binding.suffix_name ("_" ^ axiomsN) binding, []), [([Assumption.assume pred_ctxt (Thm.cterm_of pred_ctxt (the asm))], [(Attrib.internal o K) Locale.witness_add])])])] else []; val notes' = body_elems |> map (defines_to_notes pred_ctxt) |> map (Element.transform_ctxt a_satisfy) |> (fn elems => fold_map assumes_to_notes elems (map (Element.conclude_witness pred_ctxt) a_axioms)) |> fst |> map (Element.transform_ctxt b_satisfy) |> map_filter (fn Notes notes => SOME notes | _ => NONE); val deps' = map (fn (l, morph) => (l, morph $> b_satisfy)) deps; val axioms = map (Element.conclude_witness pred_ctxt) b_axioms; val loc_ctxt = thy' |> Locale.register_locale binding (extra_tfrees, params) (asm, rev defs) (a_intro, b_intro) axioms hyp_spec [] (rev notes) (rev deps') |> Named_Target.init includes name |> fold (fn (kind, facts) => Local_Theory.notes_kind kind facts #> snd) notes'; in (name, loc_ctxt) end; in val add_locale = gen_add_locale (K I) cert_declaration; val add_locale_cmd = gen_add_locale Bundle.check read_declaration; end; end; diff --git a/src/Pure/Isar/generic_target.ML b/src/Pure/Isar/generic_target.ML --- a/src/Pure/Isar/generic_target.ML +++ b/src/Pure/Isar/generic_target.ML @@ -1,476 +1,478 @@ (* 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_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_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_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_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 -> 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.list_set; 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) (sort_by #1 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 = 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)]) | 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)])))) end else context; (** background primitives **) structure Foundation_Interpretations = Theory_Data ( type T = (binding * (term * term list) -> Context.generic -> Context.generic) Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun add_foundation_interpretation f = Foundation_Interpretations.map (Inttab.update_new (serial (), f)); fun foundation_interpretation binding_const_params lthy = let val interps = Foundation_Interpretations.get (Proof_Context.theory_of lthy); val interp = Inttab.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.list_set; 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.list_set_rev |> 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.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) instT []; + val cinstT = TVars.map (K (Thm.ctyp_of ctxt)) instT; val cinst = - map_filter - (fn (Var (xi, T), t) => - SOME ((xi, Term_Subst.instantiateT instT T), - Thm.cterm_of ctxt (Term.map_types (Term_Subst.instantiateT instT) t)) - | _ => NONE) (vars ~~ frees); + 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/obtain.ML b/src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML +++ b/src/Pure/Isar/obtain.ML @@ -1,425 +1,425 @@ (* Title: Pure/Isar/obtain.ML Author: Markus Wenzel, TU Muenchen Generalized existence and cases rules within Isar proof text. *) signature OBTAIN = sig val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list val obtains_attribs: ('typ, 'term) Element.obtain list -> Token.src list val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state val obtain: binding -> (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> (term * term list) list list -> (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state val obtain_cmd: binding -> (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> (string * string list) list list -> (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state val result: (Proof.context -> tactic) -> thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state end; structure Obtain: OBTAIN = struct (** specification elements **) (* obtain_export *) (* [x, A x] : B -------- B *) fun eliminate_term ctxt xs tm = let val vs = map (dest_Free o Thm.term_of) xs; val bads = Term.fold_aterms (fn t as Free v => if member (op =) vs v then insert (op aconv) t else I | _ => I) tm []; val _ = null bads orelse error ("Result contains obtained parameters: " ^ space_implode " " (map (Syntax.string_of_term ctxt) bads)); in tm end; fun eliminate ctxt rule xs As thm = let val _ = eliminate_term ctxt xs (Thm.full_prop_of thm); val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse error "Conclusion in obtained context must be object-logic judgment"; val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt; val prems = Drule.strip_imp_prems (Thm.cprop_of thm'); in ((Drule.implies_elim_list thm' (map Thm.assume prems) |> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As) |> Drule.forall_intr_list xs) COMP rule) |> Drule.implies_intr_list prems |> singleton (Variable.export ctxt' ctxt) end; fun obtain_export ctxt rule xs _ As = (eliminate ctxt rule xs As, eliminate_term ctxt xs); (* result declaration *) fun case_names (obtains: ('typ, 'term) Element.obtain list) = obtains |> map_index (fn (i, (b, _)) => if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b); fun obtains_attributes obtains = [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names (case_names obtains)]; fun obtains_attribs obtains = [Attrib.consumes (~ (length obtains)), Attrib.case_names (case_names obtains)]; (* obtain thesis *) fun obtain_thesis ctxt = let val ([x], ctxt') = Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt; val t = Object_Logic.fixed_judgment ctxt x; val v = dest_Free (Object_Logic.drop_judgment ctxt t); in ((v, t), ctxt') end; (* obtain clauses *) local val mk_all_external = Logic.all_constraint o Variable.default_type; fun mk_all_internal ctxt (y, z) t = let val frees = Frees.build (t |> Term.fold_aterms (fn Free v => Frees.add_set v | _ => I)); val T = (case Frees.get_first (fn ((x, T), _) => if x = z then SOME T else NONE) frees of SOME T => T | NONE => the_default dummyT (Variable.default_type ctxt z)); in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end; fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props = let val ((xs', vars), ctxt') = ctxt |> fold_map prep_var raw_vars |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars); val xs = map (Variable.check_name o #1) vars; in Logic.list_implies (map (parse_prop ctxt') raw_props, thesis) |> fold_rev (mk_all ctxt') (xs ~~ xs') end; fun prepare_obtains prep_clause check_terms ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) = let val clauses = raw_obtains |> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props) |> check_terms ctxt; in map fst raw_obtains ~~ clauses end; val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external; val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal; in val read_obtains = prepare_obtains parse_clause Syntax.check_terms; val cert_obtains = prepare_obtains cert_clause (K I); val parse_obtains = prepare_obtains parse_clause (K I); end; (** consider: generalized elimination and cases rule **) (* consider (a) x where "A x" | (b) y where "B y" | ... \ have thesis if a [intro?]: "\x. A x \ thesis" and b [intro?]: "\y. B y \ thesis" and ... for thesis apply (insert that) *) local fun gen_consider prep_obtains raw_obtains int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt; val obtains = prep_obtains thesis_ctxt thesis raw_obtains; val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains; in state |> Proof.have true NONE (K I) [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] (map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains) [((Binding.empty, atts), [(thesis, [])])] int |-> Proof.refine_insert end; in val consider = gen_consider cert_obtains; val consider_cmd = gen_consider read_obtains; end; (** obtain: augmented context based on generalized existence rule **) (* obtain (a) x where "A x" \ have thesis if a [intro?]: "\x. A x \ thesis" for thesis apply (insert that) fix x assm <> "A x" *) local fun gen_obtain prep_stmt prep_att that_binding raw_decls raw_fixes raw_prems raw_concls int state = let val _ = Proof.assert_forward_or_chain state; val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state); val ({vars, propss, binds, result_binds, ...}, params_ctxt) = prep_stmt (raw_decls @ raw_fixes) (raw_prems @ map #2 raw_concls) thesis_ctxt; val (decls, fixes) = chop (length raw_decls) vars ||> map #2; val (premss, conclss) = chop (length raw_prems) propss; val propss' = (map o map) (Logic.close_prop fixes (flat premss)) conclss; val that_prop = Logic.list_rename_params (map (#1 o #2) decls) (fold_rev (Logic.all o #2 o #2) decls (Logic.list_implies (flat propss', thesis))); val cparams = map (Thm.cterm_of params_ctxt o #2 o #2) decls; val asms = map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_concls ~~ map (map (rpair [])) propss'; fun after_qed (result_ctxt, results) state' = let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.fix (map #1 decls) |> Proof.map_context (fold (Variable.bind_term o apsnd (Logic.close_term fixes)) binds) |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms end; in state |> Proof.have true NONE after_qed [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] [((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])] [(Binding.empty_atts, [(thesis, [])])] int |-> Proof.refine_insert |> Proof.map_context (fold Variable.bind_term result_binds) end; in val obtain = gen_obtain Proof_Context.cert_stmt (K I); val obtain_cmd = gen_obtain Proof_Context.read_stmt Attrib.attribute_cmd; end; (** tactical result **) fun check_result ctxt thesis th = (case Thm.prems_of th of [prem] => if Thm.concl_of th aconv thesis andalso Logic.strip_assums_concl prem aconv thesis then th else error ("Guessed a different clause:\n" ^ Thm.string_of_thm ctxt th) | [] => error "Goal solved -- nothing guessed" | _ => error ("Guess split into several cases:\n" ^ Thm.string_of_thm ctxt th)); fun result tac facts ctxt = let val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt; val st = Goal.init (Thm.cterm_of ctxt thesis); val rule = (case SINGLE (Method.insert_tac thesis_ctxt facts 1 THEN tac thesis_ctxt) st of NONE => raise THM ("Obtain.result: tactic failed", 0, facts) | SOME th => check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th))); val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule; val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt; val obtain_rule = Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule'; val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt'; val (prems, ctxt'') = Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params)) (Drule.strip_imp_prems stmt) fix_ctxt; in ((params, prems), ctxt'') end; (** guess: obtain based on tactical result **) (* guess x \ { fix thesis have "PROP ?guess" apply magic \ \turn goal into \thesis \ #thesis\\ apply_end magic \ \turn final \(\x. P x \ thesis) \ #thesis\ into\ \ \\#((\x. A x \ thesis) \ thesis)\ which is a finished goal state\ } fix x assm <> "A x" *) local fun unify_params vars thesis_var raw_rule ctxt = let val thy = Proof_Context.theory_of ctxt; val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); fun err msg th = error (msg ^ ":\n" ^ Thm.string_of_thm ctxt th); val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1; val rule = Thm.incr_indexes (maxidx + 1) raw_rule; val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule)); val m = length vars; val n = length params; val _ = m <= n orelse err "More variables than parameters in obtained rule" rule; fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max) handle Type.TUNIFY => err ("Failed to unify variable " ^ string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ string_of_term (Syntax_Trans.mark_bound_abs (y, Envir.norm_type tyenv U)) ^ " in") rule; val (tyenv, _) = fold unify (map #1 vars ~~ take m params) (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule)); val norm_type = Envir.norm_type tyenv; val xs = map (apsnd norm_type o fst) vars; val ys = map (apsnd norm_type) (drop m params); val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; val terms = map (Drule.mk_term o Thm.cterm_of ctxt o Free) (xs @ ys'); val instT = TVars.build (params |> fold (#2 #> Term.fold_atyps (fn T => fn instT => (case T of TVar v => if TVars.defined instT v then instT else TVars.add (v, Thm.ctyp_of ctxt (norm_type T)) instT | _ => instT)))); val closed_rule = rule |> Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) - |> Thm.instantiate (TVars.dest instT, []); + |> Thm.instantiate (instT, Vars.empty); val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt; val vars' = map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~ (map snd vars @ replicate (length ys) NoSyn); val rule'' = Thm.forall_elim (Thm.cterm_of ctxt' (Logic.varify_global (Free thesis_var))) rule'; in ((vars', rule''), ctxt') end; fun inferred_type (binding, _, mx) ctxt = let val x = Variable.check_name binding; val ((_, T), ctxt') = Proof_Context.inferred_param x ctxt in ((x, T, mx), ctxt') end; fun polymorphic ctxt vars = let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars)) in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end; fun gen_guess prep_var raw_vars int state = let val _ = Proof.assert_forward_or_chain state; val ctxt = Proof.context_of state; val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else []; val (thesis_var, thesis) = #1 (obtain_thesis ctxt); val vars = ctxt |> fold_map prep_var raw_vars |-> fold_map inferred_type |> fst |> polymorphic ctxt; fun guess_context raw_rule state' = let val ((parms, rule), ctxt') = unify_params vars thesis_var raw_rule (Proof.context_of state'); val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt'; val ps = xs ~~ map (#2 o #1) parms; val ts = map Free ps; val asms = Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule)) |> map (fn asm => (Term.betapplys (fold_rev Term.abs ps asm, ts), [])); val _ = not (null asms) orelse error "Trivial result -- nothing guessed"; in state' |> Proof.map_context (K ctxt') |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms) |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm (obtain_export fix_ctxt rule (map (Thm.cterm_of ctxt) ts)) [] [] [(Binding.empty_atts, asms)]) |> Proof.map_context (fold Variable.unbind_term Auto_Bind.no_facts) end; val goal = Var (("guess", 0), propT); val pos = Position.thread_data (); fun print_result ctxt' (k, [(s, [_, th])]) = Proof_Display.print_results int pos ctxt' (k, [(s, [th])]); val before_qed = Method.primitive_text (fn ctxt => Goal.conclude #> Raw_Simplifier.norm_hhf ctxt #> (fn th => Goal.protect 0 (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th))); fun after_qed (result_ctxt, results) state' = let val [_, res] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in state' |> Proof.end_block |> guess_context (check_result ctxt thesis res) end; in state |> Proof.enter_forward |> Proof.begin_block |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] |> Proof.chain_facts chain_facts |> Proof.internal_goal print_result Proof_Context.mode_schematic true "guess" (SOME before_qed) after_qed [] [] [(Binding.empty_atts, [(Logic.mk_term goal, []), (goal, [])])] |> snd |> Proof.refine_singleton (Method.primitive_text (fn _ => fn _ => Goal.init (Thm.cterm_of ctxt thesis))) end; in val guess = gen_guess Proof_Context.cert_var; val guess_cmd = gen_guess Proof_Context.read_var; end; end; diff --git a/src/Pure/Isar/subgoal.ML b/src/Pure/Isar/subgoal.ML --- a/src/Pure/Isar/subgoal.ML +++ b/src/Pure/Isar/subgoal.ML @@ -1,255 +1,255 @@ (* Title: Pure/Isar/subgoal.ML Author: Makarius Author: Daniel Matichuk, NICTA/UNSW Tactical operations with explicit subgoal focus, based on canonical proof decomposition. The "visible" part of the text within the context is fixed, the remaining goal may be schematic. Isar subgoal command for proof structure within unstructured proof scripts. *) signature SUBGOAL = sig type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, - concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list} + concl: cterm, schematics: ctyp TVars.table * cterm Vars.table} val focus_params: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_params_fixed: Proof.context -> int -> binding list option -> thm -> focus * thm val focus_prems: Proof.context -> int -> binding list option -> thm -> focus * thm val focus: Proof.context -> int -> binding list option -> thm -> focus * thm val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list -> int -> thm -> thm -> thm Seq.seq val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic val subgoal: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state val subgoal_cmd: Attrib.binding -> Attrib.binding option -> bool * (string option * Position.T) list -> Proof.state -> focus * Proof.state end; structure Subgoal: SUBGOAL = struct (* focus *) type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list, asms: cterm list, - concl: cterm, schematics: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list}; + concl: cterm, schematics: ctyp TVars.table * cterm Vars.table}; fun gen_focus (do_prems, do_concl) ctxt i bindings raw_st = let val st = raw_st |> Thm.solve_constraints |> Thm.transfer' ctxt |> Raw_Simplifier.norm_hhf_protect ctxt; val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt; val ((params, goal), ctxt2) = Variable.focus_cterm bindings (Thm.cprem_of st' i) ctxt1; val (asms, concl) = if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal) else ([], goal); val text = asms @ (if do_concl then [concl] else []); val ((_, inst), ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2; - val schematic_terms = Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; + val schematic_terms = Vars.map (K (Thm.cterm_of ctxt3)) inst; - val schematics = (TVars.dest schematic_types, schematic_terms); + val schematics = (schematic_types, schematic_terms); val asms' = map (Thm.instantiate_cterm schematics) asms; val concl' = Thm.instantiate_cterm schematics concl; val (prems, context) = Assumption.add_assumes asms' ctxt3; in ({context = context, params = params, prems = prems, asms = asms', concl = concl', schematics = schematics}, Goal.init concl') end; val focus_params = gen_focus (false, false); val focus_params_fixed = gen_focus (false, true); val focus_prems = gen_focus (true, false); val focus = gen_focus (true, true); (* lift and retrofit *) (* B [?'b, ?y] ---------------- B ['b, y params] *) fun lift_import idx params th ctxt = let val ((_, [th']), ctxt') = Variable.importT [th] ctxt; val Ts = map Thm.typ_of_cterm params; val ts = map Thm.term_of params; val prop = Thm.full_prop_of th'; val concl_vars = Vars.build (Vars.add_vars (Logic.strip_imp_concl prop)); val vars = Vars.build (Vars.add_vars prop) |> Vars.list_set; val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt'; fun var_inst v y = let val ((x, i), T) = v; val (U, args) = if Vars.defined concl_vars v then (T, []) else (Ts ---> T, ts); val u = Free (y, U); in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end; val (inst1, inst2) = split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys)); - - val th'' = Thm.instantiate ([], map (apfst (Term.dest_Var o Thm.term_of)) inst1) th'; + val inst = Vars.build (fold (Vars.add o apfst (Term.dest_Var o Thm.term_of)) inst1); + val th'' = Thm.instantiate (TVars.empty, inst) th'; in ((inst2, th''), ctxt'') end; (* [x, A x] : B x \ C ------------------ [\x. A x \ B x] : C *) fun lift_subgoals ctxt params asms th = let fun lift ct = fold_rev (Thm.all_name ctxt) params (Drule.list_implies (asms, ct)); val unlift = fold (Thm.elim_implies o Thm.assume) asms o Drule.forall_elim_list (map #2 params) o Thm.assume; val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th)); val th' = fold (Thm.elim_implies o unlift) subgoals th; in (subgoals, th') end; fun retrofit ctxt1 ctxt0 params asms i st1 st0 = let val idx = Thm.maxidx_of st0 + 1; val ps = map #2 params; val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1; val (subgoals, st3) = lift_subgoals ctxt2 params asms st2; val result = st3 |> Goal.conclude |> Drule.implies_intr_list asms |> Drule.forall_intr_list ps |> Drule.implies_intr_list subgoals |> fold_rev (Thm.forall_intr o #1) subgoal_inst |> fold (Thm.forall_elim o #2) subgoal_inst |> Thm.adjust_maxidx_thm idx |> singleton (Variable.export ctxt2 ctxt0); in Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false} (false, result, Thm.nprems_of st1) i st0 end; (* tacticals *) fun GEN_FOCUS flags tac ctxt i st = if Thm.nprems_of st < i then Seq.empty else let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags (ctxt |> Variable.set_bound_focus true) i NONE st; in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end; val FOCUS_PARAMS = GEN_FOCUS (false, false); val FOCUS_PARAMS_FIXED = GEN_FOCUS (false, true); val FOCUS_PREMS = GEN_FOCUS (true, false); val FOCUS = GEN_FOCUS (true, true); fun SUBPROOF tac ctxt = FOCUS (Seq.map (Goal.check_finished ctxt) oo tac) ctxt; (* Isar subgoal command *) local fun param_bindings ctxt (param_suffix, raw_param_specs) st = let val _ = if Thm.no_prems st then error "No subgoals!" else (); val subgoal = #1 (Logic.dest_implies (Thm.prop_of st)); val subgoal_params = map (apfst (Name.internal o Name.clean)) (Term.strip_all_vars subgoal) |> Term.variant_frees subgoal |> map #1; val n = length subgoal_params; val m = length raw_param_specs; val _ = m <= n orelse error ("Excessive subgoal parameter specification" ^ Position.here_list (map snd (drop n raw_param_specs))); val param_specs = raw_param_specs |> map (fn (NONE, _) => NONE | (SOME x, pos) => let val b = #1 (#1 (Proof_Context.cert_var (Binding.make (x, pos), NONE, NoSyn) ctxt)); val _ = Variable.check_name b; in SOME b end) |> param_suffix ? append (replicate (n - m) NONE); fun bindings (SOME x :: xs) (_ :: ys) = x :: bindings xs ys | bindings (NONE :: xs) (y :: ys) = Binding.name y :: bindings xs ys | bindings _ ys = map Binding.name ys; in bindings param_specs subgoal_params end; fun gen_subgoal prep_atts raw_result_binding raw_prems_binding param_specs state = let val _ = Proof.assert_backward state; val state1 = state |> Proof.refine_insert []; val {context = ctxt, facts = facts, goal = st} = Proof.raw_goal state1; val result_binding = apsnd (map (prep_atts ctxt)) raw_result_binding; val (prems_binding, do_prems) = (case raw_prems_binding of SOME (b, raw_atts) => ((b, map (prep_atts ctxt) raw_atts), true) | NONE => (Binding.empty_atts, false)); val (subgoal_focus, _) = (if do_prems then focus else focus_params_fixed) ctxt 1 (SOME (param_bindings ctxt param_specs st)) st; fun after_qed (ctxt'', [[result]]) = Proof.end_block #> (fn state' => let val ctxt' = Proof.context_of state'; val results' = Proof_Context.export ctxt'' ctxt' (Conjunction.elim_conjunctions result); in state' |> Proof.refine_primitive (fn _ => fn _ => retrofit ctxt'' ctxt' (#params subgoal_focus) (#asms subgoal_focus) 1 (Goal.protect 0 result) st |> Seq.hd) |> Proof.map_context (#2 o Proof_Context.note_thmss "" [(result_binding, [(results', [])])]) end) #> Proof.reset_facts #> Proof.enter_backward; in state1 |> Proof.enter_forward |> Proof.using_facts [] |> Proof.begin_block |> Proof.map_context (fn _ => #context subgoal_focus |> Proof_Context.note_thmss "" [(prems_binding, [(#prems subgoal_focus, [])])] |> #2) |> Proof.internal_goal (K (K ())) (Proof_Context.get_mode ctxt) true "subgoal" NONE after_qed [] [] [(Binding.empty_atts, [(Thm.term_of (#concl subgoal_focus), [])])] |> #2 |> Proof.using_facts facts |> pair subgoal_focus end; in val subgoal = gen_subgoal Attrib.attribute; val subgoal_cmd = gen_subgoal Attrib.attribute_cmd; end; end; val SUBPROOF = Subgoal.SUBPROOF; diff --git a/src/Pure/Proof/proof_checker.ML b/src/Pure/Proof/proof_checker.ML --- a/src/Pure/Proof/proof_checker.ML +++ b/src/Pure/Proof/proof_checker.ML @@ -1,152 +1,153 @@ (* Title: Pure/Proof/proof_checker.ML Author: Stefan Berghofer, TU Muenchen Simple proof checker based only on the core inference rules of Isabelle/Pure. *) signature PROOF_CHECKER = sig val thm_of_proof : theory -> Proofterm.proof -> thm end; structure Proof_Checker : PROOF_CHECKER = struct (***** construct a theorem out of a proof term *****) fun lookup_thm thy = let val tab = Symtab.build (Global_Theory.all_thms_of thy true |> fold_rev Symtab.update) in fn s => (case Symtab.lookup tab s of NONE => error ("Unknown theorem " ^ quote s) | SOME thm => thm) end; val beta_eta_convert = Conv.fconv_rule Drule.beta_eta_conversion; (* equality modulo renaming of type variables *) fun is_equal t t' = let val atoms = fold_types (fold_atyps (insert (op =))) t []; val atoms' = fold_types (fold_atyps (insert (op =))) t' [] in length atoms = length atoms' andalso map_types (map_atyps (the o AList.lookup (op =) (atoms ~~ atoms'))) t aconv t' end; fun pretty_prf thy vs Hs prf = let val prf' = prf |> Proofterm.prf_subst_bounds (map Free vs) |> Proofterm.prf_subst_pbounds (map (Hyp o Thm.prop_of) Hs) in (Proof_Syntax.pretty_proof (Syntax.init_pretty_global thy) prf', Syntax.pretty_term_global thy (Proofterm.prop_of prf')) end; fun pretty_term thy vs _ t = let val t' = subst_bounds (map Free vs, t) in (Syntax.pretty_term_global thy t', Syntax.pretty_typ_global thy (fastype_of t')) end; fun appl_error thy prt vs Hs s f a = let val (pp_f, pp_fT) = pretty_prf thy vs Hs f; val (pp_a, pp_aT) = prt thy vs Hs a in error (cat_lines [s, "", Pretty.string_of (Pretty.block [Pretty.str "Operator:", Pretty.brk 2, pp_f, Pretty.str " ::", Pretty.brk 1, pp_fT]), Pretty.string_of (Pretty.block [Pretty.str "Operand:", Pretty.brk 3, pp_a, Pretty.str " ::", Pretty.brk 1, pp_aT]), ""]) end; fun thm_of_proof thy = let val lookup = lookup_thm thy in fn prf => let val prf_names = Proofterm.fold_proof_terms Term.declare_term_frees prf Name.context; fun thm_of_atom thm Ts = let val tvars = build_rev (Term.add_tvars (Thm.full_prop_of thm)); val (fmap, thm') = Thm.varifyT_global' TFrees.empty thm; val ctye = (tvars @ map (fn ((_, S), ixn) => (ixn, S)) fmap ~~ map (Thm.global_ctyp_of thy) Ts); in - Thm.instantiate (ctye, []) (Thm.forall_intr_vars (Thm.forall_intr_frees thm')) + Thm.instantiate (TVars.make ctye, Vars.empty) + (Thm.forall_intr_vars (Thm.forall_intr_frees thm')) end; fun thm_of _ _ (PThm ({name, prop = prop', types = SOME Ts, ...}, _)) = let val thm = Thm.unconstrainT (Drule.implies_intr_hyps (lookup name)); val prop = Thm.prop_of thm; val _ = if is_equal prop prop' then () else error ("Duplicate use of theorem name " ^ quote name ^ "\n" ^ Syntax.string_of_term_global thy prop ^ "\n\n" ^ Syntax.string_of_term_global thy prop'); in thm_of_atom thm Ts end | thm_of _ _ (PAxm (name, _, SOME Ts)) = thm_of_atom (Thm.axiom thy name) Ts | thm_of _ Hs (PBound i) = nth Hs i | thm_of (vs, names) Hs (Abst (s, SOME T, prf)) = let val (x, names') = Name.variant s names; val thm = thm_of ((x, T) :: vs, names') Hs prf in Thm.forall_intr (Thm.global_cterm_of thy (Free (x, T))) thm end | thm_of (vs, names) Hs (prf % SOME t) = let val thm = thm_of (vs, names) Hs prf; val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); in Thm.forall_elim ct thm handle THM (s, _, _) => appl_error thy pretty_term vs Hs s prf t end | thm_of (vs, names) Hs (AbsP (_, SOME t, prf)) = let val ct = Thm.global_cterm_of thy (Term.subst_bounds (map Free vs, t)); val thm = thm_of (vs, names) (Thm.assume ct :: Hs) prf; in Thm.implies_intr ct thm end | thm_of vars Hs (prf %% prf') = let val thm = beta_eta_convert (thm_of vars Hs prf); val thm' = beta_eta_convert (thm_of vars Hs prf'); in Thm.implies_elim thm thm' handle THM (s, _, _) => appl_error thy pretty_prf (fst vars) Hs s prf prf' end | thm_of _ _ (Hyp t) = Thm.assume (Thm.global_cterm_of thy t) | thm_of _ _ (PClass (T, c)) = if Sign.of_sort thy (T, [c]) then Thm.of_class (Thm.global_ctyp_of thy T, c) else error ("thm_of_proof: bad PClass proof " ^ Syntax.string_of_term_global thy (Logic.mk_of_class (T, c))) | thm_of _ _ _ = error "thm_of_proof: partial proof term"; in beta_eta_convert (thm_of ([], prf_names) [] prf) end end; end; diff --git a/src/Pure/Tools/rule_insts.ML b/src/Pure/Tools/rule_insts.ML --- a/src/Pure/Tools/rule_insts.ML +++ b/src/Pure/Tools/rule_insts.ML @@ -1,353 +1,355 @@ (* Title: Pure/Tools/rule_insts.ML Author: Makarius Rule instantiations -- operations within implicit rule / subgoal context. *) signature RULE_INSTS = sig val where_rule: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> thm val of_rule: Proof.context -> string option list * string option list -> (binding * string option * mixfix) list -> thm -> thm val read_instantiate: Proof.context -> ((indexname * Position.T) * string) list -> string list -> thm -> thm val read_term: string -> Proof.context -> term * Proof.context val goal_context: term -> Proof.context -> (string * typ) list * Proof.context val res_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val eres_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val cut_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val forw_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val dres_inst_tac: Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic val thin_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val subgoal_tac: Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic val make_elim_preserve: Proof.context -> thm -> thm val method: (Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic) -> (Proof.context -> thm list -> int -> tactic) -> (Proof.context -> Proof.method) context_parser end; structure Rule_Insts: RULE_INSTS = struct (** read instantiations **) local fun error_var msg (xi, pos) = error (msg ^ quote (Term.string_of_vname xi) ^ Position.here pos); fun the_sort tvars (ai, pos) : sort = (case TVars.get_first (fn ((bi, S), _) => if ai = bi then SOME S else NONE) tvars of SOME S => S | NONE => error_var "No such type variable in theorem: " (ai, pos)); fun the_type vars (xi, pos) : typ = (case Vartab.lookup vars xi of SOME T => T | NONE => error_var "No such variable in theorem: " (xi, pos)); fun read_type ctxt tvars ((xi, pos), s) = let val S = the_sort tvars (xi, pos); val T = Syntax.read_typ ctxt s; in if Sign.of_sort (Proof_Context.theory_of ctxt) (T, S) then ((xi, S), T) else error_var "Bad sort for instantiation of type variable: " (xi, pos) end; fun make_instT f (tvars: TVars.set) = let fun add v = let val T = TVar v; val T' = f T; in if T = T' then I else cons (v, T') end; in TVars.fold (add o #1) tvars [] end; fun make_inst f vars = let fun add v = let val t = Var v; val t' = f t; in if t aconv t' then I else cons (v, t') end; in fold add vars [] end; fun read_terms ss Ts ctxt = let fun parse T = if T = propT then Syntax.parse_prop ctxt else Syntax.parse_term ctxt; val (ts, ctxt') = fold_map Variable.fix_dummy_patterns (map2 parse Ts ss) ctxt; val ts' = map2 (Type.constraint o Type_Infer.paramify_vars) Ts ts |> Syntax.check_terms ctxt' |> Variable.polymorphic ctxt'; val Ts' = map Term.fastype_of ts'; val tyenv = Vartab.build (fold (Sign.typ_match (Proof_Context.theory_of ctxt)) (Ts ~~ Ts')); val tyenv' = Vartab.fold (fn (xi, (S, T)) => cons ((xi, S), T)) tyenv []; in ((ts', tyenv'), ctxt') end; in fun read_term s ctxt = let val (t, ctxt') = Variable.fix_dummy_patterns (Syntax.parse_term ctxt s) ctxt; val t' = Syntax.check_term ctxt' t; in (t', ctxt') end; fun read_insts thm raw_insts raw_fixes ctxt = let val (type_insts, term_insts) = List.partition (fn (((x, _), _), _) => String.isPrefix "'" x) raw_insts; val tvars = TVars.build (Thm.fold_terms {hyps = false} TVars.add_tvars thm); val vars = Vars.build (Thm.fold_terms {hyps = false} Vars.add_vars thm); (*eigen-context*) val (_, ctxt1) = ctxt |> TVars.fold (Variable.declare_internal o Logic.mk_type o TVar o #1) tvars |> Vars.fold (Variable.declare_internal o Var o #1) vars |> Proof_Context.add_fixes_cmd raw_fixes; (*explicit type instantiations*) val instT1 = Term_Subst.instantiateT (TVars.make (map (read_type ctxt1 tvars) type_insts)); val vars1 = Vartab.build (vars |> Vars.fold (fn ((v, T), _) => Vartab.insert (K true) (v, instT1 T))); (*term instantiations*) val (xs, ss) = split_list term_insts; val Ts = map (the_type vars1) xs; val ((ts, inferred), ctxt2) = read_terms ss Ts ctxt1; (*implicit type instantiations*) val instT2 = Term_Subst.instantiateT (TVars.make inferred); val vars2 = Vartab.fold (fn (v, T) => cons (v, instT2 T)) vars1 []; val inst2 = Term_Subst.instantiate (TVars.empty, Vars.build (fold2 (fn (xi, _) => fn t => Vars.add ((xi, Term.fastype_of t), t)) xs ts)) #> Envir.beta_norm; val inst_tvars = make_instT (instT2 o instT1) tvars; val inst_vars = make_inst inst2 vars2; in ((inst_tvars, inst_vars), ctxt2) end; end; (** forward rules **) fun where_rule ctxt raw_insts raw_fixes thm = let val ((inst_tvars, inst_vars), ctxt') = read_insts thm raw_insts raw_fixes ctxt; in thm |> Drule.instantiate_normalize - (map (apsnd (Thm.ctyp_of ctxt')) inst_tvars, - map (apsnd (Thm.cterm_of ctxt')) inst_vars) + (TVars.make (map (apsnd (Thm.ctyp_of ctxt')) inst_tvars), + Vars.make (map (apsnd (Thm.cterm_of ctxt')) inst_vars)) |> singleton (Variable.export ctxt' ctxt) |> Rule_Cases.save thm end; fun of_rule ctxt (args, concl_args) fixes thm = let fun zip_vars _ [] = [] | zip_vars (_ :: xs) (NONE :: rest) = zip_vars xs rest | zip_vars ((x, _) :: xs) (SOME t :: rest) = ((x, Position.none), t) :: zip_vars xs rest | zip_vars [] _ = error "More instantiations than variables in theorem"; val insts = zip_vars (Vars.build (Vars.add_vars (Thm.full_prop_of thm)) |> Vars.list_set) args @ zip_vars (Vars.build (Vars.add_vars (Thm.concl_of thm)) |> Vars.list_set) concl_args; in where_rule ctxt insts fixes thm end; fun read_instantiate ctxt insts xs = where_rule ctxt insts (map (fn x => (Binding.name x, NONE, NoSyn)) xs); (** attributes **) (* where: named instantiation *) val named_insts = Parse.and_list1 (Parse.position Args.var -- (Args.$$$ "=" |-- Parse.!!! Args.embedded_inner_syntax)) -- Parse.for_fixes; val _ = Theory.setup (Attrib.setup \<^binding>\where\ (Scan.lift named_insts >> (fn args => Thm.rule_attribute [] (fn context => uncurry (where_rule (Context.proof_of context)) args))) "named instantiation of theorem"); (* of: positional instantiation (terms only) *) local val inst = Args.maybe Args.embedded_inner_syntax; val concl = Args.$$$ "concl" -- Args.colon; val insts = Scan.repeat (Scan.unless concl inst) -- Scan.optional (concl |-- Scan.repeat inst) []; in val _ = Theory.setup (Attrib.setup \<^binding>\of\ (Scan.lift (insts -- Parse.for_fixes) >> (fn args => Thm.rule_attribute [] (fn context => uncurry (of_rule (Context.proof_of context)) args))) "positional instantiation of theorem"); end; (** tactics **) (* goal context *) fun goal_context goal ctxt = let val ((_, params), ctxt') = ctxt |> Variable.declare_constraints goal |> Variable.improper_fixes |> Variable.focus_params NONE goal ||> Variable.restore_proper_fixes ctxt; in (params, ctxt') end; (* resolution after lifting and instantiation; may refer to parameters of the subgoal *) fun bires_inst_tac bires_flag ctxt raw_insts raw_fixes thm i st = CSUBGOAL (fn (cgoal, _) => let (*goal context*) val (params, goal_ctxt) = goal_context (Thm.term_of cgoal) ctxt; val paramTs = map #2 params; (*instantiation context*) val ((inst_tvars, inst_vars), inst_ctxt) = read_insts thm raw_insts raw_fixes goal_ctxt; val fixed = map #1 (fold (Variable.add_newly_fixed inst_ctxt goal_ctxt o #2) inst_vars []); (* lift and instantiate rule *) val inc = Thm.maxidx_of st + 1; val lift_type = Logic.incr_tvar inc; fun lift_var ((a, j), T) = ((a, j + inc), paramTs ---> lift_type T); fun lift_term t = fold_rev Term.absfree params (Logic.incr_indexes (fixed, paramTs, inc) t); - val inst_tvars' = inst_tvars - |> map (fn (((a, i), S), T) => (((a, i + inc), S), Thm.ctyp_of inst_ctxt (lift_type T))); - val inst_vars' = inst_vars - |> map (fn (v, t) => (lift_var v, Thm.cterm_of inst_ctxt (lift_term t))); + val inst_tvars' = + TVars.build (inst_tvars |> fold (fn (((a, i), S), T) => + TVars.add (((a, i + inc), S), Thm.ctyp_of inst_ctxt (lift_type T)))); + val inst_vars' = + Vars.build (inst_vars |> fold (fn (v, t) => + Vars.add (lift_var v, Thm.cterm_of inst_ctxt (lift_term t)))); val thm' = Thm.lift_rule cgoal thm |> Drule.instantiate_normalize (inst_tvars', inst_vars') |> singleton (Variable.export inst_ctxt ctxt); in compose_tac ctxt (bires_flag, thm', Thm.nprems_of thm) i end) i st; val res_inst_tac = bires_inst_tac false; val eres_inst_tac = bires_inst_tac true; (* forward resolution *) fun make_elim_preserve ctxt rl = let val maxidx = Thm.maxidx_of rl; fun var x = ((x, 0), propT); fun cvar xi = Thm.cterm_of ctxt (Var (xi, propT)); val revcut_rl' = - Drule.instantiate_normalize ([], [(var "V", cvar ("V", maxidx + 1)), - (var "W", cvar ("W", maxidx + 1))]) Drule.revcut_rl; + Drule.revcut_rl |> Drule.instantiate_normalize + (TVars.empty, Vars.make [(var "V", cvar ("V", maxidx + 1)), (var "W", cvar ("W", maxidx + 1))]); in (case Seq.list_of (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false} (false, rl, Thm.nprems_of rl) 1 revcut_rl') of [th] => th | _ => raise THM ("make_elim_preserve", 1, [rl])) end; (*instantiate and cut -- for atomic fact*) fun cut_inst_tac ctxt insts fixes rule = res_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule); (*forward tactic applies a rule to an assumption without deleting it*) fun forw_inst_tac ctxt insts fixes rule = cut_inst_tac ctxt insts fixes rule THEN' assume_tac ctxt; (*dresolve tactic applies a rule to replace an assumption*) fun dres_inst_tac ctxt insts fixes rule = eres_inst_tac ctxt insts fixes (make_elim_preserve ctxt rule); (* derived tactics *) (*deletion of an assumption*) fun thin_tac ctxt s fixes = eres_inst_tac ctxt [((("V", 0), Position.none), s)] fixes Drule.thin_rl; (*Introduce the given proposition as lemma and subgoal*) fun subgoal_tac ctxt A fixes = DETERM o res_inst_tac ctxt [((("psi", 0), Position.none), A)] fixes cut_rl; (* method wrapper *) fun method inst_tac tac = Args.goal_spec -- Scan.optional (Scan.lift (named_insts --| Args.$$$ "in")) ([], []) -- Attrib.thms >> (fn ((quant, (insts, fixes)), thms) => fn ctxt => METHOD (fn facts => if null insts andalso null fixes then quant (Method.insert_tac ctxt facts THEN' tac ctxt thms) else (case thms of [thm] => quant (Method.insert_tac ctxt facts THEN' inst_tac ctxt insts fixes thm) | _ => error "Cannot have instantiations with multiple rules"))); (* setup *) (*warning: rule_tac etc. refer to dynamic subgoal context!*) val _ = Theory.setup (Method.setup \<^binding>\rule_tac\ (method res_inst_tac resolve_tac) "apply rule (dynamic instantiation)" #> Method.setup \<^binding>\erule_tac\ (method eres_inst_tac eresolve_tac) "apply rule in elimination manner (dynamic instantiation)" #> Method.setup \<^binding>\drule_tac\ (method dres_inst_tac dresolve_tac) "apply rule in destruct manner (dynamic instantiation)" #> Method.setup \<^binding>\frule_tac\ (method forw_inst_tac forward_tac) "apply rule in forward manner (dynamic instantiation)" #> Method.setup \<^binding>\cut_tac\ (method cut_inst_tac (K cut_rules_tac)) "cut rule (dynamic instantiation)" #> Method.setup \<^binding>\subgoal_tac\ (Args.goal_spec -- Scan.lift (Scan.repeat1 Args.embedded_inner_syntax -- Parse.for_fixes) >> (fn (quant, (props, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (EVERY' (map (fn prop => subgoal_tac ctxt prop fixes) props)))) "insert subgoal (dynamic instantiation)" #> Method.setup \<^binding>\thin_tac\ (Args.goal_spec -- Scan.lift (Args.embedded_inner_syntax -- Parse.for_fixes) >> (fn (quant, (prop, fixes)) => fn ctxt => SIMPLE_METHOD'' quant (thin_tac ctxt prop fixes))) "remove premise (dynamic instantiation)"); end; diff --git a/src/Pure/conjunction.ML b/src/Pure/conjunction.ML --- a/src/Pure/conjunction.ML +++ b/src/Pure/conjunction.ML @@ -1,193 +1,194 @@ (* Title: Pure/conjunction.ML Author: Makarius Meta-level conjunction. *) signature CONJUNCTION = sig val conjunction: cterm val mk_conjunction: cterm * cterm -> cterm val mk_conjunction_balanced: cterm list -> cterm val dest_conjunction: cterm -> cterm * cterm val dest_conjunctions: cterm -> cterm list val cong: thm -> thm -> thm val convs: (cterm -> thm) -> cterm -> thm val conjunctionD1: thm val conjunctionD2: thm val conjunctionI: thm val intr: thm -> thm -> thm val intr_balanced: thm list -> thm val elim: thm -> thm * thm val elim_conjunctions: thm -> thm list val elim_balanced: int -> thm -> thm list val curry_balanced: int -> thm -> thm val uncurry_balanced: int -> thm -> thm end; structure Conjunction: CONJUNCTION = struct (** abstract syntax **) fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t; val read_prop = certify o Simple_Syntax.read_prop; val true_prop = certify Logic.true_prop; val conjunction = certify Logic.conjunction; fun mk_conjunction (A, B) = Thm.apply (Thm.apply conjunction A) B; fun mk_conjunction_balanced [] = true_prop | mk_conjunction_balanced ts = Balanced_Tree.make mk_conjunction ts; fun dest_conjunction ct = (case Thm.term_of ct of (Const ("Pure.conjunction", _) $ _ $ _) => Thm.dest_binop ct | _ => raise TERM ("dest_conjunction", [Thm.term_of ct])); fun dest_conjunctions ct = (case try dest_conjunction ct of NONE => [ct] | SOME (A, B) => dest_conjunctions A @ dest_conjunctions B); (** derived rules **) (* conversion *) val cong = Thm.combination o Thm.combination (Thm.reflexive conjunction); fun convs cv ct = (case try dest_conjunction ct of NONE => cv ct | SOME (A, B) => cong (convs cv A) (convs cv B)); (* intro/elim *) local val A = read_prop "A" and vA = (("A", 0), propT); val B = read_prop "B" and vB = (("B", 0), propT); val C = read_prop "C"; val ABC = read_prop "A \ B \ C"; val A_B = read_prop "A &&& B"; val conjunction_def = Thm.unvarify_axiom (Context.the_global_context ()) "Pure.conjunction_def"; fun conjunctionD which = Drule.implies_intr_list [A, B] (Thm.assume (which (A, B))) COMP Thm.forall_elim_vars 0 (Thm.equal_elim conjunction_def (Thm.assume A_B)); in val conjunctionD1 = Drule.store_standard_thm (Binding.make ("conjunctionD1", \<^here>)) (conjunctionD #1); val conjunctionD2 = Drule.store_standard_thm (Binding.make ("conjunctionD2", \<^here>)) (conjunctionD #2); val conjunctionI = Drule.store_standard_thm (Binding.make ("conjunctionI", \<^here>)) (Drule.implies_intr_list [A, B] (Thm.equal_elim (Thm.symmetric conjunction_def) (Thm.forall_intr C (Thm.implies_intr ABC (Drule.implies_elim_list (Thm.assume ABC) [Thm.assume A, Thm.assume B]))))); fun intr tha thb = Thm.implies_elim (Thm.implies_elim - (Thm.instantiate ([], [(vA, Thm.cprop_of tha), (vB, Thm.cprop_of thb)]) conjunctionI) + (Thm.instantiate (TVars.empty, Vars.make [(vA, Thm.cprop_of tha), (vB, Thm.cprop_of thb)]) + conjunctionI) tha) thb; fun elim th = let val (A, B) = dest_conjunction (Thm.cprop_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); - val inst = Thm.instantiate ([], [(vA, A), (vB, B)]); + val inst = Thm.instantiate (TVars.empty, Vars.make [(vA, A), (vB, B)]); in (Thm.implies_elim (inst conjunctionD1) th, Thm.implies_elim (inst conjunctionD2) th) end; end; fun elim_conjunctions th = (case try elim th of NONE => [th] | SOME (th1, th2) => elim_conjunctions th1 @ elim_conjunctions th2); (* balanced conjuncts *) fun intr_balanced [] = asm_rl | intr_balanced ths = Balanced_Tree.make (uncurry intr) ths; fun elim_balanced 0 _ = [] | elim_balanced n th = Balanced_Tree.dest elim n th; (* currying *) local val bootstrap_thy = Context.the_global_context (); fun conjs n = let val As = map (fn A => Thm.global_cterm_of bootstrap_thy (Free (A, propT))) (Name.invent Name.context "A" n); in (As, mk_conjunction_balanced As) end; val B = read_prop "B"; fun comp_rule th rule = Thm.adjust_maxidx_thm ~1 (th COMP (rule |> Thm.forall_intr_frees |> Thm.forall_elim_vars (Thm.maxidx_of th + 1))); in (* A1 &&& ... &&& An \ B ----------------------- A1 \ ... \ An \ B *) fun curry_balanced n th = if n < 2 then th else let val (As, C) = conjs n; val D = Drule.mk_implies (C, B); in comp_rule th (Thm.implies_elim (Thm.assume D) (intr_balanced (map Thm.assume As)) |> Drule.implies_intr_list (D :: As)) end; (* A1 \ ... \ An \ B ----------------------- A1 &&& ... &&& An \ B *) fun uncurry_balanced n th = if n < 2 then th else let val (As, C) = conjs n; val D = Drule.list_implies (As, B); in comp_rule th (Drule.implies_elim_list (Thm.assume D) (elim_balanced n (Thm.assume C)) |> Drule.implies_intr_list [D, C]) end; end; end; diff --git a/src/Pure/drule.ML b/src/Pure/drule.ML --- a/src/Pure/drule.ML +++ b/src/Pure/drule.ML @@ -1,839 +1,838 @@ (* Title: Pure/drule.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Derived rules and other operations on theorems. *) infix 0 RL RLN MRS OF COMP INCR_COMP COMP_INCR; signature BASIC_DRULE = sig val mk_implies: cterm * cterm -> cterm val list_implies: cterm list * cterm -> cterm val strip_imp_prems: cterm -> cterm list val strip_imp_concl: cterm -> cterm val cprems_of: thm -> cterm list val forall_intr_list: cterm list -> thm -> thm val forall_elim_list: cterm list -> thm -> thm val lift_all: Proof.context -> cterm -> thm -> thm val implies_elim_list: thm -> thm list -> thm val implies_intr_list: cterm list -> thm -> thm - val instantiate_normalize: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> - thm -> thm + val instantiate_normalize: ctyp TVars.table * cterm Vars.table -> thm -> thm val instantiate'_normalize: ctyp option list -> cterm option list -> thm -> thm val infer_instantiate_types: Proof.context -> ((indexname * typ) * cterm) list -> thm -> thm val infer_instantiate: Proof.context -> (indexname * cterm) list -> thm -> thm val infer_instantiate': Proof.context -> cterm option list -> thm -> thm val zero_var_indexes_list: thm list -> thm list val zero_var_indexes: thm -> thm val implies_intr_hyps: thm -> thm val rotate_prems: int -> thm -> thm val rearrange_prems: int list -> thm -> thm val RLN: thm list * (int * thm list) -> thm list val RL: thm list * thm list -> thm list val MRS: thm list * thm -> thm val OF: thm * thm list -> thm val COMP: thm * thm -> thm val INCR_COMP: thm * thm -> thm val COMP_INCR: thm * thm -> thm val size_of_thm: thm -> int val reflexive_thm: thm val symmetric_thm: thm val transitive_thm: thm val extensional: thm -> thm val asm_rl: thm val cut_rl: thm val revcut_rl: thm val thin_rl: thm end; signature DRULE = sig include BASIC_DRULE val outer_params: term -> (string * typ) list val generalize: Names.set * Names.set -> thm -> thm val list_comb: cterm * cterm list -> cterm val strip_comb: cterm -> cterm * cterm list val beta_conv: cterm -> cterm -> cterm val flexflex_unique: Proof.context option -> thm -> thm val export_without_context: thm -> thm val export_without_context_open: thm -> thm val store_thm: binding -> thm -> thm val store_standard_thm: binding -> thm -> thm val store_thm_open: binding -> thm -> thm val store_standard_thm_open: binding -> thm -> thm val multi_resolve: Proof.context option -> thm list -> thm -> thm Seq.seq val multi_resolves: Proof.context option -> thm list -> thm list -> thm Seq.seq val compose: thm * int * thm -> thm val equals_cong: thm val imp_cong: thm val swap_prems_eq: thm val imp_cong_rule: thm -> thm -> thm val arg_cong_rule: cterm -> thm -> thm val binop_cong_rule: cterm -> thm -> thm -> thm val fun_cong_rule: thm -> cterm -> thm val beta_eta_conversion: cterm -> thm val eta_contraction_rule: thm -> thm val norm_hhf_eq: thm val norm_hhf_eqs: thm list val is_norm_hhf: term -> bool val norm_hhf: theory -> term -> term val norm_hhf_cterm: Proof.context -> cterm -> cterm val protect: cterm -> cterm val protectI: thm val protectD: thm val protect_cong: thm val implies_intr_protected: cterm list -> thm -> thm val termI: thm val mk_term: cterm -> thm val dest_term: thm -> cterm val cterm_rule: (thm -> thm) -> cterm -> cterm val add_frees_cterm: cterm -> Cterms.set -> Cterms.set val add_vars_cterm: cterm -> Cterms.set -> Cterms.set val dummy_thm: thm val free_dummy_thm: thm val is_sort_constraint: term -> bool val sort_constraintI: thm val sort_constraint_eq: thm val with_subgoal: int -> (thm -> thm) -> thm -> thm val comp_no_flatten: thm * int -> int -> thm -> thm val rename_bvars: (string * string) list -> thm -> thm val rename_bvars': string option list -> thm -> thm val incr_indexes: thm -> thm -> thm val incr_indexes2: thm -> thm -> thm -> thm val triv_forall_equality: thm val distinct_prems_rl: thm val equal_intr_rule: thm val equal_elim_rule1: thm val equal_elim_rule2: thm val remdups_rl: thm val abs_def: thm -> thm end; structure Drule: DRULE = struct (** some cterm->cterm operations: faster than calling cterm_of! **) (* A1\...An\B goes to [A1,...,An], where B is not an implication *) fun strip_imp_prems ct = let val (cA, cB) = Thm.dest_implies ct in cA :: strip_imp_prems cB end handle TERM _ => []; (* A1\...An\B goes to B, where B is not an implication *) fun strip_imp_concl ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => strip_imp_concl (Thm.dest_arg ct) | _ => ct); (*The premises of a theorem, as a cterm list*) val cprems_of = strip_imp_prems o Thm.cprop_of; fun certify t = Thm.global_cterm_of (Context.the_global_context ()) t; val implies = certify Logic.implies; fun mk_implies (A, B) = Thm.apply (Thm.apply implies A) B; (*cterm version of list_implies: [A1,...,An], B goes to \A1;...;An\\B *) fun list_implies([], B) = B | list_implies(A::As, B) = mk_implies (A, list_implies(As,B)); (*cterm version of list_comb: maps (f, [t1,...,tn]) to f(t1,...,tn) *) fun list_comb (f, []) = f | list_comb (f, t::ts) = list_comb (Thm.apply f t, ts); (*cterm version of strip_comb: maps f(t1,...,tn) to (f, [t1,...,tn]) *) fun strip_comb ct = let fun stripc (p as (ct, cts)) = let val (ct1, ct2) = Thm.dest_comb ct in stripc (ct1, ct2 :: cts) end handle CTERM _ => p in stripc (ct, []) end; (*Beta-conversion for cterms, where x is an abstraction. Simply returns the rhs of the meta-equality returned by the beta_conversion rule.*) fun beta_conv x y = Thm.dest_arg (Thm.cprop_of (Thm.beta_conversion false (Thm.apply x y))); (** Standardization of rules **) (*Generalization over a list of variables*) val forall_intr_list = fold_rev Thm.forall_intr; fun outer_params t = let val vs = Term.strip_all_vars t in Name.variant_list [] (map (Name.clean o #1) vs) ~~ map #2 vs end; (*lift vars wrt. outermost goal parameters -- reverses the effect of gen_all modulo higher-order unification*) fun lift_all ctxt raw_goal raw_th = let val thy = Proof_Context.theory_of ctxt; val goal = Thm.transfer_cterm thy raw_goal; val th = Thm.transfer thy raw_th; val maxidx = Thm.maxidx_of th; val ps = outer_params (Thm.term_of goal) |> map (fn (x, T) => Var ((x, maxidx + 1), Logic.incr_tvar (maxidx + 1) T)); val Ts = map Term.fastype_of ps; val inst = Vars.build (th |> (Thm.fold_terms {hyps = false} o Term.fold_aterms) (fn t => fn inst => (case t of Var (xi, T) => if Vars.defined inst (xi, T) then inst else let val ct = Thm.cterm_of ctxt (Term.list_comb (Var (xi, Ts ---> T), ps)) in Vars.add ((xi, T), ct) inst end | _ => inst))); in th - |> Thm.instantiate ([], Vars.dest inst) + |> Thm.instantiate (TVars.empty, inst) |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) ps end; (*direct generalization*) fun generalize names th = Thm.generalize names (Thm.maxidx_of th + 1) th; (*specialization over a list of cterms*) val forall_elim_list = fold Thm.forall_elim; (*maps A1,...,An |- B to \A1;...;An\ \ B*) val implies_intr_list = fold_rev Thm.implies_intr; (*maps \A1;...;An\ \ B and [A1,...,An] to B*) fun implies_elim_list impth ths = fold Thm.elim_implies ths impth; (*Reset Var indexes to zero, renaming to preserve distinctness*) fun zero_var_indexes_list [] = [] | zero_var_indexes_list ths = let val (instT, inst) = Term_Subst.zero_var_indexes_inst Name.context (map Thm.full_prop_of ths); val tvars = TVars.build (fold Thm.add_tvars ths); val the_tvar = the o TVars.lookup tvars; - val instT' = - build (instT |> TVars.fold (fn (v, TVar (b, _)) => - cons (v, Thm.rename_tvar b (the_tvar v)))); + val instT' = instT |> TVars.map (fn v => fn TVar (b, _) => Thm.rename_tvar b (the_tvar v)); - val vars = Vars.build (fold (Thm.add_vars o Thm.instantiate (instT', [])) ths); + val vars = Vars.build (fold (Thm.add_vars o Thm.instantiate (instT', Vars.empty)) ths); val the_var = the o Vars.lookup vars; val inst' = - build (inst |> Vars.fold (fn (v, Var (b, _)) => - cons (v, Thm.var (b, Thm.ctyp_of_cterm (the_var v))))); + inst |> Vars.map (fn v => fn Var (b, _) => Thm.var (b, Thm.ctyp_of_cterm (the_var v))); in map (Thm.adjust_maxidx_thm ~1 o Thm.instantiate (instT', inst')) ths end; val zero_var_indexes = singleton zero_var_indexes_list; (** Standard form of object-rule: no hypotheses, flexflex constraints, Frees, or outer quantifiers; all generality expressed by Vars of index 0.**) (*Discharge all hypotheses.*) fun implies_intr_hyps th = fold Thm.implies_intr (Thm.chyps_of th) th; (*Squash a theorem's flexflex constraints provided it can be done uniquely. This step can lose information.*) fun flexflex_unique opt_ctxt th = if null (Thm.tpairs_of th) then th else (case distinct Thm.eq_thm (Seq.list_of (Thm.flexflex_rule opt_ctxt th)) of [th] => th | [] => raise THM ("flexflex_unique: impossible constraints", 0, [th]) | _ => raise THM ("flexflex_unique: multiple unifiers", 0, [th])); (* old-style export without context *) val export_without_context_open = implies_intr_hyps #> Thm.forall_intr_frees #> `Thm.maxidx_of #-> (fn maxidx => Thm.forall_elim_vars (maxidx + 1) #> Thm.strip_shyps #> zero_var_indexes #> Thm.varifyT_global); val export_without_context = flexflex_unique NONE #> export_without_context_open #> Thm.close_derivation \<^here>; (*Rotates a rule's premises to the left by k*) fun rotate_prems 0 = I | rotate_prems k = Thm.permute_prems 0 k; fun with_subgoal i f = rotate_prems (i - 1) #> f #> rotate_prems (1 - i); (*Permute prems, where the i-th position in the argument list (counting from 0) gives the position within the original thm to be transferred to position i. Any remaining trailing positions are left unchanged.*) val rearrange_prems = let fun rearr new [] thm = thm | rearr new (p :: ps) thm = rearr (new + 1) (map (fn q => if new <= q andalso q < p then q + 1 else q) ps) (Thm.permute_prems (new + 1) (new - p) (Thm.permute_prems new (p - new) thm)) in rearr 0 end; (*Resolution: multiple arguments, multiple results*) local fun res opt_ctxt th i rule = (Thm.biresolution opt_ctxt false [(false, th)] i rule handle THM _ => Seq.empty) |> Seq.map Thm.solve_constraints; fun multi_res _ _ [] rule = Seq.single rule | multi_res opt_ctxt i (th :: ths) rule = Seq.maps (res opt_ctxt th i) (multi_res opt_ctxt (i + 1) ths rule); in fun multi_resolve opt_ctxt = multi_res opt_ctxt 1; fun multi_resolves opt_ctxt facts rules = Seq.maps (multi_resolve opt_ctxt facts) (Seq.of_list rules); end; (*For joining lists of rules*) fun thas RLN (i, thbs) = let val resolve = Thm.biresolution NONE false (map (pair false) thas) i fun resb thb = Seq.list_of (resolve thb) handle THM _ => [] in maps resb thbs |> map Thm.solve_constraints end; fun thas RL thbs = thas RLN (1, thbs); (*Isar-style multi-resolution*) fun bottom_rl OF rls = (case Seq.chop 2 (multi_resolve NONE rls bottom_rl) of ([th], _) => Thm.solve_constraints th | ([], _) => raise THM ("OF: no unifiers", 0, bottom_rl :: rls) | _ => raise THM ("OF: multiple unifiers", 0, bottom_rl :: rls)); (*Resolve a list of rules against bottom_rl from right to left; makes proof trees*) fun rls MRS bottom_rl = bottom_rl OF rls; (*compose Q and \...,Qi,Q(i+1),...\ \ R to \...,Q(i+1),...\ \ R with no lifting or renaming! Q may contain \ or meta-quantifiers ALWAYS deletes premise i *) fun compose (tha, i, thb) = Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, tha, 0) i thb |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("compose: unique result expected", i, [tha, thb])); (** theorem equality **) (*Useful "distance" function for BEST_FIRST*) val size_of_thm = size_of_term o Thm.full_prop_of; (*** Meta-Rewriting Rules ***) val read_prop = certify o Simple_Syntax.read_prop; fun store_thm name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm (name, th))); fun store_thm_open name th = Context.>>> (Context.map_theory_result (Global_Theory.store_thm_open (name, th))); fun store_standard_thm name th = store_thm name (export_without_context th); fun store_standard_thm_open name th = store_thm_open name (export_without_context_open th); val reflexive_thm = let val cx = certify (Var(("x",0),TVar(("'a",0),[]))) in store_standard_thm_open (Binding.make ("reflexive", \<^here>)) (Thm.reflexive cx) end; val symmetric_thm = let val xy = read_prop "x::'a \ y::'a"; val thm = Thm.implies_intr xy (Thm.symmetric (Thm.assume xy)); in store_standard_thm_open (Binding.make ("symmetric", \<^here>)) thm end; val transitive_thm = let val xy = read_prop "x::'a \ y::'a"; val yz = read_prop "y::'a \ z::'a"; val xythm = Thm.assume xy; val yzthm = Thm.assume yz; val thm = Thm.implies_intr yz (Thm.transitive xythm yzthm); in store_standard_thm_open (Binding.make ("transitive", \<^here>)) thm end; fun extensional eq = let val eq' = Thm.abstract_rule "x" (Thm.dest_arg (fst (Thm.dest_equals (Thm.cprop_of eq)))) eq in Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of eq')) eq' end; val equals_cong = store_standard_thm_open (Binding.make ("equals_cong", \<^here>)) (Thm.reflexive (read_prop "x::'a \ y::'a")); val imp_cong = let val ABC = read_prop "A \ B::prop \ C::prop" val AB = read_prop "A \ B" val AC = read_prop "A \ C" val A = read_prop "A" in store_standard_thm_open (Binding.make ("imp_cong", \<^here>)) (Thm.implies_intr ABC (Thm.equal_intr (Thm.implies_intr AB (Thm.implies_intr A (Thm.equal_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.implies_elim (Thm.assume AB) (Thm.assume A))))) (Thm.implies_intr AC (Thm.implies_intr A (Thm.equal_elim (Thm.symmetric (Thm.implies_elim (Thm.assume ABC) (Thm.assume A))) (Thm.implies_elim (Thm.assume AC) (Thm.assume A))))))) end; val swap_prems_eq = let val ABC = read_prop "A \ B \ C" val BAC = read_prop "B \ A \ C" val A = read_prop "A" val B = read_prop "B" in store_standard_thm_open (Binding.make ("swap_prems_eq", \<^here>)) (Thm.equal_intr (Thm.implies_intr ABC (Thm.implies_intr B (Thm.implies_intr A (Thm.implies_elim (Thm.implies_elim (Thm.assume ABC) (Thm.assume A)) (Thm.assume B))))) (Thm.implies_intr BAC (Thm.implies_intr A (Thm.implies_intr B (Thm.implies_elim (Thm.implies_elim (Thm.assume BAC) (Thm.assume B)) (Thm.assume A)))))) end; val imp_cong_rule = Thm.combination o Thm.combination (Thm.reflexive implies); fun arg_cong_rule ct th = Thm.combination (Thm.reflexive ct) th; (*AP_TERM in LCF/HOL*) fun fun_cong_rule th ct = Thm.combination th (Thm.reflexive ct); (*AP_THM in LCF/HOL*) fun binop_cong_rule ct th1 th2 = Thm.combination (arg_cong_rule ct th1) th2; fun beta_eta_conversion ct = let val thm = Thm.beta_conversion true ct in Thm.transitive thm (Thm.eta_conversion (Thm.rhs_of thm)) end; (*Contract all eta-redexes in the theorem, lest they give rise to needless abstractions*) fun eta_contraction_rule th = Thm.equal_elim (Thm.eta_conversion (Thm.cprop_of th)) th; (* abs_def *) (* f ?x1 ... ?xn \ u -------------------- f \ \x1 ... xn. u *) local fun contract_lhs th = Thm.transitive (Thm.symmetric (beta_eta_conversion (#1 (Thm.dest_equals (Thm.cprop_of th))))) th; fun var_args ct = (case try Thm.dest_comb ct of SOME (f, arg) => (case Thm.term_of arg of Var ((x, _), _) => update (eq_snd (op aconvc)) (x, arg) (var_args f) | _ => []) | NONE => []); in fun abs_def th = let val th' = contract_lhs th; val args = var_args (Thm.lhs_of th'); in contract_lhs (fold (uncurry Thm.abstract_rule) args th') end; end; (*** Some useful meta-theorems ***) (*The rule V/V, obtains assumption solving for eresolve_tac*) val asm_rl = store_standard_thm_open (Binding.make ("asm_rl", \<^here>)) (Thm.trivial (read_prop "?psi")); (*Meta-level cut rule: \V \ W; V\ \ W *) val cut_rl = store_standard_thm_open (Binding.make ("cut_rl", \<^here>)) (Thm.trivial (read_prop "?psi \ ?theta")); (*Generalized elim rule for one conclusion; cut_rl with reversed premises: \PROP V; PROP V \ PROP W\ \ PROP W *) val revcut_rl = let val V = read_prop "V"; val VW = read_prop "V \ W"; in store_standard_thm_open (Binding.make ("revcut_rl", \<^here>)) (Thm.implies_intr V (Thm.implies_intr VW (Thm.implies_elim (Thm.assume VW) (Thm.assume V)))) end; (*for deleting an unwanted assumption*) val thin_rl = let val V = read_prop "V"; val W = read_prop "W"; val thm = Thm.implies_intr V (Thm.implies_intr W (Thm.assume W)); in store_standard_thm_open (Binding.make ("thin_rl", \<^here>)) thm end; (* (\x. PROP ?V) \ PROP ?V Allows removal of redundant parameters*) val triv_forall_equality = let val V = read_prop "V"; val QV = read_prop "\x::'a. V"; val x = certify (Free ("x", Term.aT [])); in store_standard_thm_open (Binding.make ("triv_forall_equality", \<^here>)) (Thm.equal_intr (Thm.implies_intr QV (Thm.forall_elim x (Thm.assume QV))) (Thm.implies_intr V (Thm.forall_intr x (Thm.assume V)))) end; (* (PROP ?Phi \ PROP ?Phi \ PROP ?Psi) \ (PROP ?Phi \ PROP ?Psi) *) val distinct_prems_rl = let val AAB = read_prop "Phi \ Phi \ Psi"; val A = read_prop "Phi"; in store_standard_thm_open (Binding.make ("distinct_prems_rl", \<^here>)) (implies_intr_list [AAB, A] (implies_elim_list (Thm.assume AAB) [Thm.assume A, Thm.assume A])) end; (* \PROP ?phi \ PROP ?psi; PROP ?psi \ PROP ?phi\ \ PROP ?phi \ PROP ?psi Introduction rule for \ as a meta-theorem. *) val equal_intr_rule = let val PQ = read_prop "phi \ psi"; val QP = read_prop "psi \ phi"; in store_standard_thm_open (Binding.make ("equal_intr_rule", \<^here>)) (Thm.implies_intr PQ (Thm.implies_intr QP (Thm.equal_intr (Thm.assume PQ) (Thm.assume QP)))) end; (* PROP ?phi \ PROP ?psi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule1 = let val eq = read_prop "phi::prop \ psi::prop"; val P = read_prop "phi"; in store_standard_thm_open (Binding.make ("equal_elim_rule1", \<^here>)) (Thm.equal_elim (Thm.assume eq) (Thm.assume P) |> implies_intr_list [eq, P]) end; (* PROP ?psi \ PROP ?phi \ PROP ?phi \ PROP ?psi *) val equal_elim_rule2 = store_standard_thm_open (Binding.make ("equal_elim_rule2", \<^here>)) (symmetric_thm RS equal_elim_rule1); (* PROP ?phi \ PROP ?phi \ PROP ?psi \ PROP ?psi *) val remdups_rl = let val P = read_prop "phi"; val Q = read_prop "psi"; val thm = implies_intr_list [P, P, Q] (Thm.assume Q); in store_standard_thm_open (Binding.make ("remdups_rl", \<^here>)) thm end; (** embedded terms and types **) local val A = certify (Free ("A", propT)); val axiom = Thm.unvarify_axiom (Context.the_global_context ()); val prop_def = axiom "Pure.prop_def"; val term_def = axiom "Pure.term_def"; val sort_constraint_def = axiom "Pure.sort_constraint_def"; val C = Thm.lhs_of sort_constraint_def; val T = Thm.dest_arg C; val CA = mk_implies (C, A); in (* protect *) val protect = Thm.apply (certify Logic.protectC); val protectI = store_standard_thm (Binding.concealed (Binding.make ("protectI", \<^here>))) (Thm.equal_elim (Thm.symmetric prop_def) (Thm.assume A)); val protectD = store_standard_thm (Binding.concealed (Binding.make ("protectD", \<^here>))) (Thm.equal_elim prop_def (Thm.assume (protect A))); val protect_cong = store_standard_thm_open (Binding.make ("protect_cong", \<^here>)) (Thm.reflexive (protect A)); fun implies_intr_protected asms th = let val asms' = map protect asms in implies_elim_list (implies_intr_list asms th) (map (fn asm' => Thm.assume asm' RS protectD) asms') |> implies_intr_list asms' end; (* term *) val termI = store_standard_thm (Binding.concealed (Binding.make ("termI", \<^here>))) (Thm.equal_elim (Thm.symmetric term_def) (Thm.forall_intr A (Thm.trivial A))); fun mk_term ct = let val cT = Thm.ctyp_of_cterm ct; val T = Thm.typ_of cT; - in Thm.instantiate ([((("'a", 0), []), cT)], [((("x", 0), T), ct)]) termI end; + in + Thm.instantiate (TVars.make [((("'a", 0), []), cT)], Vars.make [((("x", 0), T), ct)]) termI + end; fun dest_term th = let val cprop = strip_imp_concl (Thm.cprop_of th) in if can Logic.dest_term (Thm.term_of cprop) then Thm.dest_arg cprop else raise THM ("dest_term", 0, [th]) end; fun cterm_rule f = dest_term o f o mk_term; val add_frees_cterm = Cterms.add_frees o mk_term; val add_vars_cterm = Cterms.add_vars o mk_term; val dummy_thm = mk_term (certify Term.dummy_prop); val free_dummy_thm = Thm.tag_free_dummy dummy_thm; (* sort_constraint *) fun is_sort_constraint (Const ("Pure.sort_constraint", _) $ Const ("Pure.type", _)) = true | is_sort_constraint _ = false; val sort_constraintI = store_standard_thm (Binding.concealed (Binding.make ("sort_constraintI", \<^here>))) (Thm.equal_elim (Thm.symmetric sort_constraint_def) (mk_term T)); val sort_constraint_eq = store_standard_thm (Binding.concealed (Binding.make ("sort_constraint_eq", \<^here>))) (Thm.equal_intr (Thm.implies_intr CA (Thm.implies_elim (Thm.assume CA) (Thm.unvarify_global (Context.the_global_context ()) sort_constraintI))) (implies_intr_list [A, C] (Thm.assume A))); end; (* HHF normalization *) (* (PROP ?phi \ (\x. PROP ?psi x)) \ (\x. PROP ?phi \ PROP ?psi x) *) val norm_hhf_eq = let val aT = TFree ("'a", []); val x = Free ("x", aT); val phi = Free ("phi", propT); val psi = Free ("psi", aT --> propT); val cx = certify x; val cphi = certify phi; val lhs = certify (Logic.mk_implies (phi, Logic.all x (psi $ x))); val rhs = certify (Logic.all x (Logic.mk_implies (phi, psi $ x))); in Thm.equal_intr (Thm.implies_elim (Thm.assume lhs) (Thm.assume cphi) |> Thm.forall_elim cx |> Thm.implies_intr cphi |> Thm.forall_intr cx |> Thm.implies_intr lhs) (Thm.implies_elim (Thm.assume rhs |> Thm.forall_elim cx) (Thm.assume cphi) |> Thm.forall_intr cx |> Thm.implies_intr cphi |> Thm.implies_intr rhs) |> store_standard_thm_open (Binding.make ("norm_hhf_eq", \<^here>)) end; val norm_hhf_prop = Logic.dest_equals (Thm.prop_of norm_hhf_eq); val norm_hhf_eqs = [norm_hhf_eq, sort_constraint_eq]; fun is_norm_hhf (Const ("Pure.sort_constraint", _)) = false | is_norm_hhf (Const ("Pure.imp", _) $ _ $ (Const ("Pure.all", _) $ _)) = false | is_norm_hhf (Abs _ $ _) = false | is_norm_hhf (t $ u) = is_norm_hhf t andalso is_norm_hhf u | is_norm_hhf (Abs (_, _, t)) = is_norm_hhf t | is_norm_hhf _ = true; fun norm_hhf thy t = if is_norm_hhf t then t else Pattern.rewrite_term thy [norm_hhf_prop] [] t; fun norm_hhf_cterm ctxt raw_ct = let val thy = Proof_Context.theory_of ctxt; val ct = Thm.transfer_cterm thy raw_ct; val t = Thm.term_of ct; in if is_norm_hhf t then ct else Thm.cterm_of ctxt (norm_hhf thy t) end; (* var indexes *) fun incr_indexes th = Thm.incr_indexes (Thm.maxidx_of th + 1); fun incr_indexes2 th1 th2 = Thm.incr_indexes (Int.max (Thm.maxidx_of th1, Thm.maxidx_of th2) + 1); local (*compose Q and \Q1,Q2,...,Qk\ \ R to \Q2,...,Qk\ \ R getting unique result*) fun comp incremented th1 th2 = Thm.bicompose NONE {flatten = true, match = false, incremented = incremented} (false, th1, 0) 1 th2 |> Seq.list_of |> distinct Thm.eq_thm |> (fn [th] => Thm.solve_constraints th | _ => raise THM ("COMP", 1, [th1, th2])); in fun th1 COMP th2 = comp false th1 th2; fun th1 INCR_COMP th2 = comp true (incr_indexes th2 th1) th2; fun th1 COMP_INCR th2 = comp true th1 (incr_indexes th1 th2); end; fun comp_no_flatten (th, n) i rule = (case distinct Thm.eq_thm (Seq.list_of (Thm.bicompose NONE {flatten = false, match = false, incremented = true} (false, th, n) i (incr_indexes th rule))) of [th'] => Thm.solve_constraints th' | [] => raise THM ("comp_no_flatten", i, [th, rule]) | _ => raise THM ("comp_no_flatten: unique result expected", i, [th, rule])); (** variations on Thm.instantiate **) fun instantiate_normalize instpair th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate instpair th COMP_INCR asm_rl); fun instantiate'_normalize Ts ts th = Thm.adjust_maxidx_thm ~1 (Thm.instantiate' Ts ts th COMP_INCR asm_rl); (*instantiation with type-inference for variables*) fun infer_instantiate_types _ [] th = th | infer_instantiate_types ctxt args raw_th = let val thy = Proof_Context.theory_of ctxt; val th = Thm.transfer thy raw_th; fun infer ((xi, T), cu) (tyenv, maxidx) = let val _ = Thm.ctyp_of ctxt T; val _ = Thm.transfer_cterm thy cu; val U = Thm.typ_of_cterm cu; val maxidx' = maxidx |> Integer.max (#2 xi) |> Term.maxidx_typ T |> Integer.max (Thm.maxidx_of_cterm cu); val (tyenv', maxidx'') = Sign.typ_unify thy (T, U) (tyenv, maxidx') handle Type.TUNIFY => let val t = Var (xi, T); val u = Thm.term_of cu; in raise THM ("infer_instantiate_types: type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv T) ^ " of variable " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) t) ^ "\ncannot be unified with type " ^ Syntax.string_of_typ ctxt (Envir.norm_type tyenv U) ^ " of term " ^ Syntax.string_of_term ctxt (Term.map_types (Envir.norm_type tyenv) u), 0, [th]) end; in (tyenv', maxidx'') end; val (tyenv, _) = fold infer args (Vartab.empty, 0); val instT = - Vartab.fold (fn (xi, (S, T)) => - cons ((xi, S), Thm.ctyp_of ctxt (Envir.norm_type tyenv T))) tyenv []; - val inst = args |> map (fn ((xi, T), cu) => - ((xi, Envir.norm_type tyenv T), - Thm.instantiate_cterm (instT, []) (Thm.transfer_cterm thy cu))); + TVars.build (tyenv |> Vartab.fold (fn (xi, (S, T)) => + TVars.add ((xi, S), Thm.ctyp_of ctxt (Envir.norm_type tyenv T)))); + val inst = + Vars.build (args |> fold (fn ((xi, T), cu) => + Vars.add ((xi, Envir.norm_type tyenv T), + Thm.instantiate_cterm (instT, Vars.empty) (Thm.transfer_cterm thy cu)))); in instantiate_normalize (instT, inst) th end handle CTERM (msg, _) => raise THM (msg, 0, [raw_th]) | TERM (msg, _) => raise THM (msg, 0, [raw_th]) | TYPE (msg, _, _) => raise THM (msg, 0, [raw_th]); fun infer_instantiate _ [] th = th | infer_instantiate ctxt args th = let val vars = Term.add_vars (Thm.full_prop_of th) []; val dups = duplicates (eq_fst op =) vars; val _ = null dups orelse raise THM ("infer_instantiate: inconsistent types for variables " ^ commas_quote (map (Syntax.string_of_term (Config.put show_types true ctxt) o Var) dups), 0, [th]); val args' = args |> map_filter (fn (xi, cu) => AList.lookup (op =) vars xi |> Option.map (fn T => ((xi, T), cu))); in infer_instantiate_types ctxt args' th end; fun infer_instantiate' ctxt args th = let val vars = build_rev (Term.add_vars (Thm.full_prop_of th)); val args' = zip_options vars args handle ListPair.UnequalLengths => raise THM ("infer_instantiate': more instantiations than variables in thm", 0, [th]); in infer_instantiate_types ctxt args' th end; (** renaming of bound variables **) (* replace bound variables x_i in thm by y_i *) (* where vs = [(x_1, y_1), ..., (x_n, y_n)] *) fun rename_bvars [] thm = thm | rename_bvars vs thm = let fun rename (Abs (x, T, t)) = Abs (AList.lookup (op =) vs x |> the_default x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; (* renaming in left-to-right order *) fun rename_bvars' xs thm = let fun rename [] t = ([], t) | rename (x' :: xs) (Abs (x, T, t)) = let val (xs', t') = rename xs t in (xs', Abs (the_default x x', T, t')) end | rename xs (t $ u) = let val (xs', t') = rename xs t; val (xs'', u') = rename xs' u; in (xs'', t' $ u') end | rename xs a = (xs, a); in (case rename xs (Thm.prop_of thm) of ([], prop') => Thm.renamed_prop prop' thm | _ => error "More names than abstractions in theorem") end; end; structure Basic_Drule: BASIC_DRULE = Drule; open Basic_Drule; diff --git a/src/Pure/goal.ML b/src/Pure/goal.ML --- a/src/Pure/goal.ML +++ b/src/Pure/goal.ML @@ -1,342 +1,342 @@ (* Title: Pure/goal.ML Author: Makarius Goals in tactical theorem proving, with support for forked proofs. *) signature BASIC_GOAL = sig val quick_and_dirty: bool Config.T val SELECT_GOAL: tactic -> int -> tactic val PREFER_GOAL: tactic -> int -> tactic val CONJUNCTS: tactic -> int -> tactic val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic end; signature GOAL = sig include BASIC_GOAL val init: cterm -> thm val protect: int -> thm -> thm val conclude: thm -> thm val check_finished: Proof.context -> thm -> thm val finish: Proof.context -> thm -> thm val norm_result: Proof.context -> thm -> thm val skip_proofs_enabled: unit -> bool val future_result: Proof.context -> thm future -> term -> thm val prove_internal: Proof.context -> cterm list -> cterm -> (thm list -> tactic) -> thm val is_schematic: term -> bool val prove_common: Proof.context -> int option -> string list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> thm list val prove_future: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_global_future: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_global: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_sorry: Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val prove_sorry_global: theory -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm val restrict: int -> int -> thm -> thm val unrestrict: int -> thm -> thm val conjunction_tac: int -> tactic val precise_conjunction_tac: int -> int -> tactic val recover_conjunction_tac: tactic val norm_hhf_tac: Proof.context -> int -> tactic val assume_rule_tac: Proof.context -> int -> tactic end; structure Goal: GOAL = struct (** goals **) (* -------- (init) C \ #C *) -fun init C = Thm.instantiate ([], [((("A", 0), propT), C)]) Drule.protectI; +fun init C = Thm.instantiate (TVars.empty, Vars.make [((("A", 0), propT), C)]) Drule.protectI; (* A1 \ ... \ An \ C ------------------------ (protect n) A1 \ ... \ An \ #C *) fun protect n th = Drule.comp_no_flatten (th, n) 1 Drule.protectI; (* A \ ... \ #C ---------------- (conclude) A \ ... \ C *) fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD; (* #C --- (finish) C *) fun check_finished ctxt th = if Thm.no_prems th then th else raise THM ("Proof failed.\n" ^ Pretty.string_of (Goal_Display.pretty_goal ctxt th), 0, [th]); fun finish ctxt = check_finished ctxt #> conclude; (** results **) (* normal form *) fun norm_result ctxt = Drule.flexflex_unique (SOME ctxt) #> Raw_Simplifier.norm_hhf_protect ctxt #> Thm.strip_shyps #> Drule.zero_var_indexes; (* scheduling parameters *) fun skip_proofs_enabled () = let val skip = Options.default_bool "skip_proofs" in if Proofterm.proofs_enabled () andalso skip then (warning "Proof terms enabled -- cannot skip proofs"; false) else skip end; (* future_result *) fun future_result ctxt result prop = let val assms = Assumption.all_assms_of ctxt; val As = map Thm.term_of assms; val frees = Frees.build (fold Frees.add_frees (prop :: As)); val xs = Frees.fold_rev (cons o Thm.cterm_of ctxt o Free o #1) frees []; val tfrees = TFrees.build (fold TFrees.add_tfrees (prop :: As)); val Ts = Names.build (TFrees.fold (Names.add_set o #1 o #1) tfrees); val instT = - build (tfrees |> TFrees.fold (fn ((a, S), _) => - cons (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S))))); + TVars.build (tfrees |> TFrees.fold (fn ((a, S), _) => + TVars.add (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S))))); val global_prop = Logic.list_implies (As, prop) |> Frees.fold_rev (Logic.all o Free o #1) frees |> Logic.varify_types_global |> Thm.cterm_of ctxt |> Thm.weaken_sorts' ctxt; val global_result = result |> Future.map (Drule.flexflex_unique (SOME ctxt) #> Drule.implies_intr_list assms #> Drule.forall_intr_list xs #> Thm.adjust_maxidx_thm ~1 #> Thm.generalize (Ts, Names.empty) 0 #> Thm.strip_shyps #> Thm.solve_constraints); val local_result = Thm.future global_result global_prop |> Thm.close_derivation \<^here> - |> Thm.instantiate (instT, []) + |> Thm.instantiate (instT, Vars.empty) |> Drule.forall_elim_list xs |> fold (Thm.elim_implies o Thm.assume) assms |> Thm.solve_constraints; in local_result end; (** tactical theorem proving **) (* prove_internal -- minimal checks, no normalization of result! *) fun prove_internal ctxt casms cprop tac = (case SINGLE (tac (map (Assumption.assume ctxt) casms)) (init cprop) of SOME th => Drule.implies_intr_list casms (finish ctxt th) | NONE => error "Tactic failed"); (* prove variations *) fun is_schematic t = Term.exists_subterm Term.is_Var t orelse Term.exists_type (Term.exists_subtype Term.is_TVar) t; fun prove_common ctxt fork_pri xs asms props tac = let val thy = Proof_Context.theory_of ctxt; val schematic = exists is_schematic props; val immediate = is_none fork_pri; val future = Future.proofs_enabled 1 andalso not (Proofterm.proofs_enabled ()); val skip = not immediate andalso not schematic andalso future andalso skip_proofs_enabled (); val pos = Position.thread_data (); fun err msg = cat_error msg ("The error(s) above occurred for the goal statement" ^ Position.here pos ^ ":\n" ^ Syntax.string_of_term ctxt (Logic.list_implies (asms, Logic.mk_conjunction_list props))); fun cert_safe t = Thm.cterm_of ctxt (Envir.beta_norm (Term.no_dummy_patterns t)) handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; val casms = map cert_safe asms; val cprops = map cert_safe props; val (prems, ctxt') = ctxt |> Variable.add_fixes_direct xs |> fold Variable.declare_term (asms @ props) |> Assumption.add_assumes casms ||> Variable.set_body true; val stmt = Thm.weaken_sorts' ctxt' (Conjunction.mk_conjunction_balanced cprops); fun tac' args st = if skip then ALLGOALS (Skip_Proof.cheat_tac ctxt) st before Skip_Proof.report ctxt else tac args st; fun result () = (case SINGLE (tac' {prems = prems, context = ctxt'}) (init stmt) of NONE => err "Tactic failed" | SOME st => let val _ = Context.subthy_id (Thm.theory_id st, Context.theory_id thy) orelse err "Bad background theory of goal state"; val res = (finish ctxt' st |> Drule.flexflex_unique (SOME ctxt') |> Thm.check_shyps ctxt' |> Thm.check_hyps (Context.Proof ctxt')) handle THM (msg, _, _) => err msg | ERROR msg => err msg; in if Unify.matches_list (Context.Proof ctxt') [Thm.term_of stmt] [Thm.prop_of res] then res else err ("Proved a different theorem: " ^ Syntax.string_of_term ctxt' (Thm.prop_of res)) end); val res = if immediate orelse schematic orelse not future orelse skip then result () else future_result ctxt' (Execution.fork {name = "Goal.prove", pos = Position.thread_data (), pri = the fork_pri} result) (Thm.term_of stmt); in res |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length props) |> map (Assumption.export false ctxt' ctxt) |> Variable.export ctxt' ctxt |> map Drule.zero_var_indexes end; fun prove_future_pri ctxt pri xs asms prop tac = hd (prove_common ctxt (SOME pri) xs asms [prop] tac); fun prove_future ctxt = prove_future_pri ctxt ~1; fun prove ctxt xs asms prop tac = hd (prove_common ctxt NONE xs asms [prop] tac); fun prove_global_future thy xs asms prop tac = Drule.export_without_context (prove_future (Proof_Context.init_global thy) xs asms prop tac); fun prove_global thy xs asms prop tac = Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac); (* skip proofs *) val quick_and_dirty = Config.declare_option_bool ("quick_and_dirty", \<^here>); fun prove_sorry ctxt xs asms prop tac = if Config.get ctxt quick_and_dirty then prove ctxt xs asms prop (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)) else (if Future.proofs_enabled 1 then prove_future_pri ctxt ~2 else prove ctxt) xs asms prop tac; fun prove_sorry_global thy xs asms prop tac = Drule.export_without_context (prove_sorry (Proof_Context.init_global thy) xs asms prop tac); (** goal structure **) (* rearrange subgoals *) fun restrict i n st = if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then raise THM ("Goal.restrict", i, [st]) else rotate_prems (i - 1) st |> protect n; fun unrestrict i = conclude #> rotate_prems (1 - i); (*with structural marker*) fun SELECT_GOAL tac i st = if Thm.nprems_of st = 1 andalso i = 1 then tac st else (PRIMITIVE (restrict i 1) THEN tac THEN PRIMITIVE (unrestrict i)) st; (*without structural marker*) fun PREFER_GOAL tac i st = if i < 1 orelse i > Thm.nprems_of st then Seq.empty else (PRIMITIVE (rotate_prems (i - 1)) THEN tac THEN PRIMITIVE (rotate_prems (1 - i))) st; (* multiple goals *) fun precise_conjunction_tac 0 i = eq_assume_tac i | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i | precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n)); val adhoc_conjunction_tac = REPEAT_ALL_NEW (SUBGOAL (fn (goal, i) => if can Logic.dest_conjunction goal then resolve0_tac [Conjunction.conjunctionI] i else no_tac)); val conjunction_tac = SUBGOAL (fn (goal, i) => precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE TRY (adhoc_conjunction_tac i)); val recover_conjunction_tac = PRIMITIVE (fn th => Conjunction.uncurry_balanced (Thm.nprems_of th) th); fun PRECISE_CONJUNCTS n tac = SELECT_GOAL (precise_conjunction_tac n 1 THEN tac THEN recover_conjunction_tac); fun CONJUNCTS tac = SELECT_GOAL (conjunction_tac 1 THEN tac THEN recover_conjunction_tac); (* hhf normal form *) fun norm_hhf_tac ctxt = resolve_tac ctxt [Drule.asm_rl] (*cheap approximation -- thanks to builtin Logic.flatten_params*) THEN' SUBGOAL (fn (t, i) => if Drule.is_norm_hhf t then all_tac else rewrite_goal_tac ctxt Drule.norm_hhf_eqs i); (* non-atomic goal assumptions *) fun non_atomic (Const ("Pure.imp", _) $ _ $ _) = true | non_atomic (Const ("Pure.all", _) $ _) = true | non_atomic _ = false; fun assume_rule_tac ctxt = norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) => let val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt; val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); val tacs = Rs |> map (fn R => eresolve_tac ctxt [Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)] THEN_ALL_NEW assume_tac ctxt); in fold_rev (curry op APPEND') tacs (K no_tac) i end); end; structure Basic_Goal: BASIC_GOAL = Goal; open Basic_Goal; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,738 +1,741 @@ (* Title: Pure/more_thm.ML Author: Makarius Further operations on type ctyp/cterm/thm, outside the inference kernel. *) infix aconvc; signature BASIC_THM = sig include BASIC_THM val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool val add_tvars: thm -> ctyp TVars.table -> ctyp TVars.table val add_vars: thm -> cterm Vars.table -> cterm Vars.table val dest_funT: ctyp -> ctyp * ctyp val strip_type: ctyp -> ctyp list * ctyp val all_name: Proof.context -> string * cterm -> cterm -> cterm val all: Proof.context -> cterm -> cterm -> cterm val mk_binop: cterm -> cterm -> cterm -> cterm val dest_binop: cterm -> cterm * cterm val dest_implies: cterm -> cterm * cterm val dest_equals: cterm -> cterm * cterm val dest_equals_lhs: cterm -> cterm val dest_equals_rhs: cterm -> cterm val lhs_of: thm -> cterm val rhs_of: thm -> cterm val is_reflexive: thm -> bool val eq_thm: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val eq_thm_strict: thm * thm -> bool val equiv_thm: theory -> thm * thm -> bool val class_triv: theory -> class -> thm val of_sort: ctyp * sort -> thm list val is_dummy: thm -> bool val add_thm: thm -> thm list -> thm list val del_thm: thm -> thm list -> thm list val merge_thms: thm list * thm list -> thm list val item_net: thm Item_Net.T val item_net_intro: thm Item_Net.T val item_net_elim: thm Item_Net.T val declare_hyps: cterm -> Proof.context -> Proof.context val assume_hyps: cterm -> Proof.context -> thm * Proof.context val unchecked_hyps: Proof.context -> Proof.context val restore_hyps: Proof.context -> Proof.context -> Proof.context val undeclared_hyps: Context.generic -> thm -> term list val check_hyps: Context.generic -> thm -> thm val declare_term_sorts: term -> Proof.context -> Proof.context val extra_shyps': Proof.context -> thm -> sort list val check_shyps: Proof.context -> thm -> thm val weaken_sorts': Proof.context -> cterm -> cterm val elim_implies: thm -> thm -> thm val forall_intr_name: string * cterm -> thm -> thm val forall_elim_var: int -> thm -> thm val forall_elim_vars: int -> thm -> thm - val instantiate_frees: ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> thm -> thm + val instantiate_frees: ctyp TFrees.table * cterm Frees.table -> thm -> thm val instantiate': ctyp option list -> cterm option list -> thm -> thm val forall_intr_frees: thm -> thm val forall_intr_vars: thm -> thm val unvarify_global: theory -> thm -> thm val unvarify_axiom: theory -> string -> thm val rename_params_rule: string list * int -> thm -> thm val rename_boundvars: term -> term -> thm -> thm val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory val add_axiom_global: binding * term -> theory -> (string * thm) * theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory type attribute = Context.generic * thm -> Context.generic option * thm option type binding = binding * attribute list val tag_rule: string * string -> thm -> thm val untag_rule: string -> thm -> thm val is_free_dummy: thm -> bool val tag_free_dummy: thm -> thm val def_name: string -> string val def_name_optional: string -> string -> string val def_binding: Binding.binding -> Binding.binding val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding val make_def_binding: bool -> Binding.binding -> Binding.binding val has_name_hint: thm -> bool val get_name_hint: thm -> string val put_name_hint: string -> thm -> thm val theoremK: string val legacy_get_kind: thm -> string val kind_rule: string -> thm -> thm val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic val theory_attributes: attribute list -> thm -> theory -> thm * theory val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context val no_attributes: 'a -> 'a * 'b list val simple_fact: 'a -> ('a * 'b list) list val tag: string * string -> attribute val untag: string -> attribute val kind: string -> attribute val register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val expose_theory: theory -> unit val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val pretty_thm_item: Proof.context -> thm -> Pretty.T val pretty_thm_global: theory -> thm -> Pretty.T val string_of_thm: Proof.context -> thm -> string val string_of_thm_global: theory -> thm -> string end; structure Thm: THM = struct (** basic operations **) (* collecting ctyps and cterms *) val eq_ctyp = op = o apply2 Thm.typ_of; val op aconvc = op aconv o apply2 Thm.term_of; val add_tvars = Thm.fold_atomic_ctyps {hyps = false} Term.is_TVar (fn cT => fn tab => let val v = Term.dest_TVar (Thm.typ_of cT) in tab |> not (TVars.defined tab v) ? TVars.add (v, cT) end); val add_vars = Thm.fold_atomic_cterms {hyps = false} Term.is_Var (fn ct => fn tab => let val v = Term.dest_Var (Thm.term_of ct) in tab |> not (Vars.defined tab v) ? Vars.add (v, ct) end); (* ctyp operations *) fun dest_funT cT = (case Thm.typ_of cT of Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end | T => raise TYPE ("dest_funT", [T], [])); (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) fun strip_type cT = (case Thm.typ_of cT of Type ("fun", _) => let val (cT1, cT2) = dest_funT cT; val (cTs, cT') = strip_type cT2 in (cT1 :: cTs, cT') end | _ => ([], cT)); (* cterm operations *) fun all_name ctxt (x, t) A = let val T = Thm.typ_of_cterm t; val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT)); in Thm.apply all_const (Thm.lambda_name (x, t) A) end; fun all ctxt t A = all_name ctxt ("", t) A; fun mk_binop c a b = Thm.apply (Thm.apply c a) b; fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); fun dest_implies ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_implies", [Thm.term_of ct])); fun dest_equals ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_equals", [Thm.term_of ct])); fun dest_equals_lhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); fun dest_equals_rhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); val lhs_of = dest_equals_lhs o Thm.cprop_of; val rhs_of = dest_equals_rhs o Thm.cprop_of; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o Thm.thm_ord; val eq_thm_prop = op aconv o apply2 Thm.full_prop_of; fun eq_thm_strict ths = eq_thm ths andalso Context.eq_thy_id (apply2 Thm.theory_id ths) andalso op = (apply2 Thm.maxidx_of ths) andalso op = (apply2 Thm.get_tags ths); (* pattern equivalence *) fun equiv_thm thy ths = Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths); (* type classes and sorts *) fun class_triv thy c = Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; (* misc operations *) fun is_dummy thm = (case try Logic.dest_term (Thm.concl_of thm) of NONE => false | SOME t => Term.is_dummy_pattern (Term.head_of t)); (* collections of theorems in canonical order *) val add_thm = update eq_thm_prop; val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; val item_net = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val item_net_intro = Item_Net.init eq_thm_prop (single o Thm.concl_of); val item_net_elim = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); (** declared hyps and sort hyps **) structure Hyps = Proof_Data ( type T = {checked_hyps: bool, hyps: Termtab.set, shyps: sort Ord_List.T}; fun init _ : T = {checked_hyps = true, hyps = Termtab.empty, shyps = []}; ); fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} => let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end); (* hyps *) fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct; val hyps' = Termtab.update (Thm.term_of ct, ()) hyps; in (checked_hyps, hyps', shyps) end); fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); fun restore_hyps ctxt = map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps)); fun undeclared_hyps context th = Thm.hyps_of th |> filter_out (case context of Context.Theory _ => K false | Context.Proof ctxt => (case Hyps.get ctxt of {checked_hyps = false, ...} => K true | {hyps, ...} => Termtab.defined hyps)); fun check_hyps context th = (case undeclared_hyps context th of [] => th | undeclared => error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" (map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared)))); (* shyps *) fun declare_term_sorts t = map_hyps (fn (checked_hyps, hyps, shyps) => (checked_hyps, hyps, Sorts.insert_term t shyps)); fun extra_shyps' ctxt th = Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); fun check_shyps ctxt raw_th = let val th = Thm.strip_shyps raw_th; val extra_shyps = extra_shyps' ctxt th; in if null extra_shyps then th else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) end; val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; (** basic derived rules **) (*Elimination of implication A A \ B ------------ B *) fun elim_implies thA thAB = Thm.implies_elim thAB thA; (* forall_intr_name *) fun forall_intr_name (a, x) th = let val th' = Thm.forall_intr x th; val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); in Thm.renamed_prop prop' th' end; (* forall_elim_var(s) *) local fun dest_all ct = (case Thm.term_of ct of Const ("Pure.all", _) $ Abs (a, _, _) => let val (x, ct') = Thm.dest_abs NONE (Thm.dest_arg ct) in SOME ((a, Thm.ctyp_of_cterm x), ct') end | _ => NONE); fun dest_all_list ct = (case dest_all ct of NONE => [] | SOME (v, ct') => v :: dest_all_list ct'); fun forall_elim_vars_list vars i th = let val used = (Thm.fold_terms {hyps = false} o Term.fold_aterms) (fn Var ((x, j), _) => if i = j then insert (op =) x else I | _ => I) th []; val vars' = (Name.variant_list used (map #1 vars), vars) |> ListPair.map (fn (x, (_, T)) => Thm.var ((x, i), T)); in fold Thm.forall_elim vars' th end; in fun forall_elim_vars i th = forall_elim_vars_list (dest_all_list (Thm.cprop_of th)) i th; fun forall_elim_var i th = let val vars = (case dest_all (Thm.cprop_of th) of SOME (v, _) => [v] | NONE => raise THM ("forall_elim_var", i, [th])); in forall_elim_vars_list vars i th end; end; (* instantiate frees *) -fun instantiate_frees ([], []) th = th - | instantiate_frees (instT, inst) th = - let - val idx = Thm.maxidx_of th + 1; - fun index ((a, A), b) = (((a, idx), A), b); - val insts = (map index instT, map index inst); - val tfrees = Names.build (fold (Names.add_set o #1 o #1) instT); - val frees = Names.build (fold (Names.add_set o #1 o #1) inst); +fun instantiate_frees (instT, inst) th = + if TFrees.is_empty instT andalso Frees.is_empty inst then th + else + let + val idx = Thm.maxidx_of th + 1; + fun index ((a, A), b) = (((a, idx), A), b); + val insts = + (TVars.build (TFrees.fold (TVars.add o index) instT), + Vars.build (Frees.fold (Vars.add o index) inst)); + val tfrees = Names.build (TFrees.fold (Names.add_set o #1 o #1) instT); + val frees = Names.build (Frees.fold (Names.add_set o #1 o #1) inst); - val hyps = Thm.chyps_of th; - val inst_cterm = - Thm.generalize_cterm (tfrees, frees) idx #> - Thm.instantiate_cterm insts; - in - th - |> fold_rev Thm.implies_intr hyps - |> Thm.generalize (tfrees, frees) idx - |> Thm.instantiate insts - |> fold (elim_implies o Thm.assume o inst_cterm) hyps - end; + val hyps = Thm.chyps_of th; + val inst_cterm = + Thm.generalize_cterm (tfrees, frees) idx #> + Thm.instantiate_cterm insts; + in + th + |> fold_rev Thm.implies_intr hyps + |> Thm.generalize (tfrees, frees) idx + |> Thm.instantiate insts + |> fold (elim_implies o Thm.assume o inst_cterm) hyps + end; (* instantiate by left-to-right occurrence of variables *) fun instantiate' cTs cts thm = let fun err msg = raise TYPE ("instantiate': " ^ msg, map_filter (Option.map Thm.typ_of) cTs, map_filter (Option.map Thm.term_of) cts); fun zip_vars xs ys = zip_options xs ys handle ListPair.UnequalLengths => err "more instantiations than variables in thm"; val instT = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_tvars thm)) cTs; - val thm' = Thm.instantiate (instT, []) thm; + val thm' = Thm.instantiate (TVars.make instT, Vars.empty) thm; val inst = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_vars thm')) cts; - in Thm.instantiate ([], inst) thm' end; + in Thm.instantiate (TVars.empty, Vars.make inst) thm' end; (* implicit generalization over variables -- canonical order *) fun forall_intr_vars th = let val vars = Cterms.build (Cterms.add_vars th) in fold_rev Thm.forall_intr (Cterms.list_set vars) th end; fun forall_intr_frees th = let val fixed = Frees.build (fold Frees.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> fold Frees.add_frees (Thm.hyps_of th)); val is_fixed = Frees.defined fixed o Term.dest_Free o Thm.term_of; val frees = Cterms.build (th |> Thm.fold_atomic_cterms {hyps = false} Term.is_Free (fn ct => not (is_fixed ct) ? Cterms.add_set ct)); in fold_rev Thm.forall_intr (Cterms.list_set frees) th end; (* unvarify_global: global schematic variables *) fun unvarify_global thy th = let val prop = Thm.full_prop_of th; val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); val cert = Thm.global_cterm_of thy; val certT = Thm.global_ctyp_of thy; val instT = TVars.build (prop |> (Term.fold_types o Term.fold_atyps) (fn T => fn instT => (case T of TVar (v as ((a, _), S)) => if TVars.defined instT v then instT else TVars.add (v, TFree (a, S)) instT | _ => instT))); val cinstT = TVars.map (K certT) instT; val cinst = Vars.build (prop |> Term.fold_aterms (fn t => fn inst => (case t of Var ((x, i), T) => let val T' = Term_Subst.instantiateT instT T in Vars.add (((x, i), T'), cert (Free ((x, T')))) inst end | _ => inst))); - in Thm.instantiate (TVars.dest cinstT, Vars.dest cinst) th end; + in Thm.instantiate (cinstT, cinst) th end; fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; (* user renaming of parameters in a subgoal *) (*The names, if distinct, are used for the innermost parameters of subgoal i; preceding parameters may be renamed to make all parameters distinct.*) fun rename_params_rule (names, i) st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val params = map #1 (Logic.strip_params Bi); val short = length params - length names; val names' = if short < 0 then error "More names than parameters in subgoal!" else Name.variant_list names (take short params) @ names; val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val Bi' = Logic.list_rename_params names' Bi; in (case duplicates (op =) names of a :: _ => (warning ("Can't rename. Bound variables not distinct: " ^ a); st) | [] => (case inter (op =) names free_names of a :: _ => (warning ("Can't rename. Bound/Free variable clash: " ^ a); st) | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) end; (* preservation of bound variable names *) fun rename_boundvars pat obj th = (case Term.rename_abs pat obj (Thm.prop_of th) of NONE => th | SOME prop' => Thm.renamed_prop prop' th); (** specification primitives **) (* rules *) fun stripped_sorts thy t = let val tfrees = build_rev (Term.add_tfrees t); val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); val recover = map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S)))) tfrees' tfrees; val strip = map (apply2 TFree) (tfrees ~~ tfrees'); val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; fun add_axiom ctxt (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val (strip, recover, prop') = stripped_sorts thy prop; val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of ctxt T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = - Thm.instantiate (recover, []) axm' + Thm.instantiate (TVars.make recover, Vars.empty) axm' |> unvarify_global thy' |> fold elim_implies of_sorts; in ((axm_name, thm), thy') end; fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def context unchecked overloaded (b, concl') thy; val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = - Thm.instantiate (recover, []) axm' + Thm.instantiate (TVars.make recover, Vars.empty) axm' |> unvarify_global thy' |> fold_rev Thm.implies_intr prems; in ((axm_name, thm), thy') end; fun add_def_global unchecked overloaded arg thy = add_def (Defs.global_context thy) unchecked overloaded arg thy; (** theorem tags **) (* add / delete tags *) fun tag_rule tg = Thm.map_tags (insert (op =) tg); fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); (* free dummy thm -- for abstract closure *) val free_dummyN = "free_dummy"; fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; val tag_free_dummy = tag_rule (free_dummyN, ""); (* def_name *) fun def_name c = c ^ "_def"; fun def_name_optional c "" = def_name c | def_name_optional _ name = name; val def_binding = Binding.map_name def_name #> Binding.reset_pos; fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); (* theorem kinds *) val theoremK = "theorem"; fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; (** attributes **) (*attributes subsume any kind of rules or context modifiers*) type attribute = Context.generic * thm -> Context.generic option * thm option; type binding = binding * attribute list; fun rule_attribute ths f (x, th) = (NONE, (case find_first is_free_dummy (th :: ths) of SOME th' => SOME th' | NONE => SOME (f x th))); fun declaration_attribute f (x, th) = (if is_free_dummy th then NONE else SOME (f th x), NONE); fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; fun apply_attribute (att: attribute) th x = let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) in (the_default th th', the_default x x') end; fun attribute_declaration att th x = #2 (apply_attribute att th x); fun apply_attributes mk dest = let fun app [] th x = (th, x) | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; in app end; val theory_attributes = apply_attributes Context.Theory Context.the_theory; val proof_attributes = apply_attributes Context.Proof Context.the_proof; fun no_attributes x = (x, []); fun simple_fact x = [(x, [])]; fun tag tg = rule_attribute [] (K (tag_rule tg)); fun untag s = rule_attribute [] (K (untag_rule s)); fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k)); (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy Inttab.table; val empty = Inttab.empty; val extend = I; val merge = Inttab.merge (K true); ); fun reset_proofs thy = if Inttab.is_empty (Proofs.get thy) then NONE else SOME (Proofs.put Inttab.empty thy); val _ = Theory.setup (Theory.at_begin reset_proofs); fun register_proofs ths thy = let val entry = (serial (), Lazy.map_finished (map Thm.trim_context) ths) in (Proofs.map o Inttab.update) entry thy end; fun force_proofs thy = Proofs.get thy |> Inttab.dest |> maps (map (Thm.transfer thy) o Lazy.force o #2); val consolidate_theory = Thm.consolidate o force_proofs; fun expose_theory thy = if Proofterm.export_enabled () then Thm.expose_proofs thy (force_proofs thy) else (); (** print theorems **) (* options *) val show_consts = Config.declare_option_bool ("show_consts", \<^here>); val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false); val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false); (* pretty_thm etc. *) fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th = let val show_tags = Config.get ctxt show_tags; val show_hyps = Config.get ctxt show_hyps; val th = raw_th |> perhaps (try (Thm.transfer' ctxt)) |> perhaps (try Thm.strip_shyps); val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; val extra_shyps = extra_shyps' ctxt th; val tags = Thm.get_tags th; val tpairs = Thm.tpairs_of th; val q = if quote then Pretty.quote else I; val prt_term = q o Syntax.pretty_term ctxt; val hlen = length extra_shyps + length hyps + length tpairs; val hsymbs = if hlen = 0 then [] else if show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" (map (q o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ map (Syntax.pretty_sort ctxt) extra_shyps)] else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")]; val tsymbs = if null tags orelse not show_tags then [] else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true}; fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; fun pretty_thm_global thy = pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false}; val string_of_thm = Pretty.string_of oo pretty_thm; val string_of_thm_global = Pretty.string_of oo pretty_thm_global; open Thm; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/morphism.ML b/src/Pure/morphism.ML --- a/src/Pure/morphism.ML +++ b/src/Pure/morphism.ML @@ -1,170 +1,170 @@ (* 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 morphism: string -> {binding: (binding -> binding) list, typ: (typ -> typ) list, term: (term -> term) list, fact: (thm list -> thm list) list} -> morphism 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 compose: morphism -> morphism -> morphism val transform: morphism -> (morphism -> 'a) -> morphism -> 'a val form: (morphism -> 'a) -> 'a val binding_morphism: string -> (binding -> binding) -> morphism val typ_morphism: string -> (typ -> typ) -> 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 instantiate_frees_morphism: - ((string * sort) * ctyp) list * ((string * typ) * cterm) list -> morphism - val instantiate_morphism: - ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> 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 * ('a -> 'a)) list; exception MORPHISM of string * exn; fun app (name, f) x = f x handle exn => if Exn.is_interrupt exn then Exn.reraise exn else raise MORPHISM (name, exn); fun apply fs = fold_rev app fs; (* type morphism *) datatype morphism = Morphism of {names: string list, binding: binding funs, typ: typ funs, term: term funs, fact: thm list funs}; type declaration = morphism -> Context.generic -> Context.generic; fun morphism a {binding, typ, term, fact} = Morphism { 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}; fun pretty (Morphism {names, ...}) = Pretty.enum ";" "{" "}" (map Pretty.str (rev names)); val _ = ML_system_pp (fn _ => fn _ => Pretty.to_polyml o pretty); fun binding (Morphism {binding, ...}) = apply binding; fun binding_prefix morph = Binding.name "x" |> binding morph |> Binding.prefix_of; fun typ (Morphism {typ, ...}) = apply typ; fun term (Morphism {term, ...}) = apply term; fun fact (Morphism {fact, ...}) = apply fact; val thm = singleton o fact; val cterm = Drule.cterm_rule o thm; (* morphism combinators *) val identity = morphism "" {binding = [], typ = [], term = [], fact = []}; fun compose (Morphism {names = names1, binding = binding1, typ = typ1, term = term1, fact = fact1}) (Morphism {names = names2, binding = binding2, typ = typ2, term = term2, fact = fact2}) = Morphism { names = names1 @ names2, binding = binding1 @ binding2, typ = typ1 @ typ2, term = term1 @ term2, fact = fact1 @ fact2}; fun phi1 $> phi2 = compose phi2 phi1; fun transform phi f = fn psi => f (phi $> psi); fun form f = f identity; (* concrete morphisms *) fun binding_morphism a binding = morphism a {binding = [binding], typ = [], term = [], fact = []}; fun typ_morphism a typ = morphism a {binding = [], typ = [typ], term = [], fact = []}; fun term_morphism a term = morphism a {binding = [], typ = [], term = [term], fact = []}; fun fact_morphism a fact = morphism a {binding = [], typ = [], term = [], fact = [fact]}; fun thm_morphism a thm = morphism a {binding = [], typ = [], term = [], fact = [map thm]}; val transfer_morphism = thm_morphism "transfer" o Thm.join_transfer; 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; (* instantiate *) -fun instantiate_frees_morphism ([], []) = identity - | instantiate_frees_morphism (cinstT, cinst) = - let - val instT = TFrees.build (cinstT |> fold (fn (v, cT) => TFrees.add (v, Thm.typ_of cT))); - val inst = Frees.build (cinst |> fold (fn (v, ct) => Frees.add (v, Thm.term_of ct))); - in - morphism "instantiate_frees" - {binding = [], - typ = - if TFrees.is_empty instT then [] - else [Term_Subst.instantiateT_frees instT], - term = [Term_Subst.instantiate_frees (instT, inst)], - fact = [map (Thm.instantiate_frees (cinstT, cinst))]} - end; +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 [Term_Subst.instantiateT_frees instT], + term = [Term_Subst.instantiate_frees (instT, inst)], + fact = [map (Thm.instantiate_frees (cinstT, cinst))]} + end; -fun instantiate_morphism ([], []) = identity - | instantiate_morphism (cinstT, cinst) = - let - val instT = TVars.build (cinstT |> fold (fn (v, cT) => TVars.add (v, Thm.typ_of cT))); - val inst = Vars.build (cinst |> fold (fn (v, ct) => Vars.add (v, Thm.term_of ct))); - in - morphism "instantiate" - {binding = [], - typ = - if TVars.is_empty instT then [] - else [Term_Subst.instantiateT instT], - term = [Term_Subst.instantiate (instT, inst)], - fact = [map (Thm.instantiate (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 [Term_Subst.instantiateT instT], + term = [Term_Subst.instantiate (instT, inst)], + fact = [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,1463 +1,1466 @@ (* 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 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 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 extend = I; 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, 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}; val extend = I; 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.count_prems prop' = 0); 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, T, _) => let val (v, ctxt') = Variable.next_bound (a, T) ctxt; val b = #1 (Term.dest_Free v); val (v', t') = Thm.dest_abs (SOME b) t0; val b' = #1 (Term.dest_Free (Thm.term_of v')); val _ = if b <> b' then warning ("Bad Simplifier context: renamed bound variable " ^ quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ())) else (); 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 ([], [(vA, prem), (vB, lhs), (vC, rhs)]) Drule.imp_cong) + (Thm.instantiate (TVars.empty, Vars.make [(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 ([], [(vA, prem'), (vB, prem), (vC, concl)]) Drule.swap_prems_eq) + (Thm.instantiate (TVars.empty, Vars.make [(vA, prem'), (vB, prem), (vC, concl)]) + Drule.swap_prems_eq) eq') - (Thm.instantiate ([], [(vA, prem), (vB, prem''), (vC, concl)]) Drule.swap_prems_eq) + (Thm.instantiate (TVars.empty, Vars.make [(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 ss ctxt0 th0 = let val (ctxt, th) = Thm.join_transfer_context (ctxt0, th0); val th' = if Drule.is_norm_hhf (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 hhf_ss; val norm_hhf_protect = gen_norm_hhf hhf_protect_ss; end; end; structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier; open Basic_Meta_Simplifier; diff --git a/src/Pure/tactic.ML b/src/Pure/tactic.ML --- a/src/Pure/tactic.ML +++ b/src/Pure/tactic.ML @@ -1,328 +1,329 @@ (* Title: Pure/tactic.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Fundamental tactics. *) signature BASIC_TACTIC = sig val trace_goalno_tac: (int -> tactic) -> int -> tactic val rule_by_tactic: Proof.context -> tactic -> thm -> thm val assume_tac: Proof.context -> int -> tactic val eq_assume_tac: int -> tactic val compose_tac: Proof.context -> (bool * thm * int) -> int -> tactic val make_elim: thm -> thm val biresolve0_tac: (bool * thm) list -> int -> tactic val biresolve_tac: Proof.context -> (bool * thm) list -> int -> tactic val resolve0_tac: thm list -> int -> tactic val resolve_tac: Proof.context -> thm list -> int -> tactic val eresolve0_tac: thm list -> int -> tactic val eresolve_tac: Proof.context -> thm list -> int -> tactic val forward_tac: Proof.context -> thm list -> int -> tactic val dresolve0_tac: thm list -> int -> tactic val dresolve_tac: Proof.context -> thm list -> int -> tactic val ares_tac: Proof.context -> thm list -> int -> tactic val solve_tac: Proof.context -> thm list -> int -> tactic val bimatch_tac: Proof.context -> (bool * thm) list -> int -> tactic val match_tac: Proof.context -> thm list -> int -> tactic val ematch_tac: Proof.context -> thm list -> int -> tactic val dmatch_tac: Proof.context -> thm list -> int -> tactic val flexflex_tac: Proof.context -> tactic val distinct_subgoals_tac: tactic val cut_tac: thm -> int -> tactic val cut_rules_tac: thm list -> int -> tactic val cut_facts_tac: thm list -> int -> tactic val filter_thms: (term * term -> bool) -> int * term * thm list -> thm list val biresolution_from_nets_tac: Proof.context -> ('a list -> (bool * thm) list) -> bool -> 'a Net.net * 'a Net.net -> int -> tactic val biresolve_from_nets_tac: Proof.context -> (int * (bool * thm)) Net.net * (int * (bool * thm)) Net.net -> int -> tactic val bimatch_from_nets_tac: Proof.context -> (int * (bool * thm)) Net.net * (int * (bool * thm)) Net.net -> int -> tactic val filt_resolve_from_net_tac: Proof.context -> int -> (int * thm) Net.net -> int -> tactic val resolve_from_net_tac: Proof.context -> (int * thm) Net.net -> int -> tactic val match_from_net_tac: Proof.context -> (int * thm) Net.net -> int -> tactic val subgoals_of_brl: bool * thm -> int val lessb: (bool * thm) * (bool * thm) -> bool val rename_tac: string list -> int -> tactic val rotate_tac: int -> int -> tactic val defer_tac: int -> tactic val prefer_tac: int -> tactic val filter_prems_tac: Proof.context -> (term -> bool) -> int -> tactic end; signature TACTIC = sig include BASIC_TACTIC val insert_tagged_brl: 'a * (bool * thm) -> ('a * (bool * thm)) Net.net * ('a * (bool * thm)) Net.net -> ('a * (bool * thm)) Net.net * ('a * (bool * thm)) Net.net val delete_tagged_brl: bool * thm -> ('a * (bool * thm)) Net.net * ('a * (bool * thm)) Net.net -> ('a * (bool * thm)) Net.net * ('a * (bool * thm)) Net.net val eq_kbrl: ('a * (bool * thm)) * ('a * (bool * thm)) -> bool val build_net: thm list -> (int * thm) Net.net end; structure Tactic: TACTIC = struct (*Discover which goal is chosen: SOMEGOAL(trace_goalno_tac tac) *) fun trace_goalno_tac tac i st = case Seq.pull(tac i st) of NONE => Seq.empty | seqcell => (tracing ("Subgoal " ^ string_of_int i ^ " selected"); Seq.make(fn()=> seqcell)); (*Makes a rule by applying a tactic to an existing rule*) fun rule_by_tactic ctxt tac rl = let val thy = Proof_Context.theory_of ctxt; val ctxt' = Variable.declare_thm rl ctxt; val ((_, [st]), ctxt'') = Variable.import true [Thm.transfer thy rl] ctxt'; in (case Seq.pull (tac st) of NONE => raise THM ("rule_by_tactic", 0, [rl]) | SOME (st', _) => zero_var_indexes (singleton (Variable.export ctxt'' ctxt') st')) end; (*** Basic tactics ***) (*** The following fail if the goal number is out of range: thus (REPEAT (resolve_tac rules i)) stops once subgoal i disappears. *) (*Solve subgoal i by assumption*) fun assume_tac ctxt i = PRIMSEQ (Thm.assumption (SOME ctxt) i); (*Solve subgoal i by assumption, using no unification*) fun eq_assume_tac i = PRIMITIVE (Thm.eq_assumption i); (** Resolution/matching tactics **) (*The composition rule/state: no lifting or var renaming. The arg = (bires_flg, orule, m); see Thm.bicompose for explanation.*) fun compose_tac ctxt arg i = PRIMSEQ (Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false} arg i); (*Converts a "destruct" rule like P \ Q \ P to an "elimination" rule like \P \ Q; P \ R\ \ R *) fun make_elim rl = zero_var_indexes (rl RS revcut_rl); (*Attack subgoal i by resolution, using flags to indicate elimination rules*) fun biresolve0_tac brules i = PRIMSEQ (Thm.biresolution NONE false brules i); fun biresolve_tac ctxt brules i = PRIMSEQ (Thm.biresolution (SOME ctxt) false brules i); (*Resolution: the simple case, works for introduction rules*) fun resolve0_tac rules = biresolve0_tac (map (pair false) rules); fun resolve_tac ctxt rules = biresolve_tac ctxt (map (pair false) rules); (*Resolution with elimination rules only*) fun eresolve0_tac rules = biresolve0_tac (map (pair true) rules); fun eresolve_tac ctxt rules = biresolve_tac ctxt (map (pair true) rules); (*Forward reasoning using destruction rules.*) fun forward_tac ctxt rls = resolve_tac ctxt (map make_elim rls) THEN' assume_tac ctxt; (*Like forward_tac, but deletes the assumption after use.*) fun dresolve0_tac rls = eresolve0_tac (map make_elim rls); fun dresolve_tac ctxt rls = eresolve_tac ctxt (map make_elim rls); (*Use an assumption or some rules*) fun ares_tac ctxt rules = assume_tac ctxt ORELSE' resolve_tac ctxt rules; fun solve_tac ctxt rules = resolve_tac ctxt rules THEN_ALL_NEW assume_tac ctxt; (*Matching tactics -- as above, but forbid updating of state*) fun bimatch_tac ctxt brules i = PRIMSEQ (Thm.biresolution (SOME ctxt) true brules i); fun match_tac ctxt rules = bimatch_tac ctxt (map (pair false) rules); fun ematch_tac ctxt rules = bimatch_tac ctxt (map (pair true) rules); fun dmatch_tac ctxt rls = ematch_tac ctxt (map make_elim rls); (*Smash all flex-flex disagreement pairs in the proof state.*) fun flexflex_tac ctxt = PRIMSEQ (Thm.flexflex_rule (SOME ctxt)); (*Remove duplicate subgoals.*) fun distinct_subgoals_tac st = let val subgoals = Thm.cprems_of st; val (tab, n) = (subgoals, (Ctermtab.empty, 0)) |-> fold (fn ct => fn (tab, i) => if Ctermtab.defined tab ct then (tab, i) else (Ctermtab.update (ct, i) tab, i + 1)); val st' = if n = length subgoals then st else let val thy = Thm.theory_of_thm st; fun cert_prop i = Thm.global_cterm_of thy (Free (Name.bound i, propT)); val As = map (cert_prop o the o Ctermtab.lookup tab) subgoals; val As' = map cert_prop (0 upto (n - 1)); val C = cert_prop n; val template = Drule.list_implies (As, C); val inst = - (dest_Free (Thm.term_of C), Thm.cconcl_of st) :: - Ctermtab.fold (fn (ct, i) => cons ((Name.bound i, propT), ct)) tab []; + Frees.build + (Frees.add (dest_Free (Thm.term_of C), Thm.cconcl_of st) #> + Ctermtab.fold (fn (ct, i) => Frees.add ((Name.bound i, propT), ct)) tab); in Thm.assume template |> fold (Thm.elim_implies o Thm.assume) As |> fold_rev Thm.implies_intr As' |> Thm.implies_intr template - |> Thm.instantiate_frees ([], inst) + |> Thm.instantiate_frees (TFrees.empty, inst) |> Thm.elim_implies st end; in Seq.single st' end; (*** Applications of cut_rl ***) (*The conclusion of the rule gets assumed in subgoal i, while subgoal i+1,... are the premises of the rule.*) fun cut_tac rule i = resolve0_tac [cut_rl] i THEN resolve0_tac [rule] (i + 1); (*"Cut" a list of rules into the goal. Their premises will become new subgoals.*) fun cut_rules_tac ths i = EVERY (map (fn th => cut_tac th i) ths); (*As above, but inserts only facts (unconditional theorems); generates no additional subgoals. *) fun cut_facts_tac ths = cut_rules_tac (filter Thm.no_prems ths); (**** Indexing and filtering of theorems ****) (*Returns the list of potentially resolvable theorems for the goal "prem", using the predicate could(subgoal,concl). Resulting list is no longer than "limit"*) fun filter_thms could (limit, prem, ths) = let val pb = Logic.strip_assums_concl prem; (*delete assumptions*) fun filtr (limit, []) = [] | filtr (limit, th::ths) = if limit=0 then [] else if could(pb, Thm.concl_of th) then th :: filtr(limit-1, ths) else filtr(limit,ths) in filtr(limit,ths) end; (*** biresolution and resolution using nets ***) (** To preserve the order of the rules, tag them with increasing integers **) (*insert one tagged brl into the pair of nets*) fun insert_tagged_brl (kbrl as (k, (eres, th))) (inet, enet) = if eres then (case try Thm.major_prem_of th of SOME prem => (inet, Net.insert_term (K false) (prem, kbrl) enet) | NONE => error "insert_tagged_brl: elimination rule with no premises") else (Net.insert_term (K false) (Thm.concl_of th, kbrl) inet, enet); (*delete one kbrl from the pair of nets*) fun eq_kbrl ((_, (_, th)), (_, (_, th'))) = Thm.eq_thm_prop (th, th') fun delete_tagged_brl (brl as (eres, th)) (inet, enet) = (if eres then (case try Thm.major_prem_of th of SOME prem => (inet, Net.delete_term eq_kbrl (prem, ((), brl)) enet) | NONE => (inet, enet)) (*no major premise: ignore*) else (Net.delete_term eq_kbrl (Thm.concl_of th, ((), brl)) inet, enet)) handle Net.DELETE => (inet,enet); (*biresolution using a pair of nets rather than rules. function "order" must sort and possibly filter the list of brls. boolean "match" indicates matching or unification.*) fun biresolution_from_nets_tac ctxt order match (inet, enet) = SUBGOAL (fn (prem, i) => let val hyps = Logic.strip_assums_hyp prem; val concl = Logic.strip_assums_concl prem; val kbrls = Net.unify_term inet concl @ maps (Net.unify_term enet) hyps; in PRIMSEQ (Thm.biresolution (SOME ctxt) match (order kbrls) i) end); (*versions taking pre-built nets. No filtering of brls*) fun biresolve_from_nets_tac ctxt = biresolution_from_nets_tac ctxt order_list false; fun bimatch_from_nets_tac ctxt = biresolution_from_nets_tac ctxt order_list true; (*** Simpler version for resolve_tac -- only one net, and no hyps ***) (*insert one tagged rl into the net*) fun insert_krl (krl as (k,th)) = Net.insert_term (K false) (Thm.concl_of th, krl); (*build a net of rules for resolution*) fun build_net rls = fold_rev insert_krl (tag_list 1 rls) Net.empty; (*resolution using a net rather than rules; pred supports filt_resolve_tac*) fun filt_resolution_from_net_tac ctxt match pred net = SUBGOAL (fn (prem, i) => let val krls = Net.unify_term net (Logic.strip_assums_concl prem) in if pred krls then PRIMSEQ (Thm.biresolution (SOME ctxt) match (map (pair false) (order_list krls)) i) else no_tac end); (*Resolve the subgoal using the rules (making a net) unless too flexible, which means more than maxr rules are unifiable. *) fun filt_resolve_from_net_tac ctxt maxr net = let fun pred krls = length krls <= maxr in filt_resolution_from_net_tac ctxt false pred net end; (*versions taking pre-built nets*) fun resolve_from_net_tac ctxt = filt_resolution_from_net_tac ctxt false (K true); fun match_from_net_tac ctxt = filt_resolution_from_net_tac ctxt true (K true); (*** For Natural Deduction using (bires_flg, rule) pairs ***) (*The number of new subgoals produced by the brule*) fun subgoals_of_brl (true, rule) = Thm.nprems_of rule - 1 | subgoals_of_brl (false, rule) = Thm.nprems_of rule; (*Less-than test: for sorting to minimize number of new subgoals*) fun lessb (brl1,brl2) = subgoals_of_brl brl1 < subgoals_of_brl brl2; (*Renaming of parameters in a subgoal*) fun rename_tac xs i = case find_first (not o Symbol_Pos.is_identifier) xs of SOME x => error ("Not an identifier: " ^ x) | NONE => PRIMITIVE (Thm.rename_params_rule (xs, i)); (*rotate_tac n i: rotate the assumptions of subgoal i by n positions, from right to left if n is positive, and from left to right if n is negative.*) fun rotate_tac 0 i = all_tac | rotate_tac k i = PRIMITIVE (Thm.rotate_rule k i); (*Rotate the given subgoal to be the last.*) fun defer_tac i = PRIMITIVE (Thm.permute_prems (i - 1) 1); (*Rotate the given subgoal to be the first.*) fun prefer_tac i = PRIMITIVE (Thm.permute_prems (i - 1) 1 #> Thm.permute_prems 0 ~1); (*Remove premises that do not satisfy pred; fails if all prems satisfy pred.*) fun filter_prems_tac ctxt pred = let fun Then NONE tac = SOME tac | Then (SOME tac) tac' = SOME (tac THEN' tac'); fun thins H (tac, n) = if pred H then (tac, n + 1) else (Then tac (rotate_tac n THEN' eresolve_tac ctxt [thin_rl]), 0); in SUBGOAL (fn (goal, i) => let val Hs = Logic.strip_assums_hyp goal in (case fst (fold thins Hs (NONE, 0)) of NONE => no_tac | SOME tac => tac i) end) end; end; structure Basic_Tactic: BASIC_TACTIC = Tactic; open Basic_Tactic; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2390 +1,2392 @@ (* Title: Pure/thm.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius The very core of Isabelle's Meta Logic: certified types and terms, derivations, theorems, inference rules (including lifting and resolution), oracles. *) infix 0 RS RSN; signature BASIC_THM = sig type ctyp type cterm exception CTERM of string * cterm list type thm type conv = cterm -> thm exception THM of string * int * thm list val RSN: thm * (int * thm) -> thm val RS: thm * thm -> thm end; signature THM = sig include BASIC_THM (*certified types*) val typ_of: ctyp -> typ val global_ctyp_of: theory -> typ -> ctyp val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list val dest_ctypN: int -> ctyp -> ctyp val dest_ctyp0: ctyp -> ctyp val dest_ctyp1: ctyp -> ctyp val make_ctyp: ctyp -> ctyp list -> ctyp (*certified terms*) val term_of: cterm -> term val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int val global_cterm_of: theory -> term -> cterm val cterm_of: Proof.context -> term -> cterm val renamed_term: term -> cterm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm val dest_fun2: cterm -> cterm val dest_arg1: cterm -> cterm val dest_abs: string option -> cterm -> cterm * cterm val rename_tvar: indexname -> ctyp -> ctyp val var: indexname * ctyp -> cterm val apply: cterm -> cterm -> cterm val lambda_name: string * cterm -> cterm -> cterm val lambda: cterm -> cterm -> cterm val adjust_maxidx_cterm: int -> cterm -> cterm val incr_indexes_cterm: int -> cterm -> cterm - val match: cterm * cterm -> - ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list - val first_order_match: cterm * cterm -> - ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list + val match: cterm * cterm -> ctyp TVars.table * cterm Vars.table + val first_order_match: cterm * cterm -> ctyp TVars.table * cterm Vars.table (*theorems*) val fold_terms: {hyps: bool} -> (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: {hyps: bool} -> (typ -> bool) -> (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: {hyps: bool} -> (term -> bool) -> (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int val shyps_of: thm -> sort Ord_List.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list val thm_ord: thm ord exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm val weaken_sorts: sort list -> cterm -> cterm val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T val extra_shyps: thm -> sort list val strip_shyps: thm -> thm val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: Names.set * Names.set -> int -> thm -> thm val generalize_cterm: Names.set * Names.set -> int -> cterm -> cterm val generalize_ctyp: Names.set -> int -> ctyp -> ctyp - val instantiate: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> - thm -> thm - val instantiate_cterm: ((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> - cterm -> cterm + val instantiate: ctyp TVars.table * cterm Vars.table -> thm -> thm + val instantiate_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm val varifyT_global': TFrees.set -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val thynames_of_arity: theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; fun typ_of (Ctyp {T, ...}) = T; fun global_ctyp_of thy raw_T = let val T = Sign.certify_typ thy raw_T; val maxidx = Term.maxidx_of_typ T; val sorts = Sorts.insert_typ T []; in Ctyp {cert = Context.Certificate thy, T = T, maxidx = maxidx, sorts = sorts} end; val ctyp_of = global_ctyp_of o Proof_Context.theory_of; fun dest_ctyp (Ctyp {cert, T = Type (_, Ts), maxidx, sorts}) = map (fn T => Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}) Ts | dest_ctyp cT = raise TYPE ("dest_ctyp", [typ_of cT], []); fun dest_ctypN n (Ctyp {cert, T, maxidx, sorts}) = let fun err () = raise TYPE ("dest_ctypN", [T], []) in (case T of Type (_, Ts) => Ctyp {cert = cert, T = nth Ts n handle General.Subscript => err (), maxidx = maxidx, sorts = sorts} | _ => err ()) end; val dest_ctyp0 = dest_ctypN 0; val dest_ctyp1 = dest_ctypN 1; fun join_certificate_ctyp (Ctyp {cert, ...}) cert0 = Context.join_certificate (cert0, cert); fun union_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sorts.union sorts0 sorts; fun maxidx_ctyp (Ctyp {maxidx, ...}) maxidx0 = Int.max (maxidx0, maxidx); fun make_ctyp (Ctyp {cert, T, maxidx = _, sorts = _}) cargs = let val As = map typ_of cargs; fun err () = raise TYPE ("make_ctyp", T :: As, []); in (case T of Type (a, args) => Ctyp { cert = fold join_certificate_ctyp cargs cert, maxidx = fold maxidx_ctyp cargs ~1, sorts = fold union_sorts_ctyp cargs [], T = if length args = length cargs then Type (a, As) else err ()} | _ => err ()) end; (** certified terms **) (*certified terms with checked typ, maxidx, and sorts*) datatype cterm = Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: sort Ord_List.T}; exception CTERM of string * cterm list; fun term_of (Cterm {t, ...}) = t; fun typ_of_cterm (Cterm {T, ...}) = T; fun ctyp_of_cterm (Cterm {cert, T, maxidx, sorts, ...}) = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = sorts}; fun maxidx_of_cterm (Cterm {maxidx, ...}) = maxidx; fun global_cterm_of thy tm = let val (t, T, maxidx) = Sign.certify_term thy tm; val sorts = Sorts.insert_term t []; in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); val fast_term_ord = Term_Ord.fast_term_ord o apply2 term_of; val term_ord = Term_Ord.term_ord o apply2 term_of; (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun dest_abs a (Cterm {t = Abs (x, T, t), T = Type ("fun", [_, U]), cert, maxidx, sorts}) = let val (y', t') = Term.dest_abs (the_default x a, T, t) in (Cterm {t = Free (y', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_abs _ ct = raise CTERM ("dest_abs", [ct]); (* constructors *) fun rename_tvar (a, i) (Ctyp {cert, T, maxidx, sorts}) = let val S = (case T of TFree (_, S) => S | TVar (_, S) => S | _ => raise TYPE ("rename_tvar: no variable", [T], [])); val _ = if i < 0 then raise TYPE ("rename_tvar: bad index", [TVar ((a, i), S)], []) else (); in Ctyp {cert = cert, T = TVar ((a, i), S), maxidx = Int.max (i, maxidx), sorts = sorts} end; fun var ((x, i), Ctyp {cert, T, maxidx, sorts}) = if i < 0 then raise TERM ("var: bad index", [Var ((x, i), T)]) else Cterm {cert = cert, t = Var ((x, i), T), T = T, maxidx = Int.max (i, maxidx), sorts = sorts}; fun apply (cf as Cterm {t = f, T = Type ("fun", [dty, rty]), maxidx = maxidx1, sorts = sorts1, ...}) (cx as Cterm {t = x, T, maxidx = maxidx2, sorts = sorts2, ...}) = if T = dty then Cterm {cert = join_certificate0 (cf, cx), t = f $ x, T = rty, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} else raise CTERM ("apply: types don't agree", [cf, cx]) | apply cf cx = raise CTERM ("apply: first arg is not a function", [cf, cx]); fun lambda_name (x, ct1 as Cterm {t = t1, T = T1, maxidx = maxidx1, sorts = sorts1, ...}) (ct2 as Cterm {t = t2, T = T2, maxidx = maxidx2, sorts = sorts2, ...}) = let val t = Term.lambda_name (x, t1) t2 in Cterm {cert = join_certificate0 (ct1, ct2), t = t, T = T1 --> T2, maxidx = Int.max (maxidx1, maxidx2), sorts = Sorts.union sorts1 sorts2} end; fun lambda t u = lambda_name ("", t) u; (* indexes *) fun adjust_maxidx_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if maxidx = i then ct else if maxidx < i then Cterm {maxidx = i, cert = cert, t = t, T = T, sorts = sorts} else Cterm {maxidx = Int.max (maxidx_of_term t, i), cert = cert, t = t, T = T, sorts = sorts}; fun incr_indexes_cterm i (ct as Cterm {cert, t, T, maxidx, sorts}) = if i < 0 then raise CTERM ("negative increment", [ct]) else if i = 0 then ct else Cterm {cert = cert, t = Logic.incr_indexes ([], [], i) t, T = Logic.incr_tvar i T, maxidx = maxidx + i, sorts = sorts}; (*** Derivations and Theorems ***) (* sort constraints *) type constraint = {theory: theory, typ: typ, sort: sort}; local val constraint_ord : constraint ord = Context.theory_id_ord o apply2 (Context.theory_id o #theory) ||| Term_Ord.typ_ord o apply2 #typ ||| Term_Ord.sort_ord o apply2 #sort; val smash_atyps = map_atyps (fn TVar (_, S) => Term.aT S | TFree (_, S) => Term.aT S | T => T); in val union_constraints = Ord_List.union constraint_ord; fun insert_constraints thy (T, S) = let val ignored = S = [] orelse (case T of TFree (_, S') => S = S' | TVar (_, S') => S = S' | _ => false); in if ignored then I else Ord_List.insert constraint_ord {theory = thy, typ = smash_atyps T, sort = S} end; fun insert_constraints_env thy env = let val tyenv = Envir.type_env env; fun insert ([], _) = I | insert (S, T) = insert_constraints thy (Envir.norm_type tyenv T, S); in tyenv |> Vartab.fold (insert o #2) end; end; (* datatype thm *) datatype thm = Thm of deriv * (*derivation*) {cert: Context.certificate, (*background theory certificate*) tags: Properties.T, (*additional annotations/comments*) maxidx: int, (*maximum index of any Var or TVar*) constraints: constraint Ord_List.T, (*implicit proof obligations for sort constraints*) shyps: sort Ord_List.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps; fun fold_atomic_ctyps h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms h o fold_types o fold_atyps) (fn T => if g T then f (ctyp T) else I) th end; fun fold_atomic_cterms h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps}; fun apply t T = if g t then f (cterm t T) else I; in (fold_terms h o fold_aterms) (fn t as Const (_, T) => apply t T | t as Free (_, T) => apply t T | t as Var (_, T) => apply t T | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; val theory_name = Context.theory_id_name o theory_id; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; fun no_prems th = nprems_of th = 0; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* thm order: ignores theory context! *) val thm_ord = Term_Ord.fast_term_ord o apply2 prop_of ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of ||| list_ord Term_Ord.fast_term_ord o apply2 hyps_of ||| list_ord Term_Ord.sort_ord o apply2 shyps_of; (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = (Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt)) then (ctxt, transfer' ctxt th) else (Context.raw_transfer (theory_of_thm th) ctxt, th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); val sorts = Sorts.union sorts1 sorts2; fun mk_cTinst ((a, i), (S, T)) = (((a, i), S), Ctyp {T = T, cert = cert, maxidx = maxidx2, sorts = sorts}); fun mk_ctinst ((x, i), (U, t)) = let val T = Envir.subst_type Tinsts U in (((x, i), T), Cterm {t = t, T = T, cert = cert, maxidx = maxidx2, sorts = sorts}) end; - in (Vartab.fold (cons o mk_cTinst) Tinsts [], Vartab.fold (cons o mk_ctinst) tinsts []) end; + in + (TVars.build (Vartab.fold (TVars.add o mk_cTinst) Tinsts), + Vars.build (Vartab.fold (Vars.add o mk_ctinst) tinsts)) + end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; fun weaken_sorts raw_sorts ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val thy = theory_of_cterm ct; val more_sorts = Sorts.make (map (Sign.certify_sort thy) raw_sorts); val sorts' = Sorts.union sorts more_sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.unions_oracles [oracles1, oracles2]; val thms = Proofterm.unions_thms [thms1, thms2]; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; val _ = forall (fn (j, _) => i <> j) promises orelse err "bad dependencies"; val _ = join_promises promises; in thm end; fun future future_thm ct = let val Cterm {cert = cert, t = prop, T, maxidx, sorts} = ct; val _ = T <> propT andalso raise CTERM ("future: prop expected", [ct]); val _ = if Proofterm.proofs_enabled () then raise CTERM ("future: proof terms enabled", [ct]) else (); val i = serial (); val future = future_thm |> Future.map (future_result i cert sorts prop); in Thm (make_deriv [(i, future)] [] [] MinProof, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end; (** Axioms **) fun axiom thy name = (case Name_Space.lookup (Theory.axiom_table thy) name of SOME prop => let val der = deriv_rule0 (fn () => Proofterm.axm_proof name prop); val cert = Context.Certificate thy; val maxidx = maxidx_of_term prop; val shyps = Sorts.insert_term prop []; in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 ||| fast_string_ord o apply2 #3 ||| list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes); val extend = I; fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; val thy = Context.certificate_theory cert; val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*Dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps; (*Remove extra sorts that are witnessed by type signature information*) fun strip_shyps thm = (case thm of Thm (_, {shyps = [], ...}) => thm | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) => let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val minimize = Sorts.minimize_sort algebra; val le = Sorts.sort_le algebra; fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1)); fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)]; val present = (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; val extra = fold (Sorts.remove_sort o #2) present shyps; val witnessed = Sign.witness_sorts thy present extra; val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize); val extra' = non_witnessed |> map_filter (fn (S, _) => if non_witnessed |> exists (fn (S', _) => lt (S', S)) then NONE else SOME S) |> Sorts.make; val constrs' = non_witnessed |> maps (fn (S1, S2) => let val S0 = the (find_first (fn S => le (S, S1)) extra') in rel (S0, S1) @ rel (S1, S2) end); val constraints' = fold (insert_constraints thy) (witnessed @ constrs') constraints; val shyps' = fold (Sorts.insert_sort o #2) present extra'; in Thm (deriv_rule_unconditional (Proofterm.strip_shyps_proof algebra present witnessed extra') der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs, prop = prop}) end) |> solve_constraints; (*** Closed theorems with official name ***) (*non-deterministic, depends on unknown promises*) fun derivation_closed (Thm (Deriv {body, ...}, _)) = Proofterm.compact_proof (Proofterm.proof_of body); (*non-deterministic, depends on unknown promises*) fun raw_derivation_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (Proofterm.proof_of body); fun expand_name (Thm (Deriv {body, ...}, {shyps, hyps, prop, ...})) = let val self_id = (case Proofterm.get_identity shyps hyps prop (Proofterm.proof_of body) of NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = if self_id header orelse #name header = "" then SOME "" else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_approximative_name shyps hyps prop (proof_of thm); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = Proofterm.get_id shyps hyps prop (proof_of thm); (*dependencies of PThm node*) fun thm_deps (thm as Thm (Deriv {promises = [], body = PBody {thms, ...}, ...}, _)) = (case (derivation_id thm, thms) of (SOME {serial = i, ...}, [(j, thm_node)]) => if i = j then Proofterm.thm_node_thms thm_node else thms | _ => thms) | thm_deps thm = raise THM ("thm_deps: bad promises", 0, [thm]); fun name_derivation name_pos = strip_shyps #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop) | 1 => (((name, Position.thread_data ()), SOME prop), MinProof) | 0 => (((name, Position.none), NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = remove_hyps A hyps, tpairs = tpairs, prop = Logic.mk_implies (A, prop)}); (*Implication elimination A \ B A ------------ B *) fun implies_elim thAB thA = let val Thm (derA, {maxidx = maxidx1, hyps = hypsA, constraints = constraintsA, shyps = shypsA, tpairs = tpairsA, prop = propA, ...}) = thA and Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...}) = thAB; fun err () = raise THM ("implies_elim: major premise", 0, [thAB, thA]); in (case prop of Const ("Pure.imp", _) $ A $ B => if A aconv propA then Thm (deriv_rule2 (curry Proofterm.%%) der derA, {cert = join_certificate2 (thAB, thA), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraintsA constraints, shyps = Sorts.union shypsA shyps, hyps = union_hyps hypsA hyps, tpairs = union_tpairs tpairsA tpairs, prop = B}) else err () | _ => err ()) end; (*Forall introduction. The Free or Var x must not be free in the hypotheses. [x] : A ------ \x. A *) fun forall_intr (ct as Cterm {maxidx = maxidx1, t = x, T, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = let fun result a = Thm (deriv_rule1 (Proofterm.forall_intr_proof (a, x) NONE) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.all_const T $ Abs (a, T, abstract_over (x, prop))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("forall_intr: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result a) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result a) | _ => raise THM ("forall_intr: not a variable", 0, [th])) end; (*Forall elimination \x. A ------ A[t/x] *) fun forall_elim (ct as Cterm {t, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of Const ("Pure.all", Type ("fun", [Type ("fun", [qary, _]), _])) $ A => if T <> qary then raise THM ("forall_elim: type mismatch", 0, [th]) else Thm (deriv_rule1 (Proofterm.% o rpair (SOME t)) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Term.betapply (A, t)}) | _ => raise THM ("forall_elim: not quantified", 0, [th])); (* Equality *) (*Reflexivity t \ t *) fun reflexive (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t)}); (*Symmetry t \ u ------ u \ t *) fun symmetric (th as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = (case prop of (eq as Const ("Pure.eq", _)) $ t $ u => Thm (deriv_rule1 Proofterm.symmetric_proof der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = eq $ u $ t}) | _ => raise THM ("symmetric", 0, [th])); (*Transitivity t1 \ u u \ t2 ------------------ t1 \ t2 *) fun transitive th1 th2 = let val Thm (der1, {maxidx = maxidx1, hyps = hyps1, constraints = constraints1, shyps = shyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, hyps = hyps2, constraints = constraints2, shyps = shyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("transitive: " ^ msg, 0, [th1, th2]); in case (prop1, prop2) of ((eq as Const ("Pure.eq", Type (_, [U, _]))) $ t1 $ u, Const ("Pure.eq", _) $ u' $ t2) => if not (u aconv u') then err "middle term" else Thm (deriv_rule2 (Proofterm.transitive_proof U u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = eq $ t1 $ t2}) | _ => err "premises" end; (*Beta-conversion (\x. t) u \ t[u/x] fully beta-reduces the term if full = true *) fun beta_conversion full (Cterm {cert, t, T = _, maxidx, sorts}) = let val t' = if full then Envir.beta_norm t else (case t of Abs (_, _, bodt) $ u => subst_bound (u, bodt) | _ => raise THM ("beta_conversion: not a redex", 0, [])); in Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, t')}) end; fun eta_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_contract t)}); fun eta_long_conversion (Cterm {cert, t, T = _, maxidx, sorts}) = Thm (deriv_rule0 (fn () => Proofterm.reflexive_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_equals (t, Envir.eta_long [] t)}); (*The abstraction rule. The Free or Var x must not be free in the hypotheses. The bound variable will be named "a" (since x will be something like x320) t \ u -------------- \x. t \ \x. u *) fun abstract_rule a (Cterm {t = x, T, sorts, ...}) (th as Thm (der, {cert, maxidx, hyps, constraints, shyps, tpairs, prop, ...})) = let val (t, u) = Logic.dest_equals prop handle TERM _ => raise THM ("abstract_rule: premise not an equality", 0, [th]); val result = Thm (deriv_rule1 (Proofterm.abstract_rule_proof (a, x)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = Sorts.union sorts shyps, hyps = hyps, tpairs = tpairs, prop = Logic.mk_equals (Abs (a, T, abstract_over (x, t)), Abs (a, T, abstract_over (x, u)))}); fun check_occs a x ts = if exists (fn t => Logic.occs (x, t)) ts then raise THM ("abstract_rule: variable " ^ quote a ^ " free in assumptions", 0, [th]) else (); in (case x of Free (a, _) => (check_occs a x hyps; check_occs a x (terms_of_tpairs tpairs); result) | Var ((a, _), _) => (check_occs a x (terms_of_tpairs tpairs); result) | _ => raise THM ("abstract_rule: not a variable", 0, [th])) end; (*The combination rule f \ g t \ u ------------- f t \ g u *) fun combination th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun chktypes fT tT = (case fT of Type ("fun", [T1, _]) => if T1 <> tT then raise THM ("combination: types", 0, [th1, th2]) else () | _ => raise THM ("combination: not function type", 0, [th1, th2])); in (case (prop1, prop2) of (Const ("Pure.eq", Type ("fun", [fT, _])) $ f $ g, Const ("Pure.eq", Type ("fun", [tT, _])) $ t $ u) => (chktypes fT tT; Thm (deriv_rule2 (Proofterm.combination_proof f g t u) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (f $ t, g $ u)})) | _ => raise THM ("combination: premises", 0, [th1, th2])) end; (*Equality introduction A \ B B \ A ---------------- A \ B *) fun equal_intr th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_intr: " ^ msg, 0, [th1, th2]); in (case (prop1, prop2) of (Const("Pure.imp", _) $ A $ B, Const("Pure.imp", _) $ B' $ A') => if A aconv A' andalso B aconv B' then Thm (deriv_rule2 (Proofterm.equal_intr_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = Logic.mk_equals (A, B)}) else err "not equal" | _ => err "premises") end; (*The equal propositions rule A \ B A --------- B *) fun equal_elim th1 th2 = let val Thm (der1, {maxidx = maxidx1, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs = tpairs1, prop = prop1, ...}) = th1 and Thm (der2, {maxidx = maxidx2, constraints = constraints2, shyps = shyps2, hyps = hyps2, tpairs = tpairs2, prop = prop2, ...}) = th2; fun err msg = raise THM ("equal_elim: " ^ msg, 0, [th1, th2]); in (case prop1 of Const ("Pure.eq", _) $ A $ B => if prop2 aconv A then Thm (deriv_rule2 (Proofterm.equal_elim_proof A B) der1 der2, {cert = join_certificate2 (th1, th2), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = union_constraints constraints1 constraints2, shyps = Sorts.union shyps1 shyps2, hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; val constraints' = insert_constraints_env (Context.certificate_theory cert') env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize (tfrees, frees) idx th = if Names.is_empty tfrees andalso Names.is_empty frees then th else let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if Names.is_empty tfrees then K false else Term.exists_subtype (fn TFree (a, _) => Names.defined tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse Names.defined frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if Names.is_empty tfrees andalso Names.is_empty frees then ct else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) = if Names.is_empty tfrees then cT else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c); val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert); val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert); fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts; val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts); val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts); -fun add_instT thy (v as (_, S), Ctyp {T = U, maxidx, ...}) = - if Sign.of_sort thy (U, S) then TVars.add (v, (U, maxidx)) +fun make_instT thy (v as (_, S)) (Ctyp {T = U, maxidx, ...}) = + if Sign.of_sort thy (U, S) then (U, maxidx) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []); -fun add_inst thy (v as (_, T), Cterm {t = u, T = U, maxidx, ...}) = - if T = U then Vars.add (v, (u, maxidx)) +fun make_inst thy (v as (_, T)) (Cterm {t = u, T = U, maxidx, ...}) = + if T = U then (u, maxidx) else let fun pretty_typing t ty = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy ty]; val msg = Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing (Var v) T, Pretty.fbrk, pretty_typing u U]) in raise TYPE (msg, [T, U], [Var v, u]) end; fun prep_insts (instT, inst) (cert, sorts) = let val cert' = cert - |> fold add_instT_cert instT - |> fold add_inst_cert inst; - val thy' = Context.certificate_theory cert' - handle ERROR msg => raise CONTEXT (msg, map #2 instT, map #2 inst, [], NONE); + |> TVars.fold add_instT_cert instT + |> Vars.fold add_inst_cert inst; + val thy' = + Context.certificate_theory cert' handle ERROR msg => + raise CONTEXT (msg, TVars.dest instT |> map #2, Vars.dest inst |> map #2, [], NONE); val sorts' = sorts - |> fold add_instT_sorts instT - |> fold add_inst_sorts inst; + |> TVars.fold add_instT_sorts instT + |> Vars.fold add_inst_sorts inst; - val instT' = TVars.build (fold (add_instT thy') instT); - val inst' = Vars.build (fold (add_inst thy') inst); + val instT' = TVars.map (make_instT thy') instT; + val inst' = Vars.map (make_inst thy') inst; in ((instT', inst'), (cert', sorts')) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) -fun instantiate ([], []) th = th - | instantiate (instT, inst) th = - let - val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; - val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) - handle CONTEXT (msg, cTs, cts, ths, context) => - raise CONTEXT (msg, cTs, cts, th :: ths, context); - - val subst = Term_Subst.instantiate_maxidx (instT', inst'); - val (prop', maxidx1) = subst prop ~1; - val (tpairs', maxidx') = - fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; +fun instantiate (instT, inst) th = + if TVars.is_empty instT andalso Vars.is_empty inst then th + else + let + val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; + val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) + handle CONTEXT (msg, cTs, cts, ths, context) => + raise CONTEXT (msg, cTs, cts, th :: ths, context); - val thy' = Context.certificate_theory cert'; - val constraints' = - TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) - instT' constraints; - in - Thm (deriv_rule1 - (fn d => Proofterm.instantiate (TVars.map (K #1) instT', Vars.map (K #1) inst') d) der, - {cert = cert', - tags = [], - maxidx = maxidx', - constraints = constraints', - shyps = shyps', - hyps = hyps, - tpairs = tpairs', - prop = prop'}) - |> solve_constraints - end - handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); + val subst = Term_Subst.instantiate_maxidx (instT', inst'); + val (prop', maxidx1) = subst prop ~1; + val (tpairs', maxidx') = + fold_map (fn (t, u) => fn i => subst t i ||>> subst u) tpairs maxidx1; -fun instantiate_cterm ([], []) ct = ct - | instantiate_cterm (instT, inst) ct = - let - val Cterm {cert, t, T, sorts, ...} = ct; - val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); - val subst = Term_Subst.instantiate_maxidx (instT', inst'); - val substT = Term_Subst.instantiateT_maxidx instT'; - val (t', maxidx1) = subst t ~1; - val (T', maxidx') = substT T maxidx1; - in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end - handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); + val thy' = Context.certificate_theory cert'; + val constraints' = + TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) + instT' constraints; + in + Thm (deriv_rule1 + (fn d => Proofterm.instantiate (TVars.map (K #1) instT', Vars.map (K #1) inst') d) der, + {cert = cert', + tags = [], + maxidx = maxidx', + constraints = constraints', + shyps = shyps', + hyps = hyps, + tpairs = tpairs', + prop = prop'}) + |> solve_constraints + end + handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); + +fun instantiate_cterm (instT, inst) ct = + if TVars.is_empty instT andalso Vars.is_empty inst then ct + else + let + val Cterm {cert, t, T, sorts, ...} = ct; + val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); + val subst = Term_Subst.instantiate_maxidx (instT', inst'); + val substT = Term_Subst.instantiateT_maxidx instT'; + val (t', maxidx1) = subst t ~1; + val (T', maxidx') = substT T maxidx1; + in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end + handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Internalize sort constraints of type variables*) val unconstrainT = strip_shyps #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; val tfrees = build_rev (Term.add_tfree_names prop); val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy) shyps prop ps body; val der' = make_deriv [] [] [pthm] proof; val prop' = Proofterm.thm_node_prop (#2 pthm); in Thm (der', {cert = cert, tags = [], maxidx = maxidx_of_term prop', constraints = [], shyps = [[]], (*potentially redundant*) hyps = [], tpairs = [], prop = prop'}) end); (*Replace all TFrees not fixed or in the hyps by new TVars*) fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = let val tfrees = fold TFrees.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; val varifyT_global = #2 o varifyT_global' TFrees.empty; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" else if not (null (extra_shyps thm)) then err "theorem may not contain sort hypotheses" else if not (null (tpairs_of thm)) then err "theorem may not contain flex-flex pairs" else prop_of thm end; (*** Inference rules for tactics ***) (*Destruct proof state into constraints, other goals, goal(i), rest *) fun dest_state (state as Thm (_, {prop,tpairs,...}), i) = (case Logic.strip_prems(i, [], prop) of (B::rBs, C) => (tpairs, rev rBs, B, C) | _ => raise THM("dest_state", i, [state])) handle TERM _ => raise THM("dest_state", i, [state]); (*Prepare orule for resolution by lifting it over the parameters and assumptions of goal.*) fun lift_rule goal orule = let val Cterm {t = gprop, T, maxidx = gmax, sorts, ...} = goal; val inc = gmax + 1; val lift_abs = Logic.lift_abs inc gprop; val lift_all = Logic.lift_all inc gprop; val Thm (der, {maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = orule; val (As, B) = Logic.strip_horn prop; in if T <> propT then raise THM ("lift_rule: the term must have type prop", 0, []) else Thm (deriv_rule1 (Proofterm.lift_proof gprop inc prop) der, {cert = join_certificate1 (goal, orule), tags = [], maxidx = maxidx + inc, constraints = constraints, shyps = Sorts.union shyps sorts, (*sic!*) hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); val vars = Vartab.build (add_vars th1 #> add_vars th2); in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end val constraints' = union_constraints constraints1 constraints2 |> insert_constraints_env (Context.certificate_theory cert) env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; val tvars = build_rev (Term.add_tvars (prop_of thm)); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; - in instantiate (tinst, []) thm end + in instantiate (TVars.make tinst, Vars.empty) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun thynames_of_arity thy (a, c) = build (get_arities thy |> Aritytab.fold (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser))) |> sort (int_ord o apply2 #2) |> map #1; fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/variable.ML b/src/Pure/variable.ML --- a/src/Pure/variable.ML +++ b/src/Pure/variable.ML @@ -1,745 +1,745 @@ (* Title: Pure/variable.ML Author: Makarius Fixed type/term variables and polymorphic term abbreviations. *) signature VARIABLE = sig val names_of: Proof.context -> Name.context val binds_of: Proof.context -> (typ * term) Vartab.table val maxidx_of: Proof.context -> int val constraints_of: Proof.context -> typ Vartab.table * sort Vartab.table val is_declared: Proof.context -> string -> bool val check_name: binding -> string val default_type: Proof.context -> string -> typ option val def_type: Proof.context -> bool -> indexname -> typ option val def_sort: Proof.context -> indexname -> sort option val declare_maxidx: int -> Proof.context -> Proof.context val declare_names: term -> Proof.context -> Proof.context val declare_constraints: term -> Proof.context -> Proof.context val declare_internal: term -> Proof.context -> Proof.context val declare_term: term -> Proof.context -> Proof.context val declare_typ: typ -> Proof.context -> Proof.context val declare_prf: Proofterm.proof -> Proof.context -> Proof.context val declare_thm: thm -> Proof.context -> Proof.context val variant_frees: Proof.context -> term list -> (string * 'a) list -> (string * 'a) list val bind_term: indexname * term -> Proof.context -> Proof.context val unbind_term: indexname -> Proof.context -> Proof.context val maybe_bind_term: indexname * term option -> Proof.context -> Proof.context val expand_binds: Proof.context -> term -> term val lookup_const: Proof.context -> string -> string option val is_const: Proof.context -> string -> bool val declare_const: string * string -> Proof.context -> Proof.context val next_bound: string * typ -> Proof.context -> term * Proof.context val revert_bounds: Proof.context -> term -> term val is_body: Proof.context -> bool val set_body: bool -> Proof.context -> Proof.context val restore_body: Proof.context -> Proof.context -> Proof.context val improper_fixes: Proof.context -> Proof.context val restore_proper_fixes: Proof.context -> Proof.context -> Proof.context val is_improper: Proof.context -> string -> bool val is_fixed: Proof.context -> string -> bool val is_newly_fixed: Proof.context -> Proof.context -> string -> bool val fixed_ord: Proof.context -> string ord val intern_fixed: Proof.context -> string -> string val lookup_fixed: Proof.context -> string -> string option val revert_fixed: Proof.context -> string -> string val markup_fixed: Proof.context -> string -> Markup.T val markup: Proof.context -> string -> Markup.T val markup_entity_def: Proof.context -> string -> Markup.T val dest_fixes: Proof.context -> (string * string) list val add_fixed_names: Proof.context -> term -> string list -> string list val add_fixed: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_newly_fixed: Proof.context -> Proof.context -> term -> (string * typ) list -> (string * typ) list val add_free_names: Proof.context -> term -> string list -> string list val add_frees: Proof.context -> term -> (string * typ) list -> (string * typ) list val add_fixes_binding: binding list -> Proof.context -> string list * Proof.context val add_fixes: string list -> Proof.context -> string list * Proof.context val add_fixes_direct: string list -> Proof.context -> Proof.context val add_fixes_implicit: term -> Proof.context -> Proof.context val fix_dummy_patterns: term -> Proof.context -> term * Proof.context val variant_fixes: string list -> Proof.context -> string list * Proof.context val gen_all: Proof.context -> thm -> thm val export_terms: Proof.context -> Proof.context -> term list -> term list val exportT_terms: Proof.context -> Proof.context -> term list -> term list val exportT: Proof.context -> Proof.context -> thm list -> thm list val export_prf: Proof.context -> Proof.context -> Proofterm.proof -> Proofterm.proof val export: Proof.context -> Proof.context -> thm list -> thm list val export_morphism: Proof.context -> Proof.context -> morphism val invent_types: sort list -> Proof.context -> (string * sort) list * Proof.context val importT_inst: term list -> Proof.context -> typ TVars.table * Proof.context val import_inst: bool -> term list -> Proof.context -> (typ TVars.table * term Vars.table) * Proof.context val importT_terms: term list -> Proof.context -> term list * Proof.context val import_terms: bool -> term list -> Proof.context -> term list * Proof.context val importT: thm list -> Proof.context -> (ctyp TVars.table * thm list) * Proof.context val import_prf: bool -> Proofterm.proof -> Proof.context -> Proofterm.proof * Proof.context val import: bool -> thm list -> Proof.context -> ((ctyp TVars.table * cterm Vars.table) * thm list) * Proof.context val import_vars: Proof.context -> thm -> thm val tradeT: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list val trade: (Proof.context -> thm list -> thm list) -> Proof.context -> thm list -> thm list val is_bound_focus: Proof.context -> bool val set_bound_focus: bool -> Proof.context -> Proof.context val restore_bound_focus: Proof.context -> Proof.context -> Proof.context val focus_params: binding list option -> term -> Proof.context -> (string list * (string * typ) list) * Proof.context val focus: binding list option -> term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context val focus_cterm: binding list option -> cterm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val focus_subgoal: binding list option -> int -> thm -> Proof.context -> ((string * cterm) list * cterm) * Proof.context val warn_extra_tfrees: Proof.context -> Proof.context -> unit val polymorphic_types: Proof.context -> term list -> (indexname * sort) list * term list val polymorphic: Proof.context -> term list -> term list end; structure Variable: VARIABLE = struct (** local context data **) type fixes = (string * bool) Name_Space.table; val empty_fixes: fixes = Name_Space.empty_table Markup.fixedN; datatype data = Data of {names: Name.context, (*type/term variable names*) consts: string Symtab.table, (*consts within the local scope*) bounds: int * ((string * typ) * string) list, (*next index, internal name, type, external name*) fixes: fixes, (*term fixes -- global name space, intern ~> extern*) binds: (typ * term) Vartab.table, (*term bindings*) type_occs: string list Symtab.table, (*type variables -- possibly within term variables*) maxidx: int, (*maximum var index*) constraints: typ Vartab.table * (*type constraints*) sort Vartab.table}; (*default sorts*) fun make_data (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) = Data {names = names, consts = consts, bounds = bounds, fixes = fixes, binds = binds, type_occs = type_occs, maxidx = maxidx, constraints = constraints}; val empty_data = make_data (Name.context, Symtab.empty, (0, []), empty_fixes, Vartab.empty, Symtab.empty, ~1, (Vartab.empty, Vartab.empty)); structure Data = Proof_Data ( type T = data; fun init _ = empty_data; ); fun map_data f = Data.map (fn Data {names, consts, bounds, fixes, binds, type_occs, maxidx, constraints} => make_data (f (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints))); fun map_names f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (f names, consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_consts f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, f consts, bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_bounds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, f bounds, fixes, binds, type_occs, maxidx, constraints)); fun map_fixes f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, f fixes, binds, type_occs, maxidx, constraints)); fun map_binds f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, f binds, type_occs, maxidx, constraints)); fun map_type_occs f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, f type_occs, maxidx, constraints)); fun map_maxidx f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, f maxidx, constraints)); fun map_constraints f = map_data (fn (names, consts, bounds, fixes, binds, type_occs, maxidx, constraints) => (names, consts, bounds, fixes, binds, type_occs, maxidx, f constraints)); fun rep_data ctxt = Data.get ctxt |> (fn Data rep => rep); val names_of = #names o rep_data; val fixes_of = #fixes o rep_data; val fixes_space = Name_Space.space_of_table o fixes_of; val binds_of = #binds o rep_data; val type_occs_of = #type_occs o rep_data; val maxidx_of = #maxidx o rep_data; val constraints_of = #constraints o rep_data; val is_declared = Name.is_declared o names_of; val check_name = Name_Space.base_name o tap Binding.check; (** declarations **) (* default sorts and types *) fun default_type ctxt x = Vartab.lookup (#1 (constraints_of ctxt)) (x, ~1); fun def_type ctxt pattern xi = let val {binds, constraints = (types, _), ...} = rep_data ctxt in (case Vartab.lookup types xi of NONE => if pattern then NONE else Vartab.lookup binds xi |> Option.map (Type.mark_polymorphic o #1) | some => some) end; val def_sort = Vartab.lookup o #2 o constraints_of; (* maxidx *) val declare_maxidx = map_maxidx o Integer.max; (* names *) fun declare_type_names t = map_names (fold_types (fold_atyps Term.declare_typ_names) t) #> map_maxidx (fold_types Term.maxidx_typ t); fun declare_names t = declare_type_names t #> map_names (fold_aterms Term.declare_term_frees t) #> map_maxidx (Term.maxidx_term t); (* type occurrences *) fun decl_type_occsT T = fold_atyps (fn TFree (a, _) => Symtab.default (a, []) | _ => I) T; val decl_type_occs = fold_term_types (fn Free (x, _) => fold_atyps (fn TFree (a, _) => Symtab.insert_list (op =) (a, x) | _ => I) | _ => decl_type_occsT); val declare_type_occsT = map_type_occs o fold_types decl_type_occsT; val declare_type_occs = map_type_occs o decl_type_occs; (* constraints *) fun constrain_tvar (xi, raw_S) = let val S = #2 (Term_Position.decode_positionS raw_S) in if S = dummyS then Vartab.delete_safe xi else Vartab.update (xi, S) end; fun declare_constraints t = map_constraints (fn (types, sorts) => let val types' = fold_aterms (fn Free (x, T) => Vartab.update ((x, ~1), T) | Var v => Vartab.update v | _ => I) t types; val sorts' = (fold_types o fold_atyps) (fn TFree (x, S) => constrain_tvar ((x, ~1), S) | TVar v => constrain_tvar v | _ => I) t sorts; in (types', sorts') end) #> declare_type_occsT t #> declare_type_names t; (* common declarations *) fun declare_internal t = declare_names t #> declare_type_occs t #> Thm.declare_term_sorts t; fun declare_term t = declare_internal t #> declare_constraints t; val declare_typ = declare_term o Logic.mk_type; val declare_prf = Proofterm.fold_proof_terms_types declare_internal (declare_internal o Logic.mk_type); val declare_thm = Thm.fold_terms {hyps = true} declare_internal; (* renaming term/type frees *) fun variant_frees ctxt ts frees = let val names = names_of (fold declare_names ts ctxt); val xs = fst (fold_map Name.variant (map #1 frees) names); in xs ~~ map snd frees end; (** term bindings **) fun bind_term ((x, i), t) = let val u = Term.close_schematic_term t; val U = Term.fastype_of u; in declare_term u #> map_binds (Vartab.update ((x, i), (U, u))) end; val unbind_term = map_binds o Vartab.delete_safe; fun maybe_bind_term (xi, SOME t) = bind_term (xi, t) | maybe_bind_term (xi, NONE) = unbind_term xi; fun expand_binds ctxt = let val binds = binds_of ctxt; val get = fn Var (xi, _) => Vartab.lookup binds xi | _ => NONE; in Envir.beta_norm o Envir.expand_term get end; (** consts **) val lookup_const = Symtab.lookup o #consts o rep_data; val is_const = is_some oo lookup_const; val declare_fixed = map_consts o Symtab.delete_safe; val declare_const = map_consts o Symtab.update; (** bounds **) fun next_bound (a, T) ctxt = let val b = Name.bound (#1 (#bounds (rep_data ctxt))); val ctxt' = ctxt |> map_bounds (fn (next, bounds) => (next + 1, ((b, T), a) :: bounds)); in (Free (b, T), ctxt') end; fun revert_bounds ctxt t = (case #2 (#bounds (rep_data ctxt)) of [] => t | bounds => let val names = Term.declare_term_names t (names_of ctxt); val xs = rev (#1 (fold_map Name.variant (rev (map #2 bounds)) names)); fun subst ((b, T), _) x' = (Free (b, T), Syntax_Trans.mark_bound_abs (x', T)); in Term.subst_atomic (map2 subst bounds xs) t end); (** fixes **) (* inner body mode *) val inner_body = Config.declare_bool ("inner_body", \<^here>) (K false); val is_body = Config.apply inner_body; val set_body = Config.put inner_body; val restore_body = set_body o is_body; (* proper mode *) val proper_fixes = Config.declare_bool ("proper_fixes", \<^here>) (K true); val improper_fixes = Config.put proper_fixes false; val restore_proper_fixes = Config.put proper_fixes o Config.apply proper_fixes; fun is_improper ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (_, proper) => not proper | NONE => false); (* specialized name space *) val is_fixed = Name_Space.defined o fixes_of; fun is_newly_fixed inner outer = is_fixed inner andf (not o is_fixed outer); val fixed_ord = Name_Space.entry_ord o fixes_space; val intern_fixed = Name_Space.intern o fixes_space; fun lookup_fixed ctxt x = let val x' = intern_fixed ctxt x in if is_fixed ctxt x' then SOME x' else NONE end; fun revert_fixed ctxt x = (case Name_Space.lookup (fixes_of ctxt) x of SOME (x', _) => if intern_fixed ctxt x' = x then x' else x | NONE => x); fun markup_fixed ctxt x = Name_Space.markup (fixes_space ctxt) x |> Markup.name (revert_fixed ctxt x); fun markup ctxt x = if is_improper ctxt x then Markup.improper else if Name.is_skolem x then Markup.skolem else Markup.free; val markup_entity_def = Name_Space.markup_def o fixes_space; fun dest_fixes ctxt = Name_Space.fold_table (fn (x, (y, _)) => cons (y, x)) (fixes_of ctxt) [] |> sort (Name_Space.entry_ord (fixes_space ctxt) o apply2 #2); (* collect variables *) fun add_free_names ctxt = fold_aterms (fn Free (x, _) => not (is_fixed ctxt x) ? insert (op =) x | _ => I); fun add_frees ctxt = fold_aterms (fn Free (x, T) => not (is_fixed ctxt x) ? insert (op =) (x, T) | _ => I); fun add_fixed_names ctxt = fold_aterms (fn Free (x, _) => is_fixed ctxt x ? insert (op =) x | _ => I); fun add_fixed ctxt = fold_aterms (fn Free (x, T) => is_fixed ctxt x ? insert (op =) (x, T) | _ => I); fun add_newly_fixed ctxt' ctxt = fold_aterms (fn Free (x, T) => is_newly_fixed ctxt' ctxt x ? insert (op =) (x, T) | _ => I); (* declarations *) local fun err_dups dups = error ("Duplicate fixed variable(s): " ^ commas (map Binding.print dups)); fun new_fixed ((x, x'), pos) ctxt = if is_some (lookup_fixed ctxt x') then err_dups [Binding.make (x, pos)] else let val proper = Config.get ctxt proper_fixes; val context = Context.Proof ctxt |> Name_Space.map_naming (K Name_Space.global_naming) |> Context_Position.set_visible_generic false; in ctxt |> map_fixes (Name_Space.define context true (Binding.make (x', pos), (x, proper)) #> snd #> Name_Space.alias_table Name_Space.global_naming (Binding.make (x, pos)) x') |> declare_fixed x |> declare_constraints (Syntax.free x') end; fun new_fixes names' args = map_names (K names') #> fold new_fixed args #> pair (map (#2 o #1) args); in fun add_fixes_binding bs ctxt = let val _ = (case filter (Name.is_skolem o Binding.name_of) bs of [] => () | bads => error ("Illegal internal Skolem constant(s): " ^ commas (map Binding.print bads))); val _ = (case duplicates (op = o apply2 Binding.name_of) bs of [] => () | dups => err_dups dups); val xs = map check_name bs; val names = names_of ctxt; val (xs', names') = if is_body ctxt then fold_map Name.variant xs names |>> map Name.skolem else (xs, fold Name.declare xs names); in ctxt |> new_fixes names' ((xs ~~ xs') ~~ map Binding.pos_of bs) end; fun variant_names ctxt raw_xs = let val names = names_of ctxt; val xs = map (fn x => Name.clean x |> Name.is_internal x ? Name.internal) raw_xs; val (xs', names') = fold_map Name.variant xs names |>> (is_body ctxt ? map Name.skolem); in (names', xs ~~ xs') end; fun variant_fixes xs ctxt = let val (names', vs) = variant_names ctxt xs; in ctxt |> new_fixes names' (map (rpair Position.none) vs) end; fun bound_fixes xs ctxt = let val (names', vs) = variant_names ctxt (map #1 xs); val (ys, ctxt') = fold_map next_bound (map2 (fn (x', _) => fn (_, T) => (x', T)) vs xs) ctxt; val fixes = map2 (fn (x, _) => fn Free (y, _) => ((x, y), Position.none)) vs ys; in ctxt' |> new_fixes names' fixes end; end; val add_fixes = add_fixes_binding o map Binding.name; fun add_fixes_direct xs ctxt = ctxt |> set_body false |> (snd o add_fixes xs) |> restore_body ctxt; fun add_fixes_implicit t ctxt = ctxt |> not (is_body ctxt) ? add_fixes_direct (rev (add_free_names ctxt t [])); (* dummy patterns *) fun fix_dummy_patterns (Const ("Pure.dummy_pattern", T)) ctxt = let val ([x], ctxt') = ctxt |> set_body true |> add_fixes [Name.uu_] ||> restore_body ctxt in (Free (x, T), ctxt') end | fix_dummy_patterns (Abs (x, T, b)) ctxt = let val (b', ctxt') = fix_dummy_patterns b ctxt in (Abs (x, T, b'), ctxt') end | fix_dummy_patterns (t $ u) ctxt = let val (t', ctxt') = fix_dummy_patterns t ctxt; val (u', ctxt'') = fix_dummy_patterns u ctxt'; in (t' $ u', ctxt'') end | fix_dummy_patterns a ctxt = (a, ctxt); (** export -- generalize type/term variables (beware of closure sizes) **) fun gen_all ctxt th = let val i = Thm.maxidx_thm th (maxidx_of ctxt) + 1; fun gen (x, T) = Thm.forall_elim (Thm.cterm_of ctxt (Var ((x, i), T))); in fold gen (Drule.outer_params (Thm.prop_of th)) th end; fun export_inst inner outer = let val declared_outer = is_declared outer; val still_fixed = not o is_newly_fixed inner outer; val gen_fixes = Names.build (fixes_of inner |> Name_Space.fold_table (fn (y, _) => not (is_fixed outer y) ? Names.add_set y)); val type_occs_inner = type_occs_of inner; fun gen_fixesT ts = Names.build (fold decl_type_occs ts type_occs_inner |> Symtab.fold (fn (a, xs) => if declared_outer a orelse exists still_fixed xs then I else Names.add_set a)); in (gen_fixesT, gen_fixes) end; fun exportT_inst inner outer = #1 (export_inst inner outer); fun exportT_terms inner outer = let val mk_tfrees = exportT_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, Names.empty) (fold (Term.fold_types Term.maxidx_typ) ts maxidx + 1)) end; fun export_terms inner outer = let val (mk_tfrees, tfrees) = export_inst inner outer; val maxidx = maxidx_of outer; in fn ts => ts |> map (Term_Subst.generalize (mk_tfrees ts, tfrees) (fold Term.maxidx_term ts maxidx + 1)) end; fun export_prf inner outer prf = let val (mk_tfrees, frees) = export_inst (declare_prf prf inner) outer; val tfrees = mk_tfrees []; val maxidx = maxidx_of outer; val idx = Proofterm.maxidx_proof prf maxidx + 1; val gen_term = Term_Subst.generalize_same (tfrees, frees) idx; val gen_typ = Term_Subst.generalizeT_same tfrees idx; in Same.commit (Proofterm.map_proof_terms_same gen_term gen_typ) prf end; fun gen_export (mk_tfrees, frees) maxidx ths = let val tfrees = mk_tfrees (map Thm.full_prop_of ths); val idx = fold Thm.maxidx_thm ths maxidx + 1; in map (Thm.generalize (tfrees, frees) idx) ths end; fun exportT inner outer = gen_export (exportT_inst inner outer, Names.empty) (maxidx_of outer); fun export inner outer = gen_export (export_inst inner outer) (maxidx_of outer); fun export_morphism inner outer = let val fact = export inner outer; val term = singleton (export_terms inner outer); val typ = Logic.type_map term; in Morphism.transfer_morphism' inner $> Morphism.transfer_morphism' outer $> Morphism.morphism "Variable.export" {binding = [], typ = [typ], term = [term], fact = [fact]} end; (** import -- fix schematic type/term variables **) fun invent_types Ss ctxt = let val tfrees = Name.invent (names_of ctxt) Name.aT (length Ss) ~~ Ss; val ctxt' = fold (declare_constraints o Logic.mk_type o TFree) tfrees ctxt; in (tfrees, ctxt') end; fun importT_inst ts ctxt = let val tvars = TVars.build (fold TVars.add_tvars ts) |> TVars.list_set; val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; val instT = TVars.build (fold2 (fn a => fn b => TVars.add (a, TFree b)) tvars tfrees); in (instT, ctxt') end; fun import_inst is_open ts ctxt = let val ren = Name.clean #> (if is_open then I else Name.internal); val (instT, ctxt') = importT_inst ts ctxt; val vars = Vars.build (fold Vars.add_vars ts) |> Vars.list_set |> map (apsnd (Term_Subst.instantiateT instT)); val (ys, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; val inst = Vars.build (fold2 (fn (x, T) => fn y => Vars.add ((x, T), Free (y, T))) vars ys); in ((instT, inst), ctxt'') end; fun importT_terms ts ctxt = let val (instT, ctxt') = importT_inst ts ctxt in (map (Term_Subst.instantiate (instT, Vars.empty)) ts, ctxt') end; fun import_terms is_open ts ctxt = let val (inst, ctxt') = import_inst is_open ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end; fun importT ths ctxt = let val (instT, ctxt') = importT_inst (map Thm.full_prop_of ths) ctxt; val instT' = TVars.map (K (Thm.ctyp_of ctxt')) instT; - val ths' = map (Thm.instantiate (TVars.dest instT', [])) ths; + val ths' = map (Thm.instantiate (instT', Vars.empty)) ths; in ((instT', ths'), ctxt') end; fun import_prf is_open prf ctxt = let val ts = rev (Proofterm.fold_proof_terms_types cons (cons o Logic.mk_type) prf []); val (insts, ctxt') = import_inst is_open ts ctxt; in (Proofterm.instantiate insts prf, ctxt') end; fun import is_open ths ctxt = let val ((instT, inst), ctxt') = import_inst is_open (map Thm.full_prop_of ths) ctxt; val instT' = TVars.map (K (Thm.ctyp_of ctxt')) instT; val inst' = Vars.map (K (Thm.cterm_of ctxt')) inst; - val ths' = map (Thm.instantiate (TVars.dest instT', Vars.dest inst')) ths; + val ths' = map (Thm.instantiate (instT', inst')) ths; in (((instT', inst'), ths'), ctxt') end; fun import_vars ctxt th = let val ((_, [th']), _) = ctxt |> set_body false |> import true [th]; in th' end; (* import/export *) fun gen_trade imp exp f ctxt ths = let val ((_, ths'), ctxt') = imp ths ctxt in exp ctxt' ctxt (f ctxt' ths') end; val tradeT = gen_trade importT exportT; val trade = gen_trade (import true) export; (* focus on outermost parameters: \x y z. B *) val bound_focus = Config.declare_bool ("bound_focus", \<^here>) (K false); val is_bound_focus = Config.apply bound_focus; val set_bound_focus = Config.put bound_focus; val restore_bound_focus = set_bound_focus o is_bound_focus; fun focus_params bindings t ctxt = let val ps = Term.variant_frees t (Term.strip_all_vars t); (*as they are printed :-*) val (xs, Ts) = split_list ps; val (xs', ctxt') = (case bindings of SOME bs => ctxt |> set_body true |> add_fixes_binding bs ||> restore_body ctxt | NONE => if is_bound_focus ctxt then bound_fixes ps ctxt else variant_fixes xs ctxt); val ps' = xs' ~~ Ts; val ctxt'' = ctxt' |> fold (declare_constraints o Free) ps'; in ((xs, ps'), ctxt'') end; fun focus bindings t ctxt = let val ((xs, ps), ctxt') = focus_params bindings t ctxt; val t' = Term.subst_bounds (rev (map Free ps), Term.strip_all_body t); in (((xs ~~ ps), t'), ctxt') end; fun forall_elim_prop t prop = Thm.beta_conversion false (Thm.apply (Thm.dest_arg prop) t) |> Thm.cprop_of |> Thm.dest_arg; fun focus_cterm bindings goal ctxt = let val ((xs, ps), ctxt') = focus_params bindings (Thm.term_of goal) ctxt; val ps' = map (Thm.cterm_of ctxt' o Free) ps; val goal' = fold forall_elim_prop ps' goal; in ((xs ~~ ps', goal'), ctxt') end; fun focus_subgoal bindings i st = let val all_vars = Vars.build (Thm.fold_terms {hyps = false} Vars.add_vars st); in Vars.fold (unbind_term o #1 o #1) all_vars #> Vars.fold (declare_constraints o Var o #1) all_vars #> focus_cterm bindings (Thm.cprem_of st i) end; (** implicit polymorphism **) (* warn_extra_tfrees *) fun warn_extra_tfrees ctxt1 ctxt2 = let fun occs_typ a = Term.exists_subtype (fn TFree (b, _) => a = b | _ => false); fun occs_free a x = (case def_type ctxt1 false (x, ~1) of SOME T => if occs_typ a T then I else cons (a, x) | NONE => cons (a, x)); val occs1 = type_occs_of ctxt1; val occs2 = type_occs_of ctxt2; val extras = Symtab.fold (fn (a, xs) => if Symtab.defined occs1 a then I else fold (occs_free a) xs) occs2 []; val tfrees = map #1 extras |> sort_distinct string_ord; val frees = map #2 extras |> sort_distinct string_ord; in if null extras orelse not (Context_Position.is_visible ctxt2) then () else warning ("Introduced fixed type variable(s): " ^ commas tfrees ^ " in " ^ space_implode " or " (map quote frees)) end; (* polymorphic terms *) fun polymorphic_types ctxt ts = let val ctxt' = fold declare_term ts ctxt; val occs = type_occs_of ctxt; val occs' = type_occs_of ctxt'; val types = Names.build (occs' |> Symtab.fold (fn (a, _) => if Symtab.defined occs a then I else Names.add_set a)); val idx = maxidx_of ctxt' + 1; val Ts' = (fold o fold_types o fold_atyps) (fn T as TFree _ => (case Term_Subst.generalizeT types idx T of TVar v => insert (op =) v | _ => I) | _ => I) ts []; val ts' = map (Term_Subst.generalize (types, Names.empty) idx) ts; in (rev Ts', ts') end; fun polymorphic ctxt ts = snd (polymorphic_types ctxt ts); end; diff --git a/src/Tools/Code/code_preproc.ML b/src/Tools/Code/code_preproc.ML --- a/src/Tools/Code/code_preproc.ML +++ b/src/Tools/Code/code_preproc.ML @@ -1,650 +1,650 @@ (* Title: Tools/Code/code_preproc.ML Author: Florian Haftmann, TU Muenchen Preprocessing code equations into a well-sorted system in a graph with explicit dependencies. *) signature CODE_PREPROC = sig val map_pre: (Proof.context -> Proof.context) -> theory -> theory val map_post: (Proof.context -> Proof.context) -> theory -> theory val add_functrans: string * (Proof.context -> (thm * bool) list -> (thm * bool) list option) -> theory -> theory val del_functrans: string -> theory -> theory val simple_functrans: (Proof.context -> thm list -> thm list option) -> Proof.context -> (thm * bool) list -> (thm * bool) list option val print_codeproc: Proof.context -> unit type code_algebra type code_graph val cert: code_graph -> string -> Code.cert val sortargs: code_graph -> string -> sort list val all: code_graph -> string list val pretty: Proof.context -> code_graph -> Pretty.T val obtain: bool -> { ctxt: Proof.context, consts: string list, terms: term list } -> { algebra: code_algebra, eqngr: code_graph } val dynamic_conv: Proof.context -> (code_algebra -> code_graph -> term -> conv) -> conv val dynamic_value: Proof.context -> ((term -> term) -> 'a -> 'b) -> (code_algebra -> code_graph -> term -> 'a) -> term -> 'b val static_conv: { ctxt: Proof.context, consts: string list } -> ({ algebra: code_algebra, eqngr: code_graph } -> Proof.context -> term -> conv) -> Proof.context -> conv val static_value: { ctxt: Proof.context, lift_postproc: ((term -> term) -> 'a -> 'b), consts: string list } -> ({ algebra: code_algebra, eqngr: code_graph } -> Proof.context -> term -> 'a) -> Proof.context -> term -> 'b val trace_none: Context.generic -> Context.generic val trace_all: Context.generic -> Context.generic val trace_only: string list -> Context.generic -> Context.generic val trace_only_ext: string list -> Context.generic -> Context.generic val timing: bool Config.T val timed: string -> ('a -> Proof.context) -> ('a -> 'b) -> 'a -> 'b val timed_exec: string -> (unit -> 'a) -> Proof.context -> 'a val timed_conv: string -> (Proof.context -> conv) -> Proof.context -> conv val timed_value: string -> (Proof.context -> term -> 'a) -> Proof.context -> term -> 'a end structure Code_Preproc : CODE_PREPROC = struct (** timing **) val timing = Attrib.setup_config_bool \<^binding>\code_timing\ (K false); fun timed msg ctxt_of f x = if Config.get (ctxt_of x) timing then timeap_msg msg f x else f x; fun timed_exec msg f ctxt = if Config.get ctxt timing then timeap_msg msg f () else f (); fun timed' msg f ctxt x = if Config.get ctxt timing then timeap_msg msg (f ctxt) x else f ctxt x; val timed_conv = timed'; val timed_value = timed'; (** preprocessor administration **) (* theory data *) datatype thmproc = Thmproc of { pre: simpset, post: simpset, functrans: (string * (serial * (Proof.context -> (thm * bool) list -> (thm * bool) list option))) list }; fun make_thmproc ((pre, post), functrans) = Thmproc { pre = pre, post = post, functrans = functrans }; fun map_thmproc f (Thmproc { pre, post, functrans }) = make_thmproc (f ((pre, post), functrans)); fun merge_thmproc (Thmproc { pre = pre1, post = post1, functrans = functrans1 }, Thmproc { pre = pre2, post = post2, functrans = functrans2 }) = let val pre = Simplifier.merge_ss (pre1, pre2); val post = Simplifier.merge_ss (post1, post2); val functrans = AList.merge (op =) (eq_fst (op =)) (functrans1, functrans2) handle AList.DUP => error ("Duplicate function transformer"); in make_thmproc ((pre, post), functrans) end; structure Code_Preproc_Data = Theory_Data ( type T = thmproc; val empty = make_thmproc ((Simplifier.empty_ss, Simplifier.empty_ss), []); val extend = I; val merge = merge_thmproc; ); fun the_thmproc thy = case Code_Preproc_Data.get thy of Thmproc x => x; fun delete_force msg key xs = if AList.defined (op =) xs key then AList.delete (op =) key xs else error ("No such " ^ msg ^ ": " ^ quote key); val map_data = Code_Preproc_Data.map o map_thmproc; val map_pre_post = map_data o apfst; fun map_simpset which f thy = map_pre_post (which (simpset_map (Proof_Context.init_global thy) f)) thy; val map_pre = map_simpset apfst; val map_post = map_simpset apsnd; fun process_unfold add_del = map_pre o add_del; fun process_post add_del = map_post o add_del; fun process_abbrev add_del raw_thm thy = let val ctxt = Proof_Context.init_global thy; val thm = Local_Defs.meta_rewrite_rule ctxt raw_thm; val thm_sym = Thm.symmetric thm; in thy |> map_pre_post (fn (pre, post) => (pre |> simpset_map ctxt (add_del thm_sym), post |> simpset_map ctxt (add_del thm))) end; fun add_functrans (name, f) = (map_data o apsnd) (AList.update (op =) (name, (serial (), f))); fun del_functrans name = (map_data o apsnd) (delete_force "function transformer" name); (* algebra of sandwiches: cterm transformations with pending postprocessors *) fun matches_transitive eq1 eq2 = Thm.rhs_of eq1 aconvc Thm.lhs_of eq2; fun trans_comb eq1 eq2 = (*explicit assertions: evaluation conversion stacks are error-prone*) if Thm.is_reflexive eq1 then (\<^assert> (matches_transitive eq1 eq2); eq2) else if Thm.is_reflexive eq2 then (\<^assert> (matches_transitive eq1 eq2); eq1) else Thm.transitive eq1 eq2; fun trans_conv_rule conv eq = trans_comb eq (conv (Thm.rhs_of eq)); structure Sandwich : sig type T = Proof.context -> cterm -> (Proof.context -> thm -> thm) * cterm; val chain: T -> T -> T val lift: (Proof.context -> cterm -> (Proof.context -> cterm -> thm) * thm) -> T val conversion: T -> (Proof.context -> term -> conv) -> Proof.context -> conv; val computation: T -> ((term -> term) -> 'a -> 'b) -> (Proof.context -> term -> 'a) -> Proof.context -> term -> 'b; end = struct type T = Proof.context -> cterm -> (Proof.context -> thm -> thm) * cterm; fun chain sandwich2 sandwich1 ctxt = sandwich1 ctxt ##>> sandwich2 ctxt #>> (fn (f, g) => fn ctxt => f ctxt o g ctxt); fun lift conv_sandwhich ctxt ct = let val (postproc_conv, eq) = conv_sandwhich ctxt ct; fun potentail_trans_comb eq1 eq2 = if matches_transitive eq1 eq2 then trans_comb eq1 eq2 else eq2; (*weakened protocol for plain term evaluation*) in (fn ctxt => trans_conv_rule (postproc_conv ctxt) o potentail_trans_comb eq, Thm.rhs_of eq) end; fun conversion sandwich conv ctxt ct = let val (postproc, ct') = sandwich ctxt ct; val thm = conv ctxt (Thm.term_of ct') ct'; val thm' = postproc ctxt thm; in thm' end; fun computation sandwich lift_postproc eval ctxt t = let val (postproc, ct') = sandwich ctxt (Thm.cterm_of ctxt t); val result = eval ctxt (Thm.term_of ct'); val result' = lift_postproc (Thm.term_of o Thm.rhs_of o postproc ctxt o Thm.reflexive o Thm.cterm_of ctxt) result in result' end; end; (* post- and preprocessing *) fun normalized_tfrees_sandwich ctxt ct = let val t = Thm.term_of ct; val vs_original = fold_term_types (K (fold_atyps (insert (eq_fst op =) o dest_TFree))) t []; val vs_normalized = Name.invent_names Name.context Name.aT (map snd vs_original); val normalize = map_type_tfree (TFree o the o AList.lookup (op =) (vs_original ~~ vs_normalized)); val normalization = map2 (fn (v, sort) => fn (v', _) => (((v', 0), sort), Thm.ctyp_of ctxt (TFree (v, sort)))) vs_original vs_normalized; in if eq_list (eq_fst (op =)) (vs_normalized, vs_original) then (K I, ct) else - (K (Thm.instantiate (normalization, []) o Thm.varifyT_global), + (K (Thm.instantiate (TVars.make normalization, Vars.empty) o Thm.varifyT_global), Thm.cterm_of ctxt (map_types normalize t)) end; fun no_variables_sandwich ctxt ct = let val all_vars = fold_aterms (fn t as Free _ => insert (op aconvc) (Thm.cterm_of ctxt t) | t as Var _ => insert (op aconvc) (Thm.cterm_of ctxt t) | _ => I) (Thm.term_of ct) []; fun apply_beta var thm = Thm.combination thm (Thm.reflexive var) |> Conv.fconv_rule (Conv.arg_conv (Conv.try_conv (Thm.beta_conversion false))) |> Conv.fconv_rule (Conv.arg1_conv (Thm.beta_conversion false)); in if null all_vars then (K I, ct) else (K (fold apply_beta all_vars), fold_rev Thm.lambda all_vars ct) end; fun simplifier_conv_sandwich ctxt = let val thy = Proof_Context.theory_of ctxt; val pre = (#pre o the_thmproc) thy; val post = (#post o the_thmproc) thy; fun pre_conv ctxt' = Simplifier.rewrite (put_simpset pre ctxt') #> trans_conv_rule (Axclass.unoverload_conv ctxt') #> trans_conv_rule (Thm.eta_conversion); fun post_conv ctxt'' = Axclass.overload_conv ctxt'' #> trans_conv_rule (Simplifier.rewrite (put_simpset post ctxt'')); in fn ctxt' => timed_conv "preprocessing term" pre_conv ctxt' #> pair (timed_conv "postprocessing term" post_conv) end; fun simplifier_sandwich ctxt = Sandwich.lift (simplifier_conv_sandwich ctxt); fun value_sandwich ctxt = normalized_tfrees_sandwich |> Sandwich.chain no_variables_sandwich |> Sandwich.chain (simplifier_sandwich ctxt); fun print_codeproc ctxt = let val thy = Proof_Context.theory_of ctxt; val pre = (#pre o the_thmproc) thy; val post = (#post o the_thmproc) thy; val functrans = (map fst o #functrans o the_thmproc) thy; in Pretty.writeln_chunks [ Pretty.block [ Pretty.str "preprocessing simpset:", Pretty.fbrk, Simplifier.pretty_simpset true (put_simpset pre ctxt) ], Pretty.block [ Pretty.str "postprocessing simpset:", Pretty.fbrk, Simplifier.pretty_simpset true (put_simpset post ctxt) ], Pretty.block ( Pretty.str "function transformers:" :: Pretty.fbrk :: (Pretty.fbreaks o map Pretty.str) functrans ) ] end; fun simple_functrans f ctxt eqns = case f ctxt (map fst eqns) of SOME thms' => SOME (map (rpair (forall snd eqns)) thms') | NONE => NONE; (** sort algebra and code equation graph types **) type code_algebra = (sort -> sort) * Sorts.algebra; type code_graph = ((string * sort) list * Code.cert) Graph.T; fun get_node eqngr const = Graph.get_node eqngr const handle Graph.UNDEF _ => error ("No such constant in code equation graph: " ^ quote const); fun cert eqngr = snd o get_node eqngr; fun sortargs eqngr = map snd o fst o get_node eqngr; fun all eqngr = Graph.keys eqngr; fun pretty ctxt eqngr = let val thy = Proof_Context.theory_of ctxt; in AList.make (snd o Graph.get_node eqngr) (Graph.keys eqngr) |> (map o apfst) (Code.string_of_const thy) |> sort (string_ord o apply2 fst) |> (map o apsnd) (Code.pretty_cert thy) |> filter_out (null o snd) |> map (fn (s, ps) => (Pretty.block o Pretty.fbreaks) (Pretty.str s :: ps)) |> Pretty.chunks end; (** simplifier tracing **) structure Trace_Switch = Generic_Data ( type T = string list option; val empty = SOME []; val extend = I; fun merge (NONE, _) = NONE | merge (_, NONE) = NONE | merge (SOME cs1, SOME cs2) = SOME (Library.merge (op =) (cs1, cs2)); ); val trace_none = Trace_Switch.put (SOME []); val trace_all = Trace_Switch.put NONE; fun gen_trace_only prep_const raw_cs context = let val cs = map (prep_const (Context.theory_of context)) raw_cs; in Trace_Switch.put (SOME cs) context end; val trace_only = gen_trace_only (K I); val trace_only_ext = gen_trace_only Code.read_const; fun switch_trace c ctxt = let val d = Trace_Switch.get (Context.Proof ctxt); val switch = case d of NONE => true | SOME cs => member (op =) cs c; val _ = if switch then tracing ("Preprocessing function equations for " ^ Code.string_of_const (Proof_Context.theory_of ctxt) c) else (); in Config.put simp_trace switch ctxt end; (** the Waisenhaus algorithm **) (* auxiliary *) fun is_proper_class thy = can (Axclass.get_info thy); fun complete_proper_sort thy = Sign.complete_sort thy #> filter (is_proper_class thy); fun inst_params thy tyco = map (fn (c, _) => Axclass.param_of_inst thy (c, tyco)) o maps (#params o Axclass.get_info thy); (* data structures *) datatype const = Fun of string | Inst of class * string; fun const_ord (Fun c1, Fun c2) = fast_string_ord (c1, c2) | const_ord (Inst class_tyco1, Inst class_tyco2) = prod_ord fast_string_ord fast_string_ord (class_tyco1, class_tyco2) | const_ord (Fun _, Inst _) = LESS | const_ord (Inst _, Fun _) = GREATER; type var = const * int; structure Vargraph = Graph(type key = var val ord = prod_ord const_ord int_ord); datatype styp = Tyco of string * styp list | Var of var | Free; fun styp_of c_lhs (Type (tyco, tys)) = Tyco (tyco, map (styp_of c_lhs) tys) | styp_of c_lhs (TFree (v, _)) = case c_lhs of SOME (c, lhs) => Var (Fun c, find_index (fn (v', _) => v = v') lhs) | NONE => Free; type vardeps_data = ((string * styp list) list * class list) Vargraph.T * (((string * sort) list * Code.cert) Symtab.table * (class * string) list); val empty_vardeps_data : vardeps_data = (Vargraph.empty, (Symtab.empty, [])); (* retrieving equations and instances from the background context *) fun obtain_eqns ctxt eqngr c = case try (Graph.get_node eqngr) c of SOME (lhs, cert) => ((lhs, []), cert) | NONE => let val thy = Proof_Context.theory_of ctxt; val functrans = (map (fn (_, (_, f)) => f ctxt) o #functrans o the_thmproc) thy; val cert = Code.get_cert (switch_trace c ctxt) functrans c; val (lhs, rhss) = Code.typargs_deps_of_cert thy cert; in ((lhs, rhss), cert) end; fun obtain_instance ctxt arities (inst as (class, tyco)) = case AList.lookup (op =) arities inst of SOME classess => (classess, ([], [])) | NONE => let val thy = Proof_Context.theory_of ctxt; val all_classes = complete_proper_sort thy [class]; val super_classes = remove (op =) class all_classes; val classess = map (complete_proper_sort thy) (Proof_Context.arity_sorts ctxt tyco [class]); val inst_params = inst_params thy tyco all_classes; in (classess, (super_classes, inst_params)) end; (* computing instantiations *) fun add_classes ctxt arities eqngr c_k new_classes vardeps_data = let val (styps, old_classes) = Vargraph.get_node (fst vardeps_data) c_k; val diff_classes = new_classes |> subtract (op =) old_classes; in if null diff_classes then vardeps_data else let val c_ks = Vargraph.immediate_succs (fst vardeps_data) c_k |> insert (op =) c_k; in vardeps_data |> (apfst o Vargraph.map_node c_k o apsnd) (append diff_classes) |> fold (fn styp => fold (ensure_typmatch_inst ctxt arities eqngr styp) new_classes) styps |> fold (fn c_k => add_classes ctxt arities eqngr c_k diff_classes) c_ks end end and add_styp ctxt arities eqngr c_k new_tyco_styps vardeps_data = let val (old_tyco_stypss, classes) = Vargraph.get_node (fst vardeps_data) c_k; in if member (op =) old_tyco_stypss new_tyco_styps then vardeps_data else vardeps_data |> (apfst o Vargraph.map_node c_k o apfst) (cons new_tyco_styps) |> fold (ensure_typmatch_inst ctxt arities eqngr new_tyco_styps) classes end and add_dep ctxt arities eqngr c_k c_k' vardeps_data = let val (_, classes) = Vargraph.get_node (fst vardeps_data) c_k; in vardeps_data |> add_classes ctxt arities eqngr c_k' classes |> apfst (Vargraph.add_edge (c_k, c_k')) end and ensure_typmatch_inst ctxt arities eqngr (tyco, styps) class vardeps_data = if can (Proof_Context.arity_sorts ctxt tyco) [class] then vardeps_data |> ensure_inst ctxt arities eqngr (class, tyco) |> fold_index (fn (k, styp) => ensure_typmatch ctxt arities eqngr styp (Inst (class, tyco), k)) styps else vardeps_data (*permissive!*) and ensure_inst ctxt arities eqngr (inst as (class, tyco)) (vardeps_data as (_, (_, insts))) = if member (op =) insts inst then vardeps_data else let val (classess, (super_classes, inst_params)) = obtain_instance ctxt arities inst; in vardeps_data |> (apsnd o apsnd) (insert (op =) inst) |> fold_index (fn (k, _) => apfst (Vargraph.new_node ((Inst (class, tyco), k), ([] ,[])))) classess |> fold (fn super_class => ensure_inst ctxt arities eqngr (super_class, tyco)) super_classes |> fold (ensure_fun ctxt arities eqngr) inst_params |> fold_index (fn (k, classes) => add_classes ctxt arities eqngr (Inst (class, tyco), k) classes #> fold (fn super_class => add_dep ctxt arities eqngr (Inst (super_class, tyco), k) (Inst (class, tyco), k)) super_classes #> fold (fn inst_param => add_dep ctxt arities eqngr (Fun inst_param, k) (Inst (class, tyco), k) ) inst_params ) classess end and ensure_typmatch ctxt arities eqngr (Tyco tyco_styps) c_k vardeps_data = vardeps_data |> add_styp ctxt arities eqngr c_k tyco_styps | ensure_typmatch ctxt arities eqngr (Var c_k') c_k vardeps_data = vardeps_data |> add_dep ctxt arities eqngr c_k c_k' | ensure_typmatch ctxt arities eqngr Free c_k vardeps_data = vardeps_data and ensure_rhs ctxt arities eqngr (c', styps) vardeps_data = vardeps_data |> ensure_fun ctxt arities eqngr c' |> fold_index (fn (k, styp) => ensure_typmatch ctxt arities eqngr styp (Fun c', k)) styps and ensure_fun ctxt arities eqngr c (vardeps_data as (_, (eqntab, _))) = if Symtab.defined eqntab c then vardeps_data else let val ((lhs, rhss), eqns) = obtain_eqns ctxt eqngr c; val rhss' = (map o apsnd o map) (styp_of (SOME (c, lhs))) rhss; in vardeps_data |> (apsnd o apfst) (Symtab.update_new (c, (lhs, eqns))) |> fold_index (fn (k, _) => apfst (Vargraph.new_node ((Fun c, k), ([] ,[])))) lhs |> fold_index (fn (k, (_, sort)) => add_classes ctxt arities eqngr (Fun c, k) (complete_proper_sort (Proof_Context.theory_of ctxt) sort)) lhs |> fold (ensure_rhs ctxt arities eqngr) rhss' end; (* applying instantiations *) fun dicts_of ctxt (proj_sort, algebra) (T, sort) = let val thy = Proof_Context.theory_of ctxt; fun class_relation _ (x, _) _ = x; fun type_constructor (tyco, _) xs class = inst_params thy tyco (Sorts.complete_sort algebra [class]) @ (maps o maps) fst xs; fun type_variable (TFree (_, sort)) = map (pair []) (proj_sort sort); in flat (Sorts.of_sort_derivation algebra { class_relation = K class_relation, type_constructor = type_constructor, type_variable = type_variable } (T, proj_sort sort) handle Sorts.CLASS_ERROR _ => [] (*permissive!*)) end; fun add_arity ctxt vardeps (class, tyco) = AList.default (op =) ((class, tyco), map_range (fn k => (snd o Vargraph.get_node vardeps) (Inst (class, tyco), k)) (Sign.arity_number (Proof_Context.theory_of ctxt) tyco)); fun add_cert ctxt vardeps (c, (proto_lhs, proto_cert)) (rhss, eqngr) = if can (Graph.get_node eqngr) c then (rhss, eqngr) else let val thy = Proof_Context.theory_of ctxt; val lhs = map_index (fn (k, (v, _)) => (v, snd (Vargraph.get_node vardeps (Fun c, k)))) proto_lhs; val cert = proto_cert |> Code.constrain_cert thy (map (Sign.minimize_sort thy o snd) lhs) |> Code.conclude_cert; val (vs, rhss') = Code.typargs_deps_of_cert thy cert; val eqngr' = Graph.new_node (c, (vs, cert)) eqngr; in (map (pair c) rhss' @ rhss, eqngr') end; fun extend_arities_eqngr raw_ctxt cs ts (arities, (eqngr : code_graph)) = let val thy = Proof_Context.theory_of raw_ctxt; val {pre, ...} = the_thmproc thy; val ctxt = put_simpset pre raw_ctxt; val cs_rhss = (fold o fold_aterms) (fn Const (c_ty as (c, _)) => insert (op =) (c, (map (styp_of NONE) o Sign.const_typargs thy) c_ty) | _ => I) ts []; val (vardeps, (eqntab, insts)) = empty_vardeps_data |> fold (ensure_fun ctxt arities eqngr) cs |> fold (ensure_rhs ctxt arities eqngr) cs_rhss; val arities' = fold (add_arity ctxt vardeps) insts arities; val algebra = Sorts.subalgebra (Context.Theory thy) (is_proper_class thy) (AList.lookup (op =) arities') (Sign.classes_of thy); val (rhss, eqngr') = Symtab.fold (add_cert ctxt vardeps) eqntab ([], eqngr); fun deps_of (c, rhs) = c :: maps (dicts_of ctxt algebra) (rhs ~~ sortargs eqngr' c); val eqngr'' = fold (fn (c, rhs) => fold (curry Graph.add_edge c) (deps_of rhs)) rhss eqngr'; in (algebra, (arities', eqngr'')) end; (** store for preprocessed arities and code equations **) structure Wellsorted = Code_Data ( type T = ((string * class) * sort list) list * code_graph; val empty = ([], Graph.empty); ); (** retrieval and evaluation interfaces **) (* naming conventions * evaluator "eval" is either * conversion "conv" * value computation "comp" * "evaluation" is a lifting of an evaluator *) fun obtain ignore_cache = timed "preprocessing equations" #ctxt (fn { ctxt, consts, terms } => apsnd snd (Wellsorted.change_yield (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt)) (extend_arities_eqngr ctxt consts terms))) #> (fn (algebra, eqngr) => { algebra = algebra, eqngr = eqngr }); fun dynamic_evaluation eval ctxt t = let val consts = fold_aterms (fn Const (c, _) => insert (op =) c | _ => I) t []; val { algebra, eqngr } = obtain false { ctxt = ctxt, consts = consts, terms = [t] }; in eval algebra eqngr t end; fun static_evaluation ctxt consts eval = eval (obtain true { ctxt = ctxt, consts = consts, terms = [] }); fun dynamic_conv ctxt conv = Sandwich.conversion (value_sandwich ctxt) (dynamic_evaluation conv) ctxt; fun dynamic_value ctxt lift_postproc evaluator = Sandwich.computation (value_sandwich ctxt) lift_postproc (dynamic_evaluation evaluator) ctxt; fun static_conv { ctxt, consts } conv = Sandwich.conversion (value_sandwich ctxt) (static_evaluation ctxt consts conv); fun static_value { ctxt, lift_postproc, consts } comp = Sandwich.computation (value_sandwich ctxt) lift_postproc (static_evaluation ctxt consts comp); (** setup **) val _ = Theory.setup ( let fun mk_attribute f = Thm.declaration_attribute (fn thm => Context.mapping (f thm) I); fun add_del_attribute_parser process = Attrib.add_del (mk_attribute (process Simplifier.add_simp)) (mk_attribute (process Simplifier.del_simp)); in Attrib.setup \<^binding>\code_unfold\ (add_del_attribute_parser process_unfold) "preprocessing equations for code generator" #> Attrib.setup \<^binding>\code_post\ (add_del_attribute_parser process_post) "postprocessing equations for code generator" #> Attrib.setup \<^binding>\code_abbrev\ (add_del_attribute_parser process_abbrev) "post- and preprocessing equations for code generator" #> Attrib.setup \<^binding>\code_preproc_trace\ ((Scan.lift (Args.$$$ "off" >> K trace_none) || (Scan.lift (Args.$$$ "only" |-- Args.colon |-- Scan.repeat1 Parse.term)) >> trace_only_ext || Scan.succeed trace_all) >> (Thm.declaration_attribute o K)) "tracing of the code generator preprocessor" end); val _ = Outer_Syntax.command \<^command_keyword>\print_codeproc\ "print code preprocessor setup" (Scan.succeed (Toplevel.keep (print_codeproc o Toplevel.context_of))); end; (*struct*) diff --git a/src/Tools/IsaPlanner/rw_inst.ML b/src/Tools/IsaPlanner/rw_inst.ML --- a/src/Tools/IsaPlanner/rw_inst.ML +++ b/src/Tools/IsaPlanner/rw_inst.ML @@ -1,243 +1,243 @@ (* Title: Tools/IsaPlanner/rw_inst.ML Author: Lucas Dixon, University of Edinburgh Rewriting using a conditional meta-equality theorem which supports schematic variable instantiation. *) signature RW_INST = sig val rw: Proof.context -> ((indexname * (sort * typ)) list * (* type var instantiations *) (indexname * (typ * term)) list) (* schematic var instantiations *) * (string * typ) list (* Fake named bounds + types *) * (string * typ) list (* names of bound + types *) * term -> (* outer term for instantiation *) thm -> (* rule with indexes lifted *) thm -> (* target thm *) thm (* rewritten theorem possibly with additional premises for rule conditions *) end; structure RW_Inst: RW_INST = struct (* Given (string,type) pairs capturing the free vars that need to be allified in the assumption, and a theorem with assumptions possibly containing the free vars, then we give back the assumptions allified as hidden hyps. Given: x th: A vs ==> B vs Results in: "B vs" [!!x. A x] *) fun allify_conditions ctxt Ts th = let fun allify (x, T) t = Logic.all_const T $ Abs (x, T, Term.abstract_over (Free (x, T), t)); val cTs = map (Thm.cterm_of ctxt o Free) Ts; val cterm_asms = map (Thm.cterm_of ctxt o fold_rev allify Ts) (Thm.prems_of th); val allifyied_asm_thms = map (Drule.forall_elim_list cTs o Thm.assume) cterm_asms; in (fold (curry op COMP) allifyied_asm_thms th, cterm_asms) end; (* Given a list of variables that were bound, and a that has been instantiated with free variable placeholders for the bound vars, it creates an abstracted version of the theorem, with local bound vars as lambda-params: Ts: ("x", ty) rule:: C :x ==> P :x = Q :x results in: ("!! x. C x", (%x. p x = %y. p y) [!! x. C x]) note: assumes rule is instantiated *) (* Note, we take abstraction in the order of last abstraction first *) fun mk_abstractedrule ctxt TsFake Ts rule = let (* now we change the names of temporary free vars that represent bound vars with binders outside the redex *) val ns = IsaND.variant_names ctxt (Thm.full_prop_of rule :: Thm.hyps_of rule) (map fst Ts); val (fromnames, tonames, Ts') = fold (fn (((faken, _), (n, ty)), n2) => fn (rnf, rnt, Ts'') => (Thm.cterm_of ctxt (Free(faken,ty)) :: rnf, Thm.cterm_of ctxt (Free(n2,ty)) :: rnt, (n2,ty) :: Ts'')) (TsFake ~~ Ts ~~ ns) ([], [], []); (* rename conflicting free's in the rule to avoid cconflicts with introduced vars from bounds outside in redex *) val rule' = rule |> Drule.forall_intr_list fromnames |> Drule.forall_elim_list tonames; (* make unconditional rule and prems *) val (uncond_rule, cprems) = allify_conditions ctxt (rev Ts') rule'; (* using these names create lambda-abstracted version of the rule *) val abstractions = rev (Ts' ~~ tonames); val abstract_rule = fold (fn ((n, ty), ct) => Thm.abstract_rule n ct) abstractions uncond_rule; in (cprems, abstract_rule) end; (* given names to avoid, and vars that need to be fixed, it gives unique new names to the vars so that they can be fixed as free variables *) (* make fixed unique free variable instantiations for non-ground vars *) (* Create a table of vars to be renamed after instantiation - ie other uninstantiated vars in the hyps of the rule ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) fun mk_renamings ctxt tgt rule_inst = let val rule_conds = Thm.prems_of rule_inst; val (_, cond_vs) = fold (fn t => fn (tyvs, vs) => (union (op =) (Misc_Legacy.term_tvars t) tyvs, union (op =) (map Term.dest_Var (Misc_Legacy.term_vars t)) vs)) rule_conds ([], []); val termvars = map Term.dest_Var (Misc_Legacy.term_vars tgt); val vars_to_fix = union (op =) termvars cond_vs; val ys = IsaND.variant_names ctxt (tgt :: rule_conds) (map (fst o fst) vars_to_fix); in map2 (fn (xi, T) => fn y => ((xi, T), Free (y, T))) vars_to_fix ys end; (* make a new fresh typefree instantiation for the given tvar *) fun new_tfree (tv as (ix,sort)) (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix,(sort,TFree(v,sort)))::pairs, v::used) end; (* make instantiations to fix type variables that are not already instantiated (in ignore_ixs) from the list of terms. *) fun mk_fixtvar_tyinsts ignore_insts ts = let val ignore_ixs = map fst ignore_insts; val (tvars, tfrees) = fold_rev (fn t => fn (varixs, tfrees) => (Misc_Legacy.add_term_tvars (t,varixs), Misc_Legacy.add_term_tfrees (t,tfrees))) ts ([], []); val unfixed_tvars = filter (fn (ix,s) => not (member (op =) ignore_ixs ix)) tvars; val (fixtyinsts, _) = fold_rev new_tfree unfixed_tvars ([], map fst tfrees) in (fixtyinsts, tfrees) end; (* cross-instantiate the instantiations - ie for each instantiation replace all occurrences in other instantiations - no loops are possible and thus only one-parsing of the instantiations is necessary. *) fun cross_inst insts = let fun instL (ix, (ty,t)) = map (fn (ix2,(ty2,t2)) => (ix2, (ty2,Term.subst_vars ([], [(ix, t)]) t2))); fun cross_instL ([], l) = rev l | cross_instL ((ix, t) :: insts, l) = cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); in cross_instL (insts, []) end; (* as above but for types -- I don't know if this is needed, will we ever incur mixed up types? *) fun cross_inst_typs insts = let fun instL (ix, (srt,ty)) = map (fn (ix2,(srt2,ty2)) => (ix2, (srt2,Term.typ_subst_TVars [(ix, ty)] ty2))); fun cross_instL ([], l) = rev l | cross_instL ((ix, t) :: insts, l) = cross_instL (instL (ix, t) insts, (ix, t) :: (instL (ix, t) l)); in cross_instL (insts, []) end; (* assume that rule and target_thm have distinct var names. THINK: efficient version with tables for vars for: target vars, introduced vars, and rule vars, for quicker instantiation? The outerterm defines which part of the target_thm was modified. Note: we take Ts in the upterm order, ie last abstraction first., and with an outeterm where the abstracted subterm has the arguments in the revered order, ie first abstraction first. FakeTs has abstractions using the fake name - ie the name distinct from all other abstractions. *) fun rw ctxt ((nonfixed_typinsts, unprepinsts), FakeTs, Ts, outerterm) rule target_thm = let (* fix all non-instantiated tvars *) val (fixtyinsts, othertfrees) = (* FIXME proper context!? *) mk_fixtvar_tyinsts nonfixed_typinsts [Thm.prop_of rule, Thm.prop_of target_thm]; val typinsts = cross_inst_typs (nonfixed_typinsts @ fixtyinsts); (* certified instantiations for types *) - val ctyp_insts = map (fn (ix, (s, ty)) => ((ix, s), Thm.ctyp_of ctxt ty)) typinsts; + val ctyp_insts = TVars.make (map (fn (ix, (s, ty)) => ((ix, s), Thm.ctyp_of ctxt ty)) typinsts); (* type instantiated versions *) - val tgt_th_tyinst = Thm.instantiate (ctyp_insts,[]) target_thm; - val rule_tyinst = Thm.instantiate (ctyp_insts,[]) rule; + val tgt_th_tyinst = Thm.instantiate (ctyp_insts,Vars.empty) target_thm; + val rule_tyinst = Thm.instantiate (ctyp_insts,Vars.empty) rule; val term_typ_inst = map (fn (ix,(_,ty)) => (ix,ty)) typinsts; (* type instanitated outer term *) val outerterm_tyinst = Term.subst_TVars term_typ_inst outerterm; val FakeTs_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) FakeTs; val Ts_tyinst = map (apsnd (Term.typ_subst_TVars term_typ_inst)) Ts; (* type-instantiate the var instantiations *) val insts_tyinst = fold_rev (fn (ix, (ty, t)) => fn insts_tyinst => (ix, (Term.typ_subst_TVars term_typ_inst ty, Term.subst_TVars term_typ_inst t)) :: insts_tyinst) unprepinsts []; (* cross-instantiate *) val insts_tyinst_inst = cross_inst insts_tyinst; (* create certms of instantiations *) val cinsts_tyinst = - map (fn (ix, (ty, t)) => ((ix, ty), Thm.cterm_of ctxt t)) insts_tyinst_inst; + Vars.make (map (fn (ix, (ty, t)) => ((ix, ty), Thm.cterm_of ctxt t)) insts_tyinst_inst); (* The instantiated rule *) - val rule_inst = rule_tyinst |> Thm.instantiate ([], cinsts_tyinst); + val rule_inst = rule_tyinst |> Thm.instantiate (TVars.empty, cinsts_tyinst); (* Create a table of vars to be renamed after instantiation - ie other uninstantiated vars in the hyps the *instantiated* rule ie ?z in C ?z ?x ==> A ?x ?y = B ?x ?y *) val renamings = mk_renamings ctxt (Thm.prop_of tgt_th_tyinst) rule_inst; val cterm_renamings = map (fn (x, y) => apply2 (Thm.cterm_of ctxt) (Var x, y)) renamings; (* Create the specific version of the rule for this target application *) val outerterm_inst = outerterm_tyinst |> Term.subst_Vars (map (fn (ix, (ty, t)) => (ix, t)) insts_tyinst_inst) |> Term.subst_Vars (map (fn ((ix, ty), t) => (ix, t)) renamings); val couter_inst = Thm.reflexive (Thm.cterm_of ctxt outerterm_inst); val (cprems, abstract_rule_inst) = rule_inst - |> Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) cterm_renamings) + |> Thm.instantiate (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) cterm_renamings)) |> mk_abstractedrule ctxt FakeTs_tyinst Ts_tyinst; val specific_tgt_rule = Conv.fconv_rule Drule.beta_eta_conversion (Thm.combination couter_inst abstract_rule_inst); (* create an instantiated version of the target thm *) val tgt_th_inst = tgt_th_tyinst - |> Thm.instantiate ([], cinsts_tyinst) - |> Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) cterm_renamings); + |> Thm.instantiate (TVars.empty, cinsts_tyinst) + |> Thm.instantiate (TVars.empty, Vars.make (map (apfst (dest_Var o Thm.term_of)) cterm_renamings)); val (vars,frees_of_fixed_vars) = Library.split_list cterm_renamings; in Conv.fconv_rule Drule.beta_eta_conversion tgt_th_inst |> Thm.equal_elim specific_tgt_rule |> Drule.implies_intr_list cprems |> Drule.forall_intr_list frees_of_fixed_vars |> Drule.forall_elim_list vars |> Thm.varifyT_global' (TFrees.make_set othertfrees) |-> K Drule.zero_var_indexes end; end; diff --git a/src/Tools/coherent.ML b/src/Tools/coherent.ML --- a/src/Tools/coherent.ML +++ b/src/Tools/coherent.ML @@ -1,234 +1,235 @@ (* Title: Tools/coherent.ML Author: Stefan Berghofer, TU Muenchen Author: Marc Bezem, Institutt for Informatikk, Universitetet i Bergen Prover for coherent logic, see e.g. Marc Bezem and Thierry Coquand, Automating Coherent Logic, LPAR 2005 for a description of the algorithm. *) signature COHERENT_DATA = sig val atomize_elimL: thm val atomize_exL: thm val atomize_conjL: thm val atomize_disjL: thm val operator_names: string list end; signature COHERENT = sig val trace: bool Config.T val coherent_tac: Proof.context -> thm list -> int -> tactic end; functor Coherent(Data: COHERENT_DATA) : COHERENT = struct (** misc tools **) val (trace, trace_setup) = Attrib.config_bool \<^binding>\coherent_trace\ (K false); fun cond_trace ctxt msg = if Config.get ctxt trace then tracing (msg ()) else (); datatype cl_prf = ClPrf of thm * (Type.tyenv * Envir.tenv) * ((indexname * typ) * term) list * int list * (term list * cl_prf) list; val is_atomic = not o exists_Const (member (op =) Data.operator_names o #1); fun rulify_elim_conv ctxt ct = if is_atomic (Logic.strip_imp_concl (Thm.term_of ct)) then Conv.all_conv ct else Conv.concl_conv (length (Logic.strip_imp_prems (Thm.term_of ct))) (Conv.rewr_conv (Thm.symmetric Data.atomize_elimL) then_conv Raw_Simplifier.rewrite ctxt true (map Thm.symmetric [Data.atomize_exL, Data.atomize_conjL, Data.atomize_disjL])) ct fun rulify_elim ctxt th = Simplifier.norm_hhf ctxt (Conv.fconv_rule (rulify_elim_conv ctxt) th); (* Decompose elimination rule of the form A1 ==> ... ==> Am ==> (!!xs1. Bs1 ==> P) ==> ... ==> (!!xsn. Bsn ==> P) ==> P *) fun dest_elim prop = let val prems = Logic.strip_imp_prems prop; val concl = Logic.strip_imp_concl prop; val (prems1, prems2) = chop_suffix (fn t => Logic.strip_assums_concl t = concl) prems; in (prems1, if null prems2 then [([], [concl])] else map (fn t => (map snd (Logic.strip_params t), Logic.strip_assums_hyp t)) prems2) end; fun mk_rule ctxt th = let val th' = rulify_elim ctxt th; val (prems, cases) = dest_elim (Thm.prop_of th') in (th', prems, cases) end; fun mk_dom ts = fold (fn t => Typtab.map_default (fastype_of t, []) (fn us => us @ [t])) ts Typtab.empty; val empty_env = (Vartab.empty, Vartab.empty); (* Find matcher that makes conjunction valid in given state *) fun valid_conj _ _ env [] = Seq.single (env, []) | valid_conj ctxt facts env (t :: ts) = Seq.maps (fn (u, x) => Seq.map (apsnd (cons x)) (valid_conj ctxt facts (Pattern.match (Proof_Context.theory_of ctxt) (t, u) env) ts handle Pattern.MATCH => Seq.empty)) (Seq.of_list (sort (int_ord o apply2 snd) (Net.unify_term facts t))); (* Instantiate variables that only occur free in conlusion *) fun inst_extra_vars ctxt dom cs = let val vs = fold Term.add_vars (maps snd cs) []; fun insts [] inst = Seq.single inst | insts ((ixn, T) :: vs') inst = Seq.maps (fn t => insts vs' (((ixn, T), t) :: inst)) (Seq.of_list (case Typtab.lookup dom T of NONE => error ("Unknown domain: " ^ Syntax.string_of_typ ctxt T ^ "\nfor term(s) " ^ commas (maps (map (Syntax.string_of_term ctxt) o snd) cs)) | SOME ts => ts)) in Seq.map (fn inst => (inst, map (apsnd (map (subst_Vars (map (apfst fst) inst)))) cs)) (insts vs []) end; (* Check whether disjunction is valid in given state *) fun is_valid_disj _ _ [] = false | is_valid_disj ctxt facts ((Ts, ts) :: ds) = let val vs = map_index (fn (i, T) => Var (("x", i), T)) Ts in (case Seq.pull (valid_conj ctxt facts empty_env (map (fn t => subst_bounds (rev vs, t)) ts)) of SOME _ => true | NONE => is_valid_disj ctxt facts ds) end; fun string_of_facts ctxt s facts = Pretty.string_of (Pretty.big_list s (map (Syntax.pretty_term ctxt) (map fst (sort (int_ord o apply2 snd) (Net.content facts))))); fun valid ctxt rules goal dom facts nfacts nparams = let val seq = Seq.of_list rules |> Seq.maps (fn (th, ps, cs) => valid_conj ctxt facts empty_env ps |> Seq.maps (fn (env as (tye, _), is) => let val cs' = map (fn (Ts, ts) => (map (Envir.subst_type tye) Ts, map (Envir.subst_term env) ts)) cs in inst_extra_vars ctxt dom cs' |> Seq.map_filter (fn (inst, cs'') => if is_valid_disj ctxt facts cs'' then NONE else SOME (th, env, inst, is, cs'')) end)); in (case Seq.pull seq of NONE => (if Context_Position.is_visible ctxt then warning (string_of_facts ctxt "Countermodel found:" facts) else (); NONE) | SOME ((th, env, inst, is, cs), _) => if cs = [([], [goal])] then SOME (ClPrf (th, env, inst, is, [])) else (case valid_cases ctxt rules goal dom facts nfacts nparams cs of NONE => NONE | SOME prfs => SOME (ClPrf (th, env, inst, is, prfs)))) end and valid_cases _ _ _ _ _ _ _ [] = SOME [] | valid_cases ctxt rules goal dom facts nfacts nparams ((Ts, ts) :: ds) = let val _ = cond_trace ctxt (fn () => Pretty.string_of (Pretty.block (Pretty.str "case" :: Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_term ctxt) ts)))); val ps = map_index (fn (i, T) => ("par" ^ string_of_int (nparams + i), T)) Ts; val (params, ctxt') = fold_map Variable.next_bound ps ctxt; val ts' = map_index (fn (i, t) => (subst_bounds (rev params, t), nfacts + i)) ts; val dom' = fold (fn (T, p) => Typtab.map_default (T, []) (fn ps => ps @ [p])) (Ts ~~ params) dom; val facts' = fold (fn (t, i) => Net.insert_term op = (t, (t, i))) ts' facts; in (case valid ctxt' rules goal dom' facts' (nfacts + length ts) (nparams + length Ts) of NONE => NONE | SOME prf => (case valid_cases ctxt rules goal dom facts nfacts nparams ds of NONE => NONE | SOME prfs => SOME ((params, prf) :: prfs))) end; (** proof replaying **) fun thm_of_cl_prf ctxt goal asms (ClPrf (th, (tye, env), insts, is, prfs)) = let val _ = cond_trace ctxt (fn () => Pretty.string_of (Pretty.big_list "asms:" (map (Thm.pretty_thm ctxt) asms))); + val instT = map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T)) (Vartab.dest tye); + val inst = + map (fn (ixn, (T, t)) => + ((ixn, Envir.subst_type tye T), Thm.cterm_of ctxt t)) (Vartab.dest env) @ + map (fn (ixnT, t) => (ixnT, Thm.cterm_of ctxt t)) insts; val th' = Drule.implies_elim_list - (Thm.instantiate - (map (fn (ixn, (S, T)) => ((ixn, S), Thm.ctyp_of ctxt T)) (Vartab.dest tye), - map (fn (ixn, (T, t)) => - ((ixn, Envir.subst_type tye T), Thm.cterm_of ctxt t)) (Vartab.dest env) @ - map (fn (ixnT, t) => (ixnT, Thm.cterm_of ctxt t)) insts) th) + (Thm.instantiate (TVars.make instT, Vars.make inst) th) (map (nth asms) is); val (_, cases) = dest_elim (Thm.prop_of th'); in (case (cases, prfs) of ([([], [_])], []) => th' | ([([], [_])], [([], prf)]) => thm_of_cl_prf ctxt goal (asms @ [th']) prf | _ => Drule.implies_elim_list (Thm.instantiate (Thm.match (Drule.strip_imp_concl (Thm.cprop_of th'), goal)) th') (map (thm_of_case_prf ctxt goal asms) (prfs ~~ cases))) end and thm_of_case_prf ctxt goal asms ((params, prf), (_, asms')) = let val cparams = map (Thm.cterm_of ctxt) params; val asms'' = map (Thm.cterm_of ctxt o curry subst_bounds (rev params)) asms'; val (prems'', ctxt') = fold_map Thm.assume_hyps asms'' ctxt; in Drule.forall_intr_list cparams (Drule.implies_intr_list asms'' (thm_of_cl_prf ctxt' goal (asms @ prems'') prf)) end; (** external interface **) fun coherent_tac ctxt rules = SUBPROOF (fn {prems, concl, params, context = ctxt', ...} => resolve_tac ctxt' [rulify_elim_conv ctxt' concl RS Drule.equal_elim_rule2] 1 THEN SUBPROOF (fn {prems = prems', concl, context = ctxt'', ...} => let val xs = map (Thm.term_of o #2) params @ map (fn (_, s) => Free (s, the (Variable.default_type ctxt'' s))) (rev (Variable.dest_fixes ctxt'')) (* FIXME !? *) in (case valid ctxt'' (map (mk_rule ctxt'') (prems' @ prems @ rules)) (Thm.term_of concl) (mk_dom xs) Net.empty 0 0 of NONE => no_tac | SOME prf => resolve_tac ctxt'' [thm_of_cl_prf ctxt'' concl [] prf] 1) end) ctxt' 1) ctxt; val _ = Theory.setup (trace_setup #> Method.setup \<^binding>\coherent\ (Attrib.thms >> (fn rules => fn ctxt => METHOD (fn facts => HEADGOAL (coherent_tac ctxt (facts @ rules))))) "prove coherent formula"); end; diff --git a/src/Tools/induct.ML b/src/Tools/induct.ML --- a/src/Tools/induct.ML +++ b/src/Tools/induct.ML @@ -1,972 +1,975 @@ (* Title: Tools/induct.ML Author: Markus Wenzel, TU Muenchen Proof by cases, induction, and coinduction. *) signature INDUCT_ARGS = sig val cases_default: thm val atomize: thm list val rulify: thm list val rulify_fallback: thm list val equal_def: thm val dest_def: term -> (term * term) option val trivial_tac: Proof.context -> int -> tactic end; signature INDUCT = sig (*rule declarations*) val vars_of: term -> term list val dest_rules: Proof.context -> {type_cases: (string * thm) list, pred_cases: (string * thm) list, type_induct: (string * thm) list, pred_induct: (string * thm) list, type_coinduct: (string * thm) list, pred_coinduct: (string * thm) list} val print_rules: Proof.context -> unit val lookup_casesT: Proof.context -> string -> thm option val lookup_casesP: Proof.context -> string -> thm option val lookup_inductT: Proof.context -> string -> thm option val lookup_inductP: Proof.context -> string -> thm option val lookup_coinductT: Proof.context -> string -> thm option val lookup_coinductP: Proof.context -> string -> thm option val find_casesT: Proof.context -> typ -> thm list val find_casesP: Proof.context -> term -> thm list val find_inductT: Proof.context -> typ -> thm list val find_inductP: Proof.context -> term -> thm list val find_coinductT: Proof.context -> typ -> thm list val find_coinductP: Proof.context -> term -> thm list val cases_type: string -> attribute val cases_pred: string -> attribute val cases_del: attribute val induct_type: string -> attribute val induct_pred: string -> attribute val induct_del: attribute val coinduct_type: string -> attribute val coinduct_pred: string -> attribute val coinduct_del: attribute val map_simpset: (Proof.context -> Proof.context) -> Context.generic -> Context.generic val induct_simp_add: attribute val induct_simp_del: attribute val no_simpN: string val casesN: string val inductN: string val coinductN: string val typeN: string val predN: string val setN: string (*proof methods*) val arbitrary_tac: Proof.context -> int -> (string * typ) list -> int -> tactic val add_defs: (binding option * (term * bool)) option list -> Proof.context -> (term option list * thm list) * Proof.context val atomize_term: Proof.context -> term -> term val atomize_cterm: Proof.context -> conv val atomize_tac: Proof.context -> int -> tactic val inner_atomize_tac: Proof.context -> int -> tactic val rulified_term: Proof.context -> thm -> term val rulify_tac: Proof.context -> int -> tactic val simplified_rule: Proof.context -> thm -> thm val simplify_tac: Proof.context -> int -> tactic val trivial_tac: Proof.context -> int -> tactic val rotate_tac: int -> int -> int -> tactic val internalize: Proof.context -> int -> thm -> thm val guess_instance: Proof.context -> thm -> int -> thm -> thm Seq.seq val cases_context_tactic: bool -> term option list list -> thm option -> thm list -> int -> context_tactic val cases_tac: Proof.context -> bool -> term option list list -> thm option -> thm list -> int -> tactic val get_inductT: Proof.context -> term option list list -> thm list list val gen_induct_context_tactic: ((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) -> bool -> (binding option * (term * bool)) option list list -> (string * typ) list list -> term option list -> thm list option -> thm list -> int -> context_tactic val gen_induct_tac: Proof.context -> ((Rule_Cases.info * int) * thm -> (Rule_Cases.info * int) * thm) -> bool -> (binding option * (term * bool)) option list list -> (string * typ) list list -> term option list -> thm list option -> thm list -> int -> tactic val induct_context_tactic: bool -> (binding option * (term * bool)) option list list -> (string * typ) list list -> term option list -> thm list option -> thm list -> int -> context_tactic val induct_tac: Proof.context -> bool -> (binding option * (term * bool)) option list list -> (string * typ) list list -> term option list -> thm list option -> thm list -> int -> tactic val coinduct_context_tactic: term option list -> term option list -> thm option -> thm list -> int -> context_tactic val coinduct_tac: Proof.context -> term option list -> term option list -> thm option -> thm list -> int -> tactic val gen_induct_setup: binding -> (bool -> (binding option * (term * bool)) option list list -> (string * typ) list list -> term option list -> thm list option -> thm list -> int -> context_tactic) -> local_theory -> local_theory end; functor Induct(Induct_Args: INDUCT_ARGS): INDUCT = struct (** variables -- ordered left-to-right, preferring right **) fun vars_of tm = rev (distinct (op =) (Term.fold_aterms (fn t as Var _ => cons t | _ => I) tm [])); local val mk_var = Net.encode_type o #2 o Term.dest_Var; fun concl_var which thm = mk_var (which (vars_of (Thm.concl_of thm))) handle List.Empty => raise THM ("No variables in conclusion of rule", 0, [thm]); in fun left_var_prem thm = mk_var (hd (vars_of (hd (Thm.prems_of thm)))) handle List.Empty => raise THM ("No variables in major premise of rule", 0, [thm]); val left_var_concl = concl_var hd; val right_var_concl = concl_var List.last; end; (** constraint simplification **) (* rearrange parameters and premises to allow application of one-point-rules *) fun swap_params_conv ctxt i j cv = let fun conv1 0 ctxt = Conv.forall_conv (cv o snd) ctxt | conv1 k ctxt = Conv.rewr_conv @{thm swap_params} then_conv Conv.forall_conv (conv1 (k - 1) o snd) ctxt; fun conv2 0 ctxt = conv1 j ctxt | conv2 k ctxt = Conv.forall_conv (conv2 (k - 1) o snd) ctxt; in conv2 i ctxt end; fun swap_prems_conv 0 = Conv.all_conv | swap_prems_conv i = Conv.implies_concl_conv (swap_prems_conv (i - 1)) then_conv Conv.rewr_conv Drule.swap_prems_eq; fun find_eq ctxt t = let val l = length (Logic.strip_params t); val Hs = Logic.strip_assums_hyp t; fun find (i, t) = (case Induct_Args.dest_def (Object_Logic.drop_judgment ctxt t) of SOME (Bound j, _) => SOME (i, j) | SOME (_, Bound j) => SOME (i, j) | _ => NONE); in (case get_first find (map_index I Hs) of NONE => NONE | SOME (0, 0) => NONE | SOME (i, j) => SOME (i, l - j - 1, j)) end; fun mk_swap_rrule ctxt ct = (case find_eq ctxt (Thm.term_of ct) of NONE => NONE | SOME (i, k, j) => SOME (swap_params_conv ctxt k j (K (swap_prems_conv i)) ct)); val rearrange_eqs_simproc = Simplifier.make_simproc \<^context> "rearrange_eqs" {lhss = [\<^term>\Pure.all (t :: 'a::{} \ prop)\], proc = fn _ => fn ctxt => fn ct => mk_swap_rrule ctxt ct}; (* rotate k premises to the left by j, skipping over first j premises *) fun rotate_conv 0 _ 0 = Conv.all_conv | rotate_conv 0 j k = swap_prems_conv j then_conv rotate_conv 1 j (k - 1) | rotate_conv i j k = Conv.implies_concl_conv (rotate_conv (i - 1) j k); fun rotate_tac _ 0 = K all_tac | rotate_tac j k = SUBGOAL (fn (goal, i) => CONVERSION (rotate_conv j (length (Logic.strip_assums_hyp goal) - j - k) k) i); (* rulify operators around definition *) fun rulify_defs_conv ctxt ct = if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) andalso not (is_some (Induct_Args.dest_def (Object_Logic.drop_judgment ctxt (Thm.term_of ct)))) then (Conv.forall_conv (rulify_defs_conv o snd) ctxt else_conv Conv.implies_conv (Conv.try_conv (rulify_defs_conv ctxt)) (Conv.try_conv (rulify_defs_conv ctxt)) else_conv Conv.first_conv (map Conv.rewr_conv Induct_Args.rulify) then_conv Conv.try_conv (rulify_defs_conv ctxt)) ct else Conv.no_conv ct; (** induct data **) (* rules *) type rules = (string * thm) Item_Net.T; fun init_rules index : rules = Item_Net.init (fn ((s1, th1), (s2, th2)) => s1 = s2 andalso Thm.eq_thm_prop (th1, th2)) (single o index); fun filter_rules (rs: rules) th = filter (fn (_, th') => Thm.eq_thm_prop (th, th')) (Item_Net.content rs); fun pretty_rules ctxt kind rs = let val thms = map snd (Item_Net.content rs) in Pretty.big_list kind (map (Thm.pretty_thm_item ctxt) thms) end; (* context data *) structure Data = Generic_Data ( type T = (rules * rules) * (rules * rules) * (rules * rules) * simpset; val empty = ((init_rules (left_var_prem o #2), init_rules (Thm.major_prem_of o #2)), (init_rules (right_var_concl o #2), init_rules (Thm.major_prem_of o #2)), (init_rules (left_var_concl o #2), init_rules (Thm.concl_of o #2)), simpset_of (empty_simpset \<^context> addsimprocs [rearrange_eqs_simproc] addsimps [Drule.norm_hhf_eq])); val extend = I; fun merge (((casesT1, casesP1), (inductT1, inductP1), (coinductT1, coinductP1), simpset1), ((casesT2, casesP2), (inductT2, inductP2), (coinductT2, coinductP2), simpset2)) = ((Item_Net.merge (casesT1, casesT2), Item_Net.merge (casesP1, casesP2)), (Item_Net.merge (inductT1, inductT2), Item_Net.merge (inductP1, inductP2)), (Item_Net.merge (coinductT1, coinductT2), Item_Net.merge (coinductP1, coinductP2)), merge_ss (simpset1, simpset2)); ); val get_local = Data.get o Context.Proof; fun dest_rules ctxt = let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in {type_cases = Item_Net.content casesT, pred_cases = Item_Net.content casesP, type_induct = Item_Net.content inductT, pred_induct = Item_Net.content inductP, type_coinduct = Item_Net.content coinductT, pred_coinduct = Item_Net.content coinductP} end; fun print_rules ctxt = let val ((casesT, casesP), (inductT, inductP), (coinductT, coinductP), _) = get_local ctxt in [pretty_rules ctxt "coinduct type:" coinductT, pretty_rules ctxt "coinduct pred:" coinductP, pretty_rules ctxt "induct type:" inductT, pretty_rules ctxt "induct pred:" inductP, pretty_rules ctxt "cases type:" casesT, pretty_rules ctxt "cases pred:" casesP] |> Pretty.writeln_chunks end; val _ = Outer_Syntax.command \<^command_keyword>\print_induct_rules\ "print induction and cases rules" (Scan.succeed (Toplevel.keep (print_rules o Toplevel.context_of))); (* access rules *) local fun lookup_rule which ctxt = AList.lookup (op =) (Item_Net.content (which (get_local ctxt))) #> Option.map (Thm.transfer' ctxt); fun find_rules which how ctxt x = Item_Net.retrieve (which (get_local ctxt)) (how x) |> map (Thm.transfer' ctxt o snd); in val lookup_casesT = lookup_rule (#1 o #1); val lookup_casesP = lookup_rule (#2 o #1); val lookup_inductT = lookup_rule (#1 o #2); val lookup_inductP = lookup_rule (#2 o #2); val lookup_coinductT = lookup_rule (#1 o #3); val lookup_coinductP = lookup_rule (#2 o #3); val find_casesT = find_rules (#1 o #1) Net.encode_type; val find_casesP = find_rules (#2 o #1) I; val find_inductT = find_rules (#1 o #2) Net.encode_type; val find_inductP = find_rules (#2 o #2) I; val find_coinductT = find_rules (#1 o #3) Net.encode_type; val find_coinductP = find_rules (#2 o #3) I; end; (** attributes **) local fun mk_att f g name = Thm.mixed_attribute (fn (context, thm) => let val thm' = g thm; val context' = if Thm.is_free_dummy thm then context else Data.map (f (name, Thm.trim_context thm')) context; in (context', thm') end); fun del_att which = Thm.declaration_attribute (fn th => Data.map (which (apply2 (fn rs => fold Item_Net.remove (filter_rules rs th) rs)))); fun add_casesT rule x = @{apply 4(1)} (apfst (Item_Net.update rule)) x; fun add_casesP rule x = @{apply 4(1)} (apsnd (Item_Net.update rule)) x; fun add_inductT rule x = @{apply 4(2)} (apfst (Item_Net.update rule)) x; fun add_inductP rule x = @{apply 4(2)} (apsnd (Item_Net.update rule)) x; fun add_coinductT rule x = @{apply 4(3)} (apfst (Item_Net.update rule)) x; fun add_coinductP rule x = @{apply 4(3)} (apsnd (Item_Net.update rule)) x; val consumes0 = Rule_Cases.default_consumes 0; val consumes1 = Rule_Cases.default_consumes 1; in val cases_type = mk_att add_casesT consumes0; val cases_pred = mk_att add_casesP consumes1; val cases_del = del_att @{apply 4(1)}; val induct_type = mk_att add_inductT consumes0; val induct_pred = mk_att add_inductP consumes1; val induct_del = del_att @{apply 4(2)}; val coinduct_type = mk_att add_coinductT consumes0; val coinduct_pred = mk_att add_coinductP consumes1; val coinduct_del = del_att @{apply 4(3)}; fun map_simpset f context = context |> (Data.map o @{apply 4(4)} o Simplifier.simpset_map (Context.proof_of context)) f; fun induct_simp f = Thm.declaration_attribute (fn thm => map_simpset (fn ctxt => f (ctxt, [thm]))); val induct_simp_add = induct_simp (op addsimps); val induct_simp_del = induct_simp (op delsimps); end; (** attribute syntax **) val no_simpN = "no_simp"; val casesN = "cases"; val inductN = "induct"; val coinductN = "coinduct"; val typeN = "type"; val predN = "pred"; val setN = "set"; local fun spec k arg = Scan.lift (Args.$$$ k -- Args.colon) |-- arg || Scan.lift (Args.$$$ k) >> K ""; fun attrib add_type add_pred del = spec typeN (Args.type_name {proper = false, strict = false}) >> add_type || spec predN (Args.const {proper = false, strict = false}) >> add_pred || spec setN (Args.const {proper = false, strict = false}) >> add_pred || Scan.lift Args.del >> K del; in val _ = Theory.local_setup (Attrib.local_setup \<^binding>\cases\ (attrib cases_type cases_pred cases_del) "declaration of cases rule" #> Attrib.local_setup \<^binding>\induct\ (attrib induct_type induct_pred induct_del) "declaration of induction rule" #> Attrib.local_setup \<^binding>\coinduct\ (attrib coinduct_type coinduct_pred coinduct_del) "declaration of coinduction rule" #> Attrib.local_setup \<^binding>\induct_simp\ (Attrib.add_del induct_simp_add induct_simp_del) "declaration of rules for simplifying induction or cases rules"); end; (** method utils **) (* alignment *) fun align_left msg xs ys = let val m = length xs and n = length ys in if m < n then error msg else (take n xs ~~ ys) end; fun align_right msg xs ys = let val m = length xs and n = length ys in if m < n then error msg else (drop (m - n) xs ~~ ys) end; (* prep_inst *) fun prep_inst ctxt align tune (tm, ts) = let fun prep_var (Var (x, xT), SOME t) = let val ct = Thm.cterm_of ctxt (tune t); val tT = Thm.typ_of_cterm ct; in if Type.could_unify (tT, xT) then SOME (x, ct) else error (Pretty.string_of (Pretty.block [Pretty.str "Ill-typed instantiation:", Pretty.fbrk, Syntax.pretty_term ctxt (Thm.term_of ct), Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt tT])) end | prep_var (_, NONE) = NONE; val xs = vars_of tm; in align "Rule has fewer variables than instantiations given" xs ts |> map_filter prep_var end; (* trace_rules *) fun trace_rules _ kind [] = error ("Unable to figure out " ^ kind ^ " rule") | trace_rules ctxt _ rules = Method.trace ctxt rules; (* mark equality constraints in cases rule *) val equal_def' = Thm.symmetric Induct_Args.equal_def; fun mark_constraints n ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (Conv.params_conv ~1 (K (Conv.prems_conv n (Raw_Simplifier.rewrite ctxt false [equal_def']))) ctxt)); fun unmark_constraints ctxt = Conv.fconv_rule (Raw_Simplifier.rewrite ctxt true [Induct_Args.equal_def]); (* simplify *) fun simplify_conv' ctxt = Simplifier.full_rewrite (put_simpset (#4 (get_local ctxt)) ctxt); fun simplify_conv ctxt ct = if exists_subterm (is_some o Induct_Args.dest_def) (Thm.term_of ct) then (Conv.try_conv (rulify_defs_conv ctxt) then_conv simplify_conv' ctxt) ct else Conv.all_conv ct; fun gen_simplified_rule cv ctxt = Conv.fconv_rule (Conv.prems_conv ~1 (cv ctxt)); val simplified_rule' = gen_simplified_rule simplify_conv'; val simplified_rule = gen_simplified_rule simplify_conv; fun simplify_tac ctxt = CONVERSION (simplify_conv ctxt); val trivial_tac = Induct_Args.trivial_tac; (** cases method **) (* rule selection scheme: cases - default case split `A t` cases ... - predicate/set cases cases t - type cases ... cases ... r - explicit rule *) local fun get_casesT ctxt ((SOME t :: _) :: _) = find_casesT ctxt (Term.fastype_of t) | get_casesT _ _ = []; fun get_casesP ctxt (fact :: _) = find_casesP ctxt (Thm.concl_of fact) | get_casesP _ _ = []; in fun cases_context_tactic simp insts opt_rule facts i : context_tactic = fn (ctxt, st) => let fun inst_rule r = (if null insts then r else (align_left "Rule has fewer premises than arguments given" (Thm.prems_of r) insts |> maps (prep_inst ctxt align_left I) |> infer_instantiate ctxt) r) |> simp ? mark_constraints (Rule_Cases.get_constraints r) ctxt |> pair (Rule_Cases.get r); val (opt_rule', facts') = (case (opt_rule, facts) of (NONE, th :: ths) => if is_some (Object_Logic.elim_concl ctxt th) then (SOME th, ths) else (opt_rule, facts) | _ => (opt_rule, facts)); val ruleq = (case opt_rule' of SOME r => Seq.single (inst_rule r) | NONE => (get_casesP ctxt facts' @ get_casesT ctxt insts @ [Induct_Args.cases_default]) |> tap (trace_rules ctxt casesN) |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); in ruleq |> Seq.maps (Rule_Cases.consume ctxt [] facts') |> Seq.maps (fn ((cases, (_, facts'')), rule) => let val rule' = rule |> simp ? (simplified_rule' ctxt #> unmark_constraints ctxt); in CONTEXT_CASES (Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params rule')) cases) ((Method.insert_tac ctxt facts'' THEN' resolve_tac ctxt [rule'] THEN_ALL_NEW (if simp then TRY o trivial_tac ctxt else K all_tac)) i) (ctxt, st) end) end; fun cases_tac ctxt x1 x2 x3 x4 x5 = NO_CONTEXT_TACTIC ctxt (cases_context_tactic x1 x2 x3 x4 x5); end; (** induct method **) val conjunction_congs = @{thms Pure.all_conjunction imp_conjunction}; (* atomize *) fun atomize_term ctxt = Raw_Simplifier.rewrite_term (Proof_Context.theory_of ctxt) Induct_Args.atomize [] #> Object_Logic.drop_judgment ctxt; fun atomize_cterm ctxt = Raw_Simplifier.rewrite ctxt true Induct_Args.atomize; fun atomize_tac ctxt = rewrite_goal_tac ctxt Induct_Args.atomize; fun inner_atomize_tac ctxt = rewrite_goal_tac ctxt (map Thm.symmetric conjunction_congs) THEN' atomize_tac ctxt; (* rulify *) fun rulify_term thy = Raw_Simplifier.rewrite_term thy (Induct_Args.rulify @ conjunction_congs) [] #> Raw_Simplifier.rewrite_term thy Induct_Args.rulify_fallback []; fun rulified_term ctxt thm = let val rulify = rulify_term (Proof_Context.theory_of ctxt); val (As, B) = Logic.strip_horn (Thm.prop_of thm); in Logic.list_implies (map rulify As, rulify B) end; fun rulify_tac ctxt = rewrite_goal_tac ctxt (Induct_Args.rulify @ conjunction_congs) THEN' rewrite_goal_tac ctxt Induct_Args.rulify_fallback THEN' Goal.conjunction_tac THEN_ALL_NEW (rewrite_goal_tac ctxt @{thms Pure.conjunction_imp} THEN' Goal.norm_hhf_tac ctxt); (* prepare rule *) fun rule_instance ctxt inst rule = infer_instantiate ctxt (prep_inst ctxt align_left I (Thm.prop_of rule, inst)) rule; fun internalize ctxt k th = th |> Thm.permute_prems 0 k |> Conv.fconv_rule (Conv.concl_conv (Thm.nprems_of th - k) (atomize_cterm ctxt)); (* guess rule instantiation -- cannot handle pending goal parameters *) local -fun dest_env ctxt env = +fun insts_env ctxt env = let val pairs = Vartab.dest (Envir.term_env env); val types = Vartab.dest (Envir.type_env env); val ts = map (Thm.cterm_of ctxt o Envir.norm_term env o #2 o #2) pairs; val xs = map #1 pairs ~~ map Thm.typ_of_cterm ts; - in (map (fn (ai, (S, T)) => ((ai, S), Thm.ctyp_of ctxt T)) types, xs ~~ ts) end; + val instT = + TVars.build (types |> fold (fn (ai, (S, T)) => TVars.add ((ai, S), Thm.ctyp_of ctxt T))); + val inst = Vars.build (fold2 (fn x => fn t => Vars.add (x, t)) xs ts); + in (instT, inst) end; in fun guess_instance ctxt rule i st = let val maxidx = Thm.maxidx_of st; val goal = Thm.term_of (Thm.cprem_of st i); (*exception Subscript*) val params = rev (Term.rename_wrt_term goal (Logic.strip_params goal)); in if not (null params) then (warning ("Cannot determine rule instantiation due to pending parameter(s): " ^ commas_quote (map (Syntax.string_of_term ctxt o Syntax_Trans.mark_bound_abs) params)); Seq.single rule) else let val rule' = Thm.incr_indexes (maxidx + 1) rule; val concl = Logic.strip_assums_concl goal; in Unify.smash_unifiers (Context.Proof ctxt) [(Thm.concl_of rule', concl)] (Envir.empty (Thm.maxidx_of rule')) - |> Seq.map (fn env => Drule.instantiate_normalize (dest_env ctxt env) rule') + |> Seq.map (fn env => Drule.instantiate_normalize (insts_env ctxt env) rule') end end handle General.Subscript => Seq.empty; end; (* special renaming of rule parameters *) fun special_rename_params ctxt [[SOME (Free (z, Type (T, _)))]] [thm] = let val x = Name.clean (Variable.revert_fixed ctxt z); fun index _ [] = [] | index i (y :: ys) = if x = y then x ^ string_of_int i :: index (i + 1) ys else y :: index i ys; fun rename_params [] = [] | rename_params ((y, Type (U, _)) :: ys) = (if U = T then x else y) :: rename_params ys | rename_params ((y, _) :: ys) = y :: rename_params ys; fun rename_asm A = let val xs = rename_params (Logic.strip_params A); val xs' = (case filter (fn x' => x' = x) xs of [] => xs | [_] => xs | _ => index 1 xs); in Logic.list_rename_params xs' A end; fun rename_prop prop = let val (As, C) = Logic.strip_horn prop in Logic.list_implies (map rename_asm As, C) end; val thm' = Thm.renamed_prop (rename_prop (Thm.prop_of thm)) thm; in [Rule_Cases.save thm thm'] end | special_rename_params _ _ ths = ths; (* arbitrary_tac *) local fun goal_prefix k ((c as Const (\<^const_name>\Pure.all\, _)) $ Abs (a, T, B)) = c $ Abs (a, T, goal_prefix k B) | goal_prefix 0 _ = Term.dummy_prop | goal_prefix k ((c as Const (\<^const_name>\Pure.imp\, _)) $ A $ B) = c $ A $ goal_prefix (k - 1) B | goal_prefix _ _ = Term.dummy_prop; fun goal_params k (Const (\<^const_name>\Pure.all\, _) $ Abs (_, _, B)) = goal_params k B + 1 | goal_params 0 _ = 0 | goal_params k (Const (\<^const_name>\Pure.imp\, _) $ _ $ B) = goal_params (k - 1) B | goal_params _ _ = 0; fun meta_spec_tac ctxt n (x, T) = SUBGOAL (fn (goal, i) => let val v = Free (x, T); fun spec_rule prfx (xs, body) = @{thm Pure.meta_spec} |> Thm.rename_params_rule ([Name.clean (Variable.revert_fixed ctxt x)], 1) |> Thm.lift_rule (Thm.cterm_of ctxt prfx) |> `(Thm.prop_of #> Logic.strip_assums_concl) |-> (fn pred $ arg => infer_instantiate ctxt [(#1 (dest_Var (head_of pred)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, body))), (#1 (dest_Var (head_of arg)), Thm.cterm_of ctxt (Logic.rlist_abs (xs, v)))]); fun goal_concl k xs (Const (\<^const_name>\Pure.all\, _) $ Abs (a, T, B)) = goal_concl k ((a, T) :: xs) B | goal_concl 0 xs B = if not (Term.exists_subterm (fn t => t aconv v) B) then NONE else SOME (xs, absfree (x, T) (Term.incr_boundvars 1 B)) | goal_concl k xs (Const (\<^const_name>\Pure.imp\, _) $ _ $ B) = goal_concl (k - 1) xs B | goal_concl _ _ _ = NONE; in (case goal_concl n [] goal of SOME concl => (compose_tac ctxt (false, spec_rule (goal_prefix n goal) concl, 1) THEN' resolve_tac ctxt [asm_rl]) i | NONE => all_tac) end); fun miniscope_tac p = CONVERSION o Conv.params_conv p (fn ctxt => Raw_Simplifier.rewrite ctxt true [Thm.symmetric Drule.norm_hhf_eq]); in fun arbitrary_tac _ _ [] = K all_tac | arbitrary_tac ctxt n xs = SUBGOAL (fn (goal, i) => (EVERY' (map (meta_spec_tac ctxt n) xs) THEN' (miniscope_tac (goal_params n goal) ctxt)) i); end; (* add_defs *) fun add_defs def_insts = let fun add (SOME (_, (t, true))) ctxt = ((SOME t, []), ctxt) | add (SOME (SOME x, (t, _))) ctxt = let val ([(lhs, (_, th))], ctxt') = Local_Defs.define [((x, NoSyn), ((Thm.def_binding x, []), t))] ctxt in ((SOME lhs, [th]), ctxt') end | add (SOME (NONE, (t as Free _, _))) ctxt = ((SOME t, []), ctxt) | add (SOME (NONE, (t, _))) ctxt = let val (s, _) = Name.variant "x" (Variable.names_of ctxt); val x = Binding.name s; val ([(lhs, (_, th))], ctxt') = ctxt |> Local_Defs.define [((x, NoSyn), ((Thm.def_binding x, []), t))]; in ((SOME lhs, [th]), ctxt') end | add NONE ctxt = ((NONE, []), ctxt); in fold_map add def_insts #> apfst (split_list #> apsnd flat) end; (* induct_tac *) (* rule selection scheme: `A x` induct ... - predicate/set induction induct x - type induction ... induct ... r - explicit rule *) fun get_inductT ctxt insts = fold_rev (map_product cons) (insts |> map ((fn [] => NONE | ts => List.last ts) #> (fn NONE => TVar (("'a", 0), []) | SOME t => Term.fastype_of t) #> find_inductT ctxt)) [[]] |> filter_out (forall Rule_Cases.is_inner_rule); fun get_inductP ctxt (fact :: _) = map single (find_inductP ctxt (Thm.concl_of fact)) | get_inductP _ _ = []; fun gen_induct_context_tactic mod_cases simp def_insts arbitrary taking opt_rule facts = CONTEXT_SUBGOAL (fn (_, i) => fn (ctxt, st) => let val ((insts, defs), defs_ctxt) = fold_map add_defs def_insts ctxt |>> split_list; val atomized_defs = map (map (Conv.fconv_rule (atomize_cterm defs_ctxt))) defs; fun inst_rule (concls, r) = (if null insts then `Rule_Cases.get r else (align_left "Rule has fewer conclusions than arguments given" (map Logic.strip_imp_concl (Logic.dest_conjunctions (Thm.concl_of r))) insts |> maps (prep_inst ctxt align_right (atomize_term ctxt)) |> infer_instantiate ctxt) r |> pair (Rule_Cases.get r)) |> mod_cases |> (fn ((cases, consumes), th) => (((cases, concls), consumes), th)); val ruleq = (case opt_rule of SOME rs => Seq.single (inst_rule (Rule_Cases.strict_mutual_rule ctxt rs)) | NONE => (get_inductP ctxt facts @ map (special_rename_params defs_ctxt insts) (get_inductT ctxt insts)) |> map_filter (Rule_Cases.mutual_rule ctxt) |> tap (trace_rules ctxt inductN o map #2) |> Seq.of_list |> Seq.maps (Seq.try inst_rule)); fun rule_cases ctxt rule cases = let val rule' = Rule_Cases.internalize_params rule; val rule'' = rule' |> simp ? simplified_rule ctxt; val nonames = map (fn ((cn, _), cls) => ((cn, []), cls)); val cases' = if Thm.eq_thm_prop (rule', rule'') then cases else nonames cases; in Rule_Cases.make_nested ctxt (Thm.prop_of rule'') (rulified_term ctxt rule'') cases' end; fun context_tac _ _ = ruleq |> Seq.maps (Rule_Cases.consume defs_ctxt (flat defs) facts) |> Seq.maps (fn (((cases, concls), (more_consumes, more_facts)), rule) => (PRECISE_CONJUNCTS (length concls) (ALLGOALS (fn j => (CONJUNCTS (ALLGOALS let val adefs = nth_list atomized_defs (j - 1); val frees = fold (Term.add_frees o Thm.prop_of) adefs []; val xs = nth_list arbitrary (j - 1); val k = nth concls (j - 1) + more_consumes; in Method.insert_tac defs_ctxt (more_facts @ adefs) THEN' (if simp then rotate_tac k (length adefs) THEN' arbitrary_tac defs_ctxt k (List.partition (member op = frees) xs |> op @) else arbitrary_tac defs_ctxt k xs) end) THEN' inner_atomize_tac defs_ctxt) j)) THEN' atomize_tac defs_ctxt) i st |> Seq.maps (fn st' => guess_instance ctxt (internalize ctxt more_consumes rule) i st' |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) |> Seq.maps (fn rule' => CONTEXT_CASES (rule_cases ctxt rule' cases) (resolve_tac ctxt [rule'] i THEN PRIMITIVE (singleton (Proof_Context.export defs_ctxt ctxt))) (ctxt, st')))); in (context_tac CONTEXT_THEN_ALL_NEW ((if simp then simplify_tac ctxt THEN' (TRY o trivial_tac ctxt) else K all_tac) THEN_ALL_NEW rulify_tac ctxt)) i (ctxt, st) end) val induct_context_tactic = gen_induct_context_tactic I; fun gen_induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 x8 = NO_CONTEXT_TACTIC ctxt (gen_induct_context_tactic x1 x2 x3 x4 x5 x6 x7 x8); fun induct_tac ctxt x1 x2 x3 x4 x5 x6 x7 = NO_CONTEXT_TACTIC ctxt (induct_context_tactic x1 x2 x3 x4 x5 x6 x7); (** coinduct method **) (* rule selection scheme: goal "A x" coinduct ... - predicate/set coinduction coinduct x - type coinduction coinduct ... r - explicit rule *) local fun get_coinductT ctxt (SOME t :: _) = find_coinductT ctxt (Term.fastype_of t) | get_coinductT _ _ = []; fun get_coinductP ctxt goal = find_coinductP ctxt (Logic.strip_assums_concl goal); fun main_prop_of th = if Rule_Cases.get_consumes th > 0 then Thm.major_prem_of th else Thm.concl_of th; in fun coinduct_context_tactic inst taking opt_rule facts = CONTEXT_SUBGOAL (fn (goal, i) => fn (ctxt, st) => let fun inst_rule r = if null inst then `Rule_Cases.get r else infer_instantiate ctxt (prep_inst ctxt align_right I (main_prop_of r, inst)) r |> pair (Rule_Cases.get r); in (case opt_rule of SOME r => Seq.single (inst_rule r) | NONE => (get_coinductP ctxt goal @ get_coinductT ctxt inst) |> tap (trace_rules ctxt coinductN) |> Seq.of_list |> Seq.maps (Seq.try inst_rule)) |> Seq.maps (Rule_Cases.consume ctxt [] facts) |> Seq.maps (fn ((cases, (_, more_facts)), rule) => guess_instance ctxt rule i st |> Seq.map (rule_instance ctxt (burrow_options (Variable.polymorphic ctxt) taking)) |> Seq.maps (fn rule' => CONTEXT_CASES (Rule_Cases.make_common ctxt (Thm.prop_of (Rule_Cases.internalize_params rule')) cases) (Method.insert_tac ctxt more_facts i THEN resolve_tac ctxt [rule'] i) (ctxt, st))) end); fun coinduct_tac ctxt x1 x2 x3 x4 x5 = NO_CONTEXT_TACTIC ctxt (coinduct_context_tactic x1 x2 x3 x4 x5); end; (** concrete syntax **) val arbitraryN = "arbitrary"; val takingN = "taking"; val ruleN = "rule"; local fun single_rule [rule] = rule | single_rule _ = error "Single rule expected"; fun named_rule k arg get = Scan.lift (Args.$$$ k -- Args.colon) |-- Scan.repeat arg :|-- (fn names => Scan.peek (fn context => Scan.succeed (names |> map (fn name => (case get (Context.proof_of context) name of SOME x => x | NONE => error ("No rule for " ^ k ^ " " ^ quote name)))))); fun rule get_type get_pred = named_rule typeN (Args.type_name {proper = false, strict = false}) get_type || named_rule predN (Args.const {proper = false, strict = false}) get_pred || named_rule setN (Args.const {proper = false, strict = false}) get_pred || Scan.lift (Args.$$$ ruleN -- Args.colon) |-- Attrib.thms; val cases_rule = rule lookup_casesT lookup_casesP >> single_rule; val induct_rule = rule lookup_inductT lookup_inductP; val coinduct_rule = rule lookup_coinductT lookup_coinductP >> single_rule; val inst = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> SOME; val inst' = Scan.lift (Args.$$$ "_") >> K NONE || Args.term >> (SOME o rpair false) || Scan.lift (Args.$$$ "(") |-- (Args.term >> (SOME o rpair true)) --| Scan.lift (Args.$$$ ")"); val def_inst = ((Scan.lift (Args.binding --| (Args.$$$ "\" || Args.$$$ "==")) >> SOME) -- (Args.term >> rpair false)) >> SOME || inst' >> Option.map (pair NONE); val free = Args.context -- Args.term >> (fn (_, Free v) => v | (ctxt, t) => error ("Bad free variable: " ^ Syntax.string_of_term ctxt t)); fun unless_more_args scan = Scan.unless (Scan.lift ((Args.$$$ arbitraryN || Args.$$$ takingN || Args.$$$ typeN || Args.$$$ predN || Args.$$$ setN || Args.$$$ ruleN) -- Args.colon)) scan; val arbitrary = Scan.optional (Scan.lift (Args.$$$ arbitraryN -- Args.colon) |-- Parse.and_list1' (Scan.repeat (unless_more_args free))) []; val taking = Scan.optional (Scan.lift (Args.$$$ takingN -- Args.colon) |-- Scan.repeat1 (unless_more_args inst)) []; in fun gen_induct_setup binding tac = Method.local_setup binding (Scan.lift (Args.mode no_simpN) -- (Parse.and_list' (Scan.repeat (unless_more_args def_inst)) -- (arbitrary -- taking -- Scan.option induct_rule)) >> (fn (no_simp, (insts, ((arbitrary, taking), opt_rule))) => fn _ => fn facts => Seq.DETERM (tac (not no_simp) insts arbitrary taking opt_rule facts 1))) "induction on types or predicates/sets"; val _ = Theory.local_setup (Method.local_setup \<^binding>\cases\ (Scan.lift (Args.mode no_simpN) -- (Parse.and_list' (Scan.repeat (unless_more_args inst)) -- Scan.option cases_rule) >> (fn (no_simp, (insts, opt_rule)) => fn _ => CONTEXT_METHOD (fn facts => Seq.DETERM (cases_context_tactic (not no_simp) insts opt_rule facts 1)))) "case analysis on types or predicates/sets" #> gen_induct_setup \<^binding>\induct\ induct_context_tactic #> Method.local_setup \<^binding>\coinduct\ (Scan.repeat (unless_more_args inst) -- taking -- Scan.option coinduct_rule >> (fn ((insts, taking), opt_rule) => fn _ => fn facts => Seq.DETERM (coinduct_context_tactic insts taking opt_rule facts 1))) "coinduction on types or predicates/sets"); end; end; diff --git a/src/Tools/misc_legacy.ML b/src/Tools/misc_legacy.ML --- a/src/Tools/misc_legacy.ML +++ b/src/Tools/misc_legacy.ML @@ -1,253 +1,256 @@ (* Title: Tools/misc_legacy.ML Misc legacy stuff -- to be phased out eventually. *) signature MISC_LEGACY = sig val add_term_names: term * string list -> string list val add_term_tvars: term * (indexname * sort) list -> (indexname * sort) list val add_term_tfrees: term * (string * sort) list -> (string * sort) list val typ_tvars: typ -> (indexname * sort) list val term_tfrees: term -> (string * sort) list val term_tvars: term -> (indexname * sort) list val add_term_vars: term * term list -> term list val term_vars: term -> term list val add_term_frees: term * term list -> term list val term_frees: term -> term list val mk_defpair: term * term -> string * term val get_def: theory -> xstring -> thm val METAHYPS: Proof.context -> (thm list -> tactic) -> int -> tactic val freeze_thaw_robust: Proof.context -> thm -> thm * (int -> thm -> thm) end; structure Misc_Legacy: MISC_LEGACY = struct (*iterate a function over all types in a term*) fun it_term_types f = let fun iter(Const(_,T), a) = f(T,a) | iter(Free(_,T), a) = f(T,a) | iter(Var(_,T), a) = f(T,a) | iter(Abs(_,T,t), a) = iter(t,f(T,a)) | iter(f$u, a) = iter(f, iter(u, a)) | iter(Bound _, a) = a in iter end (*Accumulates the names in the term, suppressing duplicates. Includes Frees and Consts. For choosing unambiguous bound var names.*) fun add_term_names (Const(a,_), bs) = insert (op =) (Long_Name.base_name a) bs | add_term_names (Free(a,_), bs) = insert (op =) a bs | add_term_names (f$u, bs) = add_term_names (f, add_term_names(u, bs)) | add_term_names (Abs(_,_,t), bs) = add_term_names(t,bs) | add_term_names (_, bs) = bs; (*Accumulates the TVars in a type, suppressing duplicates.*) fun add_typ_tvars(Type(_,Ts),vs) = List.foldr add_typ_tvars vs Ts | add_typ_tvars(TFree(_),vs) = vs | add_typ_tvars(TVar(v),vs) = insert (op =) v vs; (*Accumulates the TFrees in a type, suppressing duplicates.*) fun add_typ_tfree_names(Type(_,Ts),fs) = List.foldr add_typ_tfree_names fs Ts | add_typ_tfree_names(TFree(f,_),fs) = insert (op =) f fs | add_typ_tfree_names(TVar(_),fs) = fs; fun add_typ_tfrees(Type(_,Ts),fs) = List.foldr add_typ_tfrees fs Ts | add_typ_tfrees(TFree(f),fs) = insert (op =) f fs | add_typ_tfrees(TVar(_),fs) = fs; (*Accumulates the TVars in a term, suppressing duplicates.*) val add_term_tvars = it_term_types add_typ_tvars; (*Accumulates the TFrees in a term, suppressing duplicates.*) val add_term_tfrees = it_term_types add_typ_tfrees; val add_term_tfree_names = it_term_types add_typ_tfree_names; (*Non-list versions*) fun typ_tfrees T = add_typ_tfrees(T,[]); fun typ_tvars T = add_typ_tvars(T,[]); fun term_tfrees t = add_term_tfrees(t,[]); fun term_tvars t = add_term_tvars(t,[]); (*Accumulates the Vars in the term, suppressing duplicates.*) fun add_term_vars (t, vars: term list) = case t of Var _ => Ord_List.insert Term_Ord.term_ord t vars | Abs (_,_,body) => add_term_vars(body,vars) | f$t => add_term_vars (f, add_term_vars(t, vars)) | _ => vars; fun term_vars t = add_term_vars(t,[]); (*Accumulates the Frees in the term, suppressing duplicates.*) fun add_term_frees (t, frees: term list) = case t of Free _ => Ord_List.insert Term_Ord.term_ord t frees | Abs (_,_,body) => add_term_frees(body,frees) | f$t => add_term_frees (f, add_term_frees(t, frees)) | _ => frees; fun term_frees t = add_term_frees(t,[]); fun mk_defpair (lhs, rhs) = (case Term.head_of lhs of Const (name, _) => (Thm.def_name (Long_Name.base_name name), Logic.mk_equals (lhs, rhs)) | _ => raise TERM ("Malformed definition: head of lhs not a constant", [lhs, rhs])); fun get_def thy = Thm.axiom thy o Name_Space.intern (Theory.axiom_space thy) o Thm.def_name; (**** METAHYPS -- tactical for using hypotheses as meta-level assumptions METAHYPS (fn prems => tac prems) i converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new proof state A==>A, supplying A1,...,An as meta-level assumptions (in "prems"). The parameters x1,...,xm become free variables. If the resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An) then it is lifted back into the original context, yielding k subgoals. Replaces unknowns in the context by Frees having the prefix METAHYP_ New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm. DOES NOT HANDLE TYPE UNKNOWNS. NOTE: This version does not observe the proof context, and thus cannot work reliably. See also Subgoal.SUBPROOF and Subgoal.FOCUS for properly localized variants of the same idea. ****) local (*Strips assumptions in goal yielding ( [x1,...,xm], [H1,...,Hn], B ) H1,...,Hn are the hypotheses; x1...xm are variants of the parameters. Main difference from strip_assums concerns parameters: it replaces the bound variables by free variables. *) fun strip_context_aux (params, Hs, Const (\<^const_name>\Pure.imp\, _) $ H $ B) = strip_context_aux (params, H :: Hs, B) | strip_context_aux (params, Hs, Const (\<^const_name>\Pure.all\,_) $ Abs (a, T, t)) = let val (b, u) = Syntax_Trans.variant_abs (a, T, t) in strip_context_aux ((b, T) :: params, Hs, u) end | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B); fun strip_context A = strip_context_aux ([], [], A); (*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. Instantiates distinct free variables by terms of same type.*) fun free_instantiate ctpairs = forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs); fun free_of s ((a, i), T) = Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T) fun mk_inst v = (Var v, free_of "METAHYP1_" v) fun metahyps_split_prem prem = let (*find all vars in the hyps -- should find tvars also!*) val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) [] val insts = map mk_inst hyps_vars (*replace the hyps_vars by Frees*) val prem' = subst_atomic insts prem val (params,hyps,concl) = strip_context prem' in (insts,params,hyps,concl) end; fun metahyps_aux_tac ctxt tacf (prem,gno) state = let val (insts,params,hyps,concl) = metahyps_split_prem prem val maxidx = Thm.maxidx_of state val chyps = map (Thm.cterm_of ctxt) hyps val hypths = map Thm.assume chyps val subprems = map (Thm.forall_elim_vars 0) hypths val fparams = map Free params val cparams = map (Thm.cterm_of ctxt) fparams fun swap_ctpair (t, u) = apply2 (Thm.cterm_of ctxt) (u, t) (*Subgoal variables: make Free; lift type over params*) fun mk_subgoal_inst concl_vars (v, T) = if member (op =) concl_vars (v, T) then ((v, T), true, free_of "METAHYP2_" (v, T)) else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T)) (*Instantiate subgoal vars by Free applied to params*) fun mk_inst (v, in_concl, u) = if in_concl then (v, Thm.cterm_of ctxt u) else (v, Thm.cterm_of ctxt (list_comb (u, fparams))) (*Restore Vars with higher type and index*) fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) = if in_concl then apply2 (Thm.cterm_of ctxt) (u, Var ((a, i), T)) else apply2 (Thm.cterm_of ctxt) (u, Var ((a, i + maxidx), U)) (*Embed B in the original context of params and hyps*) fun embed B = fold_rev Logic.all fparams (Logic.list_implies (hyps, B)) (*Strip the context using elimination rules*) fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths (*A form of lifting that discharges assumptions.*) fun relift st = let val prop = Thm.prop_of st val subgoal_vars = (*Vars introduced in the subgoals*) fold Term.add_vars (Logic.strip_imp_prems prop) [] and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) [] val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars - val st' = Thm.instantiate ([], map mk_inst subgoal_insts) st + val st' = + Thm.instantiate (TVars.empty, Vars.build (fold (Vars.add o mk_inst) subgoal_insts)) st val emBs = map (Thm.cterm_of ctxt o embed) (Thm.prems_of st') val Cth = implies_elim_list st' (map (elim o Thm.assume) emBs) in (*restore the unknowns to the hypotheses*) free_instantiate (map swap_ctpair insts @ map mk_subgoal_swap_ctpair subgoal_insts) (*discharge assumptions from state in same order*) (implies_intr_list emBs (forall_intr_list cparams (implies_intr_list chyps Cth))) end (*function to replace the current subgoal*) fun next st = Thm.bicompose (SOME ctxt) {flatten = true, match = false, incremented = false} (false, relift st, Thm.nprems_of st) gno state in Seq.maps next (tacf subprems (Thm.trivial (Thm.cterm_of ctxt concl))) end; in fun METAHYPS ctxt tacf n thm = SUBGOAL (metahyps_aux_tac ctxt tacf) n thm handle THM ("assume: variables", _, _) => Seq.empty end; (* generating identifiers -- often fresh *) local (*Maps 0-61 to A-Z, a-z, 0-9; exclude _ or ' to avoid clash with internal/unusual indentifiers*) fun gensym_char i = if i<26 then chr (ord "A" + i) else if i<52 then chr (ord "a" + i - 26) else chr (ord "0" + i - 52); val char_vec = Vector.tabulate (62, gensym_char); fun newid n = implode (map (fn i => Vector.sub (char_vec, i)) (radixpand (62, n))); val gensym_seed = Synchronized.var "gensym_seed" (0: int); in fun gensym pre = Synchronized.change_result gensym_seed (fn i => (pre ^ newid i, i + 1)); end; (*Convert all Vars in a theorem to Frees. Also return a function for reversing that operation. DOES NOT WORK FOR TYPE VARIABLES.*) fun freeze_thaw_robust ctxt th = let val fth = Thm.legacy_freezeT th in case Thm.fold_terms {hyps = false} Term.add_vars fth [] of [] => (fth, fn _ => fn x => x) (*No vars: nothing to do!*) | vars => let fun newName (ix,_) = (ix, gensym (string_of_indexname ix)) val alist = map newName vars fun mk_inst (v,T) = apply2 (Thm.cterm_of ctxt) (Var (v, T), Free (the (AList.lookup (op =) alist v), T)) val insts = map mk_inst vars fun thaw i th' = (*i is non-negative increment for Var indexes*) th' |> forall_intr_list (map #2 insts) |> forall_elim_list (map (Thm.incr_indexes_cterm i o #1) insts) - in (Thm.instantiate ([], map (apfst (dest_Var o Thm.term_of)) insts) fth, thaw) end + in + (Thm.instantiate (TVars.empty, Vars.build (fold (Vars.add o apfst (dest_Var o Thm.term_of)) insts)) fth, thaw) + end end; end; diff --git a/src/Tools/nbe.ML b/src/Tools/nbe.ML --- a/src/Tools/nbe.ML +++ b/src/Tools/nbe.ML @@ -1,624 +1,624 @@ (* Title: Tools/nbe.ML Authors: Klaus Aehlig, LMU Muenchen; Tobias Nipkow, Florian Haftmann, TU Muenchen Normalization by evaluation, based on generic code generator. *) signature NBE = sig val dynamic_conv: Proof.context -> conv val dynamic_value: Proof.context -> term -> term val static_conv: { ctxt: Proof.context, consts: string list } -> Proof.context -> conv val static_value: { ctxt: Proof.context, consts: string list } -> Proof.context -> term -> term datatype Univ = Const of int * Univ list (*named (uninterpreted) constants*) | DFree of string * int (*free (uninterpreted) dictionary parameters*) | BVar of int * Univ list | Abs of (int * (Univ list -> Univ)) * Univ list val apps: Univ -> Univ list -> Univ (*explicit applications*) val abss: int -> (Univ list -> Univ) -> Univ (*abstractions as closures*) val same: Univ * Univ -> bool val put_result: (unit -> Univ list -> Univ list) -> Proof.context -> Proof.context val trace: bool Config.T val add_const_alias: thm -> theory -> theory end; structure Nbe: NBE = struct (* generic non-sense *) val trace = Attrib.setup_config_bool \<^binding>\nbe_trace\ (K false); fun traced ctxt f x = if Config.get ctxt trace then (tracing (f x); x) else x; (** certificates and oracle for "trivial type classes" **) structure Triv_Class_Data = Theory_Data ( type T = (class * thm) list; val empty = []; val extend = I; fun merge data : T = AList.merge (op =) (K true) data; ); fun add_const_alias thm thy = let val (ofclass, eqn) = case try Logic.dest_equals (Thm.prop_of thm) of SOME ofclass_eq => ofclass_eq | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); val (T, class) = case try Logic.dest_of_class ofclass of SOME T_class => T_class | _ => error ("Bad certificate: " ^ Thm.string_of_thm_global thy thm); val tvar = case try Term.dest_TVar T of SOME (tvar as (_, sort)) => if null (filter (can (Axclass.get_info thy)) sort) then tvar else error ("Bad sort: " ^ Thm.string_of_thm_global thy thm) | _ => error ("Bad type: " ^ Thm.string_of_thm_global thy thm); val _ = if Term.add_tvars eqn [] = [tvar] then () else error ("Inconsistent type: " ^ Thm.string_of_thm_global thy thm); val lhs_rhs = case try Logic.dest_equals eqn of SOME lhs_rhs => lhs_rhs | _ => error ("Not an equation: " ^ Syntax.string_of_term_global thy eqn); val c_c' = case try (apply2 (Axclass.unoverload_const thy o dest_Const)) lhs_rhs of SOME c_c' => c_c' | _ => error ("Not an equation with two constants: " ^ Syntax.string_of_term_global thy eqn); val _ = if the_list (Axclass.class_of_param thy (snd c_c')) = [class] then () else error ("Inconsistent class: " ^ Thm.string_of_thm_global thy thm); in Triv_Class_Data.map (AList.update (op =) (class, Thm.trim_context thm)) thy end; local val get_triv_classes = map fst o Triv_Class_Data.get; val (_, triv_of_class) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\triv_of_class\, fn (thy, T, class) => Thm.global_cterm_of thy (Logic.mk_of_class (T, class))))); in fun lift_triv_classes_conv orig_ctxt conv ct = let val thy = Proof_Context.theory_of orig_ctxt; val ctxt = Proof_Context.init_global thy; (*FIXME quasi-global context*) val algebra = Sign.classes_of thy; val triv_classes = get_triv_classes thy; fun additional_classes sort = filter_out (fn class => Sorts.sort_le algebra (sort, [class])) triv_classes; fun mk_entry (v, sort) = let val T = TFree (v, sort); val cT = Thm.ctyp_of ctxt T; val triv_sort = additional_classes sort; in (v, (Sorts.inter_sort algebra (sort, triv_sort), (cT, AList.make (fn class => Thm.of_class (cT, class)) sort @ AList.make (fn class => triv_of_class (thy, T, class)) triv_sort))) end; val vs_tab = map mk_entry (Term.add_tfrees (Thm.term_of ct) []); fun instantiate thm = let val tvars = Term.add_tvars (#1 (Logic.dest_equals (Logic.strip_imp_concl (Thm.prop_of thm)))) []; val instT = map2 (fn v => fn (_, (_, (cT, _))) => (v, cT)) tvars vs_tab; - in Thm.instantiate (instT, []) thm end; + in Thm.instantiate (TVars.make instT, Vars.empty) thm end; fun of_class (TFree (v, _), class) = the (AList.lookup (op =) ((snd o snd o the o AList.lookup (op =) vs_tab) v) class) | of_class (T, _) = error ("Bad type " ^ Syntax.string_of_typ ctxt T); fun strip_of_class thm = let val prems_of_class = Thm.prop_of thm |> Logic.strip_imp_prems |> map (Logic.dest_of_class #> of_class); in fold Thm.elim_implies prems_of_class thm end; in ct |> Thm.term_of |> (map_types o map_type_tfree) (fn (v, _) => TFree (v, (fst o the o AList.lookup (op =) vs_tab) v)) |> Thm.cterm_of ctxt |> conv ctxt |> Thm.strip_shyps |> Thm.varifyT_global |> Thm.unconstrainT |> instantiate |> strip_of_class end; fun lift_triv_classes_rew ctxt rew t = let val thy = Proof_Context.theory_of ctxt; val algebra = Sign.classes_of thy; val triv_classes = get_triv_classes thy; val vs = Term.add_tfrees t []; in t |> (map_types o map_type_tfree) (fn (v, sort) => TFree (v, Sorts.inter_sort algebra (sort, triv_classes))) |> rew |> (map_types o map_type_tfree) (fn (v, sort) => TFree (v, the_default sort (AList.lookup (op =) vs v))) end; end; (** the semantic universe **) (* Functions are given by their semantical function value. To avoid trouble with the ML-type system, these functions have the most generic type, that is "Univ list -> Univ". The calling convention is that the arguments come as a list, the last argument first. In other words, a function call that usually would look like f x_1 x_2 ... x_n or f(x_1,x_2, ..., x_n) would be in our convention called as f [x_n,..,x_2,x_1] Moreover, to handle functions that are still waiting for some arguments we have additionally a list of arguments collected to far and the number of arguments we're still waiting for. *) datatype Univ = Const of int * Univ list (*named (uninterpreted) constants*) | DFree of string * int (*free (uninterpreted) dictionary parameters*) | BVar of int * Univ list (*bound variables, named*) | Abs of (int * (Univ list -> Univ)) * Univ list (*abstractions as closures*); (* constructor functions *) fun abss n f = Abs ((n, f), []); fun apps (Abs ((n, f), xs)) ys = let val k = n - length ys in case int_ord (k, 0) of EQUAL => f (ys @ xs) | LESS => let val (zs, ws) = chop (~ k) ys in apps (f (ws @ xs)) zs end | GREATER => Abs ((k, f), ys @ xs) (*note: reverse convention also for apps!*) end | apps (Const (name, xs)) ys = Const (name, ys @ xs) | apps (BVar (n, xs)) ys = BVar (n, ys @ xs); fun same (Const (k, xs), Const (l, ys)) = k = l andalso eq_list same (xs, ys) | same (DFree (s, k), DFree (t, l)) = s = t andalso k = l | same (BVar (k, xs), BVar (l, ys)) = k = l andalso eq_list same (xs, ys) | same _ = false; (** assembling and compiling ML code from terms **) (* abstract ML syntax *) infix 9 `$` `$$`; fun e1 `$` e2 = "(" ^ e1 ^ " " ^ e2 ^ ")"; fun e `$$` [] = e | e `$$` es = "(" ^ e ^ " " ^ space_implode " " es ^ ")"; fun ml_abs v e = "(fn " ^ v ^ " => " ^ e ^ ")"; fun ml_cases t cs = "(case " ^ t ^ " of " ^ space_implode " | " (map (fn (p, t) => p ^ " => " ^ t) cs) ^ ")"; fun ml_Let d e = "let\n" ^ d ^ " in " ^ e ^ " end"; fun ml_as v t = "(" ^ v ^ " as " ^ t ^ ")"; fun ml_and [] = "true" | ml_and [x] = x | ml_and xs = "(" ^ space_implode " andalso " xs ^ ")"; fun ml_if b x y = "(if " ^ b ^ " then " ^ x ^ " else " ^ y ^ ")"; fun ml_list es = "[" ^ commas es ^ "]"; fun ml_fundefs ([(name, [([], e)])]) = "val " ^ name ^ " = " ^ e ^ "\n" | ml_fundefs (eqs :: eqss) = let fun fundef (name, eqs) = let fun eqn (es, e) = name ^ " " ^ space_implode " " es ^ " = " ^ e in space_implode "\n | " (map eqn eqs) end; in (prefix "fun " o fundef) eqs :: map (prefix "and " o fundef) eqss |> cat_lines |> suffix "\n" end; (* nbe specific syntax and sandbox communication *) structure Univs = Proof_Data ( type T = unit -> Univ list -> Univ list; val empty: T = fn () => raise Fail "Univs"; fun init _ = empty; ); val get_result = Univs.get; val put_result = Univs.put; local val prefix = "Nbe."; val name_put = prefix ^ "put_result"; val name_const = prefix ^ "Const"; val name_abss = prefix ^ "abss"; val name_apps = prefix ^ "apps"; val name_same = prefix ^ "same"; in val univs_cookie = (get_result, put_result, name_put); fun nbe_fun idx_of 0 (Code_Symbol.Constant "") = "nbe_value" | nbe_fun idx_of i sym = "c_" ^ string_of_int (idx_of sym) ^ "_" ^ Code_Symbol.default_base sym ^ "_" ^ string_of_int i; fun nbe_dict v n = "d_" ^ v ^ "_" ^ string_of_int n; fun nbe_bound v = "v_" ^ v; fun nbe_bound_optional NONE = "_" | nbe_bound_optional (SOME v) = nbe_bound v; fun nbe_default v = "w_" ^ v; (*note: these three are the "turning spots" where proper argument order is established!*) fun nbe_apps t [] = t | nbe_apps t ts = name_apps `$$` [t, ml_list (rev ts)]; fun nbe_apps_local idx_of i c ts = nbe_fun idx_of i c `$` ml_list (rev ts); fun nbe_apps_constr ctxt idx_of c ts = let val c' = if Config.get ctxt trace then string_of_int (idx_of c) ^ " (*" ^ Code_Symbol.default_base c ^ "*)" else string_of_int (idx_of c); in name_const `$` ("(" ^ c' ^ ", " ^ ml_list (rev ts) ^ ")") end; fun nbe_abss 0 f = f `$` ml_list [] | nbe_abss n f = name_abss `$$` [string_of_int n, f]; fun nbe_same (v1, v2) = "(" ^ name_same ^ " (" ^ nbe_bound v1 ^ ", " ^ nbe_bound v2 ^ "))"; end; open Basic_Code_Symbol; open Basic_Code_Thingol; (* code generation *) fun assemble_eqnss ctxt idx_of deps eqnss = let fun prep_eqns (c, (vs, eqns)) = let val dicts = maps (fn (v, sort) => map_index (nbe_dict v o fst) sort) vs; val num_args = length dicts + ((length o fst o hd) eqns); in (c, (num_args, (dicts, eqns))) end; val eqnss' = map prep_eqns eqnss; fun assemble_constapp sym dss ts = let val ts' = (maps o map) assemble_dict dss @ ts; in case AList.lookup (op =) eqnss' sym of SOME (num_args, _) => if num_args <= length ts' then let val (ts1, ts2) = chop num_args ts' in nbe_apps (nbe_apps_local idx_of 0 sym ts1) ts2 end else nbe_apps (nbe_abss num_args (nbe_fun idx_of 0 sym)) ts' | NONE => if member (op =) deps sym then nbe_apps (nbe_fun idx_of 0 sym) ts' else nbe_apps_constr ctxt idx_of sym ts' end and assemble_classrels classrels = fold_rev (fn classrel => assemble_constapp (Class_Relation classrel) [] o single) classrels and assemble_dict (Dict (classrels, x)) = assemble_classrels classrels (assemble_plain_dict x) and assemble_plain_dict (Dict_Const (inst, dss)) = assemble_constapp (Class_Instance inst) dss [] | assemble_plain_dict (Dict_Var { var, index, ... }) = nbe_dict var index fun assemble_iterm constapp = let fun of_iterm match_cont t = let val (t', ts) = Code_Thingol.unfold_app t in of_iapp match_cont t' (fold_rev (cons o of_iterm NONE) ts []) end and of_iapp match_cont (IConst { sym, dicts = dss, ... }) ts = constapp sym dss ts | of_iapp match_cont (IVar v) ts = nbe_apps (nbe_bound_optional v) ts | of_iapp match_cont ((v, _) `|=> t) ts = nbe_apps (nbe_abss 1 (ml_abs (ml_list [nbe_bound_optional v]) (of_iterm NONE t))) ts | of_iapp match_cont (ICase { term = t, clauses = clauses, primitive = t0, ... }) ts = nbe_apps (ml_cases (of_iterm NONE t) (map (fn (p, t) => (of_iterm NONE p, of_iterm match_cont t)) clauses @ [("_", case match_cont of SOME s => s | NONE => of_iterm NONE t0)])) ts in of_iterm end; fun subst_nonlin_vars args = let val vs = (fold o Code_Thingol.fold_varnames) (fn v => AList.map_default (op =) (v, 0) (Integer.add 1)) args []; val names = Name.make_context (map fst vs); fun declare v k ctxt = let val vs = Name.invent ctxt v k in (vs, fold Name.declare vs ctxt) end; val (vs_renames, _) = fold_map (fn (v, k) => if k > 1 then declare v (k - 1) #>> (fn vs => (v, vs)) else pair (v, [])) vs names; val samepairs = maps (fn (v, vs) => map (pair v) vs) vs_renames; fun subst_vars (t as IConst _) samepairs = (t, samepairs) | subst_vars (t as IVar NONE) samepairs = (t, samepairs) | subst_vars (t as IVar (SOME v)) samepairs = (case AList.lookup (op =) samepairs v of SOME v' => (IVar (SOME v'), AList.delete (op =) v samepairs) | NONE => (t, samepairs)) | subst_vars (t1 `$ t2) samepairs = samepairs |> subst_vars t1 ||>> subst_vars t2 |>> (op `$) | subst_vars (ICase { primitive = t, ... }) samepairs = subst_vars t samepairs; val (args', _) = fold_map subst_vars args samepairs; in (samepairs, args') end; fun assemble_eqn sym dicts default_args (i, (args, rhs)) = let val match_cont = if Code_Symbol.is_value sym then NONE else SOME (nbe_apps_local idx_of (i + 1) sym (dicts @ default_args)); val assemble_arg = assemble_iterm (fn sym' => fn dss => fn ts => nbe_apps_constr ctxt idx_of sym' ((maps o map) (K "_") dss @ ts)) NONE; val assemble_rhs = assemble_iterm assemble_constapp match_cont; val (samepairs, args') = subst_nonlin_vars args; val s_args = map assemble_arg args'; val s_rhs = if null samepairs then assemble_rhs rhs else ml_if (ml_and (map nbe_same samepairs)) (assemble_rhs rhs) (the match_cont); val eqns = case match_cont of NONE => [([ml_list (rev (dicts @ s_args))], s_rhs)] | SOME default_rhs => [([ml_list (rev (dicts @ map2 ml_as default_args s_args))], s_rhs), ([ml_list (rev (dicts @ default_args))], default_rhs)] in (nbe_fun idx_of i sym, eqns) end; fun assemble_eqns (sym, (num_args, (dicts, eqns))) = let val default_args = map nbe_default (Name.invent Name.context "a" (num_args - length dicts)); val eqns' = map_index (assemble_eqn sym dicts default_args) eqns @ (if Code_Symbol.is_value sym then [] else [(nbe_fun idx_of (length eqns) sym, [([ml_list (rev (dicts @ default_args))], nbe_apps_constr ctxt idx_of sym (dicts @ default_args))])]); in (eqns', nbe_abss num_args (nbe_fun idx_of 0 sym)) end; val (fun_vars, fun_vals) = map_split assemble_eqns eqnss'; val deps_vars = ml_list (map (nbe_fun idx_of 0) deps); in ml_abs deps_vars (ml_Let (ml_fundefs (flat fun_vars)) (ml_list fun_vals)) end; (* compilation of equations *) fun compile_eqnss ctxt nbe_program raw_deps [] = [] | compile_eqnss ctxt nbe_program raw_deps eqnss = let val (deps, deps_vals) = split_list (map_filter (fn dep => Option.map (fn univ => (dep, univ)) (fst ((Code_Symbol.Graph.get_node nbe_program dep)))) raw_deps); val idx_of = raw_deps |> map (fn dep => (dep, snd (Code_Symbol.Graph.get_node nbe_program dep))) |> AList.lookup (op =) |> (fn f => the o f); val s = assemble_eqnss ctxt idx_of deps eqnss; val cs = map fst eqnss; in s |> traced ctxt (fn s => "\n--- code to be evaluated:\n" ^ s) |> pair "" |> Code_Runtime.value ctxt univs_cookie |> (fn f => f deps_vals) |> (fn univs => cs ~~ univs) end; (* extraction of equations from statements *) fun dummy_const sym dss = IConst { sym = sym, typargs = [], dicts = dss, dom = [], annotation = NONE }; fun eqns_of_stmt (_, Code_Thingol.NoStmt) = [] | eqns_of_stmt (_, Code_Thingol.Fun ((_, []), _)) = [] | eqns_of_stmt (sym_const, Code_Thingol.Fun (((vs, _), eqns), _)) = [(sym_const, (vs, map fst eqns))] | eqns_of_stmt (_, Code_Thingol.Datatypecons _) = [] | eqns_of_stmt (_, Code_Thingol.Datatype _) = [] | eqns_of_stmt (sym_class, Code_Thingol.Class (v, (classrels, classparams))) = let val syms = map Class_Relation classrels @ map (Constant o fst) classparams; val params = Name.invent Name.context "d" (length syms); fun mk (k, sym) = (sym, ([(v, [])], [([dummy_const sym_class [] `$$ map (IVar o SOME) params], IVar (SOME (nth params k)))])); in map_index mk syms end | eqns_of_stmt (_, Code_Thingol.Classrel _) = [] | eqns_of_stmt (_, Code_Thingol.Classparam _) = [] | eqns_of_stmt (sym_inst, Code_Thingol.Classinst { class, tyco, vs, superinsts, inst_params, ... }) = [(sym_inst, (vs, [([], dummy_const (Type_Class class) [] `$$ map (fn (class, dss) => dummy_const (Class_Instance (tyco, class)) dss) superinsts @ map (IConst o fst o snd o fst) inst_params)]))]; (* compilation of whole programs *) fun ensure_const_idx name (nbe_program, (maxidx, idx_tab)) = if can (Code_Symbol.Graph.get_node nbe_program) name then (nbe_program, (maxidx, idx_tab)) else (Code_Symbol.Graph.new_node (name, (NONE, maxidx)) nbe_program, (maxidx + 1, Inttab.update_new (maxidx, name) idx_tab)); fun compile_stmts ctxt stmts_deps = let val names = map (fst o fst) stmts_deps; val names_deps = map (fn ((name, _), deps) => (name, deps)) stmts_deps; val eqnss = maps (eqns_of_stmt o fst) stmts_deps; val refl_deps = names_deps |> maps snd |> distinct (op =) |> fold (insert (op =)) names; fun compile nbe_program = eqnss |> compile_eqnss ctxt nbe_program refl_deps |> rpair nbe_program; in fold ensure_const_idx refl_deps #> apfst (fold (fn (name, deps) => fold (curry Code_Symbol.Graph.add_edge name) deps) names_deps #> compile #-> fold (fn (name, univ) => (Code_Symbol.Graph.map_node name o apfst) (K (SOME univ)))) end; fun compile_program { ctxt, program } = let fun add_stmts names (nbe_program, (maxidx, idx_tab)) = if exists ((can o Code_Symbol.Graph.get_node) nbe_program) names then (nbe_program, (maxidx, idx_tab)) else (nbe_program, (maxidx, idx_tab)) |> compile_stmts ctxt (map (fn name => ((name, Code_Symbol.Graph.get_node program name), Code_Symbol.Graph.immediate_succs program name)) names); in fold_rev add_stmts (Code_Symbol.Graph.strong_conn program) end; (** normalization **) (* compilation and reconstruction of terms *) fun compile_term { ctxt, nbe_program, deps, term = (vs, t) } = let val dict_frees = maps (fn (v, sort) => map_index (curry DFree v o fst) sort) vs; in (Code_Symbol.value, (vs, [([], t)])) |> singleton (compile_eqnss ctxt nbe_program deps) |> snd |> (fn t => apps t (rev dict_frees)) end; fun reconstruct_term ctxt (idx_tab : Code_Symbol.T Inttab.table) t = let fun take_until f [] = [] | take_until f (x :: xs) = if f x then [] else x :: take_until f xs; fun is_dict (Const (idx, _)) = (case Inttab.lookup idx_tab idx of SOME (Constant _) => false | _ => true) | is_dict (DFree _) = true | is_dict _ = false; fun const_of_idx idx = case Inttab.lookup idx_tab idx of SOME (Constant const) => const; fun of_apps bounds (t, ts) = fold_map (of_univ bounds) ts #>> (fn ts' => list_comb (t, rev ts')) and of_univ bounds (Const (idx, ts)) typidx = let val ts' = take_until is_dict ts; val const = const_of_idx idx; val T = map_type_tvar (fn ((v, i), _) => Type_Infer.param typidx (v ^ string_of_int i, [])) (Sign.the_const_type (Proof_Context.theory_of ctxt) const); val typidx' = typidx + 1; in of_apps bounds (Term.Const (const, T), ts') typidx' end | of_univ bounds (BVar (n, ts)) typidx = of_apps bounds (Bound (bounds - n - 1), ts) typidx | of_univ bounds (t as Abs _) typidx = typidx |> of_univ (bounds + 1) (apps t [BVar (bounds, [])]) |-> (fn t' => pair (Term.Abs ("u", dummyT, t'))) in of_univ 0 t 0 |> fst end; fun compile_and_reconstruct_term { ctxt, nbe_program, idx_tab, deps, term } = compile_term { ctxt = ctxt, nbe_program = nbe_program, deps = deps, term = term } |> reconstruct_term ctxt idx_tab; fun normalize_term (nbe_program, idx_tab) raw_ctxt t_original ((vs, ty) : typscheme, t) deps = let val ctxt = Syntax.init_pretty_global (Proof_Context.theory_of raw_ctxt); val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt); fun type_infer t' = Syntax.check_term (ctxt |> Config.put Type_Infer.object_logic false |> Config.put Type_Infer_Context.const_sorts false) (Type.constraint (fastype_of t_original) t'); fun check_tvars t' = if null (Term.add_tvars t' []) then t' else error ("Illegal schematic type variables in normalized term: " ^ string_of_term t'); in Code_Preproc.timed "computing NBE expression" #ctxt compile_and_reconstruct_term { ctxt = ctxt, nbe_program = nbe_program, idx_tab = idx_tab, deps = deps, term = (vs, t) } |> traced ctxt (fn t => "Normalized:\n" ^ string_of_term t) |> type_infer |> traced ctxt (fn t => "Types inferred:\n" ^ string_of_term t) |> check_tvars |> traced ctxt (fn _ => "---\n") end; (* function store *) structure Nbe_Functions = Code_Data ( type T = (Univ option * int) Code_Symbol.Graph.T * (int * Code_Symbol.T Inttab.table); val empty = (Code_Symbol.Graph.empty, (0, Inttab.empty)); ); fun compile ignore_cache ctxt program = let val (nbe_program, (_, idx_tab)) = Nbe_Functions.change (if ignore_cache then NONE else SOME (Proof_Context.theory_of ctxt)) (Code_Preproc.timed "compiling NBE program" #ctxt compile_program { ctxt = ctxt, program = program }); in (nbe_program, idx_tab) end; (* evaluation oracle *) fun mk_equals ctxt lhs raw_rhs = let val ty = Thm.typ_of_cterm lhs; val eq = Thm.cterm_of ctxt (Term.Const (\<^const_name>\Pure.eq\, ty --> ty --> propT)); val rhs = Thm.cterm_of ctxt raw_rhs; in Thm.mk_binop eq lhs rhs end; val (_, raw_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\normalization_by_evaluation\, fn (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct) => mk_equals ctxt ct (normalize_term nbe_program_idx_tab ctxt (Thm.term_of ct) vs_ty_t deps)))); fun oracle nbe_program_idx_tab ctxt vs_ty_t deps ct = raw_oracle (nbe_program_idx_tab, ctxt, vs_ty_t, deps, ct); fun dynamic_conv ctxt = lift_triv_classes_conv ctxt (fn ctxt' => Code_Thingol.dynamic_conv ctxt' (fn program => oracle (compile false ctxt program) ctxt')); fun dynamic_value ctxt = lift_triv_classes_rew ctxt (Code_Thingol.dynamic_value ctxt I (fn program => normalize_term (compile false ctxt program) ctxt)); fun static_conv (ctxt_consts as { ctxt, ... }) = let val conv = Code_Thingol.static_conv_thingol ctxt_consts (fn { program, deps = _ } => oracle (compile true ctxt program)); in fn ctxt' => lift_triv_classes_conv ctxt' conv end; fun static_value { ctxt, consts } = let val comp = Code_Thingol.static_value { ctxt = ctxt, lift_postproc = I, consts = consts } (fn { program, deps = _ } => normalize_term (compile false ctxt program)); in fn ctxt' => lift_triv_classes_rew ctxt' (comp ctxt') end; end; diff --git a/src/ZF/Tools/cartprod.ML b/src/ZF/Tools/cartprod.ML --- a/src/ZF/Tools/cartprod.ML +++ b/src/ZF/Tools/cartprod.ML @@ -1,117 +1,117 @@ (* Title: ZF/Tools/cartprod.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1996 University of Cambridge Signatures for inductive definitions. Syntactic operations for Cartesian Products. *) signature FP = (** Description of a fixed point operator **) sig val oper : term (*fixed point operator*) val bnd_mono : term (*monotonicity predicate*) val bnd_monoI : thm (*intro rule for bnd_mono*) val subs : thm (*subset theorem for fp*) val Tarski : thm (*Tarski's fixed point theorem*) val induct : thm (*induction/coinduction rule*) end; signature SU = (** Description of a disjoint sum **) sig val sum : term (*disjoint sum operator*) val inl : term (*left injection*) val inr : term (*right injection*) val elim : term (*case operator*) val case_inl : thm (*inl equality rule for case*) val case_inr : thm (*inr equality rule for case*) val inl_iff : thm (*injectivity of inl, using <->*) val inr_iff : thm (*injectivity of inr, using <->*) val distinct : thm (*distinctness of inl, inr using <->*) val distinct' : thm (*distinctness of inr, inl using <->*) val free_SEs : thm list (*elim rules for SU, and pair_iff!*) end; signature PR = (** Description of a Cartesian product **) sig val sigma : term (*Cartesian product operator*) val pair : term (*pairing operator*) val split_name : string (*name of polymorphic split*) val pair_iff : thm (*injectivity of pairing, using <->*) val split_eq : thm (*equality rule for split*) val fsplitI : thm (*intro rule for fsplit*) val fsplitD : thm (*destruct rule for fsplit*) val fsplitE : thm (*elim rule; apparently never used*) end; signature CARTPROD = (** Derived syntactic functions for produts **) sig val ap_split : typ -> typ -> term -> term val factors : typ -> typ list val mk_prod : typ * typ -> typ val mk_tuple : term -> typ -> term list -> term val pseudo_type : term -> typ val remove_split : Proof.context -> thm -> thm val split_const : typ -> term val split_rule_var : Proof.context -> term * typ * thm -> thm end; functor CartProd_Fun (Pr: PR) : CARTPROD = struct (* Some of these functions expect "pseudo-types" containing products, as in HOL; the true ZF types would just be "i" *) fun mk_prod (T1,T2) = Type("*", [T1,T2]); (*Bogus product type underlying a (possibly nested) Sigma. Lets us share HOL code*) fun pseudo_type (t $ A $ Abs(_,_,B)) = if t = Pr.sigma then mk_prod(pseudo_type A, pseudo_type B) else \<^typ>\i\ | pseudo_type _ = \<^typ>\i\; (*Maps the type T1*...*Tn to [T1,...,Tn], however nested*) fun factors (Type("*", [T1, T2])) = factors T1 @ factors T2 | factors T = [T]; (*Make a well-typed instance of "split"*) fun split_const T = Const(Pr.split_name, [[Ind_Syntax.iT, Ind_Syntax.iT]--->T, Ind_Syntax.iT] ---> T); (*In ap_split S T u, term u expects separate arguments for the factors of S, with result type T. The call creates a new term expecting one argument of type S.*) fun ap_split (Type("*", [T1,T2])) T3 u = split_const T3 $ Abs("v", Ind_Syntax.iT, (*Not T1, as it involves pseudo-product types*) ap_split T2 T3 ((ap_split T1 (factors T2 ---> T3) (incr_boundvars 1 u)) $ Bound 0)) | ap_split T T3 u = u; (*Makes a nested tuple from a list, following the product type structure*) fun mk_tuple pair (Type("*", [T1,T2])) tms = pair $ mk_tuple pair T1 tms $ mk_tuple pair T2 (drop (length (factors T1)) tms) | mk_tuple pair T (t::_) = t; (*Attempts to remove occurrences of split, and pair-valued parameters*) fun remove_split ctxt = rewrite_rule ctxt [Pr.split_eq]; (*Uncurries any Var according to its "pseudo-product type" T1 in the rule*) fun split_rule_var ctxt (Var(v,_), Type("fun",[T1,T2]), rl) = let val T' = factors T1 ---> T2 val newt = ap_split T1 T2 (Var(v,T')) in remove_split ctxt - (Drule.instantiate_normalize ([], - [((v, Ind_Syntax.iT-->T2), Thm.cterm_of ctxt newt)]) rl) + (Drule.instantiate_normalize (TVars.empty, + Vars.make [((v, Ind_Syntax.iT-->T2), Thm.cterm_of ctxt newt)]) rl) end | split_rule_var _ (t,T,rl) = rl; end; diff --git a/src/ZF/Tools/inductive_package.ML b/src/ZF/Tools/inductive_package.ML --- a/src/ZF/Tools/inductive_package.ML +++ b/src/ZF/Tools/inductive_package.ML @@ -1,612 +1,612 @@ (* Title: ZF/Tools/inductive_package.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Fixedpoint definition module -- for Inductive/Coinductive Definitions The functor will be instantiated for normal sums/products (inductive defs) and non-standard sums/products (coinductive defs) Sums are used only for mutual recursion; Products are used only to derive "streamlined" induction rules for relations *) type inductive_result = {defs : thm list, (*definitions made in thy*) bnd_mono : thm, (*monotonicity for the lfp definition*) dom_subset : thm, (*inclusion of recursive set in dom*) intrs : thm list, (*introduction rules*) elim : thm, (*case analysis theorem*) induct : thm, (*main induction rule*) mutual_induct : thm}; (*mutual induction rule*) (*Functor's result signature*) signature INDUCTIVE_PACKAGE = sig (*Insert definitions for the recursive sets, which must *already* be declared as constants in parent theory!*) val add_inductive_i: bool -> term list * term -> ((binding * term) * attribute list) list -> thm list * thm list * thm list * thm list -> theory -> theory * inductive_result val add_inductive: string list * string -> ((binding * string) * Token.src list) list -> (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list * (Facts.ref * Token.src list) list -> theory -> theory * inductive_result end; (*Declares functions to add fixedpoint/constructor defs to a theory. Recursive sets must *already* be declared as constants.*) functor Add_inductive_def_Fun (structure Fp: FP and Pr : PR and CP: CARTPROD and Su : SU val coind: bool) : INDUCTIVE_PACKAGE = struct val co_prefix = if coind then "co" else ""; (* utils *) (*make distinct individual variables a1, a2, a3, ..., an. *) fun mk_frees a [] = [] | mk_frees a (T::Ts) = Free(a,T) :: mk_frees (Symbol.bump_string a) Ts; (* add_inductive(_i) *) (*internal version, accepting terms*) fun add_inductive_i verbose (rec_tms, dom_sum) raw_intr_specs (monos, con_defs, type_intrs, type_elims) thy0 = let val ctxt0 = Proof_Context.init_global thy0; val intr_specs = map (apfst (apfst Binding.name_of)) raw_intr_specs; val (intr_names, intr_tms) = split_list (map fst intr_specs); val case_names = Rule_Cases.case_names intr_names; (*recT and rec_params should agree for all mutually recursive components*) val rec_hds = map head_of rec_tms; val dummy = rec_hds |> forall (fn t => is_Const t orelse error ("Recursive set not previously declared as constant: " ^ Syntax.string_of_term ctxt0 t)); (*Now we know they are all Consts, so get their names, type and params*) val rec_names = map (#1 o dest_Const) rec_hds and (Const(_,recT),rec_params) = strip_comb (hd rec_tms); val rec_base_names = map Long_Name.base_name rec_names; val dummy = rec_base_names |> forall (fn a => Symbol_Pos.is_identifier a orelse error ("Base name of recursive set not an identifier: " ^ a)); local (*Checking the introduction rules*) val intr_sets = map (#2 o Ind_Syntax.rule_concl_msg thy0) intr_tms; fun intr_ok set = case head_of set of Const(a,recT) => member (op =) rec_names a | _ => false; in val dummy = intr_sets |> forall (fn t => intr_ok t orelse error ("Conclusion of rule does not name a recursive set: " ^ Syntax.string_of_term ctxt0 t)); end; val dummy = rec_params |> forall (fn t => is_Free t orelse error ("Param in recursion term not a free variable: " ^ Syntax.string_of_term ctxt0 t)); (*** Construct the fixedpoint definition ***) val mk_variant = singleton (Name.variant_list (List.foldr Misc_Legacy.add_term_names [] intr_tms)); val z' = mk_variant"z" and X' = mk_variant"X" and w' = mk_variant"w"; fun dest_tprop (Const(\<^const_name>\Trueprop\,_) $ P) = P | dest_tprop Q = error ("Ill-formed premise of introduction rule: " ^ Syntax.string_of_term ctxt0 Q); (*Makes a disjunct from an introduction rule*) fun fp_part intr = (*quantify over rule's free vars except parameters*) let val prems = map dest_tprop (Logic.strip_imp_prems intr) val dummy = List.app (fn rec_hd => List.app (Ind_Syntax.chk_prem rec_hd) prems) rec_hds val exfrees = subtract (op =) rec_params (Misc_Legacy.term_frees intr) val zeq = FOLogic.mk_eq (Free(z', Ind_Syntax.iT), #1 (Ind_Syntax.rule_concl intr)) in List.foldr FOLogic.mk_exists (Balanced_Tree.make FOLogic.mk_conj (zeq::prems)) exfrees end; (*The Part(A,h) terms -- compose injections to make h*) fun mk_Part (Bound 0) = Free(X', Ind_Syntax.iT) (*no mutual rec, no Part needed*) | mk_Part h = \<^const>\Part\ $ Free(X', Ind_Syntax.iT) $ Abs (w', Ind_Syntax.iT, h); (*Access to balanced disjoint sums via injections*) val parts = map mk_Part (Balanced_Tree.accesses {left = fn t => Su.inl $ t, right = fn t => Su.inr $ t, init = Bound 0} (length rec_tms)); (*replace each set by the corresponding Part(A,h)*) val part_intrs = map (subst_free (rec_tms ~~ parts) o fp_part) intr_tms; val fp_abs = absfree (X', Ind_Syntax.iT) (Ind_Syntax.mk_Collect (z', dom_sum, Balanced_Tree.make FOLogic.mk_disj part_intrs)); val fp_rhs = Fp.oper $ dom_sum $ fp_abs val dummy = List.app (fn rec_hd => (Logic.occs (rec_hd, fp_rhs) andalso error "Illegal occurrence of recursion operator"; ())) rec_hds; (*** Make the new theory ***) (*A key definition: If no mutual recursion then it equals the one recursive set. If mutual recursion then it differs from all the recursive sets. *) val big_rec_base_name = space_implode "_" rec_base_names; val big_rec_name = Proof_Context.intern_const ctxt0 big_rec_base_name; val _ = if verbose then writeln ((if coind then "Coind" else "Ind") ^ "uctive definition " ^ quote big_rec_name) else (); (*Big_rec... is the union of the mutually recursive sets*) val big_rec_tm = list_comb(Const(big_rec_name,recT), rec_params); (*The individual sets must already be declared*) val axpairs = map Misc_Legacy.mk_defpair ((big_rec_tm, fp_rhs) :: (case parts of [_] => [] (*no mutual recursion*) | _ => rec_tms ~~ (*define the sets as Parts*) map (subst_atomic [(Free (X', Ind_Syntax.iT), big_rec_tm)]) parts)); (*tracing: print the fixedpoint definition*) val dummy = if !Ind_Syntax.trace then writeln (cat_lines (map (Syntax.string_of_term ctxt0 o #2) axpairs)) else () (*add definitions of the inductive sets*) val (_, thy1) = thy0 |> Sign.add_path big_rec_base_name |> Global_Theory.add_defs false (map (Thm.no_attributes o apfst Binding.name) axpairs); (*fetch fp definitions from the theory*) val big_rec_def::part_rec_defs = map (Misc_Legacy.get_def thy1) (case rec_names of [_] => rec_names | _ => big_rec_base_name::rec_names); (********) val dummy = writeln " Proving monotonicity..."; val bnd_mono0 = Goal.prove_global thy1 [] [] (FOLogic.mk_Trueprop (Fp.bnd_mono $ dom_sum $ fp_abs)) (fn {context = ctxt, ...} => EVERY [resolve_tac ctxt [@{thm Collect_subset} RS @{thm bnd_monoI}] 1, REPEAT (ares_tac ctxt (@{thms basic_monos} @ monos) 1)]); val dom_subset0 = Drule.export_without_context (big_rec_def RS Fp.subs); val ([bnd_mono, dom_subset], thy2) = thy1 |> Global_Theory.add_thms [((Binding.name "bnd_mono", bnd_mono0), []), ((Binding.name "dom_subset", dom_subset0), [])]; val unfold = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.Tarski); (********) val dummy = writeln " Proving the introduction rules..."; (*Mutual recursion? Helps to derive subset rules for the individual sets.*) val Part_trans = case rec_names of [_] => asm_rl | _ => Drule.export_without_context (@{thm Part_subset} RS @{thm subset_trans}); (*To type-check recursive occurrences of the inductive sets, possibly enclosed in some monotonic operator M.*) val rec_typechecks = [dom_subset] RL (asm_rl :: ([Part_trans] RL monos)) RL [@{thm subsetD}]; (*Type-checking is hardest aspect of proof; disjIn selects the correct disjunct after unfolding*) fun intro_tacsf disjIn ctxt = [DETERM (stac ctxt unfold 1), REPEAT (resolve_tac ctxt [@{thm Part_eqI}, @{thm CollectI}] 1), (*Now 2-3 subgoals: typechecking, the disjunction, perhaps equality.*) resolve_tac ctxt [disjIn] 2, (*Not ares_tac, since refl must be tried before equality assumptions; backtracking may occur if the premises have extra variables!*) DEPTH_SOLVE_1 (resolve_tac ctxt [@{thm refl}, @{thm exI}, @{thm conjI}] 2 APPEND assume_tac ctxt 2), (*Now solve the equations like Tcons(a,f) = Inl(?b4)*) rewrite_goals_tac ctxt con_defs, REPEAT (resolve_tac ctxt @{thms refl} 2), (*Typechecking; this can fail*) if !Ind_Syntax.trace then print_tac ctxt "The type-checking subgoal:" else all_tac, REPEAT (FIRSTGOAL (dresolve_tac ctxt rec_typechecks ORELSE' eresolve_tac ctxt (asm_rl :: @{thm PartE} :: @{thm SigmaE2} :: type_elims) ORELSE' hyp_subst_tac ctxt)), if !Ind_Syntax.trace then print_tac ctxt "The subgoal after monos, type_elims:" else all_tac, DEPTH_SOLVE (swap_res_tac ctxt (@{thm SigmaI} :: @{thm subsetI} :: type_intrs) 1)]; (*combines disjI1 and disjI2 to get the corresponding nested disjunct...*) val mk_disj_rls = Balanced_Tree.accesses {left = fn rl => rl RS @{thm disjI1}, right = fn rl => rl RS @{thm disjI2}, init = @{thm asm_rl}}; val intrs0 = (intr_tms, map intro_tacsf (mk_disj_rls (length intr_tms))) |> ListPair.map (fn (t, tacs) => Goal.prove_global thy2 [] [] t (fn {context = ctxt, ...} => EVERY (rewrite_goals_tac ctxt part_rec_defs :: tacs ctxt))); val ([intrs], thy3) = thy2 |> Global_Theory.add_thmss [((Binding.name "intros", intrs0), [])]; val ctxt3 = Proof_Context.init_global thy3; (********) val dummy = writeln " Proving the elimination rule..."; (*Breaks down logical connectives in the monotonic function*) fun basic_elim_tac ctxt = REPEAT (SOMEGOAL (eresolve_tac ctxt (Ind_Syntax.elim_rls @ Su.free_SEs) ORELSE' bound_hyp_subst_tac ctxt)) THEN prune_params_tac ctxt (*Mutual recursion: collapse references to Part(D,h)*) THEN (PRIMITIVE (fold_rule ctxt part_rec_defs)); (*Elimination*) val (elim, thy4) = thy3 |> Global_Theory.add_thm ((Binding.name "elim", rule_by_tactic ctxt3 (basic_elim_tac ctxt3) (unfold RS Ind_Syntax.equals_CollectD)), []); val ctxt4 = Proof_Context.init_global thy4; (*Applies freeness of the given constructors, which *must* be unfolded by the given defs. Cannot simply use the local con_defs because con_defs=[] for inference systems. Proposition A should have the form t:Si where Si is an inductive set*) fun make_cases ctxt A = rule_by_tactic ctxt (basic_elim_tac ctxt THEN ALLGOALS (asm_full_simp_tac ctxt) THEN basic_elim_tac ctxt) (Thm.assume A RS elim) |> Drule.export_without_context_open; fun induction_rules raw_induct = let val dummy = writeln " Proving the induction rule..."; (*** Prove the main induction rule ***) val pred_name = "P"; (*name for predicate variables*) (*Used to make induction rules; ind_alist = [(rec_tm1,pred1),...] associates predicates with rec ops prem is a premise of an intr rule*) fun add_induct_prem ind_alist (prem as Const (\<^const_name>\Trueprop\, _) $ (Const (\<^const_name>\mem\, _) $ t $ X), iprems) = (case AList.lookup (op aconv) ind_alist X of SOME pred => prem :: FOLogic.mk_Trueprop (pred $ t) :: iprems | NONE => (*possibly membership in M(rec_tm), for M monotone*) let fun mk_sb (rec_tm,pred) = (rec_tm, \<^const>\Collect\ $ rec_tm $ pred) in subst_free (map mk_sb ind_alist) prem :: iprems end) | add_induct_prem ind_alist (prem,iprems) = prem :: iprems; (*Make a premise of the induction rule.*) fun induct_prem ind_alist intr = let val xs = subtract (op =) rec_params (Misc_Legacy.term_frees intr) val iprems = List.foldr (add_induct_prem ind_alist) [] (Logic.strip_imp_prems intr) val (t,X) = Ind_Syntax.rule_concl intr val (SOME pred) = AList.lookup (op aconv) ind_alist X val concl = FOLogic.mk_Trueprop (pred $ t) in fold_rev Logic.all xs (Logic.list_implies (iprems,concl)) end handle Bind => error"Recursion term not found in conclusion"; (*Minimizes backtracking by delivering the correct premise to each goal. Intro rules with extra Vars in premises still cause some backtracking *) fun ind_tac _ [] 0 = all_tac | ind_tac ctxt (prem::prems) i = DEPTH_SOLVE_1 (ares_tac ctxt [prem, @{thm refl}] i) THEN ind_tac ctxt prems (i-1); val pred = Free(pred_name, Ind_Syntax.iT --> FOLogic.oT); val ind_prems = map (induct_prem (map (rpair pred) rec_tms)) intr_tms; val dummy = if ! Ind_Syntax.trace then writeln (cat_lines (["ind_prems:"] @ map (Syntax.string_of_term ctxt4) ind_prems @ ["raw_induct:", Thm.string_of_thm ctxt4 raw_induct])) else (); (*We use a MINIMAL simpset. Even FOL_ss contains too many simpules. If the premises get simplified, then the proofs could fail.*) val min_ss = (empty_simpset ctxt4 |> Simplifier.set_mksimps (fn ctxt => map mk_eq o ZF_atomize o Variable.gen_all ctxt)) setSolver (mk_solver "minimal" (fn ctxt => resolve_tac ctxt (triv_rls @ Simplifier.prems_of ctxt) ORELSE' assume_tac ctxt ORELSE' eresolve_tac ctxt @{thms FalseE})); val quant_induct = Goal.prove_global thy4 [] ind_prems (FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, pred))) (fn {context = ctxt, prems} => EVERY [rewrite_goals_tac ctxt part_rec_defs, resolve_tac ctxt [@{thm impI} RS @{thm allI}] 1, DETERM (eresolve_tac ctxt [raw_induct] 1), (*Push Part inside Collect*) full_simp_tac (min_ss addsimps [@{thm Part_Collect}]) 1, (*This CollectE and disjE separates out the introduction rules*) REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm disjE}])), (*Now break down the individual cases. No disjE here in case some premise involves disjunction.*) REPEAT (FIRSTGOAL (eresolve_tac ctxt [@{thm CollectE}, @{thm exE}, @{thm conjE}] ORELSE' (bound_hyp_subst_tac ctxt))), ind_tac ctxt (rev (map (rewrite_rule ctxt part_rec_defs) prems)) (length prems)]); val dummy = if ! Ind_Syntax.trace then writeln ("quant_induct:\n" ^ Thm.string_of_thm ctxt4 quant_induct) else (); (*** Prove the simultaneous induction rule ***) (*Make distinct predicates for each inductive set*) (*The components of the element type, several if it is a product*) val elem_type = CP.pseudo_type dom_sum; val elem_factors = CP.factors elem_type; val elem_frees = mk_frees "za" elem_factors; val elem_tuple = CP.mk_tuple Pr.pair elem_type elem_frees; (*Given a recursive set and its domain, return the "fsplit" predicate and a conclusion for the simultaneous induction rule. NOTE. This will not work for mutually recursive predicates. Previously a summand 'domt' was also an argument, but this required the domain of mutual recursion to invariably be a disjoint sum.*) fun mk_predpair rec_tm = let val rec_name = (#1 o dest_Const o head_of) rec_tm val pfree = Free(pred_name ^ "_" ^ Long_Name.base_name rec_name, elem_factors ---> FOLogic.oT) val qconcl = List.foldr FOLogic.mk_all (FOLogic.imp $ (\<^const>\mem\ $ elem_tuple $ rec_tm) $ (list_comb (pfree, elem_frees))) elem_frees in (CP.ap_split elem_type FOLogic.oT pfree, qconcl) end; val (preds,qconcls) = split_list (map mk_predpair rec_tms); (*Used to form simultaneous induction lemma*) fun mk_rec_imp (rec_tm,pred) = FOLogic.imp $ (\<^const>\mem\ $ Bound 0 $ rec_tm) $ (pred $ Bound 0); (*To instantiate the main induction rule*) val induct_concl = FOLogic.mk_Trueprop (Ind_Syntax.mk_all_imp (big_rec_tm, Abs("z", Ind_Syntax.iT, Balanced_Tree.make FOLogic.mk_conj (ListPair.map mk_rec_imp (rec_tms, preds))))) and mutual_induct_concl = FOLogic.mk_Trueprop (Balanced_Tree.make FOLogic.mk_conj qconcls); val dummy = if !Ind_Syntax.trace then (writeln ("induct_concl = " ^ Syntax.string_of_term ctxt4 induct_concl); writeln ("mutual_induct_concl = " ^ Syntax.string_of_term ctxt4 mutual_induct_concl)) else (); fun lemma_tac ctxt = FIRST' [eresolve_tac ctxt [@{thm asm_rl}, @{thm conjE}, @{thm PartE}, @{thm mp}], resolve_tac ctxt [@{thm allI}, @{thm impI}, @{thm conjI}, @{thm Part_eqI}], dresolve_tac ctxt [@{thm spec}, @{thm mp}, Pr.fsplitD]]; val need_mutual = length rec_names > 1; val lemma = (*makes the link between the two induction rules*) if need_mutual then (writeln " Proving the mutual induction rule..."; Goal.prove_global thy4 [] [] (Logic.mk_implies (induct_concl, mutual_induct_concl)) (fn {context = ctxt, ...} => EVERY [rewrite_goals_tac ctxt part_rec_defs, REPEAT (rewrite_goals_tac ctxt [Pr.split_eq] THEN lemma_tac ctxt 1)])) else (writeln " [ No mutual induction rule needed ]"; @{thm TrueI}); val dummy = if ! Ind_Syntax.trace then writeln ("lemma: " ^ Thm.string_of_thm ctxt4 lemma) else (); (*Mutual induction follows by freeness of Inl/Inr.*) (*Simplification largely reduces the mutual induction rule to the standard rule*) val mut_ss = min_ss addsimps [Su.distinct, Su.distinct', Su.inl_iff, Su.inr_iff]; val all_defs = con_defs @ part_rec_defs; (*Removes Collects caused by M-operators in the intro rules. It is very hard to simplify list({v: tf. (v : t --> P_t(v)) & (v : f --> P_f(v))}) where t==Part(tf,Inl) and f==Part(tf,Inr) to list({v: tf. P_t(v)}). Instead the following rules extract the relevant conjunct. *) val cmonos = [@{thm subset_refl} RS @{thm Collect_mono}] RL monos RLN (2,[@{thm rev_subsetD}]); (*Minimizes backtracking by delivering the correct premise to each goal*) fun mutual_ind_tac _ [] 0 = all_tac | mutual_ind_tac ctxt (prem::prems) i = DETERM (SELECT_GOAL ( (*Simplify the assumptions and goal by unfolding Part and using freeness of the Sum constructors; proves all but one conjunct by contradiction*) rewrite_goals_tac ctxt all_defs THEN simp_tac (mut_ss addsimps [@{thm Part_iff}]) 1 THEN IF_UNSOLVED (*simp_tac may have finished it off!*) ((*simplify assumptions*) (*some risk of excessive simplification here -- might have to identify the bare minimum set of rewrites*) full_simp_tac (mut_ss addsimps @{thms conj_simps} @ @{thms imp_simps} @ @{thms quant_simps}) 1 THEN (*unpackage and use "prem" in the corresponding place*) REPEAT (resolve_tac ctxt @{thms impI} 1) THEN resolve_tac ctxt [rewrite_rule ctxt all_defs prem] 1 THEN (*prem must not be REPEATed below: could loop!*) DEPTH_SOLVE (FIRSTGOAL (ares_tac ctxt [@{thm impI}] ORELSE' eresolve_tac ctxt (@{thm conjE} :: @{thm mp} :: cmonos)))) ) i) THEN mutual_ind_tac ctxt prems (i-1); val mutual_induct_fsplit = if need_mutual then Goal.prove_global thy4 [] (map (induct_prem (rec_tms~~preds)) intr_tms) mutual_induct_concl (fn {context = ctxt, prems} => EVERY [resolve_tac ctxt [quant_induct RS lemma] 1, mutual_ind_tac ctxt (rev prems) (length prems)]) else @{thm TrueI}; (** Uncurrying the predicate in the ordinary induction rule **) (*instantiate the variable to a tuple, if it is non-trivial, in order to allow the predicate to be "opened up". The name "x.1" comes from the "RS spec" !*) val inst = case elem_frees of [_] => I - | _ => Drule.instantiate_normalize ([], [(((("x",1), Ind_Syntax.iT)), + | _ => Drule.instantiate_normalize (TVars.empty, Vars.make [(((("x",1), Ind_Syntax.iT)), Thm.global_cterm_of thy4 elem_tuple)]); (*strip quantifier and the implication*) val induct0 = inst (quant_induct RS @{thm spec} RSN (2, @{thm rev_mp})); val Const (\<^const_name>\Trueprop\, _) $ (pred_var $ _) = Thm.concl_of induct0 val induct0 = CP.split_rule_var ctxt4 (pred_var, elem_type-->FOLogic.oT, induct0) |> Drule.export_without_context and mutual_induct = CP.remove_split ctxt4 mutual_induct_fsplit val ([induct, mutual_induct], thy5) = thy4 |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), induct0), [case_names, Induct.induct_pred big_rec_name]), ((Binding.name "mutual_induct", mutual_induct), [case_names])]; in ((induct, mutual_induct), thy5) end; (*of induction_rules*) val raw_induct = Drule.export_without_context ([big_rec_def, bnd_mono] MRS Fp.induct) val ((induct, mutual_induct), thy5) = if not coind then induction_rules raw_induct else thy4 |> Global_Theory.add_thms [((Binding.name (co_prefix ^ "induct"), raw_induct), [])] |> apfst (hd #> pair @{thm TrueI}); val (([cases], [defs]), thy6) = thy5 |> IndCases.declare big_rec_name make_cases |> Global_Theory.add_thms [((Binding.name "cases", elim), [case_names, Induct.cases_pred big_rec_name])] ||>> (Global_Theory.add_thmss o map Thm.no_attributes) [(Binding.name "defs", big_rec_def :: part_rec_defs)]; val (named_intrs, thy7) = thy6 |> Global_Theory.add_thms ((map Binding.name intr_names ~~ intrs) ~~ map #2 intr_specs) ||> Sign.parent_path; in (thy7, {defs = defs, bnd_mono = bnd_mono, dom_subset = dom_subset, intrs = named_intrs, elim = cases, induct = induct, mutual_induct = mutual_induct}) end; (*source version*) fun add_inductive (srec_tms, sdom_sum) intr_srcs (raw_monos, raw_con_defs, raw_type_intrs, raw_type_elims) thy = let val ctxt = Proof_Context.init_global thy; val read_terms = map (Syntax.parse_term ctxt #> Type.constraint Ind_Syntax.iT) #> Syntax.check_terms ctxt; val intr_atts = map (map (Attrib.attribute_cmd ctxt) o snd) intr_srcs; val sintrs = map fst intr_srcs ~~ intr_atts; val rec_tms = read_terms srec_tms; val dom_sum = singleton read_terms sdom_sum; val intr_tms = Syntax.read_props ctxt (map (snd o fst) sintrs); val intr_specs = (map (fst o fst) sintrs ~~ intr_tms) ~~ map snd sintrs; val monos = Attrib.eval_thms ctxt raw_monos; val con_defs = Attrib.eval_thms ctxt raw_con_defs; val type_intrs = Attrib.eval_thms ctxt raw_type_intrs; val type_elims = Attrib.eval_thms ctxt raw_type_elims; in thy |> add_inductive_i true (rec_tms, dom_sum) intr_specs (monos, con_defs, type_intrs, type_elims) end; (* outer syntax *) fun mk_ind (((((doms, intrs), monos), con_defs), type_intrs), type_elims) = #1 o add_inductive doms (map (fn ((x, y), z) => ((x, z), y)) intrs) (monos, con_defs, type_intrs, type_elims); val ind_decl = (\<^keyword>\domains\ |-- Parse.!!! (Parse.enum1 "+" Parse.term -- ((\<^keyword>\\\ || \<^keyword>\<=\) |-- Parse.term))) -- (\<^keyword>\intros\ |-- Parse.!!! (Scan.repeat1 (Parse_Spec.opt_thm_name ":" -- Parse.prop))) -- Scan.optional (\<^keyword>\monos\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\con_defs\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\type_intros\ |-- Parse.!!! Parse.thms1) [] -- Scan.optional (\<^keyword>\type_elims\ |-- Parse.!!! Parse.thms1) [] >> (Toplevel.theory o mk_ind); val _ = Outer_Syntax.command (if coind then \<^command_keyword>\coinductive\ else \<^command_keyword>\inductive\) ("define " ^ co_prefix ^ "inductive sets") ind_decl; end;