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,1258 +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} @{index_ML_type class: string} \\ @{index_ML_type sort: "class list"} \\ @{index_ML_type arity: "string * sort list * sort"} \\ @{index_ML_type typ} \\ @{index_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ @{index_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{index_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @{index_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ @{index_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\ @{index_ML Sign.add_type_abbrev: "Proof.context -> binding * string list * typ -> theory -> theory"} \\ @{index_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\ @{index_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ @{index_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} @{index_ML_type term} \\ @{index_ML_op "aconv": "term * term -> bool"} \\ @{index_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{index_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{index_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @{index_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{index_ML fastype_of: "term -> typ"} \\ @{index_ML lambda: "term -> term -> term"} \\ @{index_ML betapply: "term * term -> term"} \\ @{index_ML incr_boundvars: "int -> term -> term"} \\ @{index_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{index_ML Sign.add_abbrev: "string -> binding * term -> theory -> (term * term) * theory"} \\ @{index_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ @{index_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 @{index_ML Bound}, @{index_ML Free}, @{index_ML Var}, @{index_ML Const}, @{index_ML Abs}, @{index_ML_op "$"}. \<^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} @{index_ML Logic.all: "term -> term -> term"} \\ @{index_ML Logic.mk_implies: "term * term -> term"} \\ \end{mldecls} \begin{mldecls} @{index_ML_type ctyp} \\ @{index_ML_type cterm} \\ @{index_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ @{index_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{index_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{index_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{index_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\ @{index_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{index_ML_type thm} \\ @{index_ML Thm.transfer: "theory -> thm -> thm"} \\ @{index_ML Thm.assume: "cterm -> thm"} \\ @{index_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{index_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @{index_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ @{index_ML Thm.implies_elim: "thm -> thm -> thm"} \\ @{index_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\ @{index_ML Thm.instantiate: "((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm"} \\ @{index_ML Thm.add_axiom: "Proof.context -> binding * term -> theory -> (string * thm) * theory"} \\ @{index_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory"} \\ @{index_ML Thm.add_def: "Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} @{index_ML Theory.add_deps: "Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory"} \\ - @{index_ML Thm_Deps.all_oracles: "thm list -> (string * term option) list"} \\ + @{index_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_op>\$\, \<^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 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 a name, plus the original - proposition, if @{ML Proofterm.proofs} was at least @{ML 1} during the - oracle inference. See also the command @{command_ref "thm_oracles"}. + 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} @{index_ML Conjunction.intr: "thm -> thm -> thm"} \\ @{index_ML Conjunction.elim: "thm -> thm * thm"} \\ @{index_ML Drule.mk_term: "cterm -> thm"} \\ @{index_ML Drule.dest_term: "thm -> cterm"} \\ @{index_ML Logic.mk_type: "typ -> term"} \\ @{index_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} @{index_ML Thm.extra_shyps: "thm -> sort list"} \\ @{index_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} @{index_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} @{index_ML_op "RSN": "thm * (int * thm) -> thm"} \\ @{index_ML_op "RS": "thm * thm -> thm"} \\ @{index_ML_op "RLN": "thm list * (int * thm list) -> thm list"} \\ @{index_ML_op "RL": "thm list * thm list -> thm list"} \\ @{index_ML_op "MRS": "thm list * thm -> thm"} \\ @{index_ML_op "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} @{index_ML_type proof} \\ @{index_ML_type proof_body} \\ @{index_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{index_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{index_ML Proofterm.expand_proof: "theory -> (Proofterm.thm_header -> string option) -> proof -> proof"} \\ @{index_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{index_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{index_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 @{index_ML Abst}, @{index_ML AbsP}, @{index_ML_op "%"}, @{index_ML_op "%%"}, @{index_ML PBound}, @{index_ML MinProof}, @{index_ML Hyp}, @{index_ML PAxm}, @{index_ML Oracle}, @{index_ML PThm} as explained above. %FIXME OfClass (!?) \<^descr> Type \<^ML_type>\proof_body\ represents the nested proof information of a named theorem, consisting of a digest of oracles and named theorem over some proof term. The digest only covers the directly visible part of the proof: in order to get the full information, the implicit graph of nested theorems needs to be traversed (e.g.\ using \<^ML>\Proofterm.fold_body_thms\). \<^descr> \<^ML>\Thm.proof_of\~\thm\ and \<^ML>\Thm.proof_body_of\~\thm\ produce the proof term or proof body (with digest of oracles and theorems) from a given theorem. Note that this involves a full join of internal futures that fulfill pending proof promises, and thus disrupts the natural bottom-up construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. \<^descr> \<^ML>\Proofterm.proofs\ specifies the detail of proof recording within \<^ML_type>\thm\ values produced by the inference kernel: \<^ML>\0\ records only the names of oracles, \<^ML>\1\ records oracle names and propositions, \<^ML>\2\ additionally records full proof terms. Officially named theorems that contribute to a result are recorded in any case. \<^descr> \<^ML>\Proofterm.reconstruct_proof\~\thy prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. Reconstruction may fail if \prf\ is not a proof of \prop\, or if it does not contain sufficient information for reconstruction. Failure may only happen for proofs that are constructed manually, but not for those produced automatically by the inference kernel. \<^descr> \<^ML>\Proofterm.expand_proof\~\thy expand prf\ reconstructs and expands the proofs of nested theorems according to the given \expand\ function: a result of @{ML \SOME ""\} means full expansion, @{ML \SOME\}~\name\ means to keep the theorem node but replace its internal name, @{ML NONE} means no change. \<^descr> \<^ML>\Proof_Checker.thm_of_proof\~\thy prf\ turns the given (full) proof into a theorem, by replaying it using only primitive rules of the inference kernel. \<^descr> \<^ML>\Proof_Syntax.read_proof\~\thy b\<^sub>1 b\<^sub>2 s\ reads in a proof term. The Boolean flags indicate the use of sort and type information. Usually, typing information is left implicit and is inferred during proof reconstruction. %FIXME eliminate flags!? \<^descr> \<^ML>\Proof_Syntax.pretty_proof\~\ctxt prf\ pretty-prints the given proof term. \ text %mlex \ \<^item> \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\ provides basic examples involving proof terms. \<^item> \<^file>\~~/src/HOL/Proofs/ex/XML_Data.thy\ demonstrates export and import of proof terms via XML/ML data representation. \ end diff --git a/src/HOL/ex/Iff_Oracle.thy b/src/HOL/ex/Iff_Oracle.thy --- a/src/HOL/ex/Iff_Oracle.thy +++ b/src/HOL/ex/Iff_Oracle.thy @@ -1,78 +1,78 @@ (* Title: HOL/ex/Iff_Oracle.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius *) section \Example of Declaring an Oracle\ theory Iff_Oracle imports Main begin subsection \Oracle declaration\ text \ This oracle makes tautologies of the form \<^prop>\P \ P \ P \ P\. The length is specified by an integer, which is checked to be even and positive. \ oracle iff_oracle = \ let fun mk_iff 1 = Var (("P", 0), \<^typ>\bool\) | mk_iff n = HOLogic.mk_eq (Var (("P", 0), \<^typ>\bool\), mk_iff (n - 1)); in fn (thy, n) => if n > 0 andalso n mod 2 = 0 then Thm.global_cterm_of thy (HOLogic.mk_Trueprop (mk_iff n)) else raise Fail ("iff_oracle: " ^ string_of_int n) end \ subsection \Oracle as low-level rule\ ML \iff_oracle (\<^theory>, 2)\ ML \iff_oracle (\<^theory>, 10)\ ML \ - \<^assert> (map #1 (Thm_Deps.all_oracles [iff_oracle (\<^theory>, 10)]) = [\<^oracle_name>\iff_oracle\]); + \<^assert> (map (#1 o #1) (Thm_Deps.all_oracles [iff_oracle (\<^theory>, 10)]) = [\<^oracle_name>\iff_oracle\]); \ text \These oracle calls had better fail.\ ML \ (iff_oracle (\<^theory>, 5); error "Bad oracle") handle Fail _ => writeln "Oracle failed, as expected" \ ML \ (iff_oracle (\<^theory>, 1); error "Bad oracle") handle Fail _ => writeln "Oracle failed, as expected" \ subsection \Oracle as proof method\ method_setup iff = \Scan.lift Parse.nat >> (fn n => fn ctxt => SIMPLE_METHOD (HEADGOAL (resolve_tac ctxt [iff_oracle (Proof_Context.theory_of ctxt, n)]) handle Fail _ => no_tac))\ lemma "A \ A" by (iff 2) lemma "A \ A \ A \ A \ A \ A \ A \ A \ A \ A" by (iff 10) lemma "A \ A \ A \ A \ A" apply (iff 5)? oops lemma A apply (iff 1)? oops end diff --git a/src/Pure/General/position.ML b/src/Pure/General/position.ML --- a/src/Pure/General/position.ML +++ b/src/Pure/General/position.ML @@ -1,281 +1,298 @@ (* Title: Pure/General/position.ML Author: Markus Wenzel, TU Muenchen Source positions: counting Isabelle symbols, starting from 1. *) signature POSITION = sig eqtype T + val ord: T ord val make: Thread_Position.T -> T val dest: T -> Thread_Position.T val line_of: T -> int option val offset_of: T -> int option val end_offset_of: T -> int option val file_of: T -> string option val advance: Symbol.symbol -> T -> T val advance_offsets: int -> T -> T val distance_of: T * T -> int option val none: T val start: T val file_name: string -> Properties.T val file_only: string -> T val file: string -> T val line_file_only: int -> string -> T val line_file: int -> string -> T val line: int -> T val get_props: T -> Properties.T val id: string -> T val id_only: string -> T val get_id: T -> string option val put_id: string -> T -> T val copy_id: T -> T -> T val id_properties_of: T -> Properties.T val parse_id: T -> int option val adjust_offsets: (int -> int option) -> T -> T val of_properties: Properties.T -> T val properties_of: T -> Properties.T val offset_properties_of: T -> Properties.T val def_properties_of: T -> Properties.T val entity_properties_of: bool -> serial -> T -> Properties.T val default_properties: T -> Properties.T -> Properties.T val markup: T -> Markup.T -> Markup.T val is_reported: T -> bool val is_reported_range: T -> bool val reported_text: T -> Markup.T -> string -> string val report_text: T -> Markup.T -> string -> unit val report: T -> Markup.T -> unit type report = T * Markup.T type report_text = report * string val reports_text: report_text list -> unit val reports: report list -> unit val store_reports: report_text list Unsynchronized.ref -> T list -> ('a -> Markup.T list) -> 'a -> unit val append_reports: report_text list Unsynchronized.ref -> report list -> unit + val here_strs: T -> string * string val here: T -> string val here_list: T list -> string type range = T * T val no_range: range val no_range_position: T -> T val range_position: range -> T val range: T * T -> range val range_of_properties: Properties.T -> range val properties_of_range: range -> Properties.T val thread_data: unit -> T val setmp_thread_data: T -> ('a -> 'b) -> 'a -> 'b val default: T -> bool * T end; structure Position: POSITION = struct (* datatype position *) datatype T = Pos of (int * int * int) * Properties.T; +fun ord (pos1 as Pos ((i, j, k), props), pos2 as Pos ((i', j', k'), props')) = + if pointer_eq (pos1, pos2) then EQUAL + else + (case int_ord (i, i') of + EQUAL => + (case int_ord (j, j') of + EQUAL => + (case int_ord (k, k') of + EQUAL => Properties.ord (props, props') + | ord => ord) + | ord => ord) + | ord => ord); + fun norm_props (props: Properties.T) = maps (fn a => the_list (find_first (fn (b, _) => a = b) props)) Markup.position_properties'; fun make {line = i, offset = j, end_offset = k, props} = Pos ((i, j, k), norm_props props); fun dest (Pos ((i, j, k), props)) = {line = i, offset = j, end_offset = k, props = props}; fun valid (i: int) = i > 0; fun if_valid i i' = if valid i then i' else i; (* fields *) fun line_of (Pos ((i, _, _), _)) = if valid i then SOME i else NONE; fun offset_of (Pos ((_, j, _), _)) = if valid j then SOME j else NONE; fun end_offset_of (Pos ((_, _, k), _)) = if valid k then SOME k else NONE; fun file_of (Pos (_, props)) = Properties.get props Markup.fileN; (* advance *) fun advance_count "\n" (i: int, j: int, k: int) = (if_valid i (i + 1), if_valid j (j + 1), k) | advance_count s (i, j, k) = if Symbol.not_eof s then (i, if_valid j (j + 1), k) else (i, j, k); fun invalid_count (i, j, _: int) = not (valid i orelse valid j); fun advance sym (pos as (Pos (count, props))) = if invalid_count count then pos else Pos (advance_count sym count, props); fun advance_offsets offset (pos as (Pos (count as (i, j, k), props))) = if offset = 0 orelse invalid_count count then pos else if offset < 0 then raise Fail "Illegal offset" else if valid i then raise Fail "Illegal line position" else Pos ((i, if_valid j (j + offset), if_valid k (k + offset)), props); (* distance of adjacent positions *) fun distance_of (Pos ((_, j, _), _), Pos ((_, j', _), _)) = if valid j andalso valid j' then SOME (j' - j) else NONE; (* make position *) val none = Pos ((0, 0, 0), []); val start = Pos ((1, 1, 0), []); fun file_name "" = [] | file_name name = [(Markup.fileN, name)]; fun file_only name = Pos ((0, 0, 0), file_name name); fun file name = Pos ((1, 1, 0), file_name name); fun line_file_only i name = Pos ((i, 0, 0), file_name name); fun line_file i name = Pos ((i, 1, 0), file_name name); fun line i = line_file i ""; fun get_props (Pos (_, props)) = props; fun id id = Pos ((0, 1, 0), [(Markup.idN, id)]); fun id_only id = Pos ((0, 0, 0), [(Markup.idN, id)]); fun get_id (Pos (_, props)) = Properties.get props Markup.idN; fun put_id id (Pos (count, props)) = Pos (count, norm_props (Properties.put (Markup.idN, id) props)); fun copy_id pos = (case get_id pos of NONE => I | SOME id => put_id id); fun parse_id pos = Option.map Value.parse_int (get_id pos); fun id_properties_of pos = (case get_id pos of SOME id => [(Markup.idN, id)] | NONE => []); fun adjust_offsets adjust (pos as Pos (_, props)) = if is_none (file_of pos) then (case parse_id pos of SOME id => (case adjust id of SOME offset => let val Pos (count, _) = advance_offsets offset pos in Pos (count, Properties.remove Markup.idN props) end | NONE => pos) | NONE => pos) else pos; (* markup properties *) fun get props name = (case Properties.get props name of NONE => 0 | SOME s => Value.parse_int s); fun of_properties props = make {line = get props Markup.lineN, offset = get props Markup.offsetN, end_offset = get props Markup.end_offsetN, props = props}; fun value k i = if valid i then [(k, Value.print_int i)] else []; fun properties_of (Pos ((i, j, k), props)) = value Markup.lineN i @ value Markup.offsetN j @ value Markup.end_offsetN k @ props; fun offset_properties_of (Pos ((_, j, k), _)) = value Markup.offsetN j @ value Markup.end_offsetN k; val def_properties_of = properties_of #> map (fn (x, y) => ("def_" ^ x, y)); fun entity_properties_of def serial pos = if def then (Markup.defN, Value.print_int serial) :: properties_of pos else (Markup.refN, Value.print_int serial) :: def_properties_of pos; fun default_properties default props = if exists (member (op =) Markup.position_properties o #1) props then props else properties_of default @ props; val markup = Markup.properties o properties_of; (* reports *) fun is_reported pos = is_some (offset_of pos) andalso is_some (get_id pos); fun is_reported_range pos = is_reported pos andalso is_some (end_offset_of pos); fun reported_text pos m txt = if is_reported pos then Markup.markup (markup pos m) txt else ""; fun report_text pos markup txt = Output.report [reported_text pos markup txt]; fun report pos markup = report_text pos markup ""; type report = T * Markup.T; type report_text = report * string; val reports_text = map (fn ((pos, m), txt) => if is_reported pos then Markup.markup (markup pos m) txt else "") #> Output.report; val reports = map (rpair "") #> reports_text; fun store_reports _ [] _ _ = () | store_reports (r: report_text list Unsynchronized.ref) ps markup x = let val ms = markup x in Unsynchronized.change r (fold (fn p => fold (fn m => cons ((p, m), "")) ms) ps) end; fun append_reports (r: report_text list Unsynchronized.ref) reports = Unsynchronized.change r (append (map (rpair "") reports)); (* here: user output *) +fun here_strs pos = + (case (line_of pos, file_of pos) of + (SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")") + | (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")") + | (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")") + | _ => if is_reported pos then ("", "\092<^here>") else ("", "")); + fun here pos = let val props = properties_of pos; - val (s1, s2) = - (case (line_of pos, file_of pos) of - (SOME i, NONE) => (" ", "(line " ^ Value.print_int i ^ ")") - | (SOME i, SOME name) => (" ", "(line " ^ Value.print_int i ^ " of " ^ quote name ^ ")") - | (NONE, SOME name) => (" ", "(file " ^ quote name ^ ")") - | _ => if is_reported pos then ("", "\092<^here>") else ("", "")); + val (s1, s2) = here_strs pos; in if s2 = "" then "" else s1 ^ Markup.markup (Markup.properties props Markup.position) s2 end; val here_list = map here #> distinct (op =) #> implode; (* range *) type range = T * T; val no_range = (none, none); fun no_range_position (Pos ((i, j, _), props)) = Pos ((i, j, 0), props); fun range_position (Pos ((i, j, _), props), Pos ((_, j', _), _)) = Pos ((i, j, j'), props); fun range (pos, pos') = (range_position (pos, pos'), no_range_position pos'); fun range_of_properties props = let val pos = of_properties props; val pos' = make {line = get props Markup.end_lineN, offset = get props Markup.end_offsetN, end_offset = 0, props = props}; in (pos, pos') end; fun properties_of_range (pos, pos') = properties_of pos @ value Markup.end_lineN (the_default 0 (line_of pos')); (* thread data *) val thread_data = make o Thread_Position.get; fun setmp_thread_data pos = Thread_Position.setmp (dest pos); fun default pos = if pos = none then (false, thread_data ()) else (true, pos); end; diff --git a/src/Pure/General/pretty.ML b/src/Pure/General/pretty.ML --- a/src/Pure/General/pretty.ML +++ b/src/Pure/General/pretty.ML @@ -1,414 +1,424 @@ (* Title: Pure/General/pretty.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Markus Wenzel, TU Munich Generic pretty printing module. Loosely based on D. C. Oppen, "Pretty Printing", ACM Transactions on Programming Languages and Systems (1980), 465-483. The object to be printed is given as a tree with indentation and line breaking information. A "break" inserts a newline if the text until the next break is too long to fit on the current line. After the newline, text is indented to the level of the enclosing block. Normally, if a block is broken then all enclosing blocks will also be broken. The stored length of a block is used in break_dist (to treat each inner block as a unit for breaking). *) signature PRETTY = sig val default_indent: string -> int -> Output.output val add_mode: string -> (string -> int -> Output.output) -> unit type T val make_block: {markup: Markup.output, consistent: bool, indent: int} -> T list -> T val markup_block: {markup: Markup.T, consistent: bool, indent: int} -> T list -> T val str: string -> T val brk: int -> T val brk_indent: int -> int -> T val fbrk: T val breaks: T list -> T list val fbreaks: T list -> T list val blk: int * T list -> T val block: T list -> T val strs: string list -> T val markup: Markup.T -> T list -> T val mark: Markup.T -> T -> T val mark_str: Markup.T * string -> T val marks_str: Markup.T list * string -> T val item: T list -> T val text_fold: T list -> T val keyword1: string -> T val keyword2: string -> T val text: string -> T list val paragraph: T list -> T val para: string -> T val quote: T -> T val cartouche: T -> T val separate: string -> T list -> T list val commas: T list -> T list val enclose: string -> string -> T list -> T val enum: string -> string -> string -> T list -> T val position: Position.T -> T + val here: Position.T -> T list val list: string -> string -> T list -> T val str_list: string -> string -> string list -> T val big_list: string -> T list -> T val indent: int -> T -> T val unbreakable: T -> T val margin_default: int Unsynchronized.ref val symbolicN: string val output_buffer: int option -> T -> Buffer.T val output: int option -> T -> Output.output val string_of_margin: int -> T -> string val string_of: T -> string val writeln: T -> unit val symbolic_output: T -> Output.output val symbolic_string_of: T -> string val unformatted_string_of: T -> string val markup_chunks: Markup.T -> T list -> T val chunks: T list -> T val chunks2: T list -> T val block_enclose: T * T -> T list -> T val writeln_chunks: T list -> unit val writeln_chunks2: T list -> unit val to_ML: FixedInt.int -> T -> ML_Pretty.pretty val from_ML: ML_Pretty.pretty -> T val to_polyml: T -> PolyML.pretty val from_polyml: PolyML.pretty -> T end; structure Pretty: PRETTY = struct (** print mode operations **) fun default_indent (_: string) = Symbol.spaces; local val default = {indent = default_indent}; val modes = Synchronized.var "Pretty.modes" (Symtab.make [("", default)]); in fun add_mode name indent = Synchronized.change modes (fn tab => (if not (Symtab.defined tab name) then () else warning ("Redefining pretty mode " ^ quote name); Symtab.update (name, {indent = indent}) tab)); fun get_mode () = the_default default (Library.get_first (Symtab.lookup (Synchronized.value modes)) (print_mode_value ())); end; fun mode_indent x y = #indent (get_mode ()) x y; val output_spaces = Output.output o Symbol.spaces; val add_indent = Buffer.add o output_spaces; (** printing items: compound phrases, strings, and breaks **) val force_nat = Integer.max 0; abstype T = Block of Markup.output * bool * int * T list * int (*markup output, consistent, indentation, body, length*) | Break of bool * int * int (*mandatory flag, width if not taken, extra indentation if taken*) | Str of Output.output * int (*text, length*) with fun length (Block (_, _, _, _, len)) = len | length (Break (_, wd, _)) = wd | length (Str (_, len)) = len; fun make_block {markup, consistent, indent} body = let val indent' = force_nat indent; fun body_length prts len = let val (line, rest) = chop_prefix (fn Break (true, _, _) => false | _ => true) prts; val len' = Int.max (fold (Integer.add o length) line 0, len); in (case rest of Break (true, _, ind) :: rest' => body_length (Break (false, indent' + ind, 0) :: rest') len' | [] => len') end; in Block (markup, consistent, indent', body, body_length body 0) end; fun markup_block {markup, consistent, indent} es = make_block {markup = Markup.output markup, consistent = consistent, indent = indent} es; (** derived operations to create formatting expressions **) val str = Output.output_width ##> force_nat #> Str; fun brk wd = Break (false, force_nat wd, 0); fun brk_indent wd ind = Break (false, force_nat wd, ind); val fbrk = Break (true, 1, 0); fun breaks prts = Library.separate (brk 1) prts; fun fbreaks prts = Library.separate fbrk prts; fun blk (indent, es) = markup_block {markup = Markup.empty, consistent = false, indent = indent} es; fun block prts = blk (2, prts); val strs = block o breaks o map str; fun markup m = markup_block {markup = m, consistent = false, indent = 0}; fun mark m prt = if m = Markup.empty then prt else markup m [prt]; fun mark_str (m, s) = mark m (str s); fun marks_str (ms, s) = fold_rev mark ms (str s); val item = markup Markup.item; val text_fold = markup Markup.text_fold; fun keyword1 name = mark_str (Markup.keyword1, name); fun keyword2 name = mark_str (Markup.keyword2, name); val text = breaks o map str o Symbol.explode_words; val paragraph = markup Markup.paragraph; val para = paragraph o text; fun quote prt = blk (1, [str "\"", prt, str "\""]); fun cartouche prt = blk (1, [str Symbol.open_, prt, str Symbol.close]); fun separate sep prts = flat (Library.separate [str sep, brk 1] (map single prts)); val commas = separate ","; fun enclose lpar rpar prts = block (str lpar :: (prts @ [str rpar])); fun enum sep lpar rpar prts = enclose lpar rpar (separate sep prts); val position = enum "," "{" "}" o map (fn (x, y) => str (x ^ "=" ^ y)) o Position.properties_of; +fun here pos = + let + val props = Position.properties_of pos; + val (s1, s2) = Position.here_strs pos; + in + if s2 = "" then [] + else [str s1, mark_str (Markup.properties props Markup.position, s2)] + end; + val list = enum ","; fun str_list lpar rpar strs = list lpar rpar (map str strs); fun big_list name prts = block (fbreaks (str name :: prts)); fun indent 0 prt = prt | indent n prt = blk (0, [str (Symbol.spaces n), prt]); fun unbreakable (Block (m, consistent, indent, es, len)) = Block (m, consistent, indent, map unbreakable es, len) | unbreakable (Break (_, wd, _)) = Str (output_spaces wd, wd) | unbreakable (e as Str _) = e; (** formatting **) (* formatted output *) local type text = {tx: Buffer.T, ind: Buffer.T, pos: int, nl: int}; val empty: text = {tx = Buffer.empty, ind = Buffer.empty, pos = 0, nl = 0}; fun newline {tx, ind = _, pos = _, nl} : text = {tx = Buffer.add (Output.output "\n") tx, ind = Buffer.empty, pos = 0, nl = nl + 1}; fun control s {tx, ind, pos: int, nl} : text = {tx = Buffer.add s tx, ind = ind, pos = pos, nl = nl}; fun string (s, len) {tx, ind, pos: int, nl} : text = {tx = Buffer.add s tx, ind = Buffer.add s ind, pos = pos + len, nl = nl}; fun blanks wd = string (output_spaces wd, wd); fun indentation (buf, len) {tx, ind, pos, nl} : text = let val s = Buffer.content buf in {tx = Buffer.add (mode_indent s len) tx, ind = Buffer.add s ind, pos = pos + len, nl = nl} end; (*Add the lengths of the expressions until the next Break; if no Break then include "after", to account for text following this block.*) fun break_dist (Break _ :: _, _) = 0 | break_dist (e :: es, after) = length e + break_dist (es, after) | break_dist ([], after) = after; val force_break = fn Break (false, wd, ind) => Break (true, wd, ind) | e => e; val force_all = map force_break; (*Search for the next break (at this or higher levels) and force it to occur.*) fun force_next [] = [] | force_next ((e as Break _) :: es) = force_break e :: es | force_next (e :: es) = e :: force_next es; in fun formatted margin input = let val breakgain = margin div 20; (*minimum added space required of a break*) val emergencypos = margin div 2; (*position too far to right*) (*es is list of expressions to print; blockin is the indentation of the current block; after is the width of the following context until next break.*) fun format ([], _, _) text = text | format (e :: es, block as (_, blockin), after) (text as {ind, pos, nl, ...}) = (case e of Block ((bg, en), consistent, indent, bes, blen) => let val pos' = pos + indent; val pos'' = pos' mod emergencypos; val block' = if pos' < emergencypos then (ind |> add_indent indent, pos') else (add_indent pos'' Buffer.empty, pos''); val d = break_dist (es, after) val bes' = if consistent andalso pos + blen > margin - d then force_all bes else bes; val btext: text = text |> control bg |> format (bes', block', d) |> control en; (*if this block was broken then force the next break*) val es' = if nl < #nl btext then force_next es else es; in format (es', block, after) btext end | Break (force, wd, ind) => (*no break if text to next break fits on this line or if breaking would add only breakgain to space*) format (es, block, after) (if not force andalso pos + wd <= Int.max (margin - break_dist (es, after), blockin + breakgain) then text |> blanks wd (*just insert wd blanks*) else text |> newline |> indentation block |> blanks ind) | Str str => format (es, block, after) (string str text)); in #tx (format ([input], (Buffer.empty, 0), 0) empty) end; end; (* special output *) (*symbolic markup -- no formatting*) fun symbolic prt = let fun out (Block ((bg, en), _, _, [], _)) = Buffer.add bg #> Buffer.add en | out (Block ((bg, en), consistent, indent, prts, _)) = Buffer.add bg #> Buffer.markup (Markup.block consistent indent) (fold out prts) #> Buffer.add en | out (Break (false, wd, ind)) = Buffer.markup (Markup.break wd ind) (Buffer.add (output_spaces wd)) | out (Break (true, _, _)) = Buffer.add (Output.output "\n") | out (Str (s, _)) = Buffer.add s; in out prt Buffer.empty end; (*unformatted output*) fun unformatted prt = let fun out (Block ((bg, en), _, _, prts, _)) = Buffer.add bg #> fold out prts #> Buffer.add en | out (Break (_, wd, _)) = Buffer.add (output_spaces wd) | out (Str (s, _)) = Buffer.add s; in out prt Buffer.empty end; (* output interfaces *) val margin_default = Unsynchronized.ref ML_Pretty.default_margin; (*right margin, or page width*) val symbolicN = "pretty_symbolic"; fun output_buffer margin prt = if print_mode_active symbolicN then symbolic prt else formatted (the_default (! margin_default) margin) prt; val output = Buffer.content oo output_buffer; fun string_of_margin margin = Output.escape o output (SOME margin); val string_of = Output.escape o output NONE; val writeln = Output.writeln o string_of; val symbolic_output = Buffer.content o symbolic; val symbolic_string_of = Output.escape o symbolic_output; val unformatted_string_of = Output.escape o Buffer.content o unformatted; (* chunks *) fun markup_chunks m prts = markup m (fbreaks (map (text_fold o single) prts)); val chunks = markup_chunks Markup.empty; fun chunks2 prts = (case try split_last prts of NONE => blk (0, []) | SOME (prefix, last) => blk (0, maps (fn prt => [text_fold [prt, fbrk], fbrk]) prefix @ [text_fold [last]])); fun block_enclose (prt1, prt2) prts = chunks [block (fbreaks (prt1 :: prts)), prt2]; fun string_of_text_fold prt = string_of prt |> Markup.markup Markup.text_fold; fun writeln_chunks prts = Output.writelns (Library.separate "\n" (map string_of_text_fold prts)); fun writeln_chunks2 prts = (case try split_last prts of NONE => () | SOME (prefix, last) => (map (fn prt => Markup.markup Markup.text_fold (string_of prt ^ "\n") ^ "\n") prefix @ [string_of_text_fold last]) |> Output.writelns); (** toplevel pretty printing **) fun to_ML 0 (Block _) = ML_Pretty.str "..." | to_ML depth (Block (m, consistent, indent, prts, _)) = ML_Pretty.Block (m, consistent, FixedInt.fromInt indent, map (to_ML (depth - 1)) prts) | to_ML _ (Break (force, wd, ind)) = ML_Pretty.Break (force, FixedInt.fromInt wd, FixedInt.fromInt ind) | to_ML _ (Str (s, len)) = ML_Pretty.String (s, FixedInt.fromInt len); fun from_ML (ML_Pretty.Block (m, consistent, indent, prts)) = make_block {markup = m, consistent = consistent, indent = FixedInt.toInt indent} (map from_ML prts) | from_ML (ML_Pretty.Break (force, wd, ind)) = Break (force, FixedInt.toInt wd, FixedInt.toInt ind) | from_ML (ML_Pretty.String (s, len)) = Str (s, force_nat (FixedInt.toInt len)); val to_polyml = to_ML ~1 #> ML_Pretty.to_polyml; val from_polyml = ML_Pretty.from_polyml #> from_ML; end; val _ = ML_system_pp (fn d => fn _ => ML_Pretty.to_polyml o to_ML (d + 1) o quote); val _ = ML_system_pp (fn _ => fn _ => to_polyml o position); end; structure ML_Pretty = struct open ML_Pretty; val string_of_polyml = Pretty.string_of o Pretty.from_polyml; end; diff --git a/src/Pure/General/properties.ML b/src/Pure/General/properties.ML --- a/src/Pure/General/properties.ML +++ b/src/Pure/General/properties.ML @@ -1,34 +1,39 @@ (* Title: Pure/General/properties.ML Author: Makarius Property lists. *) signature PROPERTIES = sig type entry = string * string type T = entry list + val entry_ord: entry ord + val ord: T ord val defined: T -> string -> bool val get: T -> string -> string option val put: entry -> T -> T val remove: string -> T -> T val seconds: T -> string -> Time.time end; structure Properties: PROPERTIES = struct type entry = string * string; type T = entry list; +val entry_ord = prod_ord string_ord string_ord; +val ord = list_ord entry_ord; + fun defined (props: T) name = AList.defined (op =) props name; fun get (props: T) name = AList.lookup (op =) props name; fun put entry (props: T) = AList.update (op =) entry props; fun remove name (props: T) = AList.delete (op =) name props; fun seconds props name = (case AList.lookup (op =) props name of NONE => Time.zeroTime | SOME s => Time.fromReal (the_default 0.0 (Real.fromString s))); end; diff --git a/src/Pure/Proof/extraction.ML b/src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML +++ b/src/Pure/Proof/extraction.ML @@ -1,862 +1,862 @@ (* Title: Pure/Proof/extraction.ML Author: Stefan Berghofer, TU Muenchen Extraction of programs from proofs. *) signature EXTRACTION = sig val set_preprocessor : (theory -> Proofterm.proof -> Proofterm.proof) -> theory -> theory val add_realizes_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory val add_realizes_eqns : string list -> theory -> theory val add_typeof_eqns_i : ((term * term) list * (term * term)) list -> theory -> theory val add_typeof_eqns : string list -> theory -> theory val add_realizers_i : (string * (string list * term * Proofterm.proof)) list -> theory -> theory val add_realizers : (thm * (string list * string * string)) list -> theory -> theory val add_expand_thm : bool -> thm -> theory -> theory val add_types : (xstring * ((term -> term option) list * (term -> typ -> term -> typ -> term) option)) list -> theory -> theory val extract : (thm * string list) list -> theory -> theory val nullT : typ val nullt : term val mk_typ : typ -> term val etype_of : theory -> string list -> typ list -> term -> typ val realizes_of: theory -> string list -> term -> term -> term val abs_corr_shyps: theory -> thm -> string list -> term list -> Proofterm.proof -> Proofterm.proof end; structure Extraction : EXTRACTION = struct (**** tools ****) val typ = Simple_Syntax.read_typ; val add_syntax = Sign.root_path #> Sign.add_types_global [(Binding.make ("Type", \<^here>), 0, NoSyn), (Binding.make ("Null", \<^here>), 0, NoSyn)] #> Sign.add_consts [(Binding.make ("typeof", \<^here>), typ "'b \ Type", NoSyn), (Binding.make ("Type", \<^here>), typ "'a itself \ Type", NoSyn), (Binding.make ("Null", \<^here>), typ "Null", NoSyn), (Binding.make ("realizes", \<^here>), typ "'a \ 'b \ 'b", NoSyn)]; val nullT = Type ("Null", []); val nullt = Const ("Null", nullT); fun mk_typ T = Const ("Type", Term.itselfT T --> Type ("Type", [])) $ Logic.mk_type T; fun typeof_proc defaultS vs (Const ("typeof", _) $ u) = SOME (mk_typ (case strip_comb u of (Var ((a, i), _), _) => if member (op =) vs a then TFree ("'" ^ a ^ ":" ^ string_of_int i, defaultS) else nullT | (Free (a, _), _) => if member (op =) vs a then TFree ("'" ^ a, defaultS) else nullT | _ => nullT)) | typeof_proc _ _ _ = NONE; fun rlz_proc (Const ("realizes", Type (_, [Type ("Null", []), _])) $ _ $ t) = SOME t | rlz_proc (Const ("realizes", Type (_, [T, _])) $ r $ t) = (case strip_comb t of (Var (ixn, U), ts) => SOME (list_comb (Var (ixn, T --> U), r :: ts)) | (Free (s, U), ts) => SOME (list_comb (Free (s, T --> U), r :: ts)) | _ => NONE) | rlz_proc _ = NONE; val unpack_ixn = apfst implode o apsnd (fst o read_int o tl) o chop_prefix (fn s => s <> ":") o raw_explode; type rules = {next: int, rs: ((term * term) list * (term * term)) list, net: (int * ((term * term) list * (term * term))) Net.net}; val empty_rules : rules = {next = 0, rs = [], net = Net.empty}; fun add_rule (r as (_, (lhs, _))) ({next, rs, net} : rules) = {next = next - 1, rs = r :: rs, net = Net.insert_term (K false) (Envir.eta_contract lhs, (next, r)) net}; fun merge_rules ({next, rs = rs1, net} : rules) ({rs = rs2, ...} : rules) = fold_rev add_rule (subtract (op =) rs1 rs2) {next = next, rs = rs1, net = net}; fun condrew thy rules procs = let fun rew tm = Pattern.rewrite_term thy [] (condrew' :: procs) tm and condrew' tm = let val cache = Unsynchronized.ref ([] : (term * term) list); fun lookup f x = (case AList.lookup (op =) (!cache) x of NONE => let val y = f x in (cache := (x, y) :: !cache; y) end | SOME y => y); in get_first (fn (_, (prems, (tm1, tm2))) => let fun ren t = the_default t (Term.rename_abs tm1 tm t); val inc = Logic.incr_indexes ([], [], maxidx_of_term tm + 1); val env as (Tenv, tenv) = Pattern.match thy (inc tm1, tm) (Vartab.empty, Vartab.empty); val prems' = map (apply2 (Envir.subst_term env o inc o ren)) prems; val env' = Envir.Envir {maxidx = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u) prems' ~1, tenv = tenv, tyenv = Tenv}; val env'' = fold (Pattern.unify (Context.Theory thy) o apply2 (lookup rew)) prems' env'; in SOME (Envir.norm_term env'' (inc (ren tm2))) end handle Pattern.MATCH => NONE | Pattern.Unif => NONE) (sort (int_ord o apply2 fst) (Net.match_term rules (Envir.eta_contract tm))) end; in rew end; val change_types = Proofterm.change_types o SOME; fun extr_name s vs = Long_Name.append "extr" (space_implode "_" (s :: vs)); fun corr_name s vs = extr_name s vs ^ "_correctness"; fun msg d s = writeln (Symbol.spaces d ^ s); fun vars_of t = map Var (rev (Term.add_vars t [])); fun frees_of t = map Free (rev (Term.add_frees t [])); fun vfs_of t = vars_of t @ frees_of t; val mkabs = fold_rev (fn v => fn t => Abs ("x", fastype_of v, abstract_over (v, t))); val mkabsp = fold_rev (fn t => fn prf => AbsP ("H", SOME t, prf)); fun strip_abs 0 t = t | strip_abs n (Abs (_, _, t)) = strip_abs (n-1) t | strip_abs _ _ = error "strip_abs: not an abstraction"; val prf_subst_TVars = Proofterm.map_proof_types o typ_subst_TVars; fun relevant_vars types prop = List.foldr (fn (Var ((a, _), T), vs) => (case body_type T of Type (s, _) => if member (op =) types s then a :: vs else vs | _ => vs) | (_, vs) => vs) [] (vars_of prop); fun tname_of (Type (s, _)) = s | tname_of _ = ""; fun get_var_type t = let val vs = Term.add_vars t []; val fs = Term.add_frees t []; in fn Var (ixn, _) => (case AList.lookup (op =) vs ixn of NONE => error "get_var_type: no such variable in term" | SOME T => Var (ixn, T)) | Free (s, _) => (case AList.lookup (op =) fs s of NONE => error "get_var_type: no such variable in term" | SOME T => Free (s, T)) | _ => error "get_var_type: not a variable" end; fun read_term ctxt T s = let val ctxt' = ctxt |> Proof_Context.set_defsort [] |> Config.put Type_Infer.object_logic false |> Config.put Type_Infer_Context.const_sorts false; val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; in parse ctxt' s |> Type.constraint T |> Syntax.check_term ctxt' end; fun make_proof_body prf = let val (oracles, thms) = ([prf], ([], [])) |-> Proofterm.fold_proof_atoms false - (fn Oracle (name, prop, _) => apfst (cons (name, SOME prop)) + (fn Oracle (name, prop, _) => apfst (cons ((name, Position.none), SOME prop)) | PThm (header, thm_body) => apsnd (cons (Proofterm.make_thm header thm_body)) | _ => I); val body = PBody {oracles = Ord_List.make Proofterm.oracle_ord oracles, thms = Ord_List.make Proofterm.thm_ord thms, proof = prf}; in Proofterm.thm_body body end; (**** theory data ****) (* theory data *) structure ExtractionData = Theory_Data ( type T = {realizes_eqns : rules, typeof_eqns : rules, types : (string * ((term -> term option) list * (term -> typ -> term -> typ -> term) option)) list, realizers : (string list * (term * proof)) list Symtab.table, defs : thm list, expand : string list, prep : (theory -> proof -> proof) option} val empty = {realizes_eqns = empty_rules, typeof_eqns = empty_rules, types = [], realizers = Symtab.empty, defs = [], expand = [], prep = NONE}; val extend = I; fun merge ({realizes_eqns = realizes_eqns1, typeof_eqns = typeof_eqns1, types = types1, realizers = realizers1, defs = defs1, expand = expand1, prep = prep1}, {realizes_eqns = realizes_eqns2, typeof_eqns = typeof_eqns2, types = types2, realizers = realizers2, defs = defs2, expand = expand2, prep = prep2}) : T = {realizes_eqns = merge_rules realizes_eqns1 realizes_eqns2, typeof_eqns = merge_rules typeof_eqns1 typeof_eqns2, types = AList.merge (op =) (K true) (types1, types2), realizers = Symtab.merge_list (eq_set (op =) o apply2 #1) (realizers1, realizers2), defs = Library.merge Thm.eq_thm (defs1, defs2), expand = Library.merge (op =) (expand1, expand2), prep = if is_some prep1 then prep1 else prep2}; ); fun read_condeq thy = let val ctxt' = Proof_Context.init_global (add_syntax thy) in fn s => let val t = Logic.varify_global (read_term ctxt' propT s) in (map Logic.dest_equals (Logic.strip_imp_prems t), Logic.dest_equals (Logic.strip_imp_concl t)) handle TERM _ => error ("Not a (conditional) meta equality:\n" ^ s) end end; (** preprocessor **) fun set_preprocessor prep thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, ...} = ExtractionData.get thy in ExtractionData.put {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = SOME prep} thy end; (** equations characterizing realizability **) fun gen_add_realizes_eqns prep_eq eqns thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; in ExtractionData.put {realizes_eqns = fold_rev add_rule (map (prep_eq thy) eqns) realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = expand, prep = prep} thy end val add_realizes_eqns_i = gen_add_realizes_eqns (K I); val add_realizes_eqns = gen_add_realizes_eqns read_condeq; (** equations characterizing type of extracted program **) fun gen_add_typeof_eqns prep_eq eqns thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val eqns' = map (prep_eq thy) eqns in ExtractionData.put {realizes_eqns = realizes_eqns, realizers = realizers, typeof_eqns = fold_rev add_rule eqns' typeof_eqns, types = types, defs = defs, expand = expand, prep = prep} thy end val add_typeof_eqns_i = gen_add_typeof_eqns (K I); val add_typeof_eqns = gen_add_typeof_eqns read_condeq; fun thaw (T as TFree (a, S)) = if exists_string (fn s => s = ":") a then TVar (unpack_ixn a, S) else T | thaw (Type (a, Ts)) = Type (a, map thaw Ts) | thaw T = T; fun freeze (TVar ((a, i), S)) = TFree (a ^ ":" ^ string_of_int i, S) | freeze (Type (a, Ts)) = Type (a, map freeze Ts) | freeze T = T; fun freeze_thaw f x = map_types thaw (f (map_types freeze x)); fun etype_of thy vs Ts t = let val {typeof_eqns, ...} = ExtractionData.get thy; fun err () = error ("Unable to determine type of extracted program for\n" ^ Syntax.string_of_term_global thy t) in (case strip_abs_body (freeze_thaw (condrew thy (#net typeof_eqns) [typeof_proc [] vs]) (fold (Term.abs o pair "x") Ts (Const ("typeof", fastype_of1 (Ts, t) --> Type ("Type", [])) $ t))) of Const ("Type", _) $ u => (Logic.dest_type u handle TERM _ => err ()) | _ => err ()) end; (** realizers for axioms / theorems, together with correctness proofs **) fun gen_add_realizers prep_rlz rs thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy in ExtractionData.put {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = fold (Symtab.cons_list o prep_rlz thy) rs realizers, defs = defs, expand = expand, prep = prep} thy end fun prep_realizer thy = let val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy; val procs = maps (fst o snd) types; val rtypes = map fst types; val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); val thy' = add_syntax thy; val ctxt' = Proof_Context.init_global thy'; val rd = Proof_Syntax.read_proof thy' true false; in fn (thm, (vs, s1, s2)) => let val name = Thm.derivation_name thm; val _ = name <> "" orelse error "add_realizers: unnamed theorem"; val prop = Thm.unconstrainT thm |> Thm.prop_of |> Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) []; val vars = vars_of prop; val vars' = filter_out (fn v => member (op =) rtypes (tname_of (body_type (fastype_of v)))) vars; val shyps = maps (fn Var ((x, i), _) => if member (op =) vs x then Logic.mk_of_sort (TVar (("'" ^ x, i), []), Sign.defaultS thy') else []) vars; val T = etype_of thy' vs [] prop; val (T', thw) = Type.legacy_freeze_thaw_type (if T = nullT then nullT else map fastype_of vars' ---> T); val t = map_types thw (read_term ctxt' T' s1); val r' = freeze_thaw (condrew thy' eqns (procs @ [typeof_proc [] vs, rlz_proc])) (Const ("realizes", T --> propT --> propT) $ (if T = nullT then t else list_comb (t, vars')) $ prop); val r = Logic.list_implies (shyps, fold_rev Logic.all (map (get_var_type r') vars) r'); val prf = Proofterm.reconstruct_proof thy' r (rd s2); in (name, (vs, (t, prf))) end end; val add_realizers_i = gen_add_realizers (fn _ => fn (name, (vs, t, prf)) => (name, (vs, (t, prf)))); val add_realizers = gen_add_realizers prep_realizer; fun realizes_of thy vs t prop = let val thy' = add_syntax thy; val {realizes_eqns, typeof_eqns, defs, types, ...} = ExtractionData.get thy'; val procs = maps (rev o fst o snd) types; val eqns = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); val prop' = Pattern.rewrite_term thy' (map (Logic.dest_equals o Thm.prop_of) defs) [] prop; in freeze_thaw (condrew thy' eqns (procs @ [typeof_proc [] vs, rlz_proc])) (Const ("realizes", fastype_of t --> propT --> propT) $ t $ prop') end; fun abs_corr_shyps thy thm vs xs prf = let val S = Sign.defaultS thy; val (ucontext, prop') = Logic.unconstrainT (Thm.shyps_of thm) (Thm.prop_of thm); val atyps = fold_types (fold_atyps (insert (op =))) (Thm.prop_of thm) []; val Ts = map_filter (fn ((v, i), _) => if member (op =) vs v then SOME (TVar (("'" ^ v, i), [])) else NONE) (rev (Term.add_vars prop' [])); val cs = maps (fn T => map (pair T) S) Ts; val constraints' = map Logic.mk_of_class cs; fun typ_map T = Type.strip_sorts (map_atyps (fn U => if member (op =) atyps U then (#atyp_map ucontext) U else U) T); fun mk_hyp (T, c) = Hyp (Logic.mk_of_class (typ_map T, c)); val xs' = map (map_types typ_map) xs in prf |> Same.commit (Proofterm.map_proof_same (map_types typ_map) typ_map mk_hyp) |> fold_rev Proofterm.implies_intr_proof' (map snd (#constraints ucontext)) |> fold_rev Proofterm.forall_intr_proof' xs' |> fold_rev Proofterm.implies_intr_proof' constraints' end; (** expanding theorems / definitions **) fun add_expand_thm is_def thm thy = let val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val name = Thm.derivation_name thm; val _ = name <> "" orelse error "add_expand_thm: unnamed theorem"; in thy |> ExtractionData.put (if is_def then {realizes_eqns = realizes_eqns, typeof_eqns = add_rule ([], Logic.dest_equals (map_types Type.strip_sorts (Thm.prop_of (Drule.abs_def thm)))) typeof_eqns, types = types, realizers = realizers, defs = insert Thm.eq_thm_prop (Thm.trim_context thm) defs, expand = expand, prep = prep} else {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = types, realizers = realizers, defs = defs, expand = insert (op =) name expand, prep = prep}) end; fun extraction_expand is_def = Thm.declaration_attribute (fn th => Context.mapping (add_expand_thm is_def th) I); (** types with computational content **) fun add_types tys thy = ExtractionData.map (fn {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} => {realizes_eqns = realizes_eqns, typeof_eqns = typeof_eqns, types = fold (AList.update (op =) o apfst (Sign.intern_type thy)) tys types, realizers = realizers, defs = defs, expand = expand, prep = prep}) thy; (** Pure setup **) val _ = Theory.setup (add_types [("prop", ([], NONE))] #> add_typeof_eqns ["(typeof (PROP P)) \ (Type (TYPE(Null))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE('Q))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE('Q)))", "(typeof (PROP Q)) \ (Type (TYPE(Null))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE(Null)))", "(typeof (PROP P)) \ (Type (TYPE('P))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE('Q))) \ \ \ (typeof (PROP P \ PROP Q)) \ (Type (TYPE('P \ 'Q)))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE(Null))) \ \ \ (typeof (\x. PROP P (x))) \ (Type (TYPE(Null)))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE('P))) \ \ \ (typeof (\x::'a. PROP P (x))) \ (Type (TYPE('a \ 'P)))", "(\x. typeof (f (x))) \ (\x. Type (TYPE('f))) \ \ \ (typeof (f)) \ (Type (TYPE('f)))"] #> add_realizes_eqns ["(typeof (PROP P)) \ (Type (TYPE(Null))) \ \ \ (realizes (r) (PROP P \ PROP Q)) \ \ \ (PROP realizes (Null) (PROP P) \ PROP realizes (r) (PROP Q))", "(typeof (PROP P)) \ (Type (TYPE('P))) \ \ \ (typeof (PROP Q)) \ (Type (TYPE(Null))) \ \ \ (realizes (r) (PROP P \ PROP Q)) \ \ \ (\x::'P. PROP realizes (x) (PROP P) \ PROP realizes (Null) (PROP Q))", "(realizes (r) (PROP P \ PROP Q)) \ \ \ (\x. PROP realizes (x) (PROP P) \ PROP realizes (r (x)) (PROP Q))", "(\x. typeof (PROP P (x))) \ (\x. Type (TYPE(Null))) \ \ \ (realizes (r) (\x. PROP P (x))) \ \ \ (\x. PROP realizes (Null) (PROP P (x)))", "(realizes (r) (\x. PROP P (x))) \ \ \ (\x. PROP realizes (r (x)) (PROP P (x)))"] #> Attrib.setup \<^binding>\extraction_expand\ (Scan.succeed (extraction_expand false)) "specify theorems to be expanded during extraction" #> Attrib.setup \<^binding>\extraction_expand_def\ (Scan.succeed (extraction_expand true)) "specify definitions to be expanded during extraction"); (**** extract program ****) val dummyt = Const ("dummy", dummyT); fun extract thm_vss thy = let val thy' = add_syntax thy; val {realizes_eqns, typeof_eqns, types, realizers, defs, expand, prep} = ExtractionData.get thy; val procs = maps (rev o fst o snd) types; val rtypes = map fst types; val typroc = typeof_proc []; fun expand_name ({name, ...}: Proofterm.thm_header) = if name = "" orelse member (op =) expand name then SOME "" else NONE; val prep = the_default (K I) prep thy' o Proof_Rewrite_Rules.elim_defs thy' false (map (Thm.transfer thy) defs) o Proofterm.expand_proof thy' expand_name; val rrews = Net.merge (K false) (#net realizes_eqns, #net typeof_eqns); fun find_inst prop Ts ts vs = let val rvs = relevant_vars rtypes prop; val vars = vars_of prop; val n = Int.min (length vars, length ts); fun add_args (Var ((a, i), _), t) (vs', tye) = if member (op =) rvs a then let val T = etype_of thy' vs Ts t in if T = nullT then (vs', tye) else (a :: vs', (("'" ^ a, i), T) :: tye) end else (vs', tye) in fold_rev add_args (take n vars ~~ take n ts) ([], []) end; fun mk_shyps tye = maps (fn (ixn, _) => Logic.mk_of_sort (TVar (ixn, []), Sign.defaultS thy)) tye; fun mk_sprfs cs tye = maps (fn (_, T) => Proof_Rewrite_Rules.expand_of_sort_proof thy' (map SOME cs) (T, Sign.defaultS thy)) tye; fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE); fun app_rlz_rews Ts vs t = strip_abs (length Ts) (freeze_thaw (condrew thy' rrews (procs @ [typroc vs, rlz_proc])) (fold (Term.abs o pair "x") Ts t)); fun realizes_null vs prop = app_rlz_rews [] vs (Const ("realizes", nullT --> propT --> propT) $ nullt $ prop); fun corr d vs ts Ts hs cs _ (PBound i) _ defs = (PBound i, defs) | corr d vs ts Ts hs cs t (Abst (s, SOME T, prf)) (Abst (_, _, prf')) defs = let val (corr_prf, defs') = corr d vs [] (T :: Ts) (dummyt :: hs) cs (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE) prf (Proofterm.incr_pboundvars 1 0 prf') defs in (Abst (s, SOME T, corr_prf), defs') end | corr d vs ts Ts hs cs t (AbsP (s, SOME prop, prf)) (AbsP (_, _, prf')) defs = let val T = etype_of thy' vs Ts prop; val u = if T = nullT then (case t of SOME u => SOME (incr_boundvars 1 u) | NONE => NONE) else (case t of SOME (Abs (_, _, u)) => SOME u | _ => NONE); val (corr_prf, defs') = corr d vs [] (T :: Ts) (prop :: hs) (prop :: cs) u (Proofterm.incr_pboundvars 0 1 prf) (Proofterm.incr_pboundvars 0 1 prf') defs; val rlz = Const ("realizes", T --> propT --> propT) in ( if T = nullT then AbsP ("R", SOME (app_rlz_rews Ts vs (rlz $ nullt $ prop)), Proofterm.prf_subst_bounds [nullt] corr_prf) else Abst (s, SOME T, AbsP ("R", SOME (app_rlz_rews (T :: Ts) vs (rlz $ Bound 0 $ incr_boundvars 1 prop)), corr_prf)), defs') end | corr d vs ts Ts hs cs t' (prf % SOME t) (prf' % _) defs = let val (Us, T) = strip_type (fastype_of1 (Ts, t)); val (corr_prf, defs') = corr d vs (t :: ts) Ts hs cs (if member (op =) rtypes (tname_of T) then t' else (case t' of SOME (u $ _) => SOME u | _ => NONE)) prf prf' defs; val u = if not (member (op =) rtypes (tname_of T)) then t else let val eT = etype_of thy' vs Ts t; val (r, Us') = if eT = nullT then (nullt, Us) else (Bound (length Us), eT :: Us); val u = list_comb (incr_boundvars (length Us') t, map Bound (length Us - 1 downto 0)); val u' = (case AList.lookup (op =) types (tname_of T) of SOME ((_, SOME f)) => f r eT u T | _ => Const ("realizes", eT --> T --> T) $ r $ u) in app_rlz_rews Ts vs (fold_rev (Term.abs o pair "x") Us' u') end in (corr_prf % SOME u, defs') end | corr d vs ts Ts hs cs t (prf1 %% prf2) (prf1' %% prf2') defs = let val prop = Proofterm.prop_of' hs prf2'; val T = etype_of thy' vs Ts prop; val (f, u, defs1) = if T = nullT then (t, NONE, defs) else (case t of SOME (f $ u) => (SOME f, SOME u, defs) | _ => let val (u, defs1) = extr d vs [] Ts hs prf2' defs in (NONE, SOME u, defs1) end) val ((corr_prf1, corr_prf2), defs2) = defs1 |> corr d vs [] Ts hs cs f prf1 prf1' ||>> corr d vs [] Ts hs cs u prf2 prf2'; in if T = nullT then (corr_prf1 %% corr_prf2, defs2) else (corr_prf1 % u %% corr_prf2, defs2) end | corr d vs ts Ts hs cs _ (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) _ defs = let val {pos, theory_name, name, prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; val sprfs = mk_sprfs cs tye; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye; val T = etype_of thy' vs' [] prop; val defs' = if T = nullT then defs else snd (extr d vs ts Ts hs prf0 defs) in if T = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs) else (case Symtab.lookup realizers name of NONE => (case find vs' (find' name defs') of NONE => let val _ = T = nullT orelse error "corr: internal error"; val _ = msg d ("Building correctness proof for " ^ quote name ^ (if null vs' then "" else " (relevant variables: " ^ commas_quote vs' ^ ")")); val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); val (corr_prf0, defs'') = corr (d + 1) vs' [] [] [] (rev shyps) NONE prf' prf' defs'; val corr_prf = mkabsp shyps corr_prf0; val corr_prop = Proofterm.prop_of corr_prf; val corr_header = Proofterm.thm_header (serial ()) pos theory_name (corr_name name vs') corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf' = Proofterm.proof_combP (Proofterm.proof_combt (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), map PBound (length shyps - 1 downto 0)) |> fold_rev Proofterm.forall_intr_proof' (map (get_var_type corr_prop) (vfs_of prop)) |> mkabsp shyps in (Proofterm.proof_combP (prf_subst_TVars tye' corr_prf', sprfs), (name, (vs', ((nullt, nullt), (corr_prf, corr_prf')))) :: defs'') end | SOME (_, (_, prf')) => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs')) | SOME rs => (case find vs' rs of SOME (_, prf') => (Proofterm.proof_combP (prf_subst_TVars tye' prf', sprfs), defs') | NONE => error ("corr: no realizer for instance of theorem " ^ quote name ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))))) end | corr d vs ts Ts hs cs _ (prf0 as PAxm (s, prop, SOME Ts')) _ defs = let val (vs', tye) = find_inst prop Ts ts vs; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in if etype_of thy' vs' [] prop = nullT andalso realizes_null vs' prop aconv prop then (prf0, defs) else case find vs' (Symtab.lookup_list realizers s) of SOME (_, prf) => (Proofterm.proof_combP (prf_subst_TVars tye' prf, mk_sprfs cs tye), defs) | NONE => error ("corr: no realizer for instance of axiom " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) end | corr d vs ts Ts hs _ _ _ _ defs = error "corr: bad proof" and extr d vs ts Ts hs (PBound i) defs = (Bound i, defs) | extr d vs ts Ts hs (Abst (s, SOME T, prf)) defs = let val (t, defs') = extr d vs [] (T :: Ts) (dummyt :: hs) (Proofterm.incr_pboundvars 1 0 prf) defs in (Abs (s, T, t), defs') end | extr d vs ts Ts hs (AbsP (s, SOME t, prf)) defs = let val T = etype_of thy' vs Ts t; val (t, defs') = extr d vs [] (T :: Ts) (t :: hs) (Proofterm.incr_pboundvars 0 1 prf) defs in (if T = nullT then subst_bound (nullt, t) else Abs (s, T, t), defs') end | extr d vs ts Ts hs (prf % SOME t) defs = let val (u, defs') = extr d vs (t :: ts) Ts hs prf defs in (if member (op =) rtypes (tname_of (body_type (fastype_of1 (Ts, t)))) then u else u $ t, defs') end | extr d vs ts Ts hs (prf1 %% prf2) defs = let val (f, defs') = extr d vs [] Ts hs prf1 defs; val prop = Proofterm.prop_of' hs prf2; val T = etype_of thy' vs Ts prop in if T = nullT then (f, defs') else let val (t, defs'') = extr d vs [] Ts hs prf2 defs' in (f $ t, defs'') end end | extr d vs ts Ts hs (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) defs = let val {pos, theory_name, name = s, prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in case Symtab.lookup realizers s of NONE => (case find vs' (find' s defs) of NONE => let val _ = msg d ("Extracting " ^ quote s ^ (if null vs' then "" else " (relevant variables: " ^ commas_quote vs' ^ ")")); val prf' = prep (Proofterm.reconstruct_proof thy' prop prf); val (t, defs') = extr (d + 1) vs' [] [] [] prf' defs; val (corr_prf, defs'') = corr (d + 1) vs' [] [] [] (rev shyps) (SOME t) prf' prf' defs'; val nt = Envir.beta_norm t; val args = filter_out (fn v => member (op =) rtypes (tname_of (body_type (fastype_of v)))) (vfs_of prop); val args' = filter (fn v => Logic.occs (v, nt)) args; val t' = mkabs args' nt; val T = fastype_of t'; val cname = extr_name s vs'; val c = Const (cname, T); val u = mkabs args (list_comb (c, args')); val eqn = Logic.mk_equals (c, t'); val rlz = Const ("realizes", fastype_of nt --> propT --> propT); val lhs = app_rlz_rews [] vs' (rlz $ nt $ prop); val rhs = app_rlz_rews [] vs' (rlz $ list_comb (c, args') $ prop); val f = app_rlz_rews [] vs' (Abs ("x", T, rlz $ list_comb (Bound 0, args') $ prop)); val corr_prf' = mkabsp shyps (change_types [] Proofterm.equal_elim_axm %> lhs %> rhs %% (change_types [propT] Proofterm.symmetric_axm %> rhs %> lhs %% (change_types [T, propT] Proofterm.combination_axm %> f %> f %> c %> t' %% (change_types [T --> propT] Proofterm.reflexive_axm %> f) %% PAxm (Thm.def_name cname, eqn, SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf); val corr_prop = Proofterm.prop_of corr_prf'; val corr_header = Proofterm.thm_header (serial ()) pos theory_name (corr_name s vs') corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf'' = Proofterm.proof_combP (Proofterm.proof_combt (PThm (corr_header, make_proof_body corr_prf), vfs_of corr_prop), map PBound (length shyps - 1 downto 0)) |> fold_rev Proofterm.forall_intr_proof' (map (get_var_type corr_prop) (vfs_of prop)) |> mkabsp shyps in (subst_TVars tye' u, (s, (vs', ((t', u), (corr_prf', corr_prf'')))) :: defs'') end | SOME ((_, u), _) => (subst_TVars tye' u, defs)) | SOME rs => (case find vs' rs of SOME (t, _) => (subst_TVars tye' t, defs) | NONE => error ("extr: no realizer for instance of theorem " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts)))))) end | extr d vs ts Ts hs (prf0 as PAxm (s, prop, SOME Ts')) defs = let val (vs', tye) = find_inst prop Ts ts vs; val tye' = (map fst (Term.add_tvars prop [] |> rev) ~~ Ts') @ tye in case find vs' (Symtab.lookup_list realizers s) of SOME (t, _) => (subst_TVars tye' t, defs) | NONE => error ("extr: no realizer for instance of axiom " ^ quote s ^ ":\n" ^ Syntax.string_of_term_global thy' (Envir.beta_norm (Proofterm.prop_of (Proofterm.proof_combt (prf0, ts))))) end | extr d vs ts Ts hs _ defs = error "extr: bad proof"; fun prep_thm vs raw_thm = let val thm = Thm.transfer thy raw_thm; val prop = Thm.prop_of thm; val prf = Thm.proof_of thm; val name = Thm.derivation_name thm; val _ = name <> "" orelse error "extraction: unnamed theorem"; val _ = etype_of thy' vs [] prop <> nullT orelse error ("theorem " ^ quote name ^ " has no computational content") in Proofterm.reconstruct_proof thy' prop prf end; val defs = fold (fn (thm, vs) => snd o (extr 0 vs [] [] [] o prep_thm vs) thm) thm_vss []; fun add_def (s, (vs, ((t, u)))) thy = let val ft = Type.legacy_freeze t; val fu = Type.legacy_freeze u; val head = head_of (strip_abs_body fu); val b = Binding.qualified_name (extr_name s vs); in thy |> Sign.add_consts [(b, fastype_of ft, NoSyn)] |> Global_Theory.add_defs false [((Binding.qualified_name (Thm.def_name (extr_name s vs)), Logic.mk_equals (head, ft)), [])] |-> (fn [def_thm] => Spec_Rules.add_global b Spec_Rules.equational [Thm.term_of (Thm.lhs_of def_thm)] [def_thm] #> Code.declare_default_eqns_global [(def_thm, true)]) end; fun add_corr (s, (vs, prf)) thy = let val corr_prop = Proofterm.prop_of prf; in thy |> Global_Theory.store_thm (Binding.qualified_name (corr_name s vs), Thm.varifyT_global (funpow (length (vars_of corr_prop)) (Thm.forall_elim_var 0) (Thm.forall_intr_frees (Proof_Checker.thm_of_proof thy (fst (Proofterm.freeze_thaw_prf prf)))))) |> snd end; fun add_def_and_corr (s, (vs, ((t, u), (prf, _)))) thy = if is_none (Sign.const_type thy (extr_name s vs)) then thy |> not (t = nullt) ? add_def (s, (vs, ((t, u)))) |> add_corr (s, (vs, prf)) else thy; in thy |> Sign.root_path |> fold_rev add_def_and_corr defs |> Sign.restore_naming thy end; val etype_of = etype_of o add_syntax; end; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2310 +1,2312 @@ (* 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 | OfClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of - {oracles: (string * term option) Ord_List.T, + {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} - type oracle = string * term option + 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 -> (string * sort) list -> 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: string list * string list -> int -> term -> proof -> proof val instantiate: ((indexname * sort) * typ) list * ((indexname * typ) * term) list -> 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 | OfClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of - {oracles: (string * term option) Ord_List.T, + {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 * term option; -val oracle_ord = prod_ord fast_string_ord (option_ord Term_Ord.fast_term_ord); +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 OfClass (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 string (option (term consts)))) (list (thm consts)) (proof consts) - (oracles, thms, 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 OfClass (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) => OfClass (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 string (option (term consts)))) (list (thm consts)) (proof consts) x; + 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 (OfClass 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) => OfClass (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 (OfClass (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]) (OfClass (_, c)) = OfClass (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 *) fun del_conflicting_tvars envT T = Term_Subst.instantiateT (map_filter (fn ixnS as (_, S) => (Type.lookup envT ixnS; NONE) handle TYPE _ => SOME (ixnS, Logic.dummy_tfree S)) (Term.add_tvarsT T [])) T; fun del_conflicting_vars env t = Term_Subst.instantiate (map_filter (fn ixnS as (_, S) => (Type.lookup (Envir.type_env env) ixnS; NONE) handle TYPE _ => SOME (ixnS, Logic.dummy_tfree S)) (Term.add_tvars t []), map_filter (fn (ixnT as (_, T)) => (Envir.lookup env ixnT; NONE) handle TYPE _ => SOME (ixnT, Free ("dummy", T))) (Term.add_vars t [])) t; 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 (OfClass (T, c)) = OfClass (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; (* 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 = Term.fold_types (Term.fold_atyps (fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t []; 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 null frees then [] else fold_aterms (fn Free (x, T) => member (op =) frees x ? insert (op =) (x, T) | _ => I) (Term_Subst.generalize (tfrees, []) idx prop) []; in prf |> Same.commit (map_proof_terms_same (Term_Subst.generalize_same (tfrees, []) 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, map (apsnd 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' _ _ (OfClass (T, c)) = OfClass (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_sorts prf = let fun get S2 (T, S1) = if Sorts.sort_le algebra (S1, S2) then SOME T else NONE; val extra = map (`Logic.dummy_tfree) extra_sorts; val replacements = present @ extra @ witnessed; fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_first (get (Type.sort_of_atyp T)) replacements 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 (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 OfClass 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 OfClass _) = ([], 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 (OfClass (T1, c1), OfClass (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 _ _ (OfClass (T, c)) = OfClass (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 | (OfClass (_, c), OfClass (_, 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 %% OfClass _) = 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 OfClass (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 _ (OfClass (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 TVar (("'", idx), Type.sort_of_atyp T) 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_not excl x = ((is_Free x orelse is_Var x) andalso not (member (op =) excl x)) ? insert (op =) x; val visible = fold_aterms (add_not []) term []; val hidden = fold_proof_terms (fold_aterms (add_not 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 = (bodies, Inttab.empty) |-> fold export_body |> 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 OfClass (#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) (fold boxes_of proofs Inttab.empty) [] 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/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2365 +1,2365 @@ (* 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: (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: (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 extra_shyps: thm -> sort list 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 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: string list * string list -> int -> thm -> thm val generalize_cterm: string list * string list -> int -> cterm -> cterm val generalize_ctyp: string list -> 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 strip_shyps: thm -> thm val unconstrainT: thm -> thm val varifyT_global': (string * sort) list -> 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 f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> fold f hyps; fun fold_atomic_ctyps f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_types o fold_atyps) (f o ctyp) th end; fun fold_atomic_cterms f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms o fold_aterms) (fn t as Const (_, T) => f (cterm t T) | t as Free (_, T) => f (cterm t T) | t as Var (_, T) => f (cterm 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 = if Context.subthy_id (Context.theory_id thy, theory_id th) then th else 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; (*dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = Sorts.subtract (fold_terms Sorts.insert_term th []) shyps; (** 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 "oracle", 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; (*** 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 = solve_constraints #> (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, SOME prop), Proofterm.oracle_proof name prop) - | 1 => ((name, SOME prop), MinProof) - | 0 => ((name, NONE), MinProof) + 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 ([], []) _ th = th | generalize (tfrees, frees) idx th = 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 null tfrees then K false else Term.exists_subtype (fn TFree (a, _) => member (op =) tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse member (op =) 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 ([], []) _ ct = ct | generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = 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 [] _ cT = cT | generalize_ctyp tfrees idx (Ctyp {cert, T, maxidx, sorts}) = 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 pretty_typing thy t T = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy T]; fun add_inst (v as (_, T), cu) (cert, sorts) = let val Cterm {t = u, T = U, cert = cert2, sorts = sorts_u, maxidx = maxidx_u, ...} = cu; val cert' = Context.join_certificate (cert, cert2); val sorts' = Sorts.union sorts_u sorts; in if T = U then ((v, (u, maxidx_u)), (cert', sorts')) else let val msg = (case cert' of Context.Certificate thy' => Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing thy' (Var v) T, Pretty.fbrk, pretty_typing thy' u U]) | Context.Certificate_Id _ => "instantiate: type conflict"); in raise TYPE (msg, [T, U], [Var v, u]) end end; fun add_instT (v as (_, S), cU) (cert, sorts) = let val Ctyp {T = U, cert = cert2, sorts = sorts_U, maxidx = maxidx_U, ...} = cU; val cert' = Context.join_certificate (cert, cert2); val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, [cU], [], [], NONE); val sorts' = Sorts.union sorts_U sorts; in if Sign.of_sort thy' (U, S) then ((v, (U, maxidx_U)), (cert', sorts')) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy' S, [U], []) 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 (inst', (instT', (cert', shyps'))) = (cert, shyps) |> fold_map add_inst inst ||> fold_map add_instT instT 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' = fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (map (apsnd #1) instT', map (apsnd #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 (inst', (instT', (cert', sorts'))) = (cert, sorts) |> fold_map add_inst inst ||> fold_map add_instT instT; 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.OfClass (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; (*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 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; (*Internalize sort constraints of type variables*) val unconstrainT = solve_constraints #> (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 = 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.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' []; (*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.empty |> 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 = 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) = (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/thm_deps.ML b/src/Pure/thm_deps.ML --- a/src/Pure/thm_deps.ML +++ b/src/Pure/thm_deps.ML @@ -1,135 +1,135 @@ (* Title: Pure/thm_deps.ML Author: Stefan Berghofer, TU Muenchen Author: Makarius Dependencies of theorems wrt. internal derivation. *) signature THM_DEPS = sig val all_oracles: thm list -> Proofterm.oracle list val has_skip_proof: thm list -> bool val pretty_thm_oracles: Proof.context -> thm list -> Pretty.T val thm_deps: theory -> thm list -> (Proofterm.thm_id * Thm_Name.T) list val pretty_thm_deps: theory -> thm list -> Pretty.T val unused_thms_cmd: theory list * theory list -> (string * thm) list end; structure Thm_Deps: THM_DEPS = struct (* oracles *) fun all_oracles thms = let fun collect (PBody {oracles, thms, ...}) = (if null oracles then I else apfst (cons oracles)) #> (tap Proofterm.join_thms thms |> fold (fn (i, thm_node) => fn (res, seen) => if Inttab.defined seen i then (res, seen) else let val body = Future.join (Proofterm.thm_node_body thm_node) in collect body (res, Inttab.update (i, ()) seen) end)); val bodies = map Thm.proof_body_of thms; in fold collect bodies ([], Inttab.empty) |> #1 |> Proofterm.unions_oracles end; fun has_skip_proof thms = - all_oracles thms |> exists (fn (name, _) => name = \<^oracle_name>\skip_proof\); + all_oracles thms |> exists (fn ((name, _), _) => name = \<^oracle_name>\skip_proof\); fun pretty_thm_oracles ctxt thms = let val thy = Proof_Context.theory_of ctxt; - fun prt_oracle (name, NONE) = [Thm.pretty_oracle ctxt name] - | prt_oracle (name, SOME prop) = - [Thm.pretty_oracle ctxt name, Pretty.str ":", Pretty.brk 1, - Syntax.pretty_term_global thy prop]; + fun prt_ora (name, pos) = Thm.pretty_oracle ctxt name :: Pretty.here pos; + fun prt_oracle (ora, NONE) = prt_ora ora + | prt_oracle (ora, SOME prop) = + prt_ora ora @ [Pretty.str ":", Pretty.brk 1, Syntax.pretty_term_global thy prop]; in Pretty.big_list "oracles:" (map (Pretty.item o prt_oracle) (all_oracles thms)) end; (* thm_deps *) fun thm_deps thy = let val lookup = Global_Theory.lookup_thm_id thy; fun deps (i, thm_node) res = if Inttab.defined res i then res else let val thm_id = Proofterm.thm_id (i, thm_node) in (case lookup thm_id of SOME thm_name => Inttab.update (i, SOME (thm_id, thm_name)) res | NONE => Inttab.update (i, NONE) res |> fold deps (Proofterm.thm_node_thms thm_node)) end; in fn thms => (fold (fold deps o Thm.thm_deps o Thm.transfer thy) thms Inttab.empty, []) |-> Inttab.fold_rev (fn (_, SOME entry) => cons entry | _ => I) end; fun pretty_thm_deps thy thms = let val ctxt = Proof_Context.init_global thy; val items = map #2 (thm_deps thy thms) |> map (fn (name, i) => (Proof_Context.markup_extern_fact ctxt name, i)) |> sort_by (#2 o #1) |> map (fn ((marks, xname), i) => Pretty.item [Pretty.marks_str (marks, Thm_Name.print (xname, i))]); in Pretty.big_list ("dependencies: " ^ string_of_int (length items)) items end; (* unused_thms_cmd *) fun unused_thms_cmd (base_thys, thys) = let fun add_fact transfer space (name, ths) = if exists (fn thy => Global_Theory.defined_fact thy name) base_thys then I else let val {concealed, group, ...} = Name_Space.the_entry space name in fold_rev (fn th => (case Thm.derivation_name th of "" => I | a => cons (a, (transfer th, concealed, group)))) ths end; fun add_facts thy = let val transfer = Global_Theory.transfer_theories thy; val facts = Global_Theory.facts_of thy; in Facts.fold_static (add_fact transfer (Facts.space_of facts)) facts end; val new_thms = fold add_facts thys [] |> sort_distinct (string_ord o apply2 #1); val used = Proofterm.fold_body_thms (fn {name = a, ...} => a <> "" ? Symtab.update (a, ())) (map Proofterm.strip_thm_body (Thm.proof_bodies_of (map (#1 o #2) new_thms))) Symtab.empty; fun is_unused a = not (Symtab.defined used a); (*groups containing at least one used theorem*) val used_groups = fold (fn (a, (_, _, group)) => if is_unused a then I else (case group of NONE => I | SOME grp => Inttab.update (grp, ()))) new_thms Inttab.empty; val (thms', _) = fold (fn (a, (th, concealed, group)) => fn q as (thms, seen_groups) => if not concealed andalso (* FIXME replace by robust treatment of thm groups *) Thm.legacy_get_kind th = Thm.theoremK andalso is_unused a then (case group of NONE => ((a, th) :: thms, seen_groups) | SOME grp => if Inttab.defined used_groups grp orelse Inttab.defined seen_groups grp then q else ((a, th) :: thms, Inttab.update (grp, ()) seen_groups)) else q) new_thms ([], Inttab.empty); in rev thms' end; end;