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,1259 @@ (*: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: "Symtab.set * Symtab.set -> int -> 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.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/Doc/Implementation/Proof.thy b/src/Doc/Implementation/Proof.thy --- a/src/Doc/Implementation/Proof.thy +++ b/src/Doc/Implementation/Proof.thy @@ -1,473 +1,473 @@ (*:maxLineLen=78:*) theory Proof imports Base begin chapter \Structured proofs\ section \Variables \label{sec:variables}\ text \ Any variable that is not explicitly bound by \\\-abstraction is considered as ``free''. Logically, free variables act like outermost universal quantification at the sequent level: \A\<^sub>1(x), \, A\<^sub>n(x) \ B(x)\ means that the result holds \<^emph>\for all\ values of \x\. Free variables for terms (not types) can be fully internalized into the logic: \\ B(x)\ and \\ \x. B(x)\ are interchangeable, provided that \x\ does not occur elsewhere in the context. Inspecting \\ \x. B(x)\ more closely, we see that inside the quantifier, \x\ is essentially ``arbitrary, but fixed'', while from outside it appears as a place-holder for instantiation (thanks to \\\ elimination). The Pure logic represents the idea of variables being either inside or outside the current scope by providing separate syntactic categories for \<^emph>\fixed variables\ (e.g.\ \x\) vs.\ \<^emph>\schematic variables\ (e.g.\ \?x\). Incidently, a universal result \\ \x. B(x)\ has the HHF normal form \\ B(?x)\, which represents its generality without requiring an explicit quantifier. The same principle works for type variables: \\ B(?\)\ represents the idea of ``\\ \\. B(\)\'' without demanding a truly polymorphic framework. \<^medskip> Additional care is required to treat type variables in a way that facilitates type-inference. In principle, term variables depend on type variables, which means that type variables would have to be declared first. For example, a raw type-theoretic framework would demand the context to be constructed in stages as follows: \\ = \: type, x: \, a: A(x\<^sub>\)\. We allow a slightly less formalistic mode of operation: term variables \x\ are fixed without specifying a type yet (essentially \<^emph>\all\ potential occurrences of some instance \x\<^sub>\\ are fixed); the first occurrence of \x\ within a specific term assigns its most general type, which is then maintained consistently in the context. The above example becomes \\ = x: term, \: type, A(x\<^sub>\)\, where type \\\ is fixed \<^emph>\after\ term \x\, and the constraint \x :: \\ is an implicit consequence of the occurrence of \x\<^sub>\\ in the subsequent proposition. This twist of dependencies is also accommodated by the reverse operation of exporting results from a context: a type variable \\\ is considered fixed as long as it occurs in some fixed term variable of the context. For example, exporting \x: term, \: type \ x\<^sub>\ \ x\<^sub>\\ produces in the first step \x: term \ x\<^sub>\ \ x\<^sub>\\ for fixed \\\, and only in the second step \\ ?x\<^sub>?\<^sub>\ \ ?x\<^sub>?\<^sub>\\ for schematic \?x\ and \?\\. The following Isar source text illustrates this scenario. \ notepad begin { fix x \ \all potential occurrences of some \x::\\ are fixed\ { have "x::'a \ x" \ \implicit type assignment by concrete occurrence\ by (rule reflexive) } thm this \ \result still with fixed type \'a\\ } thm this \ \fully general result for arbitrary \?x::?'a\\ end text \ The Isabelle/Isar proof context manages the details of term vs.\ type variables, with high-level principles for moving the frontier between fixed and schematic variables. The \add_fixes\ operation explicitly declares fixed variables; the \declare_term\ operation absorbs a term into a context by fixing new type variables and adding syntactic constraints. The \export\ operation is able to perform the main work of generalizing term and type variables as sketched above, assuming that fixing variables and terms have been declared properly. There \import\ operation makes a generalized fact a genuine part of the context, by inventing fixed variables for the schematic ones. The effect can be reversed by using \export\ later, potentially with an extended context; the result is equivalent to the original modulo renaming of schematic variables. The \focus\ operation provides a variant of \import\ for nested propositions (with explicit quantification): \\x\<^sub>1 \ x\<^sub>n. B(x\<^sub>1, \, x\<^sub>n)\ is decomposed by inventing fixed variables \x\<^sub>1, \, x\<^sub>n\ for the body. \ text %mlref \ \begin{mldecls} @{define_ML Variable.add_fixes: " string list -> Proof.context -> string list * Proof.context"} \\ @{define_ML Variable.variant_fixes: " string list -> Proof.context -> string list * Proof.context"} \\ @{define_ML Variable.declare_term: "term -> Proof.context -> Proof.context"} \\ @{define_ML Variable.declare_constraints: "term -> Proof.context -> Proof.context"} \\ @{define_ML Variable.export: "Proof.context -> Proof.context -> thm list -> thm list"} \\ @{define_ML Variable.polymorphic: "Proof.context -> term list -> term list"} \\ @{define_ML Variable.import: "bool -> thm list -> Proof.context -> - ((ctyp Term_Subst.TVars.table * cterm Term_Subst.Vars.table) * thm list) + ((ctyp TVars.table * cterm Vars.table) * thm list) * Proof.context"} \\ @{define_ML Variable.focus: "binding list option -> term -> Proof.context -> ((string * (string * typ)) list * term) * Proof.context"} \\ \end{mldecls} \<^descr> \<^ML>\Variable.add_fixes\~\xs ctxt\ fixes term variables \xs\, returning the resulting internal names. By default, the internal representation coincides with the external one, which also means that the given variables must not be fixed already. There is a different policy within a local proof body: the given names are just hints for newly invented Skolem variables. \<^descr> \<^ML>\Variable.variant_fixes\ is similar to \<^ML>\Variable.add_fixes\, but always produces fresh variants of the given names. \<^descr> \<^ML>\Variable.declare_term\~\t ctxt\ declares term \t\ to belong to the context. This automatically fixes new type variables, but not term variables. Syntactic constraints for type and term variables are declared uniformly, though. \<^descr> \<^ML>\Variable.declare_constraints\~\t ctxt\ declares syntactic constraints from term \t\, without making it part of the context yet. \<^descr> \<^ML>\Variable.export\~\inner outer thms\ generalizes fixed type and term variables in \thms\ according to the difference of the \inner\ and \outer\ context, following the principles sketched above. \<^descr> \<^ML>\Variable.polymorphic\~\ctxt ts\ generalizes type variables in \ts\ as far as possible, even those occurring in fixed term variables. The default policy of type-inference is to fix newly introduced type variables, which is essentially reversed with \<^ML>\Variable.polymorphic\: here the given terms are detached from the context as far as possible. \<^descr> \<^ML>\Variable.import\~\open thms ctxt\ invents fixed type and term variables for the schematic ones occurring in \thms\. The \open\ flag indicates whether the fixed names should be accessible to the user, otherwise newly introduced names are marked as ``internal'' (\secref{sec:names}). \<^descr> \<^ML>\Variable.focus\~\bindings B\ decomposes the outermost \\\ prefix of proposition \B\, using the given name bindings. \ text %mlex \ The following example shows how to work with fixed term and type parameters and with type-inference. \ ML_val \ (*static compile-time context -- for testing only*) val ctxt0 = \<^context>; (*locally fixed parameters -- no type assignment yet*) val ([x, y], ctxt1) = ctxt0 |> Variable.add_fixes ["x", "y"]; (*t1: most general fixed type; t1': most general arbitrary type*) val t1 = Syntax.read_term ctxt1 "x"; val t1' = singleton (Variable.polymorphic ctxt1) t1; (*term u enforces specific type assignment*) val u = Syntax.read_term ctxt1 "(x::nat) \ y"; (*official declaration of u -- propagates constraints etc.*) val ctxt2 = ctxt1 |> Variable.declare_term u; val t2 = Syntax.read_term ctxt2 "x"; (*x::nat is enforced*) \ text \ In the above example, the starting context is derived from the toplevel theory, which means that fixed variables are internalized literally: \x\ is mapped again to \x\, and attempting to fix it again in the subsequent context is an error. Alternatively, fixed parameters can be renamed explicitly as follows: \ ML_val \ val ctxt0 = \<^context>; val ([x1, x2, x3], ctxt1) = ctxt0 |> Variable.variant_fixes ["x", "x", "x"]; \ text \ The following ML code can now work with the invented names of \x1\, \x2\, \x3\, without depending on the details on the system policy for introducing these variants. Recall that within a proof body the system always invents fresh ``Skolem constants'', e.g.\ as follows: \ notepad begin ML_prf %"ML" \val ctxt0 = \<^context>; val ([x1], ctxt1) = ctxt0 |> Variable.add_fixes ["x"]; val ([x2], ctxt2) = ctxt1 |> Variable.add_fixes ["x"]; val ([x3], ctxt3) = ctxt2 |> Variable.add_fixes ["x"]; val ([y1, y2], ctxt4) = ctxt3 |> Variable.variant_fixes ["y", "y"];\ end text \ In this situation \<^ML>\Variable.add_fixes\ and \<^ML>\Variable.variant_fixes\ are very similar, but identical name proposals given in a row are only accepted by the second version. \ section \Assumptions \label{sec:assumptions}\ text \ An \<^emph>\assumption\ is a proposition that it is postulated in the current context. Local conclusions may use assumptions as additional facts, but this imposes implicit hypotheses that weaken the overall statement. Assumptions are restricted to fixed non-schematic statements, i.e.\ all generality needs to be expressed by explicit quantifiers. Nevertheless, the result will be in HHF normal form with outermost quantifiers stripped. For example, by assuming \\x :: \. P x\ we get \\x :: \. P x \ P ?x\ for schematic \?x\ of fixed type \\\. Local derivations accumulate more and more explicit references to hypotheses: \A\<^sub>1, \, A\<^sub>n \ B\ where \A\<^sub>1, \, A\<^sub>n\ needs to be covered by the assumptions of the current context. \<^medskip> The \add_assms\ operation augments the context by local assumptions, which are parameterized by an arbitrary \export\ rule (see below). The \export\ operation moves facts from a (larger) inner context into a (smaller) outer context, by discharging the difference of the assumptions as specified by the associated export rules. Note that the discharged portion is determined by the difference of contexts, not the facts being exported! There is a separate flag to indicate a goal context, where the result is meant to refine an enclosing sub-goal of a structured proof state. \<^medskip> The most basic export rule discharges assumptions directly by means of the \\\ introduction rule: \[ \infer[(\\\intro\)]{\\ - A \ A \ B\}{\\ \ B\} \] The variant for goal refinements marks the newly introduced premises, which causes the canonical Isar goal refinement scheme to enforce unification with local premises within the goal: \[ \infer[(\#\\intro\)]{\\ - A \ #A \ B\}{\\ \ B\} \] \<^medskip> Alternative versions of assumptions may perform arbitrary transformations on export, as long as the corresponding portion of hypotheses is removed from the given facts. For example, a local definition works by fixing \x\ and assuming \x \ t\, with the following export rule to reverse the effect: \[ \infer[(\\\expand\)]{\\ - (x \ t) \ B t\}{\\ \ B x\} \] This works, because the assumption \x \ t\ was introduced in a context with \x\ being fresh, so \x\ does not occur in \\\ here. \ text %mlref \ \begin{mldecls} @{define_ML_type Assumption.export} \\ @{define_ML Assumption.assume: "Proof.context -> cterm -> thm"} \\ @{define_ML Assumption.add_assms: "Assumption.export -> cterm list -> Proof.context -> thm list * Proof.context"} \\ @{define_ML Assumption.add_assumes: " cterm list -> Proof.context -> thm list * Proof.context"} \\ @{define_ML Assumption.export: "bool -> Proof.context -> Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Assumption.export\ represents arbitrary export rules, which is any function of type \<^ML_type>\bool -> cterm list -> thm -> thm\, where the \<^ML_type>\bool\ indicates goal mode, and the \<^ML_type>\cterm list\ the collection of assumptions to be discharged simultaneously. \<^descr> \<^ML>\Assumption.assume\~\ctxt A\ turns proposition \A\ into a primitive assumption \A \ A'\, where the conclusion \A'\ is in HHF normal form. \<^descr> \<^ML>\Assumption.add_assms\~\r As\ augments the context by assumptions \As\ with export rule \r\. The resulting facts are hypothetical theorems as produced by the raw \<^ML>\Assumption.assume\. \<^descr> \<^ML>\Assumption.add_assumes\~\As\ is a special case of \<^ML>\Assumption.add_assms\ where the export rule performs \\\intro\ or \#\\intro\, depending on goal mode. \<^descr> \<^ML>\Assumption.export\~\is_goal inner outer thm\ exports result \thm\ from the the \inner\ context back into the \outer\ one; \is_goal = true\ means this is a goal context. The result is in HHF normal form. Note that \<^ML>\Proof_Context.export\ combines \<^ML>\Variable.export\ and \<^ML>\Assumption.export\ in the canonical way. \ text %mlex \ The following example demonstrates how rules can be derived by building up a context of assumptions first, and exporting some local fact afterwards. We refer to \<^theory>\Pure\ equality here for testing purposes. \ ML_val \ (*static compile-time context -- for testing only*) val ctxt0 = \<^context>; val ([eq], ctxt1) = ctxt0 |> Assumption.add_assumes [\<^cprop>\x \ y\]; val eq' = Thm.symmetric eq; (*back to original context -- discharges assumption*) val r = Assumption.export false ctxt1 ctxt0 eq'; \ text \ Note that the variables of the resulting rule are not generalized. This would have required to fix them properly in the context beforehand, and export wrt.\ variables afterwards (cf.\ \<^ML>\Variable.export\ or the combined \<^ML>\Proof_Context.export\). \ section \Structured goals and results \label{sec:struct-goals}\ text \ Local results are established by monotonic reasoning from facts within a context. This allows common combinations of theorems, e.g.\ via \\/\\ elimination, resolution rules, or equational reasoning, see \secref{sec:thms}. Unaccounted context manipulations should be avoided, notably raw \\/\\ introduction or ad-hoc references to free variables or assumptions not present in the proof context. \<^medskip> The \SUBPROOF\ combinator allows to structure a tactical proof recursively by decomposing a selected sub-goal: \(\x. A(x) \ B(x)) \ \\ is turned into \B(x) \ \\ after fixing \x\ and assuming \A(x)\. This means the tactic needs to solve the conclusion, but may use the premise as a local fact, for locally fixed variables. The family of \FOCUS\ combinators is similar to \SUBPROOF\, but allows to retain schematic variables and pending subgoals in the resulting goal state. The \prove\ operation provides an interface for structured backwards reasoning under program control, with some explicit sanity checks of the result. The goal context can be augmented by additional fixed variables (cf.\ \secref{sec:variables}) and assumptions (cf.\ \secref{sec:assumptions}), which will be available as local facts during the proof and discharged into implications in the result. Type and term variables are generalized as usual, according to the context. The \obtain\ operation produces results by eliminating existing facts by means of a given tactic. This acts like a dual conclusion: the proof demonstrates that the context may be augmented by parameters and assumptions, without affecting any conclusions that do not mention these parameters. See also @{cite "isabelle-isar-ref"} for the user-level @{command obtain} and @{command guess} elements. Final results, which may not refer to the parameters in the conclusion, need to exported explicitly into the original context.\ text %mlref \ \begin{mldecls} @{define_ML SUBPROOF: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS_PREMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.FOCUS_PARAMS: "(Subgoal.focus -> tactic) -> Proof.context -> int -> tactic"} \\ @{define_ML Subgoal.focus: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ @{define_ML Subgoal.focus_prems: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ @{define_ML Subgoal.focus_params: "Proof.context -> int -> binding list option -> thm -> Subgoal.focus * thm"} \\ \end{mldecls} \begin{mldecls} @{define_ML Goal.prove: "Proof.context -> string list -> term list -> term -> ({prems: thm list, context: Proof.context} -> tactic) -> thm"} \\ @{define_ML Goal.prove_common: "Proof.context -> int option -> string list -> term list -> term list -> ({prems: thm list, context: Proof.context} -> tactic) -> thm list"} \\ \end{mldecls} \begin{mldecls} @{define_ML Obtain.result: "(Proof.context -> tactic) -> thm list -> Proof.context -> ((string * cterm) list * thm list) * Proof.context"} \\ \end{mldecls} \<^descr> \<^ML>\SUBPROOF\~\tac ctxt i\ decomposes the structure of the specified sub-goal, producing an extended context and a reduced goal, which needs to be solved by the given tactic. All schematic parameters of the goal are imported into the context as fixed ones, which may not be instantiated in the sub-proof. \<^descr> \<^ML>\Subgoal.FOCUS\, \<^ML>\Subgoal.FOCUS_PREMS\, and \<^ML>\Subgoal.FOCUS_PARAMS\ are similar to \<^ML>\SUBPROOF\, but are slightly more flexible: only the specified parts of the subgoal are imported into the context, and the body tactic may introduce new subgoals and schematic variables. \<^descr> \<^ML>\Subgoal.focus\, \<^ML>\Subgoal.focus_prems\, \<^ML>\Subgoal.focus_params\ extract the focus information from a goal state in the same way as the corresponding tacticals above. This is occasionally useful to experiment without writing actual tactics yet. \<^descr> \<^ML>\Goal.prove\~\ctxt xs As C tac\ states goal \C\ in the context augmented by fixed variables \xs\ and assumptions \As\, and applies tactic \tac\ to solve it. The latter may depend on the local assumptions being presented as facts. The result is in HHF normal form. \<^descr> \<^ML>\Goal.prove_common\~\ctxt fork_pri\ is the common form to state and prove a simultaneous goal statement, where \<^ML>\Goal.prove\ is a convenient shorthand that is most frequently used in applications. The given list of simultaneous conclusions is encoded in the goal state by means of Pure conjunction: \<^ML>\Goal.conjunction_tac\ will turn this into a collection of individual subgoals, but note that the original multi-goal state is usually required for advanced induction. It is possible to provide an optional priority for a forked proof, typically \<^ML>\SOME ~1\, while \<^ML>\NONE\ means the proof is immediate (sequential) as for \<^ML>\Goal.prove\. Note that a forked proof does not exhibit any failures in the usual way via exceptions in ML, but accumulates error situations under the execution id of the running transaction. Thus the system is able to expose error messages ultimately to the end-user, even though the subsequent ML code misses them. \<^descr> \<^ML>\Obtain.result\~\tac thms ctxt\ eliminates the given facts using a tactic, which results in additional fixed variables and assumptions in the context. Final results need to be exported explicitly. \ text %mlex \ The following minimal example illustrates how to access the focus information of a structured goal state. \ notepad begin fix A B C :: "'a \ bool" have "\x. A x \ B x \ C x" ML_val \val {goal, context = goal_ctxt, ...} = @{Isar.goal}; val (focus as {params, asms, concl, ...}, goal') = Subgoal.focus goal_ctxt 1 (SOME [\<^binding>\x\]) goal; val [A, B] = #prems focus; val [(_, x)] = #params focus;\ sorry end text \ \<^medskip> The next example demonstrates forward-elimination in a local context, using \<^ML>\Obtain.result\. \ notepad begin assume ex: "\x. B x" ML_prf %"ML" \val ctxt0 = \<^context>; val (([(_, x)], [B]), ctxt1) = ctxt0 |> Obtain.result (fn _ => eresolve_tac ctxt0 @{thms exE} 1) [@{thm ex}];\ ML_prf %"ML" \singleton (Proof_Context.export ctxt1 ctxt0) @{thm refl};\ ML_prf %"ML" \Proof_Context.export ctxt1 ctxt0 [Thm.reflexive x] handle ERROR msg => (warning msg; []);\ end 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 (Symtab.empty, Symtab.make_set (map fst abs_nms)) + |> 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) |> 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 = - Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt'' t)) inst []; + 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 (schematic_types, schematic_terms') = schematics; val schematics' = (schematic_types, 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/Library/cconv.ML b/src/HOL/Library/cconv.ML --- a/src/HOL/Library/cconv.ML +++ b/src/HOL/Library/cconv.ML @@ -1,256 +1,256 @@ (* 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 (Symtab.empty, Symtab.make_set [u]) eq]) + [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 (*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 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 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/code_lazy.ML b/src/HOL/Library/code_lazy.ML --- a/src/HOL/Library/code_lazy.ML +++ b/src/HOL/Library/code_lazy.ML @@ -1,662 +1,662 @@ (* Author: Pascal Stoop, ETH Zurich Author: Andreas Lochbihler, Digital Asset *) signature CODE_LAZY = sig type lazy_info = {eagerT: typ, lazyT: typ, ctr: term, destr: term, lazy_ctrs: term list, case_lazy: term, active: bool, activate: theory -> theory, deactivate: theory -> theory}; val code_lazy_type: string -> theory -> theory val activate_lazy_type: string -> theory -> theory val deactivate_lazy_type: string -> theory -> theory val activate_lazy_types: theory -> theory val deactivate_lazy_types: theory -> theory val get_lazy_types: theory -> (string * lazy_info) list val print_lazy_types: theory -> unit val transform_code_eqs: Proof.context -> (thm * bool) list -> (thm * bool) list option end; structure Code_Lazy : CODE_LAZY = struct type lazy_info = {eagerT: typ, lazyT: typ, ctr: term, destr: term, lazy_ctrs: term list, case_lazy: term, active: bool, activate: theory -> theory, deactivate: theory -> theory}; fun map_active f {eagerT, lazyT, ctr, destr, lazy_ctrs, case_lazy, active, activate, deactivate} = { eagerT = eagerT, lazyT = lazyT, ctr = ctr, destr = destr, lazy_ctrs = lazy_ctrs, case_lazy = case_lazy, active = f active, activate = activate, deactivate = deactivate } structure Laziness_Data = Theory_Data( type T = lazy_info Symtab.table; val empty = Symtab.empty; val merge = Symtab.join (fn _ => fn (_, record) => record); val extend = I; ); fun fold_type type' tfree tvar typ = let fun go (Type (s, Ts)) = type' (s, map go Ts) | go (TFree T) = tfree T | go (TVar T) = tvar T in go typ end; fun read_typ lthy name = let val (s, Ts) = Proof_Context.read_type_name {proper = true, strict = true} lthy name |> dest_Type val (Ts', _) = Ctr_Sugar_Util.mk_TFrees (length Ts) lthy in Type (s, Ts') end fun mk_name prefix suffix name ctxt = Ctr_Sugar_Util.mk_fresh_names ctxt 1 (prefix ^ name ^ suffix) |>> hd; fun generate_typedef_name name ctxt = mk_name "" "_lazy" name ctxt; fun add_syntax_definition short_type_name eagerT lazyT lazy_ctr lthy = let val (name, _) = mk_name "lazy_" "" short_type_name lthy val freeT = HOLogic.unitT --> lazyT val lazyT' = Type (\<^type_name>\lazy\, [lazyT]) val def = Logic.all_const freeT $ absdummy freeT (Logic.mk_equals ( Free (name, (freeT --> eagerT)) $ Bound 0, lazy_ctr $ (Const (\<^const_name>\delay\, (freeT --> lazyT')) $ Bound 0))) val lthy' = (snd o Local_Theory.begin_nested) lthy val ((t, (_, thm)), lthy') = Specification.definition NONE [] [] ((Thm.def_binding (Binding.name name), []), def) lthy' val lthy' = Specification.notation true ("", false) [(t, Mixfix.mixfix "_")] lthy' val lthy = Local_Theory.end_nested lthy' val def_thm = singleton (Proof_Context.export lthy' lthy) thm in (def_thm, lthy) end; fun add_ctr_code raw_ctrs case_thms thy = let fun mk_case_certificate ctxt raw_thms = let val thms = raw_thms |> Conjunction.intr_balanced |> Thm.unvarify_global (Proof_Context.theory_of ctxt) |> Conjunction.elim_balanced (length raw_thms) |> map Simpdata.mk_meta_eq |> map Drule.zero_var_indexes val thm1 = case thms of thm :: _ => thm | _ => raise Empty val params = Term.add_free_names (Thm.prop_of thm1) []; val rhs = thm1 |> Thm.prop_of |> Logic.dest_equals |> fst |> strip_comb ||> fst o split_last |> list_comb val lhs = Free (singleton (Name.variant_list params) "case", fastype_of rhs); val assum = Thm.cterm_of ctxt (Logic.mk_equals (lhs, rhs)) in thms |> Conjunction.intr_balanced |> rewrite_rule ctxt [Thm.symmetric (Thm.assume assum)] |> Thm.implies_intr assum - |> Thm.generalize (Symtab.empty, Symtab.make_set params) 0 + |> Thm.generalize (Names.empty, Names.make_set params) 0 |> Axclass.unoverload ctxt |> Thm.varifyT_global end val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs in if can (Code.constrset_of_consts thy) unover_ctrs then thy |> Code.declare_datatype_global ctrs |> fold_rev (Code.add_eqn_global o rpair true) case_thms |> Code.declare_case_global (mk_case_certificate (Proof_Context.init_global thy) case_thms) else thy end; fun not_found s = error (s ^ " has not been added as lazy type"); fun validate_type_name thy type_name = let val lthy = Named_Target.theory_init thy val eager_type = read_typ lthy type_name val type_name = case eager_type of Type (s, _) => s | _ => raise Match in type_name end; fun set_active_lazy_type value eager_type_string thy = let val type_name = validate_type_name thy eager_type_string val x = case Symtab.lookup (Laziness_Data.get thy) type_name of NONE => not_found type_name | SOME x => x val new_x = map_active (K value) x val thy1 = if value = #active x then thy else if value then #activate x thy else #deactivate x thy in Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 end; fun set_active_lazy_types value thy = let val lazy_type_names = Symtab.keys (Laziness_Data.get thy) fun fold_fun value type_name thy = let val x = case Symtab.lookup (Laziness_Data.get thy) type_name of SOME x => x | NONE => raise Match val new_x = map_active (K value) x val thy1 = if value = #active x then thy else if value then #activate x thy else #deactivate x thy in Laziness_Data.map (Symtab.update (type_name, new_x)) thy1 end in fold (fold_fun value) lazy_type_names thy end; (* code_lazy_type : string -> theory -> theory *) fun code_lazy_type eager_type_string thy = let val lthy = Named_Target.theory_init thy val eagerT = read_typ lthy eager_type_string val (type_name, type_args) = dest_Type eagerT val short_type_name = Long_Name.base_name type_name val _ = if Symtab.defined (Laziness_Data.get thy) type_name then error (type_name ^ " has already been added as lazy type.") else () val {case_thms, casex, ctrs, ...} = case Ctr_Sugar.ctr_sugar_of lthy type_name of SOME x => x | _ => error (type_name ^ " is not registered with free constructors.") fun substitute_ctr (old_T, new_T) ctr_T lthy = let val old_ctr_vars = map TVar (Term.add_tvarsT ctr_T []) val old_ctr_Ts = map TFree (Term.add_tfreesT ctr_T []) @ old_ctr_vars val (new_ctr_Ts, lthy') = Ctr_Sugar_Util.mk_TFrees (length old_ctr_Ts) lthy fun double_type_fold Ts = case Ts of (Type (_, Ts1), Type (_, Ts2)) => flat (map double_type_fold (Ts1 ~~ Ts2)) | (Type _, _) => raise Match | (_, Type _) => raise Match | Ts => [Ts] fun map_fun1 f = List.foldr (fn ((T1, T2), f) => fn T => if T = T1 then T2 else f T) f (double_type_fold (old_T, new_T)) val map_fun2 = AList.lookup (op =) (old_ctr_Ts ~~ new_ctr_Ts) #> the val map_fun = map_fun1 map_fun2 val new_ctr_type = fold_type Type (map_fun o TFree) (map_fun o TVar) ctr_T in (new_ctr_type, lthy') end val (short_lazy_type_name, lthy1) = generate_typedef_name short_type_name lthy fun mk_lazy_typedef (name, eager_type) lthy = let val binding = Binding.name name val (result, lthy1) = (Typedef.add_typedef { overloaded=false } (binding, rev (Term.add_tfreesT eager_type []), Mixfix.NoSyn) (Const (\<^const_name>\top\, Type (\<^type_name>\set\, [eager_type]))) NONE (fn ctxt => resolve_tac ctxt [@{thm UNIV_witness}] 1) o (snd o Local_Theory.begin_nested)) lthy in (binding, result, Local_Theory.end_nested lthy1) end; val (typedef_binding, (_, info), lthy2) = mk_lazy_typedef (short_lazy_type_name, eagerT) lthy1 val lazy_type = Type (Local_Theory.full_name lthy2 typedef_binding, type_args) val (Abs_lazy, Rep_lazy) = let val info = fst info val Abs_name = #Abs_name info val Rep_name = #Rep_name info val Abs_type = eagerT --> lazy_type val Rep_type = lazy_type --> eagerT in (Const (Abs_name, Abs_type), Const (Rep_name, Rep_type)) end val Abs_inverse = #Abs_inverse (snd info) val Rep_inverse = #Rep_inverse (snd info) val (ctrs', lthy3) = List.foldr (fn (Const (s, T), (ctrs, lthy)) => let val (T', lthy') = substitute_ctr (body_type T, eagerT) T lthy in ((Const (s, T')) :: ctrs, lthy') end ) ([], lthy2) ctrs fun to_destr_eagerT typ = case typ of Type (\<^type_name>\fun\, [_, Type (\<^type_name>\fun\, Ts)]) => to_destr_eagerT (Type (\<^type_name>\fun\, Ts)) | Type (\<^type_name>\fun\, [T, _]) => T | _ => raise Match val (case', lthy4) = let val (eager_case, caseT) = dest_Const casex val (caseT', lthy') = substitute_ctr (to_destr_eagerT caseT, eagerT) caseT lthy3 in (Const (eager_case, caseT'), lthy') end val ctr_names = map (Long_Name.base_name o fst o dest_Const) ctrs val ((((lazy_ctr_name, lazy_sel_name), lazy_ctrs_name), lazy_case_name), ctxt) = lthy4 |> mk_name "Lazy_" "" short_type_name ||>> mk_name "unlazy_" "" short_type_name ||>> fold_map (mk_name "" "_Lazy") ctr_names ||>> mk_name "case_" "_lazy" short_type_name fun mk_def (name, T, rhs) lthy = let val binding = Binding.name name val ((_, (_, thm)), lthy1) = (snd o Local_Theory.begin_nested) lthy |> Specification.definition NONE [] [] ((Thm.def_binding binding, []), rhs) val lthy2 = Local_Theory.end_nested lthy1 val def_thm = hd (Proof_Context.export lthy1 lthy2 [thm]) in ({binding = binding, const = Const (Local_Theory.full_name lthy2 binding, T), thm = def_thm}, lthy2) end; val lazy_datatype = Type (\<^type_name>\lazy\, [lazy_type]) val Lazy_type = lazy_datatype --> eagerT val unstr_type = eagerT --> lazy_datatype fun apply_bounds i n term = if n > i then apply_bounds i (n-1) (term $ Bound (n-1)) else term fun all_abs Ts t = Logic.list_all (map (pair Name.uu) Ts, t) fun mk_force t = Const (\<^const_name>\force\, lazy_datatype --> lazy_type) $ t fun mk_delay t = Const (\<^const_name>\delay\, (\<^typ>\unit\ --> lazy_type) --> lazy_datatype) $ t val lazy_ctr = all_abs [lazy_datatype] (Logic.mk_equals (Free (lazy_ctr_name, Lazy_type) $ Bound 0, Rep_lazy $ mk_force (Bound 0))) val (lazy_ctr_def, lthy5) = mk_def (lazy_ctr_name, Lazy_type, lazy_ctr) lthy4 val lazy_sel = all_abs [eagerT] (Logic.mk_equals (Free (lazy_sel_name, unstr_type) $ Bound 0, mk_delay (Abs (Name.uu, \<^typ>\unit\, Abs_lazy $ Bound 1)))) val (lazy_sel_def, lthy6) = mk_def (lazy_sel_name, unstr_type, lazy_sel) lthy5 fun mk_lazy_ctr (name, eager_ctr) = let val (_, ctrT) = dest_Const eager_ctr fun to_lazy_ctrT (Type (\<^type_name>\fun\, [T1, T2])) = T1 --> to_lazy_ctrT T2 | to_lazy_ctrT T = if T = eagerT then lazy_type else raise Match val lazy_ctrT = to_lazy_ctrT ctrT val argsT = binder_types ctrT val lhs = apply_bounds 0 (length argsT) (Free (name, lazy_ctrT)) val rhs = Abs_lazy $ apply_bounds 0 (length argsT) eager_ctr in mk_def (name, lazy_ctrT, all_abs argsT (Logic.mk_equals (lhs, rhs))) end val (lazy_ctrs_def, lthy7) = fold_map mk_lazy_ctr (lazy_ctrs_name ~~ ctrs') lthy6 val (lazy_case_def, lthy8) = let val (_, caseT) = dest_Const case' fun to_lazy_caseT (Type (\<^type_name>\fun\, [T1, T2])) = if T1 = eagerT then lazy_type --> T2 else T1 --> to_lazy_caseT T2 val lazy_caseT = to_lazy_caseT caseT val argsT = binder_types lazy_caseT val n = length argsT val lhs = apply_bounds 0 n (Free (lazy_case_name, lazy_caseT)) val rhs = apply_bounds 1 n case' $ (Rep_lazy $ Bound 0) in mk_def (lazy_case_name, lazy_caseT, all_abs argsT (Logic.mk_equals (lhs, rhs))) lthy7 end fun mk_thm ((name, term), thms) lthy = let val binding = Binding.name name fun tac {context, ...} = Simplifier.simp_tac (put_simpset HOL_basic_ss context addsimps thms) 1 val thm = Goal.prove lthy [] [] term tac val (_, lthy1) = lthy |> (snd o Local_Theory.begin_nested) |> Local_Theory.note ((binding, []), [thm]) in (thm, Local_Theory.end_nested lthy1) end fun mk_thms exec_list lthy = fold_map mk_thm exec_list lthy val mk_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq val lazy_ctrs = map #const lazy_ctrs_def val eager_lazy_ctrs = ctrs' ~~ lazy_ctrs val (((((((Lazy_delay_eq_name, Rep_ctr_names), ctrs_lazy_names), force_sel_name), case_lazy_name), sel_lazy_name), case_ctrs_name), _) = lthy5 |> mk_name "Lazy_" "_delay" short_type_name ||>> fold_map (mk_name "Rep_" "_Lazy") ctr_names ||>> fold_map (mk_name "" "_conv_lazy") ctr_names ||>> mk_name "force_unlazy_" "" short_type_name ||>> mk_name "case_" "_conv_lazy" short_type_name ||>> mk_name "Lazy_" "_inverse" short_type_name ||>> fold_map (mk_name ("case_" ^ short_type_name ^ "_lazy_") "") ctr_names val mk_Lazy_delay_eq = (#const lazy_ctr_def $ mk_delay (Bound 0), Rep_lazy $ (Bound 0 $ \<^const>\Unity\)) |> mk_eq |> all_abs [\<^typ>\unit\ --> lazy_type] val (Lazy_delay_thm, lthy8a) = mk_thm ((Lazy_delay_eq_name, mk_Lazy_delay_eq), [#thm lazy_ctr_def, @{thm force_delay}]) lthy8 fun mk_lazy_ctr_eq (eager_ctr, lazy_ctr) = let val (_, ctrT) = dest_Const eager_ctr val argsT = binder_types ctrT val args = length argsT in (Rep_lazy $ apply_bounds 0 args lazy_ctr, apply_bounds 0 args eager_ctr) |> mk_eq |> all_abs argsT end val Rep_ctr_eqs = map mk_lazy_ctr_eq eager_lazy_ctrs val (Rep_ctr_thms, lthy8b) = mk_thms ((Rep_ctr_names ~~ Rep_ctr_eqs) ~~ (map (fn def => [#thm def, Abs_inverse, @{thm UNIV_I}]) lazy_ctrs_def) ) lthy8a fun mk_ctrs_lazy_eq (eager_ctr, lazy_ctr) = let val argsT = dest_Const eager_ctr |> snd |> binder_types val n = length argsT val lhs = apply_bounds 0 n eager_ctr val rhs = #const lazy_ctr_def $ (mk_delay (Abs (Name.uu, \<^typ>\unit\, apply_bounds 1 (n + 1) lazy_ctr))) in (lhs, rhs) |> mk_eq |> all_abs argsT end val ctrs_lazy_eq = map mk_ctrs_lazy_eq eager_lazy_ctrs val (ctrs_lazy_thms, lthy8c) = mk_thms ((ctrs_lazy_names ~~ ctrs_lazy_eq) ~~ map (fn thm => [Lazy_delay_thm, thm]) Rep_ctr_thms) lthy8b val force_sel_eq = (mk_force (#const lazy_sel_def $ Bound 0), Abs_lazy $ Bound 0) |> mk_eq |> all_abs [eagerT] val (force_sel_thm, lthy8d) = mk_thm ((force_sel_name, force_sel_eq), [#thm lazy_sel_def, @{thm force_delay}]) lthy8c val case_lazy_eq = let val (_, caseT) = case' |> dest_Const val argsT = binder_types caseT val n = length argsT val lhs = apply_bounds 0 n case' val rhs = apply_bounds 1 n (#const lazy_case_def) $ (mk_force (#const lazy_sel_def $ Bound 0)) in (lhs, rhs) |> mk_eq |> all_abs argsT end val (case_lazy_thm, lthy8e) = mk_thm ((case_lazy_name, case_lazy_eq), [#thm lazy_case_def, force_sel_thm, Abs_inverse, @{thm UNIV_I}]) lthy8d val sel_lazy_eq = (#const lazy_sel_def $ (#const lazy_ctr_def $ Bound 0), Bound 0) |> mk_eq |> all_abs [lazy_datatype] val (sel_lazy_thm, lthy8f) = mk_thm ((sel_lazy_name, sel_lazy_eq), [#thm lazy_sel_def, #thm lazy_ctr_def, Rep_inverse, @{thm delay_force}]) lthy8e fun mk_case_ctrs_eq (i, lazy_ctr) = let val lazy_case = #const lazy_case_def val (_, ctrT) = dest_Const lazy_ctr val ctr_argsT = binder_types ctrT val ctr_args_n = length ctr_argsT val (_, caseT) = dest_Const lazy_case val case_argsT = binder_types caseT fun n_bounds_from m n t = if n > 0 then n_bounds_from (m - 1) (n - 1) (t $ Bound (m - 1)) else t val case_argsT' = take (length case_argsT - 1) case_argsT val Ts = case_argsT' @ ctr_argsT val num_abs_types = length Ts val lhs = n_bounds_from num_abs_types (length case_argsT') lazy_case $ apply_bounds 0 ctr_args_n lazy_ctr val rhs = apply_bounds 0 ctr_args_n (Bound (num_abs_types - i - 1)) in (lhs, rhs) |> mk_eq |> all_abs Ts end val case_ctrs_eq = map_index mk_case_ctrs_eq lazy_ctrs val (case_ctrs_thms, lthy9) = mk_thms ((case_ctrs_name ~~ case_ctrs_eq) ~~ map2 (fn thm1 => fn thm2 => [#thm lazy_case_def, thm1, thm2]) Rep_ctr_thms case_thms ) lthy8f val (pat_def_thm, lthy10) = add_syntax_definition short_type_name eagerT lazy_type (#const lazy_ctr_def) lthy9 val add_lazy_ctrs = Code.declare_datatype_global [dest_Const (#const lazy_ctr_def)] val eager_ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global)) o dest_Const) ctrs val add_eager_ctrs = fold Code.del_eqn_global ctrs_lazy_thms #> Code.declare_datatype_global eager_ctrs val add_code_eqs = fold (Code.add_eqn_global o rpair true) ([case_lazy_thm, sel_lazy_thm]) val add_lazy_ctr_thms = fold (Code.add_eqn_global o rpair true) ctrs_lazy_thms val add_lazy_case_thms = fold Code.del_eqn_global case_thms #> Code.add_eqn_global (case_lazy_thm, true) val add_eager_case_thms = Code.del_eqn_global case_lazy_thm #> fold (Code.add_eqn_global o rpair true) case_thms val delay_dummy_thm = (pat_def_thm RS @{thm symmetric}) |> Drule.infer_instantiate' lthy10 [SOME (Thm.cterm_of lthy10 (Const (\<^const_name>\Pure.dummy_pattern\, HOLogic.unitT --> lazy_type)))] |> Thm.generalize - (Symtab.make_set (map (fst o dest_TFree) type_args), Symtab.empty) + (Names.make_set (map (fst o dest_TFree) type_args), Names.empty) (Variable.maxidx_of lthy10 + 1); val ctr_post = delay_dummy_thm :: map (fn thm => thm RS @{thm sym}) ctrs_lazy_thms val ctr_thms_abs = map (fn thm => Drule.abs_def (thm RS @{thm eq_reflection})) ctrs_lazy_thms val case_thm_abs = Drule.abs_def (case_lazy_thm RS @{thm eq_reflection}) val add_simps = Code_Preproc.map_pre (fn ctxt => ctxt addsimps (case_thm_abs :: ctr_thms_abs)) val del_simps = Code_Preproc.map_pre (fn ctxt => ctxt delsimps (case_thm_abs :: ctr_thms_abs)) val add_post = Code_Preproc.map_post (fn ctxt => ctxt addsimps ctr_post) val del_post = Code_Preproc.map_post (fn ctxt => ctxt delsimps ctr_post) in Local_Theory.exit_global lthy10 |> Laziness_Data.map (Symtab.update (type_name, {eagerT = eagerT, lazyT = lazy_type, ctr = #const lazy_ctr_def, destr = #const lazy_sel_def, lazy_ctrs = map #const lazy_ctrs_def, case_lazy = #const lazy_case_def, active = true, activate = add_lazy_ctrs #> add_lazy_ctr_thms #> add_lazy_case_thms #> add_simps #> add_post, deactivate = add_eager_ctrs #> add_eager_case_thms #> del_simps #> del_post})) |> add_lazy_ctrs |> add_ctr_code (map (dest_Const o #const) lazy_ctrs_def) case_ctrs_thms |> add_code_eqs |> add_lazy_ctr_thms |> add_simps |> add_post end; fun transform_code_eqs _ [] = NONE | transform_code_eqs ctxt eqs = let fun replace_ctr ctxt = let val thy = Proof_Context.theory_of ctxt val table = Laziness_Data.get thy in fn (s1, s2) => case Symtab.lookup table s1 of NONE => false | SOME x => #active x andalso s2 <> (#ctr x |> dest_Const |> fst) end val thy = Proof_Context.theory_of ctxt val table = Laziness_Data.get thy fun num_consts_fun (_, T) = let val s = body_type T |> dest_Type |> fst in if Symtab.defined table s then Ctr_Sugar.ctr_sugar_of ctxt s |> the |> #ctrs |> length else Code.get_type thy s |> fst |> snd |> length end val eqs = map (apfst (Thm.transfer thy)) eqs; val ((code_eqs, nbe_eqs), pure) = ((case hd eqs |> fst |> Thm.prop_of of Const (\<^const_name>\Pure.eq\, _) $ _ $ _ => (map (apfst (fn x => x RS @{thm meta_eq_to_obj_eq})) eqs, true) | _ => (eqs, false)) |> apfst (List.partition snd)) handle THM _ => (([], eqs), false) val to_original_eq = if pure then map (apfst (fn x => x RS @{thm eq_reflection})) else I in case Case_Converter.to_case ctxt (Case_Converter.replace_by_type replace_ctr) num_consts_fun (map fst code_eqs) of NONE => NONE | SOME thms => SOME (nbe_eqs @ map (rpair true) thms |> to_original_eq) end val activate_lazy_type = set_active_lazy_type true; val deactivate_lazy_type = set_active_lazy_type false; val activate_lazy_types = set_active_lazy_types true; val deactivate_lazy_types = set_active_lazy_types false; fun get_lazy_types thy = Symtab.dest (Laziness_Data.get thy) fun print_lazy_type thy (name, info : lazy_info) = let val ctxt = Proof_Context.init_global thy fun pretty_ctr ctr = let val argsT = dest_Const ctr |> snd |> binder_types in Pretty.block [ Syntax.pretty_term ctxt ctr, Pretty.brk 1, Pretty.block (Pretty.separate "" (map (Pretty.quote o Syntax.pretty_typ ctxt) argsT)) ] end in Pretty.block [ Pretty.str (name ^ (if #active info then "" else " (inactive)") ^ ":"), Pretty.brk 1, Pretty.block [ Syntax.pretty_typ ctxt (#eagerT info), Pretty.brk 1, Pretty.str "=", Pretty.brk 1, Syntax.pretty_term ctxt (#ctr info), Pretty.brk 1, Pretty.block [ Pretty.str "(", Syntax.pretty_term ctxt (#destr info), Pretty.str ":", Pretty.brk 1, Syntax.pretty_typ ctxt (Type (\<^type_name>\lazy\, [#lazyT info])), Pretty.str ")" ] ], Pretty.fbrk, Pretty.keyword2 "and", Pretty.brk 1, Pretty.block ([ Syntax.pretty_typ ctxt (#lazyT info), Pretty.brk 1, Pretty.str "=", Pretty.brk 1] @ Pretty.separate " |" (map pretty_ctr (#lazy_ctrs info)) @ [ Pretty.fbrk, Pretty.keyword2 "for", Pretty.brk 1, Pretty.str "case:", Pretty.brk 1, Syntax.pretty_term ctxt (#case_lazy info) ]) ] end; fun print_lazy_types thy = let fun cmp ((name1, _), (name2, _)) = string_ord (name1, name2) val infos = Laziness_Data.get thy |> Symtab.dest |> map (apfst Long_Name.base_name) |> sort cmp in Pretty.writeln_chunks (map (print_lazy_type thy) infos) end; val _ = Outer_Syntax.command \<^command_keyword>\code_lazy_type\ "make a lazy copy of the datatype and activate substitution" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> code_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\activate_lazy_type\ "activate substitution on a specific (lazy) type" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> activate_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\deactivate_lazy_type\ "deactivate substitution on a specific (lazy) type" (Parse.binding >> (fn b => Toplevel.theory (Binding.name_of b |> deactivate_lazy_type))); val _ = Outer_Syntax.command \<^command_keyword>\activate_lazy_types\ "activate substitution on all (lazy) types" (pair (Toplevel.theory activate_lazy_types)); val _ = Outer_Syntax.command \<^command_keyword>\deactivate_lazy_types\ "deactivate substitution on all (lazy) type" (pair (Toplevel.theory deactivate_lazy_types)); val _ = Outer_Syntax.command \<^command_keyword>\print_lazy_types\ "print the types that have been declared as lazy and their substitution state" (pair (Toplevel.theory (tap print_lazy_types))); end \ No newline at end of file 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 Term_Subst.TVars.table -> 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 = - Term_Subst.TVars.table (map (apfst dest_TVar) type_pairing) + 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 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/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 (Symtab.empty, Symtab.make_set [y_s]) + |> 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 (Symtab.empty, Symtab.make_set [y_s, z_s]) + |> 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 |> 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/BNF/bnf_gfp_rec_sugar_tactics.ML b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML --- a/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML +++ b/src/HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML @@ -1,216 +1,216 @@ (* Title: HOL/Tools/BNF/bnf_gfp_rec_sugar_tactics.ML Author: Jasmin Blanchette, TU Muenchen Copyright 2013 Tactics for corecursor sugar. *) signature BNF_GFP_REC_SUGAR_TACTICS = sig val mk_primcorec_assumption_tac: Proof.context -> thm list -> int -> tactic val mk_primcorec_code_tac: Proof.context -> thm list -> thm list -> thm -> tactic val mk_primcorec_ctr_tac: Proof.context -> int -> thm -> thm option -> thm list -> tactic val mk_primcorec_disc_tac: Proof.context -> thm list -> thm -> int -> int -> thm list list list -> tactic val mk_primcorec_disc_iff_tac: Proof.context -> string list -> thm -> thm list -> thm list list -> thm list -> tactic val mk_primcorec_exhaust_tac: Proof.context -> string list -> int -> thm -> tactic val mk_primcorec_nchotomy_tac: Proof.context -> thm list -> tactic val mk_primcorec_raw_code_tac: Proof.context -> thm list -> thm list -> thm list -> thm list -> int list -> thm list -> thm option -> tactic val mk_primcorec_sel_tac: Proof.context -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm list -> thm -> int -> int -> thm list list list -> tactic end; structure BNF_GFP_Rec_Sugar_Tactics : BNF_GFP_REC_SUGAR_TACTICS = struct open BNF_Util open BNF_Tactics open BNF_FP_Util val atomize_conjL = @{thm atomize_conjL}; val falseEs = @{thms not_TrueE FalseE}; val neq_eq_eq_contradict = @{thm neq_eq_eq_contradict}; val if_split = @{thm if_split}; val if_split_asm = @{thm if_split_asm}; val split_connectI = @{thms allI impI conjI}; val unfold_lets = @{thms Let_def[abs_def] split_beta} fun exhaust_inst_as_projs ctxt frees thm = let val num_frees = length frees; val fs = Term.add_vars (Thm.prop_of thm) [] |> filter (can dest_funT o snd); fun find x = find_index (curry (op =) x) frees; fun mk_inst ((x, i), T) = ((x, i), Thm.cterm_of ctxt (mk_proj T num_frees (find x))); in infer_instantiate ctxt (map mk_inst fs) thm end; val exhaust_inst_as_projs_tac = PRIMITIVE oo exhaust_inst_as_projs; fun distinct_in_prems_tac ctxt distincts = eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' assume_tac ctxt; fun mk_primcorec_nchotomy_tac ctxt exhaust_discs = HEADGOAL (Method.insert_tac ctxt exhaust_discs THEN' clean_blast_tac ctxt); fun mk_primcorec_exhaust_tac ctxt frees n nchotomy = let val ks = 1 upto n in HEADGOAL (assume_tac ctxt ORELSE' cut_tac nchotomy THEN' K (exhaust_inst_as_projs_tac ctxt frees) THEN' EVERY' (map (fn k => (if k < n then etac ctxt disjE else K all_tac) THEN' REPEAT o (dtac ctxt meta_mp THEN' assume_tac ctxt ORELSE' etac ctxt conjE THEN' dtac ctxt meta_mp THEN' assume_tac ctxt ORELSE' assume_tac ctxt)) ks)) end; fun mk_primcorec_assumption_tac ctxt discIs = SELECT_GOAL (unfold_thms_tac ctxt @{thms fst_conv snd_conv not_not not_False_eq_True not_True_eq_False de_Morgan_conj de_Morgan_disj} THEN SOLVE (HEADGOAL (REPEAT o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' etac ctxt conjE ORELSE' eresolve_tac ctxt falseEs ORELSE' resolve_tac ctxt @{thms TrueI conjI disjI1 disjI2} ORELSE' dresolve_tac ctxt discIs THEN' assume_tac ctxt ORELSE' etac ctxt notE THEN' assume_tac ctxt ORELSE' etac ctxt disjE)))); fun ss_fst_snd_conv ctxt = Raw_Simplifier.simpset_of (ss_only @{thms fst_conv snd_conv} ctxt); fun case_atac ctxt = Simplifier.full_simp_tac (put_simpset (ss_fst_snd_conv ctxt) ctxt); fun same_case_tac ctxt m = HEADGOAL (if m = 0 then rtac ctxt TrueI else REPEAT_DETERM_N (m - 1) o (rtac ctxt conjI THEN' case_atac ctxt) THEN' case_atac ctxt); fun different_case_tac ctxt m exclude = HEADGOAL (if m = 0 then mk_primcorec_assumption_tac ctxt [] else dtac ctxt exclude THEN' (REPEAT_DETERM_N (m - 1) o case_atac ctxt) THEN' mk_primcorec_assumption_tac ctxt []); fun cases_tac ctxt k m excludesss = let val n = length excludesss in EVERY (map (fn [] => if k = n then all_tac else same_case_tac ctxt m | [exclude] => different_case_tac ctxt m exclude) (take k (nth excludesss (k - 1)))) end; fun prelude_tac ctxt fun_defs thm = unfold_thms_tac ctxt fun_defs THEN HEADGOAL (rtac ctxt thm) THEN unfold_thms_tac ctxt unfold_lets; fun mk_primcorec_disc_tac ctxt fun_defs corec_disc k m excludesss = prelude_tac ctxt fun_defs corec_disc THEN cases_tac ctxt k m excludesss; fun mk_primcorec_disc_iff_tac ctxt fun_exhaust_frees fun_exhaust fun_discs fun_discss distinct_discs = HEADGOAL (rtac ctxt iffI THEN' rtac ctxt fun_exhaust THEN' K (exhaust_inst_as_projs_tac ctxt fun_exhaust_frees) THEN' EVERY' (map (fn [] => etac ctxt FalseE | fun_discs' as [fun_disc'] => if eq_list Thm.eq_thm (fun_discs', fun_discs) then REPEAT_DETERM o etac ctxt conjI THEN' (assume_tac ctxt ORELSE' rtac ctxt TrueI) else rtac ctxt FalseE THEN' (rotate_tac 1 THEN' dtac ctxt fun_disc' THEN' REPEAT o assume_tac ctxt ORELSE' cut_tac fun_disc') THEN' dresolve_tac ctxt distinct_discs THEN' etac ctxt notE THEN' assume_tac ctxt) fun_discss) THEN' (etac ctxt FalseE ORELSE' resolve_tac ctxt (map (unfold_thms ctxt [atomize_conjL]) fun_discs) THEN_MAYBE' assume_tac ctxt)); fun mk_primcorec_sel_tac ctxt fun_defs distincts splits split_asms mapsx map_ident0s map_comps fun_sel k m excludesss = prelude_tac ctxt fun_defs (fun_sel RS trans) THEN cases_tac ctxt k m excludesss THEN HEADGOAL (REPEAT_DETERM o (rtac ctxt refl ORELSE' eresolve_tac ctxt falseEs ORELSE' resolve_tac ctxt split_connectI ORELSE' Splitter.split_asm_tac ctxt (if_split_asm :: split_asms) ORELSE' Splitter.split_tac ctxt (if_split :: splits) ORELSE' eresolve_tac ctxt (map (fn thm => thm RS neq_eq_eq_contradict) distincts) THEN' assume_tac ctxt ORELSE' etac ctxt notE THEN' assume_tac ctxt ORELSE' (CHANGED o SELECT_GOAL (unfold_thms_tac ctxt (@{thms fst_conv snd_conv id_def comp_def split_def sum.case sum.sel sum.distinct[THEN eq_False[THEN iffD2]]} @ mapsx @ map_ident0s @ map_comps))) ORELSE' fo_rtac ctxt @{thm cong} ORELSE' rtac ctxt @{thm ext} ORELSE' mk_primcorec_assumption_tac ctxt [])); fun mk_primcorec_ctr_tac ctxt m collapse disc_fun_opt sel_funs = HEADGOAL (rtac ctxt ((if null sel_funs then collapse else collapse RS sym) RS trans) THEN' (the_default (K all_tac) (Option.map (rtac ctxt) disc_fun_opt)) THEN' REPEAT_DETERM_N m o assume_tac ctxt) THEN unfold_thms_tac ctxt (@{thm split_def} :: unfold_lets @ sel_funs) THEN HEADGOAL (rtac ctxt refl); fun inst_split_eq ctxt split = (case Thm.prop_of split of \<^const>\Trueprop\ $ (Const (\<^const_name>\HOL.eq\, _) $ (Var (_, Type (_, [T, _])) $ _) $ _) => let val s = Name.uu; val eq = Abs (Name.uu, T, HOLogic.mk_eq (Free (s, T), Bound 0)); in Thm.instantiate' [] [SOME (Thm.cterm_of ctxt eq)] split - |> Drule.generalize (Symtab.empty, Symtab.make_set [s]) + |> Drule.generalize (Names.empty, Names.make_set [s]) end | _ => split); fun raw_code_single_tac ctxt distincts discIs splits split_asms m fun_ctr = let val splits' = map (fn th => th RS iffD2) (@{thm if_split_eq2} :: map (inst_split_eq ctxt) splits); in HEADGOAL (REPEAT o (resolve_tac ctxt (splits' @ split_connectI))) THEN prelude_tac ctxt [] (fun_ctr RS trans) THEN HEADGOAL ((REPEAT_DETERM_N m o mk_primcorec_assumption_tac ctxt discIs) THEN' SELECT_GOAL (SOLVE (HEADGOAL (REPEAT_DETERM o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' resolve_tac ctxt (@{thm Code.abort_def} :: split_connectI) ORELSE' Splitter.split_tac ctxt (if_split :: splits) ORELSE' Splitter.split_asm_tac ctxt (if_split_asm :: split_asms) ORELSE' mk_primcorec_assumption_tac ctxt discIs ORELSE' distinct_in_prems_tac ctxt distincts ORELSE' (TRY o dresolve_tac ctxt discIs) THEN' etac ctxt notE THEN' assume_tac ctxt))))) end; fun rulify_nchotomy n = funpow (n - 1) (fn thm => thm RS @{thm Meson.make_pos_rule'}); fun mk_primcorec_raw_code_tac ctxt distincts discIs splits split_asms ms fun_ctrs nchotomy_opt = let val n = length ms; val (ms', fun_ctrs') = (case nchotomy_opt of NONE => (ms, fun_ctrs) | SOME nchotomy => (ms |> split_last ||> K [n - 1] |> op @, fun_ctrs |> split_last ||> unfold_thms ctxt [atomize_conjL] ||> (fn thm => [rulify_nchotomy n nchotomy RS thm] handle THM _ => [thm]) |> op @)); in EVERY (map2 (raw_code_single_tac ctxt distincts discIs splits split_asms) ms' fun_ctrs') THEN IF_UNSOLVED (unfold_thms_tac ctxt @{thms Code.abort_def} THEN HEADGOAL (REPEAT_DETERM o resolve_tac ctxt (refl :: split_connectI))) end; fun mk_primcorec_code_tac ctxt distincts splits raw = HEADGOAL (rtac ctxt raw ORELSE' rtac ctxt (raw RS trans) THEN' SELECT_GOAL (unfold_thms_tac ctxt unfold_lets) THEN' REPEAT_DETERM o (rtac ctxt refl ORELSE' assume_tac ctxt ORELSE' resolve_tac ctxt split_connectI ORELSE' Splitter.split_tac ctxt (if_split :: splits) ORELSE' distinct_in_prems_tac ctxt distincts ORELSE' rtac ctxt sym THEN' assume_tac ctxt ORELSE' etac ctxt notE THEN' assume_tac ctxt)); end; diff --git a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML --- a/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML +++ b/src/HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML @@ -1,139 +1,139 @@ (* Title: HOL/Tools/Ctr_Sugar/ctr_sugar_code.ML Author: Jasmin Blanchette, TU Muenchen Author: Dmitriy Traytel, TU Muenchen Author: Stefan Berghofer, TU Muenchen Author: Florian Haftmann, TU Muenchen Copyright 2001-2013 Code generation for freely generated types. *) signature CTR_SUGAR_CODE = sig val add_ctr_code: string -> typ list -> (string * typ) list -> thm list -> thm list -> thm list -> theory -> theory end; structure Ctr_Sugar_Code : CTR_SUGAR_CODE = struct open Ctr_Sugar_Util val eqN = "eq" val reflN = "refl" val simpsN = "simps" fun mk_case_certificate ctxt raw_thms = let val thms as thm1 :: _ = raw_thms |> Conjunction.intr_balanced |> Thm.unvarify_global (Proof_Context.theory_of ctxt) |> Conjunction.elim_balanced (length raw_thms) |> map Simpdata.mk_meta_eq |> map Drule.zero_var_indexes; val params = Term.add_free_names (Thm.prop_of thm1) []; val rhs = thm1 |> Thm.prop_of |> Logic.dest_equals |> fst |> Term.strip_comb ||> fst o split_last |> list_comb; val lhs = Free (singleton (Name.variant_list params) "case", Term.fastype_of rhs); val assum = Thm.cterm_of ctxt (Logic.mk_equals (lhs, rhs)); in thms |> Conjunction.intr_balanced |> rewrite_rule ctxt [Thm.symmetric (Thm.assume assum)] |> Thm.implies_intr assum - |> Thm.generalize (Symtab.empty, Symtab.make_set params) 0 + |> Thm.generalize (Names.empty, Names.make_set params) 0 |> Axclass.unoverload ctxt |> Thm.varifyT_global end; fun mk_free_ctr_equations fcT ctrs inject_thms distinct_thms thy = let fun mk_fcT_eq (t, u) = Const (\<^const_name>\HOL.equal\, fcT --> fcT --> HOLogic.boolT) $ t $ u; fun true_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, \<^term>\True\); fun false_eq tu = HOLogic.mk_eq (mk_fcT_eq tu, \<^term>\False\); val monomorphic_prop_of = Thm.prop_of o Thm.unvarify_global thy o Drule.zero_var_indexes; fun massage_inject (tp $ (eqv $ (_ $ t $ u) $ rhs)) = tp $ (eqv $ mk_fcT_eq (t, u) $ rhs); fun massage_distinct (tp $ (_ $ (_ $ t $ u))) = [tp $ false_eq (t, u), tp $ false_eq (u, t)]; val triv_inject_goals = map_filter (fn c as (_, T) => if T = fcT then SOME (HOLogic.mk_Trueprop (true_eq (Const c, Const c))) else NONE) ctrs; val inject_goals = map (massage_inject o monomorphic_prop_of) inject_thms; val distinct_goals = maps (massage_distinct o monomorphic_prop_of) distinct_thms; val refl_goal = HOLogic.mk_Trueprop (true_eq (Free ("x", fcT), Free ("x", fcT))); fun prove goal = Goal.prove_sorry_global thy [] [] goal (fn {context = ctxt, ...} => HEADGOAL Goal.conjunction_tac THEN ALLGOALS (simp_tac (put_simpset HOL_basic_ss ctxt addsimps (map Simpdata.mk_eq (@{thms equal eq_True} @ inject_thms @ distinct_thms))))); fun proves goals = goals |> Logic.mk_conjunction_balanced |> prove |> Thm.close_derivation \<^here> |> Conjunction.elim_balanced (length goals) |> map Simpdata.mk_eq; in (proves (triv_inject_goals @ inject_goals @ distinct_goals), Simpdata.mk_eq (prove refl_goal)) end; fun add_equality fcT fcT_name As ctrs inject_thms distinct_thms = let fun add_def lthy = let fun mk_side const_name = Const (const_name, fcT --> fcT --> HOLogic.boolT) $ Free ("x", fcT) $ Free ("y", fcT); val spec = mk_Trueprop_eq (mk_side \<^const_name>\HOL.equal\, mk_side \<^const_name>\HOL.eq\) |> Syntax.check_term lthy; val ((_, (_, raw_def)), lthy') = Specification.definition NONE [] [] (Binding.empty_atts, spec) lthy; val thy_ctxt = Proof_Context.init_global (Proof_Context.theory_of lthy'); val def = singleton (Proof_Context.export lthy' thy_ctxt) raw_def; in (def, lthy') end; fun tac ctxt thms = Class.intro_classes_tac ctxt [] THEN ALLGOALS (Proof_Context.fact_tac ctxt thms); val qualify = Binding.qualify true (Long_Name.base_name fcT_name) o Binding.qualify true eqN o Binding.name; in Class.instantiation ([fcT_name], map dest_TFree As, [HOLogic.class_equal]) #> add_def #-> Class.prove_instantiation_exit_result (map o Morphism.thm) tac o single #> snd #> `(mk_free_ctr_equations fcT ctrs inject_thms distinct_thms) #-> (fn (thms, thm) => Global_Theory.note_thmss Thm.theoremK [((qualify reflN, []), [([thm], [])]), ((qualify simpsN, []), [(rev thms, [])])]) #-> (fn [(_, [thm]), (_, thms)] => Code.declare_default_eqns_global ((thm, false) :: map (rpair true) thms)) end; fun add_ctr_code fcT_name raw_As raw_ctrs inject_thms distinct_thms case_thms thy = let val As = map (perhaps (try Logic.unvarifyT_global)) raw_As; val ctrs = map (apsnd (perhaps (try Logic.unvarifyT_global))) raw_ctrs; val fcT = Type (fcT_name, As); val unover_ctrs = map (fn ctr as (_, fcT) => (Axclass.unoverload_const thy ctr, fcT)) ctrs; in if can (Code.constrset_of_consts thy) unover_ctrs then thy |> Code.declare_datatype_global ctrs |> Code.declare_default_eqns_global (map (rpair true) (rev case_thms)) |> Code.declare_case_global (mk_case_certificate (Proof_Context.init_global thy) case_thms) |> not (Sorts.has_instance (Sign.classes_of thy) fcT_name [HOLogic.class_equal]) ? add_equality fcT fcT_name As ctrs inject_thms distinct_thms else thy end; end; diff --git a/src/HOL/Tools/Function/induction_schema.ML b/src/HOL/Tools/Function/induction_schema.ML --- a/src/HOL/Tools/Function/induction_schema.ML +++ b/src/HOL/Tools/Function/induction_schema.ML @@ -1,400 +1,400 @@ (* Title: HOL/Tools/Function/induction_schema.ML Author: Alexander Krauss, TU Muenchen A method to prove induction schemas. *) signature INDUCTION_SCHEMA = sig val mk_ind_tac : (int -> tactic) -> (int -> tactic) -> (int -> tactic) -> Proof.context -> thm list -> tactic val induction_schema_tac : Proof.context -> thm list -> tactic end structure Induction_Schema : INDUCTION_SCHEMA = struct open Function_Lib type rec_call_info = int * (string * typ) list * term list * term list datatype scheme_case = SchemeCase of {bidx : int, qs: (string * typ) list, oqnames: string list, gs: term list, lhs: term list, rs: rec_call_info list} datatype scheme_branch = SchemeBranch of {P : term, xs: (string * typ) list, ws: (string * typ) list, Cs: term list} datatype ind_scheme = IndScheme of {T: typ, (* sum of products *) branches: scheme_branch list, cases: scheme_case list} fun ind_atomize ctxt = Raw_Simplifier.rewrite ctxt true @{thms induct_atomize} fun ind_rulify ctxt = Raw_Simplifier.rewrite ctxt true @{thms induct_rulify} fun meta thm = thm RS eq_reflection fun sum_prod_conv ctxt = Raw_Simplifier.rewrite ctxt true (map meta (@{thm split_conv} :: @{thms sum.case})) fun term_conv ctxt cv t = cv (Thm.cterm_of ctxt t) |> Thm.prop_of |> Logic.dest_equals |> snd fun mk_relT T = HOLogic.mk_setT (HOLogic.mk_prodT (T, T)) fun dest_hhf ctxt t = let val ((params, imp), ctxt') = Variable.focus NONE t ctxt in (ctxt', map #2 params, Logic.strip_imp_prems imp, Logic.strip_imp_concl imp) end fun mk_scheme' ctxt cases concl = let fun mk_branch concl = let val (_, ws, Cs, _ $ Pxs) = dest_hhf ctxt concl val (P, xs) = strip_comb Pxs in SchemeBranch { P=P, xs=map dest_Free xs, ws=ws, Cs=Cs } end val (branches, cases') = (* correction *) case Logic.dest_conjunctions concl of [conc] => let val _ $ Pxs = Logic.strip_assums_concl conc val (P, _) = strip_comb Pxs val (cases', conds) = chop_prefix (Term.exists_subterm (curry op aconv P)) cases val concl' = fold_rev (curry Logic.mk_implies) conds conc in ([mk_branch concl'], cases') end | concls => (map mk_branch concls, cases) fun mk_case premise = let val (ctxt', qs, prems, _ $ Plhs) = dest_hhf ctxt premise val (P, lhs) = strip_comb Plhs fun bidx Q = find_index (fn SchemeBranch {P=P',...} => Q aconv P') branches fun mk_rcinfo pr = let val (_, Gvs, Gas, _ $ Phyp) = dest_hhf ctxt' pr val (P', rcs) = strip_comb Phyp in (bidx P', Gvs, Gas, rcs) end fun is_pred v = exists (fn SchemeBranch {P,...} => v aconv P) branches val (gs, rcprs) = chop_prefix (not o Term.exists_subterm is_pred) prems in SchemeCase {bidx=bidx P, qs=qs, oqnames=map fst qs(*FIXME*), gs=gs, lhs=lhs, rs=map mk_rcinfo rcprs} end fun PT_of (SchemeBranch { xs, ...}) = foldr1 HOLogic.mk_prodT (map snd xs) val ST = Balanced_Tree.make (uncurry Sum_Tree.mk_sumT) (map PT_of branches) in IndScheme {T=ST, cases=map mk_case cases', branches=branches } end fun mk_completeness ctxt (IndScheme {cases, branches, ...}) bidx = let val SchemeBranch { xs, ws, Cs, ... } = nth branches bidx val relevant_cases = filter (fn SchemeCase {bidx=bidx', ...} => bidx' = bidx) cases val allqnames = fold (fn SchemeCase {qs, ...} => fold (insert (op =) o Free) qs) relevant_cases [] val (Pbool :: xs') = map Free (Variable.variant_frees ctxt allqnames (("P", HOLogic.boolT) :: xs)) val Cs' = map (Pattern.rewrite_term (Proof_Context.theory_of ctxt) (filter_out (op aconv) (map Free xs ~~ xs')) []) Cs fun mk_case (SchemeCase {qs, oqnames, gs, lhs, ...}) = HOLogic.mk_Trueprop Pbool |> fold_rev (fn x_l => curry Logic.mk_implies (HOLogic.mk_Trueprop(HOLogic.mk_eq x_l))) (xs' ~~ lhs) |> fold_rev (curry Logic.mk_implies) gs |> fold_rev mk_forall_rename (oqnames ~~ map Free qs) in HOLogic.mk_Trueprop Pbool |> fold_rev (curry Logic.mk_implies o mk_case) relevant_cases |> fold_rev (curry Logic.mk_implies) Cs' |> fold_rev (Logic.all o Free) ws |> fold_rev mk_forall_rename (map fst xs ~~ xs') |> mk_forall_rename ("P", Pbool) end fun mk_wf R (IndScheme {T, ...}) = HOLogic.Trueprop $ (Const (\<^const_name>\wf\, mk_relT T --> HOLogic.boolT) $ R) fun mk_ineqs R thesisn (IndScheme {T, cases, branches}) = let fun inject i ts = Sum_Tree.mk_inj T (length branches) (i + 1) (foldr1 HOLogic.mk_prod ts) val thesis = Free (thesisn, HOLogic.boolT) fun mk_pres bdx args = let val SchemeBranch { xs, ws, Cs, ... } = nth branches bdx fun replace (x, v) t = betapply (lambda (Free x) t, v) val Cs' = map (fold replace (xs ~~ args)) Cs val cse = HOLogic.mk_Trueprop thesis |> fold_rev (curry Logic.mk_implies) Cs' |> fold_rev (Logic.all o Free) ws in Logic.mk_implies (cse, HOLogic.mk_Trueprop thesis) end fun f (SchemeCase {bidx, qs, oqnames, gs, lhs, rs, ...}) = let fun g (bidx', Gvs, Gas, rcarg) = let val export = fold_rev (curry Logic.mk_implies) Gas #> fold_rev (curry Logic.mk_implies) gs #> fold_rev (Logic.all o Free) Gvs #> fold_rev mk_forall_rename (oqnames ~~ map Free qs) in (HOLogic.mk_mem (HOLogic.mk_prod (inject bidx' rcarg, inject bidx lhs), R) |> HOLogic.mk_Trueprop |> export, mk_pres bidx' rcarg |> export |> Logic.all thesis) end in map g rs end in map f cases end fun mk_ind_goal ctxt branches = let fun brnch (SchemeBranch { P, xs, ws, Cs, ... }) = HOLogic.mk_Trueprop (list_comb (P, map Free xs)) |> fold_rev (curry Logic.mk_implies) Cs |> fold_rev (Logic.all o Free) ws |> term_conv ctxt (ind_atomize ctxt) |> Object_Logic.drop_judgment ctxt |> HOLogic.tupled_lambda (foldr1 HOLogic.mk_prod (map Free xs)) in Sum_Tree.mk_sumcases HOLogic.boolT (map brnch branches) end fun mk_induct_rule ctxt R x complete_thms wf_thm ineqss (IndScheme {T, cases=scases, branches}) = let val n = length branches val scases_idx = map_index I scases fun inject i ts = Sum_Tree.mk_inj T n (i + 1) (foldr1 HOLogic.mk_prod ts) val P_of = nth (map (fn (SchemeBranch { P, ... }) => P) branches) val P_comp = mk_ind_goal ctxt branches (* Inductive Hypothesis: !!z. (z,x):R ==> P z *) val ihyp = Logic.all_const T $ Abs ("z", T, Logic.mk_implies (HOLogic.mk_Trueprop ( Const (\<^const_name>\Set.member\, HOLogic.mk_prodT (T, T) --> mk_relT T --> HOLogic.boolT) $ (HOLogic.pair_const T T $ Bound 0 $ x) $ R), HOLogic.mk_Trueprop (P_comp $ Bound 0))) |> Thm.cterm_of ctxt val aihyp = Thm.assume ihyp (* Rule for case splitting along the sum types *) val xss = map (fn (SchemeBranch { xs, ... }) => map Free xs) branches val pats = map_index (uncurry inject) xss val sum_split_rule = Pat_Completeness.prove_completeness ctxt [x] (P_comp $ x) xss (map single pats) fun prove_branch (bidx, (SchemeBranch { P, xs, ws, Cs, ... }, (complete_thm, pat))) = let val fxs = map Free xs val branch_hyp = Thm.assume (Thm.cterm_of ctxt (HOLogic.mk_Trueprop (HOLogic.mk_eq (x, pat)))) val C_hyps = map (Thm.cterm_of ctxt #> Thm.assume) Cs val (relevant_cases, ineqss') = (scases_idx ~~ ineqss) |> filter (fn ((_, SchemeCase {bidx=bidx', ...}), _) => bidx' = bidx) |> split_list fun prove_case (cidx, SchemeCase {qs, gs, lhs, rs, ...}) ineq_press = let val case_hyps = map (Thm.assume o Thm.cterm_of ctxt o HOLogic.mk_Trueprop o HOLogic.mk_eq) (fxs ~~ lhs) val cqs = map (Thm.cterm_of ctxt o Free) qs val ags = map (Thm.assume o Thm.cterm_of ctxt) gs val replace_x_simpset = put_simpset HOL_basic_ss ctxt addsimps (branch_hyp :: case_hyps) val sih = full_simplify replace_x_simpset aihyp fun mk_Prec (idx, Gvs, Gas, rcargs) (ineq, pres) = let val cGas = map (Thm.assume o Thm.cterm_of ctxt) Gas val cGvs = map (Thm.cterm_of ctxt o Free) Gvs val import = fold Thm.forall_elim (cqs @ cGvs) #> fold Thm.elim_implies (ags @ cGas) val ipres = pres |> Thm.forall_elim (Thm.cterm_of ctxt (list_comb (P_of idx, rcargs))) |> import in sih |> Thm.forall_elim (Thm.cterm_of ctxt (inject idx rcargs)) |> Thm.elim_implies (import ineq) (* Psum rcargs *) |> Conv.fconv_rule (sum_prod_conv ctxt) |> Conv.fconv_rule (ind_rulify ctxt) |> (fn th => th COMP ipres) (* P rs *) |> fold_rev (Thm.implies_intr o Thm.cprop_of) cGas |> fold_rev Thm.forall_intr cGvs end val P_recs = map2 mk_Prec rs ineq_press (* [P rec1, P rec2, ... ] *) val step = HOLogic.mk_Trueprop (list_comb (P, lhs)) |> fold_rev (curry Logic.mk_implies o Thm.prop_of) P_recs |> fold_rev (curry Logic.mk_implies) gs |> fold_rev (Logic.all o Free) qs |> Thm.cterm_of ctxt val Plhs_to_Pxs_conv = foldl1 (uncurry Conv.combination_conv) (Conv.all_conv :: map (fn ch => K (Thm.symmetric (ch RS eq_reflection))) case_hyps) val res = Thm.assume step |> fold Thm.forall_elim cqs |> fold Thm.elim_implies ags |> fold Thm.elim_implies P_recs (* P lhs *) |> Conv.fconv_rule (Conv.arg_conv Plhs_to_Pxs_conv) (* P xs *) |> fold_rev (Thm.implies_intr o Thm.cprop_of) (ags @ case_hyps) |> fold_rev Thm.forall_intr cqs (* !!qs. Gas ==> xs = lhss ==> P xs *) in (res, (cidx, step)) end val (cases, steps) = split_list (map2 prove_case relevant_cases ineqss') val bstep = complete_thm |> Thm.forall_elim (Thm.cterm_of ctxt (list_comb (P, fxs))) |> fold (Thm.forall_elim o Thm.cterm_of ctxt) (fxs @ map Free ws) |> fold Thm.elim_implies C_hyps |> fold Thm.elim_implies cases (* P xs *) |> fold_rev (Thm.implies_intr o Thm.cprop_of) C_hyps |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt o Free) ws val Pxs = Thm.cterm_of ctxt (HOLogic.mk_Trueprop (P_comp $ x)) |> Goal.init |> (Simplifier.rewrite_goals_tac ctxt (map meta (branch_hyp :: @{thm split_conv} :: @{thms sum.case})) THEN CONVERSION (ind_rulify ctxt) 1) |> Seq.hd |> Thm.elim_implies (Conv.fconv_rule Drule.beta_eta_conversion bstep) |> Goal.finish ctxt |> Thm.implies_intr (Thm.cprop_of branch_hyp) |> fold_rev (Thm.forall_intr o Thm.cterm_of ctxt) fxs in (Pxs, steps) end val (branches, steps) = map_index prove_branch (branches ~~ (complete_thms ~~ pats)) |> split_list |> apsnd flat val istep = sum_split_rule |> fold (fn b => fn th => Drule.compose (b, 1, th)) branches |> Thm.implies_intr ihyp |> Thm.forall_intr (Thm.cterm_of ctxt x) (* "!!x. (!!y P x" *) val induct_rule = @{thm "wf_induct_rule"} |> (curry op COMP) wf_thm |> (curry op COMP) istep val steps_sorted = map snd (sort (int_ord o apply2 fst) steps) in (steps_sorted, induct_rule) end fun mk_ind_tac comp_tac pres_tac term_tac ctxt facts = (* FIXME proper use of facts!? *) (ALLGOALS (Method.insert_tac ctxt facts)) THEN HEADGOAL (SUBGOAL (fn (t, i) => let val (ctxt', _, cases, concl) = dest_hhf ctxt t val scheme as IndScheme {T=ST, branches, ...} = mk_scheme' ctxt' cases concl val ([Rn, xn, thesisn], ctxt'') = Variable.variant_fixes ["R", "x", "thesis"] ctxt' val R = Free (Rn, mk_relT ST) val x = Free (xn, ST) val ineqss = mk_ineqs R thesisn scheme |> map (map (apply2 (Thm.assume o Thm.cterm_of ctxt''))) val complete = map_range (mk_completeness ctxt'' scheme #> Thm.cterm_of ctxt'' #> Thm.assume) (length branches) val wf_thm = mk_wf R scheme |> Thm.cterm_of ctxt'' |> Thm.assume val (descent, pres) = split_list (flat ineqss) val newgoals = complete @ pres @ wf_thm :: descent val (steps, indthm) = mk_induct_rule ctxt'' R x complete wf_thm ineqss scheme fun project (i, SchemeBranch {xs, ...}) = let val inst = (foldr1 HOLogic.mk_prod (map Free xs)) |> Sum_Tree.mk_inj ST (length branches) (i + 1) |> Thm.cterm_of ctxt'' in indthm |> Thm.instantiate' [] [SOME inst] |> simplify (put_simpset Sum_Tree.sumcase_split_ss ctxt'') |> Conv.fconv_rule (ind_rulify ctxt'') end val res = Conjunction.intr_balanced (map_index project branches) |> fold_rev Thm.implies_intr (map Thm.cprop_of newgoals @ steps) - |> Drule.generalize (Symtab.empty, Symtab.make_set [Rn]) + |> Drule.generalize (Names.empty, Names.make_set [Rn]) val nbranches = length branches val npres = length pres in Thm.bicompose (SOME ctxt'') {flatten = false, match = false, incremented = false} (false, res, length newgoals) i THEN term_tac (i + nbranches + npres) THEN (EVERY (map (TRY o pres_tac) ((i + nbranches + npres - 1) downto (i + nbranches)))) THEN (EVERY (map (TRY o comp_tac) ((i + nbranches - 1) downto i))) end)) fun induction_schema_tac ctxt = mk_ind_tac (K all_tac) (assume_tac ctxt APPEND' Goal.assume_rule_tac ctxt) (K all_tac) ctxt; end diff --git a/src/HOL/Tools/Lifting/lifting_setup.ML b/src/HOL/Tools/Lifting/lifting_setup.ML --- a/src/HOL/Tools/Lifting/lifting_setup.ML +++ b/src/HOL/Tools/Lifting/lifting_setup.ML @@ -1,1048 +1,1048 @@ (* Title: HOL/Tools/Lifting/lifting_setup.ML Author: Ondrej Kuncar Setting up the lifting infrastructure. *) signature LIFTING_SETUP = sig exception SETUP_LIFTING_INFR of string type config = { notes: bool }; val default_config: config; val setup_by_quotient: config -> thm -> thm option -> thm option -> local_theory -> binding * local_theory val setup_by_typedef_thm: config -> thm -> local_theory -> binding * local_theory val lifting_restore: Lifting_Info.quotient -> Context.generic -> Context.generic val lifting_forget: string -> local_theory -> local_theory val update_transfer_rules: string -> local_theory -> local_theory val pointer_of_bundle_binding: Proof.context -> binding -> string end structure Lifting_Setup: LIFTING_SETUP = struct open Lifting_Util infix 0 MRSL exception SETUP_LIFTING_INFR of string (* Config *) type config = { notes: bool }; val default_config = { notes = true }; fun define_crel (config: config) rep_fun lthy = let val (qty, rty) = (dest_funT o fastype_of) rep_fun val rep_fun_graph = (HOLogic.eq_const rty) $ Bound 1 $ (rep_fun $ Bound 0) val def_term = Abs ("x", rty, Abs ("y", qty, rep_fun_graph)) val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type) qty val crel_name = Binding.prefix_name "cr_" qty_name val (fixed_def_term, lthy1) = lthy |> yield_singleton (Variable.importT_terms) def_term val ((_, (_ , def_thm)), lthy2) = if #notes config then Local_Theory.define ((crel_name, NoSyn), ((Thm.def_binding crel_name, []), fixed_def_term)) lthy1 else Local_Theory.define ((Binding.concealed crel_name, NoSyn), (Binding.empty_atts, fixed_def_term)) lthy1 in (def_thm, lthy2) end fun print_define_pcrel_warning msg = let val warning_msg = cat_lines ["Generation of a parametrized correspondence relation failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in warning warning_msg end fun define_pcrel (config: config) crel lthy0 = let val (fixed_crel, lthy1) = yield_singleton Variable.importT_terms crel lthy0 val [rty', qty] = (binder_types o fastype_of) fixed_crel val (param_rel, args) = Lifting_Term.generate_parametrized_relator lthy1 rty' val rty_raw = (domain_type o range_type o fastype_of) param_rel val tyenv_match = Sign.typ_match (Proof_Context.theory_of lthy1) (rty_raw, rty') Vartab.empty val param_rel_subst = Envir.subst_term (tyenv_match,Vartab.empty) param_rel val args_subst = map (Envir.subst_term (tyenv_match,Vartab.empty)) args val (instT, lthy2) = lthy1 |> Variable.declare_names fixed_crel |> Variable.importT_inst (param_rel_subst :: args_subst) - val args_fixed = (map (Term_Subst.instantiate (instT, Term_Subst.Vars.empty))) args_subst - val param_rel_fixed = Term_Subst.instantiate (instT, Term_Subst.Vars.empty) param_rel_subst + val args_fixed = (map (Term_Subst.instantiate (instT, Vars.empty))) args_subst + val param_rel_fixed = Term_Subst.instantiate (instT, Vars.empty) param_rel_subst val rty = (domain_type o fastype_of) param_rel_fixed val relcomp_op = Const (\<^const_name>\relcompp\, (rty --> rty' --> HOLogic.boolT) --> (rty' --> qty --> HOLogic.boolT) --> rty --> qty --> HOLogic.boolT) val qty_name = (fst o dest_Type) qty val pcrel_name = Binding.prefix_name "pcr_" ((Binding.name o Long_Name.base_name) qty_name) val relator_type = foldr1 (op -->) ((map type_of args_fixed) @ [rty, qty, HOLogic.boolT]) val lhs = Library.foldl (op $) ((Free (Binding.name_of pcrel_name, relator_type)), args_fixed) val rhs = relcomp_op $ param_rel_fixed $ fixed_crel val definition_term = Logic.mk_equals (lhs, rhs) fun note_def lthy = Specification.definition (SOME (pcrel_name, SOME relator_type, NoSyn)) [] [] (Binding.empty_atts, definition_term) lthy |>> (snd #> snd); fun raw_def lthy = let val ((_, rhs), prove) = Local_Defs.derived_def lthy (K []) {conditional = true} definition_term; val ((_, (_, raw_th)), lthy') = Local_Theory.define ((Binding.concealed pcrel_name, NoSyn), (Binding.empty_atts, rhs)) lthy; val th = prove lthy' raw_th; in (th, lthy') end val (def_thm, lthy3) = if #notes config then note_def lthy2 else raw_def lthy2 in (SOME def_thm, lthy3) end handle Lifting_Term.PARAM_QUOT_THM (_, msg) => (print_define_pcrel_warning msg; (NONE, lthy0)) local val eq_OO_meta = mk_meta_eq @{thm eq_OO} fun print_generate_pcr_cr_eq_error ctxt term = let val goal = Const (\<^const_name>\HOL.eq\, dummyT) $ term $ Const (\<^const_name>\HOL.eq\, dummyT) val error_msg = cat_lines ["Generation of a pcr_cr_eq failed.", (Pretty.string_of (Pretty.block [Pretty.str "Reason: Cannot prove this: ", Pretty.brk 2, Syntax.pretty_term ctxt goal])), "Most probably a relator_eq rule for one of the involved types is missing."] in error error_msg end in fun define_pcr_cr_eq (config: config) lthy pcr_rel_def = let val lhs = (Thm.term_of o Thm.lhs_of) pcr_rel_def val qty_name = (Binding.name o Long_Name.base_name o fst o dest_Type o List.last o binder_types o fastype_of) lhs val args = (snd o strip_comb) lhs fun make_inst var ctxt = let val typ = snd (relation_types (#2 (dest_Var var))) val sort = Type.sort_of_atyp typ val (fresh_var, ctxt') = yield_singleton Variable.invent_types sort ctxt val inst = (#1 (dest_Var var), Thm.cterm_of ctxt' (HOLogic.eq_const (TFree fresh_var))) in (inst, ctxt') end val (args_inst, args_ctxt) = fold_map make_inst args lthy val pcr_cr_eq = pcr_rel_def |> infer_instantiate args_ctxt args_inst |> Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv (Transfer.bottom_rewr_conv (Transfer.get_relator_eq args_ctxt)))) in case (Thm.term_of o Thm.rhs_of) pcr_cr_eq of Const (\<^const_name>\relcompp\, _) $ Const (\<^const_name>\HOL.eq\, _) $ _ => let val thm = pcr_cr_eq |> Conv.fconv_rule (Conv.arg_conv (Conv.rewr_conv eq_OO_meta)) |> HOLogic.mk_obj_eq |> singleton (Variable.export args_ctxt lthy) val lthy' = lthy |> #notes config ? (Local_Theory.note ((Binding.qualify_name true qty_name "pcr_cr_eq", []), [thm]) #> #2) in (thm, lthy') end | Const (\<^const_name>\relcompp\, _) $ t $ _ => print_generate_pcr_cr_eq_error args_ctxt t | _ => error "generate_pcr_cr_eq: implementation error" end end fun define_code_constr quot_thm lthy = let val abs = quot_thm_abs quot_thm in if is_Const abs then let val (fixed_abs, lthy') = yield_singleton Variable.importT_terms abs lthy in Local_Theory.background_theory (Code.declare_datatype_global [dest_Const fixed_abs]) lthy' end else lthy end fun define_abs_type quot_thm = Lifting_Def.can_generate_code_cert quot_thm ? Code.declare_abstype (quot_thm RS @{thm Quotient_abs_rep}); local exception QUOT_ERROR of Pretty.T list in fun quot_thm_sanity_check ctxt quot_thm = let val _ = if (Thm.nprems_of quot_thm > 0) then raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem has extra assumptions:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] else () val _ = quot_thm |> Thm.concl_of |> HOLogic.dest_Trueprop |> dest_Quotient handle TERM _ => raise QUOT_ERROR [Pretty.block [Pretty.str "The Quotient theorem is not of the right form:", Pretty.brk 1, Thm.pretty_thm ctxt quot_thm]] val ((_, [quot_thm_fixed]), ctxt') = Variable.importT [quot_thm] ctxt val (rty, qty) = quot_thm_rty_qty quot_thm_fixed val rty_tfreesT = Term.add_tfree_namesT rty [] val qty_tfreesT = Term.add_tfree_namesT qty [] val extra_rty_tfrees = case subtract (op =) qty_tfreesT rty_tfreesT of [] => [] | extras => [Pretty.block ([Pretty.str "Extra variables in the raw type:", Pretty.brk 1] @ ((Pretty.commas o map (Pretty.str o quote)) extras) @ [Pretty.str "."])] val not_type_constr = case qty of Type _ => [] | _ => [Pretty.block [Pretty.str "The quotient type ", Pretty.quote (Syntax.pretty_typ ctxt' qty), Pretty.brk 1, Pretty.str "is not a type constructor."]] val errs = extra_rty_tfrees @ not_type_constr in if null errs then () else raise QUOT_ERROR errs end handle QUOT_ERROR errs => error (cat_lines (["Sanity check of the quotient theorem failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end fun lifting_bundle qty_full_name qinfo lthy = let val thy = Proof_Context.theory_of lthy val binding = Binding.qualify_name true (qty_full_name |> Long_Name.base_name |> Binding.name) "lifting" val morphed_binding = Morphism.binding (Local_Theory.target_morphism lthy) binding val bundle_name = Name_Space.full_name (Name_Space.naming_of (Context.Theory thy)) morphed_binding fun phi_qinfo phi = Lifting_Info.transform_quotient phi qinfo val dummy_thm = Thm.transfer thy Drule.dummy_thm val restore_lifting_att = ([dummy_thm], [map (Token.make_string o rpair Position.none) ["Lifting.lifting_restore_internal", bundle_name]]) in lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Lifting_Info.init_restore_data bundle_name (phi_qinfo phi)) |> Bundle.bundle ((binding, [restore_lifting_att])) [] |> pair binding end fun setup_lifting_infr config quot_thm opt_reflp_thm lthy = let val _ = quot_thm_sanity_check lthy quot_thm val (_, qty) = quot_thm_rty_qty quot_thm val (pcrel_def, lthy1) = define_pcrel config (quot_thm_crel quot_thm) lthy (**) val pcrel_def = Option.map (Morphism.thm (Local_Theory.target_morphism lthy1)) pcrel_def (**) val (pcr_cr_eq, lthy2) = case pcrel_def of SOME pcrel_def => apfst SOME (define_pcr_cr_eq config lthy1 pcrel_def) | NONE => (NONE, lthy1) val pcr_info = case pcrel_def of SOME pcrel_def => SOME { pcrel_def = pcrel_def, pcr_cr_eq = the pcr_cr_eq } | NONE => NONE val quotients = { quot_thm = quot_thm, pcr_info = pcr_info } val qty_full_name = (fst o dest_Type) qty fun quot_info phi = Lifting_Info.transform_quotient phi quotients val reflexivity_rule_attr = Attrib.internal (K Lifting_Info.add_reflexivity_rule_attribute) val lthy3 = case opt_reflp_thm of SOME reflp_thm => lthy2 |> (#2 oo Local_Theory.note) ((Binding.empty, [reflexivity_rule_attr]), [reflp_thm RS @{thm reflp_ge_eq}]) |> define_code_constr quot_thm | NONE => lthy2 |> define_abs_type quot_thm in lthy3 |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Lifting_Info.update_quotients qty_full_name (quot_info phi)) |> lifting_bundle qty_full_name quotients end local fun importT_inst_exclude exclude ts ctxt = let val tvars = rev (subtract op= exclude (fold Term.add_tvars ts [])) val (tfrees, ctxt') = Variable.invent_types (map #2 tvars) ctxt - in (Term_Subst.TVars.table (tvars ~~ map TFree tfrees), ctxt') end + in (TVars.make (tvars ~~ map TFree tfrees), ctxt') end fun import_inst_exclude exclude ts ctxt = let val excludeT = fold (Term.add_tvarsT o snd) exclude [] val (instT, ctxt') = importT_inst_exclude excludeT ts ctxt val vars = map (apsnd (Term_Subst.instantiateT instT)) (rev (subtract op= exclude (fold Term.add_vars ts []))) val (xs, ctxt'') = Variable.variant_fixes (map (#1 o #1) vars) ctxt' - val inst = Term_Subst.Vars.table (vars ~~ map Free (xs ~~ map #2 vars)) + val inst = Vars.make (vars ~~ map Free (xs ~~ map #2 vars)) in ((instT, inst), ctxt'') end fun import_terms_exclude exclude ts ctxt = let val (inst, ctxt') = import_inst_exclude exclude ts ctxt in (map (Term_Subst.instantiate inst) ts, ctxt') end in fun reduce_goal not_fix goal tac ctxt = let val (fixed_goal, ctxt') = yield_singleton (import_terms_exclude not_fix) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE tac init_goal)) end end local val OO_rules = @{thms left_total_OO left_unique_OO right_total_OO right_unique_OO bi_total_OO bi_unique_OO} in fun parametrize_class_constraint ctxt0 pcr_def constraint = let fun generate_transfer_rule pcr_def constraint goal ctxt = let val (fixed_goal, ctxt') = yield_singleton (Variable.import_terms true) goal ctxt val init_goal = Goal.init (Thm.cterm_of ctxt' fixed_goal) val rules = Transfer.get_transfer_raw ctxt' val rules = constraint :: OO_rules @ rules val tac = K (Local_Defs.unfold0_tac ctxt' [pcr_def]) THEN' REPEAT_ALL_NEW (resolve_tac ctxt' rules) in (singleton (Variable.export ctxt' ctxt) o Goal.conclude) (the (SINGLE (tac 1) init_goal)) end fun make_goal pcr_def constr = let val pred_name = (fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.prop_of) constr val arg = (fst o Logic.dest_equals o Thm.prop_of) pcr_def in HOLogic.mk_Trueprop ((Const (pred_name, (fastype_of arg) --> HOLogic.boolT)) $ arg) end val check_assms = let val right_names = ["right_total", "right_unique", "left_total", "left_unique", "bi_total", "bi_unique"] fun is_right_name name = member op= right_names (Long_Name.base_name name) fun is_trivial_assm (Const (name, _) $ Var (_, _)) = is_right_name name | is_trivial_assm (Const (name, _) $ Free (_, _)) = is_right_name name | is_trivial_assm _ = false in fn thm => let val prems = map HOLogic.dest_Trueprop (Thm.prems_of thm) val thm_name = (Long_Name.base_name o fst o dest_Const o strip_args 1 o HOLogic.dest_Trueprop o Thm.concl_of) thm val non_trivial_assms = filter_out is_trivial_assm prems in if null non_trivial_assms then () else Pretty.block ([Pretty.str "Non-trivial assumptions in ", Pretty.str thm_name, Pretty.str " transfer rule found:", Pretty.brk 1] @ Pretty.commas (map (Syntax.pretty_term ctxt0) non_trivial_assms)) |> Pretty.string_of |> warning end end val goal = make_goal pcr_def constraint val thm = generate_transfer_rule pcr_def constraint goal ctxt0 val _ = check_assms thm in thm end end local val id_unfold = (Conv.rewr_conv (mk_meta_eq @{thm id_def})) in fun generate_parametric_id lthy rty id_transfer_rule = let (* it doesn't raise an exception because it would have already raised it in define_pcrel *) val (quot_thm, _, ctxt') = Lifting_Term.prove_param_quot_thm lthy rty val parametrized_relator = singleton (Variable.export_terms ctxt' lthy) (quot_thm_crel quot_thm) val id_transfer = @{thm id_transfer} |> Thm.incr_indexes (Term.maxidx_of_term parametrized_relator + 1) |> Conv.fconv_rule(HOLogic.Trueprop_conv (Conv.arg_conv id_unfold then_conv Conv.arg1_conv id_unfold)) val var = hd (Term.add_vars (Thm.prop_of id_transfer) []) val inst = [(#1 var, Thm.cterm_of lthy parametrized_relator)] val id_par_thm = infer_instantiate lthy inst id_transfer in Lifting_Def.generate_parametric_transfer_rule lthy id_transfer_rule id_par_thm end handle Lifting_Term.MERGE_TRANSFER_REL msg => let val error_msg = cat_lines ["Generation of a parametric transfer rule for the abs. or the rep. function failed.", "A non-parametric version will be used.", (Pretty.string_of (Pretty.block [Pretty.str "Reason:", Pretty.brk 2, msg]))] in (warning error_msg; id_transfer_rule) end end local fun rewrite_first_Domainp_arg rewr_thm thm = Conv.fconv_rule (Conv.concl_conv ~1 (HOLogic.Trueprop_conv (Conv.arg1_conv (Conv.arg_conv (Conv.rewr_conv rewr_thm))))) thm fun fold_Domainp_pcrel pcrel_def thm = let val ct = thm |> Thm.cprop_of |> Drule.strip_imp_concl |> Thm.dest_arg |> Thm.dest_arg1 |> Thm.dest_arg val pcrel_def = Thm.incr_indexes (Thm.maxidx_of_cterm ct + 1) pcrel_def val thm' = Thm.instantiate (Thm.match (ct, Thm.rhs_of pcrel_def)) thm handle Pattern.MATCH => raise CTERM ("fold_Domainp_pcrel", [ct, Thm.rhs_of pcrel_def]) in rewrite_first_Domainp_arg (Thm.symmetric pcrel_def) thm' end fun reduce_Domainp ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val var = goal |> HOLogic.dest_Trueprop |> dest_comb |> snd |> dest_Var val reduced_assm = reduce_goal [var] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end in fun parametrize_domain dom_thm (pcr_info : Lifting_Info.pcr) ctxt0 = let fun reduce_first_assm ctxt rules thm = let val goal = thm |> Thm.prems_of |> hd val reduced_assm = reduce_goal [] goal (TRY (REPEAT_ALL_NEW (resolve_tac ctxt rules) 1)) ctxt in reduced_assm RS thm end val pcr_cr_met_eq = #pcr_cr_eq pcr_info RS @{thm eq_reflection} val pcr_Domainp_eq = rewrite_first_Domainp_arg (Thm.symmetric pcr_cr_met_eq) dom_thm val pcrel_def = #pcrel_def pcr_info val pcr_Domainp_par_left_total = (dom_thm RS @{thm pcr_Domainp_par_left_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_first_assm ctxt0 (Lifting_Info.get_reflexivity_rules ctxt0) val pcr_Domainp_par = (dom_thm RS @{thm pcr_Domainp_par}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt0 (Transfer.get_relator_domain ctxt0) val pcr_Domainp = (dom_thm RS @{thm pcr_Domainp}) |> fold_Domainp_pcrel pcrel_def val thms = [("domain", [pcr_Domainp], @{attributes [transfer_domain_rule]}), ("domain_par", [pcr_Domainp_par], @{attributes [transfer_domain_rule]}), ("domain_par_left_total", [pcr_Domainp_par_left_total], @{attributes [transfer_domain_rule]}), ("domain_eq", [pcr_Domainp_eq], @{attributes [transfer_domain_rule]})] in thms end fun parametrize_total_domain left_total pcrel_def ctxt = let val thm = (left_total RS @{thm pcr_Domainp_total}) |> fold_Domainp_pcrel pcrel_def |> reduce_Domainp ctxt (Transfer.get_relator_domain ctxt) in [("domain", [thm], @{attributes [transfer_domain_rule]})] end end fun get_pcrel_info ctxt qty_full_name = #pcr_info (the (Lifting_Info.lookup_quotients ctxt qty_full_name)) fun get_Domainp_thm quot_thm = the (get_first (try(curry op RS quot_thm)) [@{thm eq_onp_to_Domainp}, @{thm Quotient_to_Domainp}]) fun notes names thms = let val notes = if names then map (fn (name, thms, attrs) => ((name, []), [(thms, attrs)])) thms else map_filter (fn (_, thms, attrs) => if null attrs then NONE else SOME (Binding.empty_atts, [(thms, attrs)])) thms in Local_Theory.notes notes #> snd end fun map_thms map_name map_thm thms = map (fn (name, thms, attr) => (map_name name, map map_thm thms, attr)) thms (* Sets up the Lifting package by a quotient theorem. quot_thm - a quotient theorem (Quotient R Abs Rep T) opt_reflp_thm - a theorem saying that a relation from quot_thm is reflexive (in the form "reflp R") opt_par_thm - a parametricity theorem for R *) fun setup_by_quotient (config: config) quot_thm opt_reflp_thm opt_par_thm lthy0 = let (**) val quot_thm = Morphism.thm (Local_Theory.target_morphism lthy0) quot_thm (**) val (rty, qty) = quot_thm_rty_qty quot_thm val induct_attr = Attrib.internal (K (Induct.induct_type (fst (dest_Type qty)))) val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("abs_induct", @{thms Quotient_total_abs_induct}, [induct_attr]), ("abs_eq_iff", @{thms Quotient_total_abs_eq_iff}, [])] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => let val thms = [("abs_induct", @{thms Quotient_abs_induct}, [induct_attr])] in map_thms qualify (fn thm => quot_thm RS thm) thms end val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val notes2 = map_thms qualify (fn thm => quot_thm RS thm) [("rel_eq_transfer", @{thms Quotient_rel_eq_transfer}, @{attributes [transfer_rule]}), ("right_unique", @{thms Quotient_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms Quotient_right_total}, @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun generate_parametric_rel_eq ctxt transfer_rule opt_param_thm = (case opt_param_thm of NONE => transfer_rule | SOME param_thm => (Lifting_Def.generate_parametric_transfer_rule ctxt transfer_rule param_thm handle Lifting_Term.MERGE_TRANSFER_REL msg => error ("Generation of a parametric transfer rule for the quotient relation failed:\n" ^ Pretty.string_of msg))) fun setup_transfer_rules_par ctxt notes = let val pcrel_info = the (get_pcrel_info ctxt qty_full_name) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val thms = [("id_abs_transfer", [id_abs_transfer], @{attributes [transfer_rule]}), ("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val rel_eq_transfer = generate_parametric_rel_eq ctxt (Lifting_Term.parametrize_transfer_rule ctxt (quot_thm RS @{thm Quotient_rel_eq_transfer})) opt_par_thm val right_unique = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_unique}) val right_total = parametrize_class_constraint ctxt pcrel_def (quot_thm RS @{thm Quotient_right_total}) val notes2 = map_thms qualify I [("rel_eq_transfer", [rel_eq_transfer], @{attributes [transfer_rule]}), ("right_unique", [right_unique], @{attributes [transfer_rule]}), ("right_total", [right_total], @{attributes [transfer_rule]})] in notes2 @ notes1 @ notes end fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy0 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end (* Sets up the Lifting package by a typedef theorem. gen_code - flag if an abstract type given by typedef_thm should be registred as an abstract type in the code generator typedef_thm - a typedef theorem (type_definition Rep Abs S) *) fun setup_by_typedef_thm config typedef_thm lthy0 = let val (_ $ rep_fun $ _ $ typedef_set) = (HOLogic.dest_Trueprop o Thm.prop_of) typedef_thm val (T_def, lthy1) = define_crel config rep_fun lthy0 (**) val T_def = Morphism.thm (Local_Theory.target_morphism lthy1) T_def (**) val quot_thm = case typedef_set of Const (\<^const_name>\top\, _) => [typedef_thm, T_def] MRSL @{thm UNIV_typedef_to_Quotient} | Const (\<^const_name>\Collect\, _) $ Abs (_, _, _) => [typedef_thm, T_def] MRSL @{thm open_typedef_to_Quotient} | _ => [typedef_thm, T_def] MRSL @{thm typedef_to_Quotient} val (rty, qty) = quot_thm_rty_qty quot_thm val qty_full_name = (fst o dest_Type) qty val qty_name = (Binding.name o Long_Name.base_name) qty_full_name val qualify = Binding.qualify_name true qty_name val opt_reflp_thm = case typedef_set of Const (\<^const_name>\top\, _) => SOME ((typedef_thm RS @{thm UNIV_typedef_to_equivp}) RS @{thm equivp_reflp2}) | _ => NONE val dom_thm = get_Domainp_thm quot_thm fun setup_transfer_rules_nonpar notes = let val notes1 = case opt_reflp_thm of SOME reflp_thm => let val thms = [("id_abs_transfer",@{thms Quotient_id_abs_transfer}, @{attributes [transfer_rule]}), ("left_total", @{thms Quotient_left_total}, @{attributes [transfer_rule]}), ("bi_total", @{thms Quotient_bi_total}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [quot_thm, reflp_thm] MRSL thm) thms end | NONE => map_thms qualify I [("domain", [dom_thm], @{attributes [transfer_domain_rule]})] val thms = [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]}), ("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]})] in map_thms qualify (fn thm => [typedef_thm, T_def] MRSL thm) thms @ notes1 @ notes end fun setup_transfer_rules_par ctxt notes = let val pcrel_info = (the (get_pcrel_info ctxt qty_full_name)) val pcrel_def = #pcrel_def pcrel_info val notes1 = case opt_reflp_thm of SOME reflp_thm => let val left_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_left_total}) val bi_total = ([quot_thm, reflp_thm] MRSL @{thm Quotient_bi_total}) val domain_thms = parametrize_total_domain left_total pcrel_def ctxt val left_total = parametrize_class_constraint ctxt pcrel_def left_total val bi_total = parametrize_class_constraint ctxt pcrel_def bi_total val id_abs_transfer = generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([quot_thm, reflp_thm] MRSL @{thm Quotient_id_abs_transfer})) val thms = [("left_total", [left_total], @{attributes [transfer_rule]}), ("bi_total", [bi_total], @{attributes [transfer_rule]}), ("id_abs_transfer",[id_abs_transfer], @{attributes [transfer_rule]})] in map_thms qualify I thms @ map_thms qualify I domain_thms end | NONE => let val thms = parametrize_domain dom_thm pcrel_info ctxt in map_thms qualify I thms end val notes2 = map_thms qualify (fn thm => generate_parametric_id ctxt rty (Lifting_Term.parametrize_transfer_rule ctxt ([typedef_thm, T_def] MRSL thm))) [("rep_transfer", @{thms typedef_rep_transfer}, @{attributes [transfer_rule]})]; val notes3 = map_thms qualify (fn thm => parametrize_class_constraint ctxt pcrel_def ([typedef_thm, T_def] MRSL thm)) [("left_unique", @{thms typedef_left_unique}, @{attributes [transfer_rule]}), ("right_unique", @{thms typedef_right_unique},@{attributes [transfer_rule]}), ("bi_unique", @{thms typedef_bi_unique}, @{attributes [transfer_rule]}), ("right_total", @{thms typedef_right_total}, @{attributes [transfer_rule]})] in notes3 @ notes2 @ notes1 @ notes end val notes1 = [(Binding.prefix_name "Quotient_" qty_name, [quot_thm], [])] fun setup_rules lthy = let val thms = if is_some (get_pcrel_info lthy qty_full_name) then setup_transfer_rules_par lthy notes1 else setup_transfer_rules_nonpar notes1 in notes (#notes config) thms lthy end in lthy1 |> setup_lifting_infr config quot_thm opt_reflp_thm ||> setup_rules end fun setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm lthy = let val input_thm = singleton (Attrib.eval_thms lthy) xthm val input_term = (HOLogic.dest_Trueprop o Thm.prop_of) input_thm handle TERM _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." fun sanity_check_reflp_thm reflp_thm = let val reflp_tm = (HOLogic.dest_Trueprop o Thm.prop_of) reflp_thm handle TERM _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." in case reflp_tm of Const (\<^const_name>\reflp\, _) $ _ => () | _ => error "Invalid form of the reflexivity theorem. Use \"reflp R\"." end fun check_qty qty = if not (is_Type qty) then error "The abstract type must be a type constructor." else () fun setup_quotient () = let val opt_reflp_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_reflp_xthm val _ = if is_some opt_reflp_thm then sanity_check_reflp_thm (the opt_reflp_thm) else () val opt_par_thm = Option.map (singleton (Attrib.eval_thms lthy)) opt_par_xthm val _ = check_qty (snd (quot_thm_rty_qty input_thm)) in setup_by_quotient default_config input_thm opt_reflp_thm opt_par_thm lthy |> snd end fun setup_typedef () = let val qty = (range_type o fastype_of o hd o get_args 2) input_term val _ = check_qty qty in case opt_reflp_xthm of SOME _ => error "The reflexivity theorem cannot be specified if the type_definition theorem is used." | NONE => ( case opt_par_xthm of SOME _ => error "The parametricity theorem cannot be specified if the type_definition theorem is used." | NONE => setup_by_typedef_thm default_config input_thm lthy |> snd ) end in case input_term of (Const (\<^const_name>\Quotient\, _) $ _ $ _ $ _ $ _) => setup_quotient () | (Const (\<^const_name>\type_definition\, _) $ _ $ _ $ _) => setup_typedef () | _ => error "Unsupported type of a theorem. Only Quotient or type_definition are supported." end val _ = Outer_Syntax.local_theory \<^command_keyword>\setup_lifting\ "setup lifting infrastructure" (Parse.thm -- Scan.option Parse.thm -- Scan.option (\<^keyword>\parametric\ |-- Parse.!!! Parse.thm) >> (fn ((xthm, opt_reflp_xthm), opt_par_xthm) => setup_lifting_cmd xthm opt_reflp_xthm opt_par_xthm)) (* restoring lifting infrastructure *) local exception PCR_ERROR of Pretty.T list in fun lifting_restore_sanity_check ctxt (qinfo:Lifting_Info.quotient) = let val quot_thm = (#quot_thm qinfo) val _ = quot_thm_sanity_check ctxt quot_thm val pcr_info_err = (case #pcr_info qinfo of SOME pcr => let val pcrel_def = #pcrel_def pcr val pcr_cr_eq = #pcr_cr_eq pcr val (def_lhs, _) = Logic.dest_equals (Thm.prop_of pcrel_def) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr definiton theorem is not a plain meta equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcrel_def]] val pcr_const_def = head_of def_lhs val (eq_lhs, eq_rhs) = HOLogic.dest_eq (HOLogic.dest_Trueprop (Thm.prop_of pcr_cr_eq)) handle TERM _ => raise PCR_ERROR [Pretty.block [Pretty.str "The pcr_cr equation theorem is not a plain equation:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] val (pcr_const_eq, eqs) = strip_comb eq_lhs fun is_eq (Const (\<^const_name>\HOL.eq\, _)) = true | is_eq _ = false fun eq_Const (Const (name1, _)) (Const (name2, _)) = (name1 = name2) | eq_Const _ _ = false val all_eqs = if not (forall is_eq eqs) then [Pretty.block [Pretty.str "Arguments of the lhs of the pcr_cr equation theorem are not only equalities:", Pretty.brk 1, Thm.pretty_thm ctxt pcr_cr_eq]] else [] val pcr_consts_not_equal = if not (eq_Const pcr_const_def pcr_const_eq) then [Pretty.block [Pretty.str "Parametrized correspondence relation constants in pcr_def and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_def, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt pcr_const_eq]] else [] val crel = quot_thm_crel quot_thm val cr_consts_not_equal = if not (eq_Const crel eq_rhs) then [Pretty.block [Pretty.str "Correspondence relation constants in the Quotient theorem and pcr_cr_eq are not equal:", Pretty.brk 1, Syntax.pretty_term ctxt crel, Pretty.brk 1, Pretty.str "vs.", Pretty.brk 1, Syntax.pretty_term ctxt eq_rhs]] else [] in all_eqs @ pcr_consts_not_equal @ cr_consts_not_equal end | NONE => []) val errs = pcr_info_err in if null errs then () else raise PCR_ERROR errs end handle PCR_ERROR errs => error (cat_lines (["Sanity check failed:"] @ (map (Pretty.string_of o Pretty.item o single) errs))) end (* Registers the data in qinfo in the Lifting infrastructure. *) fun lifting_restore qinfo ctxt = let val _ = lifting_restore_sanity_check (Context.proof_of ctxt) qinfo val (_, qty) = quot_thm_rty_qty (#quot_thm qinfo) val qty_full_name = (fst o dest_Type) qty val stored_qinfo = Lifting_Info.lookup_quotients (Context.proof_of ctxt) qty_full_name in if is_some (stored_qinfo) andalso not (Lifting_Info.quotient_eq (qinfo, (the stored_qinfo))) then error (Pretty.string_of (Pretty.block [Pretty.str "Lifting is already setup for the type", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Context.proof_of ctxt) qty)])) else Lifting_Info.update_quotients qty_full_name qinfo ctxt end val parse_opt_pcr = Scan.optional (Attrib.thm -- Attrib.thm >> (fn (pcrel_def, pcr_cr_eq) => SOME {pcrel_def = pcrel_def, pcr_cr_eq = pcr_cr_eq})) NONE val lifting_restore_attribute_setup = Attrib.setup \<^binding>\lifting_restore\ ((Attrib.thm -- parse_opt_pcr) >> (fn (quot_thm, opt_pcr) => let val qinfo = { quot_thm = quot_thm, pcr_info = opt_pcr} in Thm.declaration_attribute (K (lifting_restore qinfo)) end)) "restoring lifting infrastructure" val _ = Theory.setup lifting_restore_attribute_setup fun lifting_restore_internal bundle_name ctxt = let val restore_info = Lifting_Info.lookup_restore_data (Context.proof_of ctxt) bundle_name in case restore_info of SOME restore_info => ctxt |> lifting_restore (#quotient restore_info) |> fold_rev Transfer.transfer_raw_add (Item_Net.content (#transfer_rules restore_info)) | NONE => ctxt end val lifting_restore_internal_attribute_setup = Attrib.setup \<^binding>\lifting_restore_internal\ (Scan.lift Parse.string >> (fn name => Thm.declaration_attribute (K (lifting_restore_internal name)))) "restoring lifting infrastructure; internal attribute; not meant to be used directly by regular users" val _ = Theory.setup lifting_restore_internal_attribute_setup (* lifting_forget *) val monotonicity_names = [\<^const_name>\right_unique\, \<^const_name>\left_unique\, \<^const_name>\right_total\, \<^const_name>\left_total\, \<^const_name>\bi_unique\, \<^const_name>\bi_total\] fun fold_transfer_rel f (Const (\<^const_name>\Transfer.Rel\, _) $ rel $ _ $ _) = f rel | fold_transfer_rel f (Const (\<^const_name>\HOL.eq\, _) $ (Const (\<^const_name>\Domainp\, _) $ rel) $ _) = f rel | fold_transfer_rel f (Const (name, _) $ rel) = if member op= monotonicity_names name then f rel else f \<^term>\undefined\ | fold_transfer_rel f _ = f \<^term>\undefined\ fun filter_transfer_rules_by_rel transfer_rel transfer_rules = let val transfer_rel_name = transfer_rel |> dest_Const |> fst; fun has_transfer_rel thm = let val concl = thm |> Thm.concl_of |> HOLogic.dest_Trueprop in member op= (fold_transfer_rel (fn tm => Term.add_const_names tm []) concl) transfer_rel_name end handle TERM _ => false in filter has_transfer_rel transfer_rules end type restore_data = {quotient : Lifting_Info.quotient, transfer_rules: thm Item_Net.T} fun get_transfer_rel (qinfo : Lifting_Info.quotient) = let fun get_pcrel pcr_def = pcr_def |> Thm.concl_of |> Logic.dest_equals |> fst |> head_of in if is_some (#pcr_info qinfo) then get_pcrel (#pcrel_def (the (#pcr_info qinfo))) else quot_thm_crel (#quot_thm qinfo) end fun pointer_of_bundle_name bundle_name ctxt = let val bundle = Bundle.read ctxt bundle_name fun err () = error "The provided bundle is not a lifting bundle" in (case bundle of [(_, [arg_src])] => let val (name, _) = Token.syntax (Scan.lift Parse.string) arg_src ctxt handle ERROR _ => err () in name end | _ => err ()) end fun pointer_of_bundle_binding ctxt binding = Name_Space.full_name (Name_Space.naming_of (Context.Theory (Proof_Context.theory_of ctxt))) binding fun lifting_forget pointer lthy = let fun get_transfer_rules_to_delete qinfo ctxt = let val transfer_rel = get_transfer_rel qinfo in filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw ctxt) end in case Lifting_Info.lookup_restore_data lthy pointer of SOME restore_info => let val qinfo = #quotient restore_info val quot_thm = #quot_thm qinfo val transfer_rules = get_transfer_rules_to_delete qinfo lthy in Local_Theory.declaration {syntax = false, pervasive = true} (K (fold (Transfer.transfer_raw_del) transfer_rules #> Lifting_Info.delete_quotients quot_thm)) lthy end | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_forget_cmd bundle_name lthy = lifting_forget (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_forget\ "unsetup Lifting and Transfer for the given lifting bundle" (Parse.name_position >> lifting_forget_cmd) (* lifting_update *) fun update_transfer_rules pointer lthy = let fun new_transfer_rules ({ quotient = qinfo, ... }:Lifting_Info.restore_data) lthy = let val transfer_rel = get_transfer_rel qinfo val transfer_rules = filter_transfer_rules_by_rel transfer_rel (Transfer.get_transfer_raw lthy) in fn phi => fold_rev (Item_Net.update o Morphism.thm phi) transfer_rules Thm.item_net end in case Lifting_Info.lookup_restore_data lthy pointer of SOME refresh_data => Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => Lifting_Info.add_transfer_rules_in_restore_data pointer (new_transfer_rules refresh_data lthy phi)) lthy | NONE => error "The lifting bundle refers to non-existent restore data." end fun lifting_update_cmd bundle_name lthy = update_transfer_rules (pointer_of_bundle_name bundle_name lthy) lthy val _ = Outer_Syntax.local_theory \<^command_keyword>\lifting_update\ "add newly introduced transfer rules to a bundle storing the state of Lifting and Transfer" (Parse.name_position >> lifting_update_cmd) end diff --git a/src/HOL/Tools/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,1227 +1,1223 @@ (* 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 (Term_Subst.Vars.dest t_insts) - val intro'' = Thm.instantiate (Term_Subst.TVars.dest T_insts, t_insts') intro + val t_insts' = map rewrite_pat (Vars.dest t_insts) + val intro'' = Thm.instantiate (TVars.dest T_insts, 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')), (dest_TVar uninst_T', Thm.ctyp_of ctxt' T')], [((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 - Vartab.fold (fn (xi, (S, T)) => Term_Subst.TVars.add ((xi, S), T)) - subst Term_Subst.TVars.empty - end + 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 = - Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) + TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt' T)) (subst_of (pred', pred)) []; in Thm.instantiate (instT, []) 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 val outp_pred = - Term_Subst.instantiate (subst_of (inp_pred, pred), Term_Subst.Vars.empty) inp_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 Term_Subst.TFrees.empty t)) + 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/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 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' 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 (Symtab.empty, Symtab.make_set ns) + |> 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/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 (Symtab.make_set tfrees, Symtab.make_set (rnames @ frees)) idx + |> Thm.generalize (Names.make_set tfrees, Names.make_set (rnames @ frees)) idx |> Thm.instantiate (map tinst binsts, 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 (Symtab.make_set tfrees, Symtab.make_set (rnames @ frees)) idx + |> Thm.generalize (Names.make_set tfrees, Names.make_set (rnames @ frees)) idx |> Thm.instantiate (map tinst binsts, 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, []) |> 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 - (Symtab.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Symtab.empty) + (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, []) |> 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 - (Symtab.make_set (map (fst o dest_TFree o Thm.typ_of o snd) instT), Symtab.empty) + (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/choice_specification.ML b/src/HOL/Tools/choice_specification.ML --- a/src/HOL/Tools/choice_specification.ML +++ b/src/HOL/Tools/choice_specification.ML @@ -1,206 +1,206 @@ (* Title: HOL/Tools/choice_specification.ML Author: Sebastian Skalberg, TU Muenchen Package for defining constants by specification. *) signature CHOICE_SPECIFICATION = sig val close_form : term -> term val add_specification: (bstring * xstring * bool) list -> theory * thm -> theory * thm end structure Choice_Specification: CHOICE_SPECIFICATION = struct local fun mk_definitional [] arg = arg | mk_definitional ((thname, cname, covld) :: cos) (thy, thm) = (case HOLogic.dest_Trueprop (Thm.concl_of thm) of Const (\<^const_name>\Ex\, _) $ P => let val ctype = domain_type (type_of P) val cname_full = Sign.intern_const thy cname val cdefname = if thname = "" then Thm.def_name (Long_Name.base_name cname) else thname val def_eq = Logic.mk_equals (Const (cname_full, ctype), HOLogic.choice_const ctype $ P) val (thms, thy') = Global_Theory.add_defs covld [((Binding.name cdefname, def_eq), [])] thy val thm' = [thm,hd thms] MRS @{thm exE_some} in mk_definitional cos (thy',thm') end | _ => raise THM ("Bad specification theorem", 0, [thm])) in fun add_specification cos = mk_definitional cos #> apsnd Drule.export_without_context end (* Collect all intances of constants in term *) fun collect_consts (t $ u, tms) = collect_consts (u, collect_consts (t, tms)) | collect_consts (Abs (_, _, t), tms) = collect_consts (t, tms) | collect_consts (tm as Const _, tms) = insert (op aconv) tm tms | collect_consts (_, tms) = tms (* Complementing Type.varify_global... *) fun unvarify_global t fmap = let val fmap' = map Library.swap fmap fun unthaw (f as (a, S)) = (case AList.lookup (op =) fmap' a of NONE => TVar f | SOME (b, _) => TFree (b, S)) in map_types (map_type_tvar unthaw) t end (* The syntactic meddling needed to setup add_specification for work *) fun close_form t = fold_rev (fn (s, T) => fn t => HOLogic.mk_all (s, T, t)) (map dest_Free (Misc_Legacy.term_frees t)) t fun zip3 [] [] [] = [] | zip3 (x::xs) (y::ys) (z::zs) = (x,y,z)::zip3 xs ys zs | zip3 _ _ _ = raise Fail "Choice_Specification.process_spec internal error" fun myfoldr f [x] = x | myfoldr f (x::xs) = f (x,myfoldr f xs) | myfoldr f [] = raise Fail "Choice_Specification.process_spec internal error" fun process_spec cos alt_props thy = let val ctxt = Proof_Context.init_global thy val rew_imps = alt_props |> map (Object_Logic.atomize ctxt o Thm.cterm_of ctxt o Syntax.read_prop_global thy o snd) val props' = rew_imps |> map (HOLogic.dest_Trueprop o Thm.term_of o snd o Thm.dest_equals o Thm.cprop_of) fun proc_single prop = let val frees = Misc_Legacy.term_frees prop val _ = forall (fn v => Sign.of_sort thy (type_of v, \<^sort>\type\)) frees orelse error "Specificaton: Only free variables of sort 'type' allowed" val prop_closed = close_form prop in (prop_closed, frees) end val props'' = map proc_single props' val frees = map snd props'' val prop = myfoldr HOLogic.mk_conj (map fst props'') - val (vmap, prop_thawed) = Type.varify_global Term_Subst.TFrees.empty prop + val (vmap, prop_thawed) = Type.varify_global TFrees.empty prop val thawed_prop_consts = collect_consts (prop_thawed, []) val (altcos, overloaded) = split_list cos val (names, sconsts) = split_list altcos val consts = map (Syntax.read_term_global thy) sconsts val _ = not (Library.exists (not o Term.is_Const) consts) orelse error "Specification: Non-constant found as parameter" fun proc_const c = let - val (_, c') = Type.varify_global Term_Subst.TFrees.empty c + val (_, c') = Type.varify_global TFrees.empty c val (cname,ctyp) = dest_Const c' in (case filter (fn t => let val (name, typ) = dest_Const t in name = cname andalso Sign.typ_equiv thy (typ, ctyp) end) thawed_prop_consts of [] => error ("Specification: No suitable instances of constant \"" ^ Syntax.string_of_term_global thy c ^ "\" found") | [cf] => unvarify_global cf vmap | _ => error ("Specification: Several variations of \"" ^ Syntax.string_of_term_global thy c ^ "\" found (try applying explicit type constraints)")) end val proc_consts = map proc_const consts fun mk_exist c prop = let val T = type_of c val cname = Long_Name.base_name (fst (dest_Const c)) val vname = if Symbol_Pos.is_identifier cname then cname else "x" in HOLogic.exists_const T $ Abs(vname,T,Term.abstract_over (c, prop)) end val ex_prop = fold_rev mk_exist proc_consts prop val cnames = map (fst o dest_Const) proc_consts fun post_process (arg as (thy,thm)) = let fun inst_all thy v thm = let val cv = Thm.global_cterm_of thy v val cT = Thm.ctyp_of_cterm cv val spec' = Thm.instantiate' [SOME cT] [NONE, SOME cv] spec in thm RS spec' end fun remove_alls frees (thm, thy) = (fold (inst_all thy) frees thm, thy) fun process_single ((name, atts), rew_imp, frees) args = let fun undo_imps thm = Thm.equal_elim (Thm.symmetric rew_imp) thm fun add_final (thm, thy) = if name = "" then (thm, thy) else (writeln (" " ^ name ^ ": " ^ Thm.string_of_thm_global thy thm); Global_Theory.store_thm (Binding.name name, thm) thy) in swap args |> remove_alls frees |> apfst undo_imps |> apfst Drule.export_without_context |-> Thm.theory_attributes (map (Attrib.attribute_cmd_global thy) (@{attributes [nitpick_choice_spec]} @ atts)) |> add_final |> swap end fun process_all [proc_arg] args = process_single proc_arg args | process_all (proc_arg::rest) (thy,thm) = let val single_th = thm RS conjunct1 val rest_th = thm RS conjunct2 val (thy', _) = process_single proc_arg (thy, single_th) in process_all rest (thy', rest_th) end | process_all [] _ = raise Fail "Choice_Specification.process_spec internal error" val alt_names = map fst alt_props val _ = if exists (fn (name, _) => name <> "") alt_names then writeln "specification" else () in arg |> apsnd (Thm.unvarify_global thy) |> process_all (zip3 alt_names rew_imps frees) end fun after_qed [[thm]] = Proof_Context.background_theory (fn thy => #1 (post_process (add_specification (zip3 names cnames overloaded) (thy, thm)))); in thy |> Proof_Context.init_global |> Variable.declare_term ex_prop |> Proof.theorem NONE after_qed [[(HOLogic.mk_Trueprop ex_prop, [])]] end (* outer syntax *) val opt_name = Scan.optional (Parse.name --| \<^keyword>\:\) "" val opt_overloaded = Parse.opt_keyword "overloaded" val _ = Outer_Syntax.command \<^command_keyword>\specification\ "define constants by specification" (\<^keyword>\(\ |-- Scan.repeat1 (opt_name -- Parse.term -- opt_overloaded) --| \<^keyword>\)\ -- Scan.repeat1 ((Parse_Spec.opt_thm_name ":" >> apfst Binding.name_of) -- Parse.prop) >> (fn (cos, alt_props) => Toplevel.theory_to_proof (process_spec cos alt_props))) end diff --git a/src/Pure/General/table.ML b/src/Pure/General/table.ML --- a/src/Pure/General/table.ML +++ b/src/Pure/General/table.ML @@ -1,479 +1,479 @@ (* Title: Pure/General/table.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen Generic tables. Efficient purely functional implementation using balanced 2-3 trees. *) signature KEY = sig type key val ord: key ord end; signature TABLE = sig type key type 'a table exception DUP of key exception SAME exception UNDEF of key val empty: 'a table val build: ('a table -> 'a table) -> 'a table val is_empty: 'a table -> bool val is_single: 'a table -> bool val map: (key -> 'a -> 'b) -> 'a table -> 'b table val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a val fold_rev: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a val size: 'a table -> int val dest: 'a table -> (key * 'a) list val keys: 'a table -> key list val min: 'a table -> (key * 'a) option val max: 'a table -> (key * 'a) option val exists: (key * 'a -> bool) -> 'a table -> bool val forall: (key * 'a -> bool) -> 'a table -> bool val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option val lookup_key: 'a table -> key -> (key * 'a) option val lookup: 'a table -> key -> 'a option val defined: 'a table -> key -> bool val update: key * 'a -> 'a table -> 'a table val update_new: key * 'a -> 'a table -> 'a table (*exception DUP*) val default: key * 'a -> 'a table -> 'a table val map_entry: key -> ('a -> 'a) (*exception SAME*) -> 'a table -> 'a table val map_default: key * 'a -> ('a -> 'a) -> 'a table -> 'a table val make: (key * 'a) list -> 'a table (*exception DUP*) val join: (key -> 'a * 'a -> 'a) (*exception SAME*) -> 'a table * 'a table -> 'a table (*exception DUP*) val merge: ('a * 'a -> bool) -> 'a table * 'a table -> 'a table (*exception DUP*) val delete: key -> 'a table -> 'a table (*exception UNDEF*) val delete_safe: key -> 'a table -> 'a table val member: ('b * 'a -> bool) -> 'a table -> key * 'b -> bool val insert: ('a * 'a -> bool) -> key * 'a -> 'a table -> 'a table (*exception DUP*) val remove: ('b * 'a -> bool) -> key * 'b -> 'a table -> 'a table val lookup_list: 'a list table -> key -> 'a list val cons_list: key * 'a -> 'a list table -> 'a list table val insert_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table val remove_list: ('b * 'a -> bool) -> key * 'b -> 'a list table -> 'a list table val update_list: ('a * 'a -> bool) -> key * 'a -> 'a list table -> 'a list table val make_list: (key * 'a) list -> 'a list table val dest_list: 'a list table -> (key * 'a) list val merge_list: ('a * 'a -> bool) -> 'a list table * 'a list table -> 'a list table type set = unit table val insert_set: key -> set -> set val remove_set: key -> set -> set val make_set: key list -> set val subset: set * set -> bool val eq_set: set * set -> bool end; functor Table(Key: KEY): TABLE = struct (* datatype table *) type key = Key.key; datatype 'a table = Empty | Branch2 of 'a table * (key * 'a) * 'a table | Branch3 of 'a table * (key * 'a) * 'a table * (key * 'a) * 'a table; exception DUP of key; (* empty and single *) val empty = Empty; fun build (f: 'a table -> 'a table) = f empty; fun is_empty Empty = true | is_empty _ = false; fun is_single (Branch2 (Empty, _, Empty)) = true | is_single _ = false; (* map and fold combinators *) fun map_table f = let fun map Empty = Empty | map (Branch2 (left, (k, x), right)) = Branch2 (map left, (k, f k x), map right) | map (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = Branch3 (map left, (k1, f k1 x1), map mid, (k2, f k2 x2), map right); in map end; fun fold_table f = let fun fold Empty x = x | fold (Branch2 (left, p, right)) x = fold right (f p (fold left x)) | fold (Branch3 (left, p1, mid, p2, right)) x = fold right (f p2 (fold mid (f p1 (fold left x)))); in fold end; fun fold_rev_table f = let fun fold Empty x = x | fold (Branch2 (left, p, right)) x = fold left (f p (fold right x)) | fold (Branch3 (left, p1, mid, p2, right)) x = fold left (f p1 (fold mid (f p2 (fold right x)))); in fold end; fun size tab = fold_table (fn _ => fn n => n + 1) tab 0; fun dest tab = fold_rev_table cons tab []; fun keys tab = fold_rev_table (cons o #1) tab []; (* min/max entries *) fun min Empty = NONE | min (Branch2 (Empty, p, _)) = SOME p | min (Branch3 (Empty, p, _, _, _)) = SOME p | min (Branch2 (left, _, _)) = min left | min (Branch3 (left, _, _, _, _)) = min left; fun max Empty = NONE | max (Branch2 (_, p, Empty)) = SOME p | max (Branch3 (_, _, _, p, Empty)) = SOME p | max (Branch2 (_, _, right)) = max right | max (Branch3 (_, _, _, _, right)) = max right; (* exists and forall *) fun exists pred = let fun ex Empty = false | ex (Branch2 (left, (k, x), right)) = ex left orelse pred (k, x) orelse ex right | ex (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = ex left orelse pred (k1, x1) orelse ex mid orelse pred (k2, x2) orelse ex right; in ex end; fun forall pred = not o exists (not o pred); (* get_first *) fun get_first f = let fun get Empty = NONE | get (Branch2 (left, (k, x), right)) = (case get left of NONE => (case f (k, x) of NONE => get right | some => some) | some => some) | get (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case get left of NONE => (case f (k1, x1) of NONE => (case get mid of NONE => (case f (k2, x2) of NONE => get right | some => some) | some => some) | some => some) | some => some); in get end; (* lookup *) fun lookup tab key = let fun look Empty = NONE | look (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => look left | EQUAL => SOME x | GREATER => look right) | look (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => look left | EQUAL => SOME x1 | GREATER => (case Key.ord (key, k2) of LESS => look mid | EQUAL => SOME x2 | GREATER => look right)); in look tab end; fun lookup_key tab key = let fun look Empty = NONE | look (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => look left | EQUAL => SOME (k, x) | GREATER => look right) | look (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => look left | EQUAL => SOME (k1, x1) | GREATER => (case Key.ord (key, k2) of LESS => look mid | EQUAL => SOME (k2, x2) | GREATER => look right)); in look tab end; fun defined tab key = let fun def Empty = false | def (Branch2 (left, (k, x), right)) = (case Key.ord (key, k) of LESS => def left | EQUAL => true | GREATER => def right) | def (Branch3 (left, (k1, x1), mid, (k2, x2), right)) = (case Key.ord (key, k1) of LESS => def left | EQUAL => true | GREATER => (case Key.ord (key, k2) of LESS => def mid | EQUAL => true | GREATER => def right)); in def tab end; (* modify *) datatype 'a growth = Stay of 'a table | Sprout of 'a table * (key * 'a) * 'a table; exception SAME; fun modify key f tab = let fun modfy Empty = Sprout (Empty, (key, f NONE), Empty) | modfy (Branch2 (left, p as (k, x), right)) = (case Key.ord (key, k) of LESS => (case modfy left of Stay left' => Stay (Branch2 (left', p, right)) | Sprout (left1, q, left2) => Stay (Branch3 (left1, q, left2, p, right))) | EQUAL => Stay (Branch2 (left, (k, f (SOME x)), right)) | GREATER => (case modfy right of Stay right' => Stay (Branch2 (left, p, right')) | Sprout (right1, q, right2) => Stay (Branch3 (left, p, right1, q, right2)))) | modfy (Branch3 (left, p1 as (k1, x1), mid, p2 as (k2, x2), right)) = (case Key.ord (key, k1) of LESS => (case modfy left of Stay left' => Stay (Branch3 (left', p1, mid, p2, right)) | Sprout (left1, q, left2) => Sprout (Branch2 (left1, q, left2), p1, Branch2 (mid, p2, right))) | EQUAL => Stay (Branch3 (left, (k1, f (SOME x1)), mid, p2, right)) | GREATER => (case Key.ord (key, k2) of LESS => (case modfy mid of Stay mid' => Stay (Branch3 (left, p1, mid', p2, right)) | Sprout (mid1, q, mid2) => Sprout (Branch2 (left, p1, mid1), q, Branch2 (mid2, p2, right))) | EQUAL => Stay (Branch3 (left, p1, mid, (k2, f (SOME x2)), right)) | GREATER => (case modfy right of Stay right' => Stay (Branch3 (left, p1, mid, p2, right')) | Sprout (right1, q, right2) => Sprout (Branch2 (left, p1, mid), p2, Branch2 (right1, q, right2))))); in (case modfy tab of Stay tab' => tab' | Sprout br => Branch2 br) handle SAME => tab end; fun update (key, x) tab = modify key (fn _ => x) tab; fun update_new (key, x) tab = modify key (fn NONE => x | SOME _ => raise DUP key) tab; fun default (key, x) tab = modify key (fn NONE => x | SOME _ => raise SAME) tab; fun map_entry key f = modify key (fn NONE => raise SAME | SOME x => f x); fun map_default (key, x) f = modify key (fn NONE => f x | SOME y => f y); (* delete *) exception UNDEF of key; local fun compare NONE (k2, _) = LESS | compare (SOME k1) (k2, _) = Key.ord (k1, k2); fun if_eq EQUAL x y = x | if_eq _ x y = y; fun del (SOME k) Empty = raise UNDEF k | del NONE (Branch2 (Empty, p, Empty)) = (p, (true, Empty)) | del NONE (Branch3 (Empty, p, Empty, q, Empty)) = (p, (false, Branch2 (Empty, q, Empty))) | del k (Branch2 (Empty, p, Empty)) = (case compare k p of EQUAL => (p, (true, Empty)) | _ => raise UNDEF (the k)) | del k (Branch3 (Empty, p, Empty, q, Empty)) = (case compare k p of EQUAL => (p, (false, Branch2 (Empty, q, Empty))) | _ => (case compare k q of EQUAL => (q, (false, Branch2 (Empty, p, Empty))) | _ => raise UNDEF (the k))) | del k (Branch2 (l, p, r)) = (case compare k p of LESS => (case del k l of (p', (false, l')) => (p', (false, Branch2 (l', p, r))) | (p', (true, l')) => (p', case r of Branch2 (rl, rp, rr) => (true, Branch3 (l', p, rl, rp, rr)) | Branch3 (rl, rp, rm, rq, rr) => (false, Branch2 (Branch2 (l', p, rl), rp, Branch2 (rm, rq, rr))))) | ord => (case del (if_eq ord NONE k) r of (p', (false, r')) => (p', (false, Branch2 (l, if_eq ord p' p, r'))) | (p', (true, r')) => (p', case l of Branch2 (ll, lp, lr) => (true, Branch3 (ll, lp, lr, if_eq ord p' p, r')) | Branch3 (ll, lp, lm, lq, lr) => (false, Branch2 (Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, r')))))) | del k (Branch3 (l, p, m, q, r)) = (case compare k q of LESS => (case compare k p of LESS => (case del k l of (p', (false, l')) => (p', (false, Branch3 (l', p, m, q, r))) | (p', (true, l')) => (p', (false, case (m, r) of (Branch2 (ml, mp, mr), Branch2 _) => Branch2 (Branch3 (l', p, ml, mp, mr), q, r) | (Branch3 (ml, mp, mm, mq, mr), _) => Branch3 (Branch2 (l', p, ml), mp, Branch2 (mm, mq, mr), q, r) | (Branch2 (ml, mp, mr), Branch3 (rl, rp, rm, rq, rr)) => Branch3 (Branch2 (l', p, ml), mp, Branch2 (mr, q, rl), rp, Branch2 (rm, rq, rr))))) | ord => (case del (if_eq ord NONE k) m of (p', (false, m')) => (p', (false, Branch3 (l, if_eq ord p' p, m', q, r))) | (p', (true, m')) => (p', (false, case (l, r) of (Branch2 (ll, lp, lr), Branch2 _) => Branch2 (Branch3 (ll, lp, lr, if_eq ord p' p, m'), q, r) | (Branch3 (ll, lp, lm, lq, lr), _) => Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, if_eq ord p' p, m'), q, r) | (_, Branch3 (rl, rp, rm, rq, rr)) => Branch3 (l, if_eq ord p' p, Branch2 (m', q, rl), rp, Branch2 (rm, rq, rr)))))) | ord => (case del (if_eq ord NONE k) r of (q', (false, r')) => (q', (false, Branch3 (l, p, m, if_eq ord q' q, r'))) | (q', (true, r')) => (q', (false, case (l, m) of (Branch2 _, Branch2 (ml, mp, mr)) => Branch2 (l, p, Branch3 (ml, mp, mr, if_eq ord q' q, r')) | (_, Branch3 (ml, mp, mm, mq, mr)) => Branch3 (l, p, Branch2 (ml, mp, mm), mq, Branch2 (mr, if_eq ord q' q, r')) | (Branch3 (ll, lp, lm, lq, lr), Branch2 (ml, mp, mr)) => Branch3 (Branch2 (ll, lp, lm), lq, Branch2 (lr, p, ml), mp, Branch2 (mr, if_eq ord q' q, r')))))); in fun delete key tab = snd (snd (del (SOME key) tab)); fun delete_safe key tab = if defined tab key then delete key tab else tab; end; (* membership operations *) fun member eq tab (key, x) = (case lookup tab key of NONE => false | SOME y => eq (x, y)); fun insert eq (key, x) = modify key (fn NONE => x | SOME y => if eq (x, y) then raise SAME else raise DUP key); fun remove eq (key, x) tab = (case lookup tab key of NONE => tab | SOME y => if eq (x, y) then delete key tab else tab); (* simultaneous modifications *) -fun make entries = fold update_new entries empty; +fun make entries = build (fold update_new entries); fun join f (table1, table2) = let fun add (key, y) tab = modify key (fn NONE => y | SOME x => f key (x, y)) tab; in if pointer_eq (table1, table2) then table1 else if is_empty table1 then table2 else fold_table add table2 table1 end; fun merge eq = join (fn key => fn xy => if eq xy then raise SAME else raise DUP key); (* list tables *) fun lookup_list tab key = these (lookup tab key); fun cons_list (key, x) tab = modify key (fn NONE => [x] | SOME xs => x :: xs) tab; fun insert_list eq (key, x) = modify key (fn NONE => [x] | SOME xs => if Library.member eq xs x then raise SAME else x :: xs); fun remove_list eq (key, x) tab = map_entry key (fn xs => (case Library.remove eq x xs of [] => raise UNDEF key | ys => ys)) tab handle UNDEF _ => delete key tab; fun update_list eq (key, x) = modify key (fn NONE => [x] | SOME [] => [x] | SOME (xs as y :: _) => if eq (x, y) then raise SAME else Library.update eq x xs); -fun make_list args = fold_rev cons_list args empty; +fun make_list args = build (fold_rev cons_list args); fun dest_list tab = maps (fn (key, xs) => map (pair key) xs) (dest tab); fun merge_list eq = join (fn _ => Library.merge eq); (* set operations *) type set = unit table; fun insert_set x = default (x, ()); fun remove_set x : set -> set = delete_safe x; -fun make_set entries = fold insert_set entries empty; +fun make_set xs = build (fold insert_set xs); fun subset (A: set, B: set) = forall (defined B o #1) A; fun eq_set (A: set, B: set) = size A = size B andalso subset (A, B); (* ML pretty-printing *) val _ = ML_system_pp (fn depth => fn pretty => fn tab => ML_Pretty.to_polyml (ML_Pretty.enum "," "{" "}" (ML_Pretty.pair (ML_Pretty.from_polyml o ML_system_pretty) (ML_Pretty.from_polyml o pretty)) (dest tab, depth))); (*final declarations of this structure!*) val map = map_table; val fold = fold_table; val fold_rev = fold_rev_table; end; structure Inttab = Table(type key = int val ord = int_ord); structure Symtab = Table(type key = string val ord = fast_string_ord); structure Symreltab = Table(type key = string * string val ord = prod_ord fast_string_ord fast_string_ord); 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,878 +1,878 @@ (* 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)); val cert_inst = ((map #1 params ~~ - map (Term_Subst.instantiateT_frees (Term_Subst.TFrees.table instT)) parm_types) ~~ insts'') + 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)); in (Element.instantiate_normalize_morphism (map (apsnd 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 (xs, 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 [] => (Term.add_frees b' xs, (Free (y, T), b') :: env, eq :: eqs) | dups => if forall (fn (_, b'') => b' aconv b'') dups then (xs, 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')), (xs, 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', (fold Term.add_frees ts' xs, 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 = Term.add_free_names body []; val xs = filter (member (op =) env o #1) parms; val Ts = map #2 xs; val extraTs = (subtract (op =) (fold Term.add_tfreesT Ts []) (Term.add_tfrees body [])) |> sort_by #1 |> map TFree; val predT = map Term.itselfT extraTs ---> Ts ---> bodyT; val args = map Logic.mk_type extraTs @ 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 extraTs = subtract (op =) (fold Term.add_tfreesT (map snd parms) []) (fold Term.add_tfrees exts' []); val _ = if null extraTs then () else warning ("Additional type variable(s) in locale specification " ^ Binding.print binding ^ ": " ^ commas (map (Syntax.string_of_typ ctxt' o TFree) (sort_by #1 extraTs))); 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 (extraTs, 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,473 +1,472 @@ (* 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 = Term_Subst.TFrees.build (Term_Subst.add_tfreesT (Term.fastype_of u)); + val type_tfrees = TFrees.build (TFrees.add_tfreesT (Term.fastype_of u)); val extra_tfrees = build_rev (u |> (Term.fold_types o Term.fold_atyps) - (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); + (fn TFree v => not (TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); 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 = - Term_Subst.TFrees.build (Term_Subst.add_tfreesT T #> fold (Term_Subst.add_tfreesT o #2) xs); + val type_tfrees = TFrees.build (TFrees.add_tfreesT T #> fold (TFrees.add_tfreesT o #2) xs); val extra_tfrees = build_rev (rhs |> (Term.fold_types o Term.fold_atyps) - (fn TFree v => not (Term_Subst.TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); + (fn TFree v => not (TFrees.defined type_tfrees v) ? insert (op =) v | _ => I)); 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 = map TFree (Thm.fold_terms {hyps = true} Term.add_tfrees th' []); val frees = map Free (Thm.fold_terms {hyps = true} Term.add_frees th' []); 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 = - Term_Subst.TVars.build (fold2 (fn a => fn b => - (case a of TVar v => Term_Subst.TVars.add (v, b) | _ => I)) tvars tfrees); - val cinstT = Term_Subst.TVars.fold (fn (v, T) => cons (v, Thm.ctyp_of ctxt T)) 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 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); val result' = Thm.instantiate (cinstT, cinst) result; (*import assumes/defines*) val result'' = (fold (curry op COMP) asms' result' handle THM _ => raise THM ("Failed to re-import result", 0, result' :: asms')) |> Local_Defs.contract ctxt defs (Thm.cprop_of th) |> Goal.norm_result ctxt |> Global_Theory.name_thm Global_Theory.unofficial2 name; in (result'', result) end; fun bind_name lthy b = (Local_Theory.full_name lthy b, Binding.default_pos_of b); fun map_facts f = map (apsnd (map (apfst (map f)))); in fun notes target_notes kind facts lthy = let val facts' = facts |> map (fn (a, bs) => (a, Global_Theory.burrow_fact (Global_Theory.name_multi (bind_name lthy (fst a))) bs)) |> map_facts (import_export_proof lthy); val local_facts = map_facts #1 facts'; val global_facts = map_facts #2 facts'; in lthy |> target_notes kind global_facts (Attrib.partial_evaluation lthy local_facts) |> Attrib.local_notes kind local_facts end; end; (* abbrev *) fun abbrev target_abbrev prmode ((b, mx), rhs) lthy = let val (global_rhs, (extra_tfrees, (type_params, term_params))) = export_abbrev lthy I rhs; val mx' = check_mixfix lthy (b, extra_tfrees) mx; in lthy |> target_abbrev prmode (b, mx') global_rhs (type_params, term_params) |> Context_Position.set_visible false |> Proof_Context.add_abbrev Print_Mode.internal (b, rhs) |> snd |> Local_Defs.fixed_abbrev ((b, NoSyn), rhs) ||> Context_Position.restore_visible lthy end; (** theory target primitives **) fun theory_target_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) = background_foundation (((b, U), mx), (b_def, rhs)) (type_params, term_params) #-> (fn (lhs, def) => standard_const (op <>) Syntax.mode_default ((b, mx), lhs) #> pair (lhs, def)); fun theory_target_notes kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind global_facts #> snd) #> standard_notes (op <>) kind local_facts; fun theory_target_abbrev prmode (b, mx) global_rhs params = Local_Theory.background_theory_result (Sign.add_abbrev (#1 prmode) (b, global_rhs) #-> (fn (lhs, _) => (* FIXME type_params!? *) Sign.notation true prmode [(lhs, check_mixfix_global (b, null (snd params)) mx)] #> pair lhs)) #-> (fn lhs => standard_const (op <>) prmode ((b, if null (snd params) then NoSyn else mx), Term.list_comb (Logic.unvarify_global lhs, snd params))); (** theory operations **) val theory_abbrev = abbrev theory_target_abbrev; fun theory_declaration decl = background_declaration decl #> standard_declaration (K true) decl; fun target_registration lthy {inst, mixin, export} = {inst = inst, mixin = mixin, export = export $> Proof_Context.export_morphism lthy (Local_Theory.target_of lthy)}; fun theory_registration registration lthy = lthy |> (Local_Theory.raw_theory o Context.theory_map) (Locale.add_registration (target_registration lthy registration)) |> standard_registration (K true) registration; (** locale target primitives **) fun locale_target_notes locale kind global_facts local_facts = Local_Theory.background_theory (Attrib.global_notes kind (Attrib.map_facts (K []) global_facts) #> snd) #> (fn lthy => lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_facts locale kind (standard_facts lthy ctxt local_facts))) #> standard_notes (fn (this, other) => other <> 0 andalso this <> other) kind local_facts; fun locale_target_declaration locale syntax decl lthy = lthy |> Local_Theory.target (fn ctxt => ctxt |> Locale.add_declaration locale syntax (Morphism.transform (Local_Theory.standard_morphism lthy ctxt) decl)); fun locale_target_const locale phi_pred prmode ((b, mx), rhs) = locale_target_declaration locale true (const_decl phi_pred prmode ((b, mx), rhs)) (** locale operations **) fun locale_declaration locale {syntax, pervasive} decl = pervasive ? background_declaration decl #> locale_target_declaration locale syntax decl #> standard_declaration (fn (_, other) => other <> 0) decl; fun locale_const locale prmode ((b, mx), rhs) = locale_target_const locale (K true) prmode ((b, mx), rhs) #> standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), rhs); fun locale_dependency loc registration lthy = lthy |> Local_Theory.raw_theory (Locale.add_dependency loc registration) |> standard_registration (K true) registration; (** locale abbreviations **) fun locale_target_abbrev locale prmode (b, mx) global_rhs params = background_abbrev (b, global_rhs) (snd params) #-> (fn (lhs, _) => locale_const locale prmode ((b, mx), lhs)); fun locale_abbrev locale = abbrev (locale_target_abbrev locale); end; diff --git a/src/Pure/Isar/local_defs.ML b/src/Pure/Isar/local_defs.ML --- a/src/Pure/Isar/local_defs.ML +++ b/src/Pure/Isar/local_defs.ML @@ -1,271 +1,271 @@ (* Title: Pure/Isar/local_defs.ML Author: Makarius Local definitions. *) signature LOCAL_DEFS = sig val cert_def: Proof.context -> (string -> Position.T list) -> term -> (string * typ) * term val abs_def: term -> (string * typ) * term val expand: cterm list -> thm -> thm val def_export: Assumption.export val define: ((binding * mixfix) * (Thm.binding * term)) list -> Proof.context -> (term * (string * thm)) list * Proof.context val fixed_abbrev: (binding * mixfix) * term -> Proof.context -> (term * term) * Proof.context val export: Proof.context -> Proof.context -> thm -> (thm list * thm list) * thm val export_cterm: Proof.context -> Proof.context -> cterm -> (thm list * thm list) * cterm val contract: Proof.context -> thm list -> cterm -> thm -> thm val print_rules: Proof.context -> unit val defn_add: attribute val defn_del: attribute val meta_rewrite_conv: Proof.context -> conv val meta_rewrite_rule: Proof.context -> thm -> thm val abs_def_rule: Proof.context -> thm -> thm val unfold_abs_def: bool Config.T val unfold: Proof.context -> thm list -> thm -> thm val unfold_goals: Proof.context -> thm list -> thm -> thm val unfold_tac: Proof.context -> thm list -> tactic val unfold0: Proof.context -> thm list -> thm -> thm val unfold0_goals: Proof.context -> thm list -> thm -> thm val unfold0_tac: Proof.context -> thm list -> tactic val fold: Proof.context -> thm list -> thm -> thm val fold_tac: Proof.context -> thm list -> tactic val derived_def: Proof.context -> (string -> Position.T list) -> {conditional: bool} -> term -> ((string * typ) * term) * (Proof.context -> thm -> thm) end; structure Local_Defs: LOCAL_DEFS = struct (** primitive definitions **) (* prepare defs *) fun cert_def ctxt get_pos eq = let fun err msg = cat_error msg ("The error(s) above occurred in definition:\n" ^ quote (Syntax.string_of_term ctxt eq)); val ((lhs, _), args, eq') = eq |> Sign.no_vars ctxt |> Primitive_Defs.dest_def ctxt {check_head = Term.is_Free, check_free_lhs = not o Variable.is_fixed ctxt, check_free_rhs = if Variable.is_body ctxt then K true else Variable.is_fixed ctxt, check_tfree = K true} handle TERM (msg, _) => err msg | ERROR msg => err msg; val _ = Context_Position.reports ctxt (maps (fn Free (x, _) => Syntax_Phases.reports_of_scope (get_pos x) | _ => []) args); in (Term.dest_Free (Term.head_of lhs), eq') end; val abs_def = Primitive_Defs.abs_def #>> Term.dest_Free; fun mk_def ctxt args = let val (bs, rhss) = split_list args; val Ts = map Term.fastype_of rhss; val (xs, _) = ctxt |> Context_Position.set_visible false |> Proof_Context.add_fixes (map2 (fn b => fn T => (b, SOME T, NoSyn)) bs Ts); val lhss = ListPair.map Free (xs, Ts); in map Logic.mk_equals (lhss ~~ rhss) end; (* export defs *) val head_of_def = Term.dest_Free o Term.head_of o #1 o Logic.dest_equals o Term.strip_all_body; (* [x, x \ a] : B x ----------- B a *) fun expand defs = Drule.implies_intr_list defs #> Drule.generalize - (Symtab.empty, fold (Symtab.insert_set o #1 o head_of_def o Thm.term_of) defs Symtab.empty) + (Names.empty, Names.build (fold (Names.add_set o #1 o head_of_def o Thm.term_of) defs)) #> funpow (length defs) (fn th => Drule.reflexive_thm RS th); val expand_term = Envir.expand_term_frees o map (abs_def o Thm.term_of); fun def_export _ defs = (expand defs, expand_term defs); (* define *) fun define defs ctxt = let val ((xs, mxs), specs) = defs |> split_list |>> split_list; val (bs, rhss) = specs |> split_list; val eqs = mk_def ctxt (xs ~~ rhss); val lhss = map (fst o Logic.dest_equals) eqs; in ctxt |> Proof_Context.add_fixes (map2 (fn x => fn mx => (x, NONE, mx)) xs mxs) |> #2 |> fold Variable.declare_term eqs |> Proof_Context.add_assms def_export (map2 (fn b => fn eq => (b, [(eq, [])])) bs eqs) |>> map2 (fn lhs => fn (name, [th]) => (lhs, (name, th))) lhss end; (* fixed_abbrev *) fun fixed_abbrev ((x, mx), rhs) ctxt = let val T = Term.fastype_of rhs; val ([x'], ctxt') = ctxt |> Variable.declare_term rhs |> Proof_Context.add_fixes [(x, SOME T, mx)]; val lhs = Free (x', T); val _ = cert_def ctxt' (K []) (Logic.mk_equals (lhs, rhs)); fun abbrev_export _ _ = (I, Envir.expand_term_frees [((x', T), rhs)]); val (_, ctxt'') = Assumption.add_assms abbrev_export [] ctxt'; in ((lhs, rhs), ctxt'') end; (* specific export -- result based on educated guessing *) (* [xs, xs \ as] : B xs -------------- B as *) fun export inner outer th = let val defs_asms = Assumption.local_assms_of inner outer |> filter_out (Drule.is_sort_constraint o Thm.term_of) |> map (Thm.assume #> (fn asm => (case try (head_of_def o Thm.prop_of) asm of NONE => (asm, false) | SOME x => let val t = Free x in (case try (Assumption.export_term inner outer) t of NONE => (asm, false) | SOME u => if t aconv u then (asm, false) else (Drule.abs_def (Variable.gen_all outer asm), true)) end))); in (apply2 (map #1) (List.partition #2 defs_asms), Assumption.export false inner outer th) end; (* [xs, xs \ as] : TERM b xs -------------- and -------------- TERM b as b xs \ b as *) fun export_cterm inner outer ct = export inner outer (Drule.mk_term ct) ||> Drule.dest_term; fun contract ctxt defs ct th = th COMP (Raw_Simplifier.rewrite ctxt true defs ct COMP_INCR Drule.equal_elim_rule2); (** defived definitions **) (* transformation via rewrite rules *) structure Rules = Generic_Data ( type T = thm list; val empty = []; val extend = I; val merge = Thm.merge_thms; ); fun print_rules ctxt = Pretty.writeln (Pretty.big_list "definitional rewrite rules:" (map (Thm.pretty_thm_item ctxt) (Rules.get (Context.Proof ctxt)))); val defn_add = Thm.declaration_attribute (Rules.map o Thm.add_thm o Thm.trim_context); val defn_del = Thm.declaration_attribute (Rules.map o Thm.del_thm); (* meta rewrite rules *) fun meta_rewrite_conv ctxt = Raw_Simplifier.rewrite_cterm (false, false, false) (K (K NONE)) (ctxt |> Raw_Simplifier.init_simpset (Rules.get (Context.Proof ctxt)) |> Raw_Simplifier.add_eqcong Drule.equals_cong); (*protect meta-level equality*) val meta_rewrite_rule = Conv.fconv_rule o meta_rewrite_conv; fun abs_def_rule ctxt = meta_rewrite_rule ctxt #> Drule.abs_def; (* unfold object-level rules *) val unfold_abs_def = Config.declare_bool ("unfold_abs_def", \<^here>) (K true); local fun gen_unfold rewrite ctxt rews = let val meta_rews = map (meta_rewrite_rule ctxt) rews in if Config.get ctxt unfold_abs_def then rewrite ctxt meta_rews #> rewrite ctxt (map (perhaps (try Drule.abs_def)) meta_rews) else rewrite ctxt meta_rews end; val no_unfold_abs_def = Config.put unfold_abs_def false; in val unfold = gen_unfold Raw_Simplifier.rewrite_rule; val unfold_goals = gen_unfold Raw_Simplifier.rewrite_goals_rule; val unfold_tac = PRIMITIVE oo unfold_goals; val unfold0 = unfold o no_unfold_abs_def; val unfold0_goals = unfold_goals o no_unfold_abs_def; val unfold0_tac = unfold_tac o no_unfold_abs_def; end (* fold object-level rules *) fun fold ctxt rews = Raw_Simplifier.fold_rule ctxt (map (meta_rewrite_rule ctxt) rews); fun fold_tac ctxt rews = Raw_Simplifier.fold_goals_tac ctxt (map (meta_rewrite_rule ctxt) rews); (* derived defs -- potentially within the object-logic *) fun derived_def ctxt get_pos {conditional} prop = let val ((c, T), rhs) = prop |> Thm.cterm_of ctxt |> meta_rewrite_conv ctxt |> (snd o Logic.dest_equals o Thm.prop_of) |> conditional ? Logic.strip_imp_concl |> (abs_def o #2 o cert_def ctxt get_pos); fun prove def_ctxt0 def = let val def_ctxt = Proof_Context.augment prop def_ctxt0; val def_thm = Goal.prove def_ctxt [] [] prop (fn {context = goal_ctxt, ...} => ALLGOALS (CONVERSION (meta_rewrite_conv goal_ctxt) THEN' rewrite_goal_tac goal_ctxt [def] THEN' resolve_tac goal_ctxt [Drule.reflexive_thm])) handle ERROR msg => cat_error msg "Failed to prove definitional specification"; in def_thm |> singleton (Proof_Context.export def_ctxt def_ctxt0) |> Drule.zero_var_indexes end; in (((c, T), rhs), prove) end; 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,427 +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 = - Term_Subst.Frees.build (t |> Term.fold_aterms - (fn Free v => Term_Subst.Frees.add (v, ()) | _ => I)); + val frees = Frees.build (t |> Term.fold_aterms (fn Free v => Frees.add_set v | _ => I)); val T = - (case Term_Subst.Frees.get_first (fn ((x, T), ()) => if x = z then SOME T else NONE) frees of + (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 = - Term_Subst.TVars.build + TVars.build (params |> fold (#2 #> Term.fold_atyps (fn T => fn instT => (case T of TVar v => - if Term_Subst.TVars.defined instT v then instT - else Term_Subst.TVars.update (v, Thm.ctyp_of ctxt (norm_type T)) instT + 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 (Term_Subst.TVars.dest instT, []); + |> Thm.instantiate (TVars.dest instT, []); 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/specification.ML b/src/Pure/Isar/specification.ML --- a/src/Pure/Isar/specification.ML +++ b/src/Pure/Isar/specification.ML @@ -1,478 +1,478 @@ (* Title: Pure/Isar/specification.ML Author: Makarius Derived local theory specifications --- with type-inference and toplevel polymorphism. *) signature SPECIFICATION = sig val read_props: string list -> (binding * string option * mixfix) list -> Proof.context -> term list * Proof.context val check_spec_open: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> term -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context val read_spec_open: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> string -> Proof.context -> ((binding * typ option * mixfix) list * string list * (string -> Position.T list) * term) * Proof.context type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list val check_multi_specs: (binding * typ option * mixfix) list -> multi_specs -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val read_multi_specs: (binding * string option * mixfix) list -> multi_specs_cmd -> Proof.context -> (((binding * typ) * mixfix) list * (Attrib.binding * term) list) * Proof.context val axiomatization: (binding * typ option * mixfix) list -> (binding * typ option * mixfix) list -> term list -> (Attrib.binding * term) list -> theory -> (term list * thm list) * theory val axiomatization_cmd: (binding * string option * mixfix) list -> (binding * string option * mixfix) list -> string list -> (Attrib.binding * string) list -> theory -> (term list * thm list) * theory val axiom: Attrib.binding * term -> theory -> thm * theory val definition: (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> local_theory -> (term * (string * thm)) * local_theory val definition': (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term list -> Attrib.binding * term -> bool -> local_theory -> (term * (string * thm)) * local_theory val definition_cmd: (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string list -> Attrib.binding * string -> bool -> local_theory -> (term * (string * thm)) * local_theory val abbreviation: Syntax.mode -> (binding * typ option * mixfix) option -> (binding * typ option * mixfix) list -> term -> bool -> local_theory -> local_theory val abbreviation_cmd: Syntax.mode -> (binding * string option * mixfix) option -> (binding * string option * mixfix) list -> string -> bool -> local_theory -> local_theory val alias: binding * string -> local_theory -> local_theory val alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_alias: binding * string -> local_theory -> local_theory val type_alias_cmd: binding * (xstring * Position.T) -> local_theory -> local_theory val type_notation: bool -> Syntax.mode -> (typ * mixfix) list -> local_theory -> local_theory val type_notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val notation: bool -> Syntax.mode -> (term * mixfix) list -> local_theory -> local_theory val notation_cmd: bool -> Syntax.mode -> (string * mixfix) list -> local_theory -> local_theory val theorems: string -> (Attrib.binding * Attrib.thms) list -> (binding * typ option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorems_cmd: string -> (Attrib.binding * (Facts.ref * Token.src list) list) list -> (binding * string option * mixfix) list -> bool -> local_theory -> (string * thm list) list * local_theory val theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state val schematic_theorem: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> string list -> Element.context_i list -> Element.statement_i -> bool -> local_theory -> Proof.state val schematic_theorem_cmd: bool -> string -> Method.text option -> (thm list list -> local_theory -> local_theory) -> Attrib.binding -> (xstring * Position.T) list -> Element.context list -> Element.statement -> bool -> local_theory -> Proof.state end; structure Specification: SPECIFICATION = struct (* prepare propositions *) fun read_props raw_props raw_fixes ctxt = let val (_, ctxt1) = ctxt |> Proof_Context.add_fixes_cmd raw_fixes; val props1 = map (Syntax.parse_prop ctxt1) raw_props; val (props2, ctxt2) = ctxt1 |> fold_map Variable.fix_dummy_patterns props1; val props3 = Syntax.check_props ctxt2 props2; val ctxt3 = ctxt2 |> fold Variable.declare_term props3; in (props3, ctxt3) end; (* prepare specification *) fun get_positions ctxt x = let fun get Cs (Const ("_type_constraint_", C) $ t) = get (C :: Cs) t | get Cs (Free (y, T)) = if x = y then map_filter Term_Position.decode_positionT (T :: map (Type.constraint_type ctxt) Cs) else [] | get _ (t $ u) = get [] t @ get [] u | get _ (Abs (_, _, t)) = get [] t | get _ _ = []; in get [] end; local fun prep_decls prep_var raw_vars ctxt = let val (vars, ctxt') = fold_map prep_var raw_vars ctxt; val (xs, ctxt'') = ctxt' |> Context_Position.set_visible false |> Proof_Context.add_fixes vars ||> Context_Position.restore_visible ctxt'; val _ = Context_Position.reports ctxt'' (map (Binding.pos_of o #1) vars ~~ map (Variable.markup_entity_def ctxt'' ##> Properties.remove Markup.kindN) xs); in ((vars, xs), ctxt'') end; fun close_form ctxt ys prems concl = let val xs = rev (fold (Variable.add_free_names ctxt) (prems @ [concl]) (rev ys)); val pos_props = Logic.strip_imp_concl concl :: Logic.strip_imp_prems concl @ prems; fun get_pos x = maps (get_positions ctxt x) pos_props; val _ = Context_Position.reports ctxt (maps (Syntax_Phases.reports_of_scope o get_pos) xs); in Logic.close_prop_constraint (Variable.default_type ctxt) (xs ~~ xs) prems concl end; fun dummy_frees ctxt xs tss = let val names = Variable.names_of ((fold o fold) Variable.declare_term tss ctxt) |> fold Name.declare xs; val (tss', _) = (fold_map o fold_map) Term.free_dummy_patterns tss names; in tss' end; fun prep_spec_open prep_var parse_prop raw_vars raw_params raw_prems raw_concl ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val (ys, params_ctxt) = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes; val props = map (parse_prop params_ctxt) (raw_concl :: raw_prems) |> singleton (dummy_frees params_ctxt (xs @ ys)); val concl :: prems = Syntax.check_props params_ctxt props; val spec = Logic.list_implies (prems, concl); val spec_ctxt = Variable.declare_term spec params_ctxt; fun get_pos x = maps (get_positions spec_ctxt x) props; in ((vars, xs, get_pos, spec), spec_ctxt) end; fun prep_specs prep_var parse_prop prep_att raw_vars raw_specss ctxt = let val ((vars, xs), vars_ctxt) = prep_decls prep_var raw_vars ctxt; val propss0 = raw_specss |> map (fn ((_, raw_concl), raw_prems, raw_params) => let val (ys, ctxt') = vars_ctxt |> fold_map prep_var raw_params |-> Proof_Context.add_fixes in (ys, map (pair ctxt') (raw_concl :: raw_prems)) end); val props = burrow (grouped 10 Par_List.map_independent (uncurry parse_prop)) (map #2 propss0) |> dummy_frees vars_ctxt xs |> map2 (fn (ys, _) => fn concl :: prems => close_form vars_ctxt ys prems concl) propss0; val specs = Syntax.check_props vars_ctxt props; val specs_ctxt = vars_ctxt |> fold Variable.declare_term specs; val ps = specs_ctxt |> fold_map Proof_Context.inferred_param xs |> fst; val params = map2 (fn (b, _, mx) => fn (_, T) => ((b, T), mx)) vars ps; val name_atts: Attrib.binding list = map (fn ((name, atts), _) => (name, map (prep_att ctxt) atts)) (map #1 raw_specss); in ((params, name_atts ~~ specs), specs_ctxt) end; in val check_spec_open = prep_spec_open Proof_Context.cert_var (K I); val read_spec_open = prep_spec_open Proof_Context.read_var Syntax.parse_prop; type multi_specs = ((Attrib.binding * term) * term list * (binding * typ option * mixfix) list) list; type multi_specs_cmd = ((Attrib.binding * string) * string list * (binding * string option * mixfix) list) list; fun check_multi_specs xs specs = prep_specs Proof_Context.cert_var (K I) (K I) xs specs; fun read_multi_specs xs specs = prep_specs Proof_Context.read_var Syntax.parse_prop Attrib.check_src xs specs; end; (* axiomatization -- within global theory *) fun gen_axioms prep_stmt prep_att raw_decls raw_fixes raw_prems raw_concls thy = let (*specification*) val ({vars, propss = [prems, concls], ...}, vars_ctxt) = Proof_Context.init_global thy |> prep_stmt (raw_decls @ raw_fixes) ((map o map) (rpair []) [raw_prems, map snd raw_concls]); val (decls, fixes) = chop (length raw_decls) vars; val frees = rev ((fold o fold) (Variable.add_frees vars_ctxt) [prems, concls] []) |> map (fn (x, T) => (x, Free (x, T))); val close = Logic.close_prop (map #2 fixes @ frees) prems; val specs = map ((apsnd o map) (prep_att vars_ctxt) o fst) raw_concls ~~ map close concls; val spec_name = Binding.conglomerate (if null decls then map (#1 o #1) specs else map (#1 o #1) decls); (*consts*) val (consts, consts_thy) = thy |> fold_map (fn ((b, _, mx), (_, t)) => Theory.specify_const ((b, Term.type_of t), mx)) decls; val subst = Term.subst_atomic (map (#2 o #2) decls ~~ consts); (*axioms*) val (axioms, axioms_thy) = (specs, consts_thy) |-> fold_map (fn ((b, atts), prop) => Thm.add_axiom_global (b, subst prop) #>> (fn (_, th) => ((b, atts), [([th], [])]))); (*facts*) val (facts, facts_lthy) = axioms_thy |> Named_Target.theory_init |> Spec_Rules.add spec_name Spec_Rules.Unknown consts (maps (maps #1 o #2) axioms) |> Local_Theory.notes axioms; in ((consts, map (the_single o #2) facts), Local_Theory.exit_global facts_lthy) end; val axiomatization = gen_axioms Proof_Context.cert_stmt (K I); val axiomatization_cmd = gen_axioms Proof_Context.read_stmt Attrib.check_src; fun axiom (b, ax) = axiomatization [] [] [] [(b, ax)] #>> (hd o snd); (* definition *) fun gen_def prep_spec prep_att raw_var raw_params raw_prems ((a, raw_atts), raw_spec) int lthy = let val atts = map (prep_att lthy) raw_atts; val ((vars, xs, get_pos, spec), _) = lthy |> prep_spec (the_list raw_var) raw_params raw_prems raw_spec; val (((x, T), rhs), prove) = Local_Defs.derived_def lthy get_pos {conditional = true} spec; val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of definition " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val const_name = Local_Theory.full_name lthy b; val name = Thm.def_binding_optional b a; val ((lhs, (_, raw_th)), lthy2) = lthy |> Local_Theory.define_internal ((b, mx), ((Binding.suffix_name "_raw" name, []), rhs)); val th = prove lthy2 raw_th; val lthy3 = lthy2 |> Spec_Rules.add name Spec_Rules.equational [lhs] [th]; val ([(def_name, [th'])], lthy4) = lthy3 |> Local_Theory.notes [((name, atts), [([th], [])])]; val lthy5 = lthy4 |> Code.declare_default_eqns [(th', true)]; val lhs' = Morphism.term (Local_Theory.target_morphism lthy5) lhs; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy5 - (Term_Subst.Frees.defined (Term_Subst.Frees.build (Term_Subst.add_frees lhs'))) [(x, T)]; + (Frees.defined (Frees.build (Frees.add_frees lhs'))) [(x, T)]; in ((lhs, (def_name, th')), lthy5) end; val definition' = gen_def check_spec_open (K I); fun definition xs ys As B = definition' xs ys As B false; val definition_cmd = gen_def read_spec_open Attrib.check_src; (* abbreviation *) fun gen_abbrev prep_spec mode raw_var raw_params raw_spec int lthy = let val lthy1 = lthy |> Proof_Context.set_syntax_mode mode; val ((vars, xs, get_pos, spec), _) = lthy |> Proof_Context.set_mode Proof_Context.mode_abbrev |> prep_spec (the_list raw_var) raw_params [] raw_spec; val ((x, T), rhs) = Local_Defs.abs_def (#2 (Local_Defs.cert_def lthy1 get_pos spec)); val _ = Name.reject_internal (x, []); val (b, mx) = (case (vars, xs) of ([], []) => (Binding.make (x, (case get_pos x of [] => Position.none | p :: _ => p)), NoSyn) | ([(b, _, mx)], [y]) => if x = y then (b, mx) else error ("Head of abbreviation " ^ quote x ^ " differs from declaration " ^ quote y ^ Position.here (Binding.pos_of b))); val lthy2 = lthy1 |> Local_Theory.abbrev mode ((b, mx), rhs) |> snd |> Proof_Context.restore_syntax_mode lthy; val _ = Proof_Display.print_consts int (Position.thread_data ()) lthy2 (K false) [(x, T)]; in lthy2 end; val abbreviation = gen_abbrev check_spec_open; val abbreviation_cmd = gen_abbrev read_spec_open; (* alias *) fun gen_alias decl check (b, arg) lthy = let val (c, reports) = check {proper = true, strict = false} lthy arg; val _ = Context_Position.reports lthy reports; in decl b c lthy end; val alias = gen_alias Local_Theory.const_alias (K (K (fn c => (c, [])))); val alias_cmd = gen_alias Local_Theory.const_alias (fn flags => fn ctxt => fn (c, pos) => apfst (#1 o dest_Const) (Proof_Context.check_const flags ctxt (c, [pos]))); val type_alias = gen_alias Local_Theory.type_alias (K (K (fn c => (c, [])))); val type_alias_cmd = gen_alias Local_Theory.type_alias (apfst (#1 o dest_Type) ooo Proof_Context.check_type_name); (* notation *) local fun gen_type_notation prep_type add mode args lthy = lthy |> Local_Theory.type_notation add mode (map (apfst (prep_type lthy)) args); fun gen_notation prep_const add mode args lthy = lthy |> Local_Theory.notation add mode (map (apfst (prep_const lthy)) args); in val type_notation = gen_type_notation (K I); val type_notation_cmd = gen_type_notation (Proof_Context.read_type_name {proper = true, strict = false}); val notation = gen_notation (K I); val notation_cmd = gen_notation (Proof_Context.read_const {proper = false, strict = false}); end; (* fact statements *) local fun gen_theorems prep_fact prep_att add_fixes kind raw_facts raw_fixes int lthy = let val facts = raw_facts |> map (fn ((name, atts), bs) => ((name, map (prep_att lthy) atts), bs |> map (fn (b, more_atts) => (prep_fact lthy b, map (prep_att lthy) more_atts)))); val (_, ctxt') = add_fixes raw_fixes lthy; val facts' = facts |> Attrib.partial_evaluation ctxt' |> Attrib.transform_facts (Proof_Context.export_morphism ctxt' lthy); val (res, lthy') = lthy |> Local_Theory.notes_kind kind facts'; val _ = Proof_Display.print_results int (Position.thread_data ()) lthy' ((kind, ""), res); in (res, lthy') end; in val theorems = gen_theorems (K I) (K I) Proof_Context.add_fixes; val theorems_cmd = gen_theorems Proof_Context.get_fact Attrib.check_src Proof_Context.add_fixes_cmd; end; (* complex goal statements *) local fun prep_statement prep_att prep_stmt raw_elems raw_stmt ctxt = let val (stmt, elems_ctxt) = prep_stmt raw_elems raw_stmt ctxt; val prems = Assumption.local_prems_of elems_ctxt ctxt; val stmt_ctxt = fold (fold (Proof_Context.augment o fst) o snd) stmt elems_ctxt; in (case raw_stmt of Element.Shows _ => let val stmt' = Attrib.map_specs (map prep_att) stmt in (([], prems, stmt', NONE), stmt_ctxt) end | Element.Obtains raw_obtains => let val asms_ctxt = stmt_ctxt |> fold (fn ((name, _), asm) => snd o Proof_Context.add_assms Assumption.assume_export [((name, [Context_Rules.intro_query NONE]), asm)]) stmt; val that = Assumption.local_prems_of asms_ctxt stmt_ctxt; val ([(_, that')], that_ctxt) = asms_ctxt |> Proof_Context.set_stmt true |> Proof_Context.note_thmss "" [((Binding.name Auto_Bind.thatN, []), [(that, [])])] ||> Proof_Context.restore_stmt asms_ctxt; val stmt' = [(Binding.empty_atts, [(#2 (#1 (Obtain.obtain_thesis ctxt)), [])])]; in ((Obtain.obtains_attribs raw_obtains, prems, stmt', SOME that'), that_ctxt) end) end; fun gen_theorem schematic bundle_includes prep_att prep_stmt long kind before_qed after_qed (name, raw_atts) raw_includes raw_elems raw_concl int lthy = let val _ = Local_Theory.assert lthy; val elems = raw_elems |> map (Element.map_ctxt_attrib (prep_att lthy)); val ((more_atts, prems, stmt, facts), goal_ctxt) = lthy |> bundle_includes raw_includes |> prep_statement (prep_att lthy) prep_stmt elems raw_concl; val atts = more_atts @ map (prep_att lthy) raw_atts; val pos = Position.thread_data (); fun after_qed' results goal_ctxt' = let val results' = burrow (map (Goal.norm_result lthy) o Proof_Context.export goal_ctxt' lthy) results; val (res, lthy') = if forall (Binding.is_empty_atts o fst) stmt then (map (pair "") results', lthy) else Local_Theory.notes_kind kind (map2 (fn (b, _) => fn ths => (b, [(ths, [])])) stmt results') lthy; val lthy'' = if Binding.is_empty_atts (name, atts) then (Proof_Display.print_results int pos lthy' ((kind, ""), res); lthy') else let val ([(res_name, _)], lthy'') = Local_Theory.notes_kind kind [((name, atts), [(maps #2 res, [])])] lthy'; val _ = Proof_Display.print_results int pos lthy' ((kind, res_name), res); in lthy'' end; in after_qed results' lthy'' end; val prems_name = if long then Auto_Bind.assmsN else Auto_Bind.thatN; in goal_ctxt |> not (null prems) ? (Proof_Context.note_thmss "" [((Binding.name prems_name, []), [(prems, [])])] #> snd) |> Proof.theorem before_qed after_qed' (map snd stmt) |> (case facts of NONE => I | SOME ths => Proof.refine_insert ths) |> tap (fn state => not schematic andalso Proof.schematic_goal state andalso error "Illegal schematic goal statement") end; in val theorem = gen_theorem false Bundle.includes (K I) Expression.cert_statement; val theorem_cmd = gen_theorem false Bundle.includes_cmd Attrib.check_src Expression.read_statement; val schematic_theorem = gen_theorem true Bundle.includes (K I) Expression.cert_statement; val schematic_theorem_cmd = gen_theorem true Bundle.includes_cmd Attrib.check_src Expression.read_statement; 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,257 +1,257 @@ (* 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} 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}; 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 = Term_Subst.Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; + val schematic_terms = Vars.fold (fn (v, t) => cons (v, Thm.cterm_of ctxt3 t)) inst []; - val schematics = (Term_Subst.TVars.dest schematic_types, schematic_terms); + val schematics = (TVars.dest 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 = Term_Subst.Vars.build (Term_Subst.add_vars (Logic.strip_imp_concl prop)); + val concl_vars = Vars.build (Vars.add_vars (Logic.strip_imp_concl prop)); val vars = build_rev (Term.add_vars prop); 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 Term_Subst.Vars.defined concl_vars v then (T, []) + 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'; 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,152 @@ (* 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' Term_Subst.TFrees.empty 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')) 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/Proof/proof_rewrite_rules.ML b/src/Pure/Proof/proof_rewrite_rules.ML --- a/src/Pure/Proof/proof_rewrite_rules.ML +++ b/src/Pure/Proof/proof_rewrite_rules.ML @@ -1,394 +1,394 @@ (* Title: Pure/Proof/proof_rewrite_rules.ML Author: Stefan Berghofer, TU Muenchen Simplification functions for proof terms involving meta level rules. *) signature PROOF_REWRITE_RULES = sig val rew : bool -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val rprocs : bool -> (typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option) list val rewrite_terms : (term -> term) -> Proofterm.proof -> Proofterm.proof val elim_defs : theory -> bool -> thm list -> Proofterm.proof -> Proofterm.proof val elim_vars : (typ -> term) -> Proofterm.proof -> Proofterm.proof val hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val un_hhf_proof : term -> term -> Proofterm.proof -> Proofterm.proof val expand_of_sort_proof : theory -> term option list -> typ * sort -> Proofterm.proof list val expand_of_class_proof : theory -> term option list -> typ * class -> proof val expand_of_class : theory -> typ list -> term option list -> Proofterm.proof -> (Proofterm.proof * Proofterm.proof) option val standard_preproc: (proof * proof) list -> theory -> proof -> proof end; structure Proof_Rewrite_Rules : PROOF_REWRITE_RULES = struct fun rew b _ _ = let fun ?? x = if b then SOME x else NONE; fun ax (prf as PAxm (s, prop, _)) Ts = if b then PAxm (s, prop, SOME Ts) else prf; fun ty T = if b then (case T of Type (_, [Type (_, [U, _]), _]) => SOME U | _ => NONE) else NONE; val equal_intr_axm = ax Proofterm.equal_intr_axm []; val equal_elim_axm = ax Proofterm.equal_elim_axm []; val symmetric_axm = ax Proofterm.symmetric_axm [propT]; fun rew' (PThm ({name = "Pure.protectD", ...}, _) % _ %% (PThm ({name = "Pure.protectI", ...}, _) % _ %% prf)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD1", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PThm ({name = "Pure.conjunctionD2", ...}, _) % _ % _ %% (PThm ({name = "Pure.conjunctionI", ...}, _) % _ % _ %% _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % A % B %% prf1 %% prf2)) = SOME (equal_intr_axm % B % A %% prf2 %% prf1) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %% ((tg as PThm ({name = "Pure.protectI", ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf1) %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) = let val _ $ A $ C = Envir.beta_norm X; val _ $ B $ D = Envir.beta_norm Y in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? B, Proofterm.equal_elim_axm %> C %> D %% Proofterm.incr_pboundvars 2 0 prf2 %% (PBound 1 %% (equal_elim_axm %> B %> A %% (Proofterm.symmetric_axm % ?? A % ?? B %% Proofterm.incr_pboundvars 2 0 prf1) %% PBound 0))))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.imp", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2))) = let val _ $ A $ C = Envir.beta_norm Y; val _ $ B $ D = Envir.beta_norm X in SOME (AbsP ("H1", ?? X, AbsP ("H2", ?? A, equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% Proofterm.incr_pboundvars 2 0 prf2) %% (PBound 1 %% (equal_elim_axm %> A %> B %% Proofterm.incr_pboundvars 2 0 prf1 %% PBound 0))))) end | rew' (PAxm ("Pure.combination", _, _) % SOME (imp as (imp' as Const ("Pure.imp", T)) $ X) % _ % Y % Z %% (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME (ax Proofterm.combination_axm [propT, propT] %> imp % ?? imp % Y % Z %% (ax Proofterm.combination_axm [propT, propT --> propT] %> imp' % ?? imp' % ?? X % ?? X %% (ax Proofterm.reflexive_axm [T] % ?? imp') %% (ax Proofterm.reflexive_axm [propT] % ?? X)) %% prf) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> incr_boundvars 1 P $ Bound 0 %> incr_boundvars 1 Q $ Bound 0 %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME X % SOME Y %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.all", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% (PAxm ("Pure.abstract_rule", _, _) % _ % _ %% prf)))) = let val Const (_, T) $ P = Envir.beta_norm X; val _ $ Q = Envir.beta_norm Y; val t = incr_boundvars 1 P $ Bound 0; val u = incr_boundvars 1 Q $ Bound 0 in SOME (AbsP ("H", ?? X, Abst ("x", ty T, equal_elim_axm %> t %> u %% (symmetric_axm % ?? u % ?? t %% (Proofterm.incr_pboundvars 1 1 prf %> Bound 0)) %% (PBound 0 %> Bound 0)))) end | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2) %% prf3) = SOME (equal_elim_axm %> B %> C %% prf2 %% (equal_elim_axm %> A %> B %% prf1 %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME A % SOME C %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.transitive", _, _) % _ % SOME B % _ %% prf1 %% prf2)) %% prf3) = SOME (equal_elim_axm %> B %> C %% (symmetric_axm % ?? C % ?? B %% prf1) %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf2) %% prf3)) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _)) %% prf) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3) %% prf4) = SOME (equal_elim_axm %> C %> D %% prf2 %% (equal_elim_axm %> A %> C %% prf3 %% (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A $ C) % SOME (_ $ B $ D) %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2) %% prf3)) %% prf4) = SOME (equal_elim_axm %> A %> B %% prf1 %% (equal_elim_axm %> C %> A %% (symmetric_axm % ?? A % ?? C %% prf3) %% (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3) %% prf4) = SOME (equal_elim_axm %> D %> C %% (symmetric_axm % ?? C % ?? D %% prf2) %% (equal_elim_axm %> B %> D %% prf3 %% (equal_elim_axm %> A %> B %% prf1 %% prf4))) | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ B $ D) % SOME (_ $ A $ C) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % _ % _ % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.eq", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% prf2)) %% prf3)) %% prf4) = SOME (equal_elim_axm %> B %> A %% (symmetric_axm % ?? A % ?? B %% prf1) %% (equal_elim_axm %> D %> B %% (symmetric_axm % ?? B % ?? D %% prf3) %% (equal_elim_axm %> C %> D %% prf2 %% prf4))) | rew' ((prf as PAxm ("Pure.combination", _, _) % SOME ((eq as Const ("Pure.eq", T)) $ t) % _ % _ % _) %% (PAxm ("Pure.reflexive", _, _) % _)) = let val (U, V) = (case T of Type (_, [U, V]) => (U, V) | _ => (dummyT, dummyT)) in SOME (prf %% (ax Proofterm.combination_axm [U, V] %> eq % ?? eq % ?? t % ?? t %% (ax Proofterm.reflexive_axm [T] % ?? eq) %% (ax Proofterm.reflexive_axm [U] % ?? t))) end | rew' _ = NONE; in rew' #> Option.map (rpair Proofterm.no_skel) end; fun rprocs b = [rew b]; val _ = Theory.setup (fold Proofterm.add_prf_rproc (rprocs false)); (**** apply rewriting function to all terms in proof ****) fun rewrite_terms r = let fun rew_term Ts t = let val frees = map Free (Name.invent (Term.declare_term_frees t Name.context) "xa" (length Ts) ~~ Ts); val t' = r (subst_bounds (frees, t)); fun strip [] t = t | strip (_ :: xs) (Abs (_, _, t)) = strip xs t; in strip Ts (fold lambda frees t') end; fun rew Ts (prf1 %% prf2) = rew Ts prf1 %% rew Ts prf2 | rew Ts (prf % SOME t) = rew Ts prf % SOME (rew_term Ts t) | rew Ts (Abst (s, SOME T, prf)) = Abst (s, SOME T, rew (T :: Ts) prf) | rew Ts (AbsP (s, SOME t, prf)) = AbsP (s, SOME (rew_term Ts t), rew Ts prf) | rew _ prf = prf in rew [] end; (**** eliminate definitions in proof ****) fun vars_of t = rev (fold_aterms (fn v as Var _ => insert (op =) v | _ => I) t []); fun insert_refl defs Ts (prf1 %% prf2) = let val (prf1', b) = insert_refl defs Ts prf1 in if b then (prf1', true) else (prf1' %% fst (insert_refl defs Ts prf2), false) end | insert_refl defs Ts (Abst (s, SOME T, prf)) = (Abst (s, SOME T, fst (insert_refl defs (T :: Ts) prf)), false) | insert_refl defs Ts (AbsP (s, t, prf)) = (AbsP (s, t, fst (insert_refl defs Ts prf)), false) | insert_refl defs Ts prf = (case Proofterm.strip_combt prf of (PThm ({name = s, prop, types = SOME Ts, ...}, _), ts) => if member (op =) defs s then let val vs = vars_of prop; val tvars = build_rev (Term.add_tvars prop); val (_, rhs) = Logic.dest_equals (Logic.strip_imp_concl prop); val rhs' = Term.betapplys (subst_TVars (map fst tvars ~~ Ts) (fold_rev (fn x => fn b => Abs ("", dummyT, abstract_over (x, b))) vs rhs), map the ts); in (Proofterm.change_types (SOME [fastype_of1 (Ts, rhs')]) Proofterm.reflexive_axm %> rhs', true) end else (prf, false) | (_, []) => (prf, false) | (prf', ts) => (Proofterm.proof_combt' (fst (insert_refl defs Ts prf'), ts), false)); fun elim_defs thy r defs prf = let val defs' = map (Logic.dest_equals o map_types Type.strip_sorts o Thm.prop_of o Drule.abs_def) defs; val defnames = map Thm.derivation_name defs; val prf' = if r then let val cnames = map (fst o dest_Const o fst) defs'; val expand = Proofterm.fold_proof_atoms true (fn PThm ({serial, name, prop, ...}, _) => if member (op =) defnames name orelse not (exists_Const (member (op =) cnames o #1) prop) then I else Inttab.update (serial, "") | _ => I) [prf] Inttab.empty; in Proofterm.expand_proof thy (Inttab.lookup expand o #serial) prf end else prf; in rewrite_terms (Pattern.rewrite_term thy defs' []) (fst (insert_refl defnames [] prf')) end; (**** eliminate all variables that don't occur in the proposition ****) fun elim_vars mk_default prf = let val prop = Proofterm.prop_of prf; - val tv = Term_Subst.Vars.build (Term_Subst.add_vars prop); - val tf = Term_Subst.Frees.build (Term_Subst.add_frees prop); + val tv = Vars.build (Vars.add_vars prop); + val tf = Frees.build (Frees.add_frees prop); - fun hidden_variable (Var v) = not (Term_Subst.Vars.defined tv v) - | hidden_variable (Free f) = not (Term_Subst.Frees.defined tf f) + fun hidden_variable (Var v) = not (Vars.defined tv v) + | hidden_variable (Free f) = not (Frees.defined tf f) | hidden_variable _ = false; fun mk_default' T = fold_rev (Term.abs o pair "x") (binder_types T) (mk_default (body_type T)); fun elim_varst (t $ u) = elim_varst t $ elim_varst u | elim_varst (Abs (s, T, t)) = Abs (s, T, elim_varst t) - | elim_varst (t as Free (x, T)) = if Term_Subst.Frees.defined tf (x, T) then t else mk_default' T - | elim_varst (t as Var (xi, T)) = if Term_Subst.Vars.defined tv (xi, T) then t else mk_default' T + | elim_varst (t as Free (x, T)) = if Frees.defined tf (x, T) then t else mk_default' T + | elim_varst (t as Var (xi, T)) = if Vars.defined tv (xi, T) then t else mk_default' T | elim_varst t = t; in Proofterm.map_proof_terms (fn t => if Term.exists_subterm hidden_variable t then Envir.beta_norm (elim_varst t) else t) I prf end; (**** convert between hhf and non-hhf form ****) fun hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf i j Hs Hs' (Const ("Pure.all", _) $ Abs (_, _, P)) prf = mk_prf i (j - 1) Hs Hs' P (prf %> Bound j) | mk_prf i j (H :: Hs) (H' :: Hs') (Const ("Pure.imp", _) $ _ $ P) prf = mk_prf (i - 1) j Hs Hs' P (prf %% un_hhf_proof H' H (PBound i)) | mk_prf _ _ _ _ _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> mk_prf (k - 1) (l - 1) Hs Hs' P |> fold_rev (fn P => fn prf => AbsP ("H", SOME P, prf)) Hs' |> fold_rev (fn (s, T) => fn prf => Abst (s, SOME T, prf)) params end and un_hhf_proof P Q prf = let val params = Logic.strip_params Q; val Hs = Logic.strip_assums_hyp P; val Hs' = Logic.strip_assums_hyp Q; val k = length Hs; val l = length params; fun mk_prf (Const ("Pure.all", _) $ Abs (s, T, P)) prf = Abst (s, SOME T, mk_prf P prf) | mk_prf (Const ("Pure.imp", _) $ P $ Q) prf = AbsP ("H", SOME P, mk_prf Q prf) | mk_prf _ prf = prf in prf |> Proofterm.incr_pboundvars k l |> fold (fn i => fn prf => prf %> Bound i) (l - 1 downto 0) |> fold (fn ((H, H'), i) => fn prf => prf %% hhf_proof H' H (PBound i)) (Hs ~~ Hs' ~~ (k - 1 downto 0)) |> mk_prf Q end; (**** expand PClass proofs ****) fun expand_of_sort_proof thy hyps (T, S) = let val of_class_hyps = map (fn SOME t => try Logic.dest_of_class t | NONE => NONE) hyps; fun of_class_index p = (case find_index (fn SOME h => h = p | NONE => false) of_class_hyps of ~1 => raise Fail "expand_of_class_proof: missing class hypothesis" | i => PBound i); val sorts = AList.coalesce (op =) (rev (map_filter I of_class_hyps)); fun get_sort T = the_default [] (AList.lookup (op =) sorts T); val subst = map_atyps (fn T as TVar (ai, _) => TVar (ai, get_sort T) | T as TFree (a, _) => TFree (a, get_sort T)); fun reconstruct prop = Proofterm.reconstruct_proof thy prop #> Proofterm.expand_proof thy Proofterm.expand_name_empty #> Same.commit (Proofterm.map_proof_same Same.same Same.same of_class_index); in map2 reconstruct (Logic.mk_of_sort (T, S)) (Proofterm.of_sort_proof (Sign.classes_of thy) (Thm.classrel_proof thy) (Thm.arity_proof thy) (PClass o apfst Type.strip_sorts) (subst T, S)) end; fun expand_of_class_proof thy hyps (T, c) = hd (expand_of_sort_proof thy hyps (T, [c])); fun expand_of_class thy (_: typ list) hyps (PClass (T, c)) = SOME (expand_of_class_proof thy hyps (T, c), Proofterm.no_skel) | expand_of_class _ _ _ _ = NONE; (* standard preprocessor *) fun standard_preproc rews thy = Proofterm.rewrite_proof thy (rews, rprocs true @ [expand_of_class thy]); 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,354 +1,353 @@ (* 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 Term_Subst.TVars.get_first (fn ((bi, S), ()) => if ai = bi then SOME S else NONE) tvars of + (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: Term_Subst.TVars.set) = +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 Term_Subst.TVars.fold (add o #1) tvars [] 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 = Term_Subst.TVars.build (Thm.fold_terms {hyps = false} Term_Subst.add_tvars thm); - val vars = Term_Subst.Vars.build (Thm.fold_terms {hyps = false} Term_Subst.add_vars thm); + 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 - |> Term_Subst.TVars.fold (Variable.declare_internal o Logic.mk_type o TVar o #1) tvars - |> Term_Subst.Vars.fold (Variable.declare_internal o Var o #1) vars + |> 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 (Term_Subst.TVars.table (map (read_type ctxt1 tvars) type_insts)); + Term_Subst.instantiateT (TVars.make (map (read_type ctxt1 tvars) type_insts)); val vars1 = - Vartab.build (vars |> Term_Subst.Vars.fold (fn ((v, T), ()) => + 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 (Term_Subst.TVars.table inferred); + 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 (Term_Subst.TVars.empty, - Term_Subst.Vars.build - (fold2 (fn (xi, _) => fn t => Term_Subst.Vars.add ((xi, Term.fastype_of t), t)) xs ts)) + 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) |> 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 (build_rev (Term.add_vars (Thm.full_prop_of thm))) args @ zip_vars (build_rev (Term.add_vars (Thm.concl_of thm))) 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 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; 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/consts.ML b/src/Pure/consts.ML --- a/src/Pure/consts.ML +++ b/src/Pure/consts.ML @@ -1,348 +1,347 @@ (* Title: Pure/consts.ML Author: Makarius Polymorphic constants: declarations, abbreviations, additional type constraints. *) signature CONSTS = sig type T val eq_consts: T * T -> bool val change_base: bool -> T -> T val change_ignore: T -> T val retrieve_abbrevs: T -> string list -> term -> (term * term) list val dest: T -> {const_space: Name_Space.T, constants: (string * (typ * term option)) list, constraints: (string * typ) list} val the_const: T -> string -> string * typ (*exception TYPE*) val the_abbreviation: T -> string -> typ * term (*exception TYPE*) val type_scheme: T -> string -> typ (*exception TYPE*) val is_monomorphic: T -> string -> bool (*exception TYPE*) val the_constraint: T -> string -> typ (*exception TYPE*) val space_of: T -> Name_Space.T val alias: Name_Space.naming -> binding -> string -> T -> T val is_concealed: T -> string -> bool val intern: T -> xstring -> string val intern_syntax: T -> xstring -> string val check_const: Context.generic -> T -> xstring * Position.T list -> term * Position.report list val certify: Context.generic -> Type.tsig -> bool -> T -> term -> term (*exception TYPE*) val typargs: T -> string * typ -> typ list val instance: T -> string * typ list -> typ val dummy_types: T -> term -> term val declare: Context.generic -> binding * typ -> T -> T val constrain: string * typ option -> T -> T val abbreviate: Context.generic -> Type.tsig -> string -> binding * term -> T -> (term * term) * T val revert_abbrev: string -> string -> T -> T val hide: bool -> string -> T -> T val empty: T val merge: T * T -> T end; structure Consts: CONSTS = struct (** consts type **) (* datatype T *) type decl = {T: typ, typargs: int list list}; type abbrev = {rhs: term, normal_rhs: term, force_expand: bool}; datatype T = Consts of {decls: (decl * abbrev option) Name_Space.table, constraints: typ Symtab.table, rev_abbrevs: (term * term) Item_Net.T Symtab.table}; fun eq_consts (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1}, Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) = pointer_eq (decls1, decls2) andalso pointer_eq (constraints1, constraints2) andalso pointer_eq (rev_abbrevs1, rev_abbrevs2); fun make_consts (decls, constraints, rev_abbrevs) = Consts {decls = decls, constraints = constraints, rev_abbrevs = rev_abbrevs}; fun map_consts f (Consts {decls, constraints, rev_abbrevs}) = make_consts (f (decls, constraints, rev_abbrevs)); fun change_base begin = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.change_base begin decls, constraints, rev_abbrevs)); val change_ignore = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.change_ignore decls, constraints, rev_abbrevs)); (* reverted abbrevs *) val empty_abbrevs = Item_Net.init (fn ((t, u), (t', u')) => t aconv t' andalso u aconv u') (single o #1); fun update_abbrevs mode abbrs = Symtab.map_default (mode, empty_abbrevs) (Item_Net.update abbrs); fun retrieve_abbrevs (Consts {rev_abbrevs, ...}) modes = let val nets = map_filter (Symtab.lookup rev_abbrevs) modes in fn t => let val retrieve = if Term.could_beta_eta_contract t then Item_Net.retrieve else Item_Net.retrieve_matching in maps (fn net => retrieve net t) nets end end; (* dest consts *) fun dest (Consts {decls, constraints, ...}) = {const_space = Name_Space.space_of_table decls, constants = Name_Space.fold_table (fn (c, ({T, ...}, abbr)) => cons (c, (T, Option.map #rhs abbr))) decls [], constraints = Symtab.dest constraints}; (* lookup consts *) fun the_entry (Consts {decls, ...}) c = (case Name_Space.lookup_key decls c of SOME entry => entry | NONE => raise TYPE ("Unknown constant: " ^ quote c, [], [])); fun the_const consts c = (case the_entry consts c of (c', ({T, ...}, NONE)) => (c', T) | _ => raise TYPE ("Not a logical constant: " ^ quote c, [], [])); fun the_abbreviation consts c = (case the_entry consts c of (_, ({T, ...}, SOME {rhs, ...})) => (T, rhs) | _ => raise TYPE ("Not an abbreviated constant: " ^ quote c, [], [])); fun the_decl consts = #1 o #2 o the_entry consts; val type_scheme = #T oo the_decl; val type_arguments = #typargs oo the_decl; val is_monomorphic = null oo type_arguments; fun the_constraint (consts as Consts {constraints, ...}) c = (case Symtab.lookup constraints c of SOME T => T | NONE => type_scheme consts c); (* name space and syntax *) fun space_of (Consts {decls, ...}) = Name_Space.space_of_table decls; fun alias naming binding name = map_consts (fn (decls, constraints, rev_abbrevs) => ((Name_Space.alias_table naming binding name decls), constraints, rev_abbrevs)); val is_concealed = Name_Space.is_concealed o space_of; val intern = Name_Space.intern o space_of; fun intern_syntax consts s = (case try Lexicon.unmark_const s of SOME c => c | NONE => intern consts s); (* check_const *) fun check_const context consts (xname, ps) = let val Consts {decls, ...} = consts; val ((c, reports), _) = Name_Space.check_reports context decls (xname, ps); val T = type_scheme consts c handle TYPE (msg, _, _) => error (msg ^ Position.here_list ps); in (Const (c, T), reports) end; (* certify *) fun certify context tsig do_expand consts = let fun err msg (c, T) = raise TYPE (msg ^ " " ^ quote c ^ " :: " ^ Syntax.string_of_typ (Syntax.init_pretty context) T, [], []); val certT = Type.cert_typ tsig; fun cert tm = let val (head, args) = Term.strip_comb tm; val args' = map cert args; fun comb head' = Term.list_comb (head', args'); in (case head of Abs (x, T, t) => comb (Abs (x, certT T, cert t)) | Const (c, T) => let val T' = certT T; val (_, ({T = U, ...}, abbr)) = the_entry consts c; fun expand u = Term.betapplys (Envir.expand_atom T' (U, u) handle TYPE _ => err "Illegal type for abbreviation" (c, T), args'); in if not (Type.raw_instance (T', U)) then err "Illegal type for constant" (c, T) else (case abbr of SOME {rhs, normal_rhs, force_expand} => if do_expand then expand normal_rhs else if force_expand then expand rhs else comb head | _ => comb head) end | _ => comb head) end; in cert end; (* typargs -- view actual const type as instance of declaration *) local fun args_of (Type (_, Ts)) pos = args_of_list Ts 0 pos | args_of (TVar v) pos = insert (eq_fst op =) (v, rev pos) | args_of (TFree _) _ = I and args_of_list (T :: Ts) i is = args_of T (i :: is) #> args_of_list Ts (i + 1) is | args_of_list [] _ _ = I; fun subscript (Type (_, Ts)) (i :: is) = subscript (nth Ts i) is | subscript T [] = T | subscript _ _ = raise Subscript; in fun typargs_of T = map #2 (rev (args_of T [] [])); fun typargs consts (c, T) = map (subscript T) (type_arguments consts c); end; fun instance consts (c, Ts) = let val declT = type_scheme consts c; val args = typargs consts (c, declT); val inst = - Term_Subst.TVars.build - (fold2 (fn a => fn T => Term_Subst.TVars.add (Term.dest_TVar a, T)) args Ts) - handle ListPair.UnequalLengths => raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]); + TVars.build (fold2 (fn a => fn T => TVars.add (Term.dest_TVar a, T)) args Ts) + handle ListPair.UnequalLengths => raise TYPE ("Consts.instance", Ts, [Const (c, dummyT)]); in declT |> Term_Subst.instantiateT inst end; fun dummy_types consts = let fun dummy (Const (c, T)) = Const (c, instance consts (c, replicate (length (typargs consts (c, T))) dummyT)) | dummy (Free (x, _)) = Free (x, dummyT) | dummy (Var (xi, _)) = Var (xi, dummyT) | dummy (b as Bound _) = b | dummy (t $ u) = dummy t $ dummy u | dummy (Abs (a, _, b)) = Abs (a, dummyT, dummy b); in dummy end; (** build consts **) (* name space *) fun hide fully c = map_consts (fn (decls, constraints, rev_abbrevs) => (Name_Space.hide_table fully c decls, constraints, rev_abbrevs)); (* declarations *) fun declare context (b, declT) = map_consts (fn (decls, constraints, rev_abbrevs) => let val decl = {T = declT, typargs = typargs_of declT}; val _ = Binding.check b; val (_, decls') = decls |> Name_Space.define context true (b, (decl, NONE)); in (decls', constraints, rev_abbrevs) end); (* constraints *) fun constrain (c, C) consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) => (#2 (the_entry consts c) handle TYPE (msg, _, _) => error msg; (decls, constraints |> (case C of SOME T => Symtab.update (c, T) | NONE => Symtab.delete_safe c), rev_abbrevs))); (* abbreviations *) local fun strip_abss (t as Abs (x, T, b)) = if Term.is_dependent b then strip_abss b |>> cons (x, T) (* FIXME decr!? *) else ([], t) | strip_abss t = ([], t); fun rev_abbrev lhs rhs = let val (xs, body) = strip_abss (Envir.beta_eta_contract rhs); val vars = fold (fn (x, T) => cons (Var ((x, 0), T))) (Term.rename_wrt_term body xs) []; in (Term.subst_bounds (rev vars, body), Term.list_comb (lhs, vars)) end; in fun abbreviate context tsig mode (b, raw_rhs) consts = let val cert_term = certify context tsig false consts; val expand_term = certify context tsig true consts; val force_expand = mode = Print_Mode.internal; val _ = Term.exists_subterm Term.is_Var raw_rhs andalso error ("Illegal schematic variables on rhs of abbreviation " ^ Binding.print b); val rhs = raw_rhs |> Term.map_types (Type.cert_typ tsig) |> cert_term |> Term.close_schematic_term; val normal_rhs = expand_term rhs; val T = Term.fastype_of rhs; val lhs = Const (Name_Space.full_name (Name_Space.naming_of context) b, T); in consts |> map_consts (fn (decls, constraints, rev_abbrevs) => let val decl = {T = T, typargs = typargs_of T}; val abbr = {rhs = rhs, normal_rhs = normal_rhs, force_expand = force_expand}; val _ = Binding.check b; val (_, decls') = decls |> Name_Space.define context true (b, (decl, SOME abbr)); val rev_abbrevs' = rev_abbrevs |> update_abbrevs mode (rev_abbrev lhs rhs); in (decls', constraints, rev_abbrevs') end) |> pair (lhs, rhs) end; fun revert_abbrev mode c consts = consts |> map_consts (fn (decls, constraints, rev_abbrevs) => let val (T, rhs) = the_abbreviation consts c; val rev_abbrevs' = rev_abbrevs |> update_abbrevs mode (rev_abbrev (Const (c, T)) rhs); in (decls, constraints, rev_abbrevs') end); end; (* empty and merge *) val empty = make_consts (Name_Space.empty_table Markup.constantN, Symtab.empty, Symtab.empty); fun merge (Consts {decls = decls1, constraints = constraints1, rev_abbrevs = rev_abbrevs1}, Consts {decls = decls2, constraints = constraints2, rev_abbrevs = rev_abbrevs2}) = let val decls' = Name_Space.merge_tables (decls1, decls2); val constraints' = Symtab.merge (K true) (constraints1, constraints2); val rev_abbrevs' = Symtab.join (K Item_Net.merge) (rev_abbrevs1, rev_abbrevs2); in make_consts (decls', constraints', rev_abbrevs') 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,834 +1,834 @@ (* 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 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: Symtab.set * Symtab.set -> thm -> thm + 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 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 = - Term_Subst.Vars.build (th |> (Thm.fold_terms {hyps = false} o Term.fold_aterms) + Vars.build (th |> (Thm.fold_terms {hyps = false} o Term.fold_aterms) (fn t => fn inst => (case t of Var (xi, T) => - if Term_Subst.Vars.defined inst (xi, T) then inst + 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 Term_Subst.Vars.update ((xi, T), ct) inst end + in Vars.add ((xi, T), ct) inst end | _ => inst))); in th - |> Thm.instantiate ([], Term_Subst.Vars.dest inst) + |> Thm.instantiate ([], Vars.dest 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 = fold Thm.add_tvars ths Term_Subst.TVars.empty; - val the_tvar = the o Term_Subst.TVars.lookup tvars; + val tvars = TVars.build (fold Thm.add_tvars ths); + val the_tvar = the o TVars.lookup tvars; val instT' = - build (instT |> Term_Subst.TVars.fold (fn (v, TVar (b, _)) => + build (instT |> TVars.fold (fn (v, TVar (b, _)) => cons (v, Thm.rename_tvar b (the_tvar v)))); - val vars = fold (Thm.add_vars o Thm.instantiate (instT', [])) ths Term_Subst.Vars.empty; - val the_var = the o Term_Subst.Vars.lookup vars; + val vars = Vars.build (fold (Thm.add_vars o Thm.instantiate (instT', [])) ths); + val the_var = the o Vars.lookup vars; val inst' = - build (inst |> Term_Subst.Vars.fold (fn (v, Var (b, _)) => + build (inst |> Vars.fold (fn (v, Var (b, _)) => cons (v, 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; 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 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))); 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; (* 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 = Term_Subst.Frees.build (fold Term_Subst.add_frees (prop :: As)); - val xs = Term_Subst.Frees.fold_rev (cons o Thm.cterm_of ctxt o Free o #1) frees []; + 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 = Term_Subst.TFrees.build (fold Term_Subst.add_tfrees (prop :: As)); - val Ts = Symtab.build (Term_Subst.TFrees.fold (Symtab.insert_set o #1 o #1) tfrees); + 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 |> Term_Subst.TFrees.fold (fn ((a, S), ()) => + build (tfrees |> TFrees.fold (fn ((a, S), ()) => cons (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S))))); val global_prop = Logic.list_implies (As, prop) - |> Term_Subst.Frees.fold_rev (Logic.all o Free o #1) frees + |> 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, Symtab.empty) 0 #> + 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, []) |> 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,783 +1,783 @@ (* 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 structure Ctermtab: TABLE structure Thmtab: TABLE val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM structure Ctermtab: TABLE structure Thmtab: TABLE val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool - val add_tvars: thm -> ctyp Term_Subst.TVars.table -> ctyp Term_Subst.TVars.table - val add_vars: thm -> cterm Term_Subst.Vars.table -> cterm Term_Subst.Vars.table + 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 fast_term_ord: cterm ord val term_ord: cterm ord val thm_ord: thm ord val cterm_cache: (cterm -> 'a) -> cterm -> 'a val thm_cache: (thm -> 'a) -> thm -> 'a 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': 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 (Term_Subst.TVars.defined tab v) ? Term_Subst.TVars.update (v, cT) end); + 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 (Term_Subst.Vars.defined tab v) ? Term_Subst.Vars.update (v, ct) end); + 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; (* certified term order *) val fast_term_ord = Term_Ord.fast_term_ord o apply2 Thm.term_of; val term_ord = Term_Ord.term_ord o apply2 Thm.term_of; (* thm order: ignores theory context! *) val thm_ord = Term_Ord.fast_term_ord o apply2 Thm.prop_of ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 Thm.tpairs_of ||| list_ord Term_Ord.fast_term_ord o apply2 Thm.hyps_of ||| list_ord Term_Ord.sort_ord o apply2 Thm.shyps_of; (* scalable collections *) structure Ctermtab = Table(type key = cterm val ord = fast_term_ord); structure Thmtab = Table(type key = thm val ord = thm_ord); fun cterm_cache f = Cache.create Ctermtab.empty Ctermtab.lookup Ctermtab.update f; fun thm_cache f = Cache.create Thmtab.empty Thmtab.lookup Thmtab.update f; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o 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 = Symtab.build (fold (Symtab.insert_set o #1 o #1) instT); - val frees = Symtab.build (fold (Symtab.insert_set o #1 o #1) 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); 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 inst = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_vars thm')) cts; in Thm.instantiate ([], inst) thm' end; (* implicit generalization over variables -- canonical order *) fun forall_intr_vars th = let val (_, vars) = - (th, (Term_Subst.Vars.empty, [])) |-> Thm.fold_atomic_cterms {hyps = false} Term.is_Var + (th, (Vars.empty, [])) |-> Thm.fold_atomic_cterms {hyps = false} Term.is_Var (fn ct => fn (seen, vars) => let val v = Term.dest_Var (Thm.term_of ct) in - if not (Term_Subst.Vars.defined seen v) - then (Term_Subst.Vars.add (v, ()) seen, ct :: vars) + if not (Vars.defined seen v) + then (Vars.add_set v seen, ct :: vars) else (seen, vars) end); in fold Thm.forall_intr vars th end; fun forall_intr_frees th = let val fixed = - Term_Subst.Frees.build - (fold Term_Subst.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> - fold Term_Subst.add_frees (Thm.hyps_of th)); + Frees.build + (fold Frees.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> + fold Frees.add_frees (Thm.hyps_of th)); val (_, frees) = (th, (fixed, [])) |-> Thm.fold_atomic_cterms {hyps = false} Term.is_Free (fn ct => fn (seen, frees) => let val v = Term.dest_Free (Thm.term_of ct) in - if not (Term_Subst.Frees.defined seen v) - then (Term_Subst.Frees.add (v, ()) seen, ct :: frees) + if not (Frees.defined seen v) + then (Frees.add_set v seen, ct :: frees) else (seen, frees) end); in fold Thm.forall_intr 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 = - Term_Subst.TVars.build (prop |> (Term.fold_types o Term.fold_atyps) + TVars.build (prop |> (Term.fold_types o Term.fold_atyps) (fn T => fn instT => (case T of TVar (v as ((a, _), S)) => - if Term_Subst.TVars.defined instT v then instT - else Term_Subst.TVars.update (v, TFree (a, S)) instT + if TVars.defined instT v then instT + else TVars.add (v, TFree (a, S)) instT | _ => instT))); - val cinstT = Term_Subst.TVars.map (K certT) instT; + val cinstT = TVars.map (K certT) instT; val cinst = - Term_Subst.Vars.build (prop |> Term.fold_aterms + 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 Term_Subst.Vars.update (((x, i), T'), cert (Free ((x, T')))) inst end + in Vars.add (((x, i), T'), cert (Free ((x, T')))) inst end | _ => inst))); - in Thm.instantiate (Term_Subst.TVars.dest cinstT, Term_Subst.Vars.dest cinst) th end; + in Thm.instantiate (TVars.dest cinstT, Vars.dest 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' |> 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' |> 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,178 +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 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 = - Term_Subst.TFrees.build - (cinstT |> fold (fn (v, cT) => Term_Subst.TFrees.add (v, Thm.typ_of cT))); - val inst = - Term_Subst.Frees.build - (cinst |> fold (fn (v, ct) => Term_Subst.Frees.add (v, Thm.term_of ct))); + 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 Term_Subst.TFrees.is_empty instT then [] + 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 = - Term_Subst.TVars.build (cinstT |> fold (fn (v, cT) => - Term_Subst.TVars.add (v, Thm.typ_of cT))); - val inst = - Term_Subst.Vars.build (cinst |> fold (fn (v, ct) => - Term_Subst.Vars.add (v, Thm.term_of ct))); + 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 Term_Subst.TVars.is_empty instT then [] + 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/primitive_defs.ML b/src/Pure/primitive_defs.ML --- a/src/Pure/primitive_defs.ML +++ b/src/Pure/primitive_defs.ML @@ -1,85 +1,85 @@ (* Title: Pure/primitive_defs.ML Author: Makarius Primitive definition forms. *) signature PRIMITIVE_DEFS = sig val dest_def: Proof.context -> {check_head: term -> bool, check_free_lhs: string -> bool, check_free_rhs: string -> bool, check_tfree: string -> bool} -> term -> (term * term) * term list * term val abs_def: term -> term * term end; structure Primitive_Defs: PRIMITIVE_DEFS = struct fun term_kind (Const _) = "existing constant " | term_kind (Free _) = "free variable " | term_kind (Bound _) = "bound variable " | term_kind _ = ""; (*c x \ t[x] to \x. c x \ t[x]*) fun dest_def ctxt {check_head, check_free_lhs, check_free_rhs, check_tfree} eq = let fun err msg = raise TERM (msg, [eq]); val eq_vars = Term.strip_all_vars eq; val eq_body = Term.strip_all_body eq; val display_terms = commas_quote o map (Syntax.string_of_term ctxt o Syntax_Trans.bound_vars eq_vars); val display_types = commas_quote o map (Syntax.string_of_typ ctxt); val (raw_lhs, rhs) = Logic.dest_equals eq_body handle TERM _ => err "Not a meta-equality (\)"; val lhs = Envir.beta_eta_contract raw_lhs; val (head, args) = Term.strip_comb lhs; - val head_tfrees = Term_Subst.TFrees.build (Term_Subst.add_tfrees head); + val head_tfrees = TFrees.build (TFrees.add_tfrees head); fun check_arg (Bound _) = true | check_arg (Free (x, _)) = check_free_lhs x | check_arg (Const ("Pure.type", Type ("itself", [TFree _]))) = true | check_arg _ = false; fun close_arg (Bound _) t = t | close_arg x t = Logic.all x t; val lhs_bads = filter_out check_arg args; val lhs_dups = duplicates (op aconv) args; val rhs_extras = Term.fold_aterms (fn v as Free (x, _) => if check_free_rhs x orelse member (op aconv) args v then I else insert (op aconv) v | _ => I) rhs []; val rhs_extrasT = Term.fold_aterms (Term.fold_types (fn v as TFree (a, S) => - if check_tfree a orelse Term_Subst.TFrees.defined head_tfrees (a, S) then I + if check_tfree a orelse TFrees.defined head_tfrees (a, S) then I else insert (op =) v | _ => I)) rhs []; in if not (check_head head) then err ("Bad head of lhs: " ^ term_kind head ^ display_terms [head]) else if not (null lhs_bads) then err ("Bad arguments on lhs: " ^ display_terms lhs_bads) else if not (null lhs_dups) then err ("Duplicate arguments on lhs: " ^ display_terms lhs_dups) else if not (null rhs_extras) then err ("Extra variables on rhs: " ^ display_terms rhs_extras) else if not (null rhs_extrasT) then err ("Extra type variables on rhs: " ^ display_types rhs_extrasT) else if exists_subterm (fn t => t aconv head) rhs then err "Entity to be defined occurs on rhs" else ((lhs, rhs), args, fold_rev close_arg args (Logic.list_all (eq_vars, (Logic.mk_equals (lhs, rhs))))) end; (*\x. c x \ t[x] to c \ \x. t[x]*) fun abs_def eq = let val body = Term.strip_all_body eq; val vars = map Free (Term.rename_wrt_term body (Term.strip_all_vars eq)); val (lhs, rhs) = Logic.dest_equals (Term.subst_bounds (vars, body)); val (lhs', args) = Term.strip_comb lhs; val rhs' = fold_rev (absfree o dest_Free) args rhs; in (lhs', rhs') end; end; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2341 +1,2341 @@ (* Title: Pure/proofterm.ML Author: Stefan Berghofer, TU Muenchen LF style proof terms. *) infix 8 % %% %>; signature PROOFTERM = sig type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option} type thm_body type thm_node datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | % of proof * term option | %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} type oracle = (string * Position.T) * term option type thm = serial * thm_node exception MIN_PROOF of unit val proof_of: proof_body -> proof val join_proof: proof_body future -> proof val map_proof_of: (proof -> proof) -> proof_body -> proof_body val thm_header: serial -> Position.T list -> string -> string -> term -> typ list option -> thm_header val thm_body: proof_body -> thm_body val thm_body_proof_raw: thm_body -> proof val thm_body_proof_open: thm_body -> proof val thm_node_theory_name: thm_node -> string val thm_node_name: thm_node -> string val thm_node_prop: thm_node -> term val thm_node_body: thm_node -> proof_body future val thm_node_thms: thm_node -> thm list val join_thms: thm list -> proof_body list val make_thm: thm_header -> thm_body -> thm val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a val fold_body_thms: ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) -> proof_body list -> 'a -> 'a val oracle_ord: oracle ord val thm_ord: thm ord val unions_oracles: oracle Ord_List.T list -> oracle Ord_List.T val unions_thms: thm Ord_List.T list -> thm Ord_List.T val no_proof_body: proof -> proof_body val no_thm_names: proof -> proof val no_thm_proofs: proof -> proof val no_body_proofs: proof -> proof val encode: Consts.T -> proof XML.Encode.T val encode_body: Consts.T -> proof_body XML.Encode.T val encode_standard_term: Consts.T -> term XML.Encode.T val encode_standard_proof: Consts.T -> proof XML.Encode.T val decode: Consts.T -> proof XML.Decode.T val decode_body: Consts.T -> proof_body XML.Decode.T val %> : proof * term -> proof (*primitive operations*) val proofs: int Unsynchronized.ref val proofs_enabled: unit -> bool val atomic_proof: proof -> bool val compact_proof: proof -> bool val proof_combt: proof * term list -> proof val proof_combt': proof * term option list -> proof val proof_combP: proof * proof list -> proof val strip_combt: proof -> proof * term option list val strip_combP: proof -> proof * proof list val strip_thm_body: proof_body -> proof_body val map_proof_same: term Same.operation -> typ Same.operation -> (typ * class -> proof) -> proof Same.operation val map_proof_terms_same: term Same.operation -> typ Same.operation -> proof Same.operation val map_proof_types_same: typ Same.operation -> proof Same.operation val map_proof_terms: (term -> term) -> (typ -> typ) -> proof -> proof val map_proof_types: (typ -> typ) -> proof -> proof val fold_proof_terms: (term -> 'a -> 'a) -> proof -> 'a -> 'a val fold_proof_terms_types: (term -> 'a -> 'a) -> (typ -> 'a -> 'a) -> proof -> 'a -> 'a val maxidx_proof: proof -> int -> int val size_of_proof: proof -> int val change_types: typ list option -> proof -> proof val prf_abstract_over: term -> proof -> proof val prf_incr_bv: int -> int -> int -> int -> proof -> proof val incr_pboundvars: int -> int -> proof -> proof val prf_loose_bvar1: proof -> int -> bool val prf_loose_Pbvar1: proof -> int -> bool val prf_add_loose_bnos: int -> int -> proof -> int list * int list -> int list * int list val norm_proof: Envir.env -> proof -> proof val norm_proof': Envir.env -> proof -> proof val prf_subst_bounds: term list -> proof -> proof val prf_subst_pbounds: proof list -> proof -> proof val freeze_thaw_prf: proof -> proof * (proof -> proof) (*proof terms for specific inference rules*) val trivial_proof: proof val implies_intr_proof: term -> proof -> proof val implies_intr_proof': term -> proof -> proof val forall_intr_proof: string * term -> typ option -> proof -> proof val forall_intr_proof': term -> proof -> proof - val varify_proof: term -> Term_Subst.TFrees.set -> proof -> proof + val varify_proof: term -> TFrees.set -> proof -> proof val legacy_freezeT: term -> proof -> proof val rotate_proof: term list -> term -> (string * typ) list -> term list -> int -> proof -> proof val permute_prems_proof: term list -> int -> int -> proof -> proof - val generalize_proof: Symtab.set * Symtab.set -> int -> term -> proof -> proof - val instantiate: typ Term_Subst.TVars.table * term Term_Subst.Vars.table -> proof -> proof + val generalize_proof: Names.set * Names.set -> int -> term -> proof -> proof + val instantiate: typ TVars.table * term Vars.table -> proof -> proof val lift_proof: term -> int -> term -> proof -> proof val incr_indexes: int -> proof -> proof val assumption_proof: term list -> term -> int -> proof -> proof val bicompose_proof: bool -> term list -> term list -> term option -> term list -> int -> int -> proof -> proof -> proof val equality_axms: (string * term) list val reflexive_axm: proof val symmetric_axm: proof val transitive_axm: proof val equal_intr_axm: proof val equal_elim_axm: proof val abstract_rule_axm: proof val combination_axm: proof val reflexive_proof: proof val symmetric_proof: proof -> proof val transitive_proof: typ -> term -> proof -> proof -> proof val equal_intr_proof: term -> term -> proof -> proof -> proof val equal_elim_proof: term -> term -> proof -> proof -> proof val abstract_rule_proof: string * term -> proof -> proof val combination_proof: term -> term -> term -> term -> proof -> proof -> proof val strip_shyps_proof: Sorts.algebra -> (typ * sort) list -> (typ * sort) list -> sort list -> proof -> proof val of_sort_proof: Sorts.algebra -> (class * class -> proof) -> (string * class list list * class -> proof) -> (typ * class -> proof) -> typ * sort -> proof list val axm_proof: string -> term -> proof val oracle_proof: string -> term -> proof val shrink_proof: proof -> proof (*rewriting on proof terms*) val add_prf_rrule: proof * proof -> theory -> theory val add_prf_rproc: (typ list -> term option list -> proof -> (proof * proof) option) -> theory -> theory val set_preproc: (theory -> proof -> proof) -> theory -> theory val apply_preproc: theory -> proof -> proof val forall_intr_variables_term: term -> term val forall_intr_variables: term -> proof -> proof val no_skel: proof val normal_skel: proof val rewrite_proof: theory -> (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rewrite_proof_notypes: (proof * proof) list * (typ list -> term option list -> proof -> (proof * proof) option) list -> proof -> proof val rew_proof: theory -> proof -> proof val reconstruct_proof: theory -> term -> proof -> proof val prop_of': term list -> proof -> term val prop_of: proof -> term val expand_name_empty: thm_header -> string option val expand_proof: theory -> (thm_header -> string option) -> proof -> proof val standard_vars: Name.context -> term * proof option -> term * proof option val standard_vars_term: Name.context -> term -> term val add_standard_vars: proof -> (string * typ) list -> (string * typ) list val add_standard_vars_term: term -> (string * typ) list -> (string * typ) list val export_enabled: unit -> bool val export_standard_enabled: unit -> bool val export_proof_boxes_required: theory -> bool val export_proof_boxes: proof_body list -> unit val fulfill_norm_proof: theory -> (serial * proof_body) list -> proof_body -> proof_body val thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> string * Position.T -> sort list -> term list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val unconstrain_thm_proof: theory -> (class * class -> proof) -> (string * class list list * class -> proof) -> sort list -> term -> (serial * proof_body future) list -> proof_body -> thm * proof val get_identity: sort list -> term list -> term -> proof -> {serial: serial, theory_name: string, name: string} option val get_approximative_name: sort list -> term list -> term -> proof -> string type thm_id = {serial: serial, theory_name: string} val make_thm_id: serial * string -> thm_id val thm_header_id: thm_header -> thm_id val thm_id: thm -> thm_id val get_id: sort list -> term list -> term -> proof -> thm_id option val this_id: thm_id option -> thm_id -> bool val proof_boxes: {excluded: thm_id -> bool, included: thm_id -> bool} -> proof list -> (thm_header * proof) list (*exception MIN_PROOF*) end structure Proofterm : PROOFTERM = struct (** datatype proof **) type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option}; datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | op % of proof * term option | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} and thm_body = Thm_Body of {open_proof: proof -> proof, body: proof_body future} and thm_node = Thm_Node of {theory_name: string, name: string, prop: term, body: proof_body future, export: unit lazy, consolidate: unit lazy}; type oracle = (string * Position.T) * term option; val oracle_ord: oracle ord = prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord); type thm = serial * thm_node; val thm_ord: thm ord = fn ((i, _), (j, _)) => int_ord (j, i); exception MIN_PROOF of unit; fun proof_of (PBody {proof, ...}) = proof; val join_proof = Future.join #> proof_of; fun map_proof_of f (PBody {oracles, thms, proof}) = PBody {oracles = oracles, thms = thms, proof = f proof}; fun thm_header serial pos theory_name name prop types : thm_header = {serial = serial, pos = pos, theory_name = theory_name, name = name, prop = prop, types = types}; fun thm_body body = Thm_Body {open_proof = I, body = Future.value body}; fun thm_body_proof_raw (Thm_Body {body, ...}) = join_proof body; fun thm_body_proof_open (Thm_Body {open_proof, body, ...}) = open_proof (join_proof body); fun rep_thm_node (Thm_Node args) = args; val thm_node_theory_name = #theory_name o rep_thm_node; val thm_node_name = #name o rep_thm_node; val thm_node_prop = #prop o rep_thm_node; val thm_node_body = #body o rep_thm_node; val thm_node_thms = thm_node_body #> Future.join #> (fn PBody {thms, ...} => thms); val thm_node_export = #export o rep_thm_node; val thm_node_consolidate = #consolidate o rep_thm_node; fun join_thms (thms: thm list) = Future.joins (map (thm_node_body o #2) thms); val consolidate_bodies = maps (fn PBody {thms, ...} => map (thm_node_consolidate o #2) thms) #> Lazy.consolidate #> map Lazy.force #> ignore; fun make_thm_node theory_name name prop body export = let val consolidate = Lazy.lazy_name "Proofterm.make_thm_node" (fn () => let val PBody {thms, ...} = Future.join body in consolidate_bodies (join_thms thms) end); in Thm_Node {theory_name = theory_name, name = name, prop = prop, body = body, export = export, consolidate = consolidate} end; val no_export = Lazy.value (); fun make_thm ({serial, theory_name, name, prop, ...}: thm_header) (Thm_Body {body, ...}) = (serial, make_thm_node theory_name name prop body no_export); (* proof atoms *) fun fold_proof_atoms all f = let fun app (Abst (_, _, prf)) = app prf | app (AbsP (_, _, prf)) = app prf | app (prf % _) = app prf | app (prf1 %% prf2) = app prf1 #> app prf2 | app (prf as PThm ({serial = i, ...}, Thm_Body {body, ...})) = (fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val (x', seen') = (if all then app (join_proof body) else I) (x, Inttab.update (i, ()) seen) in (f prf x', seen') end) | app prf = (fn (x, seen) => (f prf x, seen)); in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end; fun fold_body_thms f = let fun app (PBody {thms, ...}) = tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val name = thm_node_name thm_node; val prop = thm_node_prop thm_node; val body = Future.join (thm_node_body thm_node); val (x', seen') = app body (x, Inttab.update (i, ()) seen); in (f {serial = i, name = name, prop = prop, body = body} x', seen') end); in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end; (* proof body *) val unions_oracles = Ord_List.unions oracle_ord; val unions_thms = Ord_List.unions thm_ord; fun no_proof_body proof = PBody {oracles = [], thms = [], proof = proof}; val no_thm_body = thm_body (no_proof_body MinProof); fun no_thm_names (Abst (x, T, prf)) = Abst (x, T, no_thm_names prf) | no_thm_names (AbsP (x, t, prf)) = AbsP (x, t, no_thm_names prf) | no_thm_names (prf % t) = no_thm_names prf % t | no_thm_names (prf1 %% prf2) = no_thm_names prf1 %% no_thm_names prf2 | no_thm_names (PThm ({serial, pos, theory_name, name = _, prop, types}, thm_body)) = PThm (thm_header serial pos theory_name "" prop types, thm_body) | no_thm_names a = a; fun no_thm_proofs (Abst (x, T, prf)) = Abst (x, T, no_thm_proofs prf) | no_thm_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_thm_proofs prf) | no_thm_proofs (prf % t) = no_thm_proofs prf % t | no_thm_proofs (prf1 %% prf2) = no_thm_proofs prf1 %% no_thm_proofs prf2 | no_thm_proofs (PThm (header, _)) = PThm (header, no_thm_body) | no_thm_proofs a = a; fun no_body_proofs (Abst (x, T, prf)) = Abst (x, T, no_body_proofs prf) | no_body_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_body_proofs prf) | no_body_proofs (prf % t) = no_body_proofs prf % t | no_body_proofs (prf1 %% prf2) = no_body_proofs prf1 %% no_body_proofs prf2 | no_body_proofs (PThm (header, Thm_Body {open_proof, body})) = let val body' = Future.value (no_proof_body (join_proof body)); val thm_body' = Thm_Body {open_proof = open_proof, body = body'}; in PThm (header, thm_body') end | no_body_proofs a = a; (** XML data representation **) (* encode *) local open XML.Encode Term_XML.Encode; fun proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, b, c) => ([a], pair (option typ) (proof consts) (b, c)), fn AbsP (a, b, c) => ([a], pair (option (term consts)) (proof consts) (b, c)), fn a % b => ([], pair (proof consts) (option (term consts)) (a, b)), fn a %% b => ([], pair (proof consts) (proof consts) (a, b)), fn Hyp a => ([], term consts a), fn PAxm (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PClass (a, b) => ([b], typ a), fn Oracle (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PThm ({serial, pos, theory_name, name, prop, types}, Thm_Body {open_proof, body, ...}) => ([int_atom serial, theory_name, name], pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) (map Position.properties_of pos, (prop, (types, map_proof_of open_proof (Future.join body)))))] and proof_body consts (PBody {oracles, thms, proof = prf}) = triple (list (pair (pair string (properties o Position.properties_of)) (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf) and thm consts (a, thm_node) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node, (Future.join (thm_node_body thm_node)))))); fun standard_term consts t = t |> variant [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), fn Free (a, _) => ([a], []), fn Var (a, _) => (indexname a, []), fn Bound a => ([], int a), fn Abs (a, b, c) => ([a], pair typ (standard_term consts) (b, c)), fn op $ a => ([], pair (standard_term consts) (standard_term consts) a)]; fun standard_proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, SOME b, c) => ([a], pair typ (standard_proof consts) (b, c)), fn AbsP (a, SOME b, c) => ([a], pair (standard_term consts) (standard_proof consts) (b, c)), fn a % SOME b => ([], pair (standard_proof consts) (standard_term consts) (a, b)), fn a %% b => ([], pair (standard_proof consts) (standard_proof consts) (a, b)), fn Hyp a => ([], standard_term consts a), fn PAxm (name, _, SOME Ts) => ([name], list typ Ts), fn PClass (T, c) => ([c], typ T), fn Oracle (name, prop, SOME Ts) => ([name], pair (standard_term consts) (list typ) (prop, Ts)), fn PThm ({serial, theory_name, name, types = SOME Ts, ...}, _) => ([int_atom serial, theory_name, name], list typ Ts)]; in val encode = proof; val encode_body = proof_body; val encode_standard_term = standard_term; val encode_standard_proof = standard_proof; end; (* decode *) local open XML.Decode Term_XML.Decode; fun proof consts prf = prf |> variant [fn ([], []) => MinProof, fn ([], a) => PBound (int a), fn ([a], b) => let val (c, d) = pair (option typ) (proof consts) b in Abst (a, c, d) end, fn ([a], b) => let val (c, d) = pair (option (term consts)) (proof consts) b in AbsP (a, c, d) end, fn ([], a) => op % (pair (proof consts) (option (term consts)) a), fn ([], a) => op %% (pair (proof consts) (proof consts) a), fn ([], a) => Hyp (term consts a), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in PAxm (a, c, d) end, fn ([b], a) => PClass (typ a, b), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in Oracle (a, c, d) end, fn ([a, b, c], d) => let val ((e, (f, (g, h)))) = pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) d; val header = thm_header (int_atom a) (map Position.of_properties e) b c f g; in PThm (header, thm_body h) end] and proof_body consts x = let val (a, b, c) = triple (list (pair (pair string (Position.of_properties o properties)) (option (term consts)))) (list (thm consts)) (proof consts) x; in PBody {oracles = a, thms = b, proof = c} end and thm consts x = let val (a, (b, (c, (d, e)))) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) x in (a, make_thm_node b c d (Future.value e) no_export) end; in val decode = proof; val decode_body = proof_body; end; (** proof objects with different levels of detail **) val proofs = Unsynchronized.ref 2; fun proofs_enabled () = ! proofs >= 2; fun atomic_proof prf = (case prf of Abst _ => false | AbsP _ => false | op % _ => false | op %% _ => false | MinProof => false | _ => true); fun compact_proof (prf % _) = compact_proof prf | compact_proof (prf1 %% prf2) = atomic_proof prf2 andalso compact_proof prf1 | compact_proof prf = atomic_proof prf; fun (prf %> t) = prf % SOME t; val proof_combt = Library.foldl (op %>); val proof_combt' = Library.foldl (op %); val proof_combP = Library.foldl (op %%); fun strip_combt prf = let fun stripc (prf % t, ts) = stripc (prf, t::ts) | stripc x = x in stripc (prf, []) end; fun strip_combP prf = let fun stripc (prf %% prf', prfs) = stripc (prf, prf'::prfs) | stripc x = x in stripc (prf, []) end; fun strip_thm_body (body as PBody {proof, ...}) = (case fst (strip_combt (fst (strip_combP proof))) of PThm (_, Thm_Body {body = body', ...}) => Future.join body' | _ => body); val mk_Abst = fold_rev (fn (x, _: typ) => fn prf => Abst (x, NONE, prf)); val mk_AbsP = fold_rev (fn _: term => fn prf => AbsP ("H", NONE, prf)); fun map_proof_same term typ ofclass = let val typs = Same.map typ; fun proof (Abst (s, T, prf)) = (Abst (s, Same.map_option typ T, Same.commit proof prf) handle Same.SAME => Abst (s, T, proof prf)) | proof (AbsP (s, t, prf)) = (AbsP (s, Same.map_option term t, Same.commit proof prf) handle Same.SAME => AbsP (s, t, proof prf)) | proof (prf % t) = (proof prf % Same.commit (Same.map_option term) t handle Same.SAME => prf % Same.map_option term t) | proof (prf1 %% prf2) = (proof prf1 %% Same.commit proof prf2 handle Same.SAME => prf1 %% proof prf2) | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) | proof (PClass T_c) = ofclass T_c | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) | proof (PThm ({serial, pos, theory_name, name, prop, types = SOME Ts}, thm_body)) = PThm (thm_header serial pos theory_name name prop (SOME (typs Ts)), thm_body) | proof _ = raise Same.SAME; in proof end; fun map_proof_terms_same term typ = map_proof_same term typ (fn (T, c) => PClass (typ T, c)); fun map_proof_types_same typ = map_proof_terms_same (Term_Subst.map_types_same typ) typ; fun same eq f x = let val x' = f x in if eq (x, x') then raise Same.SAME else x' end; fun map_proof_terms f g = Same.commit (map_proof_terms_same (same (op =) f) (same (op =) g)); fun map_proof_types f = Same.commit (map_proof_types_same (same (op =) f)); fun fold_proof_terms f (Abst (_, _, prf)) = fold_proof_terms f prf | fold_proof_terms f (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms f prf | fold_proof_terms f (AbsP (_, NONE, prf)) = fold_proof_terms f prf | fold_proof_terms f (prf % SOME t) = fold_proof_terms f prf #> f t | fold_proof_terms f (prf % NONE) = fold_proof_terms f prf | fold_proof_terms f (prf1 %% prf2) = fold_proof_terms f prf1 #> fold_proof_terms f prf2 | fold_proof_terms _ _ = I; fun fold_proof_terms_types f g (Abst (_, SOME T, prf)) = g T #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (Abst (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf % SOME t) = fold_proof_terms_types f g prf #> f t | fold_proof_terms_types f g (prf % NONE) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf1 %% prf2) = fold_proof_terms_types f g prf1 #> fold_proof_terms_types f g prf2 | fold_proof_terms_types _ g (PAxm (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PClass (T, _)) = g T | fold_proof_terms_types _ g (Oracle (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PThm ({types = SOME Ts, ...}, _)) = fold g Ts | fold_proof_terms_types _ _ _ = I; fun maxidx_proof prf = fold_proof_terms_types Term.maxidx_term Term.maxidx_typ prf; fun size_of_proof (Abst (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (AbsP (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (prf % _) = 1 + size_of_proof prf | size_of_proof (prf1 %% prf2) = size_of_proof prf1 + size_of_proof prf2 | size_of_proof _ = 1; fun change_types types (PAxm (name, prop, _)) = PAxm (name, prop, types) | change_types (SOME [T]) (PClass (_, c)) = PClass (T, c) | change_types types (Oracle (name, prop, _)) = Oracle (name, prop, types) | change_types types (PThm ({serial, pos, theory_name, name, prop, types = _}, thm_body)) = PThm (thm_header serial pos theory_name name prop types, thm_body) | change_types _ prf = prf; (* utilities *) fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t | strip_abs _ t = t; fun mk_abs Ts t = Library.foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); (*Abstraction of a proof term over its occurrences of v, which must contain no loose bound variables. The resulting proof term is ready to become the body of an Abst.*) fun prf_abstract_over v = let fun abst' lev u = if v aconv u then Bound lev else (case u of Abs (a, T, t) => Abs (a, T, abst' (lev + 1) t) | f $ t => (abst' lev f $ absth' lev t handle Same.SAME => f $ abst' lev t) | _ => raise Same.SAME) and absth' lev t = (abst' lev t handle Same.SAME => t); fun abst lev (AbsP (a, t, prf)) = (AbsP (a, Same.map_option (abst' lev) t, absth lev prf) handle Same.SAME => AbsP (a, t, abst lev prf)) | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 handle Same.SAME => prf1 %% abst lev prf2) | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t handle Same.SAME => prf % Same.map_option (abst' lev) t) | abst _ _ = raise Same.SAME and absth lev prf = (abst lev prf handle Same.SAME => prf); in absth 0 end; (*increments a proof term's non-local bound variables required when moving a proof term within abstractions inc is increment for bound variables lev is level at which a bound variable is considered 'loose'*) fun incr_bv' inct tlev t = incr_bv (inct, tlev, t); fun prf_incr_bv' incP _ Plev _ (PBound i) = if i >= Plev then PBound (i+incP) else raise Same.SAME | prf_incr_bv' incP inct Plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (same (op =) (incr_bv' inct tlev)) t, prf_incr_bv incP inct (Plev+1) tlev body) handle Same.SAME => AbsP (a, t, prf_incr_bv' incP inct (Plev+1) tlev body)) | prf_incr_bv' incP inct Plev tlev (Abst (a, T, body)) = Abst (a, T, prf_incr_bv' incP inct Plev (tlev+1) body) | prf_incr_bv' incP inct Plev tlev (prf %% prf') = (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' handle Same.SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t handle Same.SAME => prf % Same.map_option (same (op =) (incr_bv' inct tlev)) t) | prf_incr_bv' _ _ _ _ _ = raise Same.SAME and prf_incr_bv incP inct Plev tlev prf = (prf_incr_bv' incP inct Plev tlev prf handle Same.SAME => prf); fun incr_pboundvars 0 0 prf = prf | incr_pboundvars incP inct prf = prf_incr_bv incP inct 0 0 prf; fun prf_loose_bvar1 (prf1 %% prf2) k = prf_loose_bvar1 prf1 k orelse prf_loose_bvar1 prf2 k | prf_loose_bvar1 (prf % SOME t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k) | prf_loose_bvar1 (_ % NONE) _ = true | prf_loose_bvar1 (AbsP (_, SOME t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k | prf_loose_bvar1 (AbsP (_, NONE, _)) _ = true | prf_loose_bvar1 (Abst (_, _, prf)) k = prf_loose_bvar1 prf (k+1) | prf_loose_bvar1 _ _ = false; fun prf_loose_Pbvar1 (PBound i) k = i = k | prf_loose_Pbvar1 (prf1 %% prf2) k = prf_loose_Pbvar1 prf1 k orelse prf_loose_Pbvar1 prf2 k | prf_loose_Pbvar1 (prf % _) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 (AbsP (_, _, prf)) k = prf_loose_Pbvar1 prf (k+1) | prf_loose_Pbvar1 (Abst (_, _, prf)) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 _ _ = false; fun prf_add_loose_bnos plev _ (PBound i) (is, js) = if i < plev then (is, js) else (insert (op =) (i-plev) is, js) | prf_add_loose_bnos plev tlev (prf1 %% prf2) p = prf_add_loose_bnos plev tlev prf2 (prf_add_loose_bnos plev tlev prf1 p) | prf_add_loose_bnos plev tlev (prf % opt) (is, js) = prf_add_loose_bnos plev tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) = prf_add_loose_bnos (plev+1) tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p = prf_add_loose_bnos plev (tlev+1) prf p | prf_add_loose_bnos _ _ _ _ = ([], []); (* substitutions *) local fun conflicting_tvarsT envT = Term.fold_atyps (fn T => fn instT => (case T of TVar (v as (_, S)) => - if Term_Subst.TVars.defined instT v orelse can (Type.lookup envT) v then instT - else Term_Subst.TVars.update (v, Logic.dummy_tfree S) instT + if TVars.defined instT v orelse can (Type.lookup envT) v then instT + else TVars.add (v, Logic.dummy_tfree S) instT | _ => instT)); fun conflicting_tvars env = Term.fold_aterms (fn t => fn inst => (case t of Var (v as (_, T)) => - if Term_Subst.Vars.defined inst v orelse can (Envir.lookup env) v then inst - else Term_Subst.Vars.update (v, Free ("dummy", T)) inst + if Vars.defined inst v orelse can (Envir.lookup env) v then inst + else Vars.add (v, Free ("dummy", T)) inst | _ => inst)); fun del_conflicting_tvars envT ty = - Term_Subst.instantiateT (Term_Subst.TVars.build (conflicting_tvarsT envT ty)) ty; + Term_Subst.instantiateT (TVars.build (conflicting_tvarsT envT ty)) ty; fun del_conflicting_vars env tm = let val instT = - Term_Subst.TVars.build (tm |> Term.fold_types (conflicting_tvarsT (Envir.type_env env))); - val inst = Term_Subst.Vars.build (tm |> conflicting_tvars env); + TVars.build (tm |> Term.fold_types (conflicting_tvarsT (Envir.type_env env))); + val inst = Vars.build (tm |> conflicting_tvars env); in Term_Subst.instantiate (instT, inst) tm end; in fun norm_proof env = let val envT = Envir.type_env env; fun msg s = warning ("type conflict in norm_proof:\n" ^ s); fun htype f t = f env t handle TYPE (s, _, _) => (msg s; f env (del_conflicting_vars env t)); fun htypeT f T = f envT T handle TYPE (s, _, _) => (msg s; f envT (del_conflicting_tvars envT T)); fun htypeTs f Ts = f envT Ts handle TYPE (s, _, _) => (msg s; f envT (map (del_conflicting_tvars envT) Ts)); fun norm (Abst (s, T, prf)) = (Abst (s, Same.map_option (htypeT Envir.norm_type_same) T, Same.commit norm prf) handle Same.SAME => Abst (s, T, norm prf)) | norm (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (htype Envir.norm_term_same) t, Same.commit norm prf) handle Same.SAME => AbsP (s, t, norm prf)) | norm (prf % t) = (norm prf % Option.map (htype Envir.norm_term) t handle Same.SAME => prf % Same.map_option (htype Envir.norm_term_same) t) | norm (prf1 %% prf2) = (norm prf1 %% Same.commit norm prf2 handle Same.SAME => prf1 %% norm prf2) | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PClass (T, c)) = PClass (htypeT Envir.norm_type_same T, c) | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PThm ({serial = i, pos = p, theory_name, name = a, prop = t, types = Ts}, thm_body)) = PThm (thm_header i p theory_name a t (Same.map_option (htypeTs Envir.norm_types_same) Ts), thm_body) | norm _ = raise Same.SAME; in Same.commit norm end; end; (* remove some types in proof term (to save space) *) fun remove_types (Abs (s, _, t)) = Abs (s, dummyT, remove_types t) | remove_types (t $ u) = remove_types t $ remove_types u | remove_types (Const (s, _)) = Const (s, dummyT) | remove_types t = t; fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) = Envir.Envir {maxidx = maxidx, tenv = Vartab.map (K (apsnd remove_types)) tenv, tyenv = tyenv}; fun norm_proof' env prf = norm_proof (remove_types_env env) prf; (* substitution of bound variables *) fun prf_subst_bounds args prf = let val n = length args; fun subst' lev (Bound i) = (if i Bound (i-n)) (*loose: change it*) | subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body) | subst' lev (f $ t) = (subst' lev f $ substh' lev t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME and substh' lev t = (subst' lev t handle Same.SAME => t); fun subst lev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' lev) t, substh lev body) handle Same.SAME => AbsP (a, t, subst lev body)) | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' handle Same.SAME => prf %% subst lev prf') | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t handle Same.SAME => prf % Same.map_option (subst' lev) t) | subst _ _ = raise Same.SAME and substh lev prf = (subst lev prf handle Same.SAME => prf); in (case args of [] => prf | _ => substh 0 prf) end; fun prf_subst_pbounds args prf = let val n = length args; fun subst (PBound i) Plev tlev = (if i < Plev then raise Same.SAME (*var is locally bound*) else incr_pboundvars Plev tlev (nth args (i-Plev)) handle General.Subscript => PBound (i-n) (*loose: change it*)) | subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev) | subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1)) | subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev handle Same.SAME => prf %% subst prf' Plev tlev) | subst (prf % t) Plev tlev = subst prf Plev tlev % t | subst _ _ _ = raise Same.SAME and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf) in (case args of [] => prf | _ => substh prf 0 0) end; (* freezing and thawing of variables in proof terms *) local fun frzT names = map_type_tvar (fn (ixn, S) => TFree (the (AList.lookup (op =) names ixn), S)); fun thawT names = map_type_tfree (fn (a, S) => (case AList.lookup (op =) names a of NONE => TFree (a, S) | SOME ixn => TVar (ixn, S))); fun freeze names names' (t $ u) = freeze names names' t $ freeze names names' u | freeze names names' (Abs (s, T, t)) = Abs (s, frzT names' T, freeze names names' t) | freeze _ names' (Const (s, T)) = Const (s, frzT names' T) | freeze _ names' (Free (s, T)) = Free (s, frzT names' T) | freeze names names' (Var (ixn, T)) = Free (the (AList.lookup (op =) names ixn), frzT names' T) | freeze _ _ t = t; fun thaw names names' (t $ u) = thaw names names' t $ thaw names names' u | thaw names names' (Abs (s, T, t)) = Abs (s, thawT names' T, thaw names names' t) | thaw _ names' (Const (s, T)) = Const (s, thawT names' T) | thaw names names' (Free (s, T)) = let val T' = thawT names' T in (case AList.lookup (op =) names s of NONE => Free (s, T') | SOME ixn => Var (ixn, T')) end | thaw _ names' (Var (ixn, T)) = Var (ixn, thawT names' T) | thaw _ _ t = t; in fun freeze_thaw_prf prf = let val (fs, Tfs, vs, Tvs) = fold_proof_terms_types (fn t => fn (fs, Tfs, vs, Tvs) => (Term.add_free_names t fs, Term.add_tfree_names t Tfs, Term.add_var_names t vs, Term.add_tvar_names t Tvs)) (fn T => fn (fs, Tfs, vs, Tvs) => (fs, Term.add_tfree_namesT T Tfs, vs, Term.add_tvar_namesT T Tvs)) prf ([], [], [], []); val names = vs ~~ Name.variant_list fs (map fst vs); val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs); val rnames = map swap names; val rnames' = map swap names'; in (map_proof_terms (freeze names names') (frzT names') prf, map_proof_terms (thaw rnames rnames') (thawT rnames')) end; end; (** inference rules **) (* trivial implication *) val trivial_proof = AbsP ("H", NONE, PBound 0); (* implication introduction *) fun gen_implies_intr_proof f h prf = let fun abshyp i (Hyp t) = if h aconv t then PBound i else raise Same.SAME | abshyp i (Abst (s, T, prf)) = Abst (s, T, abshyp i prf) | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i + 1) prf) | abshyp i (prf % t) = abshyp i prf % t | abshyp i (prf1 %% prf2) = (abshyp i prf1 %% abshyph i prf2 handle Same.SAME => prf1 %% abshyp i prf2) | abshyp _ _ = raise Same.SAME and abshyph i prf = (abshyp i prf handle Same.SAME => prf); in AbsP ("H", f h, abshyph 0 prf) end; val implies_intr_proof = gen_implies_intr_proof (K NONE); val implies_intr_proof' = gen_implies_intr_proof SOME; (* forall introduction *) fun forall_intr_proof (a, v) opt_T prf = Abst (a, opt_T, prf_abstract_over v prf); fun forall_intr_proof' v prf = let val (a, T) = (case v of Var ((a, _), T) => (a, T) | Free (a, T) => (a, T)) in forall_intr_proof (a, v) (SOME T) prf end; (* varify *) fun varify_proof t fixed prf = let val fs = build (t |> (Term.fold_types o Term.fold_atyps) - (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); + (fn TFree v => if TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ #1 (fold_map Name.variant (map fst fs) used); fun thaw (a, S) = (case AList.lookup (op =) fmap (a, S) of NONE => TFree (a, S) | SOME b => TVar ((b, 0), S)); in map_proof_terms (map_types (map_type_tfree thaw)) (map_type_tfree thaw) prf end; local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = (case AList.lookup (op =) alist ix of NONE => TVar (ix, sort) | SOME name => TFree (name, sort)); in fun legacy_freezeT t prf = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => prf (*nothing to do!*) | _ => let val frzT = map_type_tvar (freeze_one alist) in map_proof_terms (map_types frzT) frzT prf end) end; end; (* rotate assumptions *) fun rotate_proof Bs Bi' params asms m prf = let val i = length asms; val j = length Bs; in mk_AbsP (Bs @ [Bi']) (proof_combP (prf, map PBound (j downto 1) @ [mk_Abst params (mk_AbsP asms (proof_combP (proof_combt (PBound i, map Bound ((length params - 1) downto 0)), map PBound (((i-m-1) downto 0) @ ((i-1) downto (i-m))))))])) end; (* permute premises *) fun permute_prems_proof prems' j k prf = let val n = length prems' in mk_AbsP prems' (proof_combP (prf, map PBound ((n-1 downto n-j) @ (k-1 downto 0) @ (n-j-1 downto k)))) end; (* generalization *) fun generalize_proof (tfrees, frees) idx prop prf = let val gen = - if Symtab.is_empty frees then [] + if Names.is_empty frees then [] else - fold_aterms (fn Free (x, T) => Symtab.defined frees x ? insert (op =) (x, T) | _ => I) - (Term_Subst.generalize (tfrees, Symtab.empty) idx prop) []; + fold_aterms (fn Free (x, T) => Names.defined frees x ? insert (op =) (x, T) | _ => I) + (Term_Subst.generalize (tfrees, Names.empty) idx prop) []; in prf |> Same.commit (map_proof_terms_same - (Term_Subst.generalize_same (tfrees, Symtab.empty) idx) + (Term_Subst.generalize_same (tfrees, Names.empty) idx) (Term_Subst.generalizeT_same tfrees idx)) |> fold (fn (x, T) => forall_intr_proof (x, Free (x, T)) NONE) gen |> fold_rev (fn (x, T) => fn prf' => prf' %> Var (Name.clean_index (x, idx), T)) gen end; (* instantiation *) fun instantiate (instT, inst) = Same.commit (map_proof_terms_same - (Term_Subst.instantiate_same (instT, Term_Subst.Vars.map (K remove_types) inst)) + (Term_Subst.instantiate_same (instT, Vars.map (K remove_types) inst)) (Term_Subst.instantiateT_same instT)); (* lifting *) fun lift_proof Bi inc prop prf = let fun lift'' Us Ts t = strip_abs Ts (Logic.incr_indexes ([], Us, inc) (mk_abs Ts t)); fun lift' Us Ts (Abst (s, T, prf)) = (Abst (s, Same.map_option (Logic.incr_tvar_same inc) T, lifth' Us (dummyT::Ts) prf) handle Same.SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) | lift' Us Ts (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) handle Same.SAME => AbsP (s, t, lift' Us Ts prf)) | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t handle Same.SAME => prf % Same.map_option (same (op =) (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 handle Same.SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = PAxm (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PClass (T, c)) = PClass (Logic.incr_tvar_same inc T, c) | lift' _ _ (Oracle (s, prop, Ts)) = Oracle (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PThm ({serial = i, pos = p, theory_name, name = s, prop, types = Ts}, thm_body)) = PThm (thm_header i p theory_name s prop ((Same.map_option o Same.map) (Logic.incr_tvar inc) Ts), thm_body) | lift' _ _ _ = raise Same.SAME and lifth' Us Ts prf = (lift' Us Ts prf handle Same.SAME => prf); val ps = map (Logic.lift_all inc Bi) (Logic.strip_imp_prems prop); val k = length ps; fun mk_app b (i, j, prf) = if b then (i-1, j, prf %% PBound i) else (i, j-1, prf %> Bound j); fun lift Us bs i j (Const ("Pure.imp", _) $ A $ B) = AbsP ("H", NONE (*A*), lift Us (true::bs) (i+1) j B) | lift Us bs i j (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t) | lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf, map (fn k => (#3 (fold_rev mk_app bs (i-1, j-1, PBound k)))) (i + k - 1 downto i)); in mk_AbsP ps (lift [] [] 0 0 Bi) end; fun incr_indexes i = Same.commit (map_proof_terms_same (Logic.incr_indexes_same ([], [], i)) (Logic.incr_tvar_same i)); (* proof by assumption *) fun mk_asm_prf t i m = let fun imp_prf _ i 0 = PBound i | imp_prf (Const ("Pure.imp", _) $ A $ B) i m = AbsP ("H", NONE (*A*), imp_prf B (i+1) (m-1)) | imp_prf _ i _ = PBound i; fun all_prf (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), all_prf t) | all_prf t = imp_prf t (~i) m in all_prf t end; fun assumption_proof Bs Bi n prf = mk_AbsP Bs (proof_combP (prf, map PBound (length Bs - 1 downto 0) @ [mk_asm_prf Bi n ~1])); (* composition of object rule with proof state *) fun flatten_params_proof i j n (Const ("Pure.imp", _) $ A $ B, k) = AbsP ("H", NONE (*A*), flatten_params_proof (i+1) j n (B, k)) | flatten_params_proof i j n (Const ("Pure.all", _) $ Abs (a, T, t), k) = Abst (a, NONE (*T*), flatten_params_proof i (j+1) n (t, k)) | flatten_params_proof i j n (_, k) = proof_combP (proof_combt (PBound (k+i), map Bound (j-1 downto 0)), map PBound (remove (op =) (i-n) (i-1 downto 0))); fun bicompose_proof flatten Bs As A oldAs n m rprf sprf = let val lb = length Bs; val la = length As; in mk_AbsP (Bs @ As) (proof_combP (sprf, map PBound (lb + la - 1 downto la)) %% proof_combP (rprf, (if n>0 then [mk_asm_prf (the A) n m] else []) @ map (if flatten then flatten_params_proof 0 0 n else PBound o snd) (oldAs ~~ (la - 1 downto 0)))) end; (** type classes **) fun strip_shyps_proof algebra present witnessed extra prf = let val replacements = present @ witnessed @ map (`Logic.dummy_tfree) extra; fun get_replacement S = replacements |> get_first (fn (T', S') => if Sorts.sort_le algebra (S', S) then SOME T' else NONE); fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_replacement (Type.sort_of_atyp T) of SOME T' => T' | NONE => raise Fail "strip_shyps_proof: bad type variable in proof term"); in Same.commit (map_proof_types_same (Term_Subst.map_atypsT_same replace)) prf end; fun of_sort_proof algebra classrel_proof arity_proof hyps = Sorts.of_sort_derivation algebra {class_relation = fn _ => fn _ => fn (prf, c1) => fn c2 => if c1 = c2 then prf else classrel_proof (c1, c2) %% prf, type_constructor = fn (a, _) => fn dom => fn c => let val Ss = map (map snd) dom and prfs = maps (map fst) dom in proof_combP (arity_proof (a, Ss, c), prfs) end, type_variable = fn typ => map (fn c => (hyps (typ, c), c)) (Type.sort_of_atyp typ)}; (** axioms and theorems **) val add_type_variables = (fold_types o fold_atyps) (insert (op =)); fun type_variables_of t = rev (add_type_variables t []); val add_variables = fold_aterms (fn a => (is_Var a orelse is_Free a) ? insert (op =) a); fun variables_of t = rev (add_variables t []); fun test_args _ [] = true | test_args is (Bound i :: ts) = not (member (op =) is i) andalso test_args (i :: is) ts | test_args _ _ = false; fun is_fun (Type ("fun", _)) = true | is_fun (TVar _) = true | is_fun _ = false; fun vars_of t = map Var (build_rev (Term.add_vars t)); fun add_funvars Ts (vs, t) = if is_fun (fastype_of1 (Ts, t)) then union (op =) vs (map_filter (fn Var (ixn, T) => if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t)) else vs; fun add_npvars q p Ts (vs, Const ("Pure.imp", _) $ t $ u) = add_npvars q p Ts (add_npvars q (not p) Ts (vs, t), u) | add_npvars q p Ts (vs, Const ("Pure.all", Type (_, [Type (_, [T, _]), _])) $ t) = add_npvars q p Ts (vs, if p andalso q then betapply (t, Var (("",0), T)) else t) | add_npvars q p Ts (vs, Abs (_, T, t)) = add_npvars q p (T::Ts) (vs, t) | add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t) and add_npvars' Ts (vs, t) = (case strip_comb t of (Var (ixn, _), ts) => if test_args [] ts then vs else Library.foldl (add_npvars' Ts) (AList.update (op =) (ixn, Library.foldl (add_funvars Ts) ((these ooo AList.lookup) (op =) vs ixn, ts)) vs, ts) | (Abs (_, T, u), ts) => Library.foldl (add_npvars' (T::Ts)) (vs, u :: ts) | (_, ts) => Library.foldl (add_npvars' Ts) (vs, ts)); fun prop_vars (Const ("Pure.imp", _) $ P $ Q) = union (op =) (prop_vars P) (prop_vars Q) | prop_vars (Const ("Pure.all", _) $ Abs (_, _, t)) = prop_vars t | prop_vars t = (case strip_comb t of (Var (ixn, _), _) => [ixn] | _ => []); fun is_proj t = let fun is_p i t = (case strip_comb t of (Bound _, []) => false | (Bound j, ts) => j >= i orelse exists (is_p i) ts | (Abs (_, _, u), _) => is_p (i+1) u | (_, ts) => exists (is_p i) ts) in (case strip_abs_body t of Bound _ => true | t' => is_p 0 t') end; fun prop_args prop = let val needed_vars = union (op =) (Library.foldl (uncurry (union (op =))) ([], map (uncurry (insert (op =))) (add_npvars true true [] ([], prop)))) (prop_vars prop); in variables_of prop |> map (fn var as Var (ixn, _) => if member (op =) needed_vars ixn then SOME var else NONE | free => SOME free) end; fun const_proof mk name prop = let val args = prop_args prop; val ({outer_constraints, ...}, prop1) = Logic.unconstrainT [] prop; val head = mk (name, prop1, NONE); in proof_combP (proof_combt' (head, args), map PClass outer_constraints) end; val axm_proof = const_proof PAxm; val oracle_proof = const_proof Oracle; val shrink_proof = let fun shrink ls lev (prf as Abst (a, T, body)) = let val (b, is, ch, body') = shrink ls (lev+1) body in (b, is, ch, if ch then Abst (a, T, body') else prf) end | shrink ls lev (prf as AbsP (a, t, body)) = let val (b, is, ch, body') = shrink (lev::ls) lev body in (b orelse member (op =) is 0, map_filter (fn 0 => NONE | i => SOME (i-1)) is, ch, if ch then AbsP (a, t, body') else prf) end | shrink ls lev prf = let val (is, ch, _, prf') = shrink' ls lev [] [] prf in (false, is, ch, prf') end and shrink' ls lev ts prfs (prf as prf1 %% prf2) = let val p as (_, is', ch', prf') = shrink ls lev prf2; val (is, ch, ts', prf'') = shrink' ls lev ts (p::prfs) prf1 in (union (op =) is is', ch orelse ch', ts', if ch orelse ch' then prf'' %% prf' else prf) end | shrink' ls lev ts prfs (prf as prf1 % t) = let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1 in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end | shrink' ls lev ts prfs (prf as PBound i) = (if exists (fn SOME (Bound j) => lev-j <= nth ls i | _ => true) ts orelse has_duplicates (op =) (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) | shrink' _ _ ts _ (Hyp t) = ([], false, map (pair false) ts, Hyp t) | shrink' _ _ ts _ (prf as MinProof) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts _ (prf as PClass _) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts prfs prf = let val prop = (case prf of PAxm (_, prop, _) => prop | Oracle (_, prop, _) => prop | PThm ({prop, ...}, _) => prop | _ => raise Fail "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = chop (length vs) ts; val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts'; val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => insert (op =) ixn (case AList.lookup (op =) insts ixn of SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns' | _ => union (op =) ixns ixns')) (needed prop ts'' prfs, add_npvars false true [] ([], prop)); val insts' = map (fn (ixn, x as SOME _) => if member (op =) nvs ixn then (false, x) else (true, NONE) | (_, x) => (false, x)) insts in ([], false, insts' @ map (pair false) ts'', prf) end and needed (Const ("Pure.imp", _) $ t $ u) ts ((b, _, _, _)::prfs) = union (op =) (if b then map (fst o dest_Var) (vars_of t) else []) (needed u ts prfs) | needed (Var (ixn, _)) (_::_) _ = [ixn] | needed _ _ _ = []; in fn prf => #4 (shrink [] 0 prf) end; (** axioms for equality **) val aT = TFree ("'a", []); val bT = TFree ("'b", []); val x = Free ("x", aT); val y = Free ("y", aT); val z = Free ("z", aT); val A = Free ("A", propT); val B = Free ("B", propT); val f = Free ("f", aT --> bT); val g = Free ("g", aT --> bT); val equality_axms = [("reflexive", Logic.mk_equals (x, x)), ("symmetric", Logic.mk_implies (Logic.mk_equals (x, y), Logic.mk_equals (y, x))), ("transitive", Logic.list_implies ([Logic.mk_equals (x, y), Logic.mk_equals (y, z)], Logic.mk_equals (x, z))), ("equal_intr", Logic.list_implies ([Logic.mk_implies (A, B), Logic.mk_implies (B, A)], Logic.mk_equals (A, B))), ("equal_elim", Logic.list_implies ([Logic.mk_equals (A, B), A], B)), ("abstract_rule", Logic.mk_implies (Logic.all x (Logic.mk_equals (f $ x, g $ x)), Logic.mk_equals (lambda x (f $ x), lambda x (g $ x)))), ("combination", Logic.list_implies ([Logic.mk_equals (f, g), Logic.mk_equals (x, y)], Logic.mk_equals (f $ x, g $ y)))]; val [reflexive_axm, symmetric_axm, transitive_axm, equal_intr_axm, equal_elim_axm, abstract_rule_axm, combination_axm] = map (fn (s, t) => PAxm ("Pure." ^ s, Logic.varify_global t, NONE)) equality_axms; val reflexive_proof = reflexive_axm % NONE; val is_reflexive_proof = fn PAxm ("Pure.reflexive", _, _) % _ => true | _ => false; fun symmetric_proof prf = if is_reflexive_proof prf then prf else symmetric_axm % NONE % NONE %% prf; fun transitive_proof U u prf1 prf2 = if is_reflexive_proof prf1 then prf2 else if is_reflexive_proof prf2 then prf1 else if U = propT then transitive_axm % NONE % SOME (remove_types u) % NONE %% prf1 %% prf2 else transitive_axm % NONE % NONE % NONE %% prf1 %% prf2; fun equal_intr_proof A B prf1 prf2 = equal_intr_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun equal_elim_proof A B prf1 prf2 = equal_elim_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun abstract_rule_proof (a, x) prf = abstract_rule_axm % NONE % NONE %% forall_intr_proof (a, x) NONE prf; fun check_comb (PAxm ("Pure.combination", _, _) % f % _ % _ % _ %% prf %% _) = is_some f orelse check_comb prf | check_comb (PAxm ("Pure.transitive", _, _) % _ % _ % _ %% prf1 %% prf2) = check_comb prf1 andalso check_comb prf2 | check_comb (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf) = check_comb prf | check_comb _ = false; fun combination_proof f g t u prf1 prf2 = let val f = Envir.beta_norm f; val g = Envir.beta_norm g; val prf = if check_comb prf1 then combination_axm % NONE % NONE else (case prf1 of PAxm ("Pure.reflexive", _, _) % _ => combination_axm %> remove_types f % NONE | _ => combination_axm %> remove_types f %> remove_types g) in prf % (case head_of f of Abs _ => SOME (remove_types t) | Var _ => SOME (remove_types t) | _ => NONE) % (case head_of g of Abs _ => SOME (remove_types u) | Var _ => SOME (remove_types u) | _ => NONE) %% prf1 %% prf2 end; (** rewriting on proof terms **) (* simple first order matching functions for terms and proofs (see pattern.ML) *) exception PMatch; fun flt (i: int) = filter (fn n => n < i); fun fomatch Ts tymatch j instsp p = let fun mtch (instsp as (tyinsts, insts)) = fn (Var (ixn, T), t) => if j>0 andalso not (null (flt j (loose_bnos t))) then raise PMatch else (tymatch (tyinsts, fn () => (T, fastype_of1 (Ts, t))), (ixn, t) :: insts) | (Free (a, T), Free (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (Const (a, T), Const (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (f $ t, g $ u) => mtch (mtch instsp (f, g)) (t, u) | (Bound i, Bound j) => if i=j then instsp else raise PMatch | _ => raise PMatch in mtch instsp (apply2 Envir.beta_eta_contract p) end; fun match_proof Ts tymatch = let fun optmatch _ inst (NONE, _) = inst | optmatch _ _ (SOME _, NONE) = raise PMatch | optmatch mtch inst (SOME x, SOME y) = mtch inst (x, y) fun matcht Ts j (pinst, tinst) (t, u) = (pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u)); fun matchT (pinst, (tyinsts, insts)) p = (pinst, (tymatch (tyinsts, K p), insts)); fun matchTs inst (Ts, Us) = Library.foldl (uncurry matchT) (inst, Ts ~~ Us); fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) = if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst) else (case apfst (flt i) (apsnd (flt j) (prf_add_loose_bnos 0 0 prf ([], []))) of ([], []) => ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) | ([], _) => if j = 0 then ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) else raise PMatch | _ => raise PMatch) | mtch Ts i j inst (prf1 % opt1, prf2 % opt2) = optmatch (matcht Ts j) (mtch Ts i j inst (prf1, prf2)) (opt1, opt2) | mtch Ts i j inst (prf1 %% prf2, prf1' %% prf2') = mtch Ts i j (mtch Ts i j inst (prf1, prf1')) (prf2, prf2') | mtch Ts i j inst (Abst (_, opT, prf1), Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) (optmatch matchT inst (opT, opU)) (prf1, prf2) | mtch Ts i j inst (prf1, Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) inst (incr_pboundvars 0 1 prf1 %> Bound 0, prf2) | mtch Ts i j inst (AbsP (_, opt, prf1), AbsP (_, opu, prf2)) = mtch Ts (i+1) j (optmatch (matcht Ts j) inst (opt, opu)) (prf1, prf2) | mtch Ts i j inst (prf1, AbsP (_, _, prf2)) = mtch Ts (i+1) j inst (incr_pboundvars 1 0 prf1 %% PBound 0, prf2) | mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) = if s1 = s2 then optmatch matchTs inst (opTs, opUs) else raise PMatch | mtch Ts i j inst (PClass (T1, c1), PClass (T2, c2)) = if c1 = c2 then matchT inst (T1, T2) else raise PMatch | mtch Ts i j inst (PThm ({name = name1, prop = prop1, types = types1, ...}, _), PThm ({name = name2, prop = prop2, types = types2, ...}, _)) = if name1 = name2 andalso prop1 = prop2 then optmatch matchTs inst (types1, types2) else raise PMatch | mtch _ _ _ inst (PBound i, PBound j) = if i = j then inst else raise PMatch | mtch _ _ _ _ _ = raise PMatch in mtch Ts 0 0 end; fun prf_subst (pinst, (tyinsts, insts)) = let val substT = Envir.subst_type_same tyinsts; val substTs = Same.map substT; fun subst' lev (Var (xi, _)) = (case AList.lookup (op =) insts xi of NONE => raise Same.SAME | SOME u => incr_boundvars lev u) | subst' _ (Const (s, T)) = Const (s, substT T) | subst' _ (Free (s, T)) = Free (s, substT T) | subst' lev (Abs (a, T, body)) = (Abs (a, substT T, Same.commit (subst' (lev + 1)) body) handle Same.SAME => Abs (a, T, subst' (lev + 1) body)) | subst' lev (f $ t) = (subst' lev f $ Same.commit (subst' lev) t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME; fun subst plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' tlev) t, Same.commit (subst (plev + 1) tlev) body) handle Same.SAME => AbsP (a, t, subst (plev + 1) tlev body)) | subst plev tlev (Abst (a, T, body)) = (Abst (a, Same.map_option substT T, Same.commit (subst plev (tlev + 1)) body) handle Same.SAME => Abst (a, T, subst plev (tlev + 1) body)) | subst plev tlev (prf %% prf') = (subst plev tlev prf %% Same.commit (subst plev tlev) prf' handle Same.SAME => prf %% subst plev tlev prf') | subst plev tlev (prf % t) = (subst plev tlev prf % Same.commit (Same.map_option (subst' tlev)) t handle Same.SAME => prf % Same.map_option (subst' tlev) t) | subst plev tlev (Hyp (Var (xi, _))) = (case AList.lookup (op =) pinst xi of NONE => raise Same.SAME | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (PClass (T, c)) = PClass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) | subst _ _ (PThm ({serial = i, pos = p, theory_name, name = id, prop, types}, thm_body)) = PThm (thm_header i p theory_name id prop (Same.map_option substTs types), thm_body) | subst _ _ _ = raise Same.SAME; in fn t => subst 0 0 t handle Same.SAME => t end; (*A fast unification filter: true unless the two terms cannot be unified. Terms must be NORMAL. Treats all Vars as distinct. *) fun could_unify prf1 prf2 = let fun matchrands (prf1 %% prf2) (prf1' %% prf2') = could_unify prf2 prf2' andalso matchrands prf1 prf1' | matchrands (prf % SOME t) (prf' % SOME t') = Term.could_unify (t, t') andalso matchrands prf prf' | matchrands (prf % _) (prf' % _) = matchrands prf prf' | matchrands _ _ = true fun head_of (prf %% _) = head_of prf | head_of (prf % _) = head_of prf | head_of prf = prf in case (head_of prf1, head_of prf2) of (_, Hyp (Var _)) => true | (Hyp (Var _), _) => true | (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2 | (PClass (_, c), PClass (_, d)) => c = d andalso matchrands prf1 prf2 | (PThm ({name = a, prop = propa, ...}, _), PThm ({name = b, prop = propb, ...}, _)) => a = b andalso propa = propb andalso matchrands prf1 prf2 | (PBound i, PBound j) => i = j andalso matchrands prf1 prf2 | (AbsP _, _) => true (*because of possible eta equality*) | (Abst _, _) => true | (_, AbsP _) => true | (_, Abst _) => true | _ => false end; (* rewrite proof *) val no_skel = PBound 0; val normal_skel = Hyp (Var ((Name.uu, 0), propT)); fun rewrite_prf tymatch (rules, procs) prf = let fun rew _ _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, no_skel) | rew _ _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, no_skel) | rew Ts hs prf = (case get_first (fn r => r Ts hs prf) procs of NONE => get_first (fn (prf1, prf2) => SOME (prf_subst (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2) handle PMatch => NONE) (filter (could_unify prf o fst) rules) | some => some); fun rew0 Ts hs (prf as AbsP (_, _, prf' %% PBound 0)) = if prf_loose_Pbvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars (~1) 0 prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs (prf as Abst (_, _, prf' % SOME (Bound 0))) = if prf_loose_bvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars 0 (~1) prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs prf = rew Ts hs prf; fun rew1 _ _ (Hyp (Var _)) _ = NONE | rew1 Ts hs skel prf = (case rew2 Ts hs skel prf of SOME prf1 => (case rew0 Ts hs prf1 of SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2)) | NONE => SOME prf1) | NONE => (case rew0 Ts hs prf of SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1)) | NONE => NONE)) and rew2 Ts hs skel (prf % SOME t) = (case prf of Abst (_, _, body) => let val prf' = prf_subst_bounds [t] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of SOME prf' => SOME (prf' % SOME t) | NONE => NONE)) | rew2 Ts hs skel (prf % NONE) = Option.map (fn prf' => prf' % NONE) (rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf) | rew2 Ts hs skel (prf1 %% prf2) = (case prf1 of AbsP (_, _, body) => let val prf' = prf_subst_pbounds [prf2] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => let val (skel1, skel2) = (case skel of skel1 %% skel2 => (skel1, skel2) | _ => (no_skel, no_skel)) in (case rew1 Ts hs skel1 prf1 of SOME prf1' => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1' %% prf2') | NONE => SOME (prf1' %% prf2)) | NONE => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1 %% prf2') | NONE => NONE)) end) | rew2 Ts hs skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts) hs (case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (Abst (s, T, prf')) | NONE => NONE) | rew2 Ts hs skel (AbsP (s, t, prf)) = (case rew1 Ts (t :: hs) (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (AbsP (s, t, prf')) | NONE => NONE) | rew2 _ _ _ _ = NONE; in the_default prf (rew1 [] [] no_skel prf) end; fun rewrite_proof thy = rewrite_prf (fn (tyenv, f) => Sign.typ_match thy (f ()) tyenv handle Type.TYPE_MATCH => raise PMatch); fun rewrite_proof_notypes rews = rewrite_prf fst rews; (* theory data *) structure Data = Theory_Data ( type T = ((stamp * (proof * proof)) list * (stamp * (typ list -> term option list -> proof -> (proof * proof) option)) list) * (theory -> proof -> proof) option; val empty = (([], []), NONE); val extend = I; fun merge (((rules1, procs1), preproc1), ((rules2, procs2), preproc2)) : T = ((AList.merge (op =) (K true) (rules1, rules2), AList.merge (op =) (K true) (procs1, procs2)), merge_options (preproc1, preproc2)); ); fun get_rew_data thy = let val (rules, procs) = #1 (Data.get thy) in (map #2 rules, map #2 procs) end; fun rew_proof thy = rewrite_prf fst (get_rew_data thy); fun add_prf_rrule r = (Data.map o apfst o apfst) (cons (stamp (), r)); fun add_prf_rproc p = (Data.map o apfst o apsnd) (cons (stamp (), p)); fun set_preproc f = (Data.map o apsnd) (K (SOME f)); fun apply_preproc thy = (case #2 (Data.get thy) of NONE => I | SOME f => f thy); (** reconstruction of partial proof terms **) fun forall_intr_variables_term prop = fold_rev Logic.all (variables_of prop) prop; fun forall_intr_variables prop prf = fold_rev forall_intr_proof' (variables_of prop) prf; local fun app_types shift prop Ts prf = let val inst = type_variables_of prop ~~ Ts; fun subst_same A = (case AList.lookup (op =) inst A of SOME T => T | NONE => raise Same.SAME); val subst_type_same = Term_Subst.map_atypsT_same (fn TVar ((a, i), S) => subst_same (TVar ((a, i - shift), S)) | A => subst_same A); in Same.commit (map_proof_types_same subst_type_same) prf end; fun guess_name (PThm ({name, ...}, _)) = name | guess_name (prf %% Hyp _) = guess_name prf | guess_name (prf %% PClass _) = guess_name prf | guess_name (prf % NONE) = guess_name prf | guess_name (prf % SOME (Var _)) = guess_name prf | guess_name _ = ""; (* generate constraints for proof term *) fun mk_var env Ts T = let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end; fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) = (TVar (("'t", maxidx + 1), S), Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}); val mk_abs = fold (fn T => fn u => Abs ("", T, u)); fun unifyT thy env T U = let val Envir.Envir {maxidx, tenv, tyenv} = env; val (tyenv', maxidx') = Sign.typ_unify thy (T, U) (tyenv, maxidx); in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; fun chaseT env (T as TVar v) = (case Type.lookup (Envir.type_env env) v of NONE => T | SOME T' => chaseT env T') | chaseT _ T = T; fun infer_type thy (env as Envir.Envir {maxidx, tenv, tyenv}) _ vTs (t as Const (s, T)) = if T = dummyT then (case Sign.const_type thy s of NONE => error ("reconstruct_proof: No such constant: " ^ quote s) | SOME T => let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T) in (Const (s, T'), T', vTs, Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}) end) else (t, T, vTs, env) | infer_type _ env _ vTs (t as Free (s, T)) = if T = dummyT then (case Symtab.lookup vTs s of NONE => let val (T, env') = mk_tvar [] env in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end | SOME T => (Free (s, T), T, vTs, env)) else (t, T, vTs, env) | infer_type _ _ _ _ (Var _) = error "reconstruct_proof: internal error" | infer_type thy env Ts vTs (Abs (s, T, t)) = let val (T', env') = if T = dummyT then mk_tvar [] env else (T, env); val (t', U, vTs', env'') = infer_type thy env' (T' :: Ts) vTs t in (Abs (s, T', t'), T' --> U, vTs', env'') end | infer_type thy env Ts vTs (t $ u) = let val (t', T, vTs1, env1) = infer_type thy env Ts vTs t; val (u', U, vTs2, env2) = infer_type thy env1 Ts vTs1 u; in (case chaseT env2 T of Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT thy env2 U U') | _ => let val (V, env3) = mk_tvar [] env2 in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end) end | infer_type _ env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env) handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i)); fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^ Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u); fun decompose thy Ts (p as (t, u)) env = let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify thy p else apfst flat (fold_map (decompose thy Ts) (ts ~~ us) (uT env T U)) in case apply2 (strip_comb o Envir.head_norm env) p of ((Const c, ts), (Const d, us)) => rigrig c d (unifyT thy) ts us | ((Free c, ts), (Free d, us)) => rigrig c d (unifyT thy) ts us | ((Bound i, ts), (Bound j, us)) => rigrig (i, dummyT) (j, dummyT) (K o K) ts us | ((Abs (_, T, t), []), (Abs (_, U, u), [])) => decompose thy (T::Ts) (t, u) (unifyT thy env T U) | ((Abs (_, T, t), []), _) => decompose thy (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env | (_, (Abs (_, T, u), [])) => decompose thy (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env | _ => ([(mk_abs Ts t, mk_abs Ts u)], env) end; fun make_constraints_cprf thy env cprf = let fun add_cnstrt Ts prop prf cs env vTs (t, u) = let val t' = mk_abs Ts t; val u' = mk_abs Ts u in (prop, prf, cs, Pattern.unify (Context.Theory thy) (t', u') env, vTs) handle Pattern.Pattern => let val (cs', env') = decompose thy [] (t', u') env in (prop, prf, cs @ cs', env', vTs) end | Pattern.Unif => cantunify thy (Envir.norm_term env t', Envir.norm_term env u') end; fun mk_cnstrts_atom env vTs prop opTs prf = let val prop_types = type_variables_of prop; val (Ts, env') = (case opTs of NONE => fold_map (mk_tvar o Type.sort_of_atyp) prop_types env | SOME Ts => (Ts, env)); val prop' = subst_atomic_types (prop_types ~~ Ts) (forall_intr_variables_term prop) handle ListPair.UnequalLengths => error ("Wrong number of type arguments for " ^ quote (guess_name prf)) in (prop', change_types (SOME Ts) prf, [], env', vTs) end; fun head_norm (prop, prf, cnstrts, env, vTs) = (Envir.head_norm env prop, prf, cnstrts, env, vTs); fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs) handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i)) | mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) = let val (T, env') = (case opT of NONE => mk_tvar [] env | SOME T => (T, env)); val (t, prf, cnstrts, env'', vTs') = mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf; in (Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) = let val (t', _, vTs', env') = infer_type thy env Ts vTs t; val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf; in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'') end | mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) = let val (t, env') = mk_var env Ts propT; val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf; in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) = let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2 in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of (Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') => add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) env'' vTs'' (u, u') | (t, prf1, cnstrts', env'', vTs'') => let val (v, env''') = mk_var env'' Ts propT in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts) env''' vTs'' (t, Logic.mk_implies (u, v)) end) end | mk_cnstrts env Ts Hs vTs (cprf % SOME t) = let val (t', U, vTs1, env1) = infer_type thy env Ts vTs t in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env2, vTs2) => let val env3 = unifyT thy env2 T U in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2) end | (u, prf, cnstrts, env2, vTs2) => let val (v, env3) = mk_var env2 Ts (U --> propT); in add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2 (u, Const ("Pure.all", (U --> propT) --> propT) $ v) end) end | mk_cnstrts env Ts Hs vTs (cprf % NONE) = (case head_norm (mk_cnstrts env Ts Hs vTs cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env', vTs') => let val (t, env'') = mk_var env' Ts T in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs') end | (u, prf, cnstrts, env', vTs') => let val (T, env1) = mk_tvar [] env'; val (v, env2) = mk_var env1 Ts (T --> propT); val (t, env3) = mk_var env2 Ts T in add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs' (u, Const ("Pure.all", (T --> propT) --> propT) $ v) end) | mk_cnstrts env _ _ vTs (prf as PThm ({prop, types = opTs, ...}, _)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PClass (T, c)) = mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf | mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) | mk_cnstrts _ _ _ _ MinProof = raise MIN_PROOF () in mk_cnstrts env [] [] Symtab.empty cprf end; (* update list of free variables of constraints *) fun upd_constrs env cs = let val tenv = Envir.term_env env; val tyenv = Envir.type_env env; val dom = [] |> Vartab.fold (cons o #1) tenv |> Vartab.fold (cons o #1) tyenv; val vran = [] |> Vartab.fold (Term.add_var_names o #2 o #2) tenv |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; fun check_cs [] = [] | check_cs ((u, p, vs) :: ps) = let val vs' = subtract (op =) dom vs in if vs = vs' then (u, p, vs) :: check_cs ps else (true, p, fold (insert op =) vs' vran) :: check_cs ps end; in check_cs cs end; (* solution of constraints *) fun solve _ [] bigenv = bigenv | solve thy cs bigenv = let fun search _ [] = error ("Unsolvable constraints:\n" ^ Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => Syntax.pretty_flexpair (Syntax.init_pretty_global thy) (apply2 (Envir.norm_term bigenv) p)) cs))) | search env ((u, p as (t1, t2), vs)::ps) = if u then let val tn1 = Envir.norm_term bigenv t1; val tn2 = Envir.norm_term bigenv t2 in if Pattern.pattern tn1 andalso Pattern.pattern tn2 then (Pattern.unify (Context.Theory thy) (tn1, tn2) env, ps) handle Pattern.Unif => cantunify thy (tn1, tn2) else let val (cs', env') = decompose thy [] (tn1, tn2) env in if cs' = [(tn1, tn2)] then apsnd (cons (false, (tn1, tn2), vs)) (search env ps) else search env' (map (fn q => (true, q, vs)) cs' @ ps) end end else apsnd (cons (false, p, vs)) (search env ps); val Envir.Envir {maxidx, ...} = bigenv; val (env, cs') = search (Envir.empty maxidx) cs; in solve thy (upd_constrs env cs') (Envir.merge (bigenv, env)) end; in (* reconstruction of proofs *) fun reconstruct_proof thy prop cprf = let val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop); val (t, prf, cs, env, _) = make_constraints_cprf thy (Envir.empty (maxidx_proof cprf ~1)) cprf'; val cs' = map (apply2 (Envir.norm_term env)) ((t, prop') :: cs) |> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) []))); val env' = solve thy cs' env; in thawf (norm_proof env' prf) end handle MIN_PROOF () => MinProof; fun prop_of_atom prop Ts = subst_atomic_types (type_variables_of prop ~~ Ts) (forall_intr_variables_term prop); val head_norm = Envir.head_norm Envir.init; fun prop_of0 Hs (PBound i) = nth Hs i | prop_of0 Hs (Abst (s, SOME T, prf)) = Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf)) | prop_of0 Hs (AbsP (_, SOME t, prf)) = Logic.mk_implies (t, prop_of0 (t :: Hs) prf) | prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of Const ("Pure.all", _) $ f => f $ t | _ => error "prop_of: all expected") | prop_of0 Hs (prf1 %% _) = (case head_norm (prop_of0 Hs prf1) of Const ("Pure.imp", _) $ _ $ Q => Q | _ => error "prop_of: ==> expected") | prop_of0 _ (Hyp t) = t | prop_of0 _ (PThm ({prop, types = SOME Ts, ...}, _)) = prop_of_atom prop Ts | prop_of0 _ (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ (PClass (T, c)) = Logic.mk_of_class (T, c) | prop_of0 _ (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ _ = error "prop_of: partial proof object"; val prop_of' = Envir.beta_eta_contract oo prop_of0; val prop_of = prop_of' []; (* expand and reconstruct subproofs *) fun expand_name_empty (header: thm_header) = if #name header = "" then SOME "" else NONE; fun expand_proof thy expand_name prf = let fun expand seen maxidx (AbsP (s, t, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', AbsP (s, t, prf')) end | expand seen maxidx (Abst (s, T, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', Abst (s, T, prf')) end | expand seen maxidx (prf1 %% prf2) = let val (seen', maxidx', prf1') = expand seen maxidx prf1; val (seen'', maxidx'', prf2') = expand seen' maxidx' prf2; in (seen'', maxidx'', prf1' %% prf2') end | expand seen maxidx (prf % t) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', prf' % t) end | expand seen maxidx (prf as PThm (header, thm_body)) = let val {serial, pos, theory_name, name, prop, types} = header in (case expand_name header of SOME name' => if name' = "" andalso is_some types then let val (seen', maxidx', prf') = (case Inttab.lookup seen serial of NONE => let val prf1 = thm_body_proof_open thm_body |> reconstruct_proof thy prop |> forall_intr_variables prop; val (seen1, maxidx1, prf2) = expand_init seen prf1 val seen2 = seen1 |> Inttab.update (serial, (maxidx1, prf2)); in (seen2, maxidx1, prf2) end | SOME (maxidx1, prf1) => (seen, maxidx1, prf1)); val prf'' = prf' |> incr_indexes (maxidx + 1) |> app_types (maxidx + 1) prop (the types); in (seen', maxidx' + maxidx + 1, prf'') end else if name' <> name then (seen, maxidx, PThm (thm_header serial pos theory_name name' prop types, thm_body)) else (seen, maxidx, prf) | NONE => (seen, maxidx, prf)) end | expand seen maxidx prf = (seen, maxidx, prf) and expand_init seen prf = expand seen (maxidx_proof prf ~1) prf; in #3 (expand_init Inttab.empty prf) end; end; (** promises **) fun fulfill_norm_proof thy ps body0 = let val _ = consolidate_bodies (map #2 ps @ [body0]); val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0; val oracles = unions_oracles (fold (fn (_, PBody {oracles, ...}) => not (null oracles) ? cons oracles) ps [oracles0]); val thms = unions_thms (fold (fn (_, PBody {thms, ...}) => not (null thms) ? cons thms) ps [thms0]); val proof = rew_proof thy proof0; in PBody {oracles = oracles, thms = thms, proof = proof} end; fun fulfill_proof_future thy promises (postproc: proof_body -> proof_body) body = let fun fulfill () = postproc (fulfill_norm_proof thy (map (apsnd Future.join) promises) (Future.join body)); in if null promises then Future.map postproc body else if Future.is_finished body andalso length promises = 1 then Future.map (fn _ => fulfill ()) (snd (hd promises)) else (singleton o Future.forks) {name = "Proofterm.fulfill_proof_future", group = NONE, deps = Future.task_of body :: map (Future.task_of o snd) promises, pri = 1, interrupts = true} fulfill end; (** theorems **) (* standardization of variables for export: only frees and named bounds *) local val declare_names_term = Term.declare_term_frees; val declare_names_term' = fn SOME t => declare_names_term t | NONE => I; val declare_names_proof = fold_proof_terms declare_names_term; fun variant names bs x = #1 (Name.variant x (fold Name.declare bs names)); fun variant_term bs (Abs (x, T, t)) = let val x' = variant (declare_names_term t Name.context) bs x; val t' = variant_term (x' :: bs) t; in Abs (x', T, t') end | variant_term bs (t $ u) = variant_term bs t $ variant_term bs u | variant_term _ t = t; fun variant_proof bs (Abst (x, T, prf)) = let val x' = variant (declare_names_proof prf Name.context) bs x; val prf' = variant_proof (x' :: bs) prf; in Abst (x', T, prf') end | variant_proof bs (AbsP (x, t, prf)) = let val x' = variant (declare_names_term' t (declare_names_proof prf Name.context)) bs x; val t' = Option.map (variant_term bs) t; val prf' = variant_proof (x' :: bs) prf; in AbsP (x', t', prf') end | variant_proof bs (prf % t) = variant_proof bs prf % Option.map (variant_term bs) t | variant_proof bs (prf1 %% prf2) = variant_proof bs prf1 %% variant_proof bs prf2 | variant_proof bs (Hyp t) = Hyp (variant_term bs t) | variant_proof _ prf = prf; val used_frees_type = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I); fun used_frees_term t = fold_types used_frees_type t #> Term.declare_term_frees t; val used_frees_proof = fold_proof_terms_types used_frees_term used_frees_type; val unvarifyT = Term.map_atyps (fn TVar ((a, _), S) => TFree (a, S) | T => T); val unvarify = Term.map_aterms (fn Var ((x, _), T) => Free (x, T) | t => t) #> map_types unvarifyT; val unvarify_proof = map_proof_terms unvarify unvarifyT; fun hidden_types prop proof = let val visible = (fold_types o fold_atyps) (insert (op =)) prop []; val add_hiddenT = fold_atyps (fn T => not (member (op =) visible T) ? insert (op =) T); in rev (fold_proof_terms_types (fold_types add_hiddenT) add_hiddenT proof []) end; fun standard_hidden_types term proof = let val hidden = hidden_types term proof; val idx = Term.maxidx_term term (maxidx_proof proof ~1) + 1; fun smash T = if member (op =) hidden T then (case Type.sort_of_atyp T of [] => dummyT | S => TVar (("'", idx), S)) else T; val smashT = map_atyps smash; in map_proof_terms (map_types smashT) smashT proof end; fun standard_hidden_terms term proof = let fun add_unless excluded x = ((is_Free x orelse is_Var x) andalso not (member (op =) excluded x)) ? insert (op =) x; val visible = fold_aterms (add_unless []) term []; val hidden = fold_proof_terms (fold_aterms (add_unless visible)) proof []; val dummy_term = Term.map_aterms (fn x => if member (op =) hidden x then Term.dummy_pattern (Term.fastype_of x) else x); in proof |> not (null hidden) ? map_proof_terms dummy_term I end; in fun standard_vars used (term, opt_proof) = let val proofs = opt_proof |> Option.map (standard_hidden_types term #> standard_hidden_terms term) |> the_list; val proof_terms = rev (fold (fold_proof_terms_types cons (cons o Logic.mk_type)) proofs []); val used_frees = used |> used_frees_term term |> fold used_frees_proof proofs; val inst = Term_Subst.zero_var_indexes_inst used_frees (term :: proof_terms); val term' = term |> Term_Subst.instantiate inst |> unvarify |> variant_term []; val proofs' = proofs |> map (instantiate inst #> unvarify_proof #> variant_proof []); in (term', try hd proofs') end; fun standard_vars_term used t = #1 (standard_vars used (t, NONE)); val add_standard_vars_term = fold_aterms (fn Free (x, T) => (fn env => (case AList.lookup (op =) env x of NONE => (x, T) :: env | SOME T' => if T = T' then env else raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, [T, T'], []))) | _ => I); val add_standard_vars = fold_proof_terms add_standard_vars_term; end; (* PThm nodes *) fun prune_body body = if Options.default_bool "prune_proofs" then (Future.map o map_proof_of) (K MinProof) body else body; fun export_enabled () = Options.default_bool "export_proofs"; fun export_standard_enabled () = Options.default_bool "export_standard_proofs"; fun export_proof_boxes_required thy = Context.theory_name thy = Context.PureN orelse (export_enabled () andalso not (export_standard_enabled ())); fun export_proof_boxes bodies = let fun export_thm (i, thm_node) boxes = if Inttab.defined boxes i then boxes else boxes |> Inttab.update (i, thm_node_export thm_node) |> fold export_thm (thm_node_thms thm_node); fun export_body (PBody {thms, ...}) = fold export_thm thms; val exports = Inttab.build (fold export_body bodies) |> Inttab.dest; in List.app (Lazy.force o #2) exports end; local fun unconstrainT_proof algebra classrel_proof arity_proof (ucontext: Logic.unconstrain_context) = let fun hyp_map hyp = (case AList.lookup (op =) (#constraints ucontext) hyp of SOME t => Hyp t | NONE => raise Fail "unconstrainT_proof: missing constraint"); val typ = Term_Subst.map_atypsT_same (Type.strip_sorts o #atyp_map ucontext); fun ofclass (ty, c) = let val ty' = Term.map_atyps (#atyp_map ucontext) ty; in the_single (of_sort_proof algebra classrel_proof arity_proof hyp_map (ty', [c])) end; in Same.commit (map_proof_same (Term_Subst.map_types_same typ) typ ofclass) #> fold_rev (implies_intr_proof o snd) (#constraints ucontext) end; fun export_proof thy i prop prf0 = let val prf = prf0 |> reconstruct_proof thy prop |> apply_preproc thy; val (prop', SOME prf') = (prop, SOME prf) |> standard_vars Name.context; val args = [] |> add_standard_vars_term prop' |> add_standard_vars prf' |> rev; val typargs = [] |> Term.add_tfrees prop' |> fold_proof_terms Term.add_tfrees prf' |> rev; val consts = Sign.consts_of thy; val xml = (typargs, (args, (prop', no_thm_names prf'))) |> let open XML.Encode Term_XML.Encode; val encode_vars = list (pair string typ); val encode_term = encode_standard_term consts; val encode_proof = encode_standard_proof consts; in pair (list (pair string sort)) (pair encode_vars (pair encode_term encode_proof)) end; in Export.export_params {theory = thy, binding = Path.binding0 (Path.make ["proofs", string_of_int i]), executable = false, compress = true, strict = false} xml end; fun prepare_thm_proof unconstrain thy classrel_proof arity_proof (name, pos) shyps hyps concl promises body = let val named = name <> ""; val prop = Logic.list_implies (hyps, concl); val args = prop_args prop; val (ucontext, prop1) = Logic.unconstrainT shyps prop; val PBody {oracles = oracles0, thms = thms0, proof = prf} = body; val body0 = Future.value (PBody {oracles = oracles0, thms = thms0, proof = if proofs_enabled () then fold_rev implies_intr_proof hyps prf else MinProof}); fun new_prf () = let val i = serial (); val unconstrainT = unconstrainT_proof (Sign.classes_of thy) classrel_proof arity_proof ucontext; val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); in (i, fulfill_proof_future thy promises postproc body0) end; val (i, body') = (*somewhat non-deterministic proof boxes!*) if export_enabled () then new_prf () else (case strip_combt (fst (strip_combP prf)) of (PThm ({serial = ser, name = a, prop = prop', types = NONE, ...}, thm_body'), args') => if (a = "" orelse a = name) andalso prop' = prop1 andalso args' = args then let val Thm_Body {body = body', ...} = thm_body'; val i = if a = "" andalso named then serial () else ser; in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); val open_proof = not named ? rew_proof thy; val export = if export_enabled () then Lazy.lazy (fn () => join_proof body' |> open_proof |> export_proof thy i prop1 handle exn => if Exn.is_interrupt exn then raise Fail ("Interrupt: potential resource problems while exporting proof " ^ string_of_int i) else Exn.reraise exn) else no_export; val thm_body = prune_body body'; val theory_name = Context.theory_long_name thy; val thm = (i, make_thm_node theory_name name prop1 thm_body export); val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; val head = PThm (header, Thm_Body {open_proof = open_proof, body = thm_body}); val proof = if unconstrain then proof_combt' (head, (map o Option.map o Term.map_types) (#map_atyps ucontext) args) else proof_combP (proof_combt' (head, args), map PClass (#outer_constraints ucontext) @ map Hyp hyps); in (thm, proof) end; in fun thm_proof thy = prepare_thm_proof false thy; fun unconstrain_thm_proof thy classrel_proof arity_proof shyps concl promises body = prepare_thm_proof true thy classrel_proof arity_proof ("", Position.none) shyps [] concl promises body; end; (* PThm identity *) fun get_identity shyps hyps prop prf = let val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)) in (case fst (strip_combt (fst (strip_combP prf))) of PThm ({serial, theory_name, name, prop = prop', ...}, _) => if prop = prop' then SOME {serial = serial, theory_name = theory_name, name = name} else NONE | _ => NONE) end; fun get_approximative_name shyps hyps prop prf = Option.map #name (get_identity shyps hyps prop prf) |> the_default ""; (* thm_id *) type thm_id = {serial: serial, theory_name: string}; fun make_thm_id (serial, theory_name) : thm_id = {serial = serial, theory_name = theory_name}; fun thm_header_id ({serial, theory_name, ...}: thm_header) = make_thm_id (serial, theory_name); fun thm_id (serial, thm_node) : thm_id = make_thm_id (serial, thm_node_theory_name thm_node); fun get_id shyps hyps prop prf : thm_id option = (case get_identity shyps hyps prop prf of NONE => NONE | SOME {name = "", ...} => NONE | SOME {serial, theory_name, ...} => SOME (make_thm_id (serial, theory_name))); fun this_id NONE _ = false | this_id (SOME (thm_id: thm_id)) (thm_id': thm_id) = #serial thm_id = #serial thm_id'; (* proof boxes: intermediate PThm nodes *) fun proof_boxes {included, excluded} proofs = let fun boxes_of (Abst (_, _, prf)) = boxes_of prf | boxes_of (AbsP (_, _, prf)) = boxes_of prf | boxes_of (prf % _) = boxes_of prf | boxes_of (prf1 %% prf2) = boxes_of prf1 #> boxes_of prf2 | boxes_of (PThm (header as {serial = i, ...}, thm_body)) = (fn boxes => let val thm_id = thm_header_id header in if Inttab.defined boxes i orelse (excluded thm_id andalso not (included thm_id)) then boxes else let val prf' = thm_body_proof_open thm_body; val boxes' = Inttab.update (i, (header, prf')) boxes; in boxes_of prf' boxes' end end) | boxes_of MinProof = raise MIN_PROOF () | boxes_of _ = I; in Inttab.fold_rev (cons o #2) (Inttab.build (fold boxes_of proofs)) [] end; end; structure Basic_Proofterm = struct datatype proof = datatype Proofterm.proof datatype proof_body = datatype Proofterm.proof_body val op %> = Proofterm.%> end; open Basic_Proofterm; 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,1463 @@ (* 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 = Term_Subst.TVars.build (fold Term_Subst.add_tvars elhss); - val vars = Term_Subst.Vars.build (fold Term_Subst.add_vars elhss); + 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 (Term_Subst.TVars.defined tvars v) | _ => false)) orelse + (fn TVar v => not (TVars.defined tvars v) | _ => false)) orelse erhs |> Term.exists_subterm - (fn Var v => not (Term_Subst.Vars.defined vars v) | _ => false) + (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 = Term_Subst.Vars.build o Term_Subst.add_vars; +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 Term_Subst.Vars.eq_set (apply2 vars_set (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 Term_Subst.Vars.subset (vars_set rhs, vars_set lhs) then uncond_skel args + 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.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) eq') (Thm.instantiate ([], [(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/term_ord.ML b/src/Pure/term_ord.ML --- a/src/Pure/term_ord.ML +++ b/src/Pure/term_ord.ML @@ -1,234 +1,370 @@ (* Title: Pure/term_ord.ML Author: Tobias Nipkow, TU Muenchen Author: Makarius Term orderings and scalable collections. *) +signature ITEMS = +sig + type key + type 'a table + val empty: 'a table + val build: ('a table -> 'a table) -> 'a table + val is_empty: 'a table -> bool + val map: (key -> 'a -> 'b) -> 'a table -> 'b table + val fold: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a + val fold_rev: (key * 'b -> 'a -> 'a) -> 'b table -> 'a -> 'a + val size: 'a table -> int + val dest: 'a table -> (key * 'a) list + val keys: 'a table -> key list + val exists: (key * 'a -> bool) -> 'a table -> bool + val forall: (key * 'a -> bool) -> 'a table -> bool + val get_first: (key * 'a -> 'b option) -> 'a table -> 'b option + val lookup: 'a table -> key -> 'a option + val defined: 'a table -> key -> bool + val add: key * 'a -> 'a table -> 'a table + val make: (key * 'a) list -> 'a table + type set = unit table + val add_set: key -> set -> set + val make_set: key list -> set + val subset: set * set -> bool + val eq_set: set * set -> bool +end; + +functor Items(Key: KEY): ITEMS = +struct + +structure Table = Table(Key); + +type key = Table.key; +type 'a table = 'a Table.table; + +val empty = Table.empty; +val build = Table.build; +val is_empty = Table.is_empty; +val size = Table.size; +val dest = Table.dest; +val keys = Table.keys; +val exists = Table.exists; +val forall = Table.forall; +val get_first = Table.get_first; +val lookup = Table.lookup; +val defined = Table.defined; + +fun add entry = Table.insert (K true) entry; +fun make entries = Table.build (fold add entries); + +type set = unit table; + +fun add_set x = add (x, ()); +fun make_set xs = build (fold add_set xs); + +fun subset (A: set, B: set) = forall (defined B o #1) A; +fun eq_set (A: set, B: set) = size A = size B andalso subset (A, B); + +val map = Table.map; +val fold = Table.fold; +val fold_rev = Table.fold_rev; + +end; + signature BASIC_TERM_ORD = sig structure Vartab: TABLE structure Sorttab: TABLE structure Typtab: TABLE structure Termtab: TABLE structure Var_Graph: GRAPH structure Sort_Graph: GRAPH structure Typ_Graph: GRAPH structure Term_Graph: GRAPH + structure TFrees: + sig + include ITEMS + val add_tfreesT: typ -> set -> set + val add_tfrees: term -> set -> set + end + structure TVars: + sig + include ITEMS + val add_tvarsT: typ -> set -> set + val add_tvars: term -> set -> set + end + structure Frees: + sig + include ITEMS + val add_frees: term -> set -> set + end + structure Vars: + sig + include ITEMS + val add_vars: term -> set -> set + end + structure Names: + sig + include ITEMS + val add_tfree_namesT: typ -> set -> set + val add_tfree_names: term -> set -> set + val add_free_names: term -> set -> set + end end; signature TERM_ORD = sig include BASIC_TERM_ORD val fast_indexname_ord: indexname ord val sort_ord: sort ord val typ_ord: typ ord val fast_term_ord: term ord val syntax_term_ord: term ord val indexname_ord: indexname ord val tvar_ord: (indexname * sort) ord val var_ord: (indexname * typ) ord val term_ord: term ord val hd_ord: term ord val term_lpo: (term -> int) -> term ord val term_cache: (term -> 'a) -> term -> 'a end; structure Term_Ord: TERM_ORD = struct (* fast syntactic ordering -- tuned for inequalities *) val fast_indexname_ord = pointer_eq_ord (int_ord o apply2 snd ||| fast_string_ord o apply2 fst); val sort_ord = pointer_eq_ord (dict_ord fast_string_ord); local fun cons_nr (TVar _) = 0 | cons_nr (TFree _) = 1 | cons_nr (Type _) = 2; in fun typ_ord TU = if pointer_eq TU then EQUAL else (case TU of (Type (a, Ts), Type (b, Us)) => (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord) | (TFree (a, S), TFree (b, S')) => (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord) | (TVar (xi, S), TVar (yj, S')) => (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord) | (T, U) => int_ord (cons_nr T, cons_nr U)); end; local fun cons_nr (Const _) = 0 | cons_nr (Free _) = 1 | cons_nr (Var _) = 2 | cons_nr (Bound _) = 3 | cons_nr (Abs _) = 4 | cons_nr (_ $ _) = 5; fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u) | struct_ord (t1 $ t2, u1 $ u2) = (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord) | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u); fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u) | atoms_ord (t1 $ t2, u1 $ u2) = (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord) | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b) | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y) | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj) | atoms_ord (Bound i, Bound j) = int_ord (i, j) | atoms_ord _ = EQUAL; fun types_ord (Abs (_, T, t), Abs (_, U, u)) = (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord) | types_ord (t1 $ t2, u1 $ u2) = (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord) | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U) | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U) | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U) | types_ord _ = EQUAL; fun comments_ord (Abs (x, _, t), Abs (y, _, u)) = (case fast_string_ord (x, y) of EQUAL => comments_ord (t, u) | ord => ord) | comments_ord (t1 $ t2, u1 $ u2) = (case comments_ord (t1, u1) of EQUAL => comments_ord (t2, u2) | ord => ord) | comments_ord _ = EQUAL; in val fast_term_ord = pointer_eq_ord (struct_ord ||| atoms_ord ||| types_ord); val syntax_term_ord = fast_term_ord ||| comments_ord; end; (* term_ord *) (*a linear well-founded AC-compatible ordering for terms: s < t <=> 1. size(s) < size(t) or 2. size(s) = size(t) and s=f(...) and t=g(...) and f (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) | (t, u) => (case int_ord (size_of_term t, size_of_term u) of EQUAL => (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of EQUAL => args_ord (t, u) | ord => ord) | ord => ord)) and hd_ord (f, g) = prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g) and args_ord (f $ t, g $ u) = (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord) | args_ord _ = EQUAL; end; (* Lexicographic path order on terms *) (* See Baader & Nipkow, Term rewriting, CUP 1998. Without variables. Const, Var, Bound, Free and Abs are treated all as constants. f_ord maps terms to integers and serves two purposes: - Predicate on constant symbols. Those that are not recognised by f_ord must be mapped to ~1. - Order on the recognised symbols. These must be mapped to distinct integers >= 0. The argument of f_ord is never an application. *) local fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0) | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0) | unrecognized (Var v) = ((1, v), 1) | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2) | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3); fun dest_hd f_ord t = let val ord = f_ord t in if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0) end; fun term_lpo f_ord (s, t) = let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in if forall (fn si => term_lpo f_ord (si, t) = LESS) ss then case hd_ord f_ord (f, g) of GREATER => if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts then GREATER else LESS | EQUAL => if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts then list_ord (term_lpo f_ord) (ss, ts) else LESS | LESS => LESS else GREATER end and hd_ord f_ord (f, g) = case (f, g) of (Abs (_, T, t), Abs (_, U, u)) => (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) | (_, _) => prod_ord (prod_ord int_ord (prod_ord indexname_ord typ_ord)) int_ord (dest_hd f_ord f, dest_hd f_ord g); in val term_lpo = term_lpo end; (* scalable collections *) structure Vartab = Table(type key = indexname val ord = fast_indexname_ord); structure Sorttab = Table(type key = sort val ord = sort_ord); structure Typtab = Table(type key = typ val ord = typ_ord); structure Termtab = Table(type key = term val ord = fast_term_ord); +fun term_cache f = Cache.create Termtab.empty Termtab.lookup Termtab.update f; + structure Var_Graph = Graph(type key = indexname val ord = fast_indexname_ord); structure Sort_Graph = Graph(type key = sort val ord = sort_ord); structure Typ_Graph = Graph(type key = typ val ord = typ_ord); structure Term_Graph = Graph(type key = term val ord = fast_term_ord); -fun term_cache f = Cache.create Termtab.empty Termtab.lookup Termtab.update f; +structure TFrees = +struct + structure Items = + Items(type key = string * sort val ord = pointer_eq_ord (prod_ord fast_string_ord sort_ord)); + open Items; + val add_tfreesT = fold_atyps (fn TFree v => add_set v | _ => I); + val add_tfrees = fold_types add_tfreesT; +end; + +structure TVars = +struct + structure Items = + Items(type key = indexname * sort val ord = pointer_eq_ord (prod_ord fast_indexname_ord sort_ord)); + open Items; + val add_tvarsT = fold_atyps (fn TVar v => add_set v | _ => I); + val add_tvars = fold_types add_tvarsT; +end; + +structure Frees = +struct + structure Items = + Items(type key = string * typ val ord = pointer_eq_ord (prod_ord fast_string_ord typ_ord)); + open Items; + val add_frees = fold_aterms (fn Free v => add_set v | _ => I); +end; + +structure Vars = +struct + structure Items = + Items(type key = indexname * typ val ord = pointer_eq_ord (prod_ord fast_indexname_ord typ_ord)); + open Items; + val add_vars = fold_aterms (fn Var v => add_set v | _ => I); +end; + +structure Names = +struct + structure Items = Items(type key = string val ord = fast_string_ord); + open Items; + val add_tfree_namesT = fold_atyps (fn TFree (a, _) => add_set a | _ => I); + val add_tfree_names = fold_types add_tfree_namesT; + val add_free_names = fold_aterms (fn Free (x, _) => add_set x | _ => I); +end; end; structure Basic_Term_Ord: BASIC_TERM_ORD = Term_Ord; open Basic_Term_Ord; diff --git a/src/Pure/term_subst.ML b/src/Pure/term_subst.ML --- a/src/Pure/term_subst.ML +++ b/src/Pure/term_subst.ML @@ -1,289 +1,224 @@ (* Title: Pure/term_subst.ML Author: Makarius Efficient type/term substitution. *) -signature INST_TABLE = -sig - include TABLE - val add: key * 'a -> 'a table -> 'a table - val table: (key * 'a) list -> 'a table -end; - -functor Inst_Table(Key: KEY): INST_TABLE = -struct - -structure Tab = Table(Key); - -fun add entry = Tab.insert (K true) entry; -fun table entries = Tab.build (fold add entries); - -open Tab; - -end; - signature TERM_SUBST = sig - structure TFrees: INST_TABLE - structure TVars: INST_TABLE - structure Frees: INST_TABLE - structure Vars: INST_TABLE - val add_tfreesT: typ -> TFrees.set -> TFrees.set - val add_tfrees: term -> TFrees.set -> TFrees.set - val add_tvarsT: typ -> TVars.set -> TVars.set - val add_tvars: term -> TVars.set -> TVars.set - val add_frees: term -> Frees.set -> Frees.set - val add_vars: term -> Vars.set -> Vars.set - val add_tfree_namesT: typ -> Symtab.set -> Symtab.set - val add_tfree_names: term -> Symtab.set -> Symtab.set - val add_free_names: term -> Symtab.set -> Symtab.set val map_atypsT_same: typ Same.operation -> typ Same.operation val map_types_same: typ Same.operation -> term Same.operation val map_aterms_same: term Same.operation -> term Same.operation - val generalizeT_same: Symtab.set -> int -> typ Same.operation - val generalize_same: Symtab.set * Symtab.set -> int -> term Same.operation - val generalizeT: Symtab.set -> int -> typ -> typ - val generalize: Symtab.set * Symtab.set -> int -> term -> term + val generalizeT_same: Names.set -> int -> typ Same.operation + val generalize_same: Names.set * Names.set -> int -> term Same.operation + val generalizeT: Names.set -> int -> typ -> typ + val generalize: Names.set * Names.set -> int -> term -> term val instantiateT_maxidx: (typ * int) TVars.table -> typ -> int -> typ * int val instantiate_maxidx: (typ * int) TVars.table * (term * int) Vars.table -> term -> int -> term * int val instantiateT_frees_same: typ TFrees.table -> typ Same.operation val instantiate_frees_same: typ TFrees.table * term Frees.table -> term Same.operation val instantiateT_frees: typ TFrees.table -> typ -> typ val instantiate_frees: typ TFrees.table * term Frees.table -> term -> term val instantiateT_same: typ TVars.table -> typ Same.operation val instantiate_same: typ TVars.table * term Vars.table -> term Same.operation val instantiateT: typ TVars.table -> typ -> typ val instantiate: typ TVars.table * term Vars.table -> term -> term val zero_var_indexes_inst: Name.context -> term list -> typ TVars.table * term Vars.table val zero_var_indexes: term -> term val zero_var_indexes_list: term list -> term list end; structure Term_Subst: TERM_SUBST = struct -(* instantiation as table *) - -structure TFrees = Inst_Table( - type key = string * sort - val ord = pointer_eq_ord (prod_ord fast_string_ord Term_Ord.sort_ord) -); - -structure TVars = Inst_Table( - type key = indexname * sort - val ord = pointer_eq_ord (prod_ord Term_Ord.fast_indexname_ord Term_Ord.sort_ord) -); - -structure Frees = Inst_Table( - type key = string * typ - val ord = pointer_eq_ord (prod_ord fast_string_ord Term_Ord.typ_ord) -); - -structure Vars = Inst_Table( - type key = indexname * typ - val ord = pointer_eq_ord (prod_ord Term_Ord.fast_indexname_ord Term_Ord.typ_ord) -); - -val add_tfreesT = fold_atyps (fn TFree v => TFrees.add (v, ()) | _ => I); -val add_tfrees = fold_types add_tfreesT; -val add_tvarsT = fold_atyps (fn TVar v => TVars.add (v, ()) | _ => I); -val add_tvars = fold_types add_tvarsT; -val add_frees = fold_aterms (fn Free v => Frees.add (v, ()) | _ => I); -val add_vars = fold_aterms (fn Var v => Vars.add (v, ()) | _ => I); -val add_tfree_namesT = fold_atyps (fn TFree (a, _) => Symtab.insert_set a | _ => I); -val add_tfree_names = fold_types add_tfree_namesT; -val add_free_names = fold_aterms (fn Free (x, _) => Symtab.insert_set x | _ => I); - - (* generic mapping *) fun map_atypsT_same f = let fun typ (Type (a, Ts)) = Type (a, Same.map typ Ts) | typ T = f T; in typ end; fun map_types_same f = let fun term (Const (a, T)) = Const (a, f T) | term (Free (a, T)) = Free (a, f T) | term (Var (v, T)) = Var (v, f T) | term (Bound _) = raise Same.SAME | term (Abs (x, T, t)) = (Abs (x, f T, Same.commit term t) handle Same.SAME => Abs (x, T, term t)) | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u); in term end; fun map_aterms_same f = let fun term (Abs (x, T, t)) = Abs (x, T, term t) | term (t $ u) = (term t $ Same.commit term u handle Same.SAME => t $ term u) | term a = f a; in term end; (* generalization of fixed variables *) fun generalizeT_same tfrees idx ty = - if Symtab.is_empty tfrees then raise Same.SAME + if Names.is_empty tfrees then raise Same.SAME else let fun gen (Type (a, Ts)) = Type (a, Same.map gen Ts) | gen (TFree (a, S)) = - if Symtab.defined tfrees a then TVar ((a, idx), S) + if Names.defined tfrees a then TVar ((a, idx), S) else raise Same.SAME | gen _ = raise Same.SAME; in gen ty end; fun generalize_same (tfrees, frees) idx tm = - if Symtab.is_empty tfrees andalso Symtab.is_empty frees then raise Same.SAME + if Names.is_empty tfrees andalso Names.is_empty frees then raise Same.SAME else let val genT = generalizeT_same tfrees idx; fun gen (Free (x, T)) = - if Symtab.defined frees x then + if Names.defined frees x then Var (Name.clean_index (x, idx), Same.commit genT T) else Free (x, genT T) | gen (Var (xi, T)) = Var (xi, genT T) | gen (Const (c, T)) = Const (c, genT T) | gen (Bound _) = raise Same.SAME | gen (Abs (x, T, t)) = (Abs (x, genT T, Same.commit gen t) handle Same.SAME => Abs (x, T, gen t)) | gen (t $ u) = (gen t $ Same.commit gen u handle Same.SAME => t $ gen u); in gen tm end; fun generalizeT names i ty = Same.commit (generalizeT_same names i) ty; fun generalize names i tm = Same.commit (generalize_same names i) tm; (* instantiation of free variables (types before terms) *) fun instantiateT_frees_same instT ty = if TFrees.is_empty instT then raise Same.SAME else let fun subst (Type (a, Ts)) = Type (a, Same.map subst Ts) | subst (TFree v) = (case TFrees.lookup instT v of SOME T => T | NONE => raise Same.SAME) | subst _ = raise Same.SAME; in subst ty end; fun instantiate_frees_same (instT, inst) tm = if TFrees.is_empty instT andalso Frees.is_empty inst then raise Same.SAME else let val substT = instantiateT_frees_same instT; fun subst (Const (c, T)) = Const (c, substT T) | subst (Free (x, T)) = let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in (case Frees.lookup inst (x, T') of SOME t => t | NONE => if same then raise Same.SAME else Free (x, T')) end | subst (Var (xi, T)) = Var (xi, substT T) | subst (Bound _) = raise Same.SAME | subst (Abs (x, T, t)) = (Abs (x, substT T, Same.commit subst t) handle Same.SAME => Abs (x, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst tm end; fun instantiateT_frees instT ty = Same.commit (instantiateT_frees_same instT) ty; fun instantiate_frees inst tm = Same.commit (instantiate_frees_same inst) tm; (* instantiation of schematic variables (types before terms) -- recomputes maxidx *) local fun no_indexesT instT = TVars.map (fn _ => rpair ~1) instT; fun no_indexes inst = Vars.map (fn _ => rpair ~1) inst; fun instT_same maxidx instT ty = let fun maxify i = if i > ! maxidx then maxidx := i else (); fun subst_typ (Type (a, Ts)) = Type (a, subst_typs Ts) | subst_typ (TVar ((a, i), S)) = (case TVars.lookup instT ((a, i), S) of SOME (T, j) => (maxify j; T) | NONE => (maxify i; raise Same.SAME)) | subst_typ _ = raise Same.SAME and subst_typs (T :: Ts) = (subst_typ T :: Same.commit subst_typs Ts handle Same.SAME => T :: subst_typs Ts) | subst_typs [] = raise Same.SAME; in subst_typ ty end; fun inst_same maxidx (instT, inst) tm = let fun maxify i = if i > ! maxidx then maxidx := i else (); val substT = instT_same maxidx instT; fun subst (Const (c, T)) = Const (c, substT T) | subst (Free (x, T)) = Free (x, substT T) | subst (Var ((x, i), T)) = let val (T', same) = (substT T, false) handle Same.SAME => (T, true) in (case Vars.lookup inst ((x, i), T') of SOME (t, j) => (maxify j; t) | NONE => (maxify i; if same then raise Same.SAME else Var ((x, i), T'))) end | subst (Bound _) = raise Same.SAME | subst (Abs (x, T, t)) = (Abs (x, substT T, Same.commit subst t) handle Same.SAME => Abs (x, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst tm end; in fun instantiateT_maxidx instT ty i = let val maxidx = Unsynchronized.ref i in (Same.commit (instT_same maxidx instT) ty, ! maxidx) end; fun instantiate_maxidx insts tm i = let val maxidx = Unsynchronized.ref i in (Same.commit (inst_same maxidx insts) tm, ! maxidx) end; fun instantiateT_same instT ty = if TVars.is_empty instT then raise Same.SAME else instT_same (Unsynchronized.ref ~1) (no_indexesT instT) ty; fun instantiate_same (instT, inst) tm = if TVars.is_empty instT andalso Vars.is_empty inst then raise Same.SAME else inst_same (Unsynchronized.ref ~1) (no_indexesT instT, no_indexes inst) tm; fun instantiateT instT ty = Same.commit (instantiateT_same instT) ty; fun instantiate inst tm = Same.commit (instantiate_same inst) tm; end; (* zero var indexes *) fun zero_var_inst ins mk (v as ((x, i), X)) (inst, used) = let val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used in if x = x' andalso i = 0 then (inst, used') else (ins (v, mk ((x', 0), X)) inst, used') end; fun zero_var_indexes_inst used ts = let val (instT, _) = (TVars.empty, used) |> TVars.fold (zero_var_inst TVars.add TVar o #1) (TVars.build (ts |> (fold o fold_types o fold_atyps) (fn TVar v => TVars.add (v, ()) | _ => I))); val (inst, _) = (Vars.empty, used) |> Vars.fold (zero_var_inst Vars.add Var o #1) (Vars.build (ts |> (fold o fold_aterms) (fn Var (xi, T) => Vars.add ((xi, instantiateT instT T), ()) | _ => I))); in (instT, inst) end; fun zero_var_indexes_list ts = map (instantiate (zero_var_indexes_inst Name.context ts)) ts; val zero_var_indexes = singleton zero_var_indexes_list; end; diff --git a/src/Pure/theory.ML b/src/Pure/theory.ML --- a/src/Pure/theory.ML +++ b/src/Pure/theory.ML @@ -1,342 +1,342 @@ (* Title: Pure/theory.ML Author: Lawrence C Paulson and Markus Wenzel Logical theory content: axioms, definitions, and begin/end wrappers. *) signature THEORY = sig val parents_of: theory -> theory list val ancestors_of: theory -> theory list val nodes_of: theory -> theory list val setup: (theory -> theory) -> unit val local_setup: (Proof.context -> Proof.context) -> unit val install_pure: theory -> unit val get_pure: unit -> theory val get_pure_bootstrap: unit -> theory val get_markup: theory -> Markup.T val check: {long: bool} -> Proof.context -> string * Position.T -> theory val axiom_table: theory -> term Name_Space.table val axiom_space: theory -> Name_Space.T val all_axioms_of: theory -> (string * term) list val defs_of: theory -> Defs.T val at_begin: (theory -> theory option) -> theory -> theory val at_end: (theory -> theory option) -> theory -> theory val join_theory: theory list -> theory val begin_theory: string * Position.T -> theory list -> theory val end_theory: theory -> theory val add_axiom: Proof.context -> binding * term -> theory -> theory val const_dep: theory -> string * typ -> Defs.entry val type_dep: string * typ list -> Defs.entry val add_deps: Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory val add_deps_global: string -> Defs.entry -> Defs.entry list -> theory -> theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> theory val specify_const: (binding * typ) * mixfix -> theory -> term * theory val check_overloading: Proof.context -> bool -> string * typ -> unit end structure Theory: THEORY = struct (** theory context operations **) val parents_of = Context.parents_of; val ancestors_of = Context.ancestors_of; fun nodes_of thy = thy :: ancestors_of thy; fun setup f = Context.>> (Context.map_theory f); fun local_setup f = Context.>> (Context.map_proof f); (* implicit theory Pure *) val pure: theory Single_Assignment.var = Single_Assignment.var "pure"; fun install_pure thy = Single_Assignment.assign pure thy; fun get_pure () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => raise Fail "Theory Pure not present"); fun get_pure_bootstrap () = (case Single_Assignment.peek pure of SOME thy => thy | NONE => Context.the_global_context ()); (** datatype thy **) type wrapper = (theory -> theory option) * stamp; fun apply_wrappers (wrappers: wrapper list) = perhaps (perhaps_loop (perhaps_apply (map fst wrappers))); datatype thy = Thy of {pos: Position.T, id: serial, axioms: term Name_Space.table, defs: Defs.T, wrappers: wrapper list * wrapper list}; fun make_thy (pos, id, axioms, defs, wrappers) = Thy {pos = pos, id = id, axioms = axioms, defs = defs, wrappers = wrappers}; structure Thy = Theory_Data' ( type T = thy; val extend = I; val empty = make_thy (Position.none, 0, Name_Space.empty_table Markup.axiomN, Defs.empty, ([], [])); fun merge old_thys (thy1, thy2) = let val Thy {pos, id, axioms = axioms1, defs = defs1, wrappers = (bgs1, ens1)} = thy1; val Thy {pos = _, id = _, axioms = axioms2, defs = defs2, wrappers = (bgs2, ens2)} = thy2; val axioms' = Name_Space.merge_tables (axioms1, axioms2); val defs' = Defs.merge (Defs.global_context (fst old_thys)) (defs1, defs2); val bgs' = Library.merge (eq_snd op =) (bgs1, bgs2); val ens' = Library.merge (eq_snd op =) (ens1, ens2); in make_thy (pos, id, axioms', defs', (bgs', ens')) end; ); fun rep_theory thy = Thy.get thy |> (fn Thy args => args); fun map_thy f = Thy.map (fn (Thy {pos, id, axioms, defs, wrappers}) => make_thy (f (pos, id, axioms, defs, wrappers))); fun map_axioms f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, f axioms, defs, wrappers)); fun map_defs f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, f defs, wrappers)); fun map_wrappers f = map_thy (fn (pos, id, axioms, defs, wrappers) => (pos, id, axioms, defs, f wrappers)); (* entity markup *) fun theory_markup def name id pos = if id = 0 then Markup.empty else Position.make_entity_markup def id Markup.theoryN (name, pos); fun init_markup (name, pos) thy = let val id = serial (); val _ = Context_Position.reports_global thy [(pos, theory_markup {def = true} name id pos)]; in map_thy (fn (_, _, axioms, defs, wrappers) => (pos, id, axioms, defs, wrappers)) thy end; fun get_markup thy = let val {pos, id, ...} = rep_theory thy in theory_markup {def = false} (Context.theory_long_name thy) id pos end; fun check long ctxt (name, pos) = let val thy = Proof_Context.theory_of ctxt; val thy' = Context.get_theory long thy name handle ERROR msg => let val completion_report = Completion.make_report (name, pos) (fn completed => map (Context.theory_name' long) (ancestors_of thy) |> filter (completed o Long_Name.base_name) |> sort_strings |> map (fn a => (a, (Markup.theoryN, a)))); in error (msg ^ Position.here pos ^ completion_report) end; val _ = Context_Position.report ctxt pos (get_markup thy'); in thy' end; (* basic operations *) val axiom_table = #axioms o rep_theory; val axiom_space = Name_Space.space_of_table o axiom_table; val all_axioms_of = Name_Space.dest_table o axiom_table; val defs_of = #defs o rep_theory; (* join theory *) fun join_theory [] = raise List.Empty | join_theory [thy] = thy | join_theory thys = foldl1 Context.join_thys thys; (* begin/end theory *) val begin_wrappers = rev o #1 o #wrappers o rep_theory; val end_wrappers = rev o #2 o #wrappers o rep_theory; fun at_begin f = map_wrappers (apfst (cons (f, stamp ()))); fun at_end f = map_wrappers (apsnd (cons (f, stamp ()))); fun begin_theory (name, pos) imports = if name = Context.PureN then (case imports of [thy] => init_markup (name, pos) thy | _ => error "Bad bootstrapping of theory Pure") else let val thy = Context.begin_thy name imports; val wrappers = begin_wrappers thy; in thy |> init_markup (name, pos) |> Sign.init_naming |> Sign.local_path |> apply_wrappers wrappers |> tap (Syntax.force_syntax o Sign.syn_of) end; fun end_theory thy = thy |> apply_wrappers (end_wrappers thy) |> Sign.change_check |> Context.finish_thy; (** primitive specifications **) (* raw axioms *) fun cert_axm ctxt (b, raw_tm) = let val thy = Proof_Context.theory_of ctxt; val t = Sign.cert_prop thy raw_tm handle TYPE (msg, _, _) => error msg | TERM (msg, _) => error msg; val _ = Term.no_dummy_patterns t handle TERM (msg, _) => error msg; val bad_sorts = rev ((fold_types o fold_atyps_sorts) (fn (_, []) => I | (T, _) => insert (op =) T) t []); val _ = null bad_sorts orelse error ("Illegal sort constraints in primitive specification: " ^ commas (map (Syntax.string_of_typ (Config.put show_sorts true ctxt)) bad_sorts)); in (b, Sign.no_vars ctxt t) end handle ERROR msg => cat_error msg ("The error(s) above occurred in axiom " ^ Binding.print b); fun add_axiom ctxt raw_axm thy = thy |> map_axioms (fn axioms => let val axm = apsnd Logic.varify_global (cert_axm ctxt raw_axm); val context = ctxt |> Sign.inherit_naming thy |> Context_Position.set_visible_generic false; val (_, axioms') = Name_Space.define context true axm axioms; in axioms' end); (* dependencies *) fun const_dep thy (c, T) = ((Defs.Const, c), Sign.const_typargs thy (c, T)); fun type_dep (c, args) = ((Defs.Type, c), args); fun dependencies (context as (ctxt, _)) unchecked def description lhs rhs = let fun prep (item, args) = (case fold Term.add_tvarsT args [] of [] => (item, map Logic.varifyT_global args) | vs => raise TYPE ("Illegal schematic type variable(s)", map TVar vs, [])); - val lhs_vars = Term_Subst.TFrees.build (snd lhs |> fold Term_Subst.add_tfreesT); + val lhs_vars = TFrees.build (snd lhs |> fold TFrees.add_tfreesT); val rhs_extras = build (rhs |> fold (#2 #> (fold o Term.fold_atyps) - (fn TFree v => not (Term_Subst.TFrees.defined lhs_vars v) ? insert (op =) v | _ => I))); + (fn TFree v => not (TFrees.defined lhs_vars v) ? insert (op =) v | _ => I))); val _ = if null rhs_extras then () else error ("Specification depends on extra type variables: " ^ commas_quote (map (Syntax.string_of_typ ctxt o TFree) rhs_extras) ^ "\nThe error(s) above occurred in " ^ quote description); in Defs.define context unchecked def description (prep lhs) (map prep rhs) end; fun cert_entry thy ((Defs.Const, c), args) = Sign.cert_term thy (Const (c, Sign.const_instance thy (c, args))) |> dest_Const |> const_dep thy | cert_entry thy ((Defs.Type, c), args) = Sign.certify_typ thy (Type (c, args)) |> dest_Type |> type_dep; fun add_deps context a raw_lhs raw_rhs thy = let val (lhs as ((_, lhs_name), _)) :: rhs = map (cert_entry thy) (raw_lhs :: raw_rhs); val description = if a = "" then lhs_name ^ " axiom" else a; in thy |> map_defs (dependencies context false NONE description lhs rhs) end; fun add_deps_global a x y thy = add_deps (Defs.global_context thy) a x y thy; fun specify_const decl thy = let val (t, thy') = Sign.declare_const_global decl thy; in (t, add_deps_global "" (const_dep thy' (dest_Const t)) [] thy') end; (* overloading *) fun check_overloading ctxt overloaded (c, T) = let val thy = Proof_Context.theory_of ctxt; val declT = Sign.the_const_constraint thy c handle TYPE (msg, _, _) => error msg; val T' = Logic.varifyT_global T; fun message sorts txt = [Pretty.block [Pretty.str "Specification of constant ", Pretty.str c, Pretty.str " ::", Pretty.brk 1, Pretty.quote (Syntax.pretty_typ (Config.put show_sorts sorts ctxt) T)], Pretty.str txt] |> Pretty.chunks |> Pretty.string_of; in if Sign.typ_instance thy (declT, T') then () else if Type.raw_instance (declT, T') then error (message true "imposes additional sort constraints on the constant declaration") else if overloaded then () else error (message false "is strictly less general than the declared type (overloading required)") end; (* definitional axioms *) local fun check_def (context as (ctxt, _)) thy unchecked overloaded (b, tm) defs = let val name = Sign.full_name thy b; val ((lhs, rhs), _, _) = Primitive_Defs.dest_def ctxt {check_head = Term.is_Const, check_free_lhs = K true, check_free_rhs = K false, check_tfree = K false} tm handle TERM (msg, _) => error msg; val lhs_const = Term.dest_Const (Term.head_of lhs); val rhs_consts = fold_aterms (fn Const const => insert (op =) (const_dep thy const) | _ => I) rhs []; val rhs_types = (fold_types o fold_subtypes) (fn Type t => insert (op =) (type_dep t) | _ => I) rhs []; val rhs_deps = rhs_consts @ rhs_types; val _ = check_overloading ctxt overloaded lhs_const; in defs |> dependencies context unchecked (SOME name) name (const_dep thy lhs_const) rhs_deps end handle ERROR msg => cat_error msg (Pretty.string_of (Pretty.block [Pretty.str ("The error(s) above occurred in definition " ^ Binding.print b ^ ":"), Pretty.fbrk, Pretty.quote (Syntax.pretty_term ctxt tm)])); in fun add_def (context as (ctxt, _)) unchecked overloaded raw_axm thy = let val axm = cert_axm ctxt raw_axm in thy |> map_defs (check_def context thy unchecked overloaded axm) |> add_axiom ctxt axm end; end; end; diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2377 +1,2375 @@ (* 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 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 (*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 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: Symtab.set * Symtab.set -> int -> thm -> thm - val generalize_cterm: Symtab.set * Symtab.set -> int -> cterm -> cterm - val generalize_ctyp: Symtab.set -> int -> ctyp -> ctyp + 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 trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm - val varifyT_global': Term_Subst.TFrees.set -> thm -> ((string * sort) * indexname) list * 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']); (* 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; (* 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 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 Symtab.is_empty tfrees andalso Symtab.is_empty frees then 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 Symtab.is_empty tfrees then K false - else Term.exists_subtype (fn TFree (a, _) => Symtab.defined tfrees a | _ => false); - fun bad_term (Free (x, T)) = bad_type T orelse Symtab.defined frees x + 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 Symtab.is_empty tfrees andalso Symtab.is_empty frees then ct + 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 Symtab.is_empty tfrees then cT + 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 Term_Subst.TVars.add (v, (U, maxidx)) + if Sign.of_sort thy (U, S) then TVars.add (v, (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 Term_Subst.Vars.add (v, (u, maxidx)) + if T = U then Vars.add (v, (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); val sorts' = sorts |> fold add_instT_sorts instT |> fold add_inst_sorts inst; - val instT' = Term_Subst.TVars.build (fold (add_instT thy') instT); - val inst' = Term_Subst.Vars.build (fold (add_inst thy') inst); + val instT' = TVars.build (fold (add_instT thy') instT); + val inst' = Vars.build (fold (add_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; val thy' = Context.certificate_theory cert'; val constraints' = - Term_Subst.TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) + TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 - (fn d => - Proofterm.instantiate - (Term_Subst.TVars.map (K #1) instT', Term_Subst.Vars.map (K #1) inst') d) der, + (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 ([], []) 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]); 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 Term_Subst.add_tfrees hyps fixed; + 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' Term_Subst.TFrees.empty; +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 (* 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/type.ML b/src/Pure/type.ML --- a/src/Pure/type.ML +++ b/src/Pure/type.ML @@ -1,714 +1,714 @@ (* Title: Pure/type.ML Author: Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel Type signatures and certified types, special treatment of type vars, matching and unification of types, extend and merge type signatures. *) signature TYPE = sig (*constraints*) val mark_polymorphic: typ -> typ val constraint: typ -> term -> term val constraint_type: Proof.context -> typ -> typ val strip_constraints: term -> term val appl_error: Proof.context -> term -> typ -> term -> typ -> string (*type signatures and certified types*) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal type tsig val eq_tsig: tsig * tsig -> bool val rep_tsig: tsig -> {classes: Name_Space.T * Sorts.algebra, default: sort, types: decl Name_Space.table, log_types: string list} val change_base: bool -> tsig -> tsig val change_ignore: tsig -> tsig val empty_tsig: tsig val class_space: tsig -> Name_Space.T val defaultS: tsig -> sort val logical_types: tsig -> string list val eq_sort: tsig -> sort * sort -> bool val subsort: tsig -> sort * sort -> bool val of_sort: tsig -> typ * sort -> bool val inter_sort: tsig -> sort * sort -> sort val cert_class: tsig -> class -> class val cert_sort: tsig -> sort -> sort val minimize_sort: tsig -> sort -> sort val witness_sorts: tsig -> (typ * sort) list -> sort list -> (typ * sort) list type mode val mode_default: mode val mode_syntax: mode val mode_abbrev: mode val get_mode: Proof.context -> mode val set_mode: mode -> Proof.context -> Proof.context val restore_mode: Proof.context -> Proof.context -> Proof.context val type_space: tsig -> Name_Space.T val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig val check_decl: Context.generic -> tsig -> xstring * Position.T -> (string * Position.report list) * decl val the_decl: tsig -> string * Position.T -> decl val cert_typ_mode: mode -> tsig -> typ -> typ val cert_typ: tsig -> typ -> typ val arity_number: tsig -> string -> int val arity_sorts: Context.generic -> tsig -> string -> sort -> sort list (*special treatment of type vars*) val sort_of_atyp: typ -> sort val strip_sorts: typ -> typ val strip_sorts_dummy: typ -> typ val no_tvars: typ -> typ - val varify_global: Term_Subst.TFrees.set -> term -> ((string * sort) * indexname) list * term + val varify_global: TFrees.set -> term -> ((string * sort) * indexname) list * term val legacy_freeze_thaw_type: typ -> typ * (typ -> typ) val legacy_freeze_type: typ -> typ val legacy_freeze_thaw: term -> term * (term -> term) val legacy_freeze: term -> term (*matching and unification*) exception TYPE_MATCH type tyenv = (sort * typ) Vartab.table val lookup: tyenv -> indexname * sort -> typ option val devar: tyenv -> typ -> typ val typ_match: tsig -> typ * typ -> tyenv -> tyenv val typ_instance: tsig -> typ * typ -> bool val raw_match: typ * typ -> tyenv -> tyenv val raw_matches: typ list * typ list -> tyenv -> tyenv val could_match: typ * typ -> bool val could_matches: typ list * typ list -> bool val raw_instance: typ * typ -> bool exception TUNIFY val unify: tsig -> typ * typ -> tyenv * int -> tyenv * int val raw_unify: typ * typ -> tyenv -> tyenv val raw_unifys: typ list * typ list -> tyenv -> tyenv val could_unify: typ * typ -> bool val could_unifys: typ list * typ list -> bool val unified: tyenv -> typ * typ -> bool (*extend and merge type signatures*) val add_class: Context.generic -> binding * class list -> tsig -> tsig val hide_class: bool -> string -> tsig -> tsig val set_defsort: sort -> tsig -> tsig val add_type: Context.generic -> binding * int -> tsig -> tsig val add_abbrev: Context.generic -> binding * string list * typ -> tsig -> tsig val add_nonterminal: Context.generic -> binding -> tsig -> tsig val hide_type: bool -> string -> tsig -> tsig val add_arity: Context.generic -> arity -> tsig -> tsig val add_classrel: Context.generic -> class * class -> tsig -> tsig val merge_tsig: Context.generic -> tsig * tsig -> tsig end; structure Type: TYPE = struct (** constraints **) (*indicate polymorphic Vars*) fun mark_polymorphic T = Type ("_polymorphic_", [T]); fun constraint T t = if T = dummyT then t else Const ("_type_constraint_", T --> T) $ t; fun constraint_type ctxt T = let fun err () = error ("Malformed internal type constraint: " ^ Syntax.string_of_typ ctxt T); in (case T of Type ("fun", [A, B]) => if A = B then A else err () | _ => err ()) end; fun strip_constraints (Const ("_type_constraint_", _) $ t) = strip_constraints t | strip_constraints (t $ u) = strip_constraints t $ strip_constraints u | strip_constraints (Abs (x, T, t)) = Abs (x, T, strip_constraints t) | strip_constraints a = a; fun appl_error ctxt (Const ("_type_constraint_", Type ("fun", [T, _]))) _ u U = cat_lines ["Failed to meet type constraint:", "", Pretty.string_of (Pretty.block [Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U]), Pretty.string_of (Pretty.block [Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt T])] | appl_error ctxt t T u U = cat_lines ["Type error in application: " ^ (case T of Type ("fun", _) => "incompatible operand type" | _ => "operator not of function type"), "", Pretty.string_of (Pretty.block [Pretty.str "Operator:", Pretty.brk 2, Syntax.pretty_term ctxt t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt T]), Pretty.string_of (Pretty.block [Pretty.str "Operand:", Pretty.brk 3, Syntax.pretty_term ctxt u, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U])]; (** type signatures and certified types **) (* type declarations *) datatype decl = LogicalType of int | Abbreviation of string list * typ * bool | Nonterminal; (* type tsig *) datatype tsig = TSig of { classes: Name_Space.T * Sorts.algebra, (*order-sorted algebra of type classes*) default: sort, (*default sort on input*) types: decl Name_Space.table, (*declared types*) log_types: string list}; (*logical types sorted by number of arguments*) fun eq_tsig (TSig {classes = classes1, default = default1, types = types1, log_types = _}, TSig {classes = classes2, default = default2, types = types2, log_types = _}) = pointer_eq (classes1, classes2) andalso default1 = default2 andalso pointer_eq (types1, types2); fun rep_tsig (TSig comps) = comps; fun make_tsig (classes, default, types, log_types) = TSig {classes = classes, default = default, types = types, log_types = log_types}; fun change_base begin (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_base begin types, log_types); fun change_ignore (TSig {classes, default, types, log_types}) = make_tsig (classes, default, Name_Space.change_ignore types, log_types); fun build_tsig (classes, default, types) = let val log_types = Name_Space.fold_table (fn (c, LogicalType n) => cons (c, n) | _ => I) types [] |> Library.sort (int_ord o apply2 snd) |> map fst; in make_tsig (classes, default, types, log_types) end; fun map_tsig f (TSig {classes, default, types, log_types = _}) = build_tsig (f (classes, default, types)); val empty_tsig = build_tsig ((Name_Space.empty Markup.classN, Sorts.empty_algebra), [], Name_Space.empty_table Markup.type_nameN); (* classes and sorts *) val class_space = #1 o #classes o rep_tsig; fun defaultS (TSig {default, ...}) = default; fun logical_types (TSig {log_types, ...}) = log_types; fun eq_sort (TSig {classes, ...}) = Sorts.sort_eq (#2 classes); fun subsort (TSig {classes, ...}) = Sorts.sort_le (#2 classes); fun of_sort (TSig {classes, ...}) = Sorts.of_sort (#2 classes); fun inter_sort (TSig {classes, ...}) = Sorts.inter_sort (#2 classes); fun cert_class (TSig {classes = (_, algebra), ...}) c = if can (Graph.get_entry (Sorts.classes_of algebra)) c then c else raise TYPE ("Undeclared class: " ^ quote c, [], []); val cert_sort = map o cert_class; fun minimize_sort (TSig {classes, ...}) = Sorts.minimize_sort (#2 classes); fun witness_sorts (TSig {classes, log_types, ...}) = Sorts.witness_sorts (#2 classes) log_types; (* certification mode *) datatype mode = Mode of {normalize: bool, logical: bool}; val mode_default = Mode {normalize = true, logical = true}; val mode_syntax = Mode {normalize = true, logical = false}; val mode_abbrev = Mode {normalize = false, logical = false}; structure Mode = Proof_Data ( type T = mode; fun init _ = mode_default; ); val get_mode = Mode.get; fun set_mode mode = Mode.map (K mode); fun restore_mode ctxt = set_mode (get_mode ctxt); (* types *) val type_space = Name_Space.space_of_table o #types o rep_tsig; fun type_alias naming binding name = map_tsig (fn (classes, default, types) => (classes, default, (Name_Space.alias_table naming binding name types))); fun undecl_type c = "Undeclared type constructor: " ^ quote c; fun lookup_type (TSig {types, ...}) = Name_Space.lookup types; fun check_decl context (TSig {types, ...}) (c, pos) = Name_Space.check_reports context types (c, [pos]); fun the_decl tsig (c, pos) = (case lookup_type tsig c of NONE => error (undecl_type c ^ Position.here pos) | SOME decl => decl); (* certified types *) fun bad_nargs t = "Bad number of arguments for type constructor: " ^ quote t; local fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts) | inst_typ env (T as TFree (x, _)) = the_default T (AList.lookup (op =) env x) | inst_typ _ T = T; in fun cert_typ_mode (Mode {normalize, logical}) tsig ty = let fun err msg = raise TYPE (msg, [ty], []); val check_logical = if logical then fn c => err ("Illegal occurrence of syntactic type: " ^ quote c) else fn _ => (); fun cert (T as Type (c, Ts)) = let val Ts' = map cert Ts; fun nargs n = if length Ts <> n then err (bad_nargs c) else (); in (case the_decl tsig (c, Position.none) of LogicalType n => (nargs n; Type (c, Ts')) | Abbreviation (vs, U, syn) => (nargs (length vs); if syn then check_logical c else (); if normalize then inst_typ (vs ~~ Ts') U else Type (c, Ts')) | Nonterminal => (nargs 0; check_logical c; T)) end | cert (TFree (x, S)) = TFree (x, cert_sort tsig S) | cert (TVar (xi as (_, i), S)) = if i < 0 then err ("Malformed type variable: " ^ quote (Term.string_of_vname xi)) else TVar (xi, cert_sort tsig S); val ty' = cert ty; in if ty = ty' then ty else ty' end; (*avoid copying of already normal type*) val cert_typ = cert_typ_mode mode_default; end; (* type arities *) fun arity_number tsig a = (case lookup_type tsig a of SOME (LogicalType n) => n | _ => error (undecl_type a)); fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) [] | arity_sorts context (TSig {classes, ...}) a S = Sorts.mg_domain (#2 classes) a S handle Sorts.CLASS_ERROR err => error (Sorts.class_error context err); (** special treatment of type vars **) (* sort_of_atyp *) fun sort_of_atyp (TFree (_, S)) = S | sort_of_atyp (TVar (_, S)) = S | sort_of_atyp T = raise TYPE ("sort_of_atyp", [T], []); (* strip_sorts *) val strip_sorts = map_atyps (fn TFree (x, _) => TFree (x, []) | TVar (xi, _) => TVar (xi, [])); val strip_sorts_dummy = map_atyps (fn TFree (x, _) => TFree (x, dummyS) | TVar (xi, _) => TVar (xi, dummyS)); (* no_tvars *) fun no_tvars T = (case Term.add_tvarsT T [] of [] => T | vs => raise TYPE ("Illegal schematic type variable(s): " ^ commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], [])); (* varify_global *) fun varify_global fixed t = let val fs = build (t |> (Term.fold_types o Term.fold_atyps) - (fn TFree v => if Term_Subst.TFrees.defined fixed v then I else insert (op =) v | _ => I)); + (fn TFree v => if TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.context |> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t; val fmap = fs ~~ map (rpair 0) (#1 (fold_map Name.variant (map fst fs) used)); fun thaw (f as (_, S)) = (case AList.lookup (op =) fmap f of NONE => TFree f | SOME xi => TVar (xi, S)); in (fmap, map_types (map_type_tfree thaw) t) end; (* freeze_thaw: freeze TVars in a term; return the "thaw" inverse *) local fun new_name ix (pairs, used) = let val v = singleton (Name.variant_list used) (string_of_indexname ix) in ((ix, v) :: pairs, v :: used) end; fun freeze_one alist (ix, sort) = TFree (the (AList.lookup (op =) alist ix), sort) handle Option.Option => raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []); fun thaw_one alist (a, sort) = TVar (the (AList.lookup (op =) alist a), sort) handle Option.Option => TFree (a, sort); in fun legacy_freeze_thaw_type T = let val used = Term.add_tfree_namesT T []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvarsT T [])) ([], used); in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end; val legacy_freeze_type = #1 o legacy_freeze_thaw_type; fun legacy_freeze_thaw t = let val used = Term.add_tfree_names t []; val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used); in (case alist of [] => (t, fn x => x) (*nothing to do!*) | _ => (map_types (map_type_tvar (freeze_one alist)) t, map_types (map_type_tfree (thaw_one (map swap alist))))) end; val legacy_freeze = #1 o legacy_freeze_thaw; end; (** matching and unification of types **) type tyenv = (sort * typ) Vartab.table; 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 Vartab.lookup tye ixn of NONE => NONE | SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S'); (* matching *) exception TYPE_MATCH; fun typ_match tsig = let fun match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => if of_sort tsig (T, S) then Vartab.update_new (v, (S, T)) subs else raise TYPE_MATCH | SOME U => if U = T then subs else raise TYPE_MATCH) | match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else matches (Ts, Us) subs | match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | match _ _ = raise TYPE_MATCH and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs) | matches _ subs = subs; in match end; fun typ_instance tsig (T, U) = (Vartab.build (typ_match tsig (U, T)); true) handle TYPE_MATCH => false; (*purely structural matching*) fun raw_match (TVar (v, S), T) subs = (case lookup subs (v, S) of NONE => Vartab.update_new (v, (S, T)) subs | SOME U => if U = T then subs else raise TYPE_MATCH) | raw_match (Type (a, Ts), Type (b, Us)) subs = if a <> b then raise TYPE_MATCH else raw_matches (Ts, Us) subs | raw_match (TFree x, TFree y) subs = if x = y then subs else raise TYPE_MATCH | raw_match _ _ = raise TYPE_MATCH and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs) | raw_matches ([], []) subs = subs | raw_matches _ _ = raise TYPE_MATCH; (*fast matching filter*) fun could_match (Type (a, Ts), Type (b, Us)) = a = b andalso could_matches (Ts, Us) | could_match (TFree (a, _), TFree (b, _)) = a = b | could_match (TVar _, _) = true | could_match _ = false and could_matches (T :: Ts, U :: Us) = could_match (T, U) andalso could_matches (Ts, Us) | could_matches ([], []) = true | could_matches _ = false; fun raw_instance (T, U) = if could_match (U, T) then (Vartab.build (raw_match (U, T)); true) handle TYPE_MATCH => false else false; (* unification *) exception TUNIFY; (*occurs check*) fun occurs v tye = let fun occ (Type (_, Ts)) = exists occ Ts | occ (TFree _) = false | occ (TVar (w, S)) = Term.eq_ix (v, w) orelse (case lookup tye (w, S) of NONE => false | SOME U => occ U); in occ end; (*chase variable assignments; if devar returns a type var then it must be unassigned*) fun devar tye (T as TVar v) = (case lookup tye v of SOME U => devar tye U | NONE => T) | devar _ T = T; (*order-sorted unification*) fun unify (TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) = let val tyvar_count = Unsynchronized.ref maxidx; fun gen_tyvar S = TVar ((Name.aT, Unsynchronized.inc tyvar_count), S); fun mg_domain a S = Sorts.mg_domain classes a S handle Sorts.CLASS_ERROR _ => raise TUNIFY; fun meet (_, []) tye = tye | meet (TVar (xi, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else Vartab.update_new (xi, (S', gen_tyvar (Sorts.inter_sort classes (S', S)))) tye | meet (TFree (_, S'), S) tye = if Sorts.sort_le classes (S', S) then tye else raise TUNIFY | meet (Type (a, Ts), S) tye = meets (Ts, mg_domain a S) tye and meets (T :: Ts, S :: Ss) tye = meets (Ts, Ss) (meet (devar tye T, S) tye) | meets _ tye = tye; fun unif (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), U as TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else if Sorts.sort_le classes (S1, S2) then Vartab.update_new (w, (S2, T)) tye else if Sorts.sort_le classes (S2, S1) then Vartab.update_new (v, (S1, U)) tye else let val S = gen_tyvar (Sorts.inter_sort classes (S1, S2)) in Vartab.update_new (v, (S1, S)) (Vartab.update_new (w, (S2, S)) tye) end | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else meet (T, S) (Vartab.update_new (v, (S, T)) tye) | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else unifs (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and unifs (T :: Ts, U :: Us) tye = unifs (Ts, Us) (unif (T, U) tye) | unifs _ tye = tye; in (unif TU tyenv, ! tyvar_count) end; (*purely structural unification*) fun raw_unify (ty1, ty2) tye = (case (devar tye ty1, devar tye ty2) of (T as TVar (v, S1), TVar (w, S2)) => if Term.eq_ix (v, w) then if S1 = S2 then tye else tvar_clash v S1 S2 else Vartab.update_new (w, (S2, T)) tye | (TVar (v, S), T) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (T, TVar (v, S)) => if occurs v tye T then raise TUNIFY else Vartab.update_new (v, (S, T)) tye | (Type (a, Ts), Type (b, Us)) => if a <> b then raise TUNIFY else raw_unifys (Ts, Us) tye | (T, U) => if T = U then tye else raise TUNIFY) and raw_unifys (T :: Ts, U :: Us) tye = raw_unifys (Ts, Us) (raw_unify (T, U) tye) | raw_unifys ([], []) tye = tye | raw_unifys _ _ = raise TUNIFY; (*fast unification filter*) fun could_unify (Type (a, Ts), Type (b, Us)) = a = b andalso could_unifys (Ts, Us) | could_unify (TFree (a, _), TFree (b, _)) = a = b | could_unify (TVar _, _) = true | could_unify (_, TVar _) = true | could_unify _ = false and could_unifys (T :: Ts, U :: Us) = could_unify (T, U) andalso could_unifys (Ts, Us) | could_unifys ([], []) = true | could_unifys _ = false; (*equality with respect to a type environment*) fun unified tye = let fun unif (T, T') = (case (devar tye T, devar tye T') of (Type (s, Ts), Type (s', Ts')) => s = s' andalso unifs (Ts, Ts') | (U, U') => U = U') and unifs ([], []) = true | unifs (T :: Ts, T' :: Ts') = unif (T', T') andalso unifs (Ts, Ts') | unifs _ = false; in if Vartab.is_empty tye then op = else unif end; (** extend and merge type signatures **) (* classes *) fun add_class context (c, cs) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val cs' = map (cert_class tsig) cs handle TYPE (msg, _, _) => error msg; val _ = Binding.check c; val (c', space') = space |> Name_Space.declare context true c; val classes' = classes |> Sorts.add_class context (c', cs'); in ((space', classes'), default, types) end); fun hide_class fully c = map_tsig (fn ((space, classes), default, types) => ((Name_Space.hide fully c space, classes), default, types)); (* arities *) fun add_arity context (t, Ss, S) tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val _ = (case lookup_type tsig t of SOME (LogicalType n) => if length Ss <> n then error (bad_nargs t) else () | SOME _ => error ("Logical type constructor expected: " ^ quote t) | NONE => error (undecl_type t)); val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S) handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_arities context ((t, map (fn c' => (c', Ss')) S')); in ((space, classes'), default, types) end); (* classrel *) fun add_classrel context rel tsig = tsig |> map_tsig (fn ((space, classes), default, types) => let val rel' = apply2 (cert_class tsig) rel handle TYPE (msg, _, _) => error msg; val classes' = classes |> Sorts.add_classrel context rel'; in ((space, classes'), default, types) end); (* default sort *) fun set_defsort S tsig = tsig |> map_tsig (fn (classes, _, types) => (classes, cert_sort tsig S handle TYPE (msg, _, _) => error msg, types)); (* types *) local fun new_decl context (c, decl) types = (Binding.check c; #2 (Name_Space.define context true (c, decl) types)); fun map_types f = map_tsig (fn (classes, default, types) => let val types' = f types; val _ = not (Name_Space.defined types' "dummy") orelse Name_Space.intern (Name_Space.space_of_table types') "dummy" = "dummy" orelse error "Illegal declaration of dummy type"; in (classes, default, types') end); fun syntactic tsig (Type (c, Ts)) = (case lookup_type tsig c of SOME Nonterminal => true | _ => false) orelse exists (syntactic tsig) Ts | syntactic _ _ = false; in fun add_type context (c, n) = if n < 0 then error ("Bad type constructor declaration " ^ Binding.print c) else map_types (new_decl context (c, LogicalType n)); fun add_abbrev context (a, vs, rhs) tsig = tsig |> map_types (fn types => let fun err msg = cat_error msg ("The error(s) above occurred in type abbreviation " ^ Binding.print a); val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs)) handle TYPE (msg, _, _) => err msg; val _ = (case duplicates (op =) vs of [] => [] | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups)); val _ = (case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of [] => [] | extras => err ("Extra variables on rhs: " ^ commas_quote extras)); in types |> new_decl context (a, Abbreviation (vs, rhs', syntactic tsig rhs')) end); fun add_nonterminal context = map_types o new_decl context o rpair Nonterminal; end; fun hide_type fully c = map_tsig (fn (classes, default, types) => (classes, default, Name_Space.hide_table fully c types)); (* merge type signatures *) fun merge_tsig context (tsig1, tsig2) = let val (TSig {classes = (space1, classes1), default = default1, types = types1, log_types = _}) = tsig1; val (TSig {classes = (space2, classes2), default = default2, types = types2, log_types = _}) = tsig2; val space' = Name_Space.merge (space1, space2); val classes' = Sorts.merge_algebra context (classes1, classes2); val default' = Sorts.inter_sort classes' (default1, default2); val types' = Name_Space.merge_tables (types1, types2); in build_tsig ((space', classes'), default', types') end; end; diff --git a/src/Pure/type_infer.ML b/src/Pure/type_infer.ML --- a/src/Pure/type_infer.ML +++ b/src/Pure/type_infer.ML @@ -1,123 +1,123 @@ (* Title: Pure/type_infer.ML Author: Stefan Berghofer and Markus Wenzel, TU Muenchen Basic representation of type-inference problems. *) signature TYPE_INFER = sig val is_param: indexname -> bool val is_paramT: typ -> bool val param_maxidx: term -> int -> int val param_maxidx_of: term list -> int val param: int -> string * sort -> typ val mk_param: int -> sort -> typ val anyT: sort -> typ val paramify_vars: typ -> typ val deref: typ Vartab.table -> typ -> typ val finish: Proof.context -> typ Vartab.table -> typ list * term list -> typ list * term list val object_logic: bool Config.T val fixate: Proof.context -> bool -> term list -> term list end; structure Type_Infer: TYPE_INFER = struct (** type parameters and constraints **) (* type inference parameters -- may get instantiated *) fun is_param (x, _: int) = String.isPrefix "?" x; fun is_paramT (TVar (xi, _)) = is_param xi | is_paramT _ = false; val param_maxidx = (Term.fold_types o Term.fold_atyps) (fn (TVar (xi as (_, i), _)) => if is_param xi then Integer.max i else I | _ => I); fun param_maxidx_of ts = fold param_maxidx ts ~1; fun param i (x, S) = TVar (("?" ^ x, i), S); fun mk_param i S = TVar (("?'a", i), S); (* pre-stage parameters *) fun anyT S = TFree ("'_dummy_", S); val paramify_vars = Same.commit (Term_Subst.map_atypsT_same (fn TVar ((x, i), S) => (param i (x, S)) | _ => raise Same.SAME)); (** results **) (* dereferenced views *) fun deref tye (T as TVar (xi, _)) = (case Vartab.lookup tye xi of NONE => T | SOME U => deref tye U) | deref _ T = T; fun add_parms tye T = (case deref tye T of Type (_, Ts) => fold (add_parms tye) Ts | TVar (xi, _) => if is_param xi then insert (op =) xi else I | _ => I); fun add_names tye T = (case deref tye T of Type (_, Ts) => fold (add_names tye) Ts | TFree (x, _) => Name.declare x | TVar ((x, i), _) => if is_param (x, i) then I else Name.declare x); (* finish -- standardize remaining parameters *) fun finish ctxt tye (Ts, ts) = let val used = (fold o fold_types) (add_names tye) ts (fold (add_names tye) Ts (Variable.names_of ctxt)); val parms = rev ((fold o fold_types) (add_parms tye) ts (fold (add_parms tye) Ts [])); val names = Name.invent used ("?" ^ Name.aT) (length parms); val tab = Vartab.make (parms ~~ names); fun finish_typ T = (case deref tye T of Type (a, Ts) => Type (a, map finish_typ Ts) | U as TFree _ => U | U as TVar (xi, S) => (case Vartab.lookup tab xi of NONE => U | SOME a => TVar ((a, 0), S))); in (map finish_typ Ts, map (Type.strip_constraints o Term.map_types finish_typ) ts) end; (* fixate -- introduce fresh type variables *) val object_logic = Config.declare_bool ("Type_Infer.object_logic", \<^here>) (K true); fun fixate ctxt pattern ts = let val base_sort = Object_Logic.get_base_sort ctxt; val improve_sort = if is_some base_sort andalso not pattern andalso Config.get ctxt object_logic then fn [] => the base_sort | S => S else I; fun subst_param (xi, S) (inst, used) = if is_param xi then let val [a] = Name.invent used Name.aT 1; val used' = Name.declare a used; - in (Term_Subst.TVars.add ((xi, S), TFree (a, improve_sort S)) inst, used') end + in (TVars.add ((xi, S), TFree (a, improve_sort S)) inst, used') end else (inst, used); val used = (fold o fold_types) Term.declare_typ_names ts (Variable.names_of ctxt); - val (inst, _) = fold_rev subst_param (fold Term.add_tvars ts []) (Term_Subst.TVars.empty, used); + val (inst, _) = fold_rev subst_param (fold Term.add_tvars ts []) (TVars.empty, used); in (map o map_types) (Term_Subst.instantiateT inst) ts end; end; diff --git a/src/Pure/variable.ML b/src/Pure/variable.ML --- a/src/Pure/variable.ML +++ b/src/Pure/variable.ML @@ -1,747 +1,743 @@ (* 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 Term_Subst.TVars.table * Proof.context + val importT_inst: term list -> Proof.context -> typ TVars.table * Proof.context val import_inst: bool -> term list -> Proof.context -> - (typ Term_Subst.TVars.table * term Term_Subst.Vars.table) * 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 Term_Subst.TVars.table * thm 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 Term_Subst.TVars.table * cterm Term_Subst.Vars.table) * 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 = - Symtab.build (fixes_of inner |> Name_Space.fold_table (fn (y, _) => - not (is_fixed outer y) ? Symtab.insert_set y)); + 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 = - Symtab.build (fold decl_type_occs ts type_occs_inner |> Symtab.fold (fn (a, xs) => + 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 Symtab.insert_set a)); + 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, Symtab.empty) + (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, Symtab.empty) (maxidx_of outer); +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 = build_rev (fold Term.add_tvars ts); val (tfrees, ctxt') = invent_types (map #2 tvars) ctxt; - val instT = - Term_Subst.TVars.build (fold2 (fn a => fn b => - Term_Subst.TVars.add (a, TFree b)) tvars tfrees); + 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 = map (apsnd (Term_Subst.instantiateT instT)) (build_rev (fold Term.add_vars ts)); val (ys, ctxt'') = variant_fixes (map (ren o #1 o #1) vars) ctxt'; - val inst = - Term_Subst.Vars.build (fold2 (fn (x, T) => fn y => - Term_Subst.Vars.add ((x, T), Free (y, T))) vars ys); + 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, Term_Subst.Vars.empty)) ts, ctxt') end; + 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' = Term_Subst.TVars.map (K (Thm.ctyp_of ctxt')) instT; - val ths' = map (Thm.instantiate (Term_Subst.TVars.dest instT', [])) ths; + val instT' = TVars.map (K (Thm.ctyp_of ctxt')) instT; + val ths' = map (Thm.instantiate (TVars.dest instT', [])) 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' = Term_Subst.TVars.map (K (Thm.ctyp_of ctxt')) instT; - val inst' = Term_Subst.Vars.map (K (Thm.cterm_of ctxt')) inst; - val ths' = map (Thm.instantiate (Term_Subst.TVars.dest instT', Term_Subst.Vars.dest inst')) ths; + 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; 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 = Term_Subst.Vars.build (Thm.fold_terms {hyps = false} Term_Subst.add_vars st); + val all_vars = Vars.build (Thm.fold_terms {hyps = false} Vars.add_vars st); in - Term_Subst.Vars.fold (unbind_term o #1 o #1) all_vars #> - Term_Subst.Vars.fold (declare_constraints o Var o #1) all_vars #> + 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 = - Symtab.build (occs' |> Symtab.fold (fn (a, _) => - if Symtab.defined occs a then I else Symtab.insert_set a)); + 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, Symtab.empty) idx) 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/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; (* type instantiated versions *) val tgt_th_tyinst = Thm.instantiate (ctyp_insts,[]) target_thm; val rule_tyinst = Thm.instantiate (ctyp_insts,[]) 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; (* The instantiated rule *) val rule_inst = rule_tyinst |> Thm.instantiate ([], 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) |> 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); 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' (Term_Subst.TFrees.make_set othertfrees) + |> Thm.varifyT_global' (TFrees.make_set othertfrees) |-> K Drule.zero_var_indexes 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,267 +1,267 @@ (* Title: Tools/misc_legacy.ML Misc legacy stuff -- to be phased out eventually. *) signature MISC_LEGACY = sig val thm_frees: thm -> cterm list val cterm_frees: cterm -> cterm list 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 fun thm_frees th = - (th, (Term_Subst.Frees.empty, [])) |-> Thm.fold_atomic_cterms {hyps = true} Term.is_Free + (th, (Frees.empty, [])) |-> Thm.fold_atomic_cterms {hyps = true} Term.is_Free (fn ct => fn (set, list) => let val v = Term.dest_Free (Thm.term_of ct) in - if not (Term_Subst.Frees.defined set v) - then (Term_Subst.Frees.add (v, ()) set, ct :: list) + if not (Frees.defined set v) + then (Frees.add_set v set, ct :: list) else (set, list) end) |> #2; val cterm_frees = thm_frees o Drule.mk_term; (*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 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 end; end;