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,1362 +1,1362 @@ (*:maxLineLen=78:*) theory Logic imports Base begin chapter \Primitive logic \label{ch:logic}\ text \ The logical foundations of Isabelle/Isar are that of the Pure logic, which has been introduced as a Natural Deduction framework in \<^cite>\paulson700\. This is essentially the same logic as ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) \<^cite>\"Barendregt-Geuvers:2001"\, although there are some key differences in the specific treatment of simple types in Isabelle/Pure. Following type-theoretic parlance, the Pure logic consists of three levels of \\\-calculus with corresponding arrows, \\\ for syntactic function space (terms depending on terms), \\\ for universal quantification (proofs depending on terms), and \\\ for implication (proofs depending on proofs). Derivations are relative to a logical theory, which declares type constructors, constants, and axioms. Theory declarations support schematic polymorphism, which is strictly speaking outside the logic.\<^footnote>\This is the deeper logical reason, why the theory context \\\ is separate from the proof context \\\ of the core calculus: type constructors, term constants, and facts (proof constants) may involve arbitrary type schemes, but the type of a locally fixed term parameter is also fixed!\ \ section \Types \label{sec:types}\ text \ The language of types is an uninterpreted order-sorted first-order algebra; types are qualified by ordered type classes. \<^medskip> A \<^emph>\type class\ is an abstract syntactic entity declared in the theory context. The \<^emph>\subclass relation\ \c\<^sub>1 \ c\<^sub>2\ is specified by stating an acyclic generating relation; the transitive closure is maintained internally. The resulting relation is an ordering: reflexive, transitive, and antisymmetric. A \<^emph>\sort\ is a list of type classes written as \s = {c\<^sub>1, \, c\<^sub>m}\, it represents symbolic intersection. Notationally, the curly braces are omitted for singleton intersections, i.e.\ any class \c\ may be read as a sort \{c}\. The ordering on type classes is extended to sorts according to the meaning of intersections: \{c\<^sub>1, \ c\<^sub>m} \ {d\<^sub>1, \, d\<^sub>n}\ iff \\j. \i. c\<^sub>i \ d\<^sub>j\. The empty intersection \{}\ refers to the universal sort, which is the largest element wrt.\ the sort order. Thus \{}\ represents the ``full sort'', not the empty one! The intersection of all (finitely many) classes declared in the current theory is the least element wrt.\ the sort ordering. \<^medskip> A \<^emph>\fixed type variable\ is a pair of a basic name (starting with a \'\ character) and a sort constraint, e.g.\ \('a, s)\ which is usually printed as \\\<^sub>s\. A \<^emph>\schematic type variable\ is a pair of an indexname and a sort constraint, e.g.\ \(('a, 0), s)\ which is usually printed as \?\\<^sub>s\. Note that \<^emph>\all\ syntactic components contribute to the identity of type variables: basic name, index, and sort constraint. The core logic handles type variables with the same name but different sorts as different, although the type-inference layer (which is outside the core) rejects anything like that. A \<^emph>\type constructor\ \\\ is a \k\-ary operator on types declared in the theory. Type constructor application is written postfix as \(\\<^sub>1, \, \\<^sub>k)\\. For \k = 0\ the argument tuple is omitted, e.g.\ \prop\ instead of \()prop\. For \k = 1\ the parentheses are omitted, e.g.\ \\ list\ instead of \(\)list\. Further notation is provided for specific constructors, notably the right-associative infix \\ \ \\ instead of \(\, \)fun\. The logical category \<^emph>\type\ is defined inductively over type variables and type constructors as follows: \\ = \\<^sub>s | ?\\<^sub>s | (\\<^sub>1, \, \\<^sub>k)\\. A \<^emph>\type abbreviation\ is a syntactic definition \(\<^vec>\)\ = \\ of an arbitrary type expression \\\ over variables \\<^vec>\\. Type abbreviations appear as type constructors in the syntax, but are expanded before entering the logical core. A \<^emph>\type arity\ declares the image behavior of a type constructor wrt.\ the algebra of sorts: \\ :: (s\<^sub>1, \, s\<^sub>k)s\ means that \(\\<^sub>1, \, \\<^sub>k)\\ is of sort \s\ if every argument type \\\<^sub>i\ is of sort \s\<^sub>i\. Arity declarations are implicitly completed, i.e.\ \\ :: (\<^vec>s)c\ entails \\ :: (\<^vec>s)c'\ for any \c' \ c\. \<^medskip> The sort algebra is always maintained as \<^emph>\coregular\, which means that type arities are consistent with the subclass relation: for any type constructor \\\, and classes \c\<^sub>1 \ c\<^sub>2\, and arities \\ :: (\<^vec>s\<^sub>1)c\<^sub>1\ and \\ :: (\<^vec>s\<^sub>2)c\<^sub>2\ holds \\<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ component-wise. The key property of a coregular order-sorted algebra is that sort constraints can be solved in a most general fashion: for each type constructor \\\ and sort \s\ there is a most general vector of argument sorts \(s\<^sub>1, \, s\<^sub>k)\ such that a type scheme \(\\<^bsub>s\<^sub>1\<^esub>, \, \\<^bsub>s\<^sub>k\<^esub>)\\ is of sort \s\. Consequently, type unification has most general solutions (modulo equivalence of sorts), so type-inference produces primary types as expected \<^cite>\"nipkow-prehofer"\. \ text %mlref \ \begin{mldecls} @{define_ML_type class = string} \\ @{define_ML_type sort = "class list"} \\ @{define_ML_type arity = "string * sort list * sort"} \\ @{define_ML_type typ} \\ @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\ @{define_ML Sign.add_type_abbrev: "Proof.context -> binding * string list * typ -> theory -> theory"} \\ @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\ @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\class\ represents type classes. \<^descr> Type \<^ML_type>\sort\ represents sorts, i.e.\ finite intersections of classes. The empty list \<^ML>\[]: sort\ refers to the empty class intersection, i.e.\ the ``full sort''. \<^descr> Type \<^ML_type>\arity\ represents type arities. A triple \(\, \<^vec>s, s) : arity\ represents \\ :: (\<^vec>s)s\ as described above. \<^descr> Type \<^ML_type>\typ\ represents types; this is a datatype with constructors \<^ML>\TFree\, \<^ML>\TVar\, \<^ML>\Type\. \<^descr> \<^ML>\Term.map_atyps\~\f \\ applies the mapping \f\ to all atomic types (\<^ML>\TFree\, \<^ML>\TVar\) occurring in \\\. \<^descr> \<^ML>\Term.fold_atyps\~\f \\ iterates the operation \f\ over all occurrences of atomic types (\<^ML>\TFree\, \<^ML>\TVar\) in \\\; the type structure is traversed from left to right. \<^descr> \<^ML>\Sign.subsort\~\thy (s\<^sub>1, s\<^sub>2)\ tests the subsort relation \s\<^sub>1 \ s\<^sub>2\. \<^descr> \<^ML>\Sign.of_sort\~\thy (\, s)\ tests whether type \\\ is of sort \s\. \<^descr> \<^ML>\Sign.add_type\~\ctxt (\, k, mx)\ declares a new type constructors \\\ with \k\ arguments and optional mixfix syntax. \<^descr> \<^ML>\Sign.add_type_abbrev\~\ctxt (\, \<^vec>\, \)\ defines a new type abbreviation \(\<^vec>\)\ = \\. \<^descr> \<^ML>\Sign.primitive_class\~\(c, [c\<^sub>1, \, c\<^sub>n])\ declares a new class \c\, together with class relations \c \ c\<^sub>i\, for \i = 1, \, n\. \<^descr> \<^ML>\Sign.primitive_classrel\~\(c\<^sub>1, c\<^sub>2)\ declares the class relation \c\<^sub>1 \ c\<^sub>2\. \<^descr> \<^ML>\Sign.primitive_arity\~\(\, \<^vec>s, s)\ declares the arity \\ :: (\<^vec>s)s\. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "class"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "sort"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "nonterminal"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "typ"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Type"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Type_fn"} & : & \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 ; (@@{ML_antiquotation Type} | @@{ML_antiquotation Type_fn}) embedded \ \<^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\. \<^descr> \@{Type source}\ refers to embedded source text to produce a type constructor with name (formally checked) and arguments (inlined ML text). The embedded \source\ follows the syntax category @{syntax type_const} defined below. \<^descr> \@{Type_fn source}\ is similar to \@{Type source}\, but produces a partial ML function that matches against a type constructor pattern, following syntax category @{syntax type_const_fn} below. \<^rail>\ @{syntax_def type_const}: @{syntax name} (@{syntax embedded_ml}*) ; @{syntax_def type_const_fn}: @{syntax type_const} @'=>' @{syntax embedded_ml} ; @{syntax_def embedded_ml}: @'_' | @{syntax embedded} | @{syntax control_symbol} @{syntax embedded} \ The text provided as @{syntax embedded_ml} is treated as regular Isabelle/ML source, possibly with nested antiquotations. ML text that only consists of a single antiquotation in compact control-cartouche notation, can be written without an extra level of nesting embedded text (see the last example below). \ text %mlex \ Here are some minimal examples for type constructor antiquotations. \ ML \ \ \type constructor without arguments:\ val natT = \<^Type>\nat\; \ \type constructor with arguments:\ fun mk_funT (A, B) = \<^Type>\fun A B\; \ \type constructor decomposition as partial function:\ val dest_funT = \<^Type_fn>\fun A B => \(A, B)\\; \ \nested type constructors:\ fun mk_relT A = \<^Type>\fun A \<^Type>\fun A \<^Type>\bool\\\; \<^assert> (mk_relT natT = \<^typ>\nat \ nat \ bool\); \ section \Terms \label{sec:terms}\ text \ The language of terms is that of simply-typed \\\-calculus with de-Bruijn indices for bound variables (cf.\ \<^cite>\debruijn72\ or \<^cite>\"paulson-ml2"\), with the types being determined by the corresponding binders. In contrast, free variables and constants have an explicit name and type in each occurrence. \<^medskip> A \<^emph>\bound variable\ is a natural number \b\, which accounts for the number of intermediate binders between the variable occurrence in the body and its binding position. For example, the de-Bruijn term \\\<^bsub>bool\<^esub>. \\<^bsub>bool\<^esub>. 1 \ 0\ would correspond to \\x\<^bsub>bool\<^esub>. \y\<^bsub>bool\<^esub>. x \ y\ in a named representation. Note that a bound variable may be represented by different de-Bruijn indices at different occurrences, depending on the nesting of abstractions. A \<^emph>\loose variable\ is a bound variable that is outside the scope of local binders. The types (and names) for loose variables can be managed as a separate context, that is maintained as a stack of hypothetical binders. The core logic operates on closed terms, without any loose variables. A \<^emph>\fixed variable\ is a pair of a basic name and a type, e.g.\ \(x, \)\ which is usually printed \x\<^sub>\\ here. A \<^emph>\schematic variable\ is a pair of an indexname and a type, e.g.\ \((x, 0), \)\ which is likewise printed as \?x\<^sub>\\. \<^medskip> A \<^emph>\constant\ is a pair of a basic name and a type, e.g.\ \(c, \)\ which is usually printed as \c\<^sub>\\ here. Constants are declared in the context as polymorphic families \c :: \\, meaning that all substitution instances \c\<^sub>\\ for \\ = \\\ are valid. The vector of \<^emph>\type arguments\ of constant \c\<^sub>\\ wrt.\ the declaration \c :: \\ is defined as the codomain of the matcher \\ = {?\\<^sub>1 \ \\<^sub>1, \, ?\\<^sub>n \ \\<^sub>n}\ presented in canonical order \(\\<^sub>1, \, \\<^sub>n)\, corresponding to the left-to-right occurrences of the \\\<^sub>i\ in \\\. Within a given theory context, there is a one-to-one correspondence between any constant \c\<^sub>\\ and the application \c(\\<^sub>1, \, \\<^sub>n)\ of its type arguments. For example, with \plus :: \ \ \ \ \\, the instance \plus\<^bsub>nat \ nat \ nat\<^esub>\ corresponds to \plus(nat)\. Constant declarations \c :: \\ may contain sort constraints for type variables in \\\. These are observed by type-inference as expected, but \<^emph>\ignored\ by the core logic. This means the primitive logic is able to reason with instances of polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. \<^medskip> An \<^emph>\atomic term\ is either a variable or constant. The logical category \<^emph>\term\ is defined inductively over atomic terms, with abstraction and application as follows: \t = b | x\<^sub>\ | ?x\<^sub>\ | c\<^sub>\ | \\<^sub>\. t | t\<^sub>1 t\<^sub>2\. Parsing and printing takes care of converting between an external representation with named bound variables. Subsequently, we shall use the latter notation instead of internal de-Bruijn representation. The inductive relation \t :: \\ assigns a (unique) type to a term according to the structure of atomic terms, abstractions, and applications: \[ \infer{\a\<^sub>\ :: \\}{} \qquad \infer{\(\x\<^sub>\. t) :: \ \ \\}{\t :: \\} \qquad \infer{\t u :: \\}{\t :: \ \ \\ & \u :: \\} \] A \<^emph>\well-typed term\ is a term that can be typed according to these rules. Typing information can be omitted: type-inference is able to reconstruct the most general type of a raw term, while assigning most general types to all of its variables and constants. Type-inference depends on a context of type constraints for fixed variables, and declarations for polymorphic constants. The identity of atomic terms consists both of the name and the type component. This means that different variables \x\<^bsub>\\<^sub>1\<^esub>\ and \x\<^bsub>\\<^sub>2\<^esub>\ may become the same after type instantiation. Type-inference rejects variables of the same name, but different types. In contrast, mixed instances of polymorphic constants occur routinely. \<^medskip> The \<^emph>\hidden polymorphism\ of a term \t :: \\ is the set of type variables occurring in \t\, but not in its type \\\. This means that the term implicitly depends on type arguments that are not accounted in the result type, i.e.\ there are different type instances \t\ :: \\ and \t\' :: \\ with the same type. This slightly pathological situation notoriously demands additional care. \<^medskip> A \<^emph>\term abbreviation\ is a syntactic definition \c\<^sub>\ \ t\ of a closed term \t\ of type \\\, without any hidden polymorphism. A term abbreviation looks like a constant in the syntax, but is expanded before entering the logical core. Abbreviations are usually reverted when printing terms, using \t \ c\<^sub>\\ as rules for higher-order rewriting. \<^medskip> Canonical operations on \\\-terms include \\\\\-conversion: \\\-conversion refers to capture-free renaming of bound variables; \\\-conversion contracts an abstraction applied to an argument term, substituting the argument in the body: \(\x. b)a\ becomes \b[a/x]\; \\\-conversion contracts vacuous application-abstraction: \\x. f x\ becomes \f\, provided that the bound variable does not occur in \f\. Terms are normally treated modulo \\\-conversion, which is implicit in the de-Bruijn representation. Names for bound variables in abstractions are maintained separately as (meaningless) comments, mostly for parsing and printing. Full \\\\\-conversion is commonplace in various standard operations (\secref{sec:obj-rules}) that are based on higher-order unification and matching. \ text %mlref \ \begin{mldecls} @{define_ML_type term} \\ @{define_ML_infix "aconv": "term * term -> bool"} \\ @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML fastype_of: "term -> typ"} \\ @{define_ML lambda: "term -> term -> term"} \\ @{define_ML betapply: "term * term -> term"} \\ @{define_ML incr_boundvars: "int -> term -> term"} \\ @{define_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{define_ML Sign.add_abbrev: "string -> binding * term -> theory -> (term * term) * theory"} \\ @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\term\ represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}. \<^descr> \t\~\<^ML_text>\aconv\~\u\ checks \\\-equivalence of two terms. This is the basic equality relation on type \<^ML_type>\term\; raw datatype equality should only be used for operations related to parsing or printing! \<^descr> \<^ML>\Term.map_types\~\f t\ applies the mapping \f\ to all types occurring in \t\. \<^descr> \<^ML>\Term.fold_types\~\f t\ iterates the operation \f\ over all occurrences of types in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\Term.map_aterms\~\f t\ applies the mapping \f\ to all atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) occurring in \t\. \<^descr> \<^ML>\Term.fold_aterms\~\f t\ iterates the operation \f\ over all occurrences of atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\fastype_of\~\t\ determines the type of a well-typed term. This operation is relatively slow, despite the omission of any sanity checks. \<^descr> \<^ML>\lambda\~\a b\ produces an abstraction \\a. b\, where occurrences of the atomic term \a\ in the body \b\ are replaced by bound variables. \<^descr> \<^ML>\betapply\~\(t, u)\ produces an application \t u\, with topmost \\\-conversion if \t\ is an abstraction. \<^descr> \<^ML>\incr_boundvars\~\j\ increments a term's dangling bound variables by the offset \j\. This is required when moving a subterm into a context where it is enclosed by a different number of abstractions. Bound variables with a matching abstraction are unaffected. \<^descr> \<^ML>\Sign.declare_const\~\ctxt ((c, \), mx)\ declares a new constant \c :: \\ with optional mixfix syntax. \<^descr> \<^ML>\Sign.add_abbrev\~\print_mode (c, t)\ introduces a new term abbreviation \c \ t\. \<^descr> \<^ML>\Sign.const_typargs\~\thy (c, \)\ and \<^ML>\Sign.const_instance\~\thy (c, [\\<^sub>1, \, \\<^sub>n])\ convert between two representations of polymorphic constants: full type instance vs.\ compact type arguments form. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "const_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "term"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "prop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const_"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "Const_fn"} & : & \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 ; (@@{ML_antiquotation Const} | @@{ML_antiquotation Const_} | @@{ML_antiquotation Const_fn}) embedded \ \<^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\. \<^descr> \@{Const source}\ refers to embedded source text to produce a term constructor with name (formally checked), type arguments and term arguments (inlined ML text). The embedded \source\ follows the syntax category @{syntax term_const} defined below, using @{syntax embedded_ml} from antiquotation @{ML_antiquotation Type} (\secref{sec:types}). \<^descr> \@{Const_ source}\ is similar to \@{Const source}\ for patterns instead of closed values. This distinction is required due to redundant type information within term constants: subtrees with duplicate ML pattern variable need to be suppressed (replaced by dummy patterns). The embedded \source\ follows the syntax category @{syntax term_const} defined below. \<^descr> \@{Const_fn source}\ is similar to \@{Const_ source}\, but produces a proper \<^verbatim>\fn\ expression with function body. The embedded \source\ follows the syntax category @{syntax term_const_fn} defined below. \<^rail>\ @{syntax_def term_const}: @{syntax name} (@{syntax embedded_ml}*) (@{syntax for_args})? ; @{syntax_def term_const_fn}: @{syntax term_const} @'=>' @{syntax embedded_ml} ; @{syntax_def for_args}: @'for' (@{syntax embedded_ml}+) \ \ text %mlex \ Here are some minimal examples for term constant antiquotations. Type arguments for constants are analogous to type constructors (\secref{sec:types}). Term arguments are different: a flexible number of arguments is inserted via curried applications (\<^ML>\op $\). \ ML \ \ \constant without type argument:\ fun mk_conj (A, B) = \<^Const>\conj for A B\; val dest_conj = \<^Const_fn>\conj for A B => \(A, B)\\; \ \constant with type argument:\ fun mk_eq T (t, u) = \<^Const>\HOL.eq T for t u\; val dest_eq = \<^Const_fn>\HOL.eq T for t u => \(T, (t, u))\\; \ \constant with variable number of term arguments:\ val c = Const (\<^const_name>\conj\, \<^typ>\bool \ bool \ bool\); val P = \<^term>\P::bool\; val Q = \<^term>\Q::bool\; \<^assert> (\<^Const>\conj\ = c); \<^assert> (\<^Const>\conj for P\ = c $ P); \<^assert> (\<^Const>\conj for P Q\ = c $ P $ Q); \ 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\ in "Haftmann-Wenzel:2006:classes"\. \ text %mlref \ \begin{mldecls} @{define_ML Logic.all: "term -> term -> term"} \\ @{define_ML Logic.mk_implies: "term * term -> term"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type ctyp} \\ @{define_ML_type cterm} \\ @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\ @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type thm} \\ @{define_ML Thm.transfer: "theory -> thm -> thm"} \\ @{define_ML Thm.assume: "cterm -> thm"} \\ @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\ @{define_ML Thm.generalize: "Names.set * Names.set -> int -> thm -> thm"} \\ @{define_ML Thm.instantiate: "ctyp TVars.table * cterm Vars.table -> thm -> thm"} \\ @{define_ML Thm.add_axiom: "Proof.context -> binding * term -> theory -> (string * thm) * theory"} \\ @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory"} \\ @{define_ML Thm.add_def: "Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} @{define_ML Theory.add_deps: "Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory"} \\ @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\ \end{mldecls} \<^descr> \<^ML>\Logic.all\~\a B\ produces a Pure quantification \\a. B\, where occurrences of the atomic term \a\ in the body proposition \B\ are replaced by bound variables. (See also \<^ML>\lambda\ on terms.) \<^descr> \<^ML>\Logic.mk_implies\~\(A, B)\ produces a Pure implication \A \ B\. \<^descr> Types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc.\ in the background theory. The abstract types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are part of the same inference kernel that is mainly responsible for \<^ML_type>\thm\. Thus syntactic operations on \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are located in the \<^ML_structure>\Thm\ module, even though theorems are not yet involved at that stage. \<^descr> \<^ML>\Thm.ctyp_of\~\ctxt \\ and \<^ML>\Thm.cterm_of\~\ctxt t\ explicitly check types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the underlying theory context. Full re-certification is relatively slow and should be avoided in tight reasoning loops. \<^descr> \<^ML>\Thm.apply\, \<^ML>\Thm.lambda\, \<^ML>\Thm.all\, \<^ML>\Drule.mk_implies\ etc.\ compose certified terms (or propositions) incrementally. This is equivalent to \<^ML>\Thm.cterm_of\ after unchecked \<^ML_infix>\$\, \<^ML>\lambda\, \<^ML>\Logic.all\, \<^ML>\Logic.mk_implies\ etc., but there can be a big difference in performance when large existing entities are composed by a few extra constructions on top. There are separate operations to decompose certified terms and theorems to produce certified terms again. \<^descr> Type \<^ML_type>\thm\ represents proven propositions. This is an abstract datatype that guarantees that its values have been constructed by basic principles of the \<^ML_structure>\Thm\ module. Every \<^ML_type>\thm\ value refers its background theory, cf.\ \secref{sec:context-theory}. \<^descr> \<^ML>\Thm.transfer\~\thy thm\ transfers the given theorem to a \<^emph>\larger\ theory, see also \secref{sec:context}. This formal adjustment of the background context has no logical significance, but is occasionally required for formal reasons, e.g.\ when theorems that are imported from more basic theories are used in the current situation. \<^descr> \<^ML>\Thm.assume\, \<^ML>\Thm.forall_intr\, \<^ML>\Thm.forall_elim\, \<^ML>\Thm.implies_intr\, and \<^ML>\Thm.implies_elim\ correspond to the primitive inferences of \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.generalize\~\(\<^vec>\, \<^vec>x)\ corresponds to the \generalize\ rules of \figref{fig:subst-rules}. Here collections of type and term variables are generalized simultaneously, specified by the given sets of basic names. \<^descr> \<^ML>\Thm.instantiate\~\(\<^vec>\\<^sub>s, \<^vec>x\<^sub>\)\ corresponds to the \instantiate\ rules of \figref{fig:subst-rules}. Type variables are substituted before term variables. Note that the types in \\<^vec>x\<^sub>\\ refer to the instantiated versions. \<^descr> \<^ML>\Thm.add_axiom\~\ctxt (name, A)\ declares an arbitrary proposition as axiom, and retrieves it as a theorem from the resulting theory, cf.\ \axiom\ in \figref{fig:prim-rules}. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Thm.add_oracle\~\(binding, oracle)\ produces a named oracle rule, essentially generating arbitrary axioms on the fly, cf.\ \axiom\ in \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.add_def\~\ctxt unchecked overloaded (name, c \<^vec>x \ t)\ states a definitional axiom for an existing constant \c\. Dependencies are recorded via \<^ML>\Theory.add_deps\, unless the \unchecked\ option is set. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Theory.add_deps\~\ctxt name c\<^sub>\ \<^vec>d\<^sub>\\ declares dependencies of a named specification for constant \c\<^sub>\\, relative to existing specifications for constants \\<^vec>d\<^sub>\\. This also works for type constructors. \<^descr> \<^ML>\Thm_Deps.all_oracles\~\thms\ returns all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. The result contains an authentic oracle name; if @{ML Proofterm.proofs} was at least @{ML 1} during the oracle inference, it also contains the position of the oracle invocation and its proposition. See also the command @{command_ref "thm_oracles"}. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "ctyp"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cterm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cprop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thms"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "lemma"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "oracle_name"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation ctyp} typ ; @@{ML_antiquotation cterm} term ; @@{ML_antiquotation cprop} prop ; @@{ML_antiquotation thm} thm ; @@{ML_antiquotation thms} thms ; @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \ @{syntax for_fixes} @'by' method method? ; @@{ML_antiquotation oracle_name} embedded \ \<^descr> \@{ctyp \}\ produces a certified type wrt.\ the current background theory --- as abstract value of type \<^ML_type>\ctyp\. \<^descr> \@{cterm t}\ and \@{cprop \}\ produce a certified term wrt.\ the current background theory --- as abstract value of type \<^ML_type>\cterm\. \<^descr> \@{thm a}\ produces a singleton fact --- as abstract value of type \<^ML_type>\thm\. \<^descr> \@{thms a}\ produces a general fact --- as abstract value of type \<^ML_type>\thm list\. \<^descr> \@{lemma \ by meth}\ produces a fact that is proven on the spot according to the minimal proof, which imitates a terminal Isar proof. The result is an abstract value of type \<^ML_type>\thm\ or \<^ML_type>\thm list\, depending on the number of propositions given here. The internal derivation object lacks a proper theorem name, but it is formally closed, unless the \(open)\ option is specified (this may impact performance of applications with proof terms). Since ML antiquotations are always evaluated at compile-time, there is no run-time overhead even for non-trivial proofs. Nonetheless, the justification is syntactically limited to a single @{command "by"} step. More complex Isar proofs should be done in regular theory source, before compiling the corresponding ML text that uses the result. \<^descr> \@{oracle_name a}\ inlines the internalized oracle name \a\ --- as \<^ML_type>\string\ literal. \ subsection \Auxiliary connectives \label{sec:logic-aux}\ text \ Theory \Pure\ provides a few auxiliary connectives that are defined on top of the primitive ones, see \figref{fig:pure-aux}. These special constants are useful in certain internal encodings, and are normally not directly exposed to the user. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \conjunction :: prop \ prop \ prop\ & (infix \&&&\) \\ \\ A &&& B \ (\C. (A \ B \ C) \ C)\ \\[1ex] \prop :: prop \ prop\ & (prefix \#\, suppressed) \\ \#A \ A\ \\[1ex] \term :: \ \ prop\ & (prefix \TERM\) \\ \term x \ (\A. A \ A)\ \\[1ex] \type :: \ itself\ & (prefix \TYPE\) \\ \(unspecified)\ \\ \end{tabular} \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} \end{center} \end{figure} The introduction \A \ B \ A &&& B\, and eliminations (projections) \A &&& B \ A\ and \A &&& B \ B\ are available as derived rules. Conjunction allows to treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \A \ B \ C &&& D\. In particular, the goal mechanism represents multiple claims as explicit conjunction internally, but this is refined (via backwards introduction) into separate sub-goals before the user commences the proof; the final result is projected into a list of theorems using eliminations (cf.\ \secref{sec:tactical-goals}). The \prop\ marker (\#\) makes arbitrarily complex propositions appear as atomic, without changing the meaning: \\ \ A\ and \\ \ #A\ are interchangeable. See \secref{sec:tactical-goals} for specific operations. The \term\ marker turns any well-typed term into a derivable proposition: \\ TERM t\ holds unconditionally. Although this is logically vacuous, it allows to treat terms and proofs uniformly, similar to a type-theoretic framework. The \TYPE\ constructor is the canonical representative of the unspecified type \\ itself\; it essentially injects the language of types into that of terms. There is specific notation \TYPE(\)\ for \TYPE\<^bsub>\ itself\<^esub>\. Although being devoid of any particular meaning, the term \TYPE(\)\ accounts for the type \\\ within the term language. In particular, \TYPE(\)\ may be used as formal argument in primitive definitions, in order to circumvent hidden polymorphism (cf.\ \secref{sec:terms}). For example, \c TYPE(\) \ A[\]\ defines \c :: \ itself \ prop\ in terms of a proposition \A\ that depends on an additional type argument, which is essentially a predicate on types. \ text %mlref \ \begin{mldecls} @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\ @{define_ML Conjunction.elim: "thm -> thm * thm"} \\ @{define_ML Drule.mk_term: "cterm -> thm"} \\ @{define_ML Drule.dest_term: "thm -> cterm"} \\ @{define_ML Logic.mk_type: "typ -> term"} \\ @{define_ML Logic.dest_type: "term -> typ"} \\ \end{mldecls} \<^descr> \<^ML>\Conjunction.intr\ derives \A &&& B\ from \A\ and \B\. \<^descr> \<^ML>\Conjunction.elim\ derives \A\ and \B\ from \A &&& B\. \<^descr> \<^ML>\Drule.mk_term\ derives \TERM t\. \<^descr> \<^ML>\Drule.dest_term\ recovers term \t\ from \TERM t\. \<^descr> \<^ML>\Logic.mk_type\~\\\ produces the term \TYPE(\)\. \<^descr> \<^ML>\Logic.dest_type\~\TYPE(\)\ recovers the type \\\. \ subsection \Sort hypotheses\ text \ Type variables are decorated with sorts, as explained in \secref{sec:types}. This constrains type instantiation to certain ranges of types: variable \\\<^sub>s\ may only be assigned to types \\\ that belong to sort \s\. Within the logic, sort constraints act like implicit preconditions on the result \\\\<^sub>1 : s\<^sub>1\, \, \\\<^sub>n : s\<^sub>n\, \ \ \\ where the type variables \\\<^sub>1, \, \\<^sub>n\ cover the propositions \\\, \\\, as well as the proof of \\ \ \\. These \<^emph>\sort hypothesis\ of a theorem are passed monotonically through further derivations. They are redundant, as long as the statement of a theorem still contains the type variables that are accounted here. The logical significance of sort hypotheses is limited to the boundary case where type variables disappear from the proposition, e.g.\ \\\\<^sub>s : s\ \ \\. Since such dangling type variables can be renamed arbitrarily without changing the proposition \\\, the inference kernel maintains sort hypotheses in anonymous form \s \ \\. In most practical situations, such extra sort hypotheses may be stripped in a final bookkeeping step, e.g.\ at the end of a proof: they are typically left over from intermediate reasoning with type classes that can be satisfied by some concrete type \\\ of sort \s\ to replace the hypothetical type variable \\\<^sub>s\. \ text %mlref \ \begin{mldecls} - @{define_ML Thm.extra_shyps: "thm -> sort list"} \\ + @{define_ML Thm.extra_shyps: "thm -> Sortset.T"} \\ @{define_ML Thm.strip_shyps: "thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Thm.extra_shyps\~\thm\ determines the extraneous sort hypotheses of the given theorem, i.e.\ the sorts that are not present within type variables of the statement. \<^descr> \<^ML>\Thm.strip_shyps\~\thm\ removes any extraneous sort hypotheses that can be witnessed from the type signature. \ text %mlex \ The following artificial example demonstrates the derivation of \<^prop>\False\ with a pending sort hypothesis involving a logically empty sort. \ class empty = assumes bad: "\(x::'a) y. x \ y" theorem (in empty) false: False using bad by blast -ML_val \\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\empty\])\ +ML_val \\<^assert> (Sortset.dest (Thm.extra_shyps @{thm false}) = [\<^sort>\empty\])\ text \ Thanks to the inference kernel managing sort hypothesis according to their logical significance, this example is merely an instance of \<^emph>\ex falso quodlibet consequitur\ --- not a collapse of the logical framework! \ section \Object-level rules \label{sec:obj-rules}\ text \ The primitive inferences covered so far mostly serve foundational purposes. User-level reasoning usually works via object-level rules that are represented as theorems of Pure. Composition of rules involves \<^emph>\backchaining\, \<^emph>\higher-order unification\ modulo \\\\\-conversion of \\\-terms, and so-called \<^emph>\lifting\ of rules into a context of \\\ and \\\ connectives. Thus the full power of higher-order Natural Deduction in Isabelle/Pure becomes readily available. \ subsection \Hereditary Harrop Formulae\ text \ The idea of object-level rules is to model Natural Deduction inferences in the style of Gentzen \<^cite>\"Gentzen:1935"\, but we allow arbitrary nesting similar to \<^cite>\"Schroeder-Heister:1984"\. The most basic rule format is that of a \<^emph>\Horn Clause\: \[ \infer{\A\}{\A\<^sub>1\ & \\\ & \A\<^sub>n\} \] where \A, A\<^sub>1, \, A\<^sub>n\ are atomic propositions of the framework, usually of the form \Trueprop B\, where \B\ is a (compound) object-level statement. This object-level inference corresponds to an iterated implication in Pure like this: \[ \A\<^sub>1 \ \ A\<^sub>n \ A\ \] As an example consider conjunction introduction: \A \ B \ A \ B\. Any parameters occurring in such rule statements are conceptionally treated as arbitrary: \[ \\x\<^sub>1 \ x\<^sub>m. A\<^sub>1 x\<^sub>1 \ x\<^sub>m \ \ A\<^sub>n x\<^sub>1 \ x\<^sub>m \ A x\<^sub>1 \ x\<^sub>m\ \] Nesting of rules means that the positions of \A\<^sub>i\ may again hold compound rules, not just atomic propositions. Propositions of this format are called \<^emph>\Hereditary Harrop Formulae\ in the literature \<^cite>\"Miller:1991"\. Here we give an inductive characterization as follows: \<^medskip> \begin{tabular}{ll} \\<^bold>x\ & set of variables \\ \\<^bold>A\ & set of atomic propositions \\ \\<^bold>H = \\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \ \<^bold>A\ & set of Hereditary Harrop Formulas \\ \end{tabular} \<^medskip> Thus we essentially impose nesting levels on propositions formed from \\\ and \\\. At each level there is a prefix of parameters and compound premises, concluding an atomic proposition. Typical examples are \\\-introduction \(A \ B) \ A \ B\ or mathematical induction \P 0 \ (\n. P n \ P (Suc n)) \ P n\. Even deeper nesting occurs in well-founded induction \(\x. (\y. y \ x \ P y) \ P x) \ P x\, but this already marks the limit of rule complexity that is usually seen in practice. \<^medskip> Regular user-level inferences in Isabelle/Pure always maintain the following canonical form of results: \<^item> Normalization by \(A \ (\x. B x)) \ (\x. A \ B x)\, which is a theorem of Pure, means that quantifiers are pushed in front of implication at each level of nesting. The normal form is a Hereditary Harrop Formula. \<^item> The outermost prefix of parameters is represented via schematic variables: instead of \\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x\ we have \\<^vec>H ?\<^vec>x \ A ?\<^vec>x\. Note that this representation looses information about the order of parameters, and vacuous quantifiers vanish automatically. \ text %mlref \ \begin{mldecls} @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Simplifier.norm_hhf\~\ctxt thm\ normalizes the given theorem according to the canonical form specified above. This is occasionally helpful to repair some low-level tools that do not handle Hereditary Harrop Formulae properly. \ subsection \Rule composition\ text \ The rule calculus of Isabelle/Pure provides two main inferences: @{inference resolution} (i.e.\ back-chaining of rules) and @{inference assumption} (i.e.\ closing a branch), both modulo higher-order unification. There are also combined variants, notably @{inference elim_resolution} and @{inference dest_resolution}. To understand the all-important @{inference resolution} principle, we first consider raw @{inference_def composition} (modulo higher-order unification with substitution \\\): \[ \infer[(@{inference_def composition})]{\\<^vec>A\ \ C\\} {\\<^vec>A \ B\ & \B' \ C\ & \B\ = B'\\} \] Here the conclusion of the first rule is unified with the premise of the second; the resulting rule instance inherits the premises of the first and conclusion of the second. Note that \C\ can again consist of iterated implications. We can also permute the premises of the second rule back-and-forth in order to compose with \B'\ in any position (subsequently we shall always refer to position 1 w.l.o.g.). In @{inference composition} the internal structure of the common part \B\ and \B'\ is not taken into account. For proper @{inference resolution} we require \B\ to be atomic, and explicitly observe the structure \\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x\ of the premise of the second rule. The idea is to adapt the first rule by ``lifting'' it into this context, by means of iterated application of the following inferences: \[ \infer[(@{inference_def imp_lift})]{\(\<^vec>H \ \<^vec>A) \ (\<^vec>H \ B)\}{\\<^vec>A \ B\} \] \[ \infer[(@{inference_def all_lift})]{\(\\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \ (\\<^vec>x. B (?\<^vec>a \<^vec>x))\}{\\<^vec>A ?\<^vec>a \ B ?\<^vec>a\} \] By combining raw composition with lifting, we get full @{inference resolution} as follows: \[ \infer[(@{inference_def resolution})] {\(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (?\<^vec>a \<^vec>x))\ \ C\\} {\begin{tabular}{l} \\<^vec>A ?\<^vec>a \ B ?\<^vec>a\ \\ \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \(\\<^vec>x. B (?\<^vec>a \<^vec>x))\ = B'\\ \\ \end{tabular}} \] Continued resolution of rules allows to back-chain a problem towards more and sub-problems. Branches are closed either by resolving with a rule of 0 premises, or by producing a ``short-circuit'' within a solved situation (again modulo unification): \[ \infer[(@{inference_def assumption})]{\C\\} {\(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C\ & \A\ = H\<^sub>i\\~~\mbox{(for some~\i\)}} \] %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution} \ text %mlref \ \begin{mldecls} @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\ @{define_ML_infix "RS": "thm * thm -> thm"} \\ @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\ @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\ @{define_ML_infix "MRS": "thm list * thm -> thm"} \\ @{define_ML_infix "OF": "thm * thm list -> thm"} \\ \end{mldecls} \<^descr> \rule\<^sub>1 RSN (i, rule\<^sub>2)\ resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, according to the @{inference resolution} principle explained above. Unless there is precisely one resolvent it raises exception \<^ML>\THM\. This corresponds to the rule attribute @{attribute THEN} in Isar source language. \<^descr> \rule\<^sub>1 RS rule\<^sub>2\ abbreviates \rule\<^sub>1 RSN (1, rule\<^sub>2)\. \<^descr> \rules\<^sub>1 RLN (i, rules\<^sub>2)\ joins lists of rules. For every \rule\<^sub>1\ in \rules\<^sub>1\ and \rule\<^sub>2\ in \rules\<^sub>2\, it resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, accumulating multiple results in one big list. Note that such strict enumerations of higher-order unifications can be inefficient compared to the lazy variant seen in elementary tactics like \<^ML>\resolve_tac\. \<^descr> \rules\<^sub>1 RL rules\<^sub>2\ abbreviates \rules\<^sub>1 RLN (1, rules\<^sub>2)\. \<^descr> \[rule\<^sub>1, \, rule\<^sub>n] MRS rule\ resolves \rule\<^sub>i\ against premise \i\ of \rule\, for \i = n, \, 1\. By working from right to left, newly emerging premises are concatenated in the result, without interfering. \<^descr> \rule OF rules\ is an alternative notation for \rules MRS rule\, which makes rule composition look more like function application. Note that the argument \rules\ need not be atomic. This corresponds to the rule attribute @{attribute OF} in Isar source language. \ section \Proof terms \label{sec:proof-terms}\ text \ The Isabelle/Pure inference kernel can record the proof of each theorem as a proof term that contains all logical inferences in detail. Rule composition by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken down to primitive rules of the logical framework. The proof term can be inspected by a separate proof-checker, for example. According to the well-known \<^emph>\Curry-Howard isomorphism\, a proof can be viewed as a \\\-term. Following this idea, proofs in Isabelle are internally represented by a datatype similar to the one for terms described in \secref{sec:terms}. On top of these syntactic terms, two more layers of \\\-calculus are added, which correspond to \\x :: \. B x\ and \A \ B\ according to the propositions-as-types principle. The resulting 3-level \\\-calculus resembles ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) \<^cite>\"Barendregt-Geuvers:2001"\, if some fine points like schematic polymorphism and type classes are ignored. \<^medskip> \<^emph>\Proof abstractions\ of the form \\<^bold>\x :: \. prf\ or \\<^bold>\p : A. prf\ correspond to introduction of \\\/\\\, and \<^emph>\proof applications\ of the form \p \ t\ or \p \ q\ correspond to elimination of \\\/\\\. Actual types \\\, propositions \A\, and terms \t\ might be suppressed and reconstructed from the overall proof term. \<^medskip> Various atomic proofs indicate special situations within the proof construction as follows. A \<^emph>\bound proof variable\ is a natural number \b\ that acts as de-Bruijn index for proof term abstractions. A \<^emph>\minimal proof\ ``\?\'' is a dummy proof term. This indicates some unrecorded part of the proof. \Hyp A\ refers to some pending hypothesis by giving its proposition. This indicates an open context of implicit hypotheses, similar to loose bound variables or free variables within a term (\secref{sec:terms}). An \<^emph>\axiom\ or \<^emph>\oracle\ \a : A[\<^vec>\]\ refers some postulated \proof constant\, which is subject to schematic polymorphism of theory content, and the particular type instantiation may be given explicitly. The vector of types \\<^vec>\\ refers to the schematic type variables in the generic proposition \A\ in canonical order. A \<^emph>\proof promise\ \a : A[\<^vec>\]\ is a placeholder for some proof of polymorphic proposition \A\, with explicit type instantiation as given by the vector \\<^vec>\\, as above. Unlike axioms or oracles, proof promises may be \<^emph>\fulfilled\ eventually, by substituting \a\ by some particular proof \q\ at the corresponding type instance. This acts like Hindley-Milner \let\-polymorphism: a generic local proof definition may get used at different type instances, and is replaced by the concrete instance eventually. A \<^emph>\named theorem\ wraps up some concrete proof as a closed formal entity, in the manner of constant definitions for proof terms. The \<^emph>\proof body\ of such boxed theorems involves some digest about oracles and promises occurring in the original proof. This allows the inference kernel to manage this critical information without the full overhead of explicit proof terms. \ subsection \Reconstructing and checking proof terms\ text \ Fully explicit proof terms can be large, but most of this information is redundant and can be reconstructed from the context. Therefore, the Isabelle/Pure inference kernel records only \<^emph>\implicit\ proof terms, by omitting all typing information in terms, all term and type labels of proof abstractions, and some argument terms of applications \p \ t\ (if possible). There are separate operations to reconstruct the full proof term later on, using \<^emph>\higher-order pattern unification\ \<^cite>\"nipkow-patterns" and "Berghofer-Nipkow:2000:TPHOL"\. The \<^emph>\proof checker\ expects a fully reconstructed proof term, and can turn it into a theorem by replaying its primitive inferences within the kernel. \ subsection \Concrete syntax of proof terms\ text \ The concrete syntax of proof terms is a slight extension of the regular inner syntax of Isabelle/Pure \<^cite>\"isabelle-isar-ref"\. Its main syntactic category @{syntax (inner) proof} is defined as follows: \begin{center} \begin{supertabular}{rclr} @{syntax_def (inner) proof} & = & \\<^bold>\\ \params\ \<^verbatim>\.\ \proof\ \\ & \|\ & \proof\ \\\ \any\ \\ & \|\ & \proof\ \\\ \proof\ \\ & \|\ & \id | longid\ \\ \\ \param\ & = & \idt\ \\ & \|\ & \idt\ \<^verbatim>\:\ \prop\ \\ & \|\ & \<^verbatim>\(\ \param\ \<^verbatim>\)\ \\ \\ \params\ & = & \param\ \\ & \|\ & \param\ \params\ \\ \end{supertabular} \end{center} Implicit term arguments in partial proofs are indicated by ``\_\''. Type arguments for theorems and axioms may be specified using \p \ TYPE(type)\ (they must appear before any other term argument of a theorem or axiom, but may be omitted altogether). \<^medskip> There are separate read and print operations for proof terms, in order to avoid conflicts with the regular term language. \ text %mlref \ \begin{mldecls} @{define_ML_type proof} \\ @{define_ML_type proof_body} \\ @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{define_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{define_ML Proofterm.expand_proof: "theory -> (Proofterm.thm_header -> string option) -> proof -> proof"} \\ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\proof\ represents proof terms; this is a datatype with constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"}, @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained above. %FIXME PClass (!?) \<^descr> Type \<^ML_type>\proof_body\ represents the nested proof information of a named theorem, consisting of a digest of oracles and named theorem over some proof term. The digest only covers the directly visible part of the proof: in order to get the full information, the implicit graph of nested theorems needs to be traversed (e.g.\ using \<^ML>\Proofterm.fold_body_thms\). \<^descr> \<^ML>\Thm.proof_of\~\thm\ and \<^ML>\Thm.proof_body_of\~\thm\ produce the proof term or proof body (with digest of oracles and theorems) from a given theorem. Note that this involves a full join of internal futures that fulfill pending proof promises, and thus disrupts the natural bottom-up construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. \<^descr> \<^ML>\Proofterm.proofs\ specifies the detail of proof recording within \<^ML_type>\thm\ values produced by the inference kernel: \<^ML>\0\ records only the names of oracles, \<^ML>\1\ records oracle names and propositions, \<^ML>\2\ additionally records full proof terms. Officially named theorems that contribute to a result are recorded in any case. \<^descr> \<^ML>\Proofterm.reconstruct_proof\~\thy prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. Reconstruction may fail if \prf\ is not a proof of \prop\, or if it does not contain sufficient information for reconstruction. Failure may only happen for proofs that are constructed manually, but not for those produced automatically by the inference kernel. \<^descr> \<^ML>\Proofterm.expand_proof\~\thy expand prf\ reconstructs and expands the proofs of nested theorems according to the given \expand\ function: a result of @{ML \SOME ""\} means full expansion, @{ML \SOME\}~\name\ means to keep the theorem node but replace its internal name, @{ML NONE} means no change. \<^descr> \<^ML>\Proof_Checker.thm_of_proof\~\thy prf\ turns the given (full) proof into a theorem, by replaying it using only primitive rules of the inference kernel. \<^descr> \<^ML>\Proof_Syntax.read_proof\~\thy b\<^sub>1 b\<^sub>2 s\ reads in a proof term. The Boolean flags indicate the use of sort and type information. Usually, typing information is left implicit and is inferred during proof reconstruction. %FIXME eliminate flags!? \<^descr> \<^ML>\Proof_Syntax.pretty_proof\~\ctxt prf\ pretty-prints the given proof term. \ text %mlex \ \<^item> \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\ provides basic examples involving proof terms. \<^item> \<^file>\~~/src/HOL/Proofs/ex/XML_Data.thy\ demonstrates export and import of proof terms via XML/ML data representation. \ end diff --git a/src/HOL/Matrix_LP/Compute_Oracle/compute.ML b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML --- a/src/HOL/Matrix_LP/Compute_Oracle/compute.ML +++ b/src/HOL/Matrix_LP/Compute_Oracle/compute.ML @@ -1,645 +1,639 @@ (* Title: HOL/Matrix_LP/Compute_Oracle/compute.ML Author: Steven Obua *) signature COMPUTE = sig type computer type theorem type naming = int -> string datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML (* Functions designated with a ! in front of them actually update the computer parameter *) exception Make of string val make : machine -> theory -> thm list -> computer val make_with_cache : machine -> theory -> term list -> thm list -> computer val theory_of : computer -> theory val hyps_of : computer -> term list - val shyps_of : computer -> sort list + val shyps_of : computer -> Sortset.T (* ! *) val update : computer -> thm list -> unit (* ! *) val update_with_cache : computer -> term list -> thm list -> unit (* ! *) val set_naming : computer -> naming -> unit val naming_of : computer -> naming exception Compute of string val simplify : computer -> theorem -> thm val rewrite : computer -> cterm -> thm val make_theorem : computer -> thm -> string list -> theorem (* ! *) val instantiate : computer -> (string * cterm) list -> theorem -> theorem (* ! *) val evaluate_prem : computer -> int -> theorem -> theorem (* ! *) val modus_ponens : computer -> int -> thm -> theorem -> theorem end structure Compute :> COMPUTE = struct open Report; datatype machine = BARRAS | BARRAS_COMPILED | HASKELL | SML (* Terms are mapped to integer codes *) structure Encode :> sig type encoding val empty : encoding val insert : term -> encoding -> int * encoding val lookup_code : term -> encoding -> int option val lookup_term : int -> encoding -> term option val remove_code : int -> encoding -> encoding val remove_term : term -> encoding -> encoding end = struct type encoding = int * (int Termtab.table) * (term Inttab.table) val empty = (0, Termtab.empty, Inttab.empty) fun insert t (e as (count, term2int, int2term)) = (case Termtab.lookup term2int t of NONE => (count, (count+1, Termtab.update_new (t, count) term2int, Inttab.update_new (count, t) int2term)) | SOME code => (code, e)) fun lookup_code t (_, term2int, _) = Termtab.lookup term2int t fun lookup_term c (_, _, int2term) = Inttab.lookup int2term c fun remove_code c (e as (count, term2int, int2term)) = (case lookup_term c e of NONE => e | SOME t => (count, Termtab.delete t term2int, Inttab.delete c int2term)) fun remove_term t (e as (count, term2int, int2term)) = (case lookup_code t e of NONE => e | SOME c => (count, Termtab.delete t term2int, Inttab.delete c int2term)) end exception Make of string; exception Compute of string; local fun make_constant t encoding = let val (code, encoding) = Encode.insert t encoding in (encoding, AbstractMachine.Const code) end in fun remove_types encoding t = case t of Var _ => make_constant t encoding | Free _ => make_constant t encoding | Const _ => make_constant t encoding | Abs (_, _, t') => let val (encoding, t'') = remove_types encoding t' in (encoding, AbstractMachine.Abs t'') end | a $ b => let val (encoding, a) = remove_types encoding a val (encoding, b) = remove_types encoding b in (encoding, AbstractMachine.App (a,b)) end | Bound b => (encoding, AbstractMachine.Var b) end local fun type_of (Free (_, ty)) = ty | type_of (Const (_, ty)) = ty | type_of (Var (_, ty)) = ty | type_of _ = raise Fail "infer_types: type_of error" in fun infer_types naming encoding = let fun infer_types _ bounds _ (AbstractMachine.Var v) = (Bound v, nth bounds v) | infer_types _ bounds _ (AbstractMachine.Const code) = let val c = the (Encode.lookup_term code encoding) in (c, type_of c) end | infer_types level bounds _ (AbstractMachine.App (a, b)) = let val (a, aty) = infer_types level bounds NONE a val (adom, arange) = case aty of Type ("fun", [dom, range]) => (dom, range) | _ => raise Fail "infer_types: function type expected" val (b, _) = infer_types level bounds (SOME adom) b in (a $ b, arange) end | infer_types level bounds (SOME (ty as Type ("fun", [dom, range]))) (AbstractMachine.Abs m) = let val (m, _) = infer_types (level+1) (dom::bounds) (SOME range) m in (Abs (naming level, dom, m), ty) end | infer_types _ _ NONE (AbstractMachine.Abs _) = raise Fail "infer_types: cannot infer type of abstraction" fun infer ty term = let val (term', _) = infer_types 0 [] (SOME ty) term in term' end in infer end end datatype prog = ProgBarras of AM_Interpreter.program | ProgBarrasC of AM_Compiler.program | ProgHaskell of AM_GHC.program | ProgSML of AM_SML.program fun machine_of_prog (ProgBarras _) = BARRAS | machine_of_prog (ProgBarrasC _) = BARRAS_COMPILED | machine_of_prog (ProgHaskell _) = HASKELL | machine_of_prog (ProgSML _) = SML type naming = int -> string fun default_naming i = "v_" ^ string_of_int i datatype computer = Computer of - (theory * Encode.encoding * term list * unit Sorttab.table * prog * unit Unsynchronized.ref * naming) + (theory * Encode.encoding * term list * Sortset.T * prog * unit Unsynchronized.ref * naming) option Unsynchronized.ref fun theory_of (Computer (Unsynchronized.ref (SOME (thy,_,_,_,_,_,_)))) = thy fun hyps_of (Computer (Unsynchronized.ref (SOME (_,_,hyps,_,_,_,_)))) = hyps -fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = Sorttab.keys (shyptable) -fun shyptab_of (Computer (Unsynchronized.ref (SOME (_,_,_,shyptable,_,_,_)))) = shyptable +fun shyps_of (Computer (Unsynchronized.ref (SOME (_,_,_,shypset,_,_,_)))) = shypset fun stamp_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,stamp,_)))) = stamp fun prog_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,prog,_,_)))) = prog fun encoding_of (Computer (Unsynchronized.ref (SOME (_,encoding,_,_,_,_,_)))) = encoding fun set_encoding (Computer (r as Unsynchronized.ref (SOME (p1,_,p2,p3,p4,p5,p6)))) encoding' = (r := SOME (p1,encoding',p2,p3,p4,p5,p6)) fun naming_of (Computer (Unsynchronized.ref (SOME (_,_,_,_,_,_,n)))) = n fun set_naming (Computer (r as Unsynchronized.ref (SOME (p1,p2,p3,p4,p5,p6,_)))) naming'= (r := SOME (p1,p2,p3,p4,p5,p6,naming')) fun ref_of (Computer r) = r -datatype cthm = ComputeThm of term list * sort list * term +datatype cthm = ComputeThm of term list * Sortset.T * term fun thm2cthm th = (if not (null (Thm.tpairs_of th)) then raise Make "theorems may not contain tpairs" else (); ComputeThm (Thm.hyps_of th, Thm.shyps_of th, Thm.prop_of th)) fun make_internal machine thy stamp encoding cache_pattern_terms raw_ths = let fun transfer (x:thm) = Thm.transfer thy x val ths = map (thm2cthm o Thm.strip_shyps o transfer) raw_ths fun make_pattern encoding n vars (AbstractMachine.Abs _) = raise (Make "no lambda abstractions allowed in pattern") | make_pattern encoding n vars (AbstractMachine.Var _) = raise (Make "no bound variables allowed in pattern") | make_pattern encoding n vars (AbstractMachine.Const code) = (case the (Encode.lookup_term code encoding) of Var _ => ((n+1, Inttab.update_new (code, n) vars, AbstractMachine.PVar) handle Inttab.DUP _ => raise (Make "no duplicate variable in pattern allowed")) | _ => (n, vars, AbstractMachine.PConst (code, []))) | make_pattern encoding n vars (AbstractMachine.App (a, b)) = let val (n, vars, pa) = make_pattern encoding n vars a val (n, vars, pb) = make_pattern encoding n vars b in case pa of AbstractMachine.PVar => raise (Make "patterns may not start with a variable") | AbstractMachine.PConst (c, args) => (n, vars, AbstractMachine.PConst (c, args@[pb])) end - fun thm2rule (encoding, hyptable, shyptable) th = + fun thm2rule (encoding, hyptable, shypset) th = let val (ComputeThm (hyps, shyps, prop)) = th val hyptable = fold (fn h => Termtab.update (h, ())) hyps hyptable - val shyptable = fold (fn sh => Sorttab.update (sh, ())) shyps shyptable + val shypset = Sortset.merge (shyps, shypset) val (prems, prop) = (Logic.strip_imp_prems prop, Logic.strip_imp_concl prop) val (a, b) = Logic.dest_equals prop handle TERM _ => raise (Make "theorems must be meta-level equations (with optional guards)") val a = Envir.eta_contract a val b = Envir.eta_contract b val prems = map Envir.eta_contract prems val (encoding, left) = remove_types encoding a val (encoding, right) = remove_types encoding b fun remove_types_of_guard encoding g = (let val (t1, t2) = Logic.dest_equals g val (encoding, t1) = remove_types encoding t1 val (encoding, t2) = remove_types encoding t2 in (encoding, AbstractMachine.Guard (t1, t2)) end handle TERM _ => raise (Make "guards must be meta-level equations")) val (encoding, prems) = fold_rev (fn p => fn (encoding, ps) => let val (e, p) = remove_types_of_guard encoding p in (e, p::ps) end) prems (encoding, []) (* Principally, a check should be made here to see if the (meta-) hyps contain any of the variables of the rule. As it is, all variables of the rule are schematic, and there are no schematic variables in meta-hyps, therefore this check can be left out. *) val (vcount, vars, pattern) = make_pattern encoding 0 Inttab.empty left val _ = (case pattern of AbstractMachine.PVar => raise (Make "patterns may not start with a variable") | _ => ()) (* finally, provide a function for renaming the pattern bound variables on the right hand side *) fun rename level vars (var as AbstractMachine.Var _) = var | rename level vars (c as AbstractMachine.Const code) = (case Inttab.lookup vars code of NONE => c | SOME n => AbstractMachine.Var (vcount-n-1+level)) | rename level vars (AbstractMachine.App (a, b)) = AbstractMachine.App (rename level vars a, rename level vars b) | rename level vars (AbstractMachine.Abs m) = AbstractMachine.Abs (rename (level+1) vars m) fun rename_guard (AbstractMachine.Guard (a,b)) = AbstractMachine.Guard (rename 0 vars a, rename 0 vars b) in - ((encoding, hyptable, shyptable), (map rename_guard prems, pattern, rename 0 vars right)) + ((encoding, hyptable, shypset), (map rename_guard prems, pattern, rename 0 vars right)) end - val ((encoding, hyptable, shyptable), rules) = + val ((encoding, hyptable, shypset), rules) = fold_rev (fn th => fn (encoding_hyptable, rules) => let val (encoding_hyptable, rule) = thm2rule encoding_hyptable th in (encoding_hyptable, rule::rules) end) - ths ((encoding, Termtab.empty, Sorttab.empty), []) + ths ((encoding, Termtab.empty, Sortset.empty), []) fun make_cache_pattern t (encoding, cache_patterns) = let val (encoding, a) = remove_types encoding t val (_,_,p) = make_pattern encoding 0 Inttab.empty a in (encoding, p::cache_patterns) end val (encoding, _) = fold_rev make_cache_pattern cache_pattern_terms (encoding, []) val prog = case machine of BARRAS => ProgBarras (AM_Interpreter.compile rules) | BARRAS_COMPILED => ProgBarrasC (AM_Compiler.compile rules) | HASKELL => ProgHaskell (AM_GHC.compile rules) | SML => ProgSML (AM_SML.compile rules) fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) - val shyptable = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptable))) shyptable + val shypset = Sortset.fold (fn s => has_witness s ? Sortset.remove s) shypset shypset - in (thy, encoding, Termtab.keys hyptable, shyptable, prog, stamp, default_naming) end + in (thy, encoding, Termtab.keys hyptable, shypset, prog, stamp, default_naming) end fun make_with_cache machine thy cache_patterns raw_thms = Computer (Unsynchronized.ref (SOME (make_internal machine thy (Unsynchronized.ref ()) Encode.empty cache_patterns raw_thms))) fun make machine thy raw_thms = make_with_cache machine thy [] raw_thms fun update_with_cache computer cache_patterns raw_thms = let val c = make_internal (machine_of_prog (prog_of computer)) (theory_of computer) (stamp_of computer) (encoding_of computer) cache_patterns raw_thms val _ = (ref_of computer) := SOME c in () end fun update computer raw_thms = update_with_cache computer [] raw_thms fun runprog (ProgBarras p) = AM_Interpreter.run p | runprog (ProgBarrasC p) = AM_Compiler.run p | runprog (ProgHaskell p) = AM_GHC.run p | runprog (ProgSML p) = AM_SML.run p (* ------------------------------------------------------------------------------------- *) (* An oracle for exporting theorems; must only be accessible from inside this structure! *) (* ------------------------------------------------------------------------------------- *) fun merge_hyps hyps1 hyps2 = let fun add hyps tab = fold (fn h => fn tab => Termtab.update (h, ()) tab) hyps tab in Termtab.keys (add hyps2 (add hyps1 Termtab.empty)) end -fun add_shyps shyps tab = fold (fn h => fn tab => Sorttab.update (h, ()) tab) shyps tab - -fun merge_shyps shyps1 shyps2 = Sorttab.keys (add_shyps shyps2 (add_shyps shyps1 Sorttab.empty)) - val (_, export_oracle) = Context.>>> (Context.map_theory_result (Thm.add_oracle (\<^binding>\compute\, fn (thy, hyps, shyps, prop) => let - val shyptab = add_shyps shyps Sorttab.empty - fun delete s shyptab = Sorttab.delete s shyptab handle Sorttab.UNDEF _ => shyptab - fun delete_term t shyptab = fold delete (Sorts.insert_term t []) shyptab + fun remove_term t = Sortset.subtract (Sortset.build (Sortset.insert_term t)) fun has_witness s = not (null (Sign.witness_sorts thy [] [s])) - val shyptab = fold Sorttab.delete (filter has_witness (Sorttab.keys (shyptab))) shyptab - val shyps = if Sorttab.is_empty shyptab then [] else Sorttab.keys (fold delete_term (prop::hyps) shyptab) + val shyps = Sortset.fold (fn s => has_witness s ? Sortset.remove s) shyps shyps + val shyps = + if Sortset.is_empty shyps then Sortset.empty + else fold remove_term (prop::hyps) shyps val _ = - if not (null shyps) then + if not (Sortset.is_empty shyps) then raise Compute ("dangling sort hypotheses: " ^ - commas (map (Syntax.string_of_sort_global thy) shyps)) + commas (map (Syntax.string_of_sort_global thy) (Sortset.dest shyps))) else () in Thm.global_cterm_of thy (fold_rev (fn hyp => fn p => Logic.mk_implies (hyp, p)) hyps prop) end))); fun export_thm thy hyps shyps prop = let val th = export_oracle (thy, hyps, shyps, prop) val hyps = map (fn h => Thm.assume (Thm.global_cterm_of thy h)) hyps in fold (fn h => fn p => Thm.implies_elim p h) hyps th end (* --------- Rewrite ----------- *) fun rewrite computer raw_ct = let val thy = theory_of computer val ct = Thm.transfer_cterm thy raw_ct val t' = Thm.term_of ct val ty = Thm.typ_of_cterm ct val naming = naming_of computer val (encoding, t) = remove_types (encoding_of computer) t' val t = runprog (prog_of computer) t val t = infer_types naming encoding ty t val eq = Logic.mk_equals (t', t) in - export_thm thy (hyps_of computer) (Sorttab.keys (shyptab_of computer)) eq + export_thm thy (hyps_of computer) (shyps_of computer) eq end (* --------- Simplify ------------ *) datatype prem = EqPrem of AbstractMachine.term * AbstractMachine.term * Term.typ * int | Prem of AbstractMachine.term datatype theorem = Theorem of theory * unit Unsynchronized.ref * (int * typ) Symtab.table * (AbstractMachine.term option) Inttab.table - * prem list * AbstractMachine.term * term list * sort list + * prem list * AbstractMachine.term * term list * Sortset.T exception ParamSimplify of computer * theorem fun make_theorem computer raw_th vars = let val thy = theory_of computer val th = Thm.transfer thy raw_th val (ComputeThm (hyps, shyps, prop)) = thm2cthm th val encoding = encoding_of computer (* variables in the theorem are identified upfront *) fun collect_vars (Abs (_, _, t)) tab = collect_vars t tab | collect_vars (a $ b) tab = collect_vars b (collect_vars a tab) | collect_vars (Const _) tab = tab | collect_vars (Free _) tab = tab | collect_vars (Var ((s, i), ty)) tab = if List.find (fn x => x=s) vars = NONE then tab else (case Symtab.lookup tab s of SOME ((s',i'),ty') => if s' <> s orelse i' <> i orelse ty <> ty' then raise Compute ("make_theorem: variable name '"^s^"' is not unique") else tab | NONE => Symtab.update (s, ((s, i), ty)) tab) val vartab = collect_vars prop Symtab.empty fun encodevar (s, t as (_, ty)) (encoding, tab) = let val (x, encoding) = Encode.insert (Var t) encoding in (encoding, Symtab.update (s, (x, ty)) tab) end val (encoding, vartab) = Symtab.fold encodevar vartab (encoding, Symtab.empty) val varsubst = Inttab.make (map (fn (_, (x, _)) => (x, NONE)) (Symtab.dest vartab)) (* make the premises and the conclusion *) fun mk_prem encoding t = (let val (a, b) = Logic.dest_equals t val ty = type_of a val (encoding, a) = remove_types encoding a val (encoding, b) = remove_types encoding b val (eq, encoding) = Encode.insert (Const (\<^const_name>\Pure.eq\, ty --> ty --> \<^typ>\prop\)) encoding in (encoding, EqPrem (a, b, ty, eq)) end handle TERM _ => let val (encoding, t) = remove_types encoding t in (encoding, Prem t) end) val (encoding, prems) = (fold_rev (fn t => fn (encoding, l) => case mk_prem encoding t of (encoding, t) => (encoding, t::l)) (Logic.strip_imp_prems prop) (encoding, [])) val (encoding, concl) = remove_types encoding (Logic.strip_imp_concl prop) val _ = set_encoding computer encoding in Theorem (thy, stamp_of computer, vartab, varsubst, prems, concl, hyps, shyps) end fun theory_of_theorem (Theorem (thy,_,_,_,_,_,_,_)) = thy fun update_theory thy (Theorem (_,p0,p1,p2,p3,p4,p5,p6)) = Theorem (thy,p0,p1,p2,p3,p4,p5,p6) fun stamp_of_theorem (Theorem (_,s, _, _, _, _, _, _)) = s fun vartab_of_theorem (Theorem (_,_,vt,_,_,_,_,_)) = vt fun varsubst_of_theorem (Theorem (_,_,_,vs,_,_,_,_)) = vs fun update_varsubst vs (Theorem (p0,p1,p2,_,p3,p4,p5,p6)) = Theorem (p0,p1,p2,vs,p3,p4,p5,p6) fun prems_of_theorem (Theorem (_,_,_,_,prems,_,_,_)) = prems fun update_prems prems (Theorem (p0,p1,p2,p3,_,p4,p5,p6)) = Theorem (p0,p1,p2,p3,prems,p4,p5,p6) fun concl_of_theorem (Theorem (_,_,_,_,_,concl,_,_)) = concl fun hyps_of_theorem (Theorem (_,_,_,_,_,_,hyps,_)) = hyps fun update_hyps hyps (Theorem (p0,p1,p2,p3,p4,p5,_,p6)) = Theorem (p0,p1,p2,p3,p4,p5,hyps,p6) fun shyps_of_theorem (Theorem (_,_,_,_,_,_,_,shyps)) = shyps fun update_shyps shyps (Theorem (p0,p1,p2,p3,p4,p5,p6,_)) = Theorem (p0,p1,p2,p3,p4,p5,p6,shyps) fun check_compatible computer th s = if stamp_of computer <> stamp_of_theorem th then raise Compute (s^": computer and theorem are incompatible") else () fun instantiate computer insts th = let val _ = check_compatible computer th val thy = theory_of computer val vartab = vartab_of_theorem th fun rewrite computer t = let val (encoding, t) = remove_types (encoding_of computer) t val t = runprog (prog_of computer) t val _ = set_encoding computer encoding in t end fun assert_varfree vs t = if AbstractMachine.forall_consts (fn x => Inttab.lookup vs x = NONE) t then () else raise Compute "instantiate: assert_varfree failed" fun assert_closed t = if AbstractMachine.closed t then () else raise Compute "instantiate: not a closed term" fun compute_inst (s, raw_ct) vs = let val ct = Thm.transfer_cterm thy raw_ct val ty = Thm.typ_of_cterm ct in (case Symtab.lookup vartab s of NONE => raise Compute ("instantiate: variable '"^s^"' not found in theorem") | SOME (x, ty') => (case Inttab.lookup vs x of SOME (SOME _) => raise Compute ("instantiate: variable '"^s^"' has already been instantiated") | SOME NONE => if ty <> ty' then raise Compute ("instantiate: wrong type for variable '"^s^"'") else let val t = rewrite computer (Thm.term_of ct) val _ = assert_varfree vs t val _ = assert_closed t in Inttab.update (x, SOME t) vs end | NONE => raise Compute "instantiate: internal error")) end val vs = fold compute_inst insts (varsubst_of_theorem th) in update_varsubst vs th end fun match_aterms subst = let exception no_match open AbstractMachine fun match subst (b as (Const c)) a = if a = b then subst else (case Inttab.lookup subst c of SOME (SOME a') => if a=a' then subst else raise no_match | SOME NONE => if AbstractMachine.closed a then Inttab.update (c, SOME a) subst else raise no_match | NONE => raise no_match) | match subst (b as (Var _)) a = if a=b then subst else raise no_match | match subst (App (u, v)) (App (u', v')) = match (match subst u u') v v' | match subst (Abs u) (Abs u') = match subst u u' | match subst _ _ = raise no_match in fn b => fn a => (SOME (match subst b a) handle no_match => NONE) end fun apply_subst vars_allowed subst = let open AbstractMachine fun app (t as (Const c)) = (case Inttab.lookup subst c of NONE => t | SOME (SOME t) => Computed t | SOME NONE => if vars_allowed then t else raise Compute "apply_subst: no vars allowed") | app (t as (Var _)) = t | app (App (u, v)) = App (app u, app v) | app (Abs m) = Abs (app m) in app end fun splicein n l L = List.take (L, n) @ l @ List.drop (L, n+1) fun evaluate_prem computer prem_no th = let val _ = check_compatible computer th val prems = prems_of_theorem th val varsubst = varsubst_of_theorem th fun run vars_allowed t = runprog (prog_of computer) (apply_subst vars_allowed varsubst t) in case nth prems prem_no of Prem _ => raise Compute "evaluate_prem: no equality premise" | EqPrem (a, b, ty, _) => let val a' = run false a val b' = run true b in case match_aterms varsubst b' a' of NONE => let fun mk s = Syntax.string_of_term_global \<^theory>\Pure\ (infer_types (naming_of computer) (encoding_of computer) ty s) val left = "computed left side: "^(mk a') val right = "computed right side: "^(mk b') in raise Compute ("evaluate_prem: cannot assign computed left to right hand side\n"^left^"\n"^right^"\n") end | SOME varsubst => update_prems (splicein prem_no [] prems) (update_varsubst varsubst th) end end fun prem2term (Prem t) = t | prem2term (EqPrem (a,b,_,eq)) = AbstractMachine.App (AbstractMachine.App (AbstractMachine.Const eq, a), b) fun modus_ponens computer prem_no raw_th' th = let val thy = theory_of computer val _ = check_compatible computer th val _ = Context.subthy (theory_of_theorem th, thy) orelse raise Compute "modus_ponens: bad theory" val th' = make_theorem computer (Thm.transfer thy raw_th') [] val varsubst = varsubst_of_theorem th fun run vars_allowed t = runprog (prog_of computer) (apply_subst vars_allowed varsubst t) val prems = prems_of_theorem th val prem = run true (prem2term (nth prems prem_no)) val concl = run false (concl_of_theorem th') in case match_aterms varsubst prem concl of NONE => raise Compute "modus_ponens: conclusion does not match premise" | SOME varsubst => let val th = update_varsubst varsubst th val th = update_prems (splicein prem_no (prems_of_theorem th') prems) th val th = update_hyps (merge_hyps (hyps_of_theorem th) (hyps_of_theorem th')) th - val th = update_shyps (merge_shyps (shyps_of_theorem th) (shyps_of_theorem th')) th + val th = update_shyps (Sortset.merge (shyps_of_theorem th, shyps_of_theorem th')) th in update_theory thy th end end fun simplify computer th = let val _ = check_compatible computer th val varsubst = varsubst_of_theorem th val encoding = encoding_of computer val naming = naming_of computer fun infer t = infer_types naming encoding \<^typ>\prop\ t fun run t = infer (runprog (prog_of computer) (apply_subst true varsubst t)) fun runprem p = run (prem2term p) val prop = Logic.list_implies (map runprem (prems_of_theorem th), run (concl_of_theorem th)) val hyps = merge_hyps (hyps_of computer) (hyps_of_theorem th) - val shyps = merge_shyps (shyps_of_theorem th) (Sorttab.keys (shyptab_of computer)) + val shyps = Sortset.merge (shyps_of_theorem th, shyps_of computer) in export_thm (theory_of_theorem th) hyps shyps prop end end - diff --git a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML --- a/src/HOL/Matrix_LP/Compute_Oracle/linker.ML +++ b/src/HOL/Matrix_LP/Compute_Oracle/linker.ML @@ -1,467 +1,467 @@ (* Title: HOL/Matrix_LP/Compute_Oracle/linker.ML Author: Steven Obua This module solves the problem that the computing oracle does not instantiate polymorphic rules. By going through the PCompute interface, all possible instantiations are resolved by compiling new programs, if necessary. The obvious disadvantage of this approach is that in the worst case for each new term to be rewritten, a new program may be compiled. *) (* Given constants/frees c_1::t_1, c_2::t_2, ...., c_n::t_n, and constants/frees d_1::d_1, d_2::s_2, ..., d_m::s_m Find all substitutions S such that a) the domain of S is tvars (t_1, ..., t_n) b) there are indices i_1, ..., i_k, and j_1, ..., j_k with 1. S (c_i_1::t_i_1) = d_j_1::s_j_1, ..., S (c_i_k::t_i_k) = d_j_k::s_j_k 2. tvars (t_i_1, ..., t_i_k) = tvars (t_1, ..., t_n) *) signature LINKER = sig exception Link of string datatype constant = Constant of bool * string * typ val constant_of : term -> constant type instances type subst = Type.tyenv val empty : constant list -> instances val typ_of_constant : constant -> typ val add_instances : theory -> instances -> constant list -> subst list * instances val substs_of : instances -> subst list val is_polymorphic : constant -> bool val distinct_constants : constant list -> constant list val collect_consts : term list -> constant list end structure Linker : LINKER = struct exception Link of string; type subst = Type.tyenv datatype constant = Constant of bool * string * typ fun constant_of (Const (name, ty)) = Constant (false, name, ty) | constant_of (Free (name, ty)) = Constant (true, name, ty) | constant_of _ = raise Link "constant_of" fun bool_ord (x,y) = if x then (if y then EQUAL else GREATER) else (if y then LESS else EQUAL) fun constant_ord (Constant (x1,x2,x3), Constant (y1,y2,y3)) = (prod_ord (prod_ord bool_ord fast_string_ord) Term_Ord.typ_ord) (((x1,x2),x3), ((y1,y2),y3)) fun constant_modty_ord (Constant (x1,x2,_), Constant (y1,y2,_)) = (prod_ord bool_ord fast_string_ord) ((x1,x2), (y1,y2)) structure Consttab = Table(type key = constant val ord = constant_ord); structure ConsttabModTy = Table(type key = constant val ord = constant_modty_ord); fun typ_of_constant (Constant (_, _, ty)) = ty val empty_subst = (Vartab.empty : Type.tyenv) fun merge_subst (A:Type.tyenv) (B:Type.tyenv) = SOME (Vartab.fold (fn (v, t) => fn tab => (case Vartab.lookup tab v of NONE => Vartab.update (v, t) tab | SOME t' => if t = t' then tab else raise Type.TYPE_MATCH)) A B) handle Type.TYPE_MATCH => NONE fun subst_ord (A:Type.tyenv, B:Type.tyenv) = (list_ord (prod_ord Term_Ord.fast_indexname_ord (prod_ord Term_Ord.sort_ord Term_Ord.typ_ord))) (Vartab.dest A, Vartab.dest B) structure Substtab = Table(type key = Type.tyenv val ord = subst_ord); fun substtab_union c = Substtab.fold Substtab.update c fun substtab_unions [] = Substtab.empty | substtab_unions [c] = c | substtab_unions (c::cs) = substtab_union c (substtab_unions cs) datatype instances = Instances of unit ConsttabModTy.table * Type.tyenv Consttab.table Consttab.table * constant list list * unit Substtab.table fun is_polymorphic (Constant (_, _, ty)) = not (null (Term.add_tvarsT ty [])) fun distinct_constants cs = Consttab.keys (fold (fn c => Consttab.update (c, ())) cs Consttab.empty) fun empty cs = let val cs = distinct_constants (filter is_polymorphic cs) val old_cs = cs (* fun collect_tvars ty tab = fold (fn v => fn tab => Typtab.update (TVar v, ()) tab) (Misc_Legacy.typ_tvars ty) tab val tvars_count = length (Typtab.keys (fold (fn c => fn tab => collect_tvars (typ_of_constant c) tab) cs Typtab.empty)) fun tvars_of ty = collect_tvars ty Typtab.empty val cs = map (fn c => (c, tvars_of (typ_of_constant c))) cs fun tyunion A B = Typtab.fold (fn (v,()) => fn tab => Typtab.update (v, case Typtab.lookup tab v of NONE => 1 | SOME n => n+1) tab) A B fun is_essential A B = Typtab.fold (fn (v, ()) => fn essential => essential orelse (case Typtab.lookup B v of NONE => raise Link "is_essential" | SOME n => n=1)) A false fun add_minimal (c', tvs') (tvs, cs) = let val tvs = tyunion tvs' tvs val cs = (c', tvs')::cs in if forall (fn (c',tvs') => is_essential tvs' tvs) cs then SOME (tvs, cs) else NONE end fun is_spanning (tvs, _) = (length (Typtab.keys tvs) = tvars_count) fun generate_minimal_subsets subsets [] = subsets | generate_minimal_subsets subsets (c::cs) = let val subsets' = map_filter (add_minimal c) subsets in generate_minimal_subsets (subsets@subsets') cs end*) val minimal_subsets = [old_cs] (*map (fn (tvs, cs) => map fst cs) (filter is_spanning (generate_minimal_subsets [(Typtab.empty, [])] cs))*) val constants = Consttab.keys (fold (fold (fn c => Consttab.update (c, ()))) minimal_subsets Consttab.empty) in Instances ( fold (fn c => fn tab => ConsttabModTy.update (c, ()) tab) constants ConsttabModTy.empty, Consttab.make (map (fn c => (c, Consttab.empty : Type.tyenv Consttab.table)) constants), minimal_subsets, Substtab.empty) end local fun calc ctab substtab [] = substtab | calc ctab substtab (c::cs) = let val csubsts = map snd (Consttab.dest (the (Consttab.lookup ctab c))) fun merge_substs substtab subst = Substtab.fold (fn (s,_) => fn tab => (case merge_subst subst s of NONE => tab | SOME s => Substtab.update (s, ()) tab)) substtab Substtab.empty val substtab = substtab_unions (map (merge_substs substtab) csubsts) in calc ctab substtab cs end in fun calc_substs ctab (cs:constant list) = calc ctab (Substtab.update (empty_subst, ()) Substtab.empty) cs end fun add_instances thy (Instances (cfilter, ctab,minsets,substs)) cs = let (* val _ = writeln (makestring ("add_instances: ", length_cs, length cs, length (Consttab.keys ctab)))*) fun calc_instantiations (constant as Constant (free, name, ty)) instantiations = Consttab.fold (fn (constant' as Constant (free', name', ty'), insttab) => fn instantiations => if free <> free' orelse name <> name' then instantiations else case Consttab.lookup insttab constant of SOME _ => instantiations | NONE => ((constant', (constant, Sign.typ_match thy (ty', ty) empty_subst))::instantiations handle Type.TYPE_MATCH => instantiations)) ctab instantiations val instantiations = fold calc_instantiations cs [] (*val _ = writeln ("instantiations = "^(makestring (length instantiations)))*) fun update_ctab (constant', entry) ctab = (case Consttab.lookup ctab constant' of NONE => raise Link "internal error: update_ctab" | SOME tab => Consttab.update (constant', Consttab.update entry tab) ctab) val ctab = fold update_ctab instantiations ctab val new_substs = fold (fn minset => fn substs => substtab_union (calc_substs ctab minset) substs) minsets Substtab.empty val (added_substs, substs) = Substtab.fold (fn (ns, _) => fn (added, substtab) => (case Substtab.lookup substs ns of NONE => (ns::added, Substtab.update (ns, ()) substtab) | SOME () => (added, substtab))) new_substs ([], substs) in (added_substs, Instances (cfilter, ctab, minsets, substs)) end fun substs_of (Instances (_,_,_,substs)) = Substtab.keys substs local fun collect (Var _) tab = tab | collect (Bound _) tab = tab | collect (a $ b) tab = collect b (collect a tab) | collect (Abs (_, _, body)) tab = collect body tab | collect t tab = Consttab.update (constant_of t, ()) tab in fun collect_consts tms = Consttab.keys (fold collect tms Consttab.empty) end end signature PCOMPUTE = sig type pcomputer val make : Compute.machine -> theory -> thm list -> Linker.constant list -> pcomputer val make_with_cache : Compute.machine -> theory -> term list -> thm list -> Linker.constant list -> pcomputer val add_instances : pcomputer -> Linker.constant list -> bool val add_instances' : pcomputer -> term list -> bool val rewrite : pcomputer -> cterm list -> thm list val simplify : pcomputer -> Compute.theorem -> thm val make_theorem : pcomputer -> thm -> string list -> Compute.theorem val instantiate : pcomputer -> (string * cterm) list -> Compute.theorem -> Compute.theorem val evaluate_prem : pcomputer -> int -> Compute.theorem -> Compute.theorem val modus_ponens : pcomputer -> int -> thm -> Compute.theorem -> Compute.theorem end structure PCompute : PCOMPUTE = struct exception PCompute of string datatype theorem = MonoThm of thm | PolyThm of thm * Linker.instances * thm list datatype pattern = MonoPattern of term | PolyPattern of term * Linker.instances * term list datatype pcomputer = PComputer of theory * Compute.computer * theorem list Unsynchronized.ref * pattern list Unsynchronized.ref (*fun collect_consts (Var x) = [] | collect_consts (Bound _) = [] | collect_consts (a $ b) = (collect_consts a)@(collect_consts b) | collect_consts (Abs (_, _, body)) = collect_consts body | collect_consts t = [Linker.constant_of t]*) fun computer_of (PComputer (_,computer,_,_)) = computer fun collect_consts_of_thm th = let val th = Thm.prop_of th val (prems, th) = (Logic.strip_imp_prems th, Logic.strip_imp_concl th) val (left, right) = Logic.dest_equals th in (Linker.collect_consts [left], Linker.collect_consts (right::prems)) end fun create_theorem th = let val (left, right) = collect_consts_of_thm th val polycs = filter Linker.is_polymorphic left val tytab = fold (fn p => fn tab => fold (fn n => fn tab => Typtab.update (TVar n, ()) tab) (Misc_Legacy.typ_tvars (Linker.typ_of_constant p)) tab) polycs Typtab.empty fun check_const (c::cs) cs' = let val tvars = Misc_Legacy.typ_tvars (Linker.typ_of_constant c) val wrong = fold (fn n => fn wrong => wrong orelse is_none (Typtab.lookup tytab (TVar n))) tvars false in if wrong then raise PCompute "right hand side of theorem contains type variables which do not occur on the left hand side" else if null (tvars) then check_const cs (c::cs') else check_const cs cs' end | check_const [] cs' = cs' val monocs = check_const right [] in if null (polycs) then (monocs, MonoThm th) else (monocs, PolyThm (th, Linker.empty polycs, [])) end fun create_pattern pat = let val cs = Linker.collect_consts [pat] val polycs = filter Linker.is_polymorphic cs in if null (polycs) then MonoPattern pat else PolyPattern (pat, Linker.empty polycs, []) end fun create_computer machine thy pats ths = let fun add (MonoThm th) ths = th::ths | add (PolyThm (_, _, ths')) ths = ths'@ths fun addpat (MonoPattern p) pats = p::pats | addpat (PolyPattern (_, _, ps)) pats = ps@pats val ths = fold_rev add ths [] val pats = fold_rev addpat pats [] in Compute.make_with_cache machine thy pats ths end fun update_computer computer pats ths = let fun add (MonoThm th) ths = th::ths | add (PolyThm (_, _, ths')) ths = ths'@ths fun addpat (MonoPattern p) pats = p::pats | addpat (PolyPattern (_, _, ps)) pats = ps@pats val ths = fold_rev add ths [] val pats = fold_rev addpat pats [] in Compute.update_with_cache computer pats ths end fun conv_subst thy (subst : Type.tyenv) = map (fn (iname, (sort, ty)) => ((iname, sort), Thm.global_ctyp_of thy ty)) (Vartab.dest subst) fun add_monos thy monocs pats ths = let val changed = Unsynchronized.ref false fun add monocs (th as (MonoThm _)) = ([], th) | add monocs (PolyThm (th, instances, instanceths)) = let val (newsubsts, instances) = Linker.add_instances thy instances monocs val _ = if not (null newsubsts) then changed := true else () val newths = map (fn subst => Thm.instantiate (TVars.make (conv_subst thy subst), Vars.empty) th) newsubsts (* val _ = if not (null newths) then (print ("added new theorems: ", newths); ()) else ()*) val newmonos = fold (fn th => fn monos => (snd (collect_consts_of_thm th))@monos) newths [] in (newmonos, PolyThm (th, instances, instanceths@newths)) end fun addpats monocs (pat as (MonoPattern _)) = pat | addpats monocs (PolyPattern (p, instances, instancepats)) = let val (newsubsts, instances) = Linker.add_instances thy instances monocs val _ = if not (null newsubsts) then changed := true else () val newpats = map (fn subst => Envir.subst_term_types subst p) newsubsts in PolyPattern (p, instances, instancepats@newpats) end fun step monocs ths = fold_rev (fn th => fn (newmonos, ths) => let val (newmonos', th') = add monocs th in (newmonos'@newmonos, th'::ths) end) ths ([], []) fun loop monocs pats ths = let val (monocs', ths') = step monocs ths val pats' = map (addpats monocs) pats in if null (monocs') then (pats', ths') else loop monocs' pats' ths' end val result = loop monocs pats ths in (!changed, result) end -datatype cthm = ComputeThm of term list * sort list * term +datatype cthm = ComputeThm of term list * Sortset.T * term fun thm2cthm th = ComputeThm (Thm.hyps_of th, Thm.shyps_of th, Thm.prop_of th) val cthm_ord' = - prod_ord (prod_ord (list_ord Term_Ord.term_ord) (list_ord Term_Ord.sort_ord)) Term_Ord.term_ord + prod_ord (prod_ord (list_ord Term_Ord.term_ord) Sortset.ord) Term_Ord.term_ord fun cthm_ord (ComputeThm (h1, sh1, p1), ComputeThm (h2, sh2, p2)) = cthm_ord' (((h1,sh1), p1), ((h2, sh2), p2)) structure CThmtab = Table(type key = cthm val ord = cthm_ord) fun remove_duplicates ths = let val counter = Unsynchronized.ref 0 val tab = Unsynchronized.ref (CThmtab.empty : unit CThmtab.table) val thstab = Unsynchronized.ref (Inttab.empty : thm Inttab.table) fun update th = let val key = thm2cthm th in case CThmtab.lookup (!tab) key of NONE => ((tab := CThmtab.update_new (key, ()) (!tab)); thstab := Inttab.update_new (!counter, th) (!thstab); counter := !counter + 1) | _ => () end val _ = map update ths in map snd (Inttab.dest (!thstab)) end fun make_with_cache machine thy pats ths cs = let val ths = remove_duplicates ths val (monocs, ths) = fold_rev (fn th => fn (monocs, ths) => let val (m, t) = create_theorem th in (m@monocs, t::ths) end) ths (cs, []) val pats = map create_pattern pats val (_, (pats, ths)) = add_monos thy monocs pats ths val computer = create_computer machine thy pats ths in PComputer (thy, computer, Unsynchronized.ref ths, Unsynchronized.ref pats) end fun make machine thy ths cs = make_with_cache machine thy [] ths cs fun add_instances (PComputer (thy, computer, rths, rpats)) cs = let val (changed, (pats, ths)) = add_monos thy cs (!rpats) (!rths) in if changed then (update_computer computer pats ths; rths := ths; rpats := pats; true) else false end fun add_instances' pc ts = add_instances pc (Linker.collect_consts ts) fun rewrite pc cts = let val _ = add_instances' pc (map Thm.term_of cts) val computer = (computer_of pc) in map (fn ct => Compute.rewrite computer ct) cts end fun simplify pc th = Compute.simplify (computer_of pc) th fun make_theorem pc th vars = let val _ = add_instances' pc [Thm.prop_of th] in Compute.make_theorem (computer_of pc) th vars end fun instantiate pc insts th = let val _ = add_instances' pc (map (Thm.term_of o snd) insts) in Compute.instantiate (computer_of pc) insts th end fun evaluate_prem pc prem_no th = Compute.evaluate_prem (computer_of pc) prem_no th fun modus_ponens pc prem_no th' th = let val _ = add_instances' pc [Thm.prop_of th'] in Compute.modus_ponens (computer_of pc) prem_no th' th end end diff --git a/src/HOL/Types_To_Sets/internalize_sort.ML b/src/HOL/Types_To_Sets/internalize_sort.ML --- a/src/HOL/Types_To_Sets/internalize_sort.ML +++ b/src/HOL/Types_To_Sets/internalize_sort.ML @@ -1,68 +1,68 @@ (* Title: HOL/Types_To_Sets/internalize_sort.ML Author: Ondřej Kunčar, TU München A derived rule (by using Thm.unconstrainT) that internalizes type class annotations. *) (* \['a::{c_1, ..., c_n} / 'a] --------------------------------------------------------------------- class.c_1 ops_1 \ ... \ class.c_n ops_n \ \['a::type / 'a] where class.c is the locale predicate associated with type class c and ops are operations associated with type class c. For example: If c = semigroup_add, then ops = op-, op+, 0, uminus. If c = finite, then ops = TYPE('a::type). *) signature INTERNALIZE_SORT = sig val internalize_sort: ctyp -> thm -> typ * thm val internalize_sort_attr: typ -> attribute end; structure Internalize_Sort : INTERNALIZE_SORT = struct fun internalize_sort ctvar thm = let val thy = Thm.theory_of_thm thm; val tvar = Thm.typ_of ctvar; - val (ucontext, _) = Logic.unconstrainT [] (Thm.prop_of thm); + val (ucontext, _) = Logic.unconstrainT Sortset.empty (Thm.prop_of thm); fun is_proper_class thy = can (Axclass.get_info thy); (* trick by FH *) fun reduce_to_non_proper_sort (TVar (name, sort)) = TVar (name, Sign.minimize_sort thy (filter_out (is_proper_class thy) (Sign.complete_sort thy sort))) | reduce_to_non_proper_sort _ = error "not supported"; val data = map #1 (#outer_constraints ucontext) ~~ #constraints ucontext; val new_tvar = get_first (fn (tvar', ((ren_tvar, _), _)) => if tvar = tvar' then SOME (reduce_to_non_proper_sort ren_tvar) else NONE) data |> the_default tvar; fun localify class = Class.rules thy class |> snd |> Thm.transfer thy; fun discharge_of_class tvar class = Thm.of_class (Thm.global_ctyp_of thy tvar, class); val rules = map (fn (tvar', ((ren_tvar, class), _)) => if tvar = tvar' then (if Class.is_class thy class then localify class else discharge_of_class ren_tvar class) else discharge_of_class ren_tvar class) data; in (new_tvar, (Thm.unconstrainT (Thm.strip_shyps thm) OF rules) |> Drule.zero_var_indexes) end; val tvar = Args.context -- Args.typ >> (fn (_, v as TFree _) => Logic.varifyT_global v | (ctxt, t) => error ("Not a type variable: " ^ Syntax.string_of_typ ctxt t)); fun internalize_sort_attr tvar = Thm.rule_attribute [] (fn context => fn thm => (snd (internalize_sort (Thm.ctyp_of (Context.proof_of context) tvar) thm))); val _ = Context.>> (Context.map_theory (Attrib.setup \<^binding>\internalize_sort\ (tvar >> internalize_sort_attr) "internalize a sort")); end; diff --git a/src/Pure/Thy/export_theory.ML b/src/Pure/Thy/export_theory.ML --- a/src/Pure/Thy/export_theory.ML +++ b/src/Pure/Thy/export_theory.ML @@ -1,480 +1,480 @@ (* Title: Pure/Thy/export_theory.ML Author: Makarius Export foundational theory content and locale/class structure. *) signature EXPORT_THEORY = sig val other_name_space: (theory -> Name_Space.T) -> theory -> theory val export_enabled: Thy_Info.presentation_context -> bool val export_body: theory -> string -> XML.body -> unit end; structure Export_Theory: EXPORT_THEORY = struct (* other name spaces *) fun err_dup_kind kind = error ("Duplicate name space kind " ^ quote kind); structure Data = Theory_Data ( type T = (theory -> Name_Space.T) Inttab.table; val empty = Inttab.empty; val merge = Inttab.merge (K true); ); val other_name_spaces = map #2 o Inttab.dest o Data.get; fun other_name_space get_space thy = Data.map (Inttab.update (serial (), get_space)) thy; val _ = Theory.setup (other_name_space Thm.oracle_space #> other_name_space Global_Theory.fact_space #> other_name_space (Bundle.bundle_space o Context.Theory) #> other_name_space (Attrib.attribute_space o Context.Theory) #> other_name_space (Method.method_space o Context.Theory)); (* approximative syntax *) val get_syntax = Syntax.get_approx o Proof_Context.syn_of; fun get_syntax_type ctxt = get_syntax ctxt o Lexicon.mark_type; fun get_syntax_const ctxt = get_syntax ctxt o Lexicon.mark_const; fun get_syntax_fixed ctxt = get_syntax ctxt o Lexicon.mark_fixed; fun get_syntax_param ctxt loc x = let val thy = Proof_Context.theory_of ctxt in if Class.is_class thy loc then (case AList.lookup (op =) (Class.these_params thy [loc]) x of NONE => NONE | SOME (_, (c, _)) => get_syntax_const ctxt c) else get_syntax_fixed ctxt x end; val encode_syntax = XML.Encode.variant [fn NONE => ([], []), fn SOME (Syntax.Prefix delim) => ([delim], []), fn SOME (Syntax.Infix {assoc, delim, pri}) => let val ass = (case assoc of Printer.No_Assoc => 0 | Printer.Left_Assoc => 1 | Printer.Right_Assoc => 2); open XML.Encode Term_XML.Encode; in ([], triple int string int (ass, delim, pri)) end]; (* free variables: not declared in the context *) val is_free = not oo Name.is_declared; fun add_frees used = fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); fun add_tfrees used = (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); (* locales *) fun locale_content thy loc = let val ctxt = Locale.init loc thy; val args = Locale.params_of thy loc |> map (fn ((x, T), _) => ((x, T), get_syntax_param ctxt loc x)); val axioms = let val (asm, defs) = Locale.specification_of thy loc; val cprops = map (Thm.cterm_of ctxt) (the_list asm @ defs); val (intro1, intro2) = Locale.intros_of thy loc; val intros_tac = Method.try_intros_tac ctxt (the_list intro1 @ the_list intro2) []; val res = Goal.init (Conjunction.mk_conjunction_balanced cprops) |> (ALLGOALS Goal.conjunction_tac THEN intros_tac) |> try Seq.hd; in (case res of SOME goal => Thm.prems_of goal | NONE => raise Fail ("Cannot unfold locale " ^ quote loc)) end; val typargs = build_rev (fold Term.add_tfrees (map (Free o #1) args @ axioms)); in {typargs = typargs, args = args, axioms = axioms} end; fun get_locales thy = Locale.get_locales thy |> map_filter (fn loc => if Experiment.is_experiment thy loc then NONE else SOME (loc, ())); fun get_dependencies prev_thys thy = Locale.dest_dependencies prev_thys thy |> map_filter (fn dep => if Experiment.is_experiment thy (#source dep) orelse Experiment.is_experiment thy (#target dep) then NONE else let val (type_params, params) = Locale.parameters_of thy (#source dep); val typargs = fold (Term.add_tfreesT o #2 o #1) params type_params; val substT = typargs |> map_filter (fn v => let val T = TFree v; val T' = Morphism.typ (#morphism dep) T; in if T = T' then NONE else SOME (v, T') end); val subst = params |> map_filter (fn (v, _) => let val t = Free v; val t' = Morphism.term (#morphism dep) t; in if t aconv t' then NONE else SOME (v, t') end); in SOME (dep, (substT, subst)) end); (* presentation *) fun export_enabled (context: Thy_Info.presentation_context) = Options.bool (#options context) "export_theory"; fun export_body thy name body = if XML.is_empty_body body then () else Export.export thy (Path.binding0 (Path.make ("theory" :: space_explode "/" name))) body; val _ = (Theory.setup o Thy_Info.add_presentation) (fn context => fn thy => let val rep_tsig = Type.rep_tsig (Sign.tsig_of thy); val consts = Sign.consts_of thy; val thy_ctxt = Proof_Context.init_global thy; val pos_properties = Thy_Info.adjust_pos_properties context; val enabled = export_enabled context; (* strict parents *) val parents = Theory.parents_of thy; val _ = Export.export thy \<^path_binding>\theory/parents\ (XML.Encode.string (cat_lines (map Context.theory_long_name parents) ^ "\n")); (* spec rules *) fun spec_rule_content {pos, name, rough_classification, terms, rules} = let val spec = terms @ map Thm.plain_prop_of rules |> Term_Subst.zero_var_indexes_list |> map Logic.unvarify_global; in {props = pos_properties pos, name = name, rough_classification = rough_classification, typargs = build_rev (fold Term.add_tfrees spec), args = build_rev (fold Term.add_frees spec), terms = map (fn t => (t, Term.type_of t)) (take (length terms) spec), rules = drop (length terms) spec} end; (* entities *) fun make_entity_markup name xname pos serial = let val props = pos_properties pos @ Markup.serial_properties serial; in (Markup.entityN, (Markup.nameN, name) :: (Markup.xnameN, xname) :: props) end; fun entity_markup space name = let val xname = Name_Space.extern_shortest thy_ctxt space name; val {serial, pos, ...} = Name_Space.the_entry space name; in make_entity_markup name xname pos serial end; fun export_entities export_name get_space decls export = let val parent_spaces = map get_space parents; val space = get_space thy; in build (decls |> fold (fn (name, decl) => if exists (fn space => Name_Space.declared space name) parent_spaces then I else (case export name decl of NONE => I | SOME make_body => let val i = #serial (Name_Space.the_entry space name); val body = if enabled then make_body () else []; in cons (i, XML.Elem (entity_markup space name, body)) end))) |> sort (int_ord o apply2 #1) |> map #2 |> export_body thy export_name end; (* types *) val encode_type = let open XML.Encode Term_XML.Encode in triple encode_syntax (list string) (option typ) end; val _ = export_entities "types" Sign.type_space (Name_Space.dest_table (#types rep_tsig)) (fn c => (fn Type.LogicalType n => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, Name.invent Name.context Name.aT n, NONE)) | Type.Abbreviation (args, U, false) => SOME (fn () => encode_type (get_syntax_type thy_ctxt c, args, SOME U)) | _ => NONE)); (* consts *) val encode_term = Term_XML.Encode.term consts; val encode_const = let open XML.Encode Term_XML.Encode in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end; val _ = export_entities "consts" Sign.const_space (#constants (Consts.dest consts)) (fn c => fn (T, abbrev) => SOME (fn () => let val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Type.strip_sorts U; val trim_abbrev = Proofterm.standard_vars_term Name.context #> map_types Type.strip_sorts; val abbrev' = Option.map trim_abbrev abbrev; val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); in encode_const (syntax, (args, (U0, (abbrev', propositional)))) end)); (* axioms *) fun standard_prop used extra_shyps raw_prop raw_proof = let val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); val args = rev (add_frees used prop []); val typargs = rev (add_tfrees used prop []); val used_typargs = fold (Name.declare o #1) typargs used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; in ((sorts @ typargs, args, prop), proof) end; fun standard_prop_of thm = - standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm); + standard_prop Name.context (Sortset.dest (Thm.extra_shyps thm)) (Thm.full_prop_of thm); val encode_prop = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair string typ)) encode_term end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); val _ = export_entities "axioms" Theory.axiom_space (Theory.all_axioms_of thy) (fn _ => fn prop => SOME (fn () => encode_axiom Name.context prop)); (* theorems and proof terms *) val clean_thm = Thm.check_hyps (Context.Theory thy) #> Thm.strip_shyps; val prep_thm = clean_thm #> Thm.unconstrainT #> Thm.strip_shyps; val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = if #serial header = #serial thm_id then "" else (case lookup_thm_id (Proofterm.thm_header_id header) of NONE => "" | SOME thm_name => Thm_Name.print thm_name); fun entity_markup_thm (serial, (name, i)) = let val space = Global_Theory.fact_space thy; val xname = Name_Space.extern_shortest thy_ctxt space name; val {pos, ...} = Name_Space.the_entry space name; in make_entity_markup (Thm_Name.print (name, i)) (Thm_Name.print (xname, i)) pos serial end; fun encode_thm thm_id raw_thm = let val deps = map (Thm_Name.print o #2) (Thm_Deps.thm_deps thy [raw_thm]); val thm = prep_thm raw_thm; val proof0 = if Proofterm.export_standard_enabled () then Proof_Syntax.standard_proof_of {full = true, expand_name = SOME o expand_name thm_id} thm else if Proofterm.export_enabled () then Thm.reconstruct_proof_of thm else MinProof; val (prop, SOME proof) = standard_prop_of thm (SOME proof0); val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> let open XML.Encode Term_XML.Encode; val encode_proof = Proofterm.encode_standard_proof consts; in triple encode_prop (list string) encode_proof end end; fun export_thm (thm_id, thm_name) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val body = if enabled then Global_Theory.get_thm_name thy (thm_name, Position.none) |> encode_thm thm_id else []; in XML.Elem (markup, body) end; val _ = export_body thy "thms" (map export_thm (Global_Theory.dest_thm_names thy)); (* type classes *) val encode_class = let open XML.Encode Term_XML.Encode in pair (list (pair string typ)) (list (encode_axiom Name.context)) end; val _ = export_entities "classes" Sign.class_space (map (rpair ()) (Graph.keys (Sorts.classes_of (#2 (#classes rep_tsig))))) (fn name => fn () => SOME (fn () => (case try (Axclass.get_info thy) name of NONE => ([], []) | SOME {params, axioms, ...} => (params, map (Thm.plain_prop_of o clean_thm) axioms)) |> encode_class)); (* sort algebra *) val _ = if enabled then let val prop = encode_axiom Name.context o Logic.varify_global; val encode_classrel = let open XML.Encode in list (pair prop (pair string string)) end; val encode_arities = let open XML.Encode Term_XML.Encode in list (pair prop (triple string (list sort) string)) end; val export_classrel = maps (fn (c, cs) => map (pair c) cs) #> map (`Logic.mk_classrel) #> encode_classrel; val export_arities = map (`Logic.mk_arity) #> encode_arities; val {classrel, arities} = Sorts.dest_algebra (map (#2 o #classes o Type.rep_tsig o Sign.tsig_of) parents) (#2 (#classes rep_tsig)); in if null classrel then () else export_body thy "classrel" (export_classrel classrel); if null arities then () else export_body thy "arities" (export_arities arities) end else (); (* locales *) fun encode_locale used = let open XML.Encode Term_XML.Encode in triple (list (pair string sort)) (list (pair (pair string typ) encode_syntax)) (list (encode_axiom used)) end; val _ = export_entities "locales" Locale.locale_space (get_locales thy) (fn loc => fn () => SOME (fn () => let val {typargs, args, axioms} = locale_content thy loc; val used = fold Name.declare (map #1 typargs @ map (#1 o #1) args) Name.context; in encode_locale used (typargs, args, axioms) end handle ERROR msg => cat_error msg ("The error(s) above occurred in locale " ^ quote (Locale.markup_name thy_ctxt loc)))); (* locale dependencies *) fun encode_locale_dependency (dep: Locale.locale_dependency, subst) = (#source dep, (#target dep, (#prefix dep, subst))) |> let open XML.Encode Term_XML.Encode; val encode_subst = pair (list (pair (pair string sort) typ)) (list (pair (pair string typ) (term consts))); in pair string (pair string (pair (list (pair string bool)) encode_subst)) end; val _ = if enabled then get_dependencies parents thy |> map_index (fn (i, dep) => let val xname = string_of_int (i + 1); val name = Long_Name.implode [Context.theory_name thy, xname]; val markup = make_entity_markup name xname (#pos (#1 dep)) (#serial (#1 dep)); val body = encode_locale_dependency dep; in XML.Elem (markup, body) end) |> export_body thy "locale_dependencies" else (); (* constdefs *) val _ = if enabled then let val constdefs = Defs.dest_constdefs (map Theory.defs_of (Theory.parents_of thy)) (Theory.defs_of thy) |> sort_by #1; val encode = let open XML.Encode in list (pair string string) end; in if null constdefs then () else export_body thy "constdefs" (encode constdefs) end else (); (* spec rules *) val encode_specs = let open XML.Encode Term_XML.Encode in list (fn {props, name, rough_classification, typargs, args, terms, rules} => pair properties (pair string (pair Spec_Rules.encode_rough_classification (pair (list (pair string sort)) (pair (list (pair string typ)) (pair (list (pair encode_term typ)) (list encode_term)))))) (props, (name, (rough_classification, (typargs, (args, (terms, rules))))))) end; val _ = if enabled then (case Spec_Rules.dest_theory thy of [] => () | spec_rules => export_body thy "spec_rules" (encode_specs (map spec_rule_content spec_rules))) else (); (* other entities *) fun export_other get_space = let val space = get_space thy; val export_name = "other/" ^ Name_Space.kind_of space; val decls = Name_Space.get_names space |> map (rpair ()); in export_entities export_name get_space decls (fn _ => fn () => SOME (K [])) end; val other_spaces = other_name_spaces thy; val other_kinds = map (fn get_space => Name_Space.kind_of (get_space thy)) other_spaces; val _ = if null other_kinds then () else Export.export thy \<^path_binding>\theory/other_kinds\ (XML.Encode.string (cat_lines other_kinds)); val _ = List.app export_other other_spaces; in () end); end; diff --git a/src/Pure/envir.ML b/src/Pure/envir.ML --- a/src/Pure/envir.ML +++ b/src/Pure/envir.ML @@ -1,424 +1,424 @@ (* Title: Pure/envir.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Free-form environments. The type of a term variable / sort of a type variable is part of its name. The lookup function must apply type substitutions, since they may change the identity of a variable. *) signature ENVIR = sig type tenv = (typ * term) Vartab.table datatype env = Envir of {maxidx: int, tenv: tenv, tyenv: Type.tyenv} val maxidx_of: env -> int val term_env: env -> tenv val type_env: env -> Type.tyenv val is_empty: env -> bool val empty: int -> env val init: env val merge: env * env -> env - val insert_sorts: env -> sort list -> sort list + val insert_sorts: env -> Sortset.T -> Sortset.T val genvars: string -> env * typ list -> env * term list val genvar: string -> env * typ -> env * term val lookup1: tenv -> indexname * typ -> term option val lookup: env -> indexname * typ -> term option val update: (indexname * typ) * term -> env -> env val above: env -> int -> bool val vupdate: (indexname * typ) * term -> env -> env val norm_type_same: Type.tyenv -> typ Same.operation val norm_types_same: Type.tyenv -> typ list Same.operation val norm_type: Type.tyenv -> typ -> typ val norm_term_same: env -> term Same.operation val norm_term: env -> term -> term val beta_norm: term -> term val head_norm: env -> term -> term val eta_long: typ list -> term -> term val eta_contract: term -> term val beta_eta_contract: term -> term val aeconv: term * term -> bool val body_type: env -> typ -> typ val binder_types: env -> typ -> typ list val strip_type: env -> typ -> typ list * typ val fastype: env -> typ list -> term -> typ val subst_type_same: Type.tyenv -> typ Same.operation val subst_term_types_same: Type.tyenv -> term Same.operation val subst_term_same: Type.tyenv * tenv -> term Same.operation val subst_type: Type.tyenv -> typ -> typ val subst_term_types: Type.tyenv -> term -> term val subst_term: Type.tyenv * tenv -> term -> term val expand_atom: typ -> typ * term -> term val expand_term: (term -> (typ * term) option) -> term -> term val expand_term_defs: (term -> string * typ) -> ((string * typ) * term) list -> term -> term end; structure Envir: ENVIR = struct (** datatype env **) (*Updating can destroy environment in 2 ways! (1) variables out of range (2) circular assignments *) type tenv = (typ * term) Vartab.table; datatype env = Envir of {maxidx: int, (*upper bound of maximum index of vars*) tenv: tenv, (*assignments to Vars*) tyenv: Type.tyenv}; (*assignments to TVars*) fun make_env (maxidx, tenv, tyenv) = Envir {maxidx = maxidx, tenv = tenv, tyenv = tyenv}; fun maxidx_of (Envir {maxidx, ...}) = maxidx; fun term_env (Envir {tenv, ...}) = tenv; fun type_env (Envir {tyenv, ...}) = tyenv; fun is_empty env = Vartab.is_empty (term_env env) andalso Vartab.is_empty (type_env env); (* build env *) fun empty maxidx = make_env (maxidx, Vartab.empty, Vartab.empty); val init = empty ~1; fun merge (Envir {maxidx = maxidx1, tenv = tenv1, tyenv = tyenv1}, Envir {maxidx = maxidx2, tenv = tenv2, tyenv = tyenv2}) = make_env (Int.max (maxidx1, maxidx2), Vartab.merge (op =) (tenv1, tenv2), Vartab.merge (op =) (tyenv1, tyenv2)); (*NB: type unification may invent new sorts*) (* FIXME tenv!? *) -val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sorts.insert_typ T) o type_env; +val insert_sorts = Vartab.fold (fn (_, (_, T)) => Sortset.insert_typ T) o type_env; (*Generate a list of distinct variables. Increments index to make them distinct from ALL present variables. *) fun genvars name (Envir {maxidx, tenv, tyenv}, Ts) : env * term list = let fun genvs (_, [] : typ list) : term list = [] | genvs (_, [T]) = [Var ((name, maxidx + 1), T)] | genvs (n, T :: Ts) = Var ((name ^ radixstring (26, "a" , n), maxidx + 1), T) :: genvs (n + 1, Ts); in (Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}, genvs (0, Ts)) end; (*Generate a variable.*) fun genvar name (env, T) : env * term = let val (env', [v]) = genvars name (env, [T]) in (env', v) end; fun var_clash xi T T' = raise TYPE ("Variable has two distinct types", [], [Var (xi, T'), Var (xi, T)]); fun lookup_check check tenv (xi, T) = (case Vartab.lookup tenv xi of NONE => NONE | SOME (U, t) => if check (T, U) then SOME t else var_clash xi T U); (*When dealing with environments produced by matching instead of unification, there is no need to chase assigned TVars. In this case, we can simply ignore the type substitution and use = instead of eq_type.*) fun lookup1 tenv = lookup_check (op =) tenv; fun lookup (Envir {tenv, tyenv, ...}) = lookup_check (Type.unified tyenv) tenv; fun update ((xi, T), t) (Envir {maxidx, tenv, tyenv}) = Envir {maxidx = maxidx, tenv = Vartab.update_new (xi, (T, t)) tenv, tyenv = tyenv}; (*Determine if the least index updated exceeds lim*) fun above (Envir {tenv, tyenv, ...}) lim = (case Vartab.min tenv of SOME ((_, i), _) => i > lim | NONE => true) andalso (case Vartab.min tyenv of SOME ((_, i), _) => i > lim | NONE => true); (*Update, checking Var-Var assignments: try to suppress higher indexes*) fun vupdate ((a, U), t) env = (case t of Var (nT as (name', T)) => if a = name' then env (*cycle!*) else if Term_Ord.indexname_ord (a, name') = LESS then (case lookup env nT of (*if already assigned, chase*) NONE => update (nT, Var (a, T)) env | SOME u => vupdate ((a, U), u) env) else update ((a, U), t) env | _ => update ((a, U), t) env); (** beta normalization wrt. environment **) (*Chases variables in env. Does not exploit sharing of variable bindings Does not check types, so could loop.*) local fun norm_type0 tyenv : typ Same.operation = let fun norm (Type (a, Ts)) = Type (a, Same.map norm Ts) | norm (TFree _) = raise Same.SAME | norm (TVar v) = (case Type.lookup tyenv v of SOME U => Same.commit norm U | NONE => raise Same.SAME); in norm end; fun norm_term1 tenv : term Same.operation = let fun norm (Var v) = (case lookup1 tenv v of SOME u => Same.commit norm u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = ((case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ Same.commit norm t) handle Same.SAME => f $ norm t) | norm _ = raise Same.SAME; in norm end; fun norm_term2 (envir as Envir {tyenv, ...}) : term Same.operation = let val normT = norm_type0 tyenv; fun norm (Const (a, T)) = Const (a, normT T) | norm (Free (a, T)) = Free (a, normT T) | norm (Var (xi, T)) = (case lookup envir (xi, T) of SOME u => Same.commit norm u | NONE => Var (xi, normT T)) | norm (Abs (a, T, body)) = (Abs (a, normT T, Same.commit norm body) handle Same.SAME => Abs (a, T, norm body)) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = ((case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ Same.commit norm t) handle Same.SAME => f $ norm t) | norm _ = raise Same.SAME; in norm end; in fun norm_type_same tyenv T = if Vartab.is_empty tyenv then raise Same.SAME else norm_type0 tyenv T; fun norm_types_same tyenv Ts = if Vartab.is_empty tyenv then raise Same.SAME else Same.map (norm_type0 tyenv) Ts; fun norm_type tyenv T = norm_type_same tyenv T handle Same.SAME => T; fun norm_term_same (envir as Envir {tenv, tyenv, ...}) = if Vartab.is_empty tyenv then norm_term1 tenv else norm_term2 envir; fun norm_term envir t = norm_term_same envir t handle Same.SAME => t; fun beta_norm t = if Term.could_beta_contract t then norm_term init t else t; end; (* head normal form for unification *) fun head_norm env = let fun norm (Var v) = (case lookup env v of SOME u => head_norm env u | NONE => raise Same.SAME) | norm (Abs (a, T, body)) = Abs (a, T, norm body) | norm (Abs (_, _, body) $ t) = Same.commit norm (subst_bound (t, body)) | norm (f $ t) = (case norm f of Abs (_, _, body) => Same.commit norm (subst_bound (t, body)) | nf => nf $ t) | norm _ = raise Same.SAME; in Same.commit norm end; (* eta-long beta-normal form *) fun eta_long Ts (Abs (s, T, t)) = Abs (s, T, eta_long (T :: Ts) t) | eta_long Ts t = (case strip_comb t of (Abs _, _) => eta_long Ts (beta_norm t) | (u, ts) => let val Us = binder_types (fastype_of1 (Ts, t)); val i = length Us; val long = eta_long (rev Us @ Ts); in fold_rev (Term.abs o pair "x") Us (list_comb (incr_boundvars i u, map (long o incr_boundvars i) ts @ map (long o Bound) (i - 1 downto 0))) end); (* full eta contraction *) local fun decr lev (Bound i) = if i >= lev then Bound (i - 1) else raise Same.SAME | decr lev (Abs (a, T, body)) = Abs (a, T, decr (lev + 1) body) | decr lev (t $ u) = (decr lev t $ decrh lev u handle Same.SAME => t $ decr lev u) | decr _ _ = raise Same.SAME and decrh lev t = (decr lev t handle Same.SAME => t); fun eta (Abs (a, T, body)) = ((case eta body of body' as (f $ Bound 0) => if Term.is_dependent f then Abs (a, T, body') else decrh 0 f | body' => Abs (a, T, body')) handle Same.SAME => (case body of f $ Bound 0 => if Term.is_dependent f then raise Same.SAME else decrh 0 f | _ => raise Same.SAME)) | eta (t $ u) = (eta t $ Same.commit eta u handle Same.SAME => t $ eta u) | eta _ = raise Same.SAME; in fun eta_contract t = if Term.could_eta_contract t then Same.commit eta t else t; end; val beta_eta_contract = eta_contract o beta_norm; fun aeconv (t, u) = t aconv u orelse eta_contract t aconv eta_contract u; fun body_type env (Type ("fun", [_, T])) = body_type env T | body_type env (T as TVar v) = (case Type.lookup (type_env env) v of NONE => T | SOME T' => body_type env T') | body_type _ T = T; fun binder_types env (Type ("fun", [T, U])) = T :: binder_types env U | binder_types env (TVar v) = (case Type.lookup (type_env env) v of NONE => [] | SOME T' => binder_types env T') | binder_types _ _ = []; fun strip_type env T = (binder_types env T, body_type env T); (*finds type of term without checking that combinations are consistent Ts holds types of bound variables*) fun fastype (Envir {tyenv, ...}) = let val funerr = "fastype: expected function type"; fun fast Ts (f $ u) = (case Type.devar tyenv (fast Ts f) of Type ("fun", [_, T]) => T | TVar _ => raise TERM (funerr, [f $ u]) | _ => raise TERM (funerr, [f $ u])) | fast _ (Const (_, T)) = T | fast _ (Free (_, T)) = T | fast Ts (Bound i) = (nth Ts i handle General.Subscript => raise TERM ("fastype: Bound", [Bound i])) | fast _ (Var (_, T)) = T | fast Ts (Abs (_, T, u)) = T --> fast (T :: Ts) u; in fast end; (** plain substitution -- without variable chasing **) local fun subst_type0 tyenv = Term_Subst.map_atypsT_same (fn TVar v => (case Type.lookup tyenv v of SOME U => U | NONE => raise Same.SAME) | _ => raise Same.SAME); fun subst_term1 tenv = Term_Subst.map_aterms_same (fn Var v => (case lookup1 tenv v of SOME u => u | NONE => raise Same.SAME) | _ => raise Same.SAME); fun subst_term2 tenv tyenv : term Same.operation = let val substT = subst_type0 tyenv; fun subst (Const (a, T)) = Const (a, substT T) | subst (Free (a, T)) = Free (a, substT T) | subst (Var (xi, T)) = (case lookup1 tenv (xi, T) of SOME u => u | NONE => Var (xi, substT T)) | subst (Bound _) = raise Same.SAME | subst (Abs (a, T, t)) = (Abs (a, substT T, Same.commit subst t) handle Same.SAME => Abs (a, T, subst t)) | subst (t $ u) = (subst t $ Same.commit subst u handle Same.SAME => t $ subst u); in subst end; in fun subst_type_same tyenv T = if Vartab.is_empty tyenv then raise Same.SAME else subst_type0 tyenv T; fun subst_term_types_same tyenv t = if Vartab.is_empty tyenv then raise Same.SAME else Term_Subst.map_types_same (subst_type0 tyenv) t; fun subst_term_same (tyenv, tenv) = if Vartab.is_empty tenv then subst_term_types_same tyenv else if Vartab.is_empty tyenv then subst_term1 tenv else subst_term2 tenv tyenv; fun subst_type tyenv T = subst_type_same tyenv T handle Same.SAME => T; fun subst_term_types tyenv t = subst_term_types_same tyenv t handle Same.SAME => t; fun subst_term envs t = subst_term_same envs t handle Same.SAME => t; end; (** expand defined atoms -- with local beta reduction **) fun expand_atom T (U, u) = subst_term_types (Vartab.build (Type.raw_match (U, T))) u handle Type.TYPE_MATCH => raise TYPE ("expand_atom: ill-typed replacement", [T, U], [u]); fun expand_term get = let fun expand tm = let val (head, args) = Term.strip_comb tm; val args' = map expand args; fun comb head' = Term.list_comb (head', args'); in (case head of Abs (x, T, t) => comb (Abs (x, T, expand t)) | _ => (case get head of SOME def => Term.betapplys (expand_atom (Term.fastype_of head) def, args') | NONE => comb head)) end; in expand end; fun expand_term_defs dest defs = let val eqs = map (fn ((x, U), u) => (x: string, (U, u))) defs; fun get t = (case try dest t of SOME (x, _: typ) => AList.lookup (op =) eqs x | _ => NONE); in expand_term get end; end; diff --git a/src/Pure/logic.ML b/src/Pure/logic.ML --- a/src/Pure/logic.ML +++ b/src/Pure/logic.ML @@ -1,671 +1,671 @@ (* Title: Pure/logic.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius Abstract syntax operations of the Pure meta-logic. *) signature LOGIC = sig val all_const: typ -> term val all: term -> term -> term val dependent_all_name: string * term -> term -> term val is_all: term -> bool val dest_all_global: term -> (string * typ) * term val list_all: (string * typ) list * term -> term val all_constraint: (string -> typ option) -> string * string -> term -> term val dependent_all_constraint: (string -> typ option) -> string * string -> term -> term val mk_equals: term * term -> term val dest_equals: term -> term * term val implies: term val mk_implies: term * term -> term val dest_implies: term -> term * term val list_implies: term list * term -> term val strip_imp_prems: term -> term list val strip_imp_concl: term -> term val strip_prems: int * term list * term -> term list * term val count_prems: term -> int val no_prems: term -> bool val nth_prem: int * term -> term val close_term: (string * term) list -> term -> term val close_prop: (string * term) list -> term list -> term -> term val close_prop_constraint: (string -> typ option) -> (string * string) list -> term list -> term -> term val true_prop: term val conjunction: term val mk_conjunction: term * term -> term val mk_conjunction_list: term list -> term val mk_conjunction_balanced: term list -> term val dest_conjunction: term -> term * term val dest_conjunction_list: term -> term list val dest_conjunction_balanced: int -> term -> term list val dest_conjunctions: term -> term list val strip_horn: term -> term list * term val mk_type: typ -> term val dest_type: term -> typ val type_map: (term -> term) -> typ -> typ val class_suffix: string val const_of_class: class -> string val class_of_const: string -> class val mk_of_class: typ * class -> term val dest_of_class: term -> typ * class val mk_of_sort: typ * sort -> term list val name_classrel: string * string -> string val mk_classrel: class * class -> term val dest_classrel: term -> class * class val name_arities: arity -> string list val name_arity: string * sort list * class -> string val mk_arities: arity -> term list val mk_arity: string * sort list * class -> term val dest_arity: term -> string * sort list * class val dummy_tfree: sort -> typ type unconstrain_context = {present_map: (typ * typ) list, constraints_map: (sort * typ) list, atyp_map: typ -> typ, map_atyps: typ -> typ, constraints: ((typ * class) * term) list, outer_constraints: (typ * class) list}; - val unconstrainT: sort list -> term -> unconstrain_context * term + val unconstrainT: Sortset.T -> term -> unconstrain_context * term val protectC: term val protect: term -> term val unprotect: term -> term val mk_term: term -> term val dest_term: term -> term val occs: term * term -> bool val close_form: term -> term val combound: term * int * int -> term val rlist_abs: (string * typ) list * term -> term val incr_tvar_same: int -> typ Same.operation val incr_tvar: int -> typ -> typ val incr_indexes_same: string list * typ list * int -> term Same.operation val incr_indexes: string list * typ list * int -> term -> term val lift_abs: int -> term -> term -> term val lift_all: int -> term -> term -> term val strip_assums_hyp: term -> term list val strip_assums_concl: term -> term val strip_params: term -> (string * typ) list val has_meta_prems: term -> bool val flatten_params: int -> term -> term val list_rename_params: string list -> term -> term val assum_pairs: int * term -> (term * term) list val assum_problems: int * term -> (term -> term) * term list * term val bad_schematic: indexname -> string val bad_fixed: string -> string val varifyT_global: typ -> typ val unvarifyT_global: typ -> typ val varify_types_global: term -> term val unvarify_types_global: term -> term val varify_global: term -> term val unvarify_global: term -> term val get_goal: term -> int -> term val goal_params: term -> int -> term * term list val prems_of_goal: term -> int -> term list val concl_of_goal: term -> int -> term end; structure Logic : LOGIC = struct (*** Abstract syntax operations on the meta-connectives ***) (** all **) fun all_const T = Const ("Pure.all", (T --> propT) --> propT); fun all v t = all_const (Term.fastype_of v) $ lambda v t; fun dependent_all_name (x, v) t = let val x' = if x = "" then Term.term_name v else x; val T = Term.fastype_of v; val t' = Term.abstract_over (v, t); in if Term.is_dependent t' then all_const T $ Abs (x', T, t') else t end; fun is_all (Const ("Pure.all", _) $ Abs _) = true | is_all _ = false; fun dest_all_global t = (case t of Const ("Pure.all", _) $ (u as Abs _) => Term.dest_abs_global u | _ => raise TERM ("dest_all", [t])); fun list_all ([], t) = t | list_all ((a, T) :: vars, t) = all_const T $ Abs (a, T, list_all (vars, t)); (* operations before type-inference *) local fun abs_body default_type z tm = let fun abs lev (Abs (x, T, b)) = Abs (x, T, abs (lev + 1) b) | abs lev (t $ u) = abs lev t $ abs lev u | abs lev (a as Free (x, T)) = if x = z then Type.constraint (the_default dummyT (default_type x)) (Type.constraint T (Bound lev)) else a | abs _ a = a; in abs 0 (Term.incr_boundvars 1 tm) end; in fun all_constraint default_type (y, z) t = all_const dummyT $ Abs (y, dummyT, abs_body default_type z t); fun dependent_all_constraint default_type (y, z) t = let val t' = abs_body default_type z t in if Term.is_dependent t' then all_const dummyT $ Abs (y, dummyT, t') else t end; end; (** equality **) fun mk_equals (t, u) = let val T = Term.fastype_of t in Const ("Pure.eq", T --> T --> propT) $ t $ u end; fun dest_equals (Const ("Pure.eq", _) $ t $ u) = (t, u) | dest_equals t = raise TERM ("dest_equals", [t]); (** implies **) val implies = Const ("Pure.imp", propT --> propT --> propT); fun mk_implies (A, B) = implies $ A $ B; fun dest_implies (Const ("Pure.imp", _) $ A $ B) = (A, B) | dest_implies A = raise TERM ("dest_implies", [A]); (** nested implications **) (* [A1,...,An], B goes to A1\...An\B *) fun list_implies ([], B) = B | list_implies (A::As, B) = implies $ A $ list_implies(As,B); (* A1\...An\B goes to [A1,...,An], where B is not an implication *) fun strip_imp_prems (Const("Pure.imp", _) $ A $ B) = A :: strip_imp_prems B | strip_imp_prems _ = []; (* A1\...An\B goes to B, where B is not an implication *) fun strip_imp_concl (Const("Pure.imp", _) $ A $ B) = strip_imp_concl B | strip_imp_concl A = A : term; (*Strip and return premises: (i, [], A1\...Ai\B) goes to ([Ai, A(i-1),...,A1] , B) (REVERSED) if i<0 or else i too big then raises TERM*) fun strip_prems (0, As, B) = (As, B) | strip_prems (i, As, Const("Pure.imp", _) $ A $ B) = strip_prems (i-1, A::As, B) | strip_prems (_, As, A) = raise TERM("strip_prems", A::As); (*Count premises -- quicker than (length o strip_prems) *) fun count_prems (Const ("Pure.imp", _) $ _ $ B) = 1 + count_prems B | count_prems _ = 0; fun no_prems (Const ("Pure.imp", _) $ _ $ _) = false | no_prems _ = true; (*Select Ai from A1\...Ai\B*) fun nth_prem (1, Const ("Pure.imp", _) $ A $ _) = A | nth_prem (i, Const ("Pure.imp", _) $ _ $ B) = nth_prem (i - 1, B) | nth_prem (_, A) = raise TERM ("nth_prem", [A]); (*strip a proof state (Horn clause): B1 \ ... Bn \ C goes to ([B1, ..., Bn], C) *) fun strip_horn A = (strip_imp_prems A, strip_imp_concl A); (* close -- omit vacuous quantifiers *) val close_term = fold_rev Term.dependent_lambda_name; fun close_prop xs As B = fold_rev dependent_all_name xs (list_implies (As, B)); fun close_prop_constraint default_type xs As B = fold_rev (dependent_all_constraint default_type) xs (list_implies (As, B)); (** conjunction **) val true_prop = all_const propT $ Abs ("dummy", propT, mk_implies (Bound 0, Bound 0)); val conjunction = Const ("Pure.conjunction", propT --> propT --> propT); (*A &&& B*) fun mk_conjunction (A, B) = conjunction $ A $ B; (*A &&& B &&& C -- improper*) fun mk_conjunction_list [] = true_prop | mk_conjunction_list ts = foldr1 mk_conjunction ts; (*(A &&& B) &&& (C &&& D) -- balanced*) fun mk_conjunction_balanced [] = true_prop | mk_conjunction_balanced ts = Balanced_Tree.make mk_conjunction ts; (*A &&& B*) fun dest_conjunction (Const ("Pure.conjunction", _) $ A $ B) = (A, B) | dest_conjunction t = raise TERM ("dest_conjunction", [t]); (*A &&& B &&& C -- improper*) fun dest_conjunction_list t = (case try dest_conjunction t of NONE => [t] | SOME (A, B) => A :: dest_conjunction_list B); (*(A &&& B) &&& (C &&& D) -- balanced*) fun dest_conjunction_balanced 0 _ = [] | dest_conjunction_balanced n t = Balanced_Tree.dest dest_conjunction n t; (*((A &&& B) &&& C) &&& D &&& E -- flat*) fun dest_conjunctions t = (case try dest_conjunction t of NONE => [t] | SOME (A, B) => dest_conjunctions A @ dest_conjunctions B); (** types as terms **) fun mk_type ty = Const ("Pure.type", Term.itselfT ty); fun dest_type (Const ("Pure.type", Type ("itself", [ty]))) = ty | dest_type t = raise TERM ("dest_type", [t]); fun type_map f = dest_type o f o mk_type; (** type classes **) (* const names *) val class_suffix = "_class"; val const_of_class = suffix class_suffix; fun class_of_const c = unsuffix class_suffix c handle Fail _ => raise TERM ("class_of_const: bad name " ^ quote c, []); (* class/sort membership *) fun mk_of_class (ty, c) = Const (const_of_class c, Term.itselfT ty --> propT) $ mk_type ty; fun dest_of_class (Const (c_class, _) $ ty) = (dest_type ty, class_of_const c_class) | dest_of_class t = raise TERM ("dest_of_class", [t]); fun mk_of_sort (ty, S) = map (fn c => mk_of_class (ty, c)) S; (* class relations *) fun name_classrel (c1, c2) = Long_Name.base_name c1 ^ "_" ^ Long_Name.base_name c2; fun mk_classrel (c1, c2) = mk_of_class (Term.aT [c1], c2); fun dest_classrel tm = (case dest_of_class tm of (TVar (_, [c1]), c2) => (c1, c2) | _ => raise TERM ("dest_classrel", [tm])); (* type arities *) fun name_arities (t, _, S) = let val b = Long_Name.base_name t in S |> map (fn c => Long_Name.base_name c ^ "_" ^ b) end; fun name_arity (t, dom, c) = hd (name_arities (t, dom, [c])); fun mk_arities (t, Ss, S) = let val T = Type (t, ListPair.map TFree (Name.invent Name.context Name.aT (length Ss), Ss)) in map (fn c => mk_of_class (T, c)) S end; fun mk_arity (t, Ss, c) = the_single (mk_arities (t, Ss, [c])); fun dest_arity tm = let fun err () = raise TERM ("dest_arity", [tm]); val (ty, c) = dest_of_class tm; val (t, tvars) = (case ty of Type (t, tys) => (t, map dest_TVar tys handle TYPE _ => err ()) | _ => err ()); val Ss = if has_duplicates (eq_fst (op =)) tvars then err () else map snd tvars; in (t, Ss, c) end; (* internalized sort constraints *) fun dummy_tfree S = TFree ("'dummy", S); type unconstrain_context = {present_map: (typ * typ) list, constraints_map: (sort * typ) list, atyp_map: typ -> typ, map_atyps: typ -> typ, constraints: ((typ * class) * term) list, outer_constraints: (typ * class) list}; fun unconstrainT shyps prop = let val present = rev ((fold_types o fold_atyps_sorts) (insert (eq_fst op =)) prop []); - val extra = fold (Sorts.remove_sort o #2) present shyps; + val extra = Sortset.dest (fold (Sortset.remove o #2) present shyps); val n = length present; val (names1, names2) = Name.invent Name.context Name.aT (n + length extra) |> chop n; val present_map = map2 (fn (T, S) => fn a => (T, TVar ((a, 0), S))) present names1; val constraints_map = map2 (fn (_, S) => fn a => (S, TVar ((a, 0), S))) present names1 @ map2 (fn S => fn a => (S, TVar ((a, 0), S))) extra names2; fun atyp_map T = (case AList.lookup (op =) present_map T of SOME U => U | NONE => (case AList.lookup (op =) constraints_map (Type.sort_of_atyp T) of SOME U => U | NONE => raise TYPE ("Dangling type variable ", [T], [prop]))); val map_atyps = Term.map_atyps (Type.strip_sorts o atyp_map); val constraints = constraints_map |> maps (fn (_, T as TVar (ai, S)) => map (fn c => ((T, c), mk_of_class (TVar (ai, []), c))) S); val outer_constraints = maps (fn (T, S) => map (pair T) S) (present @ map (`dummy_tfree) extra); val ucontext = {present_map = present_map, constraints_map = constraints_map, atyp_map = atyp_map, map_atyps = map_atyps, constraints = constraints, outer_constraints = outer_constraints}; val prop' = prop |> Term.map_types map_atyps |> curry list_implies (map #2 constraints); in (ucontext, prop') end; (** protected propositions and embedded terms **) val protectC = Const ("Pure.prop", propT --> propT); fun protect t = protectC $ t; fun unprotect (Const ("Pure.prop", _) $ t) = t | unprotect t = raise TERM ("unprotect", [t]); fun mk_term t = Const ("Pure.term", Term.fastype_of t --> propT) $ t; fun dest_term (Const ("Pure.term", _) $ t) = t | dest_term t = raise TERM ("dest_term", [t]); (*** Low-level term operations ***) (*Does t occur in u? Or is alpha-convertible to u? The term t must contain no loose bound variables*) fun occs (t, u) = exists_subterm (fn s => t aconv s) u; (*Close up a formula over all free variables by quantification*) fun close_form A = fold_rev (all o Free) (Frees.build (Frees.add_frees A) |> Frees.list_set) A; (*** Specialized operations for resolution... ***) (*computes t(Bound(n+k-1),...,Bound(n)) *) fun combound (t, n, k) = if k>0 then combound (t,n+1,k-1) $ (Bound n) else t; (* ([xn,...,x1], t) goes to \x1 ... xn. t *) fun rlist_abs ([], body) = body | rlist_abs ((a,T)::pairs, body) = rlist_abs(pairs, Abs(a, T, body)); fun incr_tvar_same 0 = Same.same | incr_tvar_same k = Term_Subst.map_atypsT_same (fn TVar ((a, i), S) => TVar ((a, i + k), S) | _ => raise Same.SAME); fun incr_tvar k T = incr_tvar_same k T handle Same.SAME => T; (*For all variables in the term, increment indexnames and lift over the Us result is ?Gidx(B.(lev+n-1),...,B.lev) where lev is abstraction level *) fun incr_indexes_same ([], [], 0) = Same.same | incr_indexes_same (fixed, Ts, k) = let val n = length Ts; val incrT = incr_tvar_same k; fun incr lev (Var ((x, i), T)) = combound (Var ((x, i + k), Ts ---> Same.commit incrT T), lev, n) | incr lev (Free (x, T)) = if member (op =) fixed x then combound (Free (x, Ts ---> Same.commit incrT T), lev, n) else Free (x, incrT T) | incr lev (Abs (x, T, body)) = (Abs (x, incrT T, incr (lev + 1) body handle Same.SAME => body) handle Same.SAME => Abs (x, T, incr (lev + 1) body)) | incr lev (t $ u) = (incr lev t $ (incr lev u handle Same.SAME => u) handle Same.SAME => t $ incr lev u) | incr _ (Const (c, T)) = Const (c, incrT T) | incr _ (Bound _) = raise Same.SAME; in incr 0 end; fun incr_indexes arg t = incr_indexes_same arg t handle Same.SAME => t; (* Lifting functions from subgoal and increment: lift_abs operates on terms lift_all operates on propositions *) fun lift_abs inc = let fun lift Ts (Const ("Pure.imp", _) $ _ $ B) t = lift Ts B t | lift Ts (Const ("Pure.all", _) $ Abs (a, T, B)) t = Abs (a, T, lift (T :: Ts) B t) | lift Ts _ t = incr_indexes ([], rev Ts, inc) t; in lift [] end; fun lift_all inc = let fun lift Ts ((c as Const ("Pure.imp", _)) $ A $ B) t = c $ A $ lift Ts B t | lift Ts ((c as Const ("Pure.all", _)) $ Abs (a, T, B)) t = c $ Abs (a, T, lift (T :: Ts) B t) | lift Ts _ t = incr_indexes ([], rev Ts, inc) t; in lift [] end; (*Strips assumptions in goal, yielding list of hypotheses. *) fun strip_assums_hyp B = let fun strip Hs (Const ("Pure.imp", _) $ H $ B) = strip (H :: Hs) B | strip Hs (Const ("Pure.all", _) $ Abs (a, T, t)) = strip (map (incr_boundvars 1) Hs) t | strip Hs B = rev Hs in strip [] B end; (*Strips assumptions in goal, yielding conclusion. *) fun strip_assums_concl (Const("Pure.imp", _) $ H $ B) = strip_assums_concl B | strip_assums_concl (Const("Pure.all", _) $ Abs (a, T, t)) = strip_assums_concl t | strip_assums_concl B = B; (*Make a list of all the parameters in a subgoal, even if nested*) fun strip_params (Const("Pure.imp", _) $ H $ B) = strip_params B | strip_params (Const("Pure.all", _) $ Abs (a, T, t)) = (a, T) :: strip_params t | strip_params B = []; (*test for nested meta connectives in prems*) val has_meta_prems = let fun is_meta (Const ("Pure.eq", _) $ _ $ _) = true | is_meta (Const ("Pure.imp", _) $ _ $ _) = true | is_meta (Const ("Pure.all", _) $ _) = true | is_meta _ = false; fun ex_meta (Const ("Pure.imp", _) $ A $ B) = is_meta A orelse ex_meta B | ex_meta (Const ("Pure.all", _) $ Abs (_, _, B)) = ex_meta B | ex_meta _ = false; in ex_meta end; (*Removes the parameters from a subgoal and renumber bvars in hypotheses, where j is the total number of parameters (precomputed) If n>0 then deletes assumption n. *) fun remove_params j n A = if j=0 andalso n<=0 then A (*nothing left to do...*) else case A of Const("Pure.imp", _) $ H $ B => if n=1 then (remove_params j (n-1) B) else implies $ (incr_boundvars j H) $ (remove_params j (n-1) B) | Const("Pure.all",_)$Abs(a,T,t) => remove_params (j-1) n t | _ => if n>0 then raise TERM("remove_params", [A]) else A; (*Move all parameters to the front of the subgoal, renaming them apart; if n>0 then deletes assumption n. *) fun flatten_params n A = let val params = strip_params A; val vars = ListPair.zip (Name.variant_list [] (map #1 params), map #2 params) in list_all (vars, remove_params (length vars) n A) end; (*Makes parameters in a goal have the names supplied by the list cs.*) fun list_rename_params cs (Const ("Pure.imp", _) $ A $ B) = implies $ A $ list_rename_params cs B | list_rename_params (c :: cs) ((a as Const ("Pure.all", _)) $ Abs (_, T, t)) = a $ Abs (c, T, list_rename_params cs t) | list_rename_params cs B = B; (*** Treatment of "assume", "erule", etc. ***) (*Strips assumptions in goal yielding HS = [Hn,...,H1], params = [xm,...,x1], and B, where x1...xm are the parameters. This version (21.1.2005) REQUIRES the the parameters to be flattened, but it allows erule to work on assumptions of the form \x. phi. Any \ after the outermost string will be regarded as belonging to the conclusion, and left untouched. Used ONLY by assum_pairs. Unless nasms<0, it can terminate the recursion early; that allows erule to work on assumptions of the form P\Q.*) fun strip_assums_imp (0, Hs, B) = (Hs, B) (*recursion terminated by nasms*) | strip_assums_imp (nasms, Hs, Const("Pure.imp", _) $ H $ B) = strip_assums_imp (nasms-1, H::Hs, B) | strip_assums_imp (_, Hs, B) = (Hs, B); (*recursion terminated by B*) (*Strips OUTER parameters only.*) fun strip_assums_all (params, Const("Pure.all",_)$Abs(a,T,t)) = strip_assums_all ((a,T)::params, t) | strip_assums_all (params, B) = (params, B); (*Produces disagreement pairs, one for each assumption proof, in order. A is the first premise of the lifted rule, and thus has the form H1 \ ... Hk \ B and the pairs are (H1,B),...,(Hk,B). nasms is the number of assumptions in the original subgoal, needed when B has the form B1 \ B2: it stops B1 from being taken as an assumption. *) fun assum_pairs(nasms,A) = let val (params, A') = strip_assums_all ([],A) val (Hs,B) = strip_assums_imp (nasms,[],A') fun abspar t = rlist_abs(params, t) val D = abspar B fun pairrev ([], pairs) = pairs | pairrev (H::Hs, pairs) = pairrev(Hs, (abspar H, D) :: pairs) in pairrev (Hs,[]) end; fun assum_problems (nasms, A) = let val (params, A') = strip_assums_all ([], A) val (Hs, B) = strip_assums_imp (nasms, [], A') fun abspar t = rlist_abs (params, t) in (abspar, rev Hs, B) end; (* global schematic variables *) fun bad_schematic xi = "Illegal schematic variable: " ^ quote (Term.string_of_vname xi); fun bad_fixed x = "Illegal fixed variable: " ^ quote x; fun varifyT_global_same ty = ty |> Term_Subst.map_atypsT_same (fn TFree (a, S) => TVar ((a, 0), S) | TVar (ai, _) => raise TYPE (bad_schematic ai, [ty], [])); fun unvarifyT_global_same ty = ty |> Term_Subst.map_atypsT_same (fn TVar ((a, 0), S) => TFree (a, S) | TVar (ai, _) => raise TYPE (bad_schematic ai, [ty], []) | TFree (a, _) => raise TYPE (bad_fixed a, [ty], [])); val varifyT_global = Same.commit varifyT_global_same; val unvarifyT_global = Same.commit unvarifyT_global_same; fun varify_types_global tm = tm |> Same.commit (Term_Subst.map_types_same varifyT_global_same) handle TYPE (msg, _, _) => raise TERM (msg, [tm]); fun unvarify_types_global tm = tm |> Same.commit (Term_Subst.map_types_same unvarifyT_global_same) handle TYPE (msg, _, _) => raise TERM (msg, [tm]); fun varify_global tm = tm |> Same.commit (Term_Subst.map_aterms_same (fn Free (x, T) => Var ((x, 0), T) | Var (xi, _) => raise TERM (bad_schematic xi, [tm]) | _ => raise Same.SAME)) |> varify_types_global; fun unvarify_global tm = tm |> Same.commit (Term_Subst.map_aterms_same (fn Var ((x, 0), T) => Free (x, T) | Var (xi, _) => raise TERM (bad_schematic xi, [tm]) | Free (x, _) => raise TERM (bad_fixed x, [tm]) | _ => raise Same.SAME)) |> unvarify_types_global; (* goal states *) fun get_goal st i = nth_prem (i, st) handle TERM _ => error ("Subgoal number " ^ string_of_int i ^ " out of range (a total of " ^ string_of_int (count_prems st) ^ " subgoals)"); (*reverses parameters for substitution*) fun goal_params st i = let val gi = get_goal st i val rfrees = map Free (Term.rename_wrt_term gi (strip_params gi)) in (gi, rfrees) end; fun concl_of_goal st i = let val (gi, rfrees) = goal_params st i val B = strip_assums_concl gi in subst_bounds (rfrees, B) end; fun prems_of_goal st i = let val (gi, rfrees) = goal_params st i val As = strip_assums_hyp gi in map (fn A => subst_bounds (rfrees, A)) As end; end; diff --git a/src/Pure/more_thm.ML b/src/Pure/more_thm.ML --- a/src/Pure/more_thm.ML +++ b/src/Pure/more_thm.ML @@ -1,757 +1,757 @@ (* Title: Pure/more_thm.ML Author: Makarius Further operations on type ctyp/cterm/thm, outside the inference kernel. *) infix aconvc; signature BASIC_THM = sig include BASIC_THM val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val aconvc: cterm * cterm -> bool type attribute = Context.generic * thm -> Context.generic option * thm option end; signature THM = sig include THM val eq_ctyp: ctyp * ctyp -> bool val aconvc: cterm * cterm -> bool val add_tvars: thm -> ctyp TVars.table -> ctyp TVars.table val add_vars: thm -> cterm Vars.table -> cterm Vars.table val dest_funT: ctyp -> ctyp * ctyp val strip_type: ctyp -> ctyp list * ctyp val instantiate_ctyp: ctyp TVars.table -> ctyp -> ctyp val all_name: Proof.context -> string * cterm -> cterm -> cterm val all: Proof.context -> cterm -> cterm -> cterm val mk_binop: cterm -> cterm -> cterm -> cterm val dest_binop: cterm -> cterm * cterm val dest_implies: cterm -> cterm * cterm val dest_equals: cterm -> cterm * cterm val dest_equals_lhs: cterm -> cterm val dest_equals_rhs: cterm -> cterm val lhs_of: thm -> cterm val rhs_of: thm -> cterm val is_reflexive: thm -> bool val eq_thm: thm * thm -> bool val eq_thm_prop: thm * thm -> bool val eq_thm_strict: thm * thm -> bool val equiv_thm: theory -> thm * thm -> bool val class_triv: theory -> class -> thm val of_sort: ctyp * sort -> thm list val is_dummy: thm -> bool val add_thm: thm -> thm list -> thm list val del_thm: thm -> thm list -> thm list val merge_thms: thm list * thm list -> thm list val item_net: thm Item_Net.T val item_net_intro: thm Item_Net.T val item_net_elim: thm Item_Net.T val declare_hyps: cterm -> Proof.context -> Proof.context val assume_hyps: cterm -> Proof.context -> thm * Proof.context val unchecked_hyps: Proof.context -> Proof.context val restore_hyps: Proof.context -> Proof.context -> Proof.context val undeclared_hyps: Context.generic -> thm -> term list val check_hyps: Context.generic -> thm -> thm val declare_term_sorts: term -> Proof.context -> Proof.context - val extra_shyps': Proof.context -> thm -> sort list + val extra_shyps': Proof.context -> thm -> Sortset.T val check_shyps: Proof.context -> thm -> thm val weaken_sorts': Proof.context -> cterm -> cterm val elim_implies: thm -> thm -> thm val forall_intr_name: string * cterm -> thm -> thm val forall_elim_var: int -> thm -> thm val forall_elim_vars: int -> thm -> thm val instantiate_frees: ctyp TFrees.table * cterm Frees.table -> thm -> thm val instantiate': ctyp option list -> cterm option list -> thm -> thm val forall_intr_frees: thm -> thm val forall_intr_vars: thm -> thm val unvarify_global: theory -> thm -> thm val unvarify_axiom: theory -> string -> thm val rename_params_rule: string list * int -> thm -> thm val rename_boundvars: term -> term -> thm -> thm val add_axiom: Proof.context -> binding * term -> theory -> (string * thm) * theory val add_axiom_global: binding * term -> theory -> (string * thm) * theory val add_def: Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory val add_def_global: bool -> bool -> binding * term -> theory -> (string * thm) * theory type attribute = Context.generic * thm -> Context.generic option * thm option type binding = binding * attribute list val tag_rule: string * string -> thm -> thm val untag_rule: string -> thm -> thm val is_free_dummy: thm -> bool val tag_free_dummy: thm -> thm val def_name: string -> string val def_name_optional: string -> string -> string val def_binding: Binding.binding -> Binding.binding val def_binding_optional: Binding.binding -> Binding.binding -> Binding.binding val make_def_binding: bool -> Binding.binding -> Binding.binding val has_name_hint: thm -> bool val get_name_hint: thm -> string val put_name_hint: string -> thm -> thm val theoremK: string val legacy_get_kind: thm -> string val kind_rule: string -> thm -> thm val rule_attribute: thm list -> (Context.generic -> thm -> thm) -> attribute val declaration_attribute: (thm -> Context.generic -> Context.generic) -> attribute val mixed_attribute: (Context.generic * thm -> Context.generic * thm) -> attribute val apply_attribute: attribute -> thm -> Context.generic -> thm * Context.generic val attribute_declaration: attribute -> thm -> Context.generic -> Context.generic val theory_attributes: attribute list -> thm -> theory -> thm * theory val proof_attributes: attribute list -> thm -> Proof.context -> thm * Proof.context val no_attributes: 'a -> 'a * 'b list val simple_fact: 'a -> ('a * 'b list) list val tag: string * string -> attribute val untag: string -> attribute val kind: string -> attribute val register_proofs: thm list lazy -> theory -> theory val consolidate_theory: theory -> unit val expose_theory: theory -> unit val show_consts: bool Config.T val show_hyps: bool Config.T val show_tags: bool Config.T val pretty_thm_raw: Proof.context -> {quote: bool, show_hyps: bool} -> thm -> Pretty.T val pretty_thm: Proof.context -> thm -> Pretty.T val pretty_thm_item: Proof.context -> thm -> Pretty.T val pretty_thm_global: theory -> thm -> Pretty.T val string_of_thm: Proof.context -> thm -> string val string_of_thm_global: theory -> thm -> string end; structure Thm: THM = struct (** basic operations **) (* collecting ctyps and cterms *) val eq_ctyp = op = o apply2 Thm.typ_of; val op aconvc = op aconv o apply2 Thm.term_of; val add_tvars = Thm.fold_atomic_ctyps {hyps = false} Term.is_TVar (fn cT => fn tab => let val v = Term.dest_TVar (Thm.typ_of cT) in tab |> not (TVars.defined tab v) ? TVars.add (v, cT) end); val add_vars = Thm.fold_atomic_cterms {hyps = false} Term.is_Var (fn ct => fn tab => let val v = Term.dest_Var (Thm.term_of ct) in tab |> not (Vars.defined tab v) ? Vars.add (v, ct) end); (* ctyp operations *) fun dest_funT cT = (case Thm.typ_of cT of Type ("fun", _) => let val [A, B] = Thm.dest_ctyp cT in (A, B) end | T => raise TYPE ("dest_funT", [T], [])); (* ctyp version of strip_type: maps [T1,...,Tn]--->T to ([T1,T2,...,Tn], T) *) fun strip_type cT = (case Thm.typ_of cT of Type ("fun", _) => let val (cT1, cT2) = dest_funT cT; val (cTs, cT') = strip_type cT2 in (cT1 :: cTs, cT') end | _ => ([], cT)); fun instantiate_ctyp instT cT = Thm.instantiate_cterm (instT, Vars.empty) (Thm.var (("x", 0), cT)) |> Thm.ctyp_of_cterm; (* cterm operations *) fun all_name ctxt (x, t) A = let val T = Thm.typ_of_cterm t; val all_const = Thm.cterm_of ctxt (Const ("Pure.all", (T --> propT) --> propT)); in Thm.apply all_const (Thm.lambda_name (x, t) A) end; fun all ctxt t A = all_name ctxt ("", t) A; fun mk_binop c a b = Thm.apply (Thm.apply c a) b; fun dest_binop ct = (Thm.dest_arg1 ct, Thm.dest_arg ct); fun dest_implies ct = (case Thm.term_of ct of Const ("Pure.imp", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_implies", [Thm.term_of ct])); fun dest_equals ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => dest_binop ct | _ => raise TERM ("dest_equals", [Thm.term_of ct])); fun dest_equals_lhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg1 ct | _ => raise TERM ("dest_equals_lhs", [Thm.term_of ct])); fun dest_equals_rhs ct = (case Thm.term_of ct of Const ("Pure.eq", _) $ _ $ _ => Thm.dest_arg ct | _ => raise TERM ("dest_equals_rhs", [Thm.term_of ct])); val lhs_of = dest_equals_lhs o Thm.cprop_of; val rhs_of = dest_equals_rhs o Thm.cprop_of; (* equality *) fun is_reflexive th = op aconv (Logic.dest_equals (Thm.prop_of th)) handle TERM _ => false; val eq_thm = is_equal o Thm.thm_ord; val eq_thm_prop = op aconv o apply2 Thm.full_prop_of; fun eq_thm_strict ths = eq_thm ths andalso Context.eq_thy_id (apply2 Thm.theory_id ths) andalso op = (apply2 Thm.maxidx_of ths) andalso op = (apply2 Thm.get_tags ths); (* pattern equivalence *) fun equiv_thm thy ths = Pattern.equiv thy (apply2 (Thm.full_prop_of o Thm.transfer thy) ths); (* type classes and sorts *) fun class_triv thy c = Thm.of_class (Thm.global_ctyp_of thy (TVar ((Name.aT, 0), [c])), c); fun of_sort (T, S) = map (fn c => Thm.of_class (T, c)) S; (* misc operations *) fun is_dummy thm = (case try Logic.dest_term (Thm.concl_of thm) of NONE => false | SOME t => Term.is_dummy_pattern (Term.head_of t)); (* collections of theorems in canonical order *) val add_thm = update eq_thm_prop; val del_thm = remove eq_thm_prop; val merge_thms = merge eq_thm_prop; val item_net = Item_Net.init eq_thm_prop (single o Thm.full_prop_of); val item_net_intro = Item_Net.init eq_thm_prop (single o Thm.concl_of); val item_net_elim = Item_Net.init eq_thm_prop (single o Thm.major_prem_of); (** declared hyps and sort hyps **) structure Hyps = Proof_Data ( - type T = {checked_hyps: bool, hyps: Termset.T, shyps: sort Ord_List.T}; - fun init _ : T = {checked_hyps = true, hyps = Termset.empty, shyps = []}; + type T = {checked_hyps: bool, hyps: Termset.T, shyps: Sortset.T}; + fun init _ : T = {checked_hyps = true, hyps = Termset.empty, shyps = Sortset.empty}; ); fun map_hyps f = Hyps.map (fn {checked_hyps, hyps, shyps} => let val (checked_hyps', hyps', shyps') = f (checked_hyps, hyps, shyps) in {checked_hyps = checked_hyps', hyps = hyps', shyps = shyps'} end); (* hyps *) fun declare_hyps raw_ct ctxt = ctxt |> map_hyps (fn (checked_hyps, hyps, shyps) => let val ct = Thm.transfer_cterm (Proof_Context.theory_of ctxt) raw_ct; val hyps' = Termset.insert (Thm.term_of ct) hyps; in (checked_hyps, hyps', shyps) end); fun assume_hyps ct ctxt = (Thm.assume ct, declare_hyps ct ctxt); val unchecked_hyps = map_hyps (fn (_, hyps, shyps) => (false, hyps, shyps)); fun restore_hyps ctxt = map_hyps (fn (_, hyps, shyps) => (#checked_hyps (Hyps.get ctxt), hyps, shyps)); fun undeclared_hyps context th = Thm.hyps_of th |> filter_out (case context of Context.Theory _ => K false | Context.Proof ctxt => (case Hyps.get ctxt of {checked_hyps = false, ...} => K true | {hyps, ...} => Termset.member hyps)); fun check_hyps context th = (case undeclared_hyps context th of [] => th | undeclared => error (Pretty.string_of (Pretty.big_list "Undeclared hyps:" (map (Pretty.item o single o Syntax.pretty_term (Syntax.init_pretty context)) undeclared)))); (* shyps *) fun declare_term_sorts t = map_hyps (fn (checked_hyps, hyps, shyps) => - (checked_hyps, hyps, Sorts.insert_term t shyps)); + (checked_hyps, hyps, Sortset.insert_term t shyps)); fun extra_shyps' ctxt th = - Sorts.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); + Sortset.subtract (#shyps (Hyps.get ctxt)) (Thm.extra_shyps th); fun check_shyps ctxt raw_th = let val th = Thm.strip_shyps raw_th; val extra_shyps = extra_shyps' ctxt th; in - if null extra_shyps then th + if Sortset.is_empty extra_shyps then th else error (Pretty.string_of (Pretty.block (Pretty.str "Pending sort hypotheses:" :: - Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) extra_shyps)))) + Pretty.brk 1 :: Pretty.commas (map (Syntax.pretty_sort ctxt) (Sortset.dest extra_shyps))))) end; val weaken_sorts' = Thm.weaken_sorts o #shyps o Hyps.get; (** basic derived rules **) (*Elimination of implication A A \ B ------------ B *) fun elim_implies thA thAB = Thm.implies_elim thAB thA; (* forall_intr_name *) fun forall_intr_name (a, x) th = let val th' = Thm.forall_intr x th; val prop' = (case Thm.prop_of th' of all $ Abs (_, T, b) => all $ Abs (a, T, b)); in Thm.renamed_prop prop' th' end; (* forall_elim_var(s) *) local val used_frees = Name.build_context o Thm.fold_terms {hyps = true} Term.declare_term_frees; fun used_vars i = Name.build_context o (Thm.fold_terms {hyps = false} o Term.fold_aterms) (fn Var ((x, j), _) => i = j ? Name.declare x | _ => I); fun dest_all ct used = (case Thm.term_of ct of Const ("Pure.all", _) $ Abs (x, _, _) => let val (x', used') = Name.variant x used; val (v, ct') = Thm.dest_abs_fresh x' (Thm.dest_arg ct); in SOME ((x, Thm.ctyp_of_cterm v), (ct', used')) end | _ => NONE); fun dest_all_list ct used = (case dest_all ct used of NONE => ([], used) | SOME (v, (ct', used')) => let val (vs, used'') = dest_all_list ct' used' in (v :: vs, used'') end); fun forall_elim_vars_list vars i th = let val (vars', _) = (vars, used_vars i th) |-> fold_map (fn (x, T) => fn used => let val (x', used') = Name.variant x used in (Thm.var ((x', i), T), used') end); in fold Thm.forall_elim vars' th end; in fun forall_elim_vars i th = forall_elim_vars_list (#1 (dest_all_list (Thm.cprop_of th) (used_frees th))) i th; fun forall_elim_var i th = let val vars = (case dest_all (Thm.cprop_of th) (used_frees th) of SOME (v, _) => [v] | NONE => raise THM ("forall_elim_var", i, [th])); in forall_elim_vars_list vars i th end; end; (* instantiate frees *) fun instantiate_frees (instT, inst) th = if TFrees.is_empty instT andalso Frees.is_empty inst then th else let val idx = Thm.maxidx_of th + 1; fun index ((a, A), b) = (((a, idx), A), b); val insts = (TVars.build (TFrees.fold (TVars.add o index) instT), Vars.build (Frees.fold (Vars.add o index) inst)); val tfrees = Names.build (TFrees.fold (Names.add_set o #1 o #1) instT); val frees = Names.build (Frees.fold (Names.add_set o #1 o #1) inst); val hyps = Thm.chyps_of th; val inst_cterm = Thm.generalize_cterm (tfrees, frees) idx #> Thm.instantiate_cterm insts; in th |> fold_rev Thm.implies_intr hyps |> Thm.generalize (tfrees, frees) idx |> Thm.instantiate insts |> fold (elim_implies o Thm.assume o inst_cterm) hyps end; (* instantiate by left-to-right occurrence of variables *) fun instantiate' cTs cts thm = let fun err msg = raise TYPE ("instantiate': " ^ msg, map_filter (Option.map Thm.typ_of) cTs, map_filter (Option.map Thm.term_of) cts); fun zip_vars xs ys = zip_options xs ys handle ListPair.UnequalLengths => err "more instantiations than variables in thm"; val instT = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_tvars thm)) cTs; val thm' = Thm.instantiate (TVars.make instT, Vars.empty) thm; val inst = zip_vars (build_rev (Thm.fold_terms {hyps = false} Term.add_vars thm')) cts; in Thm.instantiate (TVars.empty, Vars.make inst) thm' end; (* implicit generalization over variables -- canonical order *) fun forall_intr_vars th = let val vars = Cterms.build (Cterms.add_vars th) in fold_rev Thm.forall_intr (Cterms.list_set vars) th end; fun forall_intr_frees th = let val fixed = Frees.build (fold Frees.add_frees (Thm.terms_of_tpairs (Thm.tpairs_of th)) #> fold Frees.add_frees (Thm.hyps_of th)); val is_fixed = Frees.defined fixed o Term.dest_Free o Thm.term_of; val frees = Cterms.build (th |> Thm.fold_atomic_cterms {hyps = false} Term.is_Free (fn ct => not (is_fixed ct) ? Cterms.add_set ct)); in fold_rev Thm.forall_intr (Cterms.list_set frees) th end; (* unvarify_global: global schematic variables *) fun unvarify_global thy th = let val prop = Thm.full_prop_of th; val _ = map Logic.unvarify_global (prop :: Thm.hyps_of th) handle TERM (msg, _) => raise THM (msg, 0, [th]); val cert = Thm.global_cterm_of thy; val certT = Thm.global_ctyp_of thy; val instT = TVars.build (prop |> (Term.fold_types o Term.fold_atyps) (fn T => fn instT => (case T of TVar (v as ((a, _), S)) => if TVars.defined instT v then instT else TVars.add (v, TFree (a, S)) instT | _ => instT))); val cinstT = TVars.map (K certT) instT; val cinst = Vars.build (prop |> Term.fold_aterms (fn t => fn inst => (case t of Var ((x, i), T) => let val T' = Term_Subst.instantiateT instT T in Vars.add (((x, i), T'), cert (Free ((x, T')))) inst end | _ => inst))); in Thm.instantiate (cinstT, cinst) th end; fun unvarify_axiom thy = unvarify_global thy o Thm.axiom thy; (* user renaming of parameters in a subgoal *) (*The names, if distinct, are used for the innermost parameters of subgoal i; preceding parameters may be renamed to make all parameters distinct.*) fun rename_params_rule (names, i) st = let val (_, Bs, Bi, C) = Thm.dest_state (st, i); val params = map #1 (Logic.strip_params Bi); val short = length params - length names; val names' = if short < 0 then error "More names than parameters in subgoal!" else Name.variant_list names (take short params) @ names; val free_names = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val Bi' = Logic.list_rename_params names' Bi; in (case duplicates (op =) names of a :: _ => (warning ("Can't rename. Bound variables not distinct: " ^ a); st) | [] => (case inter (op =) names free_names of a :: _ => (warning ("Can't rename. Bound/Free variable clash: " ^ a); st) | [] => Thm.renamed_prop (Logic.list_implies (Bs @ [Bi'], C)) st)) end; (* preservation of bound variable names *) fun rename_boundvars pat obj th = (case Term.rename_abs pat obj (Thm.prop_of th) of NONE => th | SOME prop' => Thm.renamed_prop prop' th); (** specification primitives **) (* rules *) fun stripped_sorts thy t = let val tfrees = build_rev (Term.add_tfrees t); val tfrees' = map (fn a => (a, [])) (Name.variant_list [] (map #1 tfrees)); val recover = map2 (fn (a', S') => fn (a, S) => (((a', 0), S'), Thm.global_ctyp_of thy (TVar ((a, 0), S)))) tfrees' tfrees; val strip = map (apply2 TFree) (tfrees ~~ tfrees'); val t' = Term.map_types (Term.map_atyps (perhaps (AList.lookup (op =) strip))) t; in (strip, recover, t') end; fun add_axiom ctxt (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val (strip, recover, prop') = stripped_sorts thy prop; val constraints = map (fn (TFree (_, S), T) => (T, S)) strip; val of_sorts = maps (fn (T as TFree (_, S), _) => of_sort (Thm.ctyp_of ctxt T, S)) strip; val thy' = thy |> Theory.add_axiom ctxt (b, Logic.list_implies (maps Logic.mk_of_sort constraints, prop')); val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (TVars.make recover, Vars.empty) axm' |> unvarify_global thy' |> fold elim_implies of_sorts; in ((axm_name, thm), thy') end; fun add_axiom_global arg thy = add_axiom (Syntax.init_pretty_global thy) arg thy; fun add_def (context as (ctxt, _)) unchecked overloaded (b, prop) thy = let val _ = Sign.no_vars ctxt prop; val prems = map (Thm.cterm_of ctxt) (Logic.strip_imp_prems prop); val (_, recover, concl') = stripped_sorts thy (Logic.strip_imp_concl prop); val thy' = Theory.add_def context unchecked overloaded (b, concl') thy; val axm_name = Sign.full_name thy' b; val axm' = Thm.axiom thy' axm_name; val thm = Thm.instantiate (TVars.make recover, Vars.empty) axm' |> unvarify_global thy' |> fold_rev Thm.implies_intr prems; in ((axm_name, thm), thy') end; fun add_def_global unchecked overloaded arg thy = add_def (Defs.global_context thy) unchecked overloaded arg thy; (** theorem tags **) (* add / delete tags *) fun tag_rule tg = Thm.map_tags (insert (op =) tg); fun untag_rule s = Thm.map_tags (filter_out (fn (s', _) => s = s')); (* free dummy thm -- for abstract closure *) val free_dummyN = "free_dummy"; fun is_free_dummy thm = Properties.defined (Thm.get_tags thm) free_dummyN; val tag_free_dummy = tag_rule (free_dummyN, ""); (* def_name *) fun def_name c = c ^ "_def"; fun def_name_optional c "" = def_name c | def_name_optional _ name = name; val def_binding = Binding.map_name def_name #> Binding.reset_pos; fun def_binding_optional b name = if Binding.is_empty name then def_binding b else name; fun make_def_binding cond b = if cond then def_binding b else Binding.empty; (* unofficial theorem names *) fun has_name_hint thm = AList.defined (op =) (Thm.get_tags thm) Markup.nameN; fun the_name_hint thm = the (AList.lookup (op =) (Thm.get_tags thm) Markup.nameN); fun get_name_hint thm = if has_name_hint thm then the_name_hint thm else "??.unknown"; fun put_name_hint name = untag_rule Markup.nameN #> tag_rule (Markup.nameN, name); (* theorem kinds *) val theoremK = "theorem"; fun legacy_get_kind thm = the_default "" (Properties.get (Thm.get_tags thm) Markup.kindN); fun kind_rule k = tag_rule (Markup.kindN, k) o untag_rule Markup.kindN; (** attributes **) (*attributes subsume any kind of rules or context modifiers*) type attribute = Context.generic * thm -> Context.generic option * thm option; type binding = binding * attribute list; fun rule_attribute ths f (x, th) = (NONE, (case find_first is_free_dummy (th :: ths) of SOME th' => SOME th' | NONE => SOME (f x th))); fun declaration_attribute f (x, th) = (if is_free_dummy th then NONE else SOME (f th x), NONE); fun mixed_attribute f (x, th) = let val (x', th') = f (x, th) in (SOME x', SOME th') end; fun apply_attribute (att: attribute) th x = let val (x', th') = att (x, check_hyps x (Thm.transfer'' x th)) in (the_default th th', the_default x x') end; fun attribute_declaration att th x = #2 (apply_attribute att th x); fun apply_attributes mk dest = let fun app [] th x = (th, x) | app (att :: atts) th x = apply_attribute att th (mk x) ||> dest |-> app atts; in app end; val theory_attributes = apply_attributes Context.Theory Context.the_theory; val proof_attributes = apply_attributes Context.Proof Context.the_proof; fun no_attributes x = (x, []); fun simple_fact x = [(x, [])]; fun tag tg = rule_attribute [] (K (tag_rule tg)); fun untag s = rule_attribute [] (K (untag_rule s)); fun kind k = rule_attribute [] (K (k <> "" ? kind_rule k)); (** forked proofs **) structure Proofs = Theory_Data ( type T = thm list lazy Inttab.table; val empty = Inttab.empty; val merge = Inttab.merge (K true); ); fun reset_proofs thy = if Inttab.is_empty (Proofs.get thy) then NONE else SOME (Proofs.put Inttab.empty thy); val _ = Theory.setup (Theory.at_begin reset_proofs); fun register_proofs ths thy = let val entry = (serial (), Lazy.map_finished (map Thm.trim_context) ths) in (Proofs.map o Inttab.update) entry thy end; fun force_proofs thy = Proofs.get thy |> Inttab.dest |> maps (map (Thm.transfer thy) o Lazy.force o #2); val consolidate_theory = Thm.consolidate o force_proofs; fun expose_theory thy = if Proofterm.export_enabled () then Thm.expose_proofs thy (force_proofs thy) else (); (** print theorems **) (* options *) val show_consts = Config.declare_option_bool ("show_consts", \<^here>); val show_hyps = Config.declare_bool ("show_hyps", \<^here>) (K false); val show_tags = Config.declare_bool ("show_tags", \<^here>) (K false); (* pretty_thm etc. *) fun pretty_tag (name, arg) = Pretty.strs [name, quote arg]; val pretty_tags = Pretty.list "[" "]" o map pretty_tag; fun pretty_thm_raw ctxt {quote, show_hyps = show_hyps'} raw_th = let val show_tags = Config.get ctxt show_tags; val show_hyps = Config.get ctxt show_hyps; val th = raw_th |> perhaps (try (Thm.transfer' ctxt)) |> perhaps (try Thm.strip_shyps); val hyps = if show_hyps then Thm.hyps_of th else undeclared_hyps (Context.Proof ctxt) th; - val extra_shyps = extra_shyps' ctxt th; + val extra_shyps = Sortset.dest (extra_shyps' ctxt th); val tags = Thm.get_tags th; val tpairs = Thm.tpairs_of th; val q = if quote then Pretty.quote else I; val prt_term = q o Syntax.pretty_term ctxt; val hlen = length extra_shyps + length hyps + length tpairs; val hsymbs = if hlen = 0 then [] else if show_hyps orelse show_hyps' then [Pretty.brk 2, Pretty.list "[" "]" (map (q o Pretty.block o Syntax.pretty_flexpair ctxt) tpairs @ map prt_term hyps @ map (Syntax.pretty_sort ctxt) extra_shyps)] else [Pretty.brk 2, Pretty.str ("[" ^ replicate_string hlen "." ^ "]")]; val tsymbs = if null tags orelse not show_tags then [] else [Pretty.brk 1, pretty_tags tags]; in Pretty.block (prt_term (Thm.prop_of th) :: (hsymbs @ tsymbs)) end; fun pretty_thm ctxt = pretty_thm_raw ctxt {quote = false, show_hyps = true}; fun pretty_thm_item ctxt th = Pretty.item [pretty_thm ctxt th]; fun pretty_thm_global thy = pretty_thm_raw (Syntax.init_pretty_global thy) {quote = false, show_hyps = false}; val string_of_thm = Pretty.string_of oo pretty_thm; val string_of_thm_global = Pretty.string_of oo pretty_thm_global; open Thm; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm; diff --git a/src/Pure/proofterm.ML b/src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML +++ b/src/Pure/proofterm.ML @@ -1,2341 +1,2341 @@ (* Title: Pure/proofterm.ML Author: Stefan Berghofer, TU Muenchen LF style proof terms. *) infix 8 % %% %>; signature PROOFTERM = sig type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option} type thm_body type thm_node datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | % of proof * term option | %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} type oracle = (string * Position.T) * term option type thm = serial * thm_node exception MIN_PROOF of unit val proof_of: proof_body -> proof val join_proof: proof_body future -> proof val map_proof_of: (proof -> proof) -> proof_body -> proof_body val thm_header: serial -> Position.T list -> string -> string -> term -> typ list option -> thm_header val thm_body: proof_body -> thm_body val thm_body_proof_raw: thm_body -> proof val thm_body_proof_open: thm_body -> proof val thm_node_theory_name: thm_node -> string val thm_node_name: thm_node -> string val thm_node_prop: thm_node -> term val thm_node_body: thm_node -> proof_body future val thm_node_thms: thm_node -> thm list val join_thms: thm list -> proof_body list val make_thm: thm_header -> thm_body -> thm val fold_proof_atoms: bool -> (proof -> 'a -> 'a) -> proof list -> 'a -> 'a val fold_body_thms: ({serial: serial, name: string, prop: term, body: proof_body} -> 'a -> 'a) -> proof_body list -> 'a -> 'a val oracle_ord: oracle ord val thm_ord: thm ord val unions_oracles: oracle Ord_List.T list -> oracle Ord_List.T val unions_thms: thm Ord_List.T list -> thm Ord_List.T val no_proof_body: proof -> proof_body val no_thm_names: proof -> proof val no_thm_proofs: proof -> proof val no_body_proofs: proof -> proof val encode: Consts.T -> proof XML.Encode.T val encode_body: Consts.T -> proof_body XML.Encode.T val encode_standard_term: Consts.T -> term XML.Encode.T val encode_standard_proof: Consts.T -> proof XML.Encode.T val decode: Consts.T -> proof XML.Decode.T val decode_body: Consts.T -> proof_body XML.Decode.T val %> : proof * term -> proof (*primitive operations*) val proofs: int Unsynchronized.ref val proofs_enabled: unit -> bool val atomic_proof: proof -> bool val compact_proof: proof -> bool val proof_combt: proof * term list -> proof val proof_combt': proof * term option list -> proof val proof_combP: proof * proof list -> proof val strip_combt: proof -> proof * term option list val strip_combP: proof -> proof * proof list val strip_thm_body: proof_body -> proof_body val map_proof_same: term Same.operation -> typ Same.operation -> (typ * class -> proof) -> proof Same.operation val map_proof_terms_same: term Same.operation -> typ Same.operation -> proof Same.operation val map_proof_types_same: typ Same.operation -> proof Same.operation val map_proof_terms: (term -> term) -> (typ -> typ) -> proof -> proof val map_proof_types: (typ -> typ) -> proof -> proof val fold_proof_terms: (term -> 'a -> 'a) -> proof -> 'a -> 'a val fold_proof_terms_types: (term -> 'a -> 'a) -> (typ -> 'a -> 'a) -> proof -> 'a -> 'a val maxidx_proof: proof -> int -> int val size_of_proof: proof -> int val change_types: typ list option -> proof -> proof val prf_abstract_over: term -> proof -> proof val prf_incr_bv: int -> int -> int -> int -> proof -> proof val incr_pboundvars: int -> int -> proof -> proof val prf_loose_bvar1: proof -> int -> bool val prf_loose_Pbvar1: proof -> int -> bool val prf_add_loose_bnos: int -> int -> proof -> int list * int list -> int list * int list val norm_proof: Envir.env -> proof -> proof val norm_proof': Envir.env -> proof -> proof val prf_subst_bounds: term list -> proof -> proof val prf_subst_pbounds: proof list -> proof -> proof val freeze_thaw_prf: proof -> proof * (proof -> proof) (*proof terms for specific inference rules*) val trivial_proof: proof val implies_intr_proof: term -> proof -> proof val implies_intr_proof': term -> proof -> proof val forall_intr_proof: string * term -> typ option -> proof -> proof val forall_intr_proof': term -> proof -> proof val varify_proof: term -> TFrees.set -> proof -> proof val legacy_freezeT: term -> proof -> proof val rotate_proof: term list -> term -> (string * typ) list -> term list -> int -> proof -> proof val permute_prems_proof: term list -> int -> int -> proof -> proof val generalize_proof: Names.set * Names.set -> int -> term -> proof -> proof val instantiate: typ TVars.table * term Vars.table -> proof -> proof val lift_proof: term -> int -> term -> proof -> proof val incr_indexes: int -> proof -> proof val assumption_proof: term list -> term -> int -> proof -> proof val bicompose_proof: bool -> term list -> term list -> term option -> term list -> int -> int -> proof -> proof -> proof val equality_axms: (string * term) list val reflexive_axm: proof val symmetric_axm: proof val transitive_axm: proof val equal_intr_axm: proof val equal_elim_axm: proof val abstract_rule_axm: proof val combination_axm: proof val reflexive_proof: proof val symmetric_proof: proof -> proof val transitive_proof: typ -> term -> proof -> proof -> proof val equal_intr_proof: term -> term -> proof -> proof -> proof val equal_elim_proof: term -> term -> proof -> proof -> proof val abstract_rule_proof: string * term -> proof -> proof val combination_proof: term -> term -> term -> term -> proof -> proof -> proof val strip_shyps_proof: Sorts.algebra -> (typ * sort) list -> (typ * sort) list -> - sort list -> proof -> proof + Sortset.T -> 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 -> + (string * class list list * class -> proof) -> string * Position.T -> Sortset.T -> 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 -> + (string * class list list * class -> proof) -> Sortset.T -> term -> (serial * proof_body future) list -> proof_body -> thm * proof - val get_identity: sort list -> term list -> term -> proof -> + val get_identity: Sortset.T -> term list -> term -> proof -> {serial: serial, theory_name: string, name: string} option - val get_approximative_name: sort list -> term list -> term -> proof -> string + val get_approximative_name: Sortset.T -> 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 get_id: Sortset.T -> term list -> term -> proof -> thm_id option val this_id: thm_id option -> thm_id -> bool val proof_boxes: {excluded: thm_id -> bool, included: thm_id -> bool} -> proof list -> (thm_header * proof) list (*exception MIN_PROOF*) end structure Proofterm : PROOFTERM = struct (** datatype proof **) type thm_header = {serial: serial, pos: Position.T list, theory_name: string, name: string, prop: term, types: typ list option}; datatype proof = MinProof | PBound of int | Abst of string * typ option * proof | AbsP of string * term option * proof | op % of proof * term option | op %% of proof * proof | Hyp of term | PAxm of string * term * typ list option | PClass of typ * class | Oracle of string * term * typ list option | PThm of thm_header * thm_body and proof_body = PBody of {oracles: ((string * Position.T) * term option) Ord_List.T, thms: (serial * thm_node) Ord_List.T, proof: proof} and thm_body = Thm_Body of {open_proof: proof -> proof, body: proof_body future} and thm_node = Thm_Node of {theory_name: string, name: string, prop: term, body: proof_body future, export: unit lazy, consolidate: unit lazy}; type oracle = (string * Position.T) * term option; val oracle_ord: oracle ord = prod_ord (prod_ord fast_string_ord Position.ord) (option_ord Term_Ord.fast_term_ord); type thm = serial * thm_node; val thm_ord: thm ord = fn ((i, _), (j, _)) => int_ord (j, i); exception MIN_PROOF of unit; fun proof_of (PBody {proof, ...}) = proof; val join_proof = Future.join #> proof_of; fun map_proof_of f (PBody {oracles, thms, proof}) = PBody {oracles = oracles, thms = thms, proof = f proof}; fun thm_header serial pos theory_name name prop types : thm_header = {serial = serial, pos = pos, theory_name = theory_name, name = name, prop = prop, types = types}; fun thm_body body = Thm_Body {open_proof = I, body = Future.value body}; fun thm_body_proof_raw (Thm_Body {body, ...}) = join_proof body; fun thm_body_proof_open (Thm_Body {open_proof, body, ...}) = open_proof (join_proof body); fun rep_thm_node (Thm_Node args) = args; val thm_node_theory_name = #theory_name o rep_thm_node; val thm_node_name = #name o rep_thm_node; val thm_node_prop = #prop o rep_thm_node; val thm_node_body = #body o rep_thm_node; val thm_node_thms = thm_node_body #> Future.join #> (fn PBody {thms, ...} => thms); val thm_node_export = #export o rep_thm_node; val thm_node_consolidate = #consolidate o rep_thm_node; fun join_thms (thms: thm list) = Future.joins (map (thm_node_body o #2) thms); val consolidate_bodies = maps (fn PBody {thms, ...} => map (thm_node_consolidate o #2) thms) #> Lazy.consolidate #> map Lazy.force #> ignore; fun make_thm_node theory_name name prop body export = let val consolidate = Lazy.lazy_name "Proofterm.make_thm_node" (fn () => let val PBody {thms, ...} = Future.join body in consolidate_bodies (join_thms thms) end); in Thm_Node {theory_name = theory_name, name = name, prop = prop, body = body, export = export, consolidate = consolidate} end; val no_export = Lazy.value (); fun make_thm ({serial, theory_name, name, prop, ...}: thm_header) (Thm_Body {body, ...}) = (serial, make_thm_node theory_name name prop body no_export); (* proof atoms *) fun fold_proof_atoms all f = let fun app (Abst (_, _, prf)) = app prf | app (AbsP (_, _, prf)) = app prf | app (prf % _) = app prf | app (prf1 %% prf2) = app prf1 #> app prf2 | app (prf as PThm ({serial = i, ...}, Thm_Body {body, ...})) = (fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val (x', seen') = (if all then app (join_proof body) else I) (x, Inttab.update (i, ()) seen) in (f prf x', seen') end) | app prf = (fn (x, seen) => (f prf x, seen)); in fn prfs => fn x => #1 (fold app prfs (x, Inttab.empty)) end; fun fold_body_thms f = let fun app (PBody {thms, ...}) = tap join_thms thms |> fold (fn (i, thm_node) => fn (x, seen) => if Inttab.defined seen i then (x, seen) else let val name = thm_node_name thm_node; val prop = thm_node_prop thm_node; val body = Future.join (thm_node_body thm_node); val (x', seen') = app body (x, Inttab.update (i, ()) seen); in (f {serial = i, name = name, prop = prop, body = body} x', seen') end); in fn bodies => fn x => #1 (fold app bodies (x, Inttab.empty)) end; (* proof body *) val unions_oracles = Ord_List.unions oracle_ord; val unions_thms = Ord_List.unions thm_ord; fun no_proof_body proof = PBody {oracles = [], thms = [], proof = proof}; val no_thm_body = thm_body (no_proof_body MinProof); fun no_thm_names (Abst (x, T, prf)) = Abst (x, T, no_thm_names prf) | no_thm_names (AbsP (x, t, prf)) = AbsP (x, t, no_thm_names prf) | no_thm_names (prf % t) = no_thm_names prf % t | no_thm_names (prf1 %% prf2) = no_thm_names prf1 %% no_thm_names prf2 | no_thm_names (PThm ({serial, pos, theory_name, name = _, prop, types}, thm_body)) = PThm (thm_header serial pos theory_name "" prop types, thm_body) | no_thm_names a = a; fun no_thm_proofs (Abst (x, T, prf)) = Abst (x, T, no_thm_proofs prf) | no_thm_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_thm_proofs prf) | no_thm_proofs (prf % t) = no_thm_proofs prf % t | no_thm_proofs (prf1 %% prf2) = no_thm_proofs prf1 %% no_thm_proofs prf2 | no_thm_proofs (PThm (header, _)) = PThm (header, no_thm_body) | no_thm_proofs a = a; fun no_body_proofs (Abst (x, T, prf)) = Abst (x, T, no_body_proofs prf) | no_body_proofs (AbsP (x, t, prf)) = AbsP (x, t, no_body_proofs prf) | no_body_proofs (prf % t) = no_body_proofs prf % t | no_body_proofs (prf1 %% prf2) = no_body_proofs prf1 %% no_body_proofs prf2 | no_body_proofs (PThm (header, Thm_Body {open_proof, body})) = let val body' = Future.value (no_proof_body (join_proof body)); val thm_body' = Thm_Body {open_proof = open_proof, body = body'}; in PThm (header, thm_body') end | no_body_proofs a = a; (** XML data representation **) (* encode *) local open XML.Encode Term_XML.Encode; fun proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, b, c) => ([a], pair (option typ) (proof consts) (b, c)), fn AbsP (a, b, c) => ([a], pair (option (term consts)) (proof consts) (b, c)), fn a % b => ([], pair (proof consts) (option (term consts)) (a, b)), fn a %% b => ([], pair (proof consts) (proof consts) (a, b)), fn Hyp a => ([], term consts a), fn PAxm (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PClass (a, b) => ([b], typ a), fn Oracle (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PThm ({serial, pos, theory_name, name, prop, types}, Thm_Body {open_proof, body, ...}) => ([int_atom serial, theory_name, name], pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) (map Position.properties_of pos, (prop, (types, map_proof_of open_proof (Future.join body)))))] and proof_body consts (PBody {oracles, thms, proof = prf}) = triple (list (pair (pair string (properties o Position.properties_of)) (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf) and thm consts (a, thm_node) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node, (Future.join (thm_node_body thm_node)))))); fun standard_term consts t = t |> variant [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), fn Free (a, _) => ([a], []), fn Var (a, _) => (indexname a, []), fn Bound a => ([], int a), fn Abs (a, b, c) => ([a], pair typ (standard_term consts) (b, c)), fn op $ a => ([], pair (standard_term consts) (standard_term consts) a)]; fun standard_proof consts prf = prf |> variant [fn MinProof => ([], []), fn PBound a => ([], int a), fn Abst (a, SOME b, c) => ([a], pair typ (standard_proof consts) (b, c)), fn AbsP (a, SOME b, c) => ([a], pair (standard_term consts) (standard_proof consts) (b, c)), fn a % SOME b => ([], pair (standard_proof consts) (standard_term consts) (a, b)), fn a %% b => ([], pair (standard_proof consts) (standard_proof consts) (a, b)), fn Hyp a => ([], standard_term consts a), fn PAxm (name, _, SOME Ts) => ([name], list typ Ts), fn PClass (T, c) => ([c], typ T), fn Oracle (name, prop, SOME Ts) => ([name], pair (standard_term consts) (list typ) (prop, Ts)), fn PThm ({serial, theory_name, name, types = SOME Ts, ...}, _) => ([int_atom serial, theory_name, name], list typ Ts)]; in val encode = proof; val encode_body = proof_body; val encode_standard_term = standard_term; val encode_standard_proof = standard_proof; end; (* decode *) local open XML.Decode Term_XML.Decode; fun proof consts prf = prf |> variant [fn ([], []) => MinProof, fn ([], a) => PBound (int a), fn ([a], b) => let val (c, d) = pair (option typ) (proof consts) b in Abst (a, c, d) end, fn ([a], b) => let val (c, d) = pair (option (term consts)) (proof consts) b in AbsP (a, c, d) end, fn ([], a) => op % (pair (proof consts) (option (term consts)) a), fn ([], a) => op %% (pair (proof consts) (proof consts) a), fn ([], a) => Hyp (term consts a), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in PAxm (a, c, d) end, fn ([b], a) => PClass (typ a, b), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in Oracle (a, c, d) end, fn ([a, b, c], d) => let val ((e, (f, (g, h)))) = pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) d; val header = thm_header (int_atom a) (map Position.of_properties e) b c f g; in PThm (header, thm_body h) end] and proof_body consts x = let val (a, b, c) = triple (list (pair (pair string (Position.of_properties o properties)) (option (term consts)))) (list (thm consts)) (proof consts) x; in PBody {oracles = a, thms = b, proof = c} end and thm consts x = let val (a, (b, (c, (d, e)))) = pair int (pair string (pair string (pair (term consts) (proof_body consts)))) x in (a, make_thm_node b c d (Future.value e) no_export) end; in val decode = proof; val decode_body = proof_body; end; (** proof objects with different levels of detail **) val proofs = Unsynchronized.ref 2; fun proofs_enabled () = ! proofs >= 2; fun atomic_proof prf = (case prf of Abst _ => false | AbsP _ => false | op % _ => false | op %% _ => false | MinProof => false | _ => true); fun compact_proof (prf % _) = compact_proof prf | compact_proof (prf1 %% prf2) = atomic_proof prf2 andalso compact_proof prf1 | compact_proof prf = atomic_proof prf; fun (prf %> t) = prf % SOME t; val proof_combt = Library.foldl (op %>); val proof_combt' = Library.foldl (op %); val proof_combP = Library.foldl (op %%); fun strip_combt prf = let fun stripc (prf % t, ts) = stripc (prf, t::ts) | stripc x = x in stripc (prf, []) end; fun strip_combP prf = let fun stripc (prf %% prf', prfs) = stripc (prf, prf'::prfs) | stripc x = x in stripc (prf, []) end; fun strip_thm_body (body as PBody {proof, ...}) = (case fst (strip_combt (fst (strip_combP proof))) of PThm (_, Thm_Body {body = body', ...}) => Future.join body' | _ => body); val mk_Abst = fold_rev (fn (x, _: typ) => fn prf => Abst (x, NONE, prf)); val mk_AbsP = fold_rev (fn _: term => fn prf => AbsP ("H", NONE, prf)); fun map_proof_same term typ ofclass = let val typs = Same.map typ; fun proof (Abst (s, T, prf)) = (Abst (s, Same.map_option typ T, Same.commit proof prf) handle Same.SAME => Abst (s, T, proof prf)) | proof (AbsP (s, t, prf)) = (AbsP (s, Same.map_option term t, Same.commit proof prf) handle Same.SAME => AbsP (s, t, proof prf)) | proof (prf % t) = (proof prf % Same.commit (Same.map_option term) t handle Same.SAME => prf % Same.map_option term t) | proof (prf1 %% prf2) = (proof prf1 %% Same.commit proof prf2 handle Same.SAME => prf1 %% proof prf2) | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) | proof (PClass T_c) = ofclass T_c | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) | proof (PThm ({serial, pos, theory_name, name, prop, types = SOME Ts}, thm_body)) = PThm (thm_header serial pos theory_name name prop (SOME (typs Ts)), thm_body) | proof _ = raise Same.SAME; in proof end; fun map_proof_terms_same term typ = map_proof_same term typ (fn (T, c) => PClass (typ T, c)); fun map_proof_types_same typ = map_proof_terms_same (Term_Subst.map_types_same typ) typ; fun same eq f x = let val x' = f x in if eq (x, x') then raise Same.SAME else x' end; fun map_proof_terms f g = Same.commit (map_proof_terms_same (same (op =) f) (same (op =) g)); fun map_proof_types f = Same.commit (map_proof_types_same (same (op =) f)); fun fold_proof_terms f (Abst (_, _, prf)) = fold_proof_terms f prf | fold_proof_terms f (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms f prf | fold_proof_terms f (AbsP (_, NONE, prf)) = fold_proof_terms f prf | fold_proof_terms f (prf % SOME t) = fold_proof_terms f prf #> f t | fold_proof_terms f (prf % NONE) = fold_proof_terms f prf | fold_proof_terms f (prf1 %% prf2) = fold_proof_terms f prf1 #> fold_proof_terms f prf2 | fold_proof_terms _ _ = I; fun fold_proof_terms_types f g (Abst (_, SOME T, prf)) = g T #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (Abst (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, SOME t, prf)) = f t #> fold_proof_terms_types f g prf | fold_proof_terms_types f g (AbsP (_, NONE, prf)) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf % SOME t) = fold_proof_terms_types f g prf #> f t | fold_proof_terms_types f g (prf % NONE) = fold_proof_terms_types f g prf | fold_proof_terms_types f g (prf1 %% prf2) = fold_proof_terms_types f g prf1 #> fold_proof_terms_types f g prf2 | fold_proof_terms_types _ g (PAxm (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PClass (T, _)) = g T | fold_proof_terms_types _ g (Oracle (_, _, SOME Ts)) = fold g Ts | fold_proof_terms_types _ g (PThm ({types = SOME Ts, ...}, _)) = fold g Ts | fold_proof_terms_types _ _ _ = I; fun maxidx_proof prf = fold_proof_terms_types Term.maxidx_term Term.maxidx_typ prf; fun size_of_proof (Abst (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (AbsP (_, _, prf)) = 1 + size_of_proof prf | size_of_proof (prf % _) = 1 + size_of_proof prf | size_of_proof (prf1 %% prf2) = size_of_proof prf1 + size_of_proof prf2 | size_of_proof _ = 1; fun change_types types (PAxm (name, prop, _)) = PAxm (name, prop, types) | change_types (SOME [T]) (PClass (_, c)) = PClass (T, c) | change_types types (Oracle (name, prop, _)) = Oracle (name, prop, types) | change_types types (PThm ({serial, pos, theory_name, name, prop, types = _}, thm_body)) = PThm (thm_header serial pos theory_name name prop types, thm_body) | change_types _ prf = prf; (* utilities *) fun strip_abs (_::Ts) (Abs (_, _, t)) = strip_abs Ts t | strip_abs _ t = t; fun mk_abs Ts t = Library.foldl (fn (t', T) => Abs ("", T, t')) (t, Ts); (*Abstraction of a proof term over its occurrences of v, which must contain no loose bound variables. The resulting proof term is ready to become the body of an Abst.*) fun prf_abstract_over v = let fun abst' lev u = if v aconv u then Bound lev else (case u of Abs (a, T, t) => Abs (a, T, abst' (lev + 1) t) | f $ t => (abst' lev f $ absth' lev t handle Same.SAME => f $ abst' lev t) | _ => raise Same.SAME) and absth' lev t = (abst' lev t handle Same.SAME => t); fun abst lev (AbsP (a, t, prf)) = (AbsP (a, Same.map_option (abst' lev) t, absth lev prf) handle Same.SAME => AbsP (a, t, abst lev prf)) | abst lev (Abst (a, T, prf)) = Abst (a, T, abst (lev + 1) prf) | abst lev (prf1 %% prf2) = (abst lev prf1 %% absth lev prf2 handle Same.SAME => prf1 %% abst lev prf2) | abst lev (prf % t) = (abst lev prf % Option.map (absth' lev) t handle Same.SAME => prf % Same.map_option (abst' lev) t) | abst _ _ = raise Same.SAME and absth lev prf = (abst lev prf handle Same.SAME => prf); in absth 0 end; (*increments a proof term's non-local bound variables required when moving a proof term within abstractions inc is increment for bound variables lev is level at which a bound variable is considered 'loose'*) fun incr_bv' inct tlev t = incr_bv (inct, tlev, t); fun prf_incr_bv' incP _ Plev _ (PBound i) = if i >= Plev then PBound (i+incP) else raise Same.SAME | prf_incr_bv' incP inct Plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (same (op =) (incr_bv' inct tlev)) t, prf_incr_bv incP inct (Plev+1) tlev body) handle Same.SAME => AbsP (a, t, prf_incr_bv' incP inct (Plev+1) tlev body)) | prf_incr_bv' incP inct Plev tlev (Abst (a, T, body)) = Abst (a, T, prf_incr_bv' incP inct Plev (tlev+1) body) | prf_incr_bv' incP inct Plev tlev (prf %% prf') = (prf_incr_bv' incP inct Plev tlev prf %% prf_incr_bv incP inct Plev tlev prf' handle Same.SAME => prf %% prf_incr_bv' incP inct Plev tlev prf') | prf_incr_bv' incP inct Plev tlev (prf % t) = (prf_incr_bv' incP inct Plev tlev prf % Option.map (incr_bv' inct tlev) t handle Same.SAME => prf % Same.map_option (same (op =) (incr_bv' inct tlev)) t) | prf_incr_bv' _ _ _ _ _ = raise Same.SAME and prf_incr_bv incP inct Plev tlev prf = (prf_incr_bv' incP inct Plev tlev prf handle Same.SAME => prf); fun incr_pboundvars 0 0 prf = prf | incr_pboundvars incP inct prf = prf_incr_bv incP inct 0 0 prf; fun prf_loose_bvar1 (prf1 %% prf2) k = prf_loose_bvar1 prf1 k orelse prf_loose_bvar1 prf2 k | prf_loose_bvar1 (prf % SOME t) k = prf_loose_bvar1 prf k orelse loose_bvar1 (t, k) | prf_loose_bvar1 (_ % NONE) _ = true | prf_loose_bvar1 (AbsP (_, SOME t, prf)) k = loose_bvar1 (t, k) orelse prf_loose_bvar1 prf k | prf_loose_bvar1 (AbsP (_, NONE, _)) _ = true | prf_loose_bvar1 (Abst (_, _, prf)) k = prf_loose_bvar1 prf (k+1) | prf_loose_bvar1 _ _ = false; fun prf_loose_Pbvar1 (PBound i) k = i = k | prf_loose_Pbvar1 (prf1 %% prf2) k = prf_loose_Pbvar1 prf1 k orelse prf_loose_Pbvar1 prf2 k | prf_loose_Pbvar1 (prf % _) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 (AbsP (_, _, prf)) k = prf_loose_Pbvar1 prf (k+1) | prf_loose_Pbvar1 (Abst (_, _, prf)) k = prf_loose_Pbvar1 prf k | prf_loose_Pbvar1 _ _ = false; fun prf_add_loose_bnos plev _ (PBound i) (is, js) = if i < plev then (is, js) else (insert (op =) (i-plev) is, js) | prf_add_loose_bnos plev tlev (prf1 %% prf2) p = prf_add_loose_bnos plev tlev prf2 (prf_add_loose_bnos plev tlev prf1 p) | prf_add_loose_bnos plev tlev (prf % opt) (is, js) = prf_add_loose_bnos plev tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (AbsP (_, opt, prf)) (is, js) = prf_add_loose_bnos (plev+1) tlev prf (case opt of NONE => (is, insert (op =) ~1 js) | SOME t => (is, add_loose_bnos (t, tlev, js))) | prf_add_loose_bnos plev tlev (Abst (_, _, prf)) p = prf_add_loose_bnos plev (tlev+1) prf p | prf_add_loose_bnos _ _ _ _ = ([], []); (* substitutions *) local fun conflicting_tvarsT envT = Term.fold_atyps (fn T => fn instT => (case T of TVar (v as (_, S)) => if TVars.defined instT v orelse can (Type.lookup envT) v then instT else TVars.add (v, Logic.dummy_tfree S) instT | _ => instT)); fun conflicting_tvars env = Term.fold_aterms (fn t => fn inst => (case t of Var (v as (_, T)) => if Vars.defined inst v orelse can (Envir.lookup env) v then inst else Vars.add (v, Free ("dummy", T)) inst | _ => inst)); fun del_conflicting_tvars envT ty = Term_Subst.instantiateT (TVars.build (conflicting_tvarsT envT ty)) ty; fun del_conflicting_vars env tm = let val instT = TVars.build (tm |> Term.fold_types (conflicting_tvarsT (Envir.type_env env))); val inst = Vars.build (tm |> conflicting_tvars env); in Term_Subst.instantiate (instT, inst) tm end; in fun norm_proof env = let val envT = Envir.type_env env; fun msg s = warning ("type conflict in norm_proof:\n" ^ s); fun htype f t = f env t handle TYPE (s, _, _) => (msg s; f env (del_conflicting_vars env t)); fun htypeT f T = f envT T handle TYPE (s, _, _) => (msg s; f envT (del_conflicting_tvars envT T)); fun htypeTs f Ts = f envT Ts handle TYPE (s, _, _) => (msg s; f envT (map (del_conflicting_tvars envT) Ts)); fun norm (Abst (s, T, prf)) = (Abst (s, Same.map_option (htypeT Envir.norm_type_same) T, Same.commit norm prf) handle Same.SAME => Abst (s, T, norm prf)) | norm (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (htype Envir.norm_term_same) t, Same.commit norm prf) handle Same.SAME => AbsP (s, t, norm prf)) | norm (prf % t) = (norm prf % Option.map (htype Envir.norm_term) t handle Same.SAME => prf % Same.map_option (htype Envir.norm_term_same) t) | norm (prf1 %% prf2) = (norm prf1 %% Same.commit norm prf2 handle Same.SAME => prf1 %% norm prf2) | norm (PAxm (s, prop, Ts)) = PAxm (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PClass (T, c)) = PClass (htypeT Envir.norm_type_same T, c) | norm (Oracle (s, prop, Ts)) = Oracle (s, prop, Same.map_option (htypeTs Envir.norm_types_same) Ts) | norm (PThm ({serial = i, pos = p, theory_name, name = a, prop = t, types = Ts}, thm_body)) = PThm (thm_header i p theory_name a t (Same.map_option (htypeTs Envir.norm_types_same) Ts), thm_body) | norm _ = raise Same.SAME; in Same.commit norm end; end; (* remove some types in proof term (to save space) *) fun remove_types (Abs (s, _, t)) = Abs (s, dummyT, remove_types t) | remove_types (t $ u) = remove_types t $ remove_types u | remove_types (Const (s, _)) = Const (s, dummyT) | remove_types t = t; fun remove_types_env (Envir.Envir {maxidx, tenv, tyenv}) = Envir.Envir {maxidx = maxidx, tenv = Vartab.map (K (apsnd remove_types)) tenv, tyenv = tyenv}; fun norm_proof' env prf = norm_proof (remove_types_env env) prf; (* substitution of bound variables *) fun prf_subst_bounds args prf = let val n = length args; fun subst' lev (Bound i) = (if i Bound (i-n)) (*loose: change it*) | subst' lev (Abs (a, T, body)) = Abs (a, T, subst' (lev+1) body) | subst' lev (f $ t) = (subst' lev f $ substh' lev t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME and substh' lev t = (subst' lev t handle Same.SAME => t); fun subst lev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' lev) t, substh lev body) handle Same.SAME => AbsP (a, t, subst lev body)) | subst lev (Abst (a, T, body)) = Abst (a, T, subst (lev+1) body) | subst lev (prf %% prf') = (subst lev prf %% substh lev prf' handle Same.SAME => prf %% subst lev prf') | subst lev (prf % t) = (subst lev prf % Option.map (substh' lev) t handle Same.SAME => prf % Same.map_option (subst' lev) t) | subst _ _ = raise Same.SAME and substh lev prf = (subst lev prf handle Same.SAME => prf); in (case args of [] => prf | _ => substh 0 prf) end; fun prf_subst_pbounds args prf = let val n = length args; fun subst (PBound i) Plev tlev = (if i < Plev then raise Same.SAME (*var is locally bound*) else incr_pboundvars Plev tlev (nth args (i-Plev)) handle General.Subscript => PBound (i-n) (*loose: change it*)) | subst (AbsP (a, t, body)) Plev tlev = AbsP (a, t, subst body (Plev+1) tlev) | subst (Abst (a, T, body)) Plev tlev = Abst (a, T, subst body Plev (tlev+1)) | subst (prf %% prf') Plev tlev = (subst prf Plev tlev %% substh prf' Plev tlev handle Same.SAME => prf %% subst prf' Plev tlev) | subst (prf % t) Plev tlev = subst prf Plev tlev % t | subst _ _ _ = raise Same.SAME and substh prf Plev tlev = (subst prf Plev tlev handle Same.SAME => prf) in (case args of [] => prf | _ => substh prf 0 0) end; (* freezing and thawing of variables in proof terms *) local fun frzT names = map_type_tvar (fn (ixn, S) => TFree (the (AList.lookup (op =) names ixn), S)); fun thawT names = map_type_tfree (fn (a, S) => (case AList.lookup (op =) names a of NONE => TFree (a, S) | SOME ixn => TVar (ixn, S))); fun freeze names names' (t $ u) = freeze names names' t $ freeze names names' u | freeze names names' (Abs (s, T, t)) = Abs (s, frzT names' T, freeze names names' t) | freeze _ names' (Const (s, T)) = Const (s, frzT names' T) | freeze _ names' (Free (s, T)) = Free (s, frzT names' T) | freeze names names' (Var (ixn, T)) = Free (the (AList.lookup (op =) names ixn), frzT names' T) | freeze _ _ t = t; fun thaw names names' (t $ u) = thaw names names' t $ thaw names names' u | thaw names names' (Abs (s, T, t)) = Abs (s, thawT names' T, thaw names names' t) | thaw _ names' (Const (s, T)) = Const (s, thawT names' T) | thaw names names' (Free (s, T)) = let val T' = thawT names' T in (case AList.lookup (op =) names s of NONE => Free (s, T') | SOME ixn => Var (ixn, T')) end | thaw _ names' (Var (ixn, T)) = Var (ixn, thawT names' T) | thaw _ _ t = t; in fun freeze_thaw_prf prf = let val (fs, Tfs, vs, Tvs) = fold_proof_terms_types (fn t => fn (fs, Tfs, vs, Tvs) => (Term.add_free_names t fs, Term.add_tfree_names t Tfs, Term.add_var_names t vs, Term.add_tvar_names t Tvs)) (fn T => fn (fs, Tfs, vs, Tvs) => (fs, Term.add_tfree_namesT T Tfs, vs, Term.add_tvar_namesT T Tvs)) prf ([], [], [], []); val names = vs ~~ Name.variant_list fs (map fst vs); val names' = Tvs ~~ Name.variant_list Tfs (map fst Tvs); val rnames = map swap names; val rnames' = map swap names'; in (map_proof_terms (freeze names names') (frzT names') prf, map_proof_terms (thaw rnames rnames') (thawT rnames')) end; end; (** inference rules **) (* trivial implication *) val trivial_proof = AbsP ("H", NONE, PBound 0); (* implication introduction *) fun gen_implies_intr_proof f h prf = let fun abshyp i (Hyp t) = if h aconv t then PBound i else raise Same.SAME | abshyp i (Abst (s, T, prf)) = Abst (s, T, abshyp i prf) | abshyp i (AbsP (s, t, prf)) = AbsP (s, t, abshyp (i + 1) prf) | abshyp i (prf % t) = abshyp i prf % t | abshyp i (prf1 %% prf2) = (abshyp i prf1 %% abshyph i prf2 handle Same.SAME => prf1 %% abshyp i prf2) | abshyp _ _ = raise Same.SAME and abshyph i prf = (abshyp i prf handle Same.SAME => prf); in AbsP ("H", f h, abshyph 0 prf) end; val implies_intr_proof = gen_implies_intr_proof (K NONE); val implies_intr_proof' = gen_implies_intr_proof SOME; (* forall introduction *) fun forall_intr_proof (a, v) opt_T prf = Abst (a, opt_T, prf_abstract_over v prf); fun forall_intr_proof' v prf = let val (a, T) = (case v of Var ((a, _), T) => (a, T) | Free (a, T) => (a, T)) in forall_intr_proof (a, v) (SOME T) prf end; (* varify *) fun varify_proof t fixed prf = let val fs = build (t |> (Term.fold_types o Term.fold_atyps) (fn TFree v => if TFrees.defined fixed v then I else insert (op =) v | _ => I)); val used = Name.build_context (t |> (fold_types o fold_atyps) (fn TVar ((a, _), _) => Name.declare a | _ => I)); 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 Names.is_empty frees then [] else fold_aterms (fn Free (x, T) => Names.defined frees x ? insert (op =) (x, T) | _ => I) (Term_Subst.generalize (tfrees, Names.empty) idx prop) []; in prf |> Same.commit (map_proof_terms_same (Term_Subst.generalize_same (tfrees, Names.empty) idx) (Term_Subst.generalizeT_same tfrees idx)) |> fold (fn (x, T) => forall_intr_proof (x, Free (x, T)) NONE) gen |> fold_rev (fn (x, T) => fn prf' => prf' %> Var (Name.clean_index (x, idx), T)) gen end; (* instantiation *) fun instantiate (instT, inst) = Same.commit (map_proof_terms_same (Term_Subst.instantiate_same (instT, Vars.map (K remove_types) inst)) (Term_Subst.instantiateT_same instT)); (* lifting *) fun lift_proof Bi inc prop prf = let fun lift'' Us Ts t = strip_abs Ts (Logic.incr_indexes ([], Us, inc) (mk_abs Ts t)); fun lift' Us Ts (Abst (s, T, prf)) = (Abst (s, Same.map_option (Logic.incr_tvar_same inc) T, lifth' Us (dummyT::Ts) prf) handle Same.SAME => Abst (s, T, lift' Us (dummyT::Ts) prf)) | lift' Us Ts (AbsP (s, t, prf)) = (AbsP (s, Same.map_option (same (op =) (lift'' Us Ts)) t, lifth' Us Ts prf) handle Same.SAME => AbsP (s, t, lift' Us Ts prf)) | lift' Us Ts (prf % t) = (lift' Us Ts prf % Option.map (lift'' Us Ts) t handle Same.SAME => prf % Same.map_option (same (op =) (lift'' Us Ts)) t) | lift' Us Ts (prf1 %% prf2) = (lift' Us Ts prf1 %% lifth' Us Ts prf2 handle Same.SAME => prf1 %% lift' Us Ts prf2) | lift' _ _ (PAxm (s, prop, Ts)) = PAxm (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PClass (T, c)) = PClass (Logic.incr_tvar_same inc T, c) | lift' _ _ (Oracle (s, prop, Ts)) = Oracle (s, prop, (Same.map_option o Same.map) (Logic.incr_tvar_same inc) Ts) | lift' _ _ (PThm ({serial = i, pos = p, theory_name, name = s, prop, types = Ts}, thm_body)) = PThm (thm_header i p theory_name s prop ((Same.map_option o Same.map) (Logic.incr_tvar inc) Ts), thm_body) | lift' _ _ _ = raise Same.SAME and lifth' Us Ts prf = (lift' Us Ts prf handle Same.SAME => prf); val ps = map (Logic.lift_all inc Bi) (Logic.strip_imp_prems prop); val k = length ps; fun mk_app b (i, j, prf) = if b then (i-1, j, prf %% PBound i) else (i, j-1, prf %> Bound j); fun lift Us bs i j (Const ("Pure.imp", _) $ A $ B) = AbsP ("H", NONE (*A*), lift Us (true::bs) (i+1) j B) | lift Us bs i j (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), lift (T::Us) (false::bs) i (j+1) t) | lift Us bs i j _ = proof_combP (lifth' (rev Us) [] prf, map (fn k => (#3 (fold_rev mk_app bs (i-1, j-1, PBound k)))) (i + k - 1 downto i)); in mk_AbsP ps (lift [] [] 0 0 Bi) end; fun incr_indexes i = Same.commit (map_proof_terms_same (Logic.incr_indexes_same ([], [], i)) (Logic.incr_tvar_same i)); (* proof by assumption *) fun mk_asm_prf t i m = let fun imp_prf _ i 0 = PBound i | imp_prf (Const ("Pure.imp", _) $ A $ B) i m = AbsP ("H", NONE (*A*), imp_prf B (i+1) (m-1)) | imp_prf _ i _ = PBound i; fun all_prf (Const ("Pure.all", _) $ Abs (a, T, t)) = Abst (a, NONE (*T*), all_prf t) | all_prf t = imp_prf t (~i) m in all_prf t end; fun assumption_proof Bs Bi n prf = mk_AbsP Bs (proof_combP (prf, map PBound (length Bs - 1 downto 0) @ [mk_asm_prf Bi n ~1])); (* composition of object rule with proof state *) fun flatten_params_proof i j n (Const ("Pure.imp", _) $ A $ B, k) = AbsP ("H", NONE (*A*), flatten_params_proof (i+1) j n (B, k)) | flatten_params_proof i j n (Const ("Pure.all", _) $ Abs (a, T, t), k) = Abst (a, NONE (*T*), flatten_params_proof i (j+1) n (t, k)) | flatten_params_proof i j n (_, k) = proof_combP (proof_combt (PBound (k+i), map Bound (j-1 downto 0)), map PBound (remove (op =) (i-n) (i-1 downto 0))); fun bicompose_proof flatten Bs As A oldAs n m rprf sprf = let val lb = length Bs; val la = length As; in mk_AbsP (Bs @ As) (proof_combP (sprf, map PBound (lb + la - 1 downto la)) %% proof_combP (rprf, (if n>0 then [mk_asm_prf (the A) n m] else []) @ map (if flatten then flatten_params_proof 0 0 n else PBound o snd) (oldAs ~~ (la - 1 downto 0)))) end; (** type classes **) fun strip_shyps_proof algebra present witnessed extra prf = let - val replacements = present @ witnessed @ map (`Logic.dummy_tfree) extra; + val replacements = present @ witnessed @ map (`Logic.dummy_tfree) (Sortset.dest extra); fun get_replacement S = replacements |> get_first (fn (T', S') => if Sorts.sort_le algebra (S', S) then SOME T' else NONE); fun replace T = if exists (fn (T', _) => T' = T) present then raise Same.SAME else (case get_replacement (Type.sort_of_atyp T) of SOME T' => T' | NONE => raise Fail "strip_shyps_proof: bad type variable in proof term"); in Same.commit (map_proof_types_same (Term_Subst.map_atypsT_same replace)) prf end; fun of_sort_proof algebra classrel_proof arity_proof hyps = Sorts.of_sort_derivation algebra {class_relation = fn _ => fn _ => fn (prf, c1) => fn c2 => if c1 = c2 then prf else classrel_proof (c1, c2) %% prf, type_constructor = fn (a, _) => fn dom => fn c => let val Ss = map (map snd) dom and prfs = maps (map fst) dom in proof_combP (arity_proof (a, Ss, c), prfs) end, type_variable = fn typ => map (fn c => (hyps (typ, c), c)) (Type.sort_of_atyp typ)}; (** axioms and theorems **) val add_type_variables = (fold_types o fold_atyps) (insert (op =)); fun type_variables_of t = rev (add_type_variables t []); val add_variables = fold_aterms (fn a => (is_Var a orelse is_Free a) ? insert (op =) a); fun variables_of t = rev (add_variables t []); fun test_args _ [] = true | test_args is (Bound i :: ts) = not (member (op =) is i) andalso test_args (i :: is) ts | test_args _ _ = false; fun is_fun (Type ("fun", _)) = true | is_fun (TVar _) = true | is_fun _ = false; fun vars_of t = map Var (build_rev (Term.add_vars t)); fun add_funvars Ts (vs, t) = if is_fun (fastype_of1 (Ts, t)) then union (op =) vs (map_filter (fn Var (ixn, T) => if is_fun T then SOME ixn else NONE | _ => NONE) (vars_of t)) else vs; fun add_npvars q p Ts (vs, Const ("Pure.imp", _) $ t $ u) = add_npvars q p Ts (add_npvars q (not p) Ts (vs, t), u) | add_npvars q p Ts (vs, Const ("Pure.all", Type (_, [Type (_, [T, _]), _])) $ t) = add_npvars q p Ts (vs, if p andalso q then betapply (t, Var (("",0), T)) else t) | add_npvars q p Ts (vs, Abs (_, T, t)) = add_npvars q p (T::Ts) (vs, t) | add_npvars _ _ Ts (vs, t) = add_npvars' Ts (vs, t) and add_npvars' Ts (vs, t) = (case strip_comb t of (Var (ixn, _), ts) => if test_args [] ts then vs else Library.foldl (add_npvars' Ts) (AList.update (op =) (ixn, Library.foldl (add_funvars Ts) ((these ooo AList.lookup) (op =) vs ixn, ts)) vs, ts) | (Abs (_, T, u), ts) => Library.foldl (add_npvars' (T::Ts)) (vs, u :: ts) | (_, ts) => Library.foldl (add_npvars' Ts) (vs, ts)); fun prop_vars (Const ("Pure.imp", _) $ P $ Q) = union (op =) (prop_vars P) (prop_vars Q) | prop_vars (Const ("Pure.all", _) $ Abs (_, _, t)) = prop_vars t | prop_vars t = (case strip_comb t of (Var (ixn, _), _) => [ixn] | _ => []); fun is_proj t = let fun is_p i t = (case strip_comb t of (Bound _, []) => false | (Bound j, ts) => j >= i orelse exists (is_p i) ts | (Abs (_, _, u), _) => is_p (i+1) u | (_, ts) => exists (is_p i) ts) in (case strip_abs_body t of Bound _ => true | t' => is_p 0 t') end; fun prop_args prop = let val needed_vars = union (op =) (Library.foldl (uncurry (union (op =))) ([], map (uncurry (insert (op =))) (add_npvars true true [] ([], prop)))) (prop_vars prop); in variables_of prop |> map (fn var as Var (ixn, _) => if member (op =) needed_vars ixn then SOME var else NONE | free => SOME free) end; fun const_proof mk name prop = let val args = prop_args prop; - val ({outer_constraints, ...}, prop1) = Logic.unconstrainT [] prop; + val ({outer_constraints, ...}, prop1) = Logic.unconstrainT Sortset.empty prop; val head = mk (name, prop1, NONE); in proof_combP (proof_combt' (head, args), map PClass outer_constraints) end; val axm_proof = const_proof PAxm; val oracle_proof = const_proof Oracle; val shrink_proof = let fun shrink ls lev (prf as Abst (a, T, body)) = let val (b, is, ch, body') = shrink ls (lev+1) body in (b, is, ch, if ch then Abst (a, T, body') else prf) end | shrink ls lev (prf as AbsP (a, t, body)) = let val (b, is, ch, body') = shrink (lev::ls) lev body in (b orelse member (op =) is 0, map_filter (fn 0 => NONE | i => SOME (i-1)) is, ch, if ch then AbsP (a, t, body') else prf) end | shrink ls lev prf = let val (is, ch, _, prf') = shrink' ls lev [] [] prf in (false, is, ch, prf') end and shrink' ls lev ts prfs (prf as prf1 %% prf2) = let val p as (_, is', ch', prf') = shrink ls lev prf2; val (is, ch, ts', prf'') = shrink' ls lev ts (p::prfs) prf1 in (union (op =) is is', ch orelse ch', ts', if ch orelse ch' then prf'' %% prf' else prf) end | shrink' ls lev ts prfs (prf as prf1 % t) = let val (is, ch, (ch', t')::ts', prf') = shrink' ls lev (t::ts) prfs prf1 in (is, ch orelse ch', ts', if ch orelse ch' then prf' % t' else prf) end | shrink' ls lev ts prfs (prf as PBound i) = (if exists (fn SOME (Bound j) => lev-j <= nth ls i | _ => true) ts orelse has_duplicates (op =) (Library.foldl (fn (js, SOME (Bound j)) => j :: js | (js, _) => js) ([], ts)) orelse exists #1 prfs then [i] else [], false, map (pair false) ts, prf) | shrink' _ _ ts _ (Hyp t) = ([], false, map (pair false) ts, Hyp t) | shrink' _ _ ts _ (prf as MinProof) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts _ (prf as PClass _) = ([], false, map (pair false) ts, prf) | shrink' _ _ ts prfs prf = let val prop = (case prf of PAxm (_, prop, _) => prop | Oracle (_, prop, _) => prop | PThm ({prop, ...}, _) => prop | _ => raise Fail "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = chop (length vs) ts; val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts'; val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => insert (op =) ixn (case AList.lookup (op =) insts ixn of SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns' | _ => union (op =) ixns ixns')) (needed prop ts'' prfs, add_npvars false true [] ([], prop)); val insts' = map (fn (ixn, x as SOME _) => if member (op =) nvs ixn then (false, x) else (true, NONE) | (_, x) => (false, x)) insts in ([], false, insts' @ map (pair false) ts'', prf) end and needed (Const ("Pure.imp", _) $ t $ u) ts ((b, _, _, _)::prfs) = union (op =) (if b then map (fst o dest_Var) (vars_of t) else []) (needed u ts prfs) | needed (Var (ixn, _)) (_::_) _ = [ixn] | needed _ _ _ = []; in fn prf => #4 (shrink [] 0 prf) end; (** axioms for equality **) val aT = TFree ("'a", []); val bT = TFree ("'b", []); val x = Free ("x", aT); val y = Free ("y", aT); val z = Free ("z", aT); val A = Free ("A", propT); val B = Free ("B", propT); val f = Free ("f", aT --> bT); val g = Free ("g", aT --> bT); val equality_axms = [("reflexive", Logic.mk_equals (x, x)), ("symmetric", Logic.mk_implies (Logic.mk_equals (x, y), Logic.mk_equals (y, x))), ("transitive", Logic.list_implies ([Logic.mk_equals (x, y), Logic.mk_equals (y, z)], Logic.mk_equals (x, z))), ("equal_intr", Logic.list_implies ([Logic.mk_implies (A, B), Logic.mk_implies (B, A)], Logic.mk_equals (A, B))), ("equal_elim", Logic.list_implies ([Logic.mk_equals (A, B), A], B)), ("abstract_rule", Logic.mk_implies (Logic.all x (Logic.mk_equals (f $ x, g $ x)), Logic.mk_equals (lambda x (f $ x), lambda x (g $ x)))), ("combination", Logic.list_implies ([Logic.mk_equals (f, g), Logic.mk_equals (x, y)], Logic.mk_equals (f $ x, g $ y)))]; val [reflexive_axm, symmetric_axm, transitive_axm, equal_intr_axm, equal_elim_axm, abstract_rule_axm, combination_axm] = map (fn (s, t) => PAxm ("Pure." ^ s, Logic.varify_global t, NONE)) equality_axms; val reflexive_proof = reflexive_axm % NONE; val is_reflexive_proof = fn PAxm ("Pure.reflexive", _, _) % _ => true | _ => false; fun symmetric_proof prf = if is_reflexive_proof prf then prf else symmetric_axm % NONE % NONE %% prf; fun transitive_proof U u prf1 prf2 = if is_reflexive_proof prf1 then prf2 else if is_reflexive_proof prf2 then prf1 else if U = propT then transitive_axm % NONE % SOME (remove_types u) % NONE %% prf1 %% prf2 else transitive_axm % NONE % NONE % NONE %% prf1 %% prf2; fun equal_intr_proof A B prf1 prf2 = equal_intr_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun equal_elim_proof A B prf1 prf2 = equal_elim_axm %> remove_types A %> remove_types B %% prf1 %% prf2; fun abstract_rule_proof (a, x) prf = abstract_rule_axm % NONE % NONE %% forall_intr_proof (a, x) NONE prf; fun check_comb (PAxm ("Pure.combination", _, _) % f % _ % _ % _ %% prf %% _) = is_some f orelse check_comb prf | check_comb (PAxm ("Pure.transitive", _, _) % _ % _ % _ %% prf1 %% prf2) = check_comb prf1 andalso check_comb prf2 | check_comb (PAxm ("Pure.symmetric", _, _) % _ % _ %% prf) = check_comb prf | check_comb _ = false; fun combination_proof f g t u prf1 prf2 = let val f = Envir.beta_norm f; val g = Envir.beta_norm g; val prf = if check_comb prf1 then combination_axm % NONE % NONE else (case prf1 of PAxm ("Pure.reflexive", _, _) % _ => combination_axm %> remove_types f % NONE | _ => combination_axm %> remove_types f %> remove_types g) in prf % (case head_of f of Abs _ => SOME (remove_types t) | Var _ => SOME (remove_types t) | _ => NONE) % (case head_of g of Abs _ => SOME (remove_types u) | Var _ => SOME (remove_types u) | _ => NONE) %% prf1 %% prf2 end; (** rewriting on proof terms **) (* simple first order matching functions for terms and proofs (see pattern.ML) *) exception PMatch; fun flt (i: int) = filter (fn n => n < i); fun fomatch Ts tymatch j instsp p = let fun mtch (instsp as (tyinsts, insts)) = fn (Var (ixn, T), t) => if j>0 andalso not (null (flt j (loose_bnos t))) then raise PMatch else (tymatch (tyinsts, fn () => (T, fastype_of1 (Ts, t))), (ixn, t) :: insts) | (Free (a, T), Free (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (Const (a, T), Const (b, U)) => if a=b then (tymatch (tyinsts, K (T, U)), insts) else raise PMatch | (f $ t, g $ u) => mtch (mtch instsp (f, g)) (t, u) | (Bound i, Bound j) => if i=j then instsp else raise PMatch | _ => raise PMatch in mtch instsp (apply2 Envir.beta_eta_contract p) end; fun match_proof Ts tymatch = let fun optmatch _ inst (NONE, _) = inst | optmatch _ _ (SOME _, NONE) = raise PMatch | optmatch mtch inst (SOME x, SOME y) = mtch inst (x, y) fun matcht Ts j (pinst, tinst) (t, u) = (pinst, fomatch Ts tymatch j tinst (t, Envir.beta_norm u)); fun matchT (pinst, (tyinsts, insts)) p = (pinst, (tymatch (tyinsts, K p), insts)); fun matchTs inst (Ts, Us) = Library.foldl (uncurry matchT) (inst, Ts ~~ Us); fun mtch Ts i j (pinst, tinst) (Hyp (Var (ixn, _)), prf) = if i = 0 andalso j = 0 then ((ixn, prf) :: pinst, tinst) else (case apfst (flt i) (apsnd (flt j) (prf_add_loose_bnos 0 0 prf ([], []))) of ([], []) => ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) | ([], _) => if j = 0 then ((ixn, incr_pboundvars (~i) (~j) prf) :: pinst, tinst) else raise PMatch | _ => raise PMatch) | mtch Ts i j inst (prf1 % opt1, prf2 % opt2) = optmatch (matcht Ts j) (mtch Ts i j inst (prf1, prf2)) (opt1, opt2) | mtch Ts i j inst (prf1 %% prf2, prf1' %% prf2') = mtch Ts i j (mtch Ts i j inst (prf1, prf1')) (prf2, prf2') | mtch Ts i j inst (Abst (_, opT, prf1), Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) (optmatch matchT inst (opT, opU)) (prf1, prf2) | mtch Ts i j inst (prf1, Abst (_, opU, prf2)) = mtch (the_default dummyT opU :: Ts) i (j+1) inst (incr_pboundvars 0 1 prf1 %> Bound 0, prf2) | mtch Ts i j inst (AbsP (_, opt, prf1), AbsP (_, opu, prf2)) = mtch Ts (i+1) j (optmatch (matcht Ts j) inst (opt, opu)) (prf1, prf2) | mtch Ts i j inst (prf1, AbsP (_, _, prf2)) = mtch Ts (i+1) j inst (incr_pboundvars 1 0 prf1 %% PBound 0, prf2) | mtch Ts i j inst (PAxm (s1, _, opTs), PAxm (s2, _, opUs)) = if s1 = s2 then optmatch matchTs inst (opTs, opUs) else raise PMatch | mtch Ts i j inst (PClass (T1, c1), PClass (T2, c2)) = if c1 = c2 then matchT inst (T1, T2) else raise PMatch | mtch Ts i j inst (PThm ({name = name1, prop = prop1, types = types1, ...}, _), PThm ({name = name2, prop = prop2, types = types2, ...}, _)) = if name1 = name2 andalso prop1 = prop2 then optmatch matchTs inst (types1, types2) else raise PMatch | mtch _ _ _ inst (PBound i, PBound j) = if i = j then inst else raise PMatch | mtch _ _ _ _ _ = raise PMatch in mtch Ts 0 0 end; fun prf_subst (pinst, (tyinsts, insts)) = let val substT = Envir.subst_type_same tyinsts; val substTs = Same.map substT; fun subst' lev (Var (xi, _)) = (case AList.lookup (op =) insts xi of NONE => raise Same.SAME | SOME u => incr_boundvars lev u) | subst' _ (Const (s, T)) = Const (s, substT T) | subst' _ (Free (s, T)) = Free (s, substT T) | subst' lev (Abs (a, T, body)) = (Abs (a, substT T, Same.commit (subst' (lev + 1)) body) handle Same.SAME => Abs (a, T, subst' (lev + 1) body)) | subst' lev (f $ t) = (subst' lev f $ Same.commit (subst' lev) t handle Same.SAME => f $ subst' lev t) | subst' _ _ = raise Same.SAME; fun subst plev tlev (AbsP (a, t, body)) = (AbsP (a, Same.map_option (subst' tlev) t, Same.commit (subst (plev + 1) tlev) body) handle Same.SAME => AbsP (a, t, subst (plev + 1) tlev body)) | subst plev tlev (Abst (a, T, body)) = (Abst (a, Same.map_option substT T, Same.commit (subst plev (tlev + 1)) body) handle Same.SAME => Abst (a, T, subst plev (tlev + 1) body)) | subst plev tlev (prf %% prf') = (subst plev tlev prf %% Same.commit (subst plev tlev) prf' handle Same.SAME => prf %% subst plev tlev prf') | subst plev tlev (prf % t) = (subst plev tlev prf % Same.commit (Same.map_option (subst' tlev)) t handle Same.SAME => prf % Same.map_option (subst' tlev) t) | subst plev tlev (Hyp (Var (xi, _))) = (case AList.lookup (op =) pinst xi of NONE => raise Same.SAME | SOME prf' => incr_pboundvars plev tlev prf') | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (PClass (T, c)) = PClass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) | subst _ _ (PThm ({serial = i, pos = p, theory_name, name = id, prop, types}, thm_body)) = PThm (thm_header i p theory_name id prop (Same.map_option substTs types), thm_body) | subst _ _ _ = raise Same.SAME; in fn t => subst 0 0 t handle Same.SAME => t end; (*A fast unification filter: true unless the two terms cannot be unified. Terms must be NORMAL. Treats all Vars as distinct. *) fun could_unify prf1 prf2 = let fun matchrands (prf1 %% prf2) (prf1' %% prf2') = could_unify prf2 prf2' andalso matchrands prf1 prf1' | matchrands (prf % SOME t) (prf' % SOME t') = Term.could_unify (t, t') andalso matchrands prf prf' | matchrands (prf % _) (prf' % _) = matchrands prf prf' | matchrands _ _ = true fun head_of (prf %% _) = head_of prf | head_of (prf % _) = head_of prf | head_of prf = prf in case (head_of prf1, head_of prf2) of (_, Hyp (Var _)) => true | (Hyp (Var _), _) => true | (PAxm (a, _, _), PAxm (b, _, _)) => a = b andalso matchrands prf1 prf2 | (PClass (_, c), PClass (_, d)) => c = d andalso matchrands prf1 prf2 | (PThm ({name = a, prop = propa, ...}, _), PThm ({name = b, prop = propb, ...}, _)) => a = b andalso propa = propb andalso matchrands prf1 prf2 | (PBound i, PBound j) => i = j andalso matchrands prf1 prf2 | (AbsP _, _) => true (*because of possible eta equality*) | (Abst _, _) => true | (_, AbsP _) => true | (_, Abst _) => true | _ => false end; (* rewrite proof *) val no_skel = PBound 0; val normal_skel = Hyp (Var ((Name.uu, 0), propT)); fun rewrite_prf tymatch (rules, procs) prf = let fun rew _ _ (Abst (_, _, body) % SOME t) = SOME (prf_subst_bounds [t] body, no_skel) | rew _ _ (AbsP (_, _, body) %% prf) = SOME (prf_subst_pbounds [prf] body, no_skel) | rew Ts hs prf = (case get_first (fn r => r Ts hs prf) procs of NONE => get_first (fn (prf1, prf2) => SOME (prf_subst (match_proof Ts tymatch ([], (Vartab.empty, [])) (prf1, prf)) prf2, prf2) handle PMatch => NONE) (filter (could_unify prf o fst) rules) | some => some); fun rew0 Ts hs (prf as AbsP (_, _, prf' %% PBound 0)) = if prf_loose_Pbvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars (~1) 0 prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs (prf as Abst (_, _, prf' % SOME (Bound 0))) = if prf_loose_bvar1 prf' 0 then rew Ts hs prf else let val prf'' = incr_pboundvars 0 (~1) prf' in SOME (the_default (prf'', no_skel) (rew Ts hs prf'')) end | rew0 Ts hs prf = rew Ts hs prf; fun rew1 _ _ (Hyp (Var _)) _ = NONE | rew1 Ts hs skel prf = (case rew2 Ts hs skel prf of SOME prf1 => (case rew0 Ts hs prf1 of SOME (prf2, skel') => SOME (the_default prf2 (rew1 Ts hs skel' prf2)) | NONE => SOME prf1) | NONE => (case rew0 Ts hs prf of SOME (prf1, skel') => SOME (the_default prf1 (rew1 Ts hs skel' prf1)) | NONE => NONE)) and rew2 Ts hs skel (prf % SOME t) = (case prf of Abst (_, _, body) => let val prf' = prf_subst_bounds [t] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => (case rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf of SOME prf' => SOME (prf' % SOME t) | NONE => NONE)) | rew2 Ts hs skel (prf % NONE) = Option.map (fn prf' => prf' % NONE) (rew1 Ts hs (case skel of skel' % _ => skel' | _ => no_skel) prf) | rew2 Ts hs skel (prf1 %% prf2) = (case prf1 of AbsP (_, _, body) => let val prf' = prf_subst_pbounds [prf2] body in SOME (the_default prf' (rew2 Ts hs no_skel prf')) end | _ => let val (skel1, skel2) = (case skel of skel1 %% skel2 => (skel1, skel2) | _ => (no_skel, no_skel)) in (case rew1 Ts hs skel1 prf1 of SOME prf1' => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1' %% prf2') | NONE => SOME (prf1' %% prf2)) | NONE => (case rew1 Ts hs skel2 prf2 of SOME prf2' => SOME (prf1 %% prf2') | NONE => NONE)) end) | rew2 Ts hs skel (Abst (s, T, prf)) = (case rew1 (the_default dummyT T :: Ts) hs (case skel of Abst (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (Abst (s, T, prf')) | NONE => NONE) | rew2 Ts hs skel (AbsP (s, t, prf)) = (case rew1 Ts (t :: hs) (case skel of AbsP (_, _, skel') => skel' | _ => no_skel) prf of SOME prf' => SOME (AbsP (s, t, prf')) | NONE => NONE) | rew2 _ _ _ _ = NONE; in the_default prf (rew1 [] [] no_skel prf) end; fun rewrite_proof thy = rewrite_prf (fn (tyenv, f) => Sign.typ_match thy (f ()) tyenv handle Type.TYPE_MATCH => raise PMatch); fun rewrite_proof_notypes rews = rewrite_prf fst rews; (* theory data *) structure Data = Theory_Data ( type T = ((stamp * (proof * proof)) list * (stamp * (typ list -> term option list -> proof -> (proof * proof) option)) list) * (theory -> proof -> proof) option; val empty = (([], []), NONE); fun merge (((rules1, procs1), preproc1), ((rules2, procs2), preproc2)) : T = ((AList.merge (op =) (K true) (rules1, rules2), AList.merge (op =) (K true) (procs1, procs2)), merge_options (preproc1, preproc2)); ); fun get_rew_data thy = let val (rules, procs) = #1 (Data.get thy) in (map #2 rules, map #2 procs) end; fun rew_proof thy = rewrite_prf fst (get_rew_data thy); fun add_prf_rrule r = (Data.map o apfst o apfst) (cons (stamp (), r)); fun add_prf_rproc p = (Data.map o apfst o apsnd) (cons (stamp (), p)); fun set_preproc f = (Data.map o apsnd) (K (SOME f)); fun apply_preproc thy = (case #2 (Data.get thy) of NONE => I | SOME f => f thy); (** reconstruction of partial proof terms **) fun forall_intr_variables_term prop = fold_rev Logic.all (variables_of prop) prop; fun forall_intr_variables prop prf = fold_rev forall_intr_proof' (variables_of prop) prf; local fun app_types shift prop Ts prf = let val inst = type_variables_of prop ~~ Ts; fun subst_same A = (case AList.lookup (op =) inst A of SOME T => T | NONE => raise Same.SAME); val subst_type_same = Term_Subst.map_atypsT_same (fn TVar ((a, i), S) => subst_same (TVar ((a, i - shift), S)) | A => subst_same A); in Same.commit (map_proof_types_same subst_type_same) prf end; fun guess_name (PThm ({name, ...}, _)) = name | guess_name (prf %% Hyp _) = guess_name prf | guess_name (prf %% PClass _) = guess_name prf | guess_name (prf % NONE) = guess_name prf | guess_name (prf % SOME (Var _)) = guess_name prf | guess_name _ = ""; (* generate constraints for proof term *) fun mk_var env Ts T = let val (env', v) = Envir.genvar "a" (env, rev Ts ---> T) in (list_comb (v, map Bound (length Ts - 1 downto 0)), env') end; fun mk_tvar S (Envir.Envir {maxidx, tenv, tyenv}) = (TVar (("'t", maxidx + 1), S), Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}); val mk_abs = fold (fn T => fn u => Abs ("", T, u)); fun unifyT thy env T U = let val Envir.Envir {maxidx, tenv, tyenv} = env; val (tyenv', maxidx') = Sign.typ_unify thy (T, U) (tyenv, maxidx); in Envir.Envir {maxidx = maxidx', tenv = tenv, tyenv = tyenv'} end; fun chaseT env (T as TVar v) = (case Type.lookup (Envir.type_env env) v of NONE => T | SOME T' => chaseT env T') | chaseT _ T = T; fun infer_type thy (env as Envir.Envir {maxidx, tenv, tyenv}) _ vTs (t as Const (s, T)) = if T = dummyT then (case Sign.const_type thy s of NONE => error ("reconstruct_proof: No such constant: " ^ quote s) | SOME T => let val T' = Type.strip_sorts (Logic.incr_tvar (maxidx + 1) T) in (Const (s, T'), T', vTs, Envir.Envir {maxidx = maxidx + 1, tenv = tenv, tyenv = tyenv}) end) else (t, T, vTs, env) | infer_type _ env _ vTs (t as Free (s, T)) = if T = dummyT then (case Symtab.lookup vTs s of NONE => let val (T, env') = mk_tvar [] env in (Free (s, T), T, Symtab.update_new (s, T) vTs, env') end | SOME T => (Free (s, T), T, vTs, env)) else (t, T, vTs, env) | infer_type _ _ _ _ (Var _) = error "reconstruct_proof: internal error" | infer_type thy env Ts vTs (Abs (s, T, t)) = let val (T', env') = if T = dummyT then mk_tvar [] env else (T, env); val (t', U, vTs', env'') = infer_type thy env' (T' :: Ts) vTs t in (Abs (s, T', t'), T' --> U, vTs', env'') end | infer_type thy env Ts vTs (t $ u) = let val (t', T, vTs1, env1) = infer_type thy env Ts vTs t; val (u', U, vTs2, env2) = infer_type thy env1 Ts vTs1 u; in (case chaseT env2 T of Type ("fun", [U', V]) => (t' $ u', V, vTs2, unifyT thy env2 U U') | _ => let val (V, env3) = mk_tvar [] env2 in (t' $ u', V, vTs2, unifyT thy env3 T (U --> V)) end) end | infer_type _ env Ts vTs (t as Bound i) = ((t, nth Ts i, vTs, env) handle General.Subscript => error ("infer_type: bad variable index " ^ string_of_int i)); fun cantunify thy (t, u) = error ("Non-unifiable terms:\n" ^ Syntax.string_of_term_global thy t ^ "\n\n" ^ Syntax.string_of_term_global thy u); fun decompose thy Ts (p as (t, u)) env = let fun rigrig (a, T) (b, U) uT ts us = if a <> b then cantunify thy p else apfst flat (fold_map (decompose thy Ts) (ts ~~ us) (uT env T U)) in case apply2 (strip_comb o Envir.head_norm env) p of ((Const c, ts), (Const d, us)) => rigrig c d (unifyT thy) ts us | ((Free c, ts), (Free d, us)) => rigrig c d (unifyT thy) ts us | ((Bound i, ts), (Bound j, us)) => rigrig (i, dummyT) (j, dummyT) (K o K) ts us | ((Abs (_, T, t), []), (Abs (_, U, u), [])) => decompose thy (T::Ts) (t, u) (unifyT thy env T U) | ((Abs (_, T, t), []), _) => decompose thy (T::Ts) (t, incr_boundvars 1 u $ Bound 0) env | (_, (Abs (_, T, u), [])) => decompose thy (T::Ts) (incr_boundvars 1 t $ Bound 0, u) env | _ => ([(mk_abs Ts t, mk_abs Ts u)], env) end; fun make_constraints_cprf thy env cprf = let fun add_cnstrt Ts prop prf cs env vTs (t, u) = let val t' = mk_abs Ts t; val u' = mk_abs Ts u in (prop, prf, cs, Pattern.unify (Context.Theory thy) (t', u') env, vTs) handle Pattern.Pattern => let val (cs', env') = decompose thy [] (t', u') env in (prop, prf, cs @ cs', env', vTs) end | Pattern.Unif => cantunify thy (Envir.norm_term env t', Envir.norm_term env u') end; fun mk_cnstrts_atom env vTs prop opTs prf = let val prop_types = type_variables_of prop; val (Ts, env') = (case opTs of NONE => fold_map (mk_tvar o Type.sort_of_atyp) prop_types env | SOME Ts => (Ts, env)); val prop' = subst_atomic_types (prop_types ~~ Ts) (forall_intr_variables_term prop) handle ListPair.UnequalLengths => error ("Wrong number of type arguments for " ^ quote (guess_name prf)) in (prop', change_types (SOME Ts) prf, [], env', vTs) end; fun head_norm (prop, prf, cnstrts, env, vTs) = (Envir.head_norm env prop, prf, cnstrts, env, vTs); fun mk_cnstrts env _ Hs vTs (PBound i) = ((nth Hs i, PBound i, [], env, vTs) handle General.Subscript => error ("mk_cnstrts: bad variable index " ^ string_of_int i)) | mk_cnstrts env Ts Hs vTs (Abst (s, opT, cprf)) = let val (T, env') = (case opT of NONE => mk_tvar [] env | SOME T => (T, env)); val (t, prf, cnstrts, env'', vTs') = mk_cnstrts env' (T::Ts) (map (incr_boundvars 1) Hs) vTs cprf; in (Const ("Pure.all", (T --> propT) --> propT) $ Abs (s, T, t), Abst (s, SOME T, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (AbsP (s, SOME t, cprf)) = let val (t', _, vTs', env') = infer_type thy env Ts vTs t; val (u, prf, cnstrts, env'', vTs'') = mk_cnstrts env' Ts (t'::Hs) vTs' cprf; in (Logic.mk_implies (t', u), AbsP (s, SOME t', prf), cnstrts, env'', vTs'') end | mk_cnstrts env Ts Hs vTs (AbsP (s, NONE, cprf)) = let val (t, env') = mk_var env Ts propT; val (u, prf, cnstrts, env'', vTs') = mk_cnstrts env' Ts (t::Hs) vTs cprf; in (Logic.mk_implies (t, u), AbsP (s, SOME t, prf), cnstrts, env'', vTs') end | mk_cnstrts env Ts Hs vTs (cprf1 %% cprf2) = let val (u, prf2, cnstrts, env', vTs') = mk_cnstrts env Ts Hs vTs cprf2 in (case head_norm (mk_cnstrts env' Ts Hs vTs' cprf1) of (Const ("Pure.imp", _) $ u' $ t', prf1, cnstrts', env'', vTs'') => add_cnstrt Ts t' (prf1 %% prf2) (cnstrts' @ cnstrts) env'' vTs'' (u, u') | (t, prf1, cnstrts', env'', vTs'') => let val (v, env''') = mk_var env'' Ts propT in add_cnstrt Ts v (prf1 %% prf2) (cnstrts' @ cnstrts) env''' vTs'' (t, Logic.mk_implies (u, v)) end) end | mk_cnstrts env Ts Hs vTs (cprf % SOME t) = let val (t', U, vTs1, env1) = infer_type thy env Ts vTs t in (case head_norm (mk_cnstrts env1 Ts Hs vTs1 cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env2, vTs2) => let val env3 = unifyT thy env2 T U in (betapply (f, t'), prf % SOME t', cnstrts, env3, vTs2) end | (u, prf, cnstrts, env2, vTs2) => let val (v, env3) = mk_var env2 Ts (U --> propT); in add_cnstrt Ts (v $ t') (prf % SOME t') cnstrts env3 vTs2 (u, Const ("Pure.all", (U --> propT) --> propT) $ v) end) end | mk_cnstrts env Ts Hs vTs (cprf % NONE) = (case head_norm (mk_cnstrts env Ts Hs vTs cprf) of (Const ("Pure.all", Type ("fun", [Type ("fun", [T, _]), _])) $ f, prf, cnstrts, env', vTs') => let val (t, env'') = mk_var env' Ts T in (betapply (f, t), prf % SOME t, cnstrts, env'', vTs') end | (u, prf, cnstrts, env', vTs') => let val (T, env1) = mk_tvar [] env'; val (v, env2) = mk_var env1 Ts (T --> propT); val (t, env3) = mk_var env2 Ts T in add_cnstrt Ts (v $ t) (prf % SOME t) cnstrts env3 vTs' (u, Const ("Pure.all", (T --> propT) --> propT) $ v) end) | mk_cnstrts env _ _ vTs (prf as PThm ({prop, types = opTs, ...}, _)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PAxm (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (prf as PClass (T, c)) = mk_cnstrts_atom env vTs (Logic.mk_of_class (T, c)) NONE prf | mk_cnstrts env _ _ vTs (prf as Oracle (_, prop, opTs)) = mk_cnstrts_atom env vTs prop opTs prf | mk_cnstrts env _ _ vTs (Hyp t) = (t, Hyp t, [], env, vTs) | mk_cnstrts _ _ _ _ MinProof = raise MIN_PROOF () in mk_cnstrts env [] [] Symtab.empty cprf end; (* update list of free variables of constraints *) fun upd_constrs env cs = let val tenv = Envir.term_env env; val tyenv = Envir.type_env env; val dom = [] |> Vartab.fold (cons o #1) tenv |> Vartab.fold (cons o #1) tyenv; val vran = [] |> Vartab.fold (Term.add_var_names o #2 o #2) tenv |> Vartab.fold (Term.add_tvar_namesT o #2 o #2) tyenv; fun check_cs [] = [] | check_cs ((u, p, vs) :: ps) = let val vs' = subtract (op =) dom vs in if vs = vs' then (u, p, vs) :: check_cs ps else (true, p, fold (insert op =) vs' vran) :: check_cs ps end; in check_cs cs end; (* solution of constraints *) fun solve _ [] bigenv = bigenv | solve thy cs bigenv = let fun search _ [] = error ("Unsolvable constraints:\n" ^ Pretty.string_of (Pretty.chunks (map (fn (_, p, _) => Pretty.item (Syntax.pretty_flexpair (Syntax.init_pretty_global thy) (apply2 (Envir.norm_term bigenv) p))) cs))) | search env ((u, p as (t1, t2), vs)::ps) = if u then let val tn1 = Envir.norm_term bigenv t1; val tn2 = Envir.norm_term bigenv t2 in if Pattern.pattern tn1 andalso Pattern.pattern tn2 then (Pattern.unify (Context.Theory thy) (tn1, tn2) env, ps) handle Pattern.Unif => cantunify thy (tn1, tn2) else let val (cs', env') = decompose thy [] (tn1, tn2) env in if cs' = [(tn1, tn2)] then apsnd (cons (false, (tn1, tn2), vs)) (search env ps) else search env' (map (fn q => (true, q, vs)) cs' @ ps) end end else apsnd (cons (false, p, vs)) (search env ps); val Envir.Envir {maxidx, ...} = bigenv; val (env, cs') = search (Envir.empty maxidx) cs; in solve thy (upd_constrs env cs') (Envir.merge (bigenv, env)) end; in (* reconstruction of proofs *) fun reconstruct_proof thy prop cprf = let val (cprf' % SOME prop', thawf) = freeze_thaw_prf (cprf % SOME prop); val (t, prf, cs, env, _) = make_constraints_cprf thy (Envir.empty (maxidx_proof cprf ~1)) cprf'; val cs' = map (apply2 (Envir.norm_term env)) ((t, prop') :: cs) |> map (fn p => (true, p, Term.add_var_names (#1 p) (Term.add_var_names (#2 p) []))); val env' = solve thy cs' env; in thawf (norm_proof env' prf) end handle MIN_PROOF () => MinProof; fun prop_of_atom prop Ts = subst_atomic_types (type_variables_of prop ~~ Ts) (forall_intr_variables_term prop); val head_norm = Envir.head_norm Envir.init; fun prop_of0 Hs (PBound i) = nth Hs i | prop_of0 Hs (Abst (s, SOME T, prf)) = Logic.all_const T $ (Abs (s, T, prop_of0 Hs prf)) | prop_of0 Hs (AbsP (_, SOME t, prf)) = Logic.mk_implies (t, prop_of0 (t :: Hs) prf) | prop_of0 Hs (prf % SOME t) = (case head_norm (prop_of0 Hs prf) of Const ("Pure.all", _) $ f => f $ t | _ => error "prop_of: all expected") | prop_of0 Hs (prf1 %% _) = (case head_norm (prop_of0 Hs prf1) of Const ("Pure.imp", _) $ _ $ Q => Q | _ => error "prop_of: ==> expected") | prop_of0 _ (Hyp t) = t | prop_of0 _ (PThm ({prop, types = SOME Ts, ...}, _)) = prop_of_atom prop Ts | prop_of0 _ (PAxm (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ (PClass (T, c)) = Logic.mk_of_class (T, c) | prop_of0 _ (Oracle (_, prop, SOME Ts)) = prop_of_atom prop Ts | prop_of0 _ _ = error "prop_of: partial proof object"; val prop_of' = Envir.beta_eta_contract oo prop_of0; val prop_of = prop_of' []; (* expand and reconstruct subproofs *) fun expand_name_empty (header: thm_header) = if #name header = "" then SOME "" else NONE; fun expand_proof thy expand_name prf = let fun expand seen maxidx (AbsP (s, t, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', AbsP (s, t, prf')) end | expand seen maxidx (Abst (s, T, prf)) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', Abst (s, T, prf')) end | expand seen maxidx (prf1 %% prf2) = let val (seen', maxidx', prf1') = expand seen maxidx prf1; val (seen'', maxidx'', prf2') = expand seen' maxidx' prf2; in (seen'', maxidx'', prf1' %% prf2') end | expand seen maxidx (prf % t) = let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', prf' % t) end | expand seen maxidx (prf as PThm (header, thm_body)) = let val {serial, pos, theory_name, name, prop, types} = header in (case expand_name header of SOME name' => if name' = "" andalso is_some types then let val (seen', maxidx', prf') = (case Inttab.lookup seen serial of NONE => let val prf1 = thm_body_proof_open thm_body |> reconstruct_proof thy prop |> forall_intr_variables prop; val (seen1, maxidx1, prf2) = expand_init seen prf1 val seen2 = seen1 |> Inttab.update (serial, (maxidx1, prf2)); in (seen2, maxidx1, prf2) end | SOME (maxidx1, prf1) => (seen, maxidx1, prf1)); val prf'' = prf' |> incr_indexes (maxidx + 1) |> app_types (maxidx + 1) prop (the types); in (seen', maxidx' + maxidx + 1, prf'') end else if name' <> name then (seen, maxidx, PThm (thm_header serial pos theory_name name' prop types, thm_body)) else (seen, maxidx, prf) | NONE => (seen, maxidx, prf)) end | expand seen maxidx prf = (seen, maxidx, prf) and expand_init seen prf = expand seen (maxidx_proof prf ~1) prf; in #3 (expand_init Inttab.empty prf) end; end; (** promises **) fun fulfill_norm_proof thy ps body0 = let val _ = consolidate_bodies (map #2 ps @ [body0]); val PBody {oracles = oracles0, thms = thms0, proof = proof0} = body0; val oracles = unions_oracles (fold (fn (_, PBody {oracles, ...}) => not (null oracles) ? cons oracles) ps [oracles0]); val thms = unions_thms (fold (fn (_, PBody {thms, ...}) => not (null thms) ? cons thms) ps [thms0]); val proof = rew_proof thy proof0; in PBody {oracles = oracles, thms = thms, proof = proof} end; fun fulfill_proof_future thy promises (postproc: proof_body -> proof_body) body = let fun fulfill () = postproc (fulfill_norm_proof thy (map (apsnd Future.join) promises) (Future.join body)); in if null promises then Future.map postproc body else if Future.is_finished body andalso length promises = 1 then Future.map (fn _ => fulfill ()) (snd (hd promises)) else (singleton o Future.forks) {name = "Proofterm.fulfill_proof_future", group = NONE, deps = Future.task_of body :: map (Future.task_of o snd) promises, pri = 1, interrupts = true} fulfill end; (** theorems **) (* standardization of variables for export: only frees and named bounds *) local val declare_names_term = Term.declare_term_frees; val declare_names_term' = fn SOME t => declare_names_term t | NONE => I; val declare_names_proof = fold_proof_terms declare_names_term; fun variant names bs x = #1 (Name.variant x (fold Name.declare bs names)); fun variant_term bs (Abs (x, T, t)) = let val x' = variant (declare_names_term t Name.context) bs x; val t' = variant_term (x' :: bs) t; in Abs (x', T, t') end | variant_term bs (t $ u) = variant_term bs t $ variant_term bs u | variant_term _ t = t; fun variant_proof bs (Abst (x, T, prf)) = let val x' = variant (declare_names_proof prf Name.context) bs x; val prf' = variant_proof (x' :: bs) prf; in Abst (x', T, prf') end | variant_proof bs (AbsP (x, t, prf)) = let val x' = variant (declare_names_term' t (declare_names_proof prf Name.context)) bs x; val t' = Option.map (variant_term bs) t; val prf' = variant_proof (x' :: bs) prf; in AbsP (x', t', prf') end | variant_proof bs (prf % t) = variant_proof bs prf % Option.map (variant_term bs) t | variant_proof bs (prf1 %% prf2) = variant_proof bs prf1 %% variant_proof bs prf2 | variant_proof bs (Hyp t) = Hyp (variant_term bs t) | variant_proof _ prf = prf; val used_frees_type = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I); fun used_frees_term t = fold_types used_frees_type t #> Term.declare_term_frees t; val used_frees_proof = fold_proof_terms_types used_frees_term used_frees_type; val unvarifyT = Term.map_atyps (fn TVar ((a, _), S) => TFree (a, S) | T => T); val unvarify = Term.map_aterms (fn Var ((x, _), T) => Free (x, T) | t => t) #> map_types unvarifyT; val unvarify_proof = map_proof_terms unvarify unvarifyT; fun hidden_types prop proof = let val visible = (fold_types o fold_atyps) (insert (op =)) prop []; val add_hiddenT = fold_atyps (fn T => not (member (op =) visible T) ? insert (op =) T); in rev (fold_proof_terms_types (fold_types add_hiddenT) add_hiddenT proof []) end; fun standard_hidden_types term proof = let val hidden = hidden_types term proof; val idx = Term.maxidx_term term (maxidx_proof proof ~1) + 1; fun smash T = if member (op =) hidden T then (case Type.sort_of_atyp T of [] => dummyT | S => TVar (("'", idx), S)) else T; val smashT = map_atyps smash; in map_proof_terms (map_types smashT) smashT proof end; fun standard_hidden_terms term proof = let fun add_unless excluded x = ((is_Free x orelse is_Var x) andalso not (member (op =) excluded x)) ? insert (op =) x; val visible = fold_aterms (add_unless []) term []; val hidden = fold_proof_terms (fold_aterms (add_unless visible)) proof []; val dummy_term = Term.map_aterms (fn x => if member (op =) hidden x then Term.dummy_pattern (Term.fastype_of x) else x); in proof |> not (null hidden) ? map_proof_terms dummy_term I end; in fun standard_vars used (term, opt_proof) = let val proofs = opt_proof |> Option.map (standard_hidden_types term #> standard_hidden_terms term) |> the_list; val proof_terms = rev (fold (fold_proof_terms_types cons (cons o Logic.mk_type)) proofs []); val used_frees = used |> used_frees_term term |> fold used_frees_proof proofs; val inst = Term_Subst.zero_var_indexes_inst used_frees (term :: proof_terms); val term' = term |> Term_Subst.instantiate inst |> unvarify |> variant_term []; val proofs' = proofs |> map (instantiate inst #> unvarify_proof #> variant_proof []); in (term', try hd proofs') end; fun standard_vars_term used t = #1 (standard_vars used (t, NONE)); val add_standard_vars_term = fold_aterms (fn Free (x, T) => (fn env => (case AList.lookup (op =) env x of NONE => (x, T) :: env | SOME T' => if T = T' then env else raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, [T, T'], []))) | _ => I); val add_standard_vars = fold_proof_terms add_standard_vars_term; end; (* PThm nodes *) fun prune_body body = if Options.default_bool "prune_proofs" then (Future.map o map_proof_of) (K MinProof) body else body; fun export_enabled () = Options.default_bool "export_proofs"; fun export_standard_enabled () = Options.default_bool "export_standard_proofs"; fun export_proof_boxes_required thy = Context.theory_name thy = Context.PureN orelse (export_enabled () andalso not (export_standard_enabled ())); fun export_proof_boxes bodies = let fun export_thm (i, thm_node) boxes = if Inttab.defined boxes i then boxes else boxes |> Inttab.update (i, thm_node_export thm_node) |> fold export_thm (thm_node_thms thm_node); fun export_body (PBody {thms, ...}) = fold export_thm thms; val exports = Inttab.build (fold export_body bodies) |> Inttab.dest; in List.app (Lazy.force o #2) exports end; local fun unconstrainT_proof algebra classrel_proof arity_proof (ucontext: Logic.unconstrain_context) = let fun hyp_map hyp = (case AList.lookup (op =) (#constraints ucontext) hyp of SOME t => Hyp t | NONE => raise Fail "unconstrainT_proof: missing constraint"); val typ = Term_Subst.map_atypsT_same (Type.strip_sorts o #atyp_map ucontext); fun ofclass (ty, c) = let val ty' = Term.map_atyps (#atyp_map ucontext) ty; in the_single (of_sort_proof algebra classrel_proof arity_proof hyp_map (ty', [c])) end; in Same.commit (map_proof_same (Term_Subst.map_types_same typ) typ ofclass) #> fold_rev (implies_intr_proof o snd) (#constraints ucontext) end; fun export_proof thy i prop prf0 = let val prf = prf0 |> reconstruct_proof thy prop |> apply_preproc thy; val (prop', SOME prf') = (prop, SOME prf) |> standard_vars Name.context; val args = [] |> add_standard_vars_term prop' |> add_standard_vars prf' |> rev; val typargs = [] |> Term.add_tfrees prop' |> fold_proof_terms Term.add_tfrees prf' |> rev; val consts = Sign.consts_of thy; val xml = (typargs, (args, (prop', no_thm_names prf'))) |> let open XML.Encode Term_XML.Encode; val encode_vars = list (pair string typ); val encode_term = encode_standard_term consts; val encode_proof = encode_standard_proof consts; in pair (list (pair string sort)) (pair encode_vars (pair encode_term encode_proof)) end; in Export.export_params {theory = thy, binding = Path.binding0 (Path.make ["proofs", string_of_int i]), executable = false, compress = true, strict = false} xml end; fun prepare_thm_proof unconstrain thy classrel_proof arity_proof (name, pos) shyps hyps concl promises body = let val named = name <> ""; val prop = Logic.list_implies (hyps, concl); val args = prop_args prop; val (ucontext, prop1) = Logic.unconstrainT shyps prop; val PBody {oracles = oracles0, thms = thms0, proof = prf} = body; val body0 = Future.value (PBody {oracles = oracles0, thms = thms0, proof = if proofs_enabled () then fold_rev implies_intr_proof hyps prf else MinProof}); fun new_prf () = let val i = serial (); val unconstrainT = unconstrainT_proof (Sign.classes_of thy) classrel_proof arity_proof ucontext; val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); in (i, fulfill_proof_future thy promises postproc body0) end; val (i, body') = (*somewhat non-deterministic proof boxes!*) if export_enabled () then new_prf () else (case strip_combt (fst (strip_combP prf)) of (PThm ({serial = ser, name = a, prop = prop', types = NONE, ...}, thm_body'), args') => if (a = "" orelse a = name) andalso prop' = prop1 andalso args' = args then let val Thm_Body {body = body', ...} = thm_body'; val i = if a = "" andalso named then serial () else ser; in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); val open_proof = not named ? rew_proof thy; val export = if export_enabled () then Lazy.lazy (fn () => join_proof body' |> open_proof |> export_proof thy i prop1 handle exn => if Exn.is_interrupt exn then raise Fail ("Interrupt: potential resource problems while exporting proof " ^ string_of_int i) else Exn.reraise exn) else no_export; val thm_body = prune_body body'; val theory_name = Context.theory_long_name thy; val thm = (i, make_thm_node theory_name name prop1 thm_body export); val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; val head = PThm (header, Thm_Body {open_proof = open_proof, body = thm_body}); val proof = if unconstrain then proof_combt' (head, (map o Option.map o Term.map_types) (#map_atyps ucontext) args) else proof_combP (proof_combt' (head, args), map PClass (#outer_constraints ucontext) @ map Hyp hyps); in (thm, proof) end; in fun thm_proof thy = prepare_thm_proof false thy; fun unconstrain_thm_proof thy classrel_proof arity_proof shyps concl promises body = prepare_thm_proof true thy classrel_proof arity_proof ("", Position.none) shyps [] concl promises body; end; (* PThm identity *) fun get_identity shyps hyps prop prf = let val (_, prop) = Logic.unconstrainT shyps (Logic.list_implies (hyps, prop)) in (case fst (strip_combt (fst (strip_combP prf))) of PThm ({serial, theory_name, name, prop = prop', ...}, _) => if prop = prop' then SOME {serial = serial, theory_name = theory_name, name = name} else NONE | _ => NONE) end; fun get_approximative_name shyps hyps prop prf = Option.map #name (get_identity shyps hyps prop prf) |> the_default ""; (* thm_id *) type thm_id = {serial: serial, theory_name: string}; fun make_thm_id (serial, theory_name) : thm_id = {serial = serial, theory_name = theory_name}; fun thm_header_id ({serial, theory_name, ...}: thm_header) = make_thm_id (serial, theory_name); fun thm_id (serial, thm_node) : thm_id = make_thm_id (serial, thm_node_theory_name thm_node); fun get_id shyps hyps prop prf : thm_id option = (case get_identity shyps hyps prop prf of NONE => NONE | SOME {name = "", ...} => NONE | SOME {serial, theory_name, ...} => SOME (make_thm_id (serial, theory_name))); fun this_id NONE _ = false | this_id (SOME (thm_id: thm_id)) (thm_id': thm_id) = #serial thm_id = #serial thm_id'; (* proof boxes: intermediate PThm nodes *) fun proof_boxes {included, excluded} proofs = let fun boxes_of (Abst (_, _, prf)) = boxes_of prf | boxes_of (AbsP (_, _, prf)) = boxes_of prf | boxes_of (prf % _) = boxes_of prf | boxes_of (prf1 %% prf2) = boxes_of prf1 #> boxes_of prf2 | boxes_of (PThm (header as {serial = i, ...}, thm_body)) = (fn boxes => let val thm_id = thm_header_id header in if Inttab.defined boxes i orelse (excluded thm_id andalso not (included thm_id)) then boxes else let val prf' = thm_body_proof_open thm_body; val boxes' = Inttab.update (i, (header, prf')) boxes; in boxes_of prf' boxes' end end) | boxes_of MinProof = raise MIN_PROOF () | boxes_of _ = I; in Inttab.fold_rev (cons o #2) (Inttab.build (fold boxes_of proofs)) [] end; end; structure Basic_Proofterm = struct datatype proof = datatype Proofterm.proof datatype proof_body = datatype Proofterm.proof_body val op %> = Proofterm.%> end; open Basic_Proofterm; diff --git a/src/Pure/sorts.ML b/src/Pure/sorts.ML --- a/src/Pure/sorts.ML +++ b/src/Pure/sorts.ML @@ -1,492 +1,453 @@ (* Title: Pure/sorts.ML Author: Markus Wenzel and Stefan Berghofer, TU Muenchen The order-sorted algebra of type classes. Classes denote (possibly empty) collections of types that are partially ordered by class inclusion. They are represented symbolically by strings. Sorts are intersections of finitely many classes. They are represented by lists of classes. Normal forms of sorts are sorted lists of minimal classes (wrt. current class inclusion). *) signature SORTS = sig - val make: sort list -> sort Ord_List.T - val subset: sort Ord_List.T * sort Ord_List.T -> bool - val union: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T - val subtract: sort Ord_List.T -> sort Ord_List.T -> sort Ord_List.T - val remove_sort: sort -> sort Ord_List.T -> sort Ord_List.T - val insert_sort: sort -> sort Ord_List.T -> sort Ord_List.T - val insert_typ: typ -> sort Ord_List.T -> sort Ord_List.T - val insert_typs: typ list -> sort Ord_List.T -> sort Ord_List.T - val insert_term: term -> sort Ord_List.T -> sort Ord_List.T - val insert_terms: term list -> sort Ord_List.T -> sort Ord_List.T type algebra val classes_of: algebra -> serial Graph.T val arities_of: algebra -> (class * sort list) list Symtab.table val all_classes: algebra -> class list val super_classes: algebra -> class -> class list val class_less: algebra -> class * class -> bool val class_le: algebra -> class * class -> bool val sort_eq: algebra -> sort * sort -> bool val sort_le: algebra -> sort * sort -> bool val sorts_le: algebra -> sort list * sort list -> bool val inter_sort: algebra -> sort * sort -> sort val minimize_sort: algebra -> sort -> sort val complete_sort: algebra -> sort -> sort val add_class: Context.generic -> class * class list -> algebra -> algebra val add_classrel: Context.generic -> class * class -> algebra -> algebra val add_arities: Context.generic -> string * (class * sort list) list -> algebra -> algebra val empty_algebra: algebra val merge_algebra: Context.generic -> algebra * algebra -> algebra val dest_algebra: algebra list -> algebra -> {classrel: (class * class list) list, arities: (string * sort list * class) list} val subalgebra: Context.generic -> (class -> bool) -> (class * string -> sort list option) -> algebra -> (sort -> sort) * algebra type class_error val class_error: Context.generic -> class_error -> string exception CLASS_ERROR of class_error val has_instance: algebra -> string -> sort -> bool val mg_domain: algebra -> string -> sort -> sort list (*exception CLASS_ERROR*) val meet_sort: algebra -> typ * sort -> sort Vartab.table -> sort Vartab.table (*exception CLASS_ERROR*) val meet_sort_typ: algebra -> typ * sort -> typ -> typ (*exception CLASS_ERROR*) val of_sort: algebra -> typ * sort -> bool val of_sort_derivation: algebra -> {class_relation: typ -> bool -> 'a * class -> class -> 'a, type_constructor: string * typ list -> ('a * class) list list -> class -> 'a, type_variable: typ -> ('a * class) list} -> typ * sort -> 'a list (*exception CLASS_ERROR*) val classrel_derivation: algebra -> ('a * class -> class -> 'a) -> 'a * class -> class -> 'a (*exception CLASS_ERROR*) val witness_sorts: algebra -> string list -> (typ * sort) list -> sort list -> (typ * sort) list end; structure Sorts: SORTS = struct - -(** ordered lists of sorts **) - -val make = Ord_List.make Term_Ord.sort_ord; -val subset = Ord_List.subset Term_Ord.sort_ord; -val union = Ord_List.union Term_Ord.sort_ord; -val subtract = Ord_List.subtract Term_Ord.sort_ord; - -val remove_sort = Ord_List.remove Term_Ord.sort_ord; -val insert_sort = Ord_List.insert Term_Ord.sort_ord; - -fun insert_typ (TFree (_, S)) Ss = insert_sort S Ss - | insert_typ (TVar (_, S)) Ss = insert_sort S Ss - | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss -and insert_typs [] Ss = Ss - | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss); - -fun insert_term (Const (_, T)) Ss = insert_typ T Ss - | insert_term (Free (_, T)) Ss = insert_typ T Ss - | insert_term (Var (_, T)) Ss = insert_typ T Ss - | insert_term (Bound _) Ss = Ss - | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss) - | insert_term (t $ u) Ss = insert_term t (insert_term u Ss); - -fun insert_terms [] Ss = Ss - | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss); - - - (** order-sorted algebra **) (* classes: graph representing class declarations together with proper subclass relation, which needs to be transitive and acyclic. arities: table of association lists of all type arities; (t, ars) means that type constructor t has the arities ars; an element (c, Ss) of ars represents the arity t::(Ss)c. "Coregularity" of the arities structure requires that for any two declarations t::(Ss1)c1 and t::(Ss2)c2 such that c1 <= c2 holds Ss1 <= Ss2. *) datatype algebra = Algebra of {classes: serial Graph.T, arities: (class * sort list) list Symtab.table}; fun classes_of (Algebra {classes, ...}) = classes; fun arities_of (Algebra {arities, ...}) = arities; fun make_algebra (classes, arities) = Algebra {classes = classes, arities = arities}; fun map_classes f (Algebra {classes, arities}) = make_algebra (f classes, arities); fun map_arities f (Algebra {classes, arities}) = make_algebra (classes, f arities); (* classes *) fun all_classes (Algebra {classes, ...}) = Graph.all_preds classes (Graph.maximals classes); val super_classes = Graph.immediate_succs o classes_of; (* class relations *) val class_less : algebra -> class * class -> bool = Graph.is_edge o classes_of; fun class_le algebra (c1, c2) = c1 = c2 orelse class_less algebra (c1, c2); (* sort relations *) fun sort_le algebra (S1: sort, S2: sort) = S1 = S2 orelse forall (fn c2 => exists (fn c1 => class_le algebra (c1, c2)) S1) S2; fun sorts_le algebra (Ss1, Ss2) = ListPair.all (sort_le algebra) (Ss1, Ss2); fun sort_eq algebra (S1, S2) = sort_le algebra (S1, S2) andalso sort_le algebra (S2, S1); (* intersection *) fun inter_class algebra c S = let fun intr [] = [c] | intr (S' as c' :: c's) = if class_le algebra (c', c) then S' else if class_le algebra (c, c') then intr c's else c' :: intr c's in intr S end; fun inter_sort algebra (S1, S2) = sort_strings (fold (inter_class algebra) S1 S2); (* normal forms *) fun minimize_sort _ [] = [] | minimize_sort _ (S as [_]) = S | minimize_sort algebra S = filter (fn c => not (exists (fn c' => class_less algebra (c', c)) S)) S |> sort_distinct string_ord; fun complete_sort algebra = Graph.all_succs (classes_of algebra) o minimize_sort algebra; (** build algebras **) (* classes *) fun err_dup_class c = error ("Duplicate declaration of class: " ^ quote c); fun err_cyclic_classes context css = error (cat_lines (map (fn cs => "Cycle in class relation: " ^ Syntax.string_of_classrel (Syntax.init_pretty context) cs) css)); fun add_class context (c, cs) = map_classes (fn classes => let val classes' = classes |> Graph.new_node (c, serial ()) handle Graph.DUP dup => err_dup_class dup; val classes'' = classes' |> fold Graph.add_edge_trans_acyclic (map (pair c) cs) handle Graph.CYCLES css => err_cyclic_classes context css; in classes'' end); (* arities *) local fun for_classes _ NONE = "" | for_classes ctxt (SOME (c1, c2)) = " for classes " ^ Syntax.string_of_classrel ctxt [c1, c2]; fun err_conflict context t cc (c, Ss) (c', Ss') = let val ctxt = Syntax.init_pretty context in error ("Conflict of type arities" ^ for_classes ctxt cc ^ ":\n " ^ Syntax.string_of_arity ctxt (t, Ss, [c]) ^ " and\n " ^ Syntax.string_of_arity ctxt (t, Ss', [c'])) end; fun coregular context algebra t (c, Ss) ars = let fun conflict (c', Ss') = if class_le algebra (c, c') andalso not (sorts_le algebra (Ss, Ss')) then SOME ((c, c'), (c', Ss')) else if class_le algebra (c', c) andalso not (sorts_le algebra (Ss', Ss)) then SOME ((c', c), (c', Ss')) else NONE; in (case get_first conflict ars of SOME ((c1, c2), (c', Ss')) => err_conflict context t (SOME (c1, c2)) (c, Ss) (c', Ss') | NONE => (c, Ss) :: ars) end; fun complete algebra (c, Ss) = map (rpair Ss) (c :: super_classes algebra c); fun insert context algebra t (c, Ss) ars = (case AList.lookup (op =) ars c of NONE => coregular context algebra t (c, Ss) ars | SOME Ss' => if sorts_le algebra (Ss, Ss') then ars else if sorts_le algebra (Ss', Ss) then coregular context algebra t (c, Ss) (remove (op =) (c, Ss') ars) else err_conflict context t NONE (c, Ss) (c, Ss')); in fun insert_ars context algebra t = fold_rev (insert context algebra t); fun insert_complete_ars context algebra (t, ars) arities = let val ars' = Symtab.lookup_list arities t |> fold_rev (insert_ars context algebra t) (map (complete algebra) ars); in Symtab.update (t, ars') arities end; fun add_arities context arg algebra = algebra |> map_arities (insert_complete_ars context algebra arg); fun add_arities_table context algebra = Symtab.fold (fn (t, ars) => insert_complete_ars context algebra (t, ars)); end; (* classrel *) fun rebuild_arities context algebra = algebra |> map_arities (Symtab.build o add_arities_table context algebra); fun add_classrel context rel = rebuild_arities context o map_classes (fn classes => classes |> Graph.add_edge_trans_acyclic rel handle Graph.CYCLES css => err_cyclic_classes context css); (* empty and merge *) val empty_algebra = make_algebra (Graph.empty, Symtab.empty); fun merge_algebra context (Algebra {classes = classes1, arities = arities1}, Algebra {classes = classes2, arities = arities2}) = let val classes' = Graph.merge_trans_acyclic (op =) (classes1, classes2) handle Graph.DUP c => err_dup_class c | Graph.CYCLES css => err_cyclic_classes context css; val algebra0 = make_algebra (classes', Symtab.empty); val arities' = (case (pointer_eq (classes1, classes2), pointer_eq (arities1, arities2)) of (true, true) => arities1 | (true, false) => (*no completion*) (arities1, arities2) |> Symtab.join (fn t => fn (ars1, ars2) => if pointer_eq (ars1, ars2) then raise Symtab.SAME else insert_ars context algebra0 t ars2 ars1) | (false, true) => (*unary completion*) Symtab.build (add_arities_table context algebra0 arities1) | (false, false) => (*binary completion*) Symtab.build (add_arities_table context algebra0 arities1 #> add_arities_table context algebra0 arities2)); in make_algebra (classes', arities') end; (* destruct *) fun dest_algebra parents (Algebra {classes, arities}) = let fun new_classrel rel = not (exists (fn algebra => class_less algebra rel) parents); val classrel = build (classes |> Graph.fold (fn (c, (_, (_, ds))) => (case filter (fn d => new_classrel (c, d)) (Graph.Keys.dest ds) of [] => I | ds' => cons (c, sort_strings ds')))) |> sort_by #1; fun is_arity t ar algebra = member (op =) (Symtab.lookup_list (arities_of algebra) t) ar; fun add_arity t (c, Ss) = not (exists (is_arity t (c, Ss)) parents) ? cons (t, Ss, c); val arities = build (arities |> Symtab.fold (fn (t, ars) => fold_rev (add_arity t) (sort_by #1 ars))) |> sort_by #1; in {classrel = classrel, arities = arities} end; (* algebra projections *) (* FIXME potentially violates abstract type integrity *) fun subalgebra context P sargs (algebra as Algebra {classes, arities}) = let val restrict_sort = minimize_sort algebra o filter P o Graph.all_succs classes; fun restrict_arity t (c, Ss) = if P c then (case sargs (c, t) of SOME sorts => SOME (c, Ss |> map2 (curry (inter_sort algebra)) sorts |> map restrict_sort) | NONE => NONE) else NONE; val classes' = classes |> Graph.restrict P; val arities' = arities |> Symtab.map (map_filter o restrict_arity); in (restrict_sort, rebuild_arities context (make_algebra (classes', arities'))) end; (** sorts of types **) (* errors -- performance tuning via delayed message composition *) datatype class_error = No_Classrel of class * class | No_Arity of string * class | No_Subsort of sort * sort; fun class_error context = let val ctxt = Syntax.init_pretty context in fn No_Classrel (c1, c2) => "No class relation " ^ Syntax.string_of_classrel ctxt [c1, c2] | No_Arity (a, c) => "No type arity " ^ Syntax.string_of_arity ctxt (a, [], [c]) | No_Subsort (S1, S2) => "Cannot derive subsort relation " ^ Syntax.string_of_sort ctxt S1 ^ " < " ^ Syntax.string_of_sort ctxt S2 end; exception CLASS_ERROR of class_error; (* instances *) fun has_instance algebra a = forall (AList.defined (op =) (Symtab.lookup_list (arities_of algebra) a)); fun mg_domain algebra a S = let val ars = Symtab.lookup_list (arities_of algebra) a; fun dom c = (case AList.lookup (op =) ars c of NONE => raise CLASS_ERROR (No_Arity (a, c)) | SOME Ss => Ss); fun dom_inter c Ss = ListPair.map (inter_sort algebra) (dom c, Ss); in (case S of [] => raise Fail "Unknown domain of empty intersection" | c :: cs => fold dom_inter cs (dom c)) end; (* meet_sort *) fun meet_sort algebra = let fun inters S S' = inter_sort algebra (S, S'); fun meet _ [] = I | meet (TFree (_, S)) S' = if sort_le algebra (S, S') then I else raise CLASS_ERROR (No_Subsort (S, S')) | meet (TVar (v, S)) S' = if sort_le algebra (S, S') then I else Vartab.map_default (v, S) (inters S') | meet (Type (a, Ts)) S = fold2 meet Ts (mg_domain algebra a S); in uncurry meet end; fun meet_sort_typ algebra (T, S) = let val tab = Vartab.build (meet_sort algebra (T, S)); in Term.map_type_tvar (fn (v, _) => TVar (v, (the o Vartab.lookup tab) v)) end; (* of_sort *) fun of_sort algebra = let fun ofS (_, []) = true | ofS (TFree (_, S), S') = sort_le algebra (S, S') | ofS (TVar (_, S), S') = sort_le algebra (S, S') | ofS (Type (a, Ts), S) = let val Ss = mg_domain algebra a S in ListPair.all ofS (Ts, Ss) end handle CLASS_ERROR _ => false; in ofS end; (* animating derivations *) fun of_sort_derivation algebra {class_relation, type_constructor, type_variable} = let val arities = arities_of algebra; fun weaken T D1 S2 = let val S1 = map snd D1 in if S1 = S2 then map fst D1 else S2 |> map (fn c2 => (case D1 |> filter (fn (_, c1) => class_le algebra (c1, c2)) of [d1] => class_relation T true d1 c2 | (d1 :: _ :: _) => class_relation T false d1 c2 | [] => raise CLASS_ERROR (No_Subsort (S1, S2)))) end; fun derive (_, []) = [] | derive (Type (a, Us), S) = let val Ss = mg_domain algebra a S; val dom = map2 (fn U => fn S => derive (U, S) ~~ S) Us Ss; in S |> map (fn c => let val Ss' = the (AList.lookup (op =) (Symtab.lookup_list arities a) c); val dom' = map (fn ((U, d), S') => weaken U d S' ~~ S') ((Us ~~ dom) ~~ Ss'); in type_constructor (a, Us) dom' c end) end | derive (T, S) = weaken T (type_variable T) S; in derive end; fun classrel_derivation algebra class_relation = let fun path (x, c1 :: c2 :: cs) = path (class_relation (x, c1) c2, c2 :: cs) | path (x, _) = x; in fn (x, c1) => fn c2 => (case Graph.irreducible_paths (classes_of algebra) (c1, c2) of [] => raise CLASS_ERROR (No_Classrel (c1, c2)) | cs :: _ => path (x, cs)) end; (* witness_sorts *) fun witness_sorts algebra ground_types hyps sorts = let fun le S1 S2 = sort_le algebra (S1, S2); fun get S2 (T, S1) = if le S1 S2 then SOME (T, S2) else NONE; fun mg_dom t S = SOME (mg_domain algebra t S) handle CLASS_ERROR _ => NONE; fun witn_sort _ [] solved_failed = (SOME (propT, []), solved_failed) | witn_sort path S (solved, failed) = if exists (le S) failed then (NONE, (solved, failed)) else (case get_first (get S) solved of SOME w => (SOME w, (solved, failed)) | NONE => (case get_first (get S) hyps of SOME w => (SOME w, (w :: solved, failed)) | NONE => witn_types path ground_types S (solved, failed))) and witn_sorts path x = fold_map (witn_sort path) x and witn_types _ [] S (solved, failed) = (NONE, (solved, S :: failed)) | witn_types path (t :: ts) S solved_failed = (case mg_dom t S of SOME SS => (*do not descend into stronger args (achieving termination)*) if exists (fn D => le D S orelse exists (le D) path) SS then witn_types path ts S solved_failed else let val (ws, (solved', failed')) = witn_sorts (S :: path) SS solved_failed in if forall is_some ws then let val w = (Type (t, map (#1 o the) ws), S) in (SOME w, (w :: solved', failed')) end else witn_types path ts S (solved', failed') end | NONE => witn_types path ts S solved_failed); in map_filter I (#1 (witn_sorts [] sorts ([], []))) end; end; diff --git a/src/Pure/term_ord.ML b/src/Pure/term_ord.ML --- a/src/Pure/term_ord.ML +++ b/src/Pure/term_ord.ML @@ -1,232 +1,261 @@ (* Title: Pure/term_ord.ML Author: Tobias Nipkow, TU Muenchen Author: Makarius Term orderings. *) signature TERM_ORD = sig val fast_indexname_ord: indexname ord val sort_ord: sort ord val typ_ord: typ ord val fast_term_ord: term ord val syntax_term_ord: term ord val indexname_ord: indexname ord val tvar_ord: (indexname * sort) ord val var_ord: (indexname * typ) ord val term_ord: term ord val hd_ord: term ord val term_lpo: (term -> int) -> term ord end; structure Term_Ord: TERM_ORD = struct (* fast syntactic ordering -- tuned for inequalities *) val fast_indexname_ord = pointer_eq_ord (int_ord o apply2 snd ||| fast_string_ord o apply2 fst); val sort_ord = pointer_eq_ord (dict_ord fast_string_ord); local fun cons_nr (TVar _) = 0 | cons_nr (TFree _) = 1 | cons_nr (Type _) = 2; in fun typ_ord TU = if pointer_eq TU then EQUAL else (case TU of (Type (a, Ts), Type (b, Us)) => (case fast_string_ord (a, b) of EQUAL => dict_ord typ_ord (Ts, Us) | ord => ord) | (TFree (a, S), TFree (b, S')) => (case fast_string_ord (a, b) of EQUAL => sort_ord (S, S') | ord => ord) | (TVar (xi, S), TVar (yj, S')) => (case fast_indexname_ord (xi, yj) of EQUAL => sort_ord (S, S') | ord => ord) | (T, U) => int_ord (cons_nr T, cons_nr U)); end; local fun cons_nr (Const _) = 0 | cons_nr (Free _) = 1 | cons_nr (Var _) = 2 | cons_nr (Bound _) = 3 | cons_nr (Abs _) = 4 | cons_nr (_ $ _) = 5; fun struct_ord (Abs (_, _, t), Abs (_, _, u)) = struct_ord (t, u) | struct_ord (t1 $ t2, u1 $ u2) = (case struct_ord (t1, u1) of EQUAL => struct_ord (t2, u2) | ord => ord) | struct_ord (t, u) = int_ord (cons_nr t, cons_nr u); fun atoms_ord (Abs (_, _, t), Abs (_, _, u)) = atoms_ord (t, u) | atoms_ord (t1 $ t2, u1 $ u2) = (case atoms_ord (t1, u1) of EQUAL => atoms_ord (t2, u2) | ord => ord) | atoms_ord (Const (a, _), Const (b, _)) = fast_string_ord (a, b) | atoms_ord (Free (x, _), Free (y, _)) = fast_string_ord (x, y) | atoms_ord (Var (xi, _), Var (yj, _)) = fast_indexname_ord (xi, yj) | atoms_ord (Bound i, Bound j) = int_ord (i, j) | atoms_ord _ = EQUAL; fun types_ord (Abs (_, T, t), Abs (_, U, u)) = (case typ_ord (T, U) of EQUAL => types_ord (t, u) | ord => ord) | types_ord (t1 $ t2, u1 $ u2) = (case types_ord (t1, u1) of EQUAL => types_ord (t2, u2) | ord => ord) | types_ord (Const (_, T), Const (_, U)) = typ_ord (T, U) | types_ord (Free (_, T), Free (_, U)) = typ_ord (T, U) | types_ord (Var (_, T), Var (_, U)) = typ_ord (T, U) | types_ord _ = EQUAL; fun comments_ord (Abs (x, _, t), Abs (y, _, u)) = (case fast_string_ord (x, y) of EQUAL => comments_ord (t, u) | ord => ord) | comments_ord (t1 $ t2, u1 $ u2) = (case comments_ord (t1, u1) of EQUAL => comments_ord (t2, u2) | ord => ord) | comments_ord _ = EQUAL; in val fast_term_ord = pointer_eq_ord (struct_ord ||| atoms_ord ||| types_ord); val syntax_term_ord = fast_term_ord ||| comments_ord; end; (* term_ord *) (*a linear well-founded AC-compatible ordering for terms: s < t <=> 1. size(s) < size(t) or 2. size(s) = size(t) and s=f(...) and t=g(...) and f (case term_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) | (t, u) => (case int_ord (size_of_term t, size_of_term u) of EQUAL => (case prod_ord hd_ord int_ord (hd_depth (t, 0), hd_depth (u, 0)) of EQUAL => args_ord (t, u) | ord => ord) | ord => ord)) and hd_ord (f, g) = prod_ord (prod_ord indexname_ord typ_ord) int_ord (dest_hd f, dest_hd g) and args_ord (f $ t, g $ u) = (case args_ord (f, g) of EQUAL => term_ord (t, u) | ord => ord) | args_ord _ = EQUAL; end; (* Lexicographic path order on terms *) (* See Baader & Nipkow, Term rewriting, CUP 1998. Without variables. Const, Var, Bound, Free and Abs are treated all as constants. f_ord maps terms to integers and serves two purposes: - Predicate on constant symbols. Those that are not recognised by f_ord must be mapped to ~1. - Order on the recognised symbols. These must be mapped to distinct integers >= 0. The argument of f_ord is never an application. *) local fun unrecognized (Const (a, T)) = ((1, ((a, 0), T)), 0) | unrecognized (Free (a, T)) = ((1, ((a, 0), T)), 0) | unrecognized (Var v) = ((1, v), 1) | unrecognized (Bound i) = ((1, (("", i), dummyT)), 2) | unrecognized (Abs (_, T, _)) = ((1, (("", 0), T)), 3); fun dest_hd f_ord t = let val ord = f_ord t in if ord = ~1 then unrecognized t else ((0, (("", ord), fastype_of t)), 0) end; fun term_lpo f_ord (s, t) = let val (f, ss) = strip_comb s and (g, ts) = strip_comb t in if forall (fn si => term_lpo f_ord (si, t) = LESS) ss then case hd_ord f_ord (f, g) of GREATER => if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts then GREATER else LESS | EQUAL => if forall (fn ti => term_lpo f_ord (s, ti) = GREATER) ts then list_ord (term_lpo f_ord) (ss, ts) else LESS | LESS => LESS else GREATER end and hd_ord f_ord (f, g) = case (f, g) of (Abs (_, T, t), Abs (_, U, u)) => (case term_lpo f_ord (t, u) of EQUAL => typ_ord (T, U) | ord => ord) | (_, _) => prod_ord (prod_ord int_ord (prod_ord indexname_ord typ_ord)) int_ord (dest_hd f_ord f, dest_hd f_ord g); in val term_lpo = term_lpo end; end; (* scalable collections *) structure Vartab = Table(type key = indexname val ord = Term_Ord.fast_indexname_ord); structure Sorttab = Table(type key = sort val ord = Term_Ord.sort_ord); structure Typtab = Table(type key = typ val ord = Term_Ord.typ_ord); structure Termtab: sig include TABLE val term_cache: (term -> 'a) -> term -> 'a end = struct structure Table = Table(type key = term val ord = Term_Ord.fast_term_ord); open Table; fun term_cache f = Cache.create empty lookup update f; end; structure Termset = Set(type key = term val ord = Term_Ord.fast_term_ord); -structure Sortset = Set(type key = sort val ord = Term_Ord.sort_ord); +structure Sortset: +sig + include SET + val insert_typ: typ -> T -> T + val insert_typs: typ list -> T -> T + val insert_term: term -> T -> T + val insert_terms: term list -> T -> T +end = +struct + +structure Set = Set(type key = sort val ord = Term_Ord.sort_ord); +open Set; + +fun insert_typ (TFree (_, S)) Ss = insert S Ss + | insert_typ (TVar (_, S)) Ss = insert S Ss + | insert_typ (Type (_, Ts)) Ss = insert_typs Ts Ss +and insert_typs [] Ss = Ss + | insert_typs (T :: Ts) Ss = insert_typs Ts (insert_typ T Ss); + +fun insert_term (Const (_, T)) Ss = insert_typ T Ss + | insert_term (Free (_, T)) Ss = insert_typ T Ss + | insert_term (Var (_, T)) Ss = insert_typ T Ss + | insert_term (Bound _) Ss = Ss + | insert_term (Abs (_, T, t)) Ss = insert_term t (insert_typ T Ss) + | insert_term (t $ u) Ss = insert_term t (insert_term u Ss); + +fun insert_terms [] Ss = Ss + | insert_terms (t :: ts) Ss = insert_terms ts (insert_term t Ss); + +end; structure Var_Graph = Graph(type key = indexname val ord = Term_Ord.fast_indexname_ord); structure Sort_Graph = Graph(type key = sort val ord = Term_Ord.sort_ord); structure Typ_Graph = Graph(type key = typ val ord = Term_Ord.typ_ord); structure Term_Graph = Graph(type key = term val ord = Term_Ord.fast_term_ord); diff --git a/src/Pure/thm.ML b/src/Pure/thm.ML --- a/src/Pure/thm.ML +++ b/src/Pure/thm.ML @@ -1,2407 +1,2405 @@ (* Title: Pure/thm.ML Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Makarius The very core of Isabelle's Meta Logic: certified types and terms, derivations, theorems, inference rules (including lifting and resolution), oracles. *) infix 0 RS RSN; signature BASIC_THM = sig type ctyp type cterm exception CTERM of string * cterm list type thm type conv = cterm -> thm exception THM of string * int * thm list val RSN: thm * (int * thm) -> thm val RS: thm * thm -> thm end; signature THM = sig include BASIC_THM (*certified types*) val typ_of: ctyp -> typ val global_ctyp_of: theory -> typ -> ctyp val ctyp_of: Proof.context -> typ -> ctyp val dest_ctyp: ctyp -> ctyp list val dest_ctypN: int -> ctyp -> ctyp val dest_ctyp0: ctyp -> ctyp val dest_ctyp1: ctyp -> ctyp val make_ctyp: ctyp -> ctyp list -> ctyp (*certified terms*) val term_of: cterm -> term val typ_of_cterm: cterm -> typ val ctyp_of_cterm: cterm -> ctyp val maxidx_of_cterm: cterm -> int val global_cterm_of: theory -> term -> cterm val cterm_of: Proof.context -> term -> cterm val renamed_term: term -> cterm -> cterm val fast_term_ord: cterm ord val term_ord: cterm ord val dest_comb: cterm -> cterm * cterm val dest_fun: cterm -> cterm val dest_arg: cterm -> cterm val dest_fun2: cterm -> cterm val dest_arg1: cterm -> cterm val dest_abs_fresh: string -> cterm -> cterm * cterm val dest_abs_global: 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 -> ctyp TVars.table * cterm Vars.table val first_order_match: cterm * cterm -> ctyp TVars.table * cterm Vars.table (*theorems*) val fold_terms: {hyps: bool} -> (term -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_ctyps: {hyps: bool} -> (typ -> bool) -> (ctyp -> 'a -> 'a) -> thm -> 'a -> 'a val fold_atomic_cterms: {hyps: bool} -> (term -> bool) -> (cterm -> 'a -> 'a) -> thm -> 'a -> 'a val terms_of_tpairs: (term * term) list -> term list val full_prop_of: thm -> term val theory_id: thm -> Context.theory_id val theory_name: thm -> string val maxidx_of: thm -> int val maxidx_thm: thm -> int -> int - val shyps_of: thm -> sort Ord_List.T + val shyps_of: thm -> Sortset.T val hyps_of: thm -> term list val prop_of: thm -> term val tpairs_of: thm -> (term * term) list val concl_of: thm -> term val prems_of: thm -> term list val nprems_of: thm -> int val no_prems: thm -> bool val major_prem_of: thm -> term val cprop_of: thm -> cterm val cprem_of: thm -> int -> cterm val cconcl_of: thm -> cterm val cprems_of: thm -> cterm list val chyps_of: thm -> cterm list val thm_ord: thm ord exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option val theory_of_cterm: cterm -> theory val theory_of_thm: thm -> theory val trim_context_ctyp: ctyp -> ctyp val trim_context_cterm: cterm -> cterm val transfer_ctyp: theory -> ctyp -> ctyp val transfer_cterm: theory -> cterm -> cterm val transfer: theory -> thm -> thm val transfer': Proof.context -> thm -> thm val transfer'': Context.generic -> thm -> thm val join_transfer: theory -> thm -> thm val join_transfer_context: Proof.context * thm -> Proof.context * thm val renamed_prop: term -> thm -> thm val weaken: cterm -> thm -> thm - val weaken_sorts: sort list -> cterm -> cterm + val weaken_sorts: Sortset.T -> cterm -> cterm val proof_bodies_of: thm list -> proof_body list val proof_body_of: thm -> proof_body val proof_of: thm -> proof val reconstruct_proof_of: thm -> Proofterm.proof val consolidate: thm list -> unit val expose_proofs: theory -> thm list -> unit val expose_proof: theory -> thm -> unit val future: thm future -> cterm -> thm val thm_deps: thm -> Proofterm.thm Ord_List.T - val extra_shyps: thm -> sort list + val extra_shyps: thm -> Sortset.T val strip_shyps: thm -> thm val derivation_closed: thm -> bool val derivation_name: thm -> string val derivation_id: thm -> Proofterm.thm_id option val raw_derivation_name: thm -> string val expand_name: thm -> Proofterm.thm_header -> string option val name_derivation: string * Position.T -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm val axiom: theory -> string -> thm val all_axioms_of: theory -> (string * thm) list val get_tags: thm -> Properties.T val map_tags: (Properties.T -> Properties.T) -> thm -> thm val norm_proof: thm -> thm val adjust_maxidx_thm: int -> thm -> thm (*type classes*) val the_classrel: theory -> class * class -> thm val the_arity: theory -> string * sort list * class -> thm val classrel_proof: theory -> class * class -> proof val arity_proof: theory -> string * sort list * class -> proof (*oracles*) val add_oracle: binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory val oracle_space: theory -> Name_Space.T val pretty_oracle: Proof.context -> string -> Pretty.T val extern_oracles: bool -> Proof.context -> (Markup.T * xstring) list val check_oracle: Proof.context -> xstring * Position.T -> string (*inference rules*) val assume: cterm -> thm val implies_intr: cterm -> thm -> thm val implies_elim: thm -> thm -> thm val forall_intr: cterm -> thm -> thm val forall_elim: cterm -> thm -> thm val reflexive: cterm -> thm val symmetric: thm -> thm val transitive: thm -> thm -> thm val beta_conversion: bool -> conv val eta_conversion: conv val eta_long_conversion: conv val abstract_rule: string -> cterm -> thm -> thm val combination: thm -> thm -> thm val equal_intr: thm -> thm -> thm val equal_elim: thm -> thm -> thm val solve_constraints: thm -> thm val flexflex_rule: Proof.context option -> thm -> thm Seq.seq val generalize: Names.set * Names.set -> int -> thm -> thm val generalize_cterm: Names.set * Names.set -> int -> cterm -> cterm val generalize_ctyp: Names.set -> int -> ctyp -> ctyp val instantiate: ctyp TVars.table * cterm Vars.table -> thm -> thm val instantiate_beta: ctyp TVars.table * cterm Vars.table -> thm -> thm val instantiate_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm val instantiate_beta_cterm: ctyp TVars.table * cterm Vars.table -> cterm -> cterm val trivial: cterm -> thm val of_class: ctyp * class -> thm val unconstrainT: thm -> thm val varifyT_global': TFrees.set -> thm -> ((string * sort) * indexname) list * thm val varifyT_global: thm -> thm val legacy_freezeT: thm -> thm val plain_prop_of: thm -> term val dest_state: thm * int -> (term * term) list * term list * term * term val lift_rule: cterm -> thm -> thm val incr_indexes: int -> thm -> thm val assumption: Proof.context option -> int -> thm -> thm Seq.seq val eq_assumption: int -> thm -> thm val rotate_rule: int -> int -> thm -> thm val permute_prems: int -> int -> thm -> thm val bicompose: Proof.context option -> {flatten: bool, match: bool, incremented: bool} -> bool * thm * int -> int -> thm -> thm Seq.seq val biresolution: Proof.context option -> bool -> (bool * thm) list -> int -> thm -> thm Seq.seq val thynames_of_arity: theory -> string * class -> string list val add_classrel: thm -> theory -> theory val add_arity: thm -> theory -> theory end; structure Thm: THM = struct (*** Certified terms and types ***) (** certified types **) -datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: sort Ord_List.T}; +datatype ctyp = Ctyp of {cert: Context.certificate, T: typ, maxidx: int, sorts: Sortset.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 []; + val sorts = Sortset.build (Sortset.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 insert_sorts_ctyp (Ctyp {sorts, ...}) sorts0 = Sortset.merge (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 [], + sorts = Sortset.build (fold insert_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}; + Cterm of {cert: Context.certificate, t: term, T: typ, maxidx: int, sorts: Sortset.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 []; + val sorts = Sortset.build (Sortset.insert_term t); in Cterm {cert = Context.Certificate thy, t = t, T = T, maxidx = maxidx, sorts = sorts} end; val cterm_of = global_cterm_of o Proof_Context.theory_of; fun join_certificate0 (Cterm {cert = cert1, ...}, Cterm {cert = cert2, ...}) = Context.join_certificate (cert1, cert2); fun renamed_term t' (Cterm {cert, t, T, maxidx, sorts}) = if t aconv t' then Cterm {cert = cert, t = t', T = T, maxidx = maxidx, sorts = sorts} else raise TERM ("renamed_term: terms disagree", [t, t']); val fast_term_ord = Term_Ord.fast_term_ord o apply2 term_of; val term_ord = Term_Ord.term_ord o apply2 term_of; (* destructors *) fun dest_comb (Cterm {t = c $ a, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in (Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts}, Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts}) end | dest_comb ct = raise CTERM ("dest_comb", [ct]); fun dest_fun (Cterm {t = c $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = c, T = A --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun ct = raise CTERM ("dest_fun", [ct]); fun dest_arg (Cterm {t = c $ a, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg ct = raise CTERM ("dest_arg", [ct]); fun dest_fun2 (Cterm {t = c $ _ $ _, T, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0; val B = Term.argument_type_of c 1; in Cterm {t = c, T = A --> B --> T, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_fun2 ct = raise CTERM ("dest_fun2", [ct]); fun dest_arg1 (Cterm {t = c $ a $ _, T = _, cert, maxidx, sorts}) = let val A = Term.argument_type_of c 0 in Cterm {t = a, T = A, cert = cert, maxidx = maxidx, sorts = sorts} end | dest_arg1 ct = raise CTERM ("dest_arg1", [ct]); fun gen_dest_abs dest ct = (case ct of Cterm {t = t as Abs _, T = Type ("fun", [_, U]), cert, maxidx, sorts} => let val ((x', T), t') = dest t; val v = Cterm {t = Free (x', T), T = T, cert = cert, maxidx = maxidx, sorts = sorts}; val body = Cterm {t = t', T = U, cert = cert, maxidx = maxidx, sorts = sorts}; in (v, body) end | _ => raise CTERM ("dest_abs", [ct])); val dest_abs_fresh = gen_dest_abs o Term.dest_abs_fresh; val dest_abs_global = gen_dest_abs Term.dest_abs_global; (* 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} + sorts = Sortset.merge (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} + sorts = Sortset.merge (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*) + shyps: Sortset.T, (*sort hypotheses*) hyps: term Ord_List.T, (*hypotheses*) tpairs: (term * term) list, (*flex-flex pairs*) prop: term} (*conclusion*) and deriv = Deriv of {promises: (serial * thm future) Ord_List.T, body: Proofterm.proof_body}; type conv = cterm -> thm; (*errors involving theorems*) exception THM of string * int * thm list; fun rep_thm (Thm (_, args)) = args; fun fold_terms h f (Thm (_, {tpairs, prop, hyps, ...})) = fold (fn (t, u) => f t #> f u) tpairs #> f prop #> #hyps h ? fold f hyps; fun fold_atomic_ctyps h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun ctyp T = Ctyp {cert = cert, T = T, maxidx = maxidx, sorts = shyps} in (fold_terms h o fold_types o fold_atyps) (fn T => if g T then f (ctyp T) else I) th end; fun fold_atomic_cterms h g f (th as Thm (_, {cert, maxidx, shyps, ...})) = let fun cterm t T = Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = shyps}; fun apply t T = if g t then f (cterm t T) else I; in (fold_terms h o fold_aterms) (fn t as Const (_, T) => apply t T | t as Free (_, T) => apply t T | t as Var (_, T) => apply t T | _ => I) th end; fun terms_of_tpairs tpairs = fold_rev (fn (t, u) => cons t o cons u) tpairs []; fun eq_tpairs ((t, u), (t', u')) = t aconv t' andalso u aconv u'; fun union_tpairs ts us = Library.merge eq_tpairs (ts, us); val maxidx_tpairs = fold (fn (t, u) => Term.maxidx_term t #> Term.maxidx_term u); fun attach_tpairs tpairs prop = Logic.list_implies (map Logic.mk_equals tpairs, prop); fun full_prop_of (Thm (_, {tpairs, prop, ...})) = attach_tpairs tpairs prop; val union_hyps = Ord_List.union Term_Ord.fast_term_ord; val insert_hyps = Ord_List.insert Term_Ord.fast_term_ord; val remove_hyps = Ord_List.remove Term_Ord.fast_term_ord; fun join_certificate1 (Cterm {cert = cert1, ...}, Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); fun join_certificate2 (Thm (_, {cert = cert1, ...}), Thm (_, {cert = cert2, ...})) = Context.join_certificate (cert1, cert2); (* basic components *) val cert_of = #cert o rep_thm; val theory_id = Context.certificate_theory_id o cert_of; val theory_name = Context.theory_id_name o theory_id; val maxidx_of = #maxidx o rep_thm; fun maxidx_thm th i = Int.max (maxidx_of th, i); val shyps_of = #shyps o rep_thm; val hyps_of = #hyps o rep_thm; val prop_of = #prop o rep_thm; val tpairs_of = #tpairs o rep_thm; val concl_of = Logic.strip_imp_concl o prop_of; val prems_of = Logic.strip_imp_prems o prop_of; val nprems_of = Logic.count_prems o prop_of; val no_prems = Logic.no_prems o prop_of; fun major_prem_of th = (case prems_of th of prem :: _ => Logic.strip_assums_concl prem | [] => raise THM ("major_prem_of: rule with no premises", 0, [th])); fun cprop_of (Thm (_, {cert, maxidx, shyps, prop, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, t = prop, sorts = shyps}; fun cprem_of (th as Thm (_, {cert, maxidx, shyps, prop, ...})) i = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = Logic.nth_prem (i, prop) handle TERM _ => raise THM ("cprem_of", i, [th])}; fun cconcl_of (th as Thm (_, {cert, maxidx, shyps, ...})) = Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = concl_of th}; fun cprems_of (th as Thm (_, {cert, maxidx, shyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = maxidx, T = propT, sorts = shyps, t = t}) (prems_of th); fun chyps_of (Thm (_, {cert, shyps, hyps, ...})) = map (fn t => Cterm {cert = cert, maxidx = ~1, T = propT, sorts = shyps, t = t}) hyps; (* thm order: ignores theory context! *) val thm_ord = pointer_eq_ord (Term_Ord.fast_term_ord o apply2 prop_of ||| list_ord (prod_ord Term_Ord.fast_term_ord Term_Ord.fast_term_ord) o apply2 tpairs_of ||| list_ord Term_Ord.fast_term_ord o apply2 hyps_of - ||| list_ord Term_Ord.sort_ord o apply2 shyps_of); + ||| Sortset.ord o apply2 shyps_of); (* implicit theory context *) exception CONTEXT of string * ctyp list * cterm list * thm list * Context.generic option; fun theory_of_cterm (ct as Cterm {cert, ...}) = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct], [], NONE); fun theory_of_thm th = Context.certificate_theory (cert_of th) handle ERROR msg => raise CONTEXT (msg, [], [], [th], NONE); fun trim_context_ctyp cT = (case cT of Ctyp {cert = Context.Certificate_Id _, ...} => cT | Ctyp {cert = Context.Certificate thy, T, maxidx, sorts} => Ctyp {cert = Context.Certificate_Id (Context.theory_id thy), T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_cterm ct = (case ct of Cterm {cert = Context.Certificate_Id _, ...} => ct | Cterm {cert = Context.Certificate thy, t, T, maxidx, sorts} => Cterm {cert = Context.Certificate_Id (Context.theory_id thy), t = t, T = T, maxidx = maxidx, sorts = sorts}); fun trim_context_thm th = (case th of Thm (_, {constraints = _ :: _, ...}) => raise THM ("trim_context: pending sort constraints", 0, [th]) | Thm (_, {cert = Context.Certificate_Id _, ...}) => th | Thm (der, {cert = Context.Certificate thy, tags, maxidx, constraints = [], shyps, hyps, tpairs, prop}) => Thm (der, {cert = Context.Certificate_Id (Context.theory_id thy), tags = tags, maxidx = maxidx, constraints = [], shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop})); fun transfer_ctyp thy' cT = let val Ctyp {cert, T, maxidx, sorts} = cT; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [cT], [], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then cT else Ctyp {cert = cert', T = T, maxidx = maxidx, sorts = sorts} end; fun transfer_cterm thy' ct = let val Cterm {cert, t, T, maxidx, sorts} = ct; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [ct], [], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then ct else Cterm {cert = cert', t = t, T = T, maxidx = maxidx, sorts = sorts} end; fun transfer thy' th = let val Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) = th; val _ = Context.subthy_id (Context.certificate_theory_id cert, Context.theory_id thy') orelse raise CONTEXT ("Cannot transfer: not a super theory", [], [], [th], SOME (Context.Theory thy')); val cert' = Context.join_certificate (Context.Certificate thy', cert); in if Context.eq_certificate (cert, cert') then th else Thm (der, {cert = cert', tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; val transfer' = transfer o Proof_Context.theory_of; val transfer'' = transfer o Context.theory_of; fun join_transfer thy th = (Context.subthy_id (theory_id th, Context.theory_id thy) ? transfer thy) th; fun join_transfer_context (ctxt, th) = if Context.subthy_id (theory_id th, Context.theory_id (Proof_Context.theory_of ctxt)) then (ctxt, transfer' ctxt th) else (Context.raw_transfer (theory_of_thm th) ctxt, th); (* matching *) local fun gen_match match (ct1 as Cterm {t = t1, sorts = sorts1, ...}, ct2 as Cterm {t = t2, sorts = sorts2, maxidx = maxidx2, ...}) = let val cert = join_certificate0 (ct1, ct2); val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [], [ct1, ct2], [], NONE); val (Tinsts, tinsts) = match thy (t1, t2) (Vartab.empty, Vartab.empty); - val sorts = Sorts.union sorts1 sorts2; + val sorts = Sortset.merge (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 (TVars.build (Vartab.fold (TVars.add o mk_cTinst) Tinsts), Vars.build (Vartab.fold (Vars.add o mk_ctinst) tinsts)) end; in val match = gen_match Pattern.match; val first_order_match = gen_match Pattern.first_order_match; end; (*implicit alpha-conversion*) fun renamed_prop prop' (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if prop aconv prop' then Thm (der, {cert = cert, tags = tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) else raise TERM ("renamed_prop: props disagree", [prop, prop']); fun make_context ths NONE cert = (Context.Theory (Context.certificate_theory cert) handle ERROR msg => raise CONTEXT (msg, [], [], ths, NONE)) | make_context ths (SOME ctxt) cert = let val thy_id = Context.certificate_theory_id cert; val thy_id' = Context.theory_id (Proof_Context.theory_of ctxt); in if Context.subthy_id (thy_id, thy_id') then Context.Proof ctxt else raise CONTEXT ("Bad context", [], [], ths, SOME (Context.Proof ctxt)) end; fun make_context_certificate ths opt_ctxt cert = let val context = make_context ths opt_ctxt cert; val cert' = Context.Certificate (Context.theory_of context); in (context, cert') end; (*explicit weakening: maps |- B to A |- B*) fun weaken raw_ct th = let val ct as Cterm {t = A, T, sorts, maxidx = maxidxA, ...} = adjust_maxidx_cterm ~1 raw_ct; val Thm (der, {tags, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; in if T <> propT then raise THM ("weaken: assumptions must have type prop", 0, []) else if maxidxA <> ~1 then raise THM ("weaken: assumptions may not contain schematic variables", maxidxA, []) else Thm (der, {cert = join_certificate1 (ct, th), tags = tags, maxidx = maxidx, constraints = constraints, - shyps = Sorts.union sorts shyps, + shyps = Sortset.merge (sorts, shyps), hyps = insert_hyps A hyps, tpairs = tpairs, prop = prop}) end; -fun weaken_sorts raw_sorts ct = +fun weaken_sorts more_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; + val sorts' = Sortset.fold (Sortset.insert o Sign.certify_sort thy) more_sorts sorts; in Cterm {cert = cert, t = t, T = T, maxidx = maxidx, sorts = sorts'} end; (** derivations and promised proofs **) fun make_deriv promises oracles thms proof = Deriv {promises = promises, body = PBody {oracles = oracles, thms = thms, proof = proof}}; val empty_deriv = make_deriv [] [] [] MinProof; (* inference rules *) val promise_ord: (serial * thm future) ord = fn ((i, _), (j, _)) => int_ord (j, i); fun bad_proofs i = error ("Illegal level of detail for proof objects: " ^ string_of_int i); fun deriv_rule2 f (Deriv {promises = ps1, body = PBody {oracles = oracles1, thms = thms1, proof = prf1}}) (Deriv {promises = ps2, body = PBody {oracles = oracles2, thms = thms2, proof = prf2}}) = let val ps = Ord_List.union promise_ord ps1 ps2; val oracles = Proofterm.unions_oracles [oracles1, oracles2]; val thms = Proofterm.unions_thms [thms1, thms2]; val prf = (case ! Proofterm.proofs of 2 => f prf1 prf2 | 1 => MinProof | 0 => MinProof | i => bad_proofs i); in make_deriv ps oracles thms prf end; fun deriv_rule1 f = deriv_rule2 (K f) empty_deriv; fun deriv_rule0 make_prf = if ! Proofterm.proofs <= 1 then empty_deriv else deriv_rule1 I (make_deriv [] [] [] (make_prf ())); fun deriv_rule_unconditional f (Deriv {promises, body = PBody {oracles, thms, proof}}) = make_deriv promises oracles thms (f proof); (* fulfilled proofs *) fun raw_promises_of (Thm (Deriv {promises, ...}, _)) = promises; fun join_promises [] = () | join_promises promises = join_promises_of (Future.joins (map snd promises)) and join_promises_of thms = join_promises (Ord_List.make promise_ord (maps raw_promises_of thms)); fun fulfill_body (th as Thm (Deriv {promises, body}, _)) = let val fulfilled_promises = map #1 promises ~~ map fulfill_body (Future.joins (map #2 promises)) in Proofterm.fulfill_norm_proof (theory_of_thm th) fulfilled_promises body end; fun proof_bodies_of thms = (join_promises_of thms; map fulfill_body thms); val proof_body_of = singleton proof_bodies_of; val proof_of = Proofterm.proof_of o proof_body_of; fun reconstruct_proof_of thm = Proofterm.reconstruct_proof (theory_of_thm thm) (prop_of thm) (proof_of thm); val consolidate = ignore o proof_bodies_of; fun expose_proofs thy thms = if Proofterm.export_proof_boxes_required thy then Proofterm.export_proof_boxes (proof_bodies_of (map (transfer thy) thms)) else (); fun expose_proof thy = expose_proofs thy o single; (* future rule *) fun future_result i orig_cert orig_shyps orig_prop thm = let fun err msg = raise THM ("future_result: " ^ msg, 0, [thm]); val Thm (Deriv {promises, ...}, {cert, constraints, shyps, hyps, tpairs, prop, ...}) = thm; val _ = Context.eq_certificate (cert, orig_cert) orelse err "bad theory"; val _ = prop aconv orig_prop orelse err "bad prop"; val _ = null constraints orelse err "bad sort constraints"; val _ = null tpairs orelse err "bad flex-flex constraints"; val _ = null hyps orelse err "bad hyps"; - val _ = Sorts.subset (shyps, orig_shyps) orelse err "bad shyps"; + val _ = Sortset.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 []; + val shyps = Sortset.build (Sortset.insert_term prop); in Thm (der, {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = shyps, hyps = [], tpairs = [], prop = prop}) end | NONE => raise THEORY ("No axiom " ^ quote name, [thy])); fun all_axioms_of thy = map (fn (name, _) => (name, axiom thy name)) (Theory.all_axioms_of thy); (* tags *) val get_tags = #tags o rep_thm; fun map_tags f (Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = Thm (der, {cert = cert, tags = f tags, maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (* technical adjustments *) fun norm_proof (th as Thm (der, args)) = Thm (deriv_rule1 (Proofterm.rew_proof (theory_of_thm th)) der, args); fun adjust_maxidx_thm i (th as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = if maxidx = i then th else if maxidx < i then Thm (der, {maxidx = i, cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) else Thm (der, {maxidx = Int.max (maxidx_tpairs tpairs (maxidx_of_term prop), i), cert = cert, tags = tags, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}); (*** Theory data ***) (* type classes *) structure Aritytab = Table( type key = string * sort list * class; val ord = fast_string_ord o apply2 #1 ||| fast_string_ord o apply2 #3 ||| list_ord Term_Ord.sort_ord o apply2 #2; ); datatype classes = Classes of {classrels: thm Symreltab.table, arities: (thm * string * serial) Aritytab.table}; fun make_classes (classrels, arities) = Classes {classrels = classrels, arities = arities}; val empty_classes = make_classes (Symreltab.empty, Aritytab.empty); (*see Theory.at_begin hook for transitive closure of classrels and arity completion*) fun merge_classes (Classes {classrels = classrels1, arities = arities1}, Classes {classrels = classrels2, arities = arities2}) = let val classrels' = Symreltab.merge (K true) (classrels1, classrels2); val arities' = Aritytab.merge (K true) (arities1, arities2); in make_classes (classrels', arities') end; (* data *) structure Data = Theory_Data ( type T = unit Name_Space.table * (*oracles: authentic derivation names*) classes; (*type classes within the logic*) val empty : T = (Name_Space.empty_table Markup.oracleN, empty_classes); fun merge ((oracles1, sorts1), (oracles2, sorts2)) : T = (Name_Space.merge_tables (oracles1, oracles2), merge_classes (sorts1, sorts2)); ); val get_oracles = #1 o Data.get; val map_oracles = Data.map o apfst; val get_classes = (fn (_, Classes args) => args) o Data.get; val get_classrels = #classrels o get_classes; val get_arities = #arities o get_classes; fun map_classes f = (Data.map o apsnd) (fn Classes {classrels, arities} => make_classes (f (classrels, arities))); fun map_classrels f = map_classes (fn (classrels, arities) => (f classrels, arities)); fun map_arities f = map_classes (fn (classrels, arities) => (classrels, f arities)); (* type classes *) fun the_classrel thy (c1, c2) = (case Symreltab.lookup (get_classrels thy) (c1, c2) of SOME thm => transfer thy thm | NONE => error ("Unproven class relation " ^ Syntax.string_of_classrel (Proof_Context.init_global thy) [c1, c2])); fun the_arity thy (a, Ss, c) = (case Aritytab.lookup (get_arities thy) (a, Ss, c) of SOME (thm, _, _) => transfer thy thm | NONE => error ("Unproven type arity " ^ Syntax.string_of_arity (Proof_Context.init_global thy) (a, Ss, [c]))); val classrel_proof = proof_of oo the_classrel; val arity_proof = proof_of oo the_arity; (* solve sort constraints by pro-forma proof *) local fun union_digest (oracles1, thms1) (oracles2, thms2) = (Proofterm.unions_oracles [oracles1, oracles2], Proofterm.unions_thms [thms1, thms2]); fun thm_digest (Thm (Deriv {body = PBody {oracles, thms, ...}, ...}, _)) = (oracles, thms); fun constraint_digest ({theory = thy, typ, sort, ...}: constraint) = Sorts.of_sort_derivation (Sign.classes_of thy) {class_relation = fn _ => fn _ => fn (digest, c1) => fn c2 => if c1 = c2 then ([], []) else union_digest digest (thm_digest (the_classrel thy (c1, c2))), type_constructor = fn (a, _) => fn dom => fn c => let val arity_digest = thm_digest (the_arity thy (a, (map o map) #2 dom, c)) in (fold o fold) (union_digest o #1) dom arity_digest end, type_variable = fn T => map (pair ([], [])) (Type.sort_of_atyp T)} (typ, sort); in fun solve_constraints (thm as Thm (_, {constraints = [], ...})) = thm | solve_constraints (thm as Thm (der, args)) = let val {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop} = args; val thy = Context.certificate_theory cert; val bad_thys = constraints |> map_filter (fn {theory = thy', ...} => if Context.eq_thy (thy, thy') then NONE else SOME thy'); val () = if null bad_thys then () else raise THEORY ("solve_constraints: bad theories for theorem\n" ^ Syntax.string_of_term_global thy (prop_of thm), thy :: bad_thys); val Deriv {promises, body = PBody {oracles, thms, proof}} = der; val (oracles', thms') = (oracles, thms) |> fold (fold union_digest o constraint_digest) constraints; val body' = PBody {oracles = oracles', thms = thms', proof = proof}; in Thm (Deriv {promises = promises, body = body'}, {constraints = [], cert = cert, tags = tags, maxidx = maxidx, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop}) end; end; (*Dangling sort constraints of a thm*) fun extra_shyps (th as Thm (_, {shyps, ...})) = - Sorts.subtract (fold_terms {hyps = true} Sorts.insert_term th []) shyps; + Sortset.subtract (Sortset.build (fold_terms {hyps = true} Sortset.insert_term th)) shyps; (*Remove extra sorts that are witnessed by type signature information*) -fun strip_shyps thm = - (case thm of - Thm (_, {shyps = [], ...}) => thm - | Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop}) => +fun strip_shyps (thm as Thm (der, {cert, tags, maxidx, constraints, shyps, hyps, tpairs, prop})) = + solve_constraints ( + if Sortset.is_empty shyps then thm + else let val thy = theory_of_thm thm; val algebra = Sign.classes_of thy; val minimize = Sorts.minimize_sort algebra; val le = Sorts.sort_le algebra; fun lt (S1, S2) = le (S1, S2) andalso not (le (S2, S1)); fun rel (S1, S2) = if S1 = S2 then [] else [(Term.aT S1, S2)]; val present = (fold_terms {hyps = true} o fold_types o fold_atyps_sorts) (insert (eq_fst op =)) thm []; - val extra = fold (Sorts.remove_sort o #2) present shyps; - val witnessed = Sign.witness_sorts thy present extra; - val non_witnessed = fold (Sorts.remove_sort o #2) witnessed extra |> map (`minimize); + val extra = fold (Sortset.remove o #2) present shyps; + val witnessed = Sign.witness_sorts thy present (Sortset.dest extra); + val non_witnessed = + fold (Sortset.remove o #2) witnessed extra |> Sortset.dest |> 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; + |> Sortset.make; val constrs' = non_witnessed |> maps (fn (S1, S2) => - let val S0 = the (find_first (fn S => le (S, S1)) extra') + let val S0 = the (Sortset.get_first (fn S => if le (S, S1) then SOME S else NONE) 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'; + val shyps' = fold (Sortset.insert 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; - + 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 = strip_shyps #> (fn thm as Thm (der, args) => let val thy = theory_of_thm thm; val Deriv {promises, body} = der; val {shyps, hyps, prop, tpairs, ...} = args; val _ = null tpairs orelse raise THM ("name_derivation: bad flex-flex constraints", 0, [thm]); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.thm_proof thy (classrel_proof thy) (arity_proof thy) name_pos shyps hyps prop ps body; val der' = make_deriv [] [] [pthm] proof; in Thm (der', args) end); fun close_derivation pos = solve_constraints #> (fn thm => if not (null (tpairs_of thm)) orelse derivation_closed thm then thm else name_derivation ("", pos) thm); val trim_context = solve_constraints #> trim_context_thm; (*** Oracles ***) fun add_oracle (b, oracle_fn) thy = let val (name, oracles') = Name_Space.define (Context.Theory thy) true (b, ()) (get_oracles thy); val thy' = map_oracles (K oracles') thy; fun invoke_oracle arg = let val Cterm {cert = cert2, t = prop, T, maxidx, sorts} = oracle_fn arg in if T <> propT then raise THM ("Oracle's result must have type prop: " ^ name, 0, []) else let val (oracle, prf) = (case ! Proofterm.proofs of 2 => (((name, Position.thread_data ()), SOME prop), Proofterm.oracle_proof name prop) | 1 => (((name, Position.thread_data ()), SOME prop), MinProof) | 0 => (((name, Position.none), NONE), MinProof) | i => bad_proofs i); in Thm (make_deriv [] [oracle] [] prf, {cert = Context.join_certificate (Context.Certificate thy', cert2), tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) end end; in ((name, invoke_oracle), thy') end; val oracle_space = Name_Space.space_of_table o get_oracles; fun pretty_oracle ctxt = Name_Space.pretty ctxt (oracle_space (Proof_Context.theory_of ctxt)); fun extern_oracles verbose ctxt = map #1 (Name_Space.markup_table verbose ctxt (get_oracles (Proof_Context.theory_of ctxt))); fun check_oracle ctxt = Name_Space.check (Context.Proof ctxt) (get_oracles (Proof_Context.theory_of ctxt)) #> #1; (*** Meta rules ***) (** primitive rules **) (*The assumption rule A |- A*) fun assume raw_ct = let val Cterm {cert, t = prop, T, maxidx, sorts} = adjust_maxidx_cterm ~1 raw_ct in if T <> propT then raise THM ("assume: prop", 0, []) else if maxidx <> ~1 then raise THM ("assume: variables", maxidx, []) else Thm (deriv_rule0 (fn () => Proofterm.Hyp prop), {cert = cert, tags = [], maxidx = ~1, constraints = [], shyps = sorts, hyps = [prop], tpairs = [], prop = prop}) end; (*Implication introduction [A] : B ------- A \ B *) fun implies_intr (ct as Cterm {t = A, T, maxidx = maxidx1, sorts, ...}) (th as Thm (der, {maxidx = maxidx2, hyps, constraints, shyps, tpairs, prop, ...})) = if T <> propT then raise THM ("implies_intr: assumptions must have type prop", 0, [th]) else Thm (deriv_rule1 (Proofterm.implies_intr_proof A) der, {cert = join_certificate1 (ct, th), tags = [], maxidx = Int.max (maxidx1, maxidx2), constraints = constraints, - shyps = Sorts.union sorts shyps, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (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, + shyps = Sortset.merge (shyps1, shyps2), hyps = union_hyps hyps1 hyps2, tpairs = union_tpairs tpairs1 tpairs2, prop = B}) else err "not equal" | _ => err "major premise") end; (**** Derived rules ****) (*Smash unifies the list of term pairs leaving no flex-flex pairs. Instantiates the theorem and deletes trivial tpairs. Resulting sequence may contain multiple elements if the tpairs are not all flex-flex.*) fun flexflex_rule opt_ctxt = solve_constraints #> (fn th => let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val (context, cert') = make_context_certificate [th] opt_ctxt cert; in Unify.smash_unifiers context tpairs (Envir.empty maxidx) |> Seq.map (fn env => if Envir.is_empty env then th else let val tpairs' = tpairs |> map (apply2 (Envir.norm_term env)) (*remove trivial tpairs, of the form t \ t*) |> filter_out (op aconv); val der' = deriv_rule1 (Proofterm.norm_proof' env) der; val constraints' = insert_constraints_env (Context.certificate_theory cert') env constraints; val prop' = Envir.norm_term env prop; val maxidx = maxidx_tpairs tpairs' (maxidx_of_term prop'); val shyps = Envir.insert_sorts env shyps; in Thm (der', {cert = cert', tags = [], maxidx = maxidx, constraints = constraints', shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end) end); (*Generalization of fixed variables A -------------------- A[?'a/'a, ?x/x, ...] *) fun generalize (tfrees, frees) idx th = if Names.is_empty tfrees andalso Names.is_empty frees then th else let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = th; val _ = idx <= maxidx andalso raise THM ("generalize: bad index", idx, [th]); val bad_type = if Names.is_empty tfrees then K false else Term.exists_subtype (fn TFree (a, _) => Names.defined tfrees a | _ => false); fun bad_term (Free (x, T)) = bad_type T orelse Names.defined frees x | bad_term (Var (_, T)) = bad_type T | bad_term (Const (_, T)) = bad_type T | bad_term (Abs (_, T, t)) = bad_type T orelse bad_term t | bad_term (t $ u) = bad_term t orelse bad_term u | bad_term (Bound _) = false; val _ = exists bad_term hyps andalso raise THM ("generalize: variable free in assumptions", 0, [th]); val generalize = Term_Subst.generalize (tfrees, frees) idx; val prop' = generalize prop; val tpairs' = map (apply2 generalize) tpairs; val maxidx' = maxidx_tpairs tpairs' (maxidx_of_term prop'); in Thm (deriv_rule1 (Proofterm.generalize_proof (tfrees, frees) idx prop) der, {cert = cert, tags = [], maxidx = maxidx', constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs', prop = prop'}) end; fun generalize_cterm (tfrees, frees) idx (ct as Cterm {cert, t, T, maxidx, sorts}) = if Names.is_empty tfrees andalso Names.is_empty frees then ct else if idx <= maxidx then raise CTERM ("generalize_cterm: bad index", [ct]) else Cterm {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, t = Term_Subst.generalize (tfrees, frees) idx t, maxidx = Int.max (maxidx, idx)}; fun generalize_ctyp tfrees idx (cT as Ctyp {cert, T, maxidx, sorts}) = if Names.is_empty tfrees then cT else if idx <= maxidx then raise CTERM ("generalize_ctyp: bad index", []) else Ctyp {cert = cert, sorts = sorts, T = Term_Subst.generalizeT tfrees idx T, maxidx = Int.max (maxidx, idx)}; (*Instantiation of schematic variables A -------------------- A[t1/v1, ..., tn/vn] *) local fun add_cert cert_of (_, c) cert = Context.join_certificate (cert, cert_of c); val add_instT_cert = add_cert (fn Ctyp {cert, ...} => cert); val add_inst_cert = add_cert (fn Cterm {cert, ...} => cert); -fun add_sorts sorts_of (_, c) sorts = Sorts.union (sorts_of c) sorts; +fun add_sorts sorts_of (_, c) sorts = Sortset.merge (sorts_of c, sorts); val add_instT_sorts = add_sorts (fn Ctyp {sorts, ...} => sorts); val add_inst_sorts = add_sorts (fn Cterm {sorts, ...} => sorts); fun make_instT thy (v as (_, S)) (Ctyp {T = U, maxidx, ...}) = if Sign.of_sort thy (U, S) then (U, maxidx) else raise TYPE ("Type not of sort " ^ Syntax.string_of_sort_global thy S, [U], []); fun make_inst thy (v as (_, T)) (Cterm {t = u, T = U, maxidx, ...}) = if T = U then (u, maxidx) else let fun pretty_typing t ty = Pretty.block [Syntax.pretty_term_global thy t, Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ_global thy ty]; val msg = Pretty.string_of (Pretty.block [Pretty.str "instantiate: type conflict", Pretty.fbrk, pretty_typing (Var v) T, Pretty.fbrk, pretty_typing u U]) in raise TYPE (msg, [T, U], [Var v, u]) end; fun prep_insts (instT, inst) (cert, sorts) = let val cert' = cert |> TVars.fold add_instT_cert instT |> Vars.fold add_inst_cert inst; val thy' = Context.certificate_theory cert' handle ERROR msg => raise CONTEXT (msg, TVars.dest instT |> map #2, Vars.dest inst |> map #2, [], NONE); val sorts' = sorts |> TVars.fold add_instT_sorts instT |> Vars.fold add_inst_sorts inst; val instT' = TVars.map (make_instT thy') instT; val inst' = Vars.map (make_inst thy') inst; in ((instT', inst'), (cert', sorts')) end; in (*Left-to-right replacements: ctpairs = [..., (vi, ti), ...]. Instantiates distinct Vars by terms of same type. Does NOT normalize the resulting theorem!*) fun gen_instantiate inst_fn (instT, inst) th = if TVars.is_empty instT andalso Vars.is_empty inst then th else let val Thm (der, {cert, hyps, constraints, shyps, tpairs, prop, ...}) = th; val ((instT', inst'), (cert', shyps')) = prep_insts (instT, inst) (cert, shyps) handle CONTEXT (msg, cTs, cts, ths, context) => raise CONTEXT (msg, cTs, cts, th :: ths, context); val subst = inst_fn (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' = TVars.fold (fn ((_, S), (T, _)) => insert_constraints thy' (T, S)) instT' constraints; in Thm (deriv_rule1 (fn d => Proofterm.instantiate (TVars.map (K #1) instT', Vars.map (K #1) inst') d) der, {cert = cert', tags = [], maxidx = maxidx', constraints = constraints', shyps = shyps', hyps = hyps, tpairs = tpairs', prop = prop'}) |> solve_constraints end handle TYPE (msg, _, _) => raise THM (msg, 0, [th]); val instantiate = gen_instantiate Term_Subst.instantiate_maxidx; val instantiate_beta = gen_instantiate Term_Subst.instantiate_beta_maxidx; fun gen_instantiate_cterm inst_fn (instT, inst) ct = if TVars.is_empty instT andalso Vars.is_empty inst then ct else let val Cterm {cert, t, T, sorts, ...} = ct; val ((instT', inst'), (cert', sorts')) = prep_insts (instT, inst) (cert, sorts); val subst = inst_fn (instT', inst'); val substT = Term_Subst.instantiateT_maxidx instT'; val (t', maxidx1) = subst t ~1; val (T', maxidx') = substT T maxidx1; in Cterm {cert = cert', t = t', T = T', sorts = sorts', maxidx = maxidx'} end handle TYPE (msg, _, _) => raise CTERM (msg, [ct]); val instantiate_cterm = gen_instantiate_cterm Term_Subst.instantiate_maxidx; val instantiate_beta_cterm = gen_instantiate_cterm Term_Subst.instantiate_beta_maxidx; end; (*The trivial implication A \ A, justified by assume and forall rules. A can contain Vars, not so for assume!*) fun trivial (Cterm {cert, t = A, T, maxidx, sorts}) = if T <> propT then raise THM ("trivial: the term must have type prop", 0, []) else Thm (deriv_rule0 (fn () => Proofterm.trivial_proof), {cert = cert, tags = [], maxidx = maxidx, constraints = [], shyps = sorts, hyps = [], tpairs = [], prop = Logic.mk_implies (A, A)}); (*Axiom-scheme reflecting signature contents T :: c ------------------- OFCLASS(T, c_class) *) fun of_class (cT, raw_c) = let val Ctyp {cert, T, ...} = cT; val thy = Context.certificate_theory cert handle ERROR msg => raise CONTEXT (msg, [cT], [], [], NONE); val c = Sign.certify_class thy raw_c; val Cterm {t = prop, maxidx, sorts, ...} = global_cterm_of thy (Logic.mk_of_class (T, c)); in if Sign.of_sort thy (T, [c]) then Thm (deriv_rule0 (fn () => Proofterm.PClass (T, c)), {cert = cert, tags = [], maxidx = maxidx, constraints = insert_constraints thy (T, [c]) [], shyps = sorts, hyps = [], tpairs = [], prop = prop}) else raise THM ("of_class: type not of class " ^ Syntax.string_of_sort_global thy [c], 0, []) end |> solve_constraints; (*Internalize sort constraints of type variables*) val unconstrainT = strip_shyps #> (fn thm as Thm (der, args) => let val Deriv {promises, body} = der; val {cert, shyps, hyps, tpairs, prop, ...} = args; val thy = theory_of_thm thm; fun err msg = raise THM ("unconstrainT: " ^ msg, 0, [thm]); val _ = null hyps orelse err "bad hyps"; val _ = null tpairs orelse err "bad flex-flex constraints"; val tfrees = build_rev (Term.add_tfree_names prop); val _ = null tfrees orelse err ("illegal free type variables " ^ commas_quote tfrees); val ps = map (apsnd (Future.map fulfill_body)) promises; val (pthm, proof) = Proofterm.unconstrain_thm_proof thy (classrel_proof thy) (arity_proof thy) shyps prop ps body; val der' = make_deriv [] [] [pthm] proof; val prop' = Proofterm.thm_node_prop (#2 pthm); in Thm (der', {cert = cert, tags = [], maxidx = maxidx_of_term prop', constraints = [], - shyps = [[]], (*potentially redundant*) + shyps = Sortset.make [[]], (*potentially redundant*) hyps = [], tpairs = [], prop = prop'}) end); (*Replace all TFrees not fixed or in the hyps by new TVars*) fun varifyT_global' fixed (Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = let val tfrees = fold TFrees.add_tfrees hyps fixed; val prop1 = attach_tpairs tpairs prop; val (al, prop2) = Type.varify_global tfrees prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in (al, Thm (deriv_rule1 (Proofterm.varify_proof prop tfrees) der, {cert = cert, tags = [], maxidx = Int.max (0, maxidx), constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3})) end; val varifyT_global = #2 o varifyT_global' TFrees.empty; (*Replace all TVars by TFrees that are often new*) fun legacy_freezeT (Thm (der, {cert, constraints, shyps, hyps, tpairs, prop, ...})) = let val prop1 = attach_tpairs tpairs prop; val prop2 = Type.legacy_freeze prop1; val (ts, prop3) = Logic.strip_prems (length tpairs, [], prop2); in Thm (deriv_rule1 (Proofterm.legacy_freezeT prop1) der, {cert = cert, tags = [], maxidx = maxidx_of_term prop2, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = rev (map Logic.dest_equals ts), prop = prop3}) end; fun plain_prop_of raw_thm = let val thm = strip_shyps raw_thm; fun err msg = raise THM ("plain_prop_of: " ^ msg, 0, [thm]); in if not (null (hyps_of thm)) then err "theorem may not contain hypotheses" - else if not (null (extra_shyps thm)) then + else if not (Sortset.is_empty (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!*) + shyps = Sortset.merge (sorts, shyps), hyps = hyps, tpairs = map (apply2 lift_abs) tpairs, prop = Logic.list_implies (map lift_all As, lift_all B)}) end; fun incr_indexes i (thm as Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...})) = if i < 0 then raise THM ("negative increment", 0, [thm]) else if i = 0 then thm else Thm (deriv_rule1 (Proofterm.incr_indexes i) der, {cert = cert, tags = [], maxidx = maxidx + i, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = map (apply2 (Logic.incr_indexes ([], [], i))) tpairs, prop = Logic.incr_indexes ([], [], i) prop}); (*Solve subgoal Bi of proof state B1...Bn/C by assumption. *) fun assumption opt_ctxt i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (context, cert') = make_context_certificate [state] opt_ctxt cert; val (tpairs, Bs, Bi, C) = dest_state (state, i); fun newth n (env, tpairs) = let val normt = Envir.norm_term env; fun assumption_proof prf = Proofterm.assumption_proof (map normt Bs) (normt Bi) n prf; in Thm (deriv_rule1 (assumption_proof #> not (Envir.is_empty env) ? Proofterm.norm_proof' env) der, {tags = [], maxidx = Envir.maxidx_of env, constraints = insert_constraints_env (Context.certificate_theory cert') env constraints, shyps = Envir.insert_sorts env shyps, hyps = hyps, tpairs = if Envir.is_empty env then tpairs else map (apply2 normt) tpairs, prop = if Envir.is_empty env then Logic.list_implies (Bs, C) (*avoid wasted normalizations*) else normt (Logic.list_implies (Bs, C)) (*normalize the new rule fully*), cert = cert'}) end; val (close, asms, concl) = Logic.assum_problems (~1, Bi); val concl' = close concl; fun addprfs [] _ = Seq.empty | addprfs (asm :: rest) n = Seq.make (fn () => Seq.pull (Seq.mapp (newth n) (if Term.could_unify (asm, concl) then (Unify.unifiers (context, Envir.empty maxidx, (close asm, concl') :: tpairs)) else Seq.empty) (addprfs rest (n + 1)))) in addprfs asms 1 end; (*Solve subgoal Bi of proof state B1...Bn/C by assumption. Checks if Bi's conclusion is alpha/eta-convertible to one of its assumptions*) fun eq_assumption i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val (_, asms, concl) = Logic.assum_problems (~1, Bi); in (case find_index (fn asm => Envir.aeconv (asm, concl)) asms of ~1 => raise THM ("eq_assumption", 0, [state]) | n => Thm (deriv_rule1 (Proofterm.assumption_proof Bs Bi (n + 1)) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs, C)})) end; (*For rotate_tac: fast rotation of assumptions of subgoal i*) fun rotate_rule k i state = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, ...}) = state; val (tpairs, Bs, Bi, C) = dest_state (state, i); val params = Term.strip_all_vars Bi; val rest = Term.strip_all_body Bi; val asms = Logic.strip_imp_prems rest val concl = Logic.strip_imp_concl rest; val n = length asms; val m = if k < 0 then n + k else k; val Bi' = if 0 = m orelse m = n then Bi else if 0 < m andalso m < n then let val (ps, qs) = chop m asms in Logic.list_all (params, Logic.list_implies (qs @ ps, concl)) end else raise THM ("rotate_rule", k, [state]); in Thm (deriv_rule1 (Proofterm.rotate_proof Bs Bi' params asms m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = Logic.list_implies (Bs @ [Bi'], C)}) end; (*Rotates a rule's premises to the left by k, leaving the first j premises unchanged. Does nothing if k=0 or if k equals n-j, where n is the number of premises. Useful with eresolve_tac and underlies defer_tac*) fun permute_prems j k rl = let val Thm (der, {cert, maxidx, constraints, shyps, hyps, tpairs, prop, ...}) = rl; val prems = Logic.strip_imp_prems prop and concl = Logic.strip_imp_concl prop; val moved_prems = List.drop (prems, j) and fixed_prems = List.take (prems, j) handle General.Subscript => raise THM ("permute_prems: j", j, [rl]); val n_j = length moved_prems; val m = if k < 0 then n_j + k else k; val (prems', prop') = if 0 = m orelse m = n_j then (prems, prop) else if 0 < m andalso m < n_j then let val (ps, qs) = chop m moved_prems; val prems' = fixed_prems @ qs @ ps; in (prems', Logic.list_implies (prems', concl)) end else raise THM ("permute_prems: k", k, [rl]); in Thm (deriv_rule1 (Proofterm.permute_prems_proof prems' j m) der, {cert = cert, tags = [], maxidx = maxidx, constraints = constraints, shyps = shyps, hyps = hyps, tpairs = tpairs, prop = prop'}) end; (* strip_apply f B A strips off all assumptions/parameters from A introduced by lifting over B, and applies f to remaining part of A*) fun strip_apply f = let fun strip (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ A2 $ B2) = Logic.mk_implies (A2, strip B1 B2) | strip ((c as Const ("Pure.all", _)) $ Abs (_, _, t1)) ( Const ("Pure.all", _) $ Abs (a, T, t2)) = c $ Abs (a, T, strip t1 t2) | strip _ A = f A in strip end; fun strip_lifted (Const ("Pure.imp", _) $ _ $ B1) (Const ("Pure.imp", _) $ _ $ B2) = strip_lifted B1 B2 | strip_lifted (Const ("Pure.all", _) $ Abs (_, _, t1)) (Const ("Pure.all", _) $ Abs (_, _, t2)) = strip_lifted t1 t2 | strip_lifted _ A = A; (*Use the alist to rename all bound variables and some unknowns in a term dpairs = current disagreement pairs; tpairs = permanent ones (flexflex); Preserves unknowns in tpairs and on lhs of dpairs. *) fun rename_bvs [] _ _ _ _ = K I | rename_bvs al dpairs tpairs B As = let val add_var = fold_aterms (fn Var ((x, _), _) => insert (op =) x | _ => I); val vids = [] |> fold (add_var o fst) dpairs |> fold (add_var o fst) tpairs |> fold (add_var o snd) tpairs; val vids' = fold (add_var o strip_lifted B) As []; (*unknowns appearing elsewhere be preserved!*) val al' = distinct ((op =) o apply2 fst) (filter_out (fn (x, y) => not (member (op =) vids' x) orelse member (op =) vids x orelse member (op =) vids y) al); val unchanged = filter_out (AList.defined (op =) al') vids'; fun del_clashing clash xs _ [] qs = if clash then del_clashing false xs xs qs [] else qs | del_clashing clash xs ys ((p as (x, y)) :: ps) qs = if member (op =) ys y then del_clashing true (x :: xs) (x :: ys) ps qs else del_clashing clash xs (y :: ys) ps (p :: qs); val al'' = del_clashing false unchanged unchanged al' []; fun rename (t as Var ((x, i), T)) = (case AList.lookup (op =) al'' x of SOME y => Var ((y, i), T) | NONE => t) | rename (Abs (x, T, t)) = Abs (the_default x (AList.lookup (op =) al x), T, rename t) | rename (f $ t) = rename f $ rename t | rename t = t; fun strip_ren f Ai = f rename B Ai in strip_ren end; (*Function to rename bounds/unknowns in the argument, lifted over B*) fun rename_bvars dpairs = rename_bvs (fold_rev Term.match_bvars dpairs []) dpairs; (*** RESOLUTION ***) (** Lifting optimizations **) (*strip off pairs of assumptions/parameters in parallel -- they are identical because of lifting*) fun strip_assums2 (Const("Pure.imp", _) $ _ $ B1, Const("Pure.imp", _) $ _ $ B2) = strip_assums2 (B1,B2) | strip_assums2 (Const("Pure.all",_)$Abs(a,T,t1), Const("Pure.all",_)$Abs(_,_,t2)) = let val (B1,B2) = strip_assums2 (t1,t2) in (Abs(a,T,B1), Abs(a,T,B2)) end | strip_assums2 BB = BB; (*Faster normalization: skip assumptions that were lifted over*) fun norm_term_skip env 0 t = Envir.norm_term env t | norm_term_skip env n (Const ("Pure.all", _) $ Abs (a, T, t)) = let val T' = Envir.norm_type (Envir.type_env env) T (*Must instantiate types of parameters because they are flattened; this could be a NEW parameter*) in Logic.all_const T' $ Abs (a, T', norm_term_skip env n t) end | norm_term_skip env n (Const ("Pure.imp", _) $ A $ B) = Logic.mk_implies (A, norm_term_skip env (n - 1) B) | norm_term_skip _ _ _ = error "norm_term_skip: too few assumptions??"; (*unify types of schematic variables (non-lifted case)*) fun unify_var_types context (th1, th2) env = let fun unify_vars (T :: Us) = fold (fn U => Pattern.unify_types context (T, U)) Us | unify_vars _ = I; val add_vars = full_prop_of #> fold_aterms (fn Var v => Vartab.insert_list (op =) v | _ => I); val vars = Vartab.build (add_vars th1 #> add_vars th2); in SOME (Vartab.fold (unify_vars o #2) vars env) end handle Pattern.Unif => NONE; (*Composition of object rule r=(A1...Am/B) with proof state s=(B1...Bn/C) Unifies B with Bi, replacing subgoal i (1 <= i <= n) If match then forbid instantiations in proof state If lifted then shorten the dpair using strip_assums2. If eres_flg then simultaneously proves A1 by assumption. nsubgoal is the number of new subgoals (written m above). Curried so that resolution calls dest_state only once. *) local exception COMPOSE in fun bicompose_aux opt_ctxt {flatten, match, incremented} (state, (stpairs, Bs, Bi, C), lifted) (eres_flg, orule, nsubgoal) = let val Thm (sder, {maxidx=smax, constraints = constraints2, shyps = shyps2, hyps = hyps2, ...}) = state and Thm (rder, {maxidx=rmax, constraints = constraints1, shyps = shyps1, hyps = hyps1, tpairs=rtpairs, prop=rprop,...}) = orule (*How many hyps to skip over during normalization*) and nlift = Logic.count_prems (strip_all_body Bi) + (if eres_flg then ~1 else 0) val (context, cert) = make_context_certificate [state, orule] opt_ctxt (join_certificate2 (state, orule)); (*Add new theorem with prop = "\Bs; As\ \ C" to thq*) fun addth A (As, oldAs, rder', n) ((env, tpairs), thq) = let val normt = Envir.norm_term env; (*perform minimal copying here by examining env*) val (ntpairs, normp) = if Envir.is_empty env then (tpairs, (Bs @ As, C)) else let val ntps = map (apply2 normt) tpairs in if Envir.above env smax then (*no assignments in state; normalize the rule only*) if lifted then (ntps, (Bs @ map (norm_term_skip env nlift) As, C)) else (ntps, (Bs @ map normt As, C)) else if match then raise COMPOSE else (*normalize the new rule fully*) (ntps, (map normt (Bs @ As), normt C)) end val constraints' = union_constraints constraints1 constraints2 |> insert_constraints_env (Context.certificate_theory cert) env; fun bicompose_proof prf1 prf2 = Proofterm.bicompose_proof flatten (map normt Bs) (map normt As) A oldAs n (nlift+1) prf1 prf2 val th = Thm (deriv_rule2 (if Envir.is_empty env then bicompose_proof else if Envir.above env smax then bicompose_proof o Proofterm.norm_proof' env else Proofterm.norm_proof' env oo bicompose_proof) rder' sder, {tags = [], maxidx = Envir.maxidx_of env, constraints = constraints', - shyps = Envir.insert_sorts env (Sorts.union shyps1 shyps2), + shyps = Envir.insert_sorts env (Sortset.merge (shyps1, shyps2)), hyps = union_hyps hyps1 hyps2, tpairs = ntpairs, prop = Logic.list_implies normp, cert = cert}) in Seq.cons th thq end handle COMPOSE => thq; val (rAs,B) = Logic.strip_prems(nsubgoal, [], rprop) handle TERM _ => raise THM("bicompose: rule", 0, [orule,state]); (*Modify assumptions, deleting n-th if n>0 for e-resolution*) fun newAs(As0, n, dpairs, tpairs) = let val (As1, rder') = if not lifted then (As0, rder) else let val rename = rename_bvars dpairs tpairs B As0 in (map (rename strip_apply) As0, deriv_rule1 (Proofterm.map_proof_terms (rename K) I) rder) end; in (map (if flatten then (Logic.flatten_params n) else I) As1, As1, rder', n) handle TERM _ => raise THM("bicompose: 1st premise", 0, [orule]) end; val BBi = if lifted then strip_assums2(B,Bi) else (B,Bi); val dpairs = BBi :: (rtpairs@stpairs); (*elim-resolution: try each assumption in turn*) fun eres _ [] = raise THM ("bicompose: no premises", 0, [orule, state]) | eres env (A1 :: As) = let val A = SOME A1; val (close, asms, concl) = Logic.assum_problems (nlift + 1, A1); val concl' = close concl; fun tryasms [] _ = Seq.empty | tryasms (asm :: rest) n = if Term.could_unify (asm, concl) then let val asm' = close asm in (case Seq.pull (Unify.unifiers (context, env, (asm', concl') :: dpairs)) of NONE => tryasms rest (n + 1) | cell as SOME ((_, tpairs), _) => Seq.it_right (addth A (newAs (As, n, [BBi, (concl', asm')], tpairs))) (Seq.make (fn () => cell), Seq.make (fn () => Seq.pull (tryasms rest (n + 1))))) end else tryasms rest (n + 1); in tryasms asms 1 end; (*ordinary resolution*) fun res env = (case Seq.pull (Unify.unifiers (context, env, dpairs)) of NONE => Seq.empty | cell as SOME ((_, tpairs), _) => Seq.it_right (addth NONE (newAs (rev rAs, 0, [BBi], tpairs))) (Seq.make (fn () => cell), Seq.empty)); val env0 = Envir.empty (Int.max (rmax, smax)); in (case if incremented then SOME env0 else unify_var_types context (state, orule) env0 of NONE => Seq.empty | SOME env => if eres_flg then eres env (rev rAs) else res env) end; end; fun bicompose opt_ctxt flags arg i state = bicompose_aux opt_ctxt flags (state, dest_state (state,i), false) arg; (*Quick test whether rule is resolvable with the subgoal with hyps Hs and conclusion B. If eres_flg then checks 1st premise of rule also*) fun could_bires (Hs, B, eres_flg, rule) = let fun could_reshyp (A1::_) = exists (fn H => Term.could_unify (A1, H)) Hs | could_reshyp [] = false; (*no premise -- illegal*) in Term.could_unify(concl_of rule, B) andalso (not eres_flg orelse could_reshyp (prems_of rule)) end; (*Bi-resolution of a state with a list of (flag,rule) pairs. Puts the rule above: rule/state. Renames vars in the rules. *) fun biresolution opt_ctxt match brules i state = let val (stpairs, Bs, Bi, C) = dest_state(state,i); val lift = lift_rule (cprem_of state i); val B = Logic.strip_assums_concl Bi; val Hs = Logic.strip_assums_hyp Bi; val compose = bicompose_aux opt_ctxt {flatten = true, match = match, incremented = true} (state, (stpairs, Bs, Bi, C), true); fun res [] = Seq.empty | res ((eres_flg, rule)::brules) = if Config.get_generic (make_context [state] opt_ctxt (cert_of state)) Pattern.unify_trace_failure orelse could_bires (Hs, B, eres_flg, rule) then Seq.make (*delay processing remainder till needed*) (fn()=> SOME(compose (eres_flg, lift rule, nprems_of rule), res brules)) else res brules in Seq.flat (res brules) end; (*Resolution: exactly one resolvent must be produced*) fun tha RSN (i, thb) = (case Seq.chop 2 (biresolution NONE false [(false, tha)] i thb) of ([th], _) => solve_constraints th | ([], _) => raise THM ("RSN: no unifiers", i, [tha, thb]) | _ => raise THM ("RSN: multiple unifiers", i, [tha, thb])); (*Resolution: P \ Q, Q \ R gives P \ R*) fun tha RS thb = tha RSN (1,thb); (**** Type classes ****) fun standard_tvars thm = let val thy = theory_of_thm thm; val tvars = build_rev (Term.add_tvars (prop_of thm)); val names = Name.invent Name.context Name.aT (length tvars); val tinst = map2 (fn (ai, S) => fn b => ((ai, S), global_ctyp_of thy (TVar ((b, 0), S)))) tvars names; in instantiate (TVars.make tinst, Vars.empty) thm end (* class relations *) val is_classrel = Symreltab.defined o get_classrels; fun complete_classrels thy = let fun complete (c, (_, (all_preds, all_succs))) (finished1, thy1) = let fun compl c1 c2 (finished2, thy2) = if is_classrel thy2 (c1, c2) then (finished2, thy2) else (false, thy2 |> (map_classrels o Symreltab.update) ((c1, c2), (the_classrel thy2 (c1, c) RS the_classrel thy2 (c, c2)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy2) |> trim_context)); val proven = is_classrel thy1; val preds = Graph.Keys.fold (fn c1 => proven (c1, c) ? cons c1) all_preds []; val succs = Graph.Keys.fold (fn c2 => proven (c, c2) ? cons c2) all_succs []; in fold_product compl preds succs (finished1, thy1) end; in (case Graph.fold complete (Sorts.classes_of (Sign.classes_of thy)) (true, thy) of (true, _) => NONE | (_, thy') => SOME thy') end; (* type arities *) fun thynames_of_arity thy (a, c) = build (get_arities thy |> Aritytab.fold (fn ((a', _, c'), (_, name, ser)) => (a = a' andalso c = c') ? cons (name, ser))) |> sort (int_ord o apply2 #2) |> map #1; fun insert_arity_completions thy ((t, Ss, c), (th, thy_name, ser)) (finished, arities) = let val completions = Sign.super_classes thy c |> map_filter (fn c1 => if Aritytab.defined arities (t, Ss, c1) then NONE else let val th1 = (th RS the_classrel thy (c, c1)) |> standard_tvars |> close_derivation \<^here> |> tap (expose_proof thy) |> trim_context; in SOME ((t, Ss, c1), (th1, thy_name, ser)) end); val finished' = finished andalso null completions; val arities' = fold Aritytab.update completions arities; in (finished', arities') end; fun complete_arities thy = let val arities = get_arities thy; val (finished, arities') = Aritytab.fold (insert_arity_completions thy) arities (true, get_arities thy); in if finished then NONE else SOME (map_arities (K arities') thy) end; val _ = Theory.setup (Theory.at_begin complete_classrels #> Theory.at_begin complete_arities); (* primitive rules *) fun add_classrel raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (c1, c2) = Logic.dest_classrel prop; in thy |> Sign.primitive_classrel (c1, c2) |> map_classrels (Symreltab.update ((c1, c2), th')) |> perhaps complete_classrels |> perhaps complete_arities end; fun add_arity raw_th thy = let val th = strip_shyps (transfer thy raw_th); val th' = th |> unconstrainT |> tap (expose_proof thy) |> trim_context; val prop = plain_prop_of th; val (t, Ss, c) = Logic.dest_arity prop; val ar = ((t, Ss, c), (th', Context.theory_name thy, serial ())); in thy |> Sign.primitive_arity (t, Ss, [c]) |> map_arities (Aritytab.update ar #> curry (insert_arity_completions thy ar) true #> #2) end; end; structure Basic_Thm: BASIC_THM = Thm; open Basic_Thm;