diff --git a/src/Doc/Implementation/Logic.thy b/src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy +++ b/src/Doc/Implementation/Logic.thy @@ -1,1259 +1,1259 @@ (*:maxLineLen=78:*) theory Logic imports Base begin chapter \Primitive logic \label{ch:logic}\ text \ The logical foundations of Isabelle/Isar are that of the Pure logic, which has been introduced as a Natural Deduction framework in @{cite paulson700}. This is essentially the same logic as ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, although there are some key differences in the specific treatment of simple types in Isabelle/Pure. Following type-theoretic parlance, the Pure logic consists of three levels of \\\-calculus with corresponding arrows, \\\ for syntactic function space (terms depending on terms), \\\ for universal quantification (proofs depending on terms), and \\\ for implication (proofs depending on proofs). Derivations are relative to a logical theory, which declares type constructors, constants, and axioms. Theory declarations support schematic polymorphism, which is strictly speaking outside the logic.\<^footnote>\This is the deeper logical reason, why the theory context \\\ is separate from the proof context \\\ of the core calculus: type constructors, term constants, and facts (proof constants) may involve arbitrary type schemes, but the type of a locally fixed term parameter is also fixed!\ \ section \Types \label{sec:types}\ text \ The language of types is an uninterpreted order-sorted first-order algebra; types are qualified by ordered type classes. \<^medskip> A \<^emph>\type class\ is an abstract syntactic entity declared in the theory context. The \<^emph>\subclass relation\ \c\<^sub>1 \ c\<^sub>2\ is specified by stating an acyclic generating relation; the transitive closure is maintained internally. The resulting relation is an ordering: reflexive, transitive, and antisymmetric. A \<^emph>\sort\ is a list of type classes written as \s = {c\<^sub>1, \, c\<^sub>m}\, it represents symbolic intersection. Notationally, the curly braces are omitted for singleton intersections, i.e.\ any class \c\ may be read as a sort \{c}\. The ordering on type classes is extended to sorts according to the meaning of intersections: \{c\<^sub>1, \ c\<^sub>m} \ {d\<^sub>1, \, d\<^sub>n}\ iff \\j. \i. c\<^sub>i \ d\<^sub>j\. The empty intersection \{}\ refers to the universal sort, which is the largest element wrt.\ the sort order. Thus \{}\ represents the ``full sort'', not the empty one! The intersection of all (finitely many) classes declared in the current theory is the least element wrt.\ the sort ordering. \<^medskip> A \<^emph>\fixed type variable\ is a pair of a basic name (starting with a \'\ character) and a sort constraint, e.g.\ \('a, s)\ which is usually printed as \\\<^sub>s\. A \<^emph>\schematic type variable\ is a pair of an indexname and a sort constraint, e.g.\ \(('a, 0), s)\ which is usually printed as \?\\<^sub>s\. Note that \<^emph>\all\ syntactic components contribute to the identity of type variables: basic name, index, and sort constraint. The core logic handles type variables with the same name but different sorts as different, although the type-inference layer (which is outside the core) rejects anything like that. A \<^emph>\type constructor\ \\\ is a \k\-ary operator on types declared in the theory. Type constructor application is written postfix as \(\\<^sub>1, \, \\<^sub>k)\\. For \k = 0\ the argument tuple is omitted, e.g.\ \prop\ instead of \()prop\. For \k = 1\ the parentheses are omitted, e.g.\ \\ list\ instead of \(\)list\. Further notation is provided for specific constructors, notably the right-associative infix \\ \ \\ instead of \(\, \)fun\. The logical category \<^emph>\type\ is defined inductively over type variables and type constructors as follows: \\ = \\<^sub>s | ?\\<^sub>s | (\\<^sub>1, \, \\<^sub>k)\\. A \<^emph>\type abbreviation\ is a syntactic definition \(\<^vec>\)\ = \\ of an arbitrary type expression \\\ over variables \\<^vec>\\. Type abbreviations appear as type constructors in the syntax, but are expanded before entering the logical core. A \<^emph>\type arity\ declares the image behavior of a type constructor wrt.\ the algebra of sorts: \\ :: (s\<^sub>1, \, s\<^sub>k)s\ means that \(\\<^sub>1, \, \\<^sub>k)\\ is of sort \s\ if every argument type \\\<^sub>i\ is of sort \s\<^sub>i\. Arity declarations are implicitly completed, i.e.\ \\ :: (\<^vec>s)c\ entails \\ :: (\<^vec>s)c'\ for any \c' \ c\. \<^medskip> The sort algebra is always maintained as \<^emph>\coregular\, which means that type arities are consistent with the subclass relation: for any type constructor \\\, and classes \c\<^sub>1 \ c\<^sub>2\, and arities \\ :: (\<^vec>s\<^sub>1)c\<^sub>1\ and \\ :: (\<^vec>s\<^sub>2)c\<^sub>2\ holds \\<^vec>s\<^sub>1 \ \<^vec>s\<^sub>2\ component-wise. The key property of a coregular order-sorted algebra is that sort constraints can be solved in a most general fashion: for each type constructor \\\ and sort \s\ there is a most general vector of argument sorts \(s\<^sub>1, \, s\<^sub>k)\ such that a type scheme \(\\<^bsub>s\<^sub>1\<^esub>, \, \\<^bsub>s\<^sub>k\<^esub>)\\ is of sort \s\. Consequently, type unification has most general solutions (modulo equivalence of sorts), so type-inference produces primary types as expected @{cite "nipkow-prehofer"}. \ text %mlref \ \begin{mldecls} @{define_ML_type class = string} \\ @{define_ML_type sort = "class list"} \\ @{define_ML_type arity = "string * sort list * sort"} \\ @{define_ML_type typ} \\ @{define_ML Term.map_atyps: "(typ -> typ) -> typ -> typ"} \\ @{define_ML Term.fold_atyps: "(typ -> 'a -> 'a) -> typ -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML Sign.subsort: "theory -> sort * sort -> bool"} \\ @{define_ML Sign.of_sort: "theory -> typ * sort -> bool"} \\ @{define_ML Sign.add_type: "Proof.context -> binding * int * mixfix -> theory -> theory"} \\ @{define_ML Sign.add_type_abbrev: "Proof.context -> binding * string list * typ -> theory -> theory"} \\ @{define_ML Sign.primitive_class: "binding * class list -> theory -> theory"} \\ @{define_ML Sign.primitive_classrel: "class * class -> theory -> theory"} \\ @{define_ML Sign.primitive_arity: "arity -> theory -> theory"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\class\ represents type classes. \<^descr> Type \<^ML_type>\sort\ represents sorts, i.e.\ finite intersections of classes. The empty list \<^ML>\[]: sort\ refers to the empty class intersection, i.e.\ the ``full sort''. \<^descr> Type \<^ML_type>\arity\ represents type arities. A triple \(\, \<^vec>s, s) : arity\ represents \\ :: (\<^vec>s)s\ as described above. \<^descr> Type \<^ML_type>\typ\ represents types; this is a datatype with constructors \<^ML>\TFree\, \<^ML>\TVar\, \<^ML>\Type\. \<^descr> \<^ML>\Term.map_atyps\~\f \\ applies the mapping \f\ to all atomic types (\<^ML>\TFree\, \<^ML>\TVar\) occurring in \\\. \<^descr> \<^ML>\Term.fold_atyps\~\f \\ iterates the operation \f\ over all occurrences of atomic types (\<^ML>\TFree\, \<^ML>\TVar\) in \\\; the type structure is traversed from left to right. \<^descr> \<^ML>\Sign.subsort\~\thy (s\<^sub>1, s\<^sub>2)\ tests the subsort relation \s\<^sub>1 \ s\<^sub>2\. \<^descr> \<^ML>\Sign.of_sort\~\thy (\, s)\ tests whether type \\\ is of sort \s\. \<^descr> \<^ML>\Sign.add_type\~\ctxt (\, k, mx)\ declares a new type constructors \\\ with \k\ arguments and optional mixfix syntax. \<^descr> \<^ML>\Sign.add_type_abbrev\~\ctxt (\, \<^vec>\, \)\ defines a new type abbreviation \(\<^vec>\)\ = \\. \<^descr> \<^ML>\Sign.primitive_class\~\(c, [c\<^sub>1, \, c\<^sub>n])\ declares a new class \c\, together with class relations \c \ c\<^sub>i\, for \i = 1, \, n\. \<^descr> \<^ML>\Sign.primitive_classrel\~\(c\<^sub>1, c\<^sub>2)\ declares the class relation \c\<^sub>1 \ c\<^sub>2\. \<^descr> \<^ML>\Sign.primitive_arity\~\(\, \<^vec>s, s)\ declares the arity \\ :: (\<^vec>s)s\. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "class"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "sort"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "type_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "nonterminal"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "typ"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation class} embedded ; @@{ML_antiquotation sort} sort ; (@@{ML_antiquotation type_name} | @@{ML_antiquotation type_abbrev} | @@{ML_antiquotation nonterminal}) embedded ; @@{ML_antiquotation typ} type \ \<^descr> \@{class c}\ inlines the internalized class \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{sort s}\ inlines the internalized sort \s\ --- as \<^ML_type>\string list\ literal. \<^descr> \@{type_name c}\ inlines the internalized type constructor \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{type_abbrev c}\ inlines the internalized type abbreviation \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{nonterminal c}\ inlines the internalized syntactic type~/ grammar nonterminal \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{typ \}\ inlines the internalized type \\\ --- as constructor term for datatype \<^ML_type>\typ\. \ section \Terms \label{sec:terms}\ text \ The language of terms is that of simply-typed \\\-calculus with de-Bruijn indices for bound variables (cf.\ @{cite debruijn72} or @{cite "paulson-ml2"}), with the types being determined by the corresponding binders. In contrast, free variables and constants have an explicit name and type in each occurrence. \<^medskip> A \<^emph>\bound variable\ is a natural number \b\, which accounts for the number of intermediate binders between the variable occurrence in the body and its binding position. For example, the de-Bruijn term \\\<^bsub>bool\<^esub>. \\<^bsub>bool\<^esub>. 1 \ 0\ would correspond to \\x\<^bsub>bool\<^esub>. \y\<^bsub>bool\<^esub>. x \ y\ in a named representation. Note that a bound variable may be represented by different de-Bruijn indices at different occurrences, depending on the nesting of abstractions. A \<^emph>\loose variable\ is a bound variable that is outside the scope of local binders. The types (and names) for loose variables can be managed as a separate context, that is maintained as a stack of hypothetical binders. The core logic operates on closed terms, without any loose variables. A \<^emph>\fixed variable\ is a pair of a basic name and a type, e.g.\ \(x, \)\ which is usually printed \x\<^sub>\\ here. A \<^emph>\schematic variable\ is a pair of an indexname and a type, e.g.\ \((x, 0), \)\ which is likewise printed as \?x\<^sub>\\. \<^medskip> A \<^emph>\constant\ is a pair of a basic name and a type, e.g.\ \(c, \)\ which is usually printed as \c\<^sub>\\ here. Constants are declared in the context as polymorphic families \c :: \\, meaning that all substitution instances \c\<^sub>\\ for \\ = \\\ are valid. The vector of \<^emph>\type arguments\ of constant \c\<^sub>\\ wrt.\ the declaration \c :: \\ is defined as the codomain of the matcher \\ = {?\\<^sub>1 \ \\<^sub>1, \, ?\\<^sub>n \ \\<^sub>n}\ presented in canonical order \(\\<^sub>1, \, \\<^sub>n)\, corresponding to the left-to-right occurrences of the \\\<^sub>i\ in \\\. Within a given theory context, there is a one-to-one correspondence between any constant \c\<^sub>\\ and the application \c(\\<^sub>1, \, \\<^sub>n)\ of its type arguments. For example, with \plus :: \ \ \ \ \\, the instance \plus\<^bsub>nat \ nat \ nat\<^esub>\ corresponds to \plus(nat)\. Constant declarations \c :: \\ may contain sort constraints for type variables in \\\. These are observed by type-inference as expected, but \<^emph>\ignored\ by the core logic. This means the primitive logic is able to reason with instances of polymorphic constants that the user-level type-checker would reject due to violation of type class restrictions. \<^medskip> An \<^emph>\atomic term\ is either a variable or constant. The logical category \<^emph>\term\ is defined inductively over atomic terms, with abstraction and application as follows: \t = b | x\<^sub>\ | ?x\<^sub>\ | c\<^sub>\ | \\<^sub>\. t | t\<^sub>1 t\<^sub>2\. Parsing and printing takes care of converting between an external representation with named bound variables. Subsequently, we shall use the latter notation instead of internal de-Bruijn representation. The inductive relation \t :: \\ assigns a (unique) type to a term according to the structure of atomic terms, abstractions, and applications: \[ \infer{\a\<^sub>\ :: \\}{} \qquad \infer{\(\x\<^sub>\. t) :: \ \ \\}{\t :: \\} \qquad \infer{\t u :: \\}{\t :: \ \ \\ & \u :: \\} \] A \<^emph>\well-typed term\ is a term that can be typed according to these rules. Typing information can be omitted: type-inference is able to reconstruct the most general type of a raw term, while assigning most general types to all of its variables and constants. Type-inference depends on a context of type constraints for fixed variables, and declarations for polymorphic constants. The identity of atomic terms consists both of the name and the type component. This means that different variables \x\<^bsub>\\<^sub>1\<^esub>\ and \x\<^bsub>\\<^sub>2\<^esub>\ may become the same after type instantiation. Type-inference rejects variables of the same name, but different types. In contrast, mixed instances of polymorphic constants occur routinely. \<^medskip> The \<^emph>\hidden polymorphism\ of a term \t :: \\ is the set of type variables occurring in \t\, but not in its type \\\. This means that the term implicitly depends on type arguments that are not accounted in the result type, i.e.\ there are different type instances \t\ :: \\ and \t\' :: \\ with the same type. This slightly pathological situation notoriously demands additional care. \<^medskip> A \<^emph>\term abbreviation\ is a syntactic definition \c\<^sub>\ \ t\ of a closed term \t\ of type \\\, without any hidden polymorphism. A term abbreviation looks like a constant in the syntax, but is expanded before entering the logical core. Abbreviations are usually reverted when printing terms, using \t \ c\<^sub>\\ as rules for higher-order rewriting. \<^medskip> Canonical operations on \\\-terms include \\\\\-conversion: \\\-conversion refers to capture-free renaming of bound variables; \\\-conversion contracts an abstraction applied to an argument term, substituting the argument in the body: \(\x. b)a\ becomes \b[a/x]\; \\\-conversion contracts vacuous application-abstraction: \\x. f x\ becomes \f\, provided that the bound variable does not occur in \f\. Terms are normally treated modulo \\\-conversion, which is implicit in the de-Bruijn representation. Names for bound variables in abstractions are maintained separately as (meaningless) comments, mostly for parsing and printing. Full \\\\\-conversion is commonplace in various standard operations (\secref{sec:obj-rules}) that are based on higher-order unification and matching. \ text %mlref \ \begin{mldecls} @{define_ML_type term} \\ @{define_ML_infix "aconv": "term * term -> bool"} \\ @{define_ML Term.map_types: "(typ -> typ) -> term -> term"} \\ @{define_ML Term.fold_types: "(typ -> 'a -> 'a) -> term -> 'a -> 'a"} \\ @{define_ML Term.map_aterms: "(term -> term) -> term -> term"} \\ @{define_ML Term.fold_aterms: "(term -> 'a -> 'a) -> term -> 'a -> 'a"} \\ \end{mldecls} \begin{mldecls} @{define_ML fastype_of: "term -> typ"} \\ @{define_ML lambda: "term -> term -> term"} \\ @{define_ML betapply: "term * term -> term"} \\ @{define_ML incr_boundvars: "int -> term -> term"} \\ @{define_ML Sign.declare_const: "Proof.context -> (binding * typ) * mixfix -> theory -> term * theory"} \\ @{define_ML Sign.add_abbrev: "string -> binding * term -> theory -> (term * term) * theory"} \\ @{define_ML Sign.const_typargs: "theory -> string * typ -> typ list"} \\ @{define_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\term\ represents de-Bruijn terms, with comments in abstractions, and explicitly named free variables and constants; this is a datatype with constructors @{define_ML Bound}, @{define_ML Free}, @{define_ML Var}, @{define_ML Const}, @{define_ML Abs}, @{define_ML_infix "$"}. \<^descr> \t\~\<^ML_text>\aconv\~\u\ checks \\\-equivalence of two terms. This is the basic equality relation on type \<^ML_type>\term\; raw datatype equality should only be used for operations related to parsing or printing! \<^descr> \<^ML>\Term.map_types\~\f t\ applies the mapping \f\ to all types occurring in \t\. \<^descr> \<^ML>\Term.fold_types\~\f t\ iterates the operation \f\ over all occurrences of types in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\Term.map_aterms\~\f t\ applies the mapping \f\ to all atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) occurring in \t\. \<^descr> \<^ML>\Term.fold_aterms\~\f t\ iterates the operation \f\ over all occurrences of atomic terms (\<^ML>\Bound\, \<^ML>\Free\, \<^ML>\Var\, \<^ML>\Const\) in \t\; the term structure is traversed from left to right. \<^descr> \<^ML>\fastype_of\~\t\ determines the type of a well-typed term. This operation is relatively slow, despite the omission of any sanity checks. \<^descr> \<^ML>\lambda\~\a b\ produces an abstraction \\a. b\, where occurrences of the atomic term \a\ in the body \b\ are replaced by bound variables. \<^descr> \<^ML>\betapply\~\(t, u)\ produces an application \t u\, with topmost \\\-conversion if \t\ is an abstraction. \<^descr> \<^ML>\incr_boundvars\~\j\ increments a term's dangling bound variables by the offset \j\. This is required when moving a subterm into a context where it is enclosed by a different number of abstractions. Bound variables with a matching abstraction are unaffected. \<^descr> \<^ML>\Sign.declare_const\~\ctxt ((c, \), mx)\ declares a new constant \c :: \\ with optional mixfix syntax. \<^descr> \<^ML>\Sign.add_abbrev\~\print_mode (c, t)\ introduces a new term abbreviation \c \ t\. \<^descr> \<^ML>\Sign.const_typargs\~\thy (c, \)\ and \<^ML>\Sign.const_instance\~\thy (c, [\\<^sub>1, \, \\<^sub>n])\ convert between two representations of polymorphic constants: full type instance vs.\ compact type arguments form. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "const_name"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const_abbrev"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "const"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "term"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "prop"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ (@@{ML_antiquotation const_name} | @@{ML_antiquotation const_abbrev}) embedded ; @@{ML_antiquotation const} ('(' (type + ',') ')')? ; @@{ML_antiquotation term} term ; @@{ML_antiquotation prop} prop \ \<^descr> \@{const_name c}\ inlines the internalized logical constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const_abbrev c}\ inlines the internalized abbreviated constant name \c\ --- as \<^ML_type>\string\ literal. \<^descr> \@{const c(\<^vec>\)}\ inlines the internalized constant \c\ with precise type instantiation in the sense of \<^ML>\Sign.const_instance\ --- as \<^ML>\Const\ constructor term for datatype \<^ML_type>\term\. \<^descr> \@{term t}\ inlines the internalized term \t\ --- as constructor term for datatype \<^ML_type>\term\. \<^descr> \@{prop \}\ inlines the internalized proposition \\\ --- as constructor term for datatype \<^ML_type>\term\. \ section \Theorems \label{sec:thms}\ text \ A \<^emph>\proposition\ is a well-typed term of type \prop\, a \<^emph>\theorem\ is a proven proposition (depending on a context of hypotheses and the background theory). Primitive inferences include plain Natural Deduction rules for the primary connectives \\\ and \\\ of the framework. There is also a builtin notion of equality/equivalence \\\. \ subsection \Primitive connectives and rules \label{sec:prim-rules}\ text \ The theory \Pure\ contains constant declarations for the primitive connectives \\\, \\\, and \\\ of the logical framework, see \figref{fig:pure-connectives}. The derivability judgment \A\<^sub>1, \, A\<^sub>n \ B\ is defined inductively by the primitive inferences given in \figref{fig:prim-rules}, with the global restriction that the hypotheses must \<^emph>\not\ contain any schematic variables. The builtin equality is conceptually axiomatized as shown in \figref{fig:pure-equality}, although the implementation works directly with derived inferences. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \all :: (\ \ prop) \ prop\ & universal quantification (binder \\\) \\ \\ :: prop \ prop \ prop\ & implication (right associative infix) \\ \\ :: \ \ \ \ prop\ & equality relation (infix) \\ \end{tabular} \caption{Primitive connectives of Pure}\label{fig:pure-connectives} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \[ \infer[\(axiom)\]{\\ A\}{\A \ \\} \qquad \infer[\(assume)\]{\A \ A\}{} \] \[ \infer[\(\\intro)\]{\\ \ \x. B[x]\}{\\ \ B[x]\ & \x \ \\} \qquad \infer[\(\\elim)\]{\\ \ B[a]\}{\\ \ \x. B[x]\} \] \[ \infer[\(\\intro)\]{\\ - A \ A \ B\}{\\ \ B\} \qquad \infer[\(\\elim)\]{\\\<^sub>1 \ \\<^sub>2 \ B\}{\\\<^sub>1 \ A \ B\ & \\\<^sub>2 \ A\} \] \caption{Primitive inferences of Pure}\label{fig:prim-rules} \end{center} \end{figure} \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \\ (\x. b[x]) a \ b[a]\ & \\\-conversion \\ \\ x \ x\ & reflexivity \\ \\ x \ y \ P x \ P y\ & substitution \\ \\ (\x. f x \ g x) \ f \ g\ & extensionality \\ \\ (A \ B) \ (B \ A) \ A \ B\ & logical equivalence \\ \end{tabular} \caption{Conceptual axiomatization of Pure equality}\label{fig:pure-equality} \end{center} \end{figure} The introduction and elimination rules for \\\ and \\\ are analogous to formation of dependently typed \\\-terms representing the underlying proof objects. Proof terms are irrelevant in the Pure logic, though; they cannot occur within propositions. The system provides a runtime option to record explicit proof terms for primitive inferences, see also \secref{sec:proof-terms}. Thus all three levels of \\\-calculus become explicit: \\\ for terms, and \\/\\ for proofs (cf.\ @{cite "Berghofer-Nipkow:2000:TPHOL"}). Observe that locally fixed parameters (as in \\\intro\) need not be recorded in the hypotheses, because the simple syntactic types of Pure are always inhabitable. ``Assumptions'' \x :: \\ for type-membership are only present as long as some \x\<^sub>\\ occurs in the statement body.\<^footnote>\This is the key difference to ``\\HOL\'' in the PTS framework @{cite "Barendregt-Geuvers:2001"}, where hypotheses \x : A\ are treated uniformly for propositions and types.\ \<^medskip> The axiomatization of a theory is implicitly closed by forming all instances of type and term variables: \\ A\\ holds for any substitution instance of an axiom \\ A\. By pushing substitutions through derivations inductively, we also get admissible \generalize\ and \instantiate\ rules as shown in \figref{fig:subst-rules}. \begin{figure}[htb] \begin{center} \[ \infer{\\ \ B[?\]\}{\\ \ B[\]\ & \\ \ \\} \quad \infer[\quad\(generalize)\]{\\ \ B[?x]\}{\\ \ B[x]\ & \x \ \\} \] \[ \infer{\\ \ B[\]\}{\\ \ B[?\]\} \quad \infer[\quad\(instantiate)\]{\\ \ B[t]\}{\\ \ B[?x]\} \] \caption{Admissible substitution rules}\label{fig:subst-rules} \end{center} \end{figure} Note that \instantiate\ does not require an explicit side-condition, because \\\ may never contain schematic variables. In principle, variables could be substituted in hypotheses as well, but this would disrupt the monotonicity of reasoning: deriving \\\ \ B\\ from \\ \ B\ is correct, but \\\ \ \\ does not necessarily hold: the result belongs to a different proof context. \<^medskip> An \<^emph>\oracle\ is a function that produces axioms on the fly. Logically, this is an instance of the \axiom\ rule (\figref{fig:prim-rules}), but there is an operational difference. The inference kernel records oracle invocations within derivations of theorems by a unique tag. This also includes implicit type-class reasoning via the order-sorted algebra of class relations and type arities (see also @{command_ref instantiation} and @{command_ref instance}). Axiomatizations should be limited to the bare minimum, typically as part of the initial logical basis of an object-logic formalization. Later on, theories are usually developed in a strictly definitional fashion, by stating only certain equalities over new constants. A \<^emph>\simple definition\ consists of a constant declaration \c :: \\ together with an axiom \\ c \ t\, where \t :: \\ is a closed term without any hidden polymorphism. The RHS may depend on further defined constants, but not \c\ itself. Definitions of functions may be presented as \c \<^vec>x \ t\ instead of the puristic \c \ \\<^vec>x. t\. An \<^emph>\overloaded definition\ consists of a collection of axioms for the same constant, with zero or one equations \c((\<^vec>\)\) \ t\ for each type constructor \\\ (for distinct variables \\<^vec>\\). The RHS may mention previously defined constants as above, or arbitrary constants \d(\\<^sub>i)\ for some \\\<^sub>i\ projected from \\<^vec>\\. Thus overloaded definitions essentially work by primitive recursion over the syntactic structure of a single type argument. See also @{cite \\S4.3\ "Haftmann-Wenzel:2006:classes"}. \ text %mlref \ \begin{mldecls} @{define_ML Logic.all: "term -> term -> term"} \\ @{define_ML Logic.mk_implies: "term * term -> term"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type ctyp} \\ @{define_ML_type cterm} \\ @{define_ML Thm.ctyp_of: "Proof.context -> typ -> ctyp"} \\ @{define_ML Thm.cterm_of: "Proof.context -> term -> cterm"} \\ @{define_ML Thm.apply: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.lambda: "cterm -> cterm -> cterm"} \\ @{define_ML Thm.all: "Proof.context -> cterm -> cterm -> cterm"} \\ @{define_ML Drule.mk_implies: "cterm * cterm -> cterm"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type thm} \\ @{define_ML Thm.transfer: "theory -> thm -> thm"} \\ @{define_ML Thm.assume: "cterm -> thm"} \\ @{define_ML Thm.forall_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.forall_elim: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_intr: "cterm -> thm -> thm"} \\ @{define_ML Thm.implies_elim: "thm -> thm -> thm"} \\ @{define_ML Thm.generalize: "string list * string list -> int -> thm -> thm"} \\ @{define_ML Thm.instantiate: "((indexname * sort) * ctyp) list * ((indexname * typ) * cterm) list -> thm -> thm"} \\ @{define_ML Thm.add_axiom: "Proof.context -> binding * term -> theory -> (string * thm) * theory"} \\ @{define_ML Thm.add_oracle: "binding * ('a -> cterm) -> theory -> (string * ('a -> thm)) * theory"} \\ @{define_ML Thm.add_def: "Defs.context -> bool -> bool -> binding * term -> theory -> (string * thm) * theory"} \\ \end{mldecls} \begin{mldecls} @{define_ML Theory.add_deps: "Defs.context -> string -> Defs.entry -> Defs.entry list -> theory -> theory"} \\ @{define_ML Thm_Deps.all_oracles: "thm list -> Proofterm.oracle list"} \\ \end{mldecls} \<^descr> \<^ML>\Logic.all\~\a B\ produces a Pure quantification \\a. B\, where occurrences of the atomic term \a\ in the body proposition \B\ are replaced by bound variables. (See also \<^ML>\lambda\ on terms.) \<^descr> \<^ML>\Logic.mk_implies\~\(A, B)\ produces a Pure implication \A \ B\. \<^descr> Types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ represent certified types and terms, respectively. These are abstract datatypes that guarantee that its values have passed the full well-formedness (and well-typedness) checks, relative to the declarations of type constructors, constants etc.\ in the background theory. The abstract types \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are part of the same inference kernel that is mainly responsible for \<^ML_type>\thm\. Thus syntactic operations on \<^ML_type>\ctyp\ and \<^ML_type>\cterm\ are located in the \<^ML_structure>\Thm\ module, even though theorems are not yet involved at that stage. \<^descr> \<^ML>\Thm.ctyp_of\~\ctxt \\ and \<^ML>\Thm.cterm_of\~\ctxt t\ explicitly check types and terms, respectively. This also involves some basic normalizations, such expansion of type and term abbreviations from the underlying theory context. Full re-certification is relatively slow and should be avoided in tight reasoning loops. \<^descr> \<^ML>\Thm.apply\, \<^ML>\Thm.lambda\, \<^ML>\Thm.all\, \<^ML>\Drule.mk_implies\ etc.\ compose certified terms (or propositions) incrementally. This is - equivalent to \<^ML>\Thm.cterm_of\ after unchecked \<^ML_op>\$\, \<^ML>\lambda\, + 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 basic names. \<^descr> \<^ML>\Thm.instantiate\~\(\<^vec>\\<^sub>s, \<^vec>x\<^sub>\)\ corresponds to the \instantiate\ rules of \figref{fig:subst-rules}. Type variables are substituted before term variables. Note that the types in \\<^vec>x\<^sub>\\ refer to the instantiated versions. \<^descr> \<^ML>\Thm.add_axiom\~\ctxt (name, A)\ declares an arbitrary proposition as axiom, and retrieves it as a theorem from the resulting theory, cf.\ \axiom\ in \figref{fig:prim-rules}. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Thm.add_oracle\~\(binding, oracle)\ produces a named oracle rule, essentially generating arbitrary axioms on the fly, cf.\ \axiom\ in \figref{fig:prim-rules}. \<^descr> \<^ML>\Thm.add_def\~\ctxt unchecked overloaded (name, c \<^vec>x \ t)\ states a definitional axiom for an existing constant \c\. Dependencies are recorded via \<^ML>\Theory.add_deps\, unless the \unchecked\ option is set. Note that the low-level representation in the axiom table may differ slightly from the returned theorem. \<^descr> \<^ML>\Theory.add_deps\~\ctxt name c\<^sub>\ \<^vec>d\<^sub>\\ declares dependencies of a named specification for constant \c\<^sub>\\, relative to existing specifications for constants \\<^vec>d\<^sub>\\. This also works for type constructors. \<^descr> \<^ML>\Thm_Deps.all_oracles\~\thms\ returns all oracles used in the internal derivation of the given theorems; this covers the full graph of transitive dependencies. The result contains an authentic oracle name; if @{ML Proofterm.proofs} was at least @{ML 1} during the oracle inference, it also contains the position of the oracle invocation and its proposition. See also the command @{command_ref "thm_oracles"}. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "ctyp"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cterm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "cprop"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thm"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "thms"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "lemma"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "oracle_name"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation ctyp} typ ; @@{ML_antiquotation cterm} term ; @@{ML_antiquotation cprop} prop ; @@{ML_antiquotation thm} thm ; @@{ML_antiquotation thms} thms ; @@{ML_antiquotation lemma} ('(' @'open' ')')? ((prop +) + @'and') \ @'by' method method? ; @@{ML_antiquotation oracle_name} embedded \ \<^descr> \@{ctyp \}\ produces a certified type wrt.\ the current background theory --- as abstract value of type \<^ML_type>\ctyp\. \<^descr> \@{cterm t}\ and \@{cprop \}\ produce a certified term wrt.\ the current background theory --- as abstract value of type \<^ML_type>\cterm\. \<^descr> \@{thm a}\ produces a singleton fact --- as abstract value of type \<^ML_type>\thm\. \<^descr> \@{thms a}\ produces a general fact --- as abstract value of type \<^ML_type>\thm list\. \<^descr> \@{lemma \ by meth}\ produces a fact that is proven on the spot according to the minimal proof, which imitates a terminal Isar proof. The result is an abstract value of type \<^ML_type>\thm\ or \<^ML_type>\thm list\, depending on the number of propositions given here. The internal derivation object lacks a proper theorem name, but it is formally closed, unless the \(open)\ option is specified (this may impact performance of applications with proof terms). Since ML antiquotations are always evaluated at compile-time, there is no run-time overhead even for non-trivial proofs. Nonetheless, the justification is syntactically limited to a single @{command "by"} step. More complex Isar proofs should be done in regular theory source, before compiling the corresponding ML text that uses the result. \<^descr> \@{oracle_name a}\ inlines the internalized oracle name \a\ --- as \<^ML_type>\string\ literal. \ subsection \Auxiliary connectives \label{sec:logic-aux}\ text \ Theory \Pure\ provides a few auxiliary connectives that are defined on top of the primitive ones, see \figref{fig:pure-aux}. These special constants are useful in certain internal encodings, and are normally not directly exposed to the user. \begin{figure}[htb] \begin{center} \begin{tabular}{ll} \conjunction :: prop \ prop \ prop\ & (infix \&&&\) \\ \\ A &&& B \ (\C. (A \ B \ C) \ C)\ \\[1ex] \prop :: prop \ prop\ & (prefix \#\, suppressed) \\ \#A \ A\ \\[1ex] \term :: \ \ prop\ & (prefix \TERM\) \\ \term x \ (\A. A \ A)\ \\[1ex] \type :: \ itself\ & (prefix \TYPE\) \\ \(unspecified)\ \\ \end{tabular} \caption{Definitions of auxiliary connectives}\label{fig:pure-aux} \end{center} \end{figure} The introduction \A \ B \ A &&& B\, and eliminations (projections) \A &&& B \ A\ and \A &&& B \ B\ are available as derived rules. Conjunction allows to treat simultaneous assumptions and conclusions uniformly, e.g.\ consider \A \ B \ C &&& D\. In particular, the goal mechanism represents multiple claims as explicit conjunction internally, but this is refined (via backwards introduction) into separate sub-goals before the user commences the proof; the final result is projected into a list of theorems using eliminations (cf.\ \secref{sec:tactical-goals}). The \prop\ marker (\#\) makes arbitrarily complex propositions appear as atomic, without changing the meaning: \\ \ A\ and \\ \ #A\ are interchangeable. See \secref{sec:tactical-goals} for specific operations. The \term\ marker turns any well-typed term into a derivable proposition: \\ TERM t\ holds unconditionally. Although this is logically vacuous, it allows to treat terms and proofs uniformly, similar to a type-theoretic framework. The \TYPE\ constructor is the canonical representative of the unspecified type \\ itself\; it essentially injects the language of types into that of terms. There is specific notation \TYPE(\)\ for \TYPE\<^bsub>\ itself\<^esub>\. Although being devoid of any particular meaning, the term \TYPE(\)\ accounts for the type \\\ within the term language. In particular, \TYPE(\)\ may be used as formal argument in primitive definitions, in order to circumvent hidden polymorphism (cf.\ \secref{sec:terms}). For example, \c TYPE(\) \ A[\]\ defines \c :: \ itself \ prop\ in terms of a proposition \A\ that depends on an additional type argument, which is essentially a predicate on types. \ text %mlref \ \begin{mldecls} @{define_ML Conjunction.intr: "thm -> thm -> thm"} \\ @{define_ML Conjunction.elim: "thm -> thm * thm"} \\ @{define_ML Drule.mk_term: "cterm -> thm"} \\ @{define_ML Drule.dest_term: "thm -> cterm"} \\ @{define_ML Logic.mk_type: "typ -> term"} \\ @{define_ML Logic.dest_type: "term -> typ"} \\ \end{mldecls} \<^descr> \<^ML>\Conjunction.intr\ derives \A &&& B\ from \A\ and \B\. \<^descr> \<^ML>\Conjunction.elim\ derives \A\ and \B\ from \A &&& B\. \<^descr> \<^ML>\Drule.mk_term\ derives \TERM t\. \<^descr> \<^ML>\Drule.dest_term\ recovers term \t\ from \TERM t\. \<^descr> \<^ML>\Logic.mk_type\~\\\ produces the term \TYPE(\)\. \<^descr> \<^ML>\Logic.dest_type\~\TYPE(\)\ recovers the type \\\. \ subsection \Sort hypotheses\ text \ Type variables are decorated with sorts, as explained in \secref{sec:types}. This constrains type instantiation to certain ranges of types: variable \\\<^sub>s\ may only be assigned to types \\\ that belong to sort \s\. Within the logic, sort constraints act like implicit preconditions on the result \\\\<^sub>1 : s\<^sub>1\, \, \\\<^sub>n : s\<^sub>n\, \ \ \\ where the type variables \\\<^sub>1, \, \\<^sub>n\ cover the propositions \\\, \\\, as well as the proof of \\ \ \\. These \<^emph>\sort hypothesis\ of a theorem are passed monotonically through further derivations. They are redundant, as long as the statement of a theorem still contains the type variables that are accounted here. The logical significance of sort hypotheses is limited to the boundary case where type variables disappear from the proposition, e.g.\ \\\\<^sub>s : s\ \ \\. Since such dangling type variables can be renamed arbitrarily without changing the proposition \\\, the inference kernel maintains sort hypotheses in anonymous form \s \ \\. In most practical situations, such extra sort hypotheses may be stripped in a final bookkeeping step, e.g.\ at the end of a proof: they are typically left over from intermediate reasoning with type classes that can be satisfied by some concrete type \\\ of sort \s\ to replace the hypothetical type variable \\\<^sub>s\. \ text %mlref \ \begin{mldecls} @{define_ML Thm.extra_shyps: "thm -> sort list"} \\ @{define_ML Thm.strip_shyps: "thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Thm.extra_shyps\~\thm\ determines the extraneous sort hypotheses of the given theorem, i.e.\ the sorts that are not present within type variables of the statement. \<^descr> \<^ML>\Thm.strip_shyps\~\thm\ removes any extraneous sort hypotheses that can be witnessed from the type signature. \ text %mlex \ The following artificial example demonstrates the derivation of \<^prop>\False\ with a pending sort hypothesis involving a logically empty sort. \ class empty = assumes bad: "\(x::'a) y. x \ y" theorem (in empty) false: False using bad by blast ML_val \\<^assert> (Thm.extra_shyps @{thm false} = [\<^sort>\empty\])\ text \ Thanks to the inference kernel managing sort hypothesis according to their logical significance, this example is merely an instance of \<^emph>\ex falso quodlibet consequitur\ --- not a collapse of the logical framework! \ section \Object-level rules \label{sec:obj-rules}\ text \ The primitive inferences covered so far mostly serve foundational purposes. User-level reasoning usually works via object-level rules that are represented as theorems of Pure. Composition of rules involves \<^emph>\backchaining\, \<^emph>\higher-order unification\ modulo \\\\\-conversion of \\\-terms, and so-called \<^emph>\lifting\ of rules into a context of \\\ and \\\ connectives. Thus the full power of higher-order Natural Deduction in Isabelle/Pure becomes readily available. \ subsection \Hereditary Harrop Formulae\ text \ The idea of object-level rules is to model Natural Deduction inferences in the style of Gentzen @{cite "Gentzen:1935"}, but we allow arbitrary nesting similar to @{cite "Schroeder-Heister:1984"}. The most basic rule format is that of a \<^emph>\Horn Clause\: \[ \infer{\A\}{\A\<^sub>1\ & \\\ & \A\<^sub>n\} \] where \A, A\<^sub>1, \, A\<^sub>n\ are atomic propositions of the framework, usually of the form \Trueprop B\, where \B\ is a (compound) object-level statement. This object-level inference corresponds to an iterated implication in Pure like this: \[ \A\<^sub>1 \ \ A\<^sub>n \ A\ \] As an example consider conjunction introduction: \A \ B \ A \ B\. Any parameters occurring in such rule statements are conceptionally treated as arbitrary: \[ \\x\<^sub>1 \ x\<^sub>m. A\<^sub>1 x\<^sub>1 \ x\<^sub>m \ \ A\<^sub>n x\<^sub>1 \ x\<^sub>m \ A x\<^sub>1 \ x\<^sub>m\ \] Nesting of rules means that the positions of \A\<^sub>i\ may again hold compound rules, not just atomic propositions. Propositions of this format are called \<^emph>\Hereditary Harrop Formulae\ in the literature @{cite "Miller:1991"}. Here we give an inductive characterization as follows: \<^medskip> \begin{tabular}{ll} \\<^bold>x\ & set of variables \\ \\<^bold>A\ & set of atomic propositions \\ \\<^bold>H = \\<^bold>x\<^sup>*. \<^bold>H\<^sup>* \ \<^bold>A\ & set of Hereditary Harrop Formulas \\ \end{tabular} \<^medskip> Thus we essentially impose nesting levels on propositions formed from \\\ and \\\. At each level there is a prefix of parameters and compound premises, concluding an atomic proposition. Typical examples are \\\-introduction \(A \ B) \ A \ B\ or mathematical induction \P 0 \ (\n. P n \ P (Suc n)) \ P n\. Even deeper nesting occurs in well-founded induction \(\x. (\y. y \ x \ P y) \ P x) \ P x\, but this already marks the limit of rule complexity that is usually seen in practice. \<^medskip> Regular user-level inferences in Isabelle/Pure always maintain the following canonical form of results: \<^item> Normalization by \(A \ (\x. B x)) \ (\x. A \ B x)\, which is a theorem of Pure, means that quantifiers are pushed in front of implication at each level of nesting. The normal form is a Hereditary Harrop Formula. \<^item> The outermost prefix of parameters is represented via schematic variables: instead of \\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x\ we have \\<^vec>H ?\<^vec>x \ A ?\<^vec>x\. Note that this representation looses information about the order of parameters, and vacuous quantifiers vanish automatically. \ text %mlref \ \begin{mldecls} @{define_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Simplifier.norm_hhf\~\ctxt thm\ normalizes the given theorem according to the canonical form specified above. This is occasionally helpful to repair some low-level tools that do not handle Hereditary Harrop Formulae properly. \ subsection \Rule composition\ text \ The rule calculus of Isabelle/Pure provides two main inferences: @{inference resolution} (i.e.\ back-chaining of rules) and @{inference assumption} (i.e.\ closing a branch), both modulo higher-order unification. There are also combined variants, notably @{inference elim_resolution} and @{inference dest_resolution}. To understand the all-important @{inference resolution} principle, we first consider raw @{inference_def composition} (modulo higher-order unification with substitution \\\): \[ \infer[(@{inference_def composition})]{\\<^vec>A\ \ C\\} {\\<^vec>A \ B\ & \B' \ C\ & \B\ = B'\\} \] Here the conclusion of the first rule is unified with the premise of the second; the resulting rule instance inherits the premises of the first and conclusion of the second. Note that \C\ can again consist of iterated implications. We can also permute the premises of the second rule back-and-forth in order to compose with \B'\ in any position (subsequently we shall always refer to position 1 w.l.o.g.). In @{inference composition} the internal structure of the common part \B\ and \B'\ is not taken into account. For proper @{inference resolution} we require \B\ to be atomic, and explicitly observe the structure \\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x\ of the premise of the second rule. The idea is to adapt the first rule by ``lifting'' it into this context, by means of iterated application of the following inferences: \[ \infer[(@{inference_def imp_lift})]{\(\<^vec>H \ \<^vec>A) \ (\<^vec>H \ B)\}{\\<^vec>A \ B\} \] \[ \infer[(@{inference_def all_lift})]{\(\\<^vec>x. \<^vec>A (?\<^vec>a \<^vec>x)) \ (\\<^vec>x. B (?\<^vec>a \<^vec>x))\}{\\<^vec>A ?\<^vec>a \ B ?\<^vec>a\} \] By combining raw composition with lifting, we get full @{inference resolution} as follows: \[ \infer[(@{inference_def resolution})] {\(\\<^vec>x. \<^vec>H \<^vec>x \ \<^vec>A (?\<^vec>a \<^vec>x))\ \ C\\} {\begin{tabular}{l} \\<^vec>A ?\<^vec>a \ B ?\<^vec>a\ \\ \(\\<^vec>x. \<^vec>H \<^vec>x \ B' \<^vec>x) \ C\ \\ \(\\<^vec>x. B (?\<^vec>a \<^vec>x))\ = B'\\ \\ \end{tabular}} \] Continued resolution of rules allows to back-chain a problem towards more and sub-problems. Branches are closed either by resolving with a rule of 0 premises, or by producing a ``short-circuit'' within a solved situation (again modulo unification): \[ \infer[(@{inference_def assumption})]{\C\\} {\(\\<^vec>x. \<^vec>H \<^vec>x \ A \<^vec>x) \ C\ & \A\ = H\<^sub>i\\~~\mbox{(for some~\i\)}} \] %FIXME @{inference_def elim_resolution}, @{inference_def dest_resolution} \ text %mlref \ \begin{mldecls} @{define_ML_infix "RSN": "thm * (int * thm) -> thm"} \\ @{define_ML_infix "RS": "thm * thm -> thm"} \\ @{define_ML_infix "RLN": "thm list * (int * thm list) -> thm list"} \\ @{define_ML_infix "RL": "thm list * thm list -> thm list"} \\ @{define_ML_infix "MRS": "thm list * thm -> thm"} \\ @{define_ML_infix "OF": "thm * thm list -> thm"} \\ \end{mldecls} \<^descr> \rule\<^sub>1 RSN (i, rule\<^sub>2)\ resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, according to the @{inference resolution} principle explained above. Unless there is precisely one resolvent it raises exception \<^ML>\THM\. This corresponds to the rule attribute @{attribute THEN} in Isar source language. \<^descr> \rule\<^sub>1 RS rule\<^sub>2\ abbreviates \rule\<^sub>1 RSN (1, rule\<^sub>2)\. \<^descr> \rules\<^sub>1 RLN (i, rules\<^sub>2)\ joins lists of rules. For every \rule\<^sub>1\ in \rules\<^sub>1\ and \rule\<^sub>2\ in \rules\<^sub>2\, it resolves the conclusion of \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, accumulating multiple results in one big list. Note that such strict enumerations of higher-order unifications can be inefficient compared to the lazy variant seen in elementary tactics like \<^ML>\resolve_tac\. \<^descr> \rules\<^sub>1 RL rules\<^sub>2\ abbreviates \rules\<^sub>1 RLN (1, rules\<^sub>2)\. \<^descr> \[rule\<^sub>1, \, rule\<^sub>n] MRS rule\ resolves \rule\<^sub>i\ against premise \i\ of \rule\, for \i = n, \, 1\. By working from right to left, newly emerging premises are concatenated in the result, without interfering. \<^descr> \rule OF rules\ is an alternative notation for \rules MRS rule\, which makes rule composition look more like function application. Note that the argument \rules\ need not be atomic. This corresponds to the rule attribute @{attribute OF} in Isar source language. \ section \Proof terms \label{sec:proof-terms}\ text \ The Isabelle/Pure inference kernel can record the proof of each theorem as a proof term that contains all logical inferences in detail. Rule composition by resolution (\secref{sec:obj-rules}) and type-class reasoning is broken down to primitive rules of the logical framework. The proof term can be inspected by a separate proof-checker, for example. According to the well-known \<^emph>\Curry-Howard isomorphism\, a proof can be viewed as a \\\-term. Following this idea, proofs in Isabelle are internally represented by a datatype similar to the one for terms described in \secref{sec:terms}. On top of these syntactic terms, two more layers of \\\-calculus are added, which correspond to \\x :: \. B x\ and \A \ B\ according to the propositions-as-types principle. The resulting 3-level \\\-calculus resembles ``\\HOL\'' in the more abstract setting of Pure Type Systems (PTS) @{cite "Barendregt-Geuvers:2001"}, if some fine points like schematic polymorphism and type classes are ignored. \<^medskip> \<^emph>\Proof abstractions\ of the form \\<^bold>\x :: \. prf\ or \\<^bold>\p : A. prf\ correspond to introduction of \\\/\\\, and \<^emph>\proof applications\ of the form \p \ t\ or \p \ q\ correspond to elimination of \\\/\\\. Actual types \\\, propositions \A\, and terms \t\ might be suppressed and reconstructed from the overall proof term. \<^medskip> Various atomic proofs indicate special situations within the proof construction as follows. A \<^emph>\bound proof variable\ is a natural number \b\ that acts as de-Bruijn index for proof term abstractions. A \<^emph>\minimal proof\ ``\?\'' is a dummy proof term. This indicates some unrecorded part of the proof. \Hyp A\ refers to some pending hypothesis by giving its proposition. This indicates an open context of implicit hypotheses, similar to loose bound variables or free variables within a term (\secref{sec:terms}). An \<^emph>\axiom\ or \<^emph>\oracle\ \a : A[\<^vec>\]\ refers some postulated \proof constant\, which is subject to schematic polymorphism of theory content, and the particular type instantiation may be given explicitly. The vector of types \\<^vec>\\ refers to the schematic type variables in the generic proposition \A\ in canonical order. A \<^emph>\proof promise\ \a : A[\<^vec>\]\ is a placeholder for some proof of polymorphic proposition \A\, with explicit type instantiation as given by the vector \\<^vec>\\, as above. Unlike axioms or oracles, proof promises may be \<^emph>\fulfilled\ eventually, by substituting \a\ by some particular proof \q\ at the corresponding type instance. This acts like Hindley-Milner \let\-polymorphism: a generic local proof definition may get used at different type instances, and is replaced by the concrete instance eventually. A \<^emph>\named theorem\ wraps up some concrete proof as a closed formal entity, in the manner of constant definitions for proof terms. The \<^emph>\proof body\ of such boxed theorems involves some digest about oracles and promises occurring in the original proof. This allows the inference kernel to manage this critical information without the full overhead of explicit proof terms. \ subsection \Reconstructing and checking proof terms\ text \ Fully explicit proof terms can be large, but most of this information is redundant and can be reconstructed from the context. Therefore, the Isabelle/Pure inference kernel records only \<^emph>\implicit\ proof terms, by omitting all typing information in terms, all term and type labels of proof abstractions, and some argument terms of applications \p \ t\ (if possible). There are separate operations to reconstruct the full proof term later on, using \<^emph>\higher-order pattern unification\ @{cite "nipkow-patterns" and "Berghofer-Nipkow:2000:TPHOL"}. The \<^emph>\proof checker\ expects a fully reconstructed proof term, and can turn it into a theorem by replaying its primitive inferences within the kernel. \ subsection \Concrete syntax of proof terms\ text \ The concrete syntax of proof terms is a slight extension of the regular inner syntax of Isabelle/Pure @{cite "isabelle-isar-ref"}. Its main syntactic category @{syntax (inner) proof} is defined as follows: \begin{center} \begin{supertabular}{rclr} @{syntax_def (inner) proof} & = & \\<^bold>\\ \params\ \<^verbatim>\.\ \proof\ \\ & \|\ & \proof\ \\\ \any\ \\ & \|\ & \proof\ \\\ \proof\ \\ & \|\ & \id | longid\ \\ \\ \param\ & = & \idt\ \\ & \|\ & \idt\ \<^verbatim>\:\ \prop\ \\ & \|\ & \<^verbatim>\(\ \param\ \<^verbatim>\)\ \\ \\ \params\ & = & \param\ \\ & \|\ & \param\ \params\ \\ \end{supertabular} \end{center} Implicit term arguments in partial proofs are indicated by ``\_\''. Type arguments for theorems and axioms may be specified using \p \ TYPE(type)\ (they must appear before any other term argument of a theorem or axiom, but may be omitted altogether). \<^medskip> There are separate read and print operations for proof terms, in order to avoid conflicts with the regular term language. \ text %mlref \ \begin{mldecls} @{define_ML_type proof} \\ @{define_ML_type proof_body} \\ @{define_ML Proofterm.proofs: "int Unsynchronized.ref"} \\ @{define_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{define_ML Proofterm.expand_proof: "theory -> (Proofterm.thm_header -> string option) -> proof -> proof"} \\ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\proof\ represents proof terms; this is a datatype with constructors @{define_ML Abst}, @{define_ML AbsP}, @{define_ML_infix "%"}, @{define_ML_infix "%%"}, @{define_ML PBound}, @{define_ML MinProof}, @{define_ML Hyp}, @{define_ML PAxm}, @{define_ML Oracle}, @{define_ML PThm} as explained above. %FIXME PClass (!?) \<^descr> Type \<^ML_type>\proof_body\ represents the nested proof information of a named theorem, consisting of a digest of oracles and named theorem over some proof term. The digest only covers the directly visible part of the proof: in order to get the full information, the implicit graph of nested theorems needs to be traversed (e.g.\ using \<^ML>\Proofterm.fold_body_thms\). \<^descr> \<^ML>\Thm.proof_of\~\thm\ and \<^ML>\Thm.proof_body_of\~\thm\ produce the proof term or proof body (with digest of oracles and theorems) from a given theorem. Note that this involves a full join of internal futures that fulfill pending proof promises, and thus disrupts the natural bottom-up construction of proofs by introducing dynamic ad-hoc dependencies. Parallel performance may suffer by inspecting proof terms at run-time. \<^descr> \<^ML>\Proofterm.proofs\ specifies the detail of proof recording within \<^ML_type>\thm\ values produced by the inference kernel: \<^ML>\0\ records only the names of oracles, \<^ML>\1\ records oracle names and propositions, \<^ML>\2\ additionally records full proof terms. Officially named theorems that contribute to a result are recorded in any case. \<^descr> \<^ML>\Proofterm.reconstruct_proof\~\thy prop prf\ turns the implicit proof term \prf\ into a full proof of the given proposition. Reconstruction may fail if \prf\ is not a proof of \prop\, or if it does not contain sufficient information for reconstruction. Failure may only happen for proofs that are constructed manually, but not for those produced automatically by the inference kernel. \<^descr> \<^ML>\Proofterm.expand_proof\~\thy expand prf\ reconstructs and expands the proofs of nested theorems according to the given \expand\ function: a result of @{ML \SOME ""\} means full expansion, @{ML \SOME\}~\name\ means to keep the theorem node but replace its internal name, @{ML NONE} means no change. \<^descr> \<^ML>\Proof_Checker.thm_of_proof\~\thy prf\ turns the given (full) proof into a theorem, by replaying it using only primitive rules of the inference kernel. \<^descr> \<^ML>\Proof_Syntax.read_proof\~\thy b\<^sub>1 b\<^sub>2 s\ reads in a proof term. The Boolean flags indicate the use of sort and type information. Usually, typing information is left implicit and is inferred during proof reconstruction. %FIXME eliminate flags!? \<^descr> \<^ML>\Proof_Syntax.pretty_proof\~\ctxt prf\ pretty-prints the given proof term. \ text %mlex \ \<^item> \<^file>\~~/src/HOL/Proofs/ex/Proof_Terms.thy\ provides basic examples involving proof terms. \<^item> \<^file>\~~/src/HOL/Proofs/ex/XML_Data.thy\ demonstrates export and import of proof terms via XML/ML data representation. \ end diff --git a/src/Doc/Implementation/ML.thy b/src/Doc/Implementation/ML.thy --- a/src/Doc/Implementation/ML.thy +++ b/src/Doc/Implementation/ML.thy @@ -1,2211 +1,2211 @@ (*:maxLineLen=78:*) theory "ML" imports Base begin chapter \Isabelle/ML\ text \ Isabelle/ML is best understood as a certain culture based on Standard ML. Thus it is not a new programming language, but a certain way to use SML at an advanced level within the Isabelle environment. This covers a variety of aspects that are geared towards an efficient and robust platform for applications of formal logic with fully foundational proof construction --- according to the well-known \<^emph>\LCF principle\. There is specific infrastructure with library modules to address the needs of this difficult task. For example, the raw parallel programming model of Poly/ML is presented as considerably more abstract concept of \<^emph>\futures\, which is then used to augment the inference kernel, Isar theory and proof interpreter, and PIDE document management. The main aspects of Isabelle/ML are introduced below. These first-hand explanations should help to understand how proper Isabelle/ML is to be read and written, and to get access to the wealth of experience that is expressed in the source text and its history of changes.\<^footnote>\See \<^url>\https://isabelle.in.tum.de/repos/isabelle\ for the full Mercurial history. There are symbolic tags to refer to official Isabelle releases, as opposed to arbitrary \<^emph>\tip\ versions that merely reflect snapshots that are never really up-to-date.\ \ section \Style and orthography\ text \ The sources of Isabelle/Isar are optimized for \<^emph>\readability\ and \<^emph>\maintainability\. The main purpose is to tell an informed reader what is really going on and how things really work. This is a non-trivial aim, but it is supported by a certain style of writing Isabelle/ML that has emerged from long years of system development.\<^footnote>\See also the interesting style guide for OCaml \<^url>\https://caml.inria.fr/resources/doc/guides/guidelines.en.html\ which shares many of our means and ends.\ The main principle behind any coding style is \<^emph>\consistency\. For a single author of a small program this merely means ``choose your style and stick to it''. A complex project like Isabelle, with long years of development and different contributors, requires more standardization. A coding style that is changed every few years or with every new contributor is no style at all, because consistency is quickly lost. Global consistency is hard to achieve, though. Nonetheless, one should always strive at least for local consistency of modules and sub-systems, without deviating from some general principles how to write Isabelle/ML. In a sense, good coding style is like an \<^emph>\orthography\ for the sources: it helps to read quickly over the text and see through the main points, without getting distracted by accidental presentation of free-style code. \ subsection \Header and sectioning\ text \ Isabelle source files have a certain standardized header format (with precise spacing) that follows ancient traditions reaching back to the earliest versions of the system by Larry Paulson. See \<^file>\~~/src/Pure/thm.ML\, for example. The header includes at least \<^verbatim>\Title\ and \<^verbatim>\Author\ entries, followed by a prose description of the purpose of the module. The latter can range from a single line to several paragraphs of explanations. The rest of the file is divided into chapters, sections, subsections, subsubsections, paragraphs etc.\ using a simple layout via ML comments as follows. @{verbatim [display] \ (**** chapter ****) (*** section ***) (** subsection **) (* subsubsection *) (*short paragraph*) (* long paragraph, with more text *)\} As in regular typography, there is some extra space \<^emph>\before\ section headings that are adjacent to plain text, but not other headings as in the example above. \<^medskip> The precise wording of the prose text given in these headings is chosen carefully to introduce the main theme of the subsequent formal ML text. \ subsection \Naming conventions\ text \ Since ML is the primary medium to express the meaning of the source text, naming of ML entities requires special care. \ paragraph \Notation.\ text \ A name consists of 1--3 \<^emph>\words\ (rarely 4, but not more) that are separated by underscore. There are three variants concerning upper or lower case letters, which are used for certain ML categories as follows: \<^medskip> \begin{tabular}{lll} variant & example & ML categories \\\hline lower-case & \<^ML_text>\foo_bar\ & values, types, record fields \\ capitalized & \<^ML_text>\Foo_Bar\ & datatype constructors, structures, functors \\ upper-case & \<^ML_text>\FOO_BAR\ & special values, exception constructors, signatures \\ \end{tabular} \<^medskip> For historical reasons, many capitalized names omit underscores, e.g.\ old-style \<^ML_text>\FooBar\ instead of \<^ML_text>\Foo_Bar\. Genuine mixed-case names are \<^emph>\not\ used, because clear division of words is essential for readability.\<^footnote>\Camel-case was invented to workaround the lack of underscore in some early non-ASCII character sets. Later it became habitual in some language communities that are now strong in numbers.\ A single (capital) character does not count as ``word'' in this respect: some Isabelle/ML names are suffixed by extra markers like this: \<^ML_text>\foo_barT\. Name variants are produced by adding 1--3 primes, e.g.\ \<^ML_text>\foo'\, \<^ML_text>\foo''\, or \<^ML_text>\foo'''\, but not \<^ML_text>\foo''''\ or more. Decimal digits scale better to larger numbers, e.g.\ \<^ML_text>\foo0\, \<^ML_text>\foo1\, \<^ML_text>\foo42\. \ paragraph \Scopes.\ text \ Apart from very basic library modules, ML structures are not ``opened'', but names are referenced with explicit qualification, as in \<^ML>\Syntax.string_of_term\ for example. When devising names for structures and their components it is important to aim at eye-catching compositions of both parts, because this is how they are seen in the sources and documentation. For the same reasons, aliases of well-known library functions should be avoided. Local names of function abstraction or case/let bindings are typically shorter, sometimes using only rudiments of ``words'', while still avoiding cryptic shorthands. An auxiliary function called \<^ML_text>\helper\, \<^ML_text>\aux\, or \<^ML_text>\f\ is considered bad style. Example: @{verbatim [display] \ (* RIGHT *) fun print_foo ctxt foo = let fun print t = ... Syntax.string_of_term ctxt t ... in ... end; (* RIGHT *) fun print_foo ctxt foo = let val string_of_term = Syntax.string_of_term ctxt; fun print t = ... string_of_term t ... in ... end; (* WRONG *) val string_of_term = Syntax.string_of_term; fun print_foo ctxt foo = let fun aux t = ... string_of_term ctxt t ... in ... end;\} \ paragraph \Specific conventions.\ text \ Here are some specific name forms that occur frequently in the sources. \<^item> A function that maps \<^ML_text>\foo\ to \<^ML_text>\bar\ is called \<^ML_text>\foo_to_bar\ or \<^ML_text>\bar_of_foo\ (never \<^ML_text>\foo2bar\, nor \<^ML_text>\bar_from_foo\, nor \<^ML_text>\bar_for_foo\, nor \<^ML_text>\bar4foo\). \<^item> The name component \<^ML_text>\legacy\ means that the operation is about to be discontinued soon. \<^item> The name component \<^ML_text>\global\ means that this works with the background theory instead of the regular local context (\secref{sec:context}), sometimes for historical reasons, sometimes due a genuine lack of locality of the concept involved, sometimes as a fall-back for the lack of a proper context in the application code. Whenever there is a non-global variant available, the application should be migrated to use it with a proper local context. \<^item> Variables of the main context types of the Isabelle/Isar framework (\secref{sec:context} and \chref{ch:local-theory}) have firm naming conventions as follows: \<^item> theories are called \<^ML_text>\thy\, rarely \<^ML_text>\theory\ (never \<^ML_text>\thry\) \<^item> proof contexts are called \<^ML_text>\ctxt\, rarely \<^ML_text>\context\ (never \<^ML_text>\ctx\) \<^item> generic contexts are called \<^ML_text>\context\ \<^item> local theories are called \<^ML_text>\lthy\, except for local theories that are treated as proof context (which is a semantic super-type) Variations with primed or decimal numbers are always possible, as well as semantic prefixes like \<^ML_text>\foo_thy\ or \<^ML_text>\bar_ctxt\, but the base conventions above need to be preserved. This allows to emphasize their data flow via plain regular expressions in the text editor. \<^item> The main logical entities (\secref{ch:logic}) have established naming convention as follows: \<^item> sorts are called \<^ML_text>\S\ \<^item> types are called \<^ML_text>\T\, \<^ML_text>\U\, or \<^ML_text>\ty\ (never \<^ML_text>\t\) \<^item> terms are called \<^ML_text>\t\, \<^ML_text>\u\, or \<^ML_text>\tm\ (never \<^ML_text>\trm\) \<^item> certified types are called \<^ML_text>\cT\, rarely \<^ML_text>\T\, with variants as for types \<^item> certified terms are called \<^ML_text>\ct\, rarely \<^ML_text>\t\, with variants as for terms (never \<^ML_text>\ctrm\) \<^item> theorems are called \<^ML_text>\th\, or \<^ML_text>\thm\ Proper semantic names override these conventions completely. For example, the left-hand side of an equation (as a term) can be called \<^ML_text>\lhs\ (not \<^ML_text>\lhs_tm\). Or a term that is known to be a variable can be called \<^ML_text>\v\ or \<^ML_text>\x\. \<^item> Tactics (\secref{sec:tactics}) are sufficiently important to have specific naming conventions. The name of a basic tactic definition always has a \<^ML_text>\_tac\ suffix, the subgoal index (if applicable) is always called \<^ML_text>\i\, and the goal state (if made explicit) is usually called \<^ML_text>\st\ instead of the somewhat misleading \<^ML_text>\thm\. Any other arguments are given before the latter two, and the general context is given first. Example: @{verbatim [display] \ fun my_tac ctxt arg1 arg2 i st = ...\} Note that the goal state \<^ML_text>\st\ above is rarely made explicit, if tactic combinators (tacticals) are used as usual. A tactic that requires a proof context needs to make that explicit as seen in the \<^verbatim>\ctxt\ argument above. Do not refer to the background theory of \<^verbatim>\st\ -- it is not a proper context, but merely a formal certificate. \ subsection \General source layout\ text \ The general Isabelle/ML source layout imitates regular type-setting conventions, augmented by the requirements for deeply nested expressions that are commonplace in functional programming. \ paragraph \Line length\ text \ is limited to 80 characters according to ancient standards, but we allow as much as 100 characters (not more).\<^footnote>\Readability requires to keep the beginning of a line in view while watching its end. Modern wide-screen displays do not change the way how the human brain works. Sources also need to be printable on plain paper with reasonable font-size.\ The extra 20 characters acknowledge the space requirements due to qualified library references in Isabelle/ML. \ paragraph \White-space\ text \ is used to emphasize the structure of expressions, following mostly standard conventions for mathematical typesetting, as can be seen in plain {\TeX} or {\LaTeX}. This defines positioning of spaces for parentheses, punctuation, and infixes as illustrated here: @{verbatim [display] \ val x = y + z * (a + b); val pair = (a, b); val record = {foo = 1, bar = 2};\} Lines are normally broken \<^emph>\after\ an infix operator or punctuation character. For example: @{verbatim [display] \ val x = a + b + c; val tuple = (a, b, c); \} Some special infixes (e.g.\ \<^ML_text>\|>\) work better at the start of the line, but punctuation is always at the end. Function application follows the tradition of \\\-calculus, not informal mathematics. For example: \<^ML_text>\f a b\ for a curried function, or \<^ML_text>\g (a, b)\ for a tupled function. Note that the space between \<^ML_text>\g\ and the pair \<^ML_text>\(a, b)\ follows the important principle of \<^emph>\compositionality\: the layout of \<^ML_text>\g p\ does not change when \<^ML_text>\p\ is refined to the concrete pair \<^ML_text>\(a, b)\. \ paragraph \Indentation\ text \ uses plain spaces, never hard tabulators.\<^footnote>\Tabulators were invented to move the carriage of a type-writer to certain predefined positions. In software they could be used as a primitive run-length compression of consecutive spaces, but the precise result would depend on non-standardized text editor configuration.\ Each level of nesting is indented by 2 spaces, sometimes 1, very rarely 4, never 8 or any other odd number. Indentation follows a simple logical format that only depends on the nesting depth, not the accidental length of the text that initiates a level of nesting. Example: @{verbatim [display] \ (* RIGHT *) if b then expr1_part1 expr1_part2 else expr2_part1 expr2_part2 (* WRONG *) if b then expr1_part1 expr1_part2 else expr2_part1 expr2_part2\} The second form has many problems: it assumes a fixed-width font when viewing the sources, it uses more space on the line and thus makes it hard to observe its strict length limit (working against \<^emph>\readability\), it requires extra editing to adapt the layout to changes of the initial text (working against \<^emph>\maintainability\) etc. \<^medskip> For similar reasons, any kind of two-dimensional or tabular layouts, ASCII-art with lines or boxes of asterisks etc.\ should be avoided. \ paragraph \Complex expressions\ text \ that consist of multi-clausal function definitions, \<^ML_text>\handle\, \<^ML_text>\case\, \<^ML_text>\let\ (and combinations) require special attention. The syntax of Standard ML is quite ambitious and admits a lot of variance that can distort the meaning of the text. Multiple clauses of \<^ML_text>\fun\, \<^ML_text>\fn\, \<^ML_text>\handle\, \<^ML_text>\case\ get extra indentation to indicate the nesting clearly. Example: @{verbatim [display] \ (* RIGHT *) fun foo p1 = expr1 | foo p2 = expr2 (* WRONG *) fun foo p1 = expr1 | foo p2 = expr2\} Body expressions consisting of \<^ML_text>\case\ or \<^ML_text>\let\ require care to maintain compositionality, to prevent loss of logical indentation where it is especially important to see the structure of the text. Example: @{verbatim [display] \ (* RIGHT *) fun foo p1 = (case e of q1 => ... | q2 => ...) | foo p2 = let ... in ... end (* WRONG *) fun foo p1 = case e of q1 => ... | q2 => ... | foo p2 = let ... in ... end\} Extra parentheses around \<^ML_text>\case\ expressions are optional, but help to analyse the nesting based on character matching in the text editor. \<^medskip> There are two main exceptions to the overall principle of compositionality in the layout of complex expressions. \<^enum> \<^ML_text>\if\ expressions are iterated as if ML had multi-branch conditionals, e.g. @{verbatim [display] \ (* RIGHT *) if b1 then e1 else if b2 then e2 else e3\} \<^enum> \<^ML_text>\fn\ abstractions are often layed-out as if they would lack any structure by themselves. This traditional form is motivated by the possibility to shift function arguments back and forth wrt.\ additional combinators. Example: @{verbatim [display] \ (* RIGHT *) fun foo x y = fold (fn z => expr)\} Here the visual appearance is that of three arguments \<^ML_text>\x\, \<^ML_text>\y\, \<^ML_text>\z\ in a row. Such weakly structured layout should be use with great care. Here are some counter-examples involving \<^ML_text>\let\ expressions: @{verbatim [display] \ (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end (* WRONG *) fun foo x = let val y = ... in ... end\} \<^medskip> In general the source layout is meant to emphasize the structure of complex language expressions, not to pretend that SML had a completely different syntax (say that of Haskell, Scala, Java). \ section \ML embedded into Isabelle/Isar\ text \ ML and Isar are intertwined via an open-ended bootstrap process that provides more and more programming facilities and logical content in an alternating manner. Bootstrapping starts from the raw environment of existing implementations of Standard ML (mainly Poly/ML). Isabelle/Pure marks the point where the raw ML toplevel is superseded by Isabelle/ML within the Isar theory and proof language, with a uniform context for arbitrary ML values (see also \secref{sec:context}). This formal environment holds ML compiler bindings, logical entities, and many other things. Object-logics like Isabelle/HOL are built within the Isabelle/ML/Isar environment by introducing suitable theories with associated ML modules, either inlined within \<^verbatim>\.thy\ files, or as separate \<^verbatim>\.ML\ files that are loading from some theory. Thus Isabelle/HOL is defined as a regular user-space application within the Isabelle framework. Further add-on tools can be implemented in ML within the Isar context in the same manner: ML is part of the standard repertoire of Isabelle, and there is no distinction between ``users'' and ``developers'' in this respect. \ subsection \Isar ML commands\ text \ The primary Isar source language provides facilities to ``open a window'' to the underlying ML compiler. Especially see the Isar commands @{command_ref "ML_file"} and @{command_ref "ML"}: both work the same way, but the source text is provided differently, via a file vs.\ inlined, respectively. Apart from embedding ML into the main theory definition like that, there are many more commands that refer to ML source, such as @{command_ref setup} or @{command_ref declaration}. Even more fine-grained embedding of ML into Isar is encountered in the proof method @{method_ref tactic}, which refines the pending goal state via a given expression of type \<^ML_type>\tactic\. \ text %mlex \ The following artificial example demonstrates some ML toplevel declarations within the implicit Isar theory context. This is regular functional programming without referring to logical entities yet. \ ML \ fun factorial 0 = 1 | factorial n = n * factorial (n - 1) \ text \ Here the ML environment is already managed by Isabelle, i.e.\ the \<^ML>\factorial\ function is not yet accessible in the preceding paragraph, nor in a different theory that is independent from the current one in the import hierarchy. Removing the above ML declaration from the source text will remove any trace of this definition, as expected. The Isabelle/ML toplevel environment is managed in a \<^emph>\stateless\ way: in contrast to the raw ML toplevel, there are no global side-effects involved here.\<^footnote>\Such a stateless compilation environment is also a prerequisite for robust parallel compilation within independent nodes of the implicit theory development graph.\ \<^medskip> The next example shows how to embed ML into Isar proofs, using @{command_ref "ML_prf"} instead of @{command_ref "ML"}. As illustrated below, the effect on the ML environment is local to the whole proof body, but ignoring the block structure. \ notepad begin ML_prf %"ML" \val a = 1\ { ML_prf %"ML" \val b = a + 1\ } \ \Isar block structure ignored by ML environment\ ML_prf %"ML" \val c = b + 1\ end text \ By side-stepping the normal scoping rules for Isar proof blocks, embedded ML code can refer to the different contexts and manipulate corresponding entities, e.g.\ export a fact from a block context. \<^medskip> Two further ML commands are useful in certain situations: @{command_ref ML_val} and @{command_ref ML_command} are \<^emph>\diagnostic\ in the sense that there is no effect on the underlying environment, and can thus be used anywhere. The examples below produce long strings of digits by invoking \<^ML>\factorial\: @{command ML_val} takes care of printing the ML toplevel result, but @{command ML_command} is silent so we produce an explicit output message. \ ML_val \factorial 100\ ML_command \writeln (string_of_int (factorial 100))\ notepad begin ML_val \factorial 100\ ML_command \writeln (string_of_int (factorial 100))\ end subsection \Compile-time context\ text \ Whenever the ML compiler is invoked within Isabelle/Isar, the formal context is passed as a thread-local reference variable. Thus ML code may access the theory context during compilation, by reading or writing the (local) theory under construction. Note that such direct access to the compile-time context is rare. In practice it is typically done via some derived ML functions instead. \ text %mlref \ \begin{mldecls} @{define_ML Context.the_generic_context: "unit -> Context.generic"} \\ @{define_ML "Context.>>": "(Context.generic -> Context.generic) -> unit"} \\ @{define_ML ML_Thms.bind_thms: "string * thm list -> unit"} \\ @{define_ML ML_Thms.bind_thm: "string * thm -> unit"} \\ \end{mldecls} \<^descr> \<^ML>\Context.the_generic_context ()\ refers to the theory context of the ML toplevel --- at compile time. ML code needs to take care to refer to \<^ML>\Context.the_generic_context ()\ correctly. Recall that evaluation of a function body is delayed until actual run-time. \<^descr> \<^ML>\Context.>>\~\f\ applies context transformation \f\ to the implicit context of the ML toplevel. \<^descr> \<^ML>\ML_Thms.bind_thms\~\(name, thms)\ stores a list of theorems produced in ML both in the (global) theory context and the ML toplevel, associating it with the provided name. \<^descr> \<^ML>\ML_Thms.bind_thm\ is similar to \<^ML>\ML_Thms.bind_thms\ but refers to a singleton fact. It is important to note that the above functions are really restricted to the compile time, even though the ML compiler is invoked at run-time. The majority of ML code either uses static antiquotations (\secref{sec:ML-antiq}) or refers to the theory or proof context at run-time, by explicit functional abstraction. \ subsection \Antiquotations \label{sec:ML-antiq}\ text \ A very important consequence of embedding ML into Isar is the concept of \<^emph>\ML antiquotation\. The standard token language of ML is augmented by special syntactic entities of the following form: \<^rail>\ @{syntax_def antiquote}: '@{' name args '}' \ Here @{syntax name} and @{syntax args} are outer syntax categories, as defined in @{cite "isabelle-isar-ref"}. \<^medskip> A regular antiquotation \@{name args}\ processes its arguments by the usual means of the Isar source language, and produces corresponding ML source text, either as literal \<^emph>\inline\ text (e.g.\ \@{term t}\) or abstract \<^emph>\value\ (e.g. \@{thm th}\). This pre-compilation scheme allows to refer to formal entities in a robust manner, with proper static scoping and with some degree of logical checking of small portions of the code. \ subsection \Printing ML values\ text \ The ML compiler knows about the structure of values according to their static type, and can print them in the manner of its toplevel, although the details are non-portable. The antiquotations @{ML_antiquotation_def "make_string"} and @{ML_antiquotation_def "print"} provide a quasi-portable way to refer to this potential capability of the underlying ML system in generic Isabelle/ML sources. This is occasionally useful for diagnostic or demonstration purposes. Note that production-quality tools require proper user-level error messages, avoiding raw ML values in the output. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "make_string"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "print"} & : & \ML_antiquotation\ \\ \end{matharray} \<^rail>\ @@{ML_antiquotation make_string} ; @@{ML_antiquotation print} embedded? \ \<^descr> \@{make_string}\ inlines a function to print arbitrary values similar to the ML toplevel. The result is compiler dependent and may fall back on "?" in certain situations. The value of configuration option @{attribute_ref ML_print_depth} determines further details of output. \<^descr> \@{print f}\ uses the ML function \f: string -> unit\ to output the result of \@{make_string}\ above, together with the source position of the antiquotation. The default output function is \<^ML>\writeln\. \ text %mlex \ The following artificial examples show how to produce adhoc output of ML values for debugging purposes. \ ML_val \ val x = 42; val y = true; writeln (\<^make_string> {x = x, y = y}); \<^print> {x = x, y = y}; \<^print>\tracing\ {x = x, y = y}; \ section \Canonical argument order \label{sec:canonical-argument-order}\ text \ Standard ML is a language in the tradition of \\\-calculus and \<^emph>\higher-order functional programming\, similar to OCaml, Haskell, or Isabelle/Pure and HOL as logical languages. Getting acquainted with the native style of representing functions in that setting can save a lot of extra boiler-plate of redundant shuffling of arguments, auxiliary abstractions etc. Functions are usually \<^emph>\curried\: the idea of turning arguments of type \\\<^sub>i\ (for \i \ {1, \ n}\) into a result of type \\\ is represented by the iterated function space \\\<^sub>1 \ \ \ \\<^sub>n \ \\. This is isomorphic to the well-known encoding via tuples \\\<^sub>1 \ \ \ \\<^sub>n \ \\, but the curried version fits more smoothly into the basic calculus.\<^footnote>\The difference is even more significant in HOL, because the redundant tuple structure needs to be accommodated extraneous proof steps.\ Currying gives some flexibility due to \<^emph>\partial application\. A function \f: \\<^sub>1 \ \\<^sub>2 \ \\ can be applied to \x: \\<^sub>1\ and the remaining \(f x): \\<^sub>2 \ \\ passed to another function etc. How well this works in practice depends on the order of arguments. In the worst case, arguments are arranged erratically, and using a function in a certain situation always requires some glue code. Thus we would get exponentially many opportunities to decorate the code with meaningless permutations of arguments. This can be avoided by \<^emph>\canonical argument order\, which observes certain standard patterns and minimizes adhoc permutations in their application. In Isabelle/ML, large portions of text can be written without auxiliary operations like \swap: \ \ \ \ \ \ \\ or \C: (\ \ \ \ \) \ (\ \ \ \ \)\ (the latter is not present in the Isabelle/ML library). \<^medskip> The main idea is that arguments that vary less are moved further to the left than those that vary more. Two particularly important categories of functions are \<^emph>\selectors\ and \<^emph>\updates\. The subsequent scheme is based on a hypothetical set-like container of type \\\ that manages elements of type \\\. Both the names and types of the associated operations are canonical for Isabelle/ML. \begin{center} \begin{tabular}{ll} kind & canonical name and type \\\hline selector & \member: \ \ \ \ bool\ \\ update & \insert: \ \ \ \ \\ \\ \end{tabular} \end{center} Given a container \B: \\, the partially applied \member B\ is a predicate over elements \\ \ bool\, and thus represents the intended denotation directly. It is customary to pass the abstract predicate to further operations, not the concrete container. The argument order makes it easy to use other combinators: \forall (member B) list\ will check a list of elements for membership in \B\ etc. Often the explicit \list\ is pointless and can be contracted to \forall (member B)\ to get directly a predicate again. In contrast, an update operation varies the container, so it moves to the right: \insert a\ is a function \\ \ \\ to insert a value \a\. These can be composed naturally as \insert c \ insert b \ insert a\. The slightly awkward inversion of the composition order is due to conventional mathematical notation, which can be easily amended as explained below. \ subsection \Forward application and composition\ text \ Regular function application and infix notation works best for relatively deeply structured expressions, e.g.\ \h (f x y + g z)\. The important special case of \<^emph>\linear transformation\ applies a cascade of functions \f\<^sub>n (\ (f\<^sub>1 x))\. This becomes hard to read and maintain if the functions are themselves given as complex expressions. The notation can be significantly improved by introducing \<^emph>\forward\ versions of application and composition as follows: \<^medskip> \begin{tabular}{lll} \x |> f\ & \\\ & \f x\ \\ \(f #> g) x\ & \\\ & \x |> f |> g\ \\ \end{tabular} \<^medskip> This enables to write conveniently \x |> f\<^sub>1 |> \ |> f\<^sub>n\ or \f\<^sub>1 #> \ #> f\<^sub>n\ for its functional abstraction over \x\. \<^medskip> There is an additional set of combinators to accommodate multiple results (via pairs) that are passed on as multiple arguments (via currying). \<^medskip> \begin{tabular}{lll} \(x, y) |-> f\ & \\\ & \f x y\ \\ \(f #-> g) x\ & \\\ & \x |> f |-> g\ \\ \end{tabular} \<^medskip> \ text %mlref \ \begin{mldecls} @{define_ML_infix "|>" : "'a * ('a -> 'b) -> 'b"} \\ @{define_ML_infix "|->" : "('c * 'a) * ('c -> 'a -> 'b) -> 'b"} \\ @{define_ML_infix "#>" : "('a -> 'b) * ('b -> 'c) -> 'a -> 'c"} \\ @{define_ML_infix "#->" : "('a -> 'c * 'b) * ('c -> 'b -> 'd) -> 'a -> 'd"} \\ \end{mldecls} \ subsection \Canonical iteration\ text \ As explained above, a function \f: \ \ \ \ \\ can be understood as update on a configuration of type \\\, parameterized by an argument of type \\\. Given \a: \\ the partial application \(f a): \ \ \\ operates homogeneously on \\\. This can be iterated naturally over a list of parameters \[a\<^sub>1, \, a\<^sub>n]\ as \f a\<^sub>1 #> \ #> f a\<^sub>n\. The latter expression is again a function \\ \ \\. It can be applied to an initial configuration \b: \\ to start the iteration over the given list of arguments: each \a\ in \a\<^sub>1, \, a\<^sub>n\ is applied consecutively by updating a cumulative configuration. The \fold\ combinator in Isabelle/ML lifts a function \f\ as above to its iterated version over a list of arguments. Lifting can be repeated, e.g.\ \(fold \ fold) f\ iterates over a list of lists as expected. The variant \fold_rev\ works inside-out over the list of arguments, such that \fold_rev f \ fold f \ rev\ holds. The \fold_map\ combinator essentially performs \fold\ and \map\ simultaneously: each application of \f\ produces an updated configuration together with a side-result; the iteration collects all such side-results as a separate list. \ text %mlref \ \begin{mldecls} @{define_ML fold: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ @{define_ML fold_rev: "('a -> 'b -> 'b) -> 'a list -> 'b -> 'b"} \\ @{define_ML fold_map: "('a -> 'b -> 'c * 'b) -> 'a list -> 'b -> 'c list * 'b"} \\ \end{mldecls} \<^descr> \<^ML>\fold\~\f\ lifts the parametrized update function \f\ to a list of parameters. \<^descr> \<^ML>\fold_rev\~\f\ is similar to \<^ML>\fold\~\f\, but works inside-out, as if the list would be reversed. \<^descr> \<^ML>\fold_map\~\f\ lifts the parametrized update function \f\ (with side-result) to a list of parameters and cumulative side-results. \begin{warn} The literature on functional programming provides a confusing multitude of combinators called \foldl\, \foldr\ etc. SML97 provides its own variations as \<^ML>\List.foldl\ and \<^ML>\List.foldr\, while the classic Isabelle library also has the historic \<^ML>\Library.foldl\ and \<^ML>\Library.foldr\. To avoid unnecessary complication, all these historical versions should be ignored, and the canonical \<^ML>\fold\ (or \<^ML>\fold_rev\) used exclusively. \end{warn} \ text %mlex \ The following example shows how to fill a text buffer incrementally by adding strings, either individually or from a given list. \ ML_val \ val s = Buffer.empty |> Buffer.add "digits: " |> fold (Buffer.add o string_of_int) (0 upto 9) |> Buffer.content; \<^assert> (s = "digits: 0123456789"); \ text \ Note how \<^ML>\fold (Buffer.add o string_of_int)\ above saves an extra \<^ML>\map\ over the given list. This kind of peephole optimization reduces both the code size and the tree structures in memory (``deforestation''), but it requires some practice to read and write fluently. \<^medskip> The next example elaborates the idea of canonical iteration, demonstrating fast accumulation of tree content using a text buffer. \ ML \ datatype tree = Text of string | Elem of string * tree list; fun slow_content (Text txt) = txt | slow_content (Elem (name, ts)) = "<" ^ name ^ ">" ^ implode (map slow_content ts) ^ "" fun add_content (Text txt) = Buffer.add txt | add_content (Elem (name, ts)) = Buffer.add ("<" ^ name ^ ">") #> fold add_content ts #> Buffer.add (""); fun fast_content tree = Buffer.empty |> add_content tree |> Buffer.content; \ text \ The slowness of \<^ML>\slow_content\ is due to the \<^ML>\implode\ of the recursive results, because it copies previously produced strings again and again. The incremental \<^ML>\add_content\ avoids this by operating on a buffer that is passed through in a linear fashion. Using \<^ML_text>\#>\ and contraction over the actual buffer argument saves some additional boiler-plate. Of course, the two \<^ML>\Buffer.add\ invocations with concatenated strings could have been split into smaller parts, but this would have obfuscated the source without making a big difference in performance. Here we have done some peephole-optimization for the sake of readability. Another benefit of \<^ML>\add_content\ is its ``open'' form as a function on buffers that can be continued in further linear transformations, folding etc. Thus it is more compositional than the naive \<^ML>\slow_content\. As realistic example, compare the old-style \<^ML>\Term.maxidx_of_term: term -> int\ with the newer \<^ML>\Term.maxidx_term: term -> int -> int\ in Isabelle/Pure. Note that \<^ML>\fast_content\ above is only defined as example. In many practical situations, it is customary to provide the incremental \<^ML>\add_content\ only and leave the initialization and termination to the concrete application to the user. \ section \Message output channels \label{sec:message-channels}\ text \ Isabelle provides output channels for different kinds of messages: regular output, high-volume tracing information, warnings, and errors. Depending on the user interface involved, these messages may appear in different text styles or colours. The standard output for batch sessions prefixes each line of warnings by \<^verbatim>\###\ and errors by \<^verbatim>\***\, but leaves anything else unchanged. The message body may contain further markup and formatting, which is routinely used in the Prover IDE @{cite "isabelle-jedit"}. Messages are associated with the transaction context of the running Isar command. This enables the front-end to manage commands and resulting messages together. For example, after deleting a command from a given theory document version, the corresponding message output can be retracted from the display. \ text %mlref \ \begin{mldecls} @{define_ML writeln: "string -> unit"} \\ @{define_ML tracing: "string -> unit"} \\ @{define_ML warning: "string -> unit"} \\ @{define_ML error: "string -> 'a"} % FIXME Output.error_message (!?) \\ \end{mldecls} \<^descr> \<^ML>\writeln\~\text\ outputs \text\ as regular message. This is the primary message output operation of Isabelle and should be used by default. \<^descr> \<^ML>\tracing\~\text\ outputs \text\ as special tracing message, indicating potential high-volume output to the front-end (hundreds or thousands of messages issued by a single command). The idea is to allow the user-interface to downgrade the quality of message display to achieve higher throughput. Note that the user might have to take special actions to see tracing output, e.g.\ switch to a different output window. So this channel should not be used for regular output. \<^descr> \<^ML>\warning\~\text\ outputs \text\ as warning, which typically means some extra emphasis on the front-end side (color highlighting, icons, etc.). \<^descr> \<^ML>\error\~\text\ raises exception \<^ML>\ERROR\~\text\ and thus lets the Isar toplevel print \text\ on the error channel, which typically means some extra emphasis on the front-end side (color highlighting, icons, etc.). This assumes that the exception is not handled before the command terminates. Handling exception \<^ML>\ERROR\~\text\ is a perfectly legal alternative: it means that the error is absorbed without any message output. \begin{warn} The actual error channel is accessed via \<^ML>\Output.error_message\, but this is normally not used directly in user code. \end{warn} \begin{warn} Regular Isabelle/ML code should output messages exclusively by the official channels. Using raw I/O on \<^emph>\stdout\ or \<^emph>\stderr\ instead (e.g.\ via \<^ML>\TextIO.output\) is apt to cause problems in the presence of parallel and asynchronous processing of Isabelle theories. Such raw output might be displayed by the front-end in some system console log, with a low chance that the user will ever see it. Moreover, as a genuine side-effect on global process channels, there is no proper way to retract output when Isar command transactions are reset by the system. \end{warn} \begin{warn} The message channels should be used in a message-oriented manner. This means that multi-line output that logically belongs together is issued by a single invocation of \<^ML>\writeln\ etc.\ with the functional concatenation of all message constituents. \end{warn} \ text %mlex \ The following example demonstrates a multi-line warning. Note that in some situations the user sees only the first line, so the most important point should be made first. \ ML_command \ warning (cat_lines ["Beware the Jabberwock, my son!", "The jaws that bite, the claws that catch!", "Beware the Jubjub Bird, and shun", "The frumious Bandersnatch!"]); \ text \ \<^medskip> An alternative is to make a paragraph of freely-floating words as follows. \ ML_command \ warning (Pretty.string_of (Pretty.para "Beware the Jabberwock, my son! \ \The jaws that bite, the claws that catch! \ \Beware the Jubjub Bird, and shun \ \The frumious Bandersnatch!")) \ text \ This has advantages with variable window / popup sizes, but might make it harder to search for message content systematically, e.g.\ by other tools or by humans expecting the ``verse'' of a formal message in a fixed layout. \ section \Exceptions \label{sec:exceptions}\ text \ The Standard ML semantics of strict functional evaluation together with exceptions is rather well defined, but some delicate points need to be observed to avoid that ML programs go wrong despite static type-checking. Exceptions in Isabelle/ML are subsequently categorized as follows. \ paragraph \Regular user errors.\ text \ These are meant to provide informative feedback about malformed input etc. The \<^emph>\error\ function raises the corresponding \<^ML>\ERROR\ exception, with a plain text message as argument. \<^ML>\ERROR\ exceptions can be handled internally, in order to be ignored, turned into other exceptions, or cascaded by appending messages. If the corresponding Isabelle/Isar command terminates with an \<^ML>\ERROR\ exception state, the system will print the result on the error channel (see \secref{sec:message-channels}). It is considered bad style to refer to internal function names or values in ML source notation in user error messages. Do not use \@{make_string}\ nor \@{here}\! Grammatical correctness of error messages can be improved by \<^emph>\omitting\ final punctuation: messages are often concatenated or put into a larger context (e.g.\ augmented with source position). Note that punctuation after formal entities (types, terms, theorems) is particularly prone to user confusion. \ paragraph \Program failures.\ text \ There is a handful of standard exceptions that indicate general failure situations, or failures of core operations on logical entities (types, terms, theorems, theories, see \chref{ch:logic}). These exceptions indicate a genuine breakdown of the program, so the main purpose is to determine quickly what has happened where. Traditionally, the (short) exception message would include the name of an ML function, although this is no longer necessary, because the ML runtime system attaches detailed source position stemming from the corresponding \<^ML_text>\raise\ keyword. \<^medskip> User modules can always introduce their own custom exceptions locally, e.g.\ to organize internal failures robustly without overlapping with existing exceptions. Exceptions that are exposed in module signatures require extra care, though, and should \<^emph>\not\ be introduced by default. Surprise by users of a module can be often minimized by using plain user errors instead. \ paragraph \Interrupts.\ text \ These indicate arbitrary system events: both the ML runtime system and the Isabelle/ML infrastructure signal various exceptional situations by raising the special \<^ML>\Exn.Interrupt\ exception in user code. This is the one and only way that physical events can intrude an Isabelle/ML program. Such an interrupt can mean out-of-memory, stack overflow, timeout, internal signaling of threads, or a POSIX process signal. An Isabelle/ML program that intercepts interrupts becomes dependent on physical effects of the environment. Even worse, exception handling patterns that are too general by accident, e.g.\ by misspelled exception constructors, will cover interrupts unintentionally and thus render the program semantics ill-defined. Note that the Interrupt exception dates back to the original SML90 language definition. It was excluded from the SML97 version to avoid its malign impact on ML program semantics, but without providing a viable alternative. Isabelle/ML recovers physical interruptibility (which is an indispensable tool to implement managed evaluation of command transactions), but requires user code to be strictly transparent wrt.\ interrupts. \begin{warn} Isabelle/ML user code needs to terminate promptly on interruption, without guessing at its meaning to the system infrastructure. Temporary handling of interrupts for cleanup of global resources etc.\ needs to be followed immediately by re-raising of the original exception. \end{warn} \ text %mlref \ \begin{mldecls} @{define_ML try: "('a -> 'b) -> 'a -> 'b option"} \\ @{define_ML can: "('a -> 'b) -> 'a -> bool"} \\ @{define_ML_exception ERROR of string} \\ @{define_ML_exception Fail of string} \\ @{define_ML Exn.is_interrupt: "exn -> bool"} \\ @{define_ML Exn.reraise: "exn -> 'a"} \\ @{define_ML Runtime.exn_trace: "(unit -> 'a) -> 'a"} \\ \end{mldecls} \<^rail>\ (@@{ML_antiquotation try} | @@{ML_antiquotation can}) embedded \ \<^descr> \<^ML>\try\~\f x\ makes the partiality of evaluating \f x\ explicit via the option datatype. Interrupts are \<^emph>\not\ handled here, i.e.\ this form serves as safe replacement for the \<^emph>\unsafe\ version \<^ML_text>\(SOME\~\f x\~\<^ML_text>\handle _ => NONE)\ that is occasionally seen in books about SML97, but not in Isabelle/ML. \<^descr> \<^ML>\can\ is similar to \<^ML>\try\ with more abstract result. \<^descr> \<^ML>\ERROR\~\msg\ represents user errors; this exception is normally raised indirectly via the \<^ML>\error\ function (see \secref{sec:message-channels}). \<^descr> \<^ML>\Fail\~\msg\ represents general program failures. \<^descr> \<^ML>\Exn.is_interrupt\ identifies interrupts robustly, without mentioning concrete exception constructors in user code. Handled interrupts need to be re-raised promptly! \<^descr> \<^ML>\Exn.reraise\~\exn\ raises exception \exn\ while preserving its implicit position information (if possible, depending on the ML platform). \<^descr> \<^ML>\Runtime.exn_trace\~\<^ML_text>\(fn () =>\~\e\\<^ML_text>\)\ evaluates expression \e\ while printing a full trace of its stack of nested exceptions (if possible, depending on the ML platform). Inserting \<^ML>\Runtime.exn_trace\ into ML code temporarily is useful for debugging, but not suitable for production code. \ text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "try"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "can"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "assert"} & : & \ML_antiquotation\ \\ @{ML_antiquotation_def "undefined"} & : & \ML_antiquotation\ \\ \end{matharray} \<^descr> \@{try}\ and \{can}\ are similar to the corresponding functions, but the argument is taken directly as ML expression: functional abstraction and application is done implicitly. \<^descr> \@{assert}\ inlines a function \<^ML_type>\bool -> unit\ that raises \<^ML>\Fail\ if the argument is \<^ML>\false\. Due to inlining the source position of failed assertions is included in the error output. \<^descr> \@{undefined}\ inlines \<^verbatim>\raise\~\<^ML>\Match\, i.e.\ the ML program behaves as in some function application of an undefined case. \ text %mlex \ We define a total version of division: any failures are swept under the carpet and mapped to a default value. Thus division-by-zero becomes 0, but there could be other exceptions like overflow that produce the same result (for unbounded integers this does not happen). \ ML \ fun div_total x y = \<^try>\x div y\ |> the_default 0; \<^assert> (div_total 1 0 = 0); \ text \ The ML function \<^ML>\undefined\ is defined in \<^file>\~~/src/Pure/library.ML\ as follows: \ ML \fun undefined _ = raise Match\ text \ \<^medskip> The following variant uses the antiquotation @{ML_antiquotation undefined} instead: \ ML \fun undefined _ = \<^undefined>\ text \ \<^medskip> Here is the same, using control-symbol notation for the antiquotation, with special rendering of \<^verbatim>\\<^undefined>\: \ ML \fun undefined _ = \<^undefined>\ text \ \<^medskip> Semantically, all forms are equivalent to a function definition without any clauses, but that is syntactically not allowed in ML. \ section \Strings of symbols \label{sec:symbols}\ text \ A \<^emph>\symbol\ constitutes the smallest textual unit in Isabelle/ML --- raw ML characters are normally not encountered at all. Isabelle strings consist of a sequence of symbols, represented as a packed string or an exploded list of strings. Each symbol is in itself a small string, which has either one of the following forms: \<^enum> a single ASCII character ``\c\'', for example ``\<^verbatim>\a\'', \<^enum> a codepoint according to UTF-8 (non-ASCII byte sequence), \<^enum> a regular symbol ``\<^verbatim>\\\'', for example ``\<^verbatim>\\\'', \<^enum> a control symbol ``\<^verbatim>\\<^ident>\'', for example ``\<^verbatim>\\<^bold>\'', The \ident\ syntax for symbol names is \letter (letter | digit)\<^sup>*\, where \letter = A..Za..z\ and \digit = 0..9\. There are infinitely many regular symbols and control symbols, but a fixed collection of standard symbols is treated specifically. For example, ``\<^verbatim>\\\'' is classified as a letter, which means it may occur within regular Isabelle identifiers. The character set underlying Isabelle symbols is 7-bit ASCII, but 8-bit character sequences are passed-through unchanged. Unicode/UCS data in UTF-8 encoding is processed in a non-strict fashion, such that well-formed code sequences are recognized accordingly. Unicode provides its own collection of mathematical symbols, but within the core Isabelle/ML world there is no link to the standard collection of Isabelle regular symbols. \<^medskip> Output of Isabelle symbols depends on the print mode. For example, the standard {\LaTeX} setup of the Isabelle document preparation system would present ``\<^verbatim>\\\'' as \\\, and ``\<^verbatim>\\<^bold>\\'' as \\<^bold>\\. On-screen rendering usually works by mapping a finite subset of Isabelle symbols to suitable Unicode characters. \ text %mlref \ \begin{mldecls} @{define_ML_type Symbol.symbol = string} \\ @{define_ML Symbol.explode: "string -> Symbol.symbol list"} \\ @{define_ML Symbol.is_letter: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_digit: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_quasi: "Symbol.symbol -> bool"} \\ @{define_ML Symbol.is_blank: "Symbol.symbol -> bool"} \\ \end{mldecls} \begin{mldecls} @{define_ML_type "Symbol.sym"} \\ @{define_ML Symbol.decode: "Symbol.symbol -> Symbol.sym"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Symbol.symbol\ represents individual Isabelle symbols. \<^descr> \<^ML>\Symbol.explode\~\str\ produces a symbol list from the packed form. This function supersedes \<^ML>\String.explode\ for virtually all purposes of manipulating text in Isabelle!\<^footnote>\The runtime overhead for exploded strings is mainly that of the list structure: individual symbols that happen to be a singleton string do not require extra memory in Poly/ML.\ \<^descr> \<^ML>\Symbol.is_letter\, \<^ML>\Symbol.is_digit\, \<^ML>\Symbol.is_quasi\, \<^ML>\Symbol.is_blank\ classify standard symbols according to fixed syntactic conventions of Isabelle, cf.\ @{cite "isabelle-isar-ref"}. \<^descr> Type \<^ML_type>\Symbol.sym\ is a concrete datatype that represents the different kinds of symbols explicitly, with constructors \<^ML>\Symbol.Char\, \<^ML>\Symbol.UTF8\, \<^ML>\Symbol.Sym\, \<^ML>\Symbol.Control\, \<^ML>\Symbol.Malformed\. \<^descr> \<^ML>\Symbol.decode\ converts the string representation of a symbol into the datatype version. \ paragraph \Historical note.\ text \ In the original SML90 standard the primitive ML type \<^ML_type>\char\ did not exists, and \<^ML_text>\explode: string -> string list\ produced a list of singleton strings like \<^ML>\raw_explode: string -> string list\ in Isabelle/ML today. When SML97 came out, Isabelle did not adopt its somewhat anachronistic 8-bit or 16-bit characters, but the idea of exploding a string into a list of small strings was extended to ``symbols'' as explained above. Thus Isabelle sources can refer to an infinite store of user-defined symbols, without having to worry about the multitude of Unicode encodings that have emerged over the years. \ section \Basic data types\ text \ The basis library proposal of SML97 needs to be treated with caution. Many of its operations simply do not fit with important Isabelle/ML conventions (like ``canonical argument order'', see \secref{sec:canonical-argument-order}), others cause problems with the parallel evaluation model of Isabelle/ML (such as \<^ML>\TextIO.print\ or \<^ML>\OS.Process.system\). Subsequently we give a brief overview of important operations on basic ML data types. \ subsection \Characters\ text %mlref \ \begin{mldecls} @{define_ML_type char} \\ \end{mldecls} \<^descr> Type \<^ML_type>\char\ is \<^emph>\not\ used. The smallest textual unit in Isabelle is represented as a ``symbol'' (see \secref{sec:symbols}). \ subsection \Strings\ text %mlref \ \begin{mldecls} @{define_ML_type string} \\ \end{mldecls} \<^descr> Type \<^ML_type>\string\ represents immutable vectors of 8-bit characters. There are operations in SML to convert back and forth to actual byte vectors, which are seldom used. This historically important raw text representation is used for Isabelle-specific purposes with the following implicit substructures packed into the string content: \<^enum> sequence of Isabelle symbols (see also \secref{sec:symbols}), with \<^ML>\Symbol.explode\ as key operation; \<^enum> XML tree structure via YXML (see also @{cite "isabelle-system"}), with \<^ML>\YXML.parse_body\ as key operation. Note that Isabelle/ML string literals may refer Isabelle symbols like ``\<^verbatim>\\\'' natively, \<^emph>\without\ escaping the backslash. This is a consequence of Isabelle treating all source text as strings of symbols, instead of raw characters. \ text %mlex \ The subsequent example illustrates the difference of physical addressing of bytes versus logical addressing of symbols in Isabelle strings. \ ML_val \ val s = "\"; \<^assert> (length (Symbol.explode s) = 1); \<^assert> (size s = 4); \ text \ Note that in Unicode renderings of the symbol \\\, variations of encodings like UTF-8 or UTF-16 pose delicate questions about the multi-byte representations of its codepoint, which is outside of the 16-bit address space of the original Unicode standard from the 1990-ies. In Isabelle/ML it is just ``\<^verbatim>\\\'' literally, using plain ASCII characters beyond any doubts. \ subsection \Integers\ text %mlref \ \begin{mldecls} @{define_ML_type int} \\ \end{mldecls} \<^descr> Type \<^ML_type>\int\ represents regular mathematical integers, which are \<^emph>\unbounded\. Overflow is treated properly, but should never happen in practice.\<^footnote>\The size limit for integer bit patterns in memory is 64\,MB for 32-bit Poly/ML, and much higher for 64-bit systems.\ Structure \<^ML_structure>\IntInf\ of SML97 is obsolete and superseded by \<^ML_structure>\Int\. Structure \<^ML_structure>\Integer\ in \<^file>\~~/src/Pure/General/integer.ML\ provides some additional operations. \ subsection \Rational numbers\ text %mlref \ \begin{mldecls} @{define_ML_type Rat.rat} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Rat.rat\ represents rational numbers, based on the unbounded integers of Poly/ML. Literal rationals may be written with special antiquotation syntax \<^verbatim>\@\\int\\<^verbatim>\/\\nat\ or \<^verbatim>\@\\int\ (without any white space). For example \<^verbatim>\@~1/4\ or \<^verbatim>\@10\. The ML toplevel pretty printer uses the same format. Standard operations are provided via ad-hoc overloading of \<^verbatim>\+\, \<^verbatim>\-\, \<^verbatim>\*\, \<^verbatim>\/\, etc. \ subsection \Time\ text %mlref \ \begin{mldecls} @{define_ML_type Time.time} \\ @{define_ML seconds: "real -> Time.time"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\Time.time\ represents time abstractly according to the SML97 basis library definition. This is adequate for internal ML operations, but awkward in concrete time specifications. \<^descr> \<^ML>\seconds\~\s\ turns the concrete scalar \s\ (measured in seconds) into an abstract time value. Floating point numbers are easy to use as configuration options in the context (see \secref{sec:config-options}) or system options that are maintained externally. \ subsection \Options\ text %mlref \ \begin{mldecls} @{define_ML Option.map: "('a -> 'b) -> 'a option -> 'b option"} \\ @{define_ML is_some: "'a option -> bool"} \\ @{define_ML is_none: "'a option -> bool"} \\ @{define_ML the: "'a option -> 'a"} \\ @{define_ML these: "'a list option -> 'a list"} \\ @{define_ML the_list: "'a option -> 'a list"} \\ @{define_ML the_default: "'a -> 'a option -> 'a"} \\ \end{mldecls} \ text \ Apart from \<^ML>\Option.map\ most other operations defined in structure \<^ML_structure>\Option\ are alien to Isabelle/ML and never used. The operations shown above are defined in \<^file>\~~/src/Pure/General/basics.ML\. \ subsection \Lists\ text \ Lists are ubiquitous in ML as simple and light-weight ``collections'' for many everyday programming tasks. Isabelle/ML provides important additions and improvements over operations that are predefined in the SML97 library. \ text %mlref \ \begin{mldecls} @{define_ML cons: "'a -> 'a list -> 'a list"} \\ @{define_ML member: "('b * 'a -> bool) -> 'a list -> 'b -> bool"} \\ @{define_ML insert: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ @{define_ML remove: "('b * 'a -> bool) -> 'b -> 'a list -> 'a list"} \\ @{define_ML update: "('a * 'a -> bool) -> 'a -> 'a list -> 'a list"} \\ \end{mldecls} \<^descr> \<^ML>\cons\~\x xs\ evaluates to \x :: xs\. Tupled infix operators are a historical accident in Standard ML. The curried \<^ML>\cons\ amends this, but it should be only used when partial application is required. \<^descr> \<^ML>\member\, \<^ML>\insert\, \<^ML>\remove\, \<^ML>\update\ treat lists as a set-like container that maintains the order of elements. See \<^file>\~~/src/Pure/library.ML\ for the full specifications (written in ML). There are some further derived operations like \<^ML>\union\ or \<^ML>\inter\. Note that \<^ML>\insert\ is conservative about elements that are already a \<^ML>\member\ of the list, while \<^ML>\update\ ensures that the latest entry is always put in front. The latter discipline is often more appropriate in declarations of context data (\secref{sec:context-data}) that are issued by the user in Isar source: later declarations take precedence over earlier ones. \ text %mlex \ Using canonical \<^ML>\fold\ together with \<^ML>\cons\ (or similar standard operations) alternates the orientation of data. The is quite natural and should not be altered forcible by inserting extra applications of \<^ML>\rev\. The alternative \<^ML>\fold_rev\ can be used in the few situations, where alternation should be prevented. \ ML_val \ val items = [1, 2, 3, 4, 5, 6, 7, 8, 9, 10]; val list1 = fold cons items []; \<^assert> (list1 = rev items); val list2 = fold_rev cons items []; \<^assert> (list2 = items); \ text \ The subsequent example demonstrates how to \<^emph>\merge\ two lists in a natural way. \ ML_val \ fun merge_lists eq (xs, ys) = fold_rev (insert eq) ys xs; \ text \ Here the first list is treated conservatively: only the new elements from the second list are inserted. The inside-out order of insertion via \<^ML>\fold_rev\ attempts to preserve the order of elements in the result. This way of merging lists is typical for context data (\secref{sec:context-data}). See also \<^ML>\merge\ as defined in \<^file>\~~/src/Pure/library.ML\. \ subsection \Association lists\ text \ The operations for association lists interpret a concrete list of pairs as a finite function from keys to values. Redundant representations with multiple occurrences of the same key are implicitly normalized: lookup and update only take the first occurrence into account. \ text \ \begin{mldecls} @{define_ML AList.lookup: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> 'c option"} \\ @{define_ML AList.defined: "('a * 'b -> bool) -> ('b * 'c) list -> 'a -> bool"} \\ @{define_ML AList.update: "('a * 'a -> bool) -> 'a * 'b -> ('a * 'b) list -> ('a * 'b) list"} \\ \end{mldecls} \<^descr> \<^ML>\AList.lookup\, \<^ML>\AList.defined\, \<^ML>\AList.update\ implement the main ``framework operations'' for mappings in Isabelle/ML, following standard conventions for their names and types. Note that a function called \<^verbatim>\lookup\ is obliged to express its partiality via an explicit option element. There is no choice to raise an exception, without changing the name to something like \the_element\ or \get\. The \defined\ operation is essentially a contraction of \<^ML>\is_some\ and \<^verbatim>\lookup\, but this is sufficiently frequent to justify its independent existence. This also gives the implementation some opportunity for peep-hole optimization. Association lists are adequate as simple implementation of finite mappings in many practical situations. A more advanced table structure is defined in \<^file>\~~/src/Pure/General/table.ML\; that version scales easily to thousands or millions of elements. \ subsection \Unsynchronized references\ text %mlref \ \begin{mldecls} @{define_ML_type "'a Unsynchronized.ref"} \\ @{define_ML Unsynchronized.ref: "'a -> 'a Unsynchronized.ref"} \\ @{define_ML "!": "'a Unsynchronized.ref -> 'a"} \\ @{define_ML_infix ":=" : "'a Unsynchronized.ref * 'a -> unit"} \\ \end{mldecls} \ text \ Due to ubiquitous parallelism in Isabelle/ML (see also \secref{sec:multi-threading}), the mutable reference cells of Standard ML are notorious for causing problems. In a highly parallel system, both correctness \<^emph>\and\ performance are easily degraded when using mutable data. The unwieldy name of \<^ML>\Unsynchronized.ref\ for the constructor for references in Isabelle/ML emphasizes the inconveniences caused by - mutability. Existing operations \<^ML>\!\ and \<^ML_op>\:=\ are unchanged, + mutability. Existing operations \<^ML>\!\ and \<^ML_infix>\:=\ are unchanged, but should be used with special precautions, say in a strictly local situation that is guaranteed to be restricted to sequential evaluation --- now and in the future. \begin{warn} Never \<^ML_text>\open Unsynchronized\, not even in a local scope! Pretending that mutable state is no problem is a very bad idea. \end{warn} \ section \Thread-safe programming \label{sec:multi-threading}\ text \ Multi-threaded execution has become an everyday reality in Isabelle since Poly/ML 5.2.1 and Isabelle2008. Isabelle/ML provides implicit and explicit parallelism by default, and there is no way for user-space tools to ``opt out''. ML programs that are purely functional, output messages only via the official channels (\secref{sec:message-channels}), and do not intercept interrupts (\secref{sec:exceptions}) can participate in the multi-threaded environment immediately without further ado. More ambitious tools with more fine-grained interaction with the environment need to observe the principles explained below. \ subsection \Multi-threading with shared memory\ text \ Multiple threads help to organize advanced operations of the system, such as real-time conditions on command transactions, sub-components with explicit communication, general asynchronous interaction etc. Moreover, parallel evaluation is a prerequisite to make adequate use of the CPU resources that are available on multi-core systems.\<^footnote>\Multi-core computing does not mean that there are ``spare cycles'' to be wasted. It means that the continued exponential speedup of CPU performance due to ``Moore's Law'' follows different rules: clock frequency has reached its peak around 2005, and applications need to be parallelized in order to avoid a perceived loss of performance. See also @{cite "Sutter:2005"}.\ Isabelle/Isar exploits the inherent structure of theories and proofs to support \<^emph>\implicit parallelism\ to a large extent. LCF-style theorem proving provides almost ideal conditions for that, see also @{cite "Wenzel:2009"}. This means, significant parts of theory and proof checking is parallelized by default. In Isabelle2013, a maximum speedup-factor of 3.5 on 4 cores and 6.5 on 8 cores can be expected @{cite "Wenzel:2013:ITP"}. \<^medskip> ML threads lack the memory protection of separate processes, and operate concurrently on shared heap memory. This has the advantage that results of independent computations are directly available to other threads: abstract values can be passed without copying or awkward serialization that is typically required for separate processes. To make shared-memory multi-threading work robustly and efficiently, some programming guidelines need to be observed. While the ML system is responsible to maintain basic integrity of the representation of ML values in memory, the application programmer needs to ensure that multi-threaded execution does not break the intended semantics. \begin{warn} To participate in implicit parallelism, tools need to be thread-safe. A single ill-behaved tool can affect the stability and performance of the whole system. \end{warn} Apart from observing the principles of thread-safeness passively, advanced tools may also exploit parallelism actively, e.g.\ by using library functions for parallel list operations (\secref{sec:parlist}). \begin{warn} Parallel computing resources are managed centrally by the Isabelle/ML infrastructure. User programs should not fork their own ML threads to perform heavy computations. \end{warn} \ subsection \Critical shared resources\ text \ Thread-safeness is mainly concerned about concurrent read/write access to shared resources, which are outside the purely functional world of ML. This covers the following in particular. \<^item> Global references (or arrays), i.e.\ mutable memory cells that persist over several invocations of associated operations.\<^footnote>\This is independent of the visibility of such mutable values in the toplevel scope.\ \<^item> Global state of the running Isabelle/ML process, i.e.\ raw I/O channels, environment variables, current working directory. \<^item> Writable resources in the file-system that are shared among different threads or external processes. Isabelle/ML provides various mechanisms to avoid critical shared resources in most situations. As last resort there are some mechanisms for explicit synchronization. The following guidelines help to make Isabelle/ML programs work smoothly in a concurrent environment. \<^item> Avoid global references altogether. Isabelle/Isar maintains a uniform context that incorporates arbitrary data declared by user programs (\secref{sec:context-data}). This context is passed as plain value and user tools can get/map their own data in a purely functional manner. Configuration options within the context (\secref{sec:config-options}) provide simple drop-in replacements for historic reference variables. \<^item> Keep components with local state information re-entrant. Instead of poking initial values into (private) global references, a new state record can be created on each invocation, and passed through any auxiliary functions of the component. The state record contain mutable references in special situations, without requiring any synchronization, as long as each invocation gets its own copy and the tool itself is single-threaded. \<^item> Avoid raw output on \stdout\ or \stderr\. The Poly/ML library is thread-safe for each individual output operation, but the ordering of parallel invocations is arbitrary. This means raw output will appear on some system console with unpredictable interleaving of atomic chunks. Note that this does not affect regular message output channels (\secref{sec:message-channels}). An official message id is associated with the command transaction from where it originates, independently of other transactions. This means each running Isar command has effectively its own set of message channels, and interleaving can only happen when commands use parallelism internally (and only at message boundaries). \<^item> Treat environment variables and the current working directory of the running process as read-only. \<^item> Restrict writing to the file-system to unique temporary files. Isabelle already provides a temporary directory that is unique for the running process, and there is a centralized source of unique serial numbers in Isabelle/ML. Thus temporary files that are passed to to some external process will be always disjoint, and thus thread-safe. \ text %mlref \ \begin{mldecls} @{define_ML File.tmp_path: "Path.T -> Path.T"} \\ @{define_ML serial_string: "unit -> string"} \\ \end{mldecls} \<^descr> \<^ML>\File.tmp_path\~\path\ relocates the base component of \path\ into the unique temporary directory of the running Isabelle/ML process. \<^descr> \<^ML>\serial_string\~\()\ creates a new serial number that is unique over the runtime of the Isabelle/ML process. \ text %mlex \ The following example shows how to create unique temporary file names. \ ML_val \ val tmp1 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); val tmp2 = File.tmp_path (Path.basic ("foo" ^ serial_string ())); \<^assert> (tmp1 <> tmp2); \ subsection \Explicit synchronization\ text \ Isabelle/ML provides explicit synchronization for mutable variables over immutable data, which may be updated atomically and exclusively. This addresses the rare situations where mutable shared resources are really required. Synchronization in Isabelle/ML is based on primitives of Poly/ML, which have been adapted to the specific assumptions of the concurrent Isabelle environment. User code should not break this abstraction, but stay within the confines of concurrent Isabelle/ML. A \<^emph>\synchronized variable\ is an explicit state component associated with mechanisms for locking and signaling. There are operations to await a condition, change the state, and signal the change to all other waiting threads. Synchronized access to the state variable is \<^emph>\not\ re-entrant: direct or indirect nesting within the same thread will cause a deadlock! \ text %mlref \ \begin{mldecls} @{define_ML_type "'a Synchronized.var"} \\ @{define_ML Synchronized.var: "string -> 'a -> 'a Synchronized.var"} \\ @{define_ML Synchronized.guarded_access: "'a Synchronized.var -> ('a -> ('b * 'a) option) -> 'b"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a Synchronized.var\ represents synchronized variables with state of type \<^ML_type>\'a\. \<^descr> \<^ML>\Synchronized.var\~\name x\ creates a synchronized variable that is initialized with value \x\. The \name\ is used for tracing. \<^descr> \<^ML>\Synchronized.guarded_access\~\var f\ lets the function \f\ operate within a critical section on the state \x\ as follows: if \f x\ produces \<^ML>\NONE\, it continues to wait on the internal condition variable, expecting that some other thread will eventually change the content in a suitable manner; if \f x\ produces \<^ML>\SOME\~\(y, x')\ it is satisfied and assigns the new state value \x'\, broadcasts a signal to all waiting threads on the associated condition variable, and returns the result \y\. There are some further variants of the \<^ML>\Synchronized.guarded_access\ combinator, see \<^file>\~~/src/Pure/Concurrent/synchronized.ML\ for details. \ text %mlex \ The following example implements a counter that produces positive integers that are unique over the runtime of the Isabelle process: \ ML_val \ local val counter = Synchronized.var "counter" 0; in fun next () = Synchronized.guarded_access counter (fn i => let val j = i + 1 in SOME (j, j) end); end; val a = next (); val b = next (); \<^assert> (a <> b); \ text \ \<^medskip> See \<^file>\~~/src/Pure/Concurrent/mailbox.ML\ how to implement a mailbox as synchronized variable over a purely functional list. \ section \Managed evaluation\ text \ Execution of Standard ML follows the model of strict functional evaluation with optional exceptions. Evaluation happens whenever some function is applied to (sufficiently many) arguments. The result is either an explicit value or an implicit exception. \<^emph>\Managed evaluation\ in Isabelle/ML organizes expressions and results to control certain physical side-conditions, to say more specifically when and how evaluation happens. For example, the Isabelle/ML library supports lazy evaluation with memoing, parallel evaluation via futures, asynchronous evaluation via promises, evaluation with time limit etc. \<^medskip> An \<^emph>\unevaluated expression\ is represented either as unit abstraction \<^verbatim>\fn () => a\ of type \<^verbatim>\unit -> 'a\ or as regular function \<^verbatim>\fn a => b\ of type \<^verbatim>\'a -> 'b\. Both forms occur routinely, and special care is required to tell them apart --- the static type-system of SML is only of limited help here. The first form is more intuitive: some combinator \<^verbatim>\(unit -> 'a) -> 'a\ applies the given function to \<^verbatim>\()\ to initiate the postponed evaluation process. The second form is more flexible: some combinator \<^verbatim>\('a -> 'b) -> 'a -> 'b\ acts like a modified form of function application; several such combinators may be cascaded to modify a given function, before it is ultimately applied to some argument. \<^medskip> \<^emph>\Reified results\ make the disjoint sum of regular values versions exceptional situations explicit as ML datatype: \'a result = Res of 'a | Exn of exn\. This is typically used for administrative purposes, to store the overall outcome of an evaluation process. \<^emph>\Parallel exceptions\ aggregate reified results, such that multiple exceptions are digested as a collection in canonical form that identifies exceptions according to their original occurrence. This is particular important for parallel evaluation via futures \secref{sec:futures}, which are organized as acyclic graph of evaluations that depend on other evaluations: exceptions stemming from shared sub-graphs are exposed exactly once and in the order of their original occurrence (e.g.\ when printed at the toplevel). Interrupt counts as neutral element here: it is treated as minimal information about some canceled evaluation process, and is absorbed by the presence of regular program exceptions. \ text %mlref \ \begin{mldecls} @{define_ML_type "'a Exn.result"} \\ @{define_ML Exn.capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\ @{define_ML Exn.interruptible_capture: "('a -> 'b) -> 'a -> 'b Exn.result"} \\ @{define_ML Exn.release: "'a Exn.result -> 'a"} \\ @{define_ML Par_Exn.release_all: "'a Exn.result list -> 'a list"} \\ @{define_ML Par_Exn.release_first: "'a Exn.result list -> 'a list"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a Exn.result\ represents the disjoint sum of ML results explicitly, with constructor \<^ML>\Exn.Res\ for regular values and \<^ML>\Exn.Exn\ for exceptions. \<^descr> \<^ML>\Exn.capture\~\f x\ manages the evaluation of \f x\ such that exceptions are made explicit as \<^ML>\Exn.Exn\. Note that this includes physical interrupts (see also \secref{sec:exceptions}), so the same precautions apply to user code: interrupts must not be absorbed accidentally! \<^descr> \<^ML>\Exn.interruptible_capture\ is similar to \<^ML>\Exn.capture\, but interrupts are immediately re-raised as required for user code. \<^descr> \<^ML>\Exn.release\~\result\ releases the original runtime result, exposing its regular value or raising the reified exception. \<^descr> \<^ML>\Par_Exn.release_all\~\results\ combines results that were produced independently (e.g.\ by parallel evaluation). If all results are regular values, that list is returned. Otherwise, the collection of all exceptions is raised, wrapped-up as collective parallel exception. Note that the latter prevents access to individual exceptions by conventional \<^verbatim>\handle\ of ML. \<^descr> \<^ML>\Par_Exn.release_first\ is similar to \<^ML>\Par_Exn.release_all\, but only the first (meaningful) exception that has occurred in the original evaluation process is raised again, the others are ignored. That single exception may get handled by conventional means in ML. \ subsection \Parallel skeletons \label{sec:parlist}\ text \ Algorithmic skeletons are combinators that operate on lists in parallel, in the manner of well-known \map\, \exists\, \forall\ etc. Management of futures (\secref{sec:futures}) and their results as reified exceptions is wrapped up into simple programming interfaces that resemble the sequential versions. What remains is the application-specific problem to present expressions with suitable \<^emph>\granularity\: each list element corresponds to one evaluation task. If the granularity is too coarse, the available CPUs are not saturated. If it is too fine-grained, CPU cycles are wasted due to the overhead of organizing parallel processing. In the worst case, parallel performance will be less than the sequential counterpart! \ text %mlref \ \begin{mldecls} @{define_ML Par_List.map: "('a -> 'b) -> 'a list -> 'b list"} \\ @{define_ML Par_List.get_some: "('a -> 'b option) -> 'a list -> 'b option"} \\ \end{mldecls} \<^descr> \<^ML>\Par_List.map\~\f [x\<^sub>1, \, x\<^sub>n]\ is like \<^ML>\map\~\f [x\<^sub>1, \, x\<^sub>n]\, but the evaluation of \f x\<^sub>i\ for \i = 1, \, n\ is performed in parallel. An exception in any \f x\<^sub>i\ cancels the overall evaluation process. The final result is produced via \<^ML>\Par_Exn.release_first\ as explained above, which means the first program exception that happened to occur in the parallel evaluation is propagated, and all other failures are ignored. \<^descr> \<^ML>\Par_List.get_some\~\f [x\<^sub>1, \, x\<^sub>n]\ produces some \f x\<^sub>i\ that is of the form \SOME y\<^sub>i\, if that exists, otherwise \NONE\. Thus it is similar to \<^ML>\Library.get_first\, but subject to a non-deterministic parallel choice process. The first successful result cancels the overall evaluation process; other exceptions are propagated as for \<^ML>\Par_List.map\. This generic parallel choice combinator is the basis for derived forms, such as \<^ML>\Par_List.find_some\, \<^ML>\Par_List.exists\, \<^ML>\Par_List.forall\. \ text %mlex \ Subsequently, the Ackermann function is evaluated in parallel for some ranges of arguments. \ ML_val \ fun ackermann 0 n = n + 1 | ackermann m 0 = ackermann (m - 1) 1 | ackermann m n = ackermann (m - 1) (ackermann m (n - 1)); Par_List.map (ackermann 2) (500 upto 1000); Par_List.map (ackermann 3) (5 upto 10); \ subsection \Lazy evaluation\ text \ Classic lazy evaluation works via the \lazy\~/ \force\ pair of operations: \lazy\ to wrap an unevaluated expression, and \force\ to evaluate it once and store its result persistently. Later invocations of \force\ retrieve the stored result without another evaluation. Isabelle/ML refines this idea to accommodate the aspects of multi-threading, synchronous program exceptions and asynchronous interrupts. The first thread that invokes \force\ on an unfinished lazy value changes its state into a \<^emph>\promise\ of the eventual result and starts evaluating it. Any other threads that \force\ the same lazy value in the meantime need to wait for it to finish, by producing a regular result or program exception. If the evaluation attempt is interrupted, this event is propagated to all waiting threads and the lazy value is reset to its original state. This means a lazy value is completely evaluated at most once, in a thread-safe manner. There might be multiple interrupted evaluation attempts, and multiple receivers of intermediate interrupt events. Interrupts are \<^emph>\not\ made persistent: later evaluation attempts start again from the original expression. \ text %mlref \ \begin{mldecls} @{define_ML_type "'a lazy"} \\ @{define_ML Lazy.lazy: "(unit -> 'a) -> 'a lazy"} \\ @{define_ML Lazy.value: "'a -> 'a lazy"} \\ @{define_ML Lazy.force: "'a lazy -> 'a"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a lazy\ represents lazy values over type \<^verbatim>\'a\. \<^descr> \<^ML>\Lazy.lazy\~\(fn () => e)\ wraps the unevaluated expression \e\ as unfinished lazy value. \<^descr> \<^ML>\Lazy.value\~\a\ wraps the value \a\ as finished lazy value. When forced, it returns \a\ without any further evaluation. There is very low overhead for this proforma wrapping of strict values as lazy values. \<^descr> \<^ML>\Lazy.force\~\x\ produces the result of the lazy value in a thread-safe manner as explained above. Thus it may cause the current thread to wait on a pending evaluation attempt by another thread. \ subsection \Futures \label{sec:futures}\ text \ Futures help to organize parallel execution in a value-oriented manner, with \fork\~/ \join\ as the main pair of operations, and some further variants; see also @{cite "Wenzel:2009" and "Wenzel:2013:ITP"}. Unlike lazy values, futures are evaluated strictly and spontaneously on separate worker threads. Futures may be canceled, which leads to interrupts on running evaluation attempts, and forces structurally related futures to fail for all time; already finished futures remain unchanged. Exceptions between related futures are propagated as well, and turned into parallel exceptions (see above). Technically, a future is a single-assignment variable together with a \<^emph>\task\ that serves administrative purposes, notably within the \<^emph>\task queue\ where new futures are registered for eventual evaluation and the worker threads retrieve their work. The pool of worker threads is limited, in correlation with the number of physical cores on the machine. Note that allocation of runtime resources may be distorted either if workers yield CPU time (e.g.\ via system sleep or wait operations), or if non-worker threads contend for significant runtime resources independently. There is a limited number of replacement worker threads that get activated in certain explicit wait conditions, after a timeout. \<^medskip> Each future task belongs to some \<^emph>\task group\, which represents the hierarchic structure of related tasks, together with the exception status a that point. By default, the task group of a newly created future is a new sub-group of the presently running one, but it is also possible to indicate different group layouts under program control. Cancellation of futures actually refers to the corresponding task group and all its sub-groups. Thus interrupts are propagated down the group hierarchy. Regular program exceptions are treated likewise: failure of the evaluation of some future task affects its own group and all sub-groups. Given a particular task group, its \<^emph>\group status\ cumulates all relevant exceptions according to its position within the group hierarchy. Interrupted tasks that lack regular result information, will pick up parallel exceptions from the cumulative group status. \<^medskip> A \<^emph>\passive future\ or \<^emph>\promise\ is a future with slightly different evaluation policies: there is only a single-assignment variable and some expression to evaluate for the \<^emph>\failed\ case (e.g.\ to clean up resources when canceled). A regular result is produced by external means, using a separate \<^emph>\fulfill\ operation. Promises are managed in the same task queue, so regular futures may depend on them. This allows a form of reactive programming, where some promises are used as minimal elements (or guards) within the future dependency graph: when these promises are fulfilled the evaluation of subsequent futures starts spontaneously, according to their own inter-dependencies. \ text %mlref \ \begin{mldecls} @{define_ML_type "'a future"} \\ @{define_ML Future.fork: "(unit -> 'a) -> 'a future"} \\ @{define_ML Future.forks: "Future.params -> (unit -> 'a) list -> 'a future list"} \\ @{define_ML Future.join: "'a future -> 'a"} \\ @{define_ML Future.joins: "'a future list -> 'a list"} \\ @{define_ML Future.value: "'a -> 'a future"} \\ @{define_ML Future.map: "('a -> 'b) -> 'a future -> 'b future"} \\ @{define_ML Future.cancel: "'a future -> unit"} \\ @{define_ML Future.cancel_group: "Future.group -> unit"} \\[0.5ex] @{define_ML Future.promise: "(unit -> unit) -> 'a future"} \\ @{define_ML Future.fulfill: "'a future -> 'a -> unit"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\'a future\ represents future values over type \<^verbatim>\'a\. \<^descr> \<^ML>\Future.fork\~\(fn () => e)\ registers the unevaluated expression \e\ as unfinished future value, to be evaluated eventually on the parallel worker-thread farm. This is a shorthand for \<^ML>\Future.forks\ below, with default parameters and a single expression. \<^descr> \<^ML>\Future.forks\~\params exprs\ is the general interface to fork several futures simultaneously. The \params\ consist of the following fields: \<^item> \name : string\ (default \<^ML>\""\) specifies a common name for the tasks of the forked futures, which serves diagnostic purposes. \<^item> \group : Future.group option\ (default \<^ML>\NONE\) specifies an optional task group for the forked futures. \<^ML>\NONE\ means that a new sub-group of the current worker-thread task context is created. If this is not a worker thread, the group will be a new root in the group hierarchy. \<^item> \deps : Future.task list\ (default \<^ML>\[]\) specifies dependencies on other future tasks, i.e.\ the adjacency relation in the global task queue. Dependencies on already finished tasks are ignored. \<^item> \pri : int\ (default \<^ML>\0\) specifies a priority within the task queue. Typically there is only little deviation from the default priority \<^ML>\0\. As a rule of thumb, \<^ML>\~1\ means ``low priority" and \<^ML>\1\ means ``high priority''. Note that the task priority only affects the position in the queue, not the thread priority. When a worker thread picks up a task for processing, it runs with the normal thread priority to the end (or until canceled). Higher priority tasks that are queued later need to wait until this (or another) worker thread becomes free again. \<^item> \interrupts : bool\ (default \<^ML>\true\) tells whether the worker thread that processes the corresponding task is initially put into interruptible state. This state may change again while running, by modifying the thread attributes. With interrupts disabled, a running future task cannot be canceled. It is the responsibility of the programmer that this special state is retained only briefly. \<^descr> \<^ML>\Future.join\~\x\ retrieves the value of an already finished future, which may lead to an exception, according to the result of its previous evaluation. For an unfinished future there are several cases depending on the role of the current thread and the status of the future. A non-worker thread waits passively until the future is eventually evaluated. A worker thread temporarily changes its task context and takes over the responsibility to evaluate the future expression on the spot. The latter is done in a thread-safe manner: other threads that intend to join the same future need to wait until the ongoing evaluation is finished. Note that excessive use of dynamic dependencies of futures by adhoc joining may lead to bad utilization of CPU cores, due to threads waiting on other threads to finish required futures. The future task farm has a limited amount of replacement threads that continue working on unrelated tasks after some timeout. Whenever possible, static dependencies of futures should be specified explicitly when forked (see \deps\ above). Thus the evaluation can work from the bottom up, without join conflicts and wait states. \<^descr> \<^ML>\Future.joins\~\xs\ joins the given list of futures simultaneously, which is more efficient than \<^ML>\map Future.join\~\xs\. Based on the dependency graph of tasks, the current thread takes over the responsibility to evaluate future expressions that are required for the main result, working from the bottom up. Waiting on future results that are presently evaluated on other threads only happens as last resort, when no other unfinished futures are left over. \<^descr> \<^ML>\Future.value\~\a\ wraps the value \a\ as finished future value, bypassing the worker-thread farm. When joined, it returns \a\ without any further evaluation. There is very low overhead for this proforma wrapping of strict values as futures. \<^descr> \<^ML>\Future.map\~\f x\ is a fast-path implementation of \<^ML>\Future.fork\~\(fn () => f (\\<^ML>\Future.join\~\x))\, which avoids the full overhead of the task queue and worker-thread farm as far as possible. The function \f\ is supposed to be some trivial post-processing or projection of the future result. \<^descr> \<^ML>\Future.cancel\~\x\ cancels the task group of the given future, using \<^ML>\Future.cancel_group\ below. \<^descr> \<^ML>\Future.cancel_group\~\group\ cancels all tasks of the given task group for all time. Threads that are presently processing a task of the given group are interrupted: it may take some time until they are actually terminated. Tasks that are queued but not yet processed are dequeued and forced into interrupted state. Since the task group is itself invalidated, any further attempt to fork a future that belongs to it will yield a canceled result as well. \<^descr> \<^ML>\Future.promise\~\abort\ registers a passive future with the given \abort\ operation: it is invoked when the future task group is canceled. \<^descr> \<^ML>\Future.fulfill\~\x a\ finishes the passive future \x\ by the given value \a\. If the promise has already been canceled, the attempt to fulfill it causes an exception. \ end diff --git a/src/Doc/Implementation/Tactic.thy b/src/Doc/Implementation/Tactic.thy --- a/src/Doc/Implementation/Tactic.thy +++ b/src/Doc/Implementation/Tactic.thy @@ -1,817 +1,817 @@ (*:maxLineLen=78:*) theory Tactic imports Base begin chapter \Tactical reasoning\ text \ Tactical reasoning works by refining an initial claim in a backwards fashion, until a solved form is reached. A \goal\ consists of several subgoals that need to be solved in order to achieve the main statement; zero subgoals means that the proof may be finished. A \tactic\ is a refinement operation that maps a goal to a lazy sequence of potential successors. A \tactical\ is a combinator for composing tactics. \ section \Goals \label{sec:tactical-goals}\ text \ Isabelle/Pure represents a goal as a theorem stating that the subgoals imply the main goal: \A\<^sub>1 \ \ \ A\<^sub>n \ C\. The outermost goal structure is that of a Horn Clause: i.e.\ an iterated implication without any quantifiers\<^footnote>\Recall that outermost \\x. \[x]\ is always represented via schematic variables in the body: \\[?x]\. These variables may get instantiated during the course of reasoning.\. For \n = 0\ a goal is called ``solved''. The structure of each subgoal \A\<^sub>i\ is that of a general Hereditary Harrop Formula \\x\<^sub>1 \ \x\<^sub>k. H\<^sub>1 \ \ \ H\<^sub>m \ B\. Here \x\<^sub>1, \, x\<^sub>k\ are goal parameters, i.e.\ arbitrary-but-fixed entities of certain types, and \H\<^sub>1, \, H\<^sub>m\ are goal hypotheses, i.e.\ facts that may be assumed locally. Together, this forms the goal context of the conclusion \B\ to be established. The goal hypotheses may be again arbitrary Hereditary Harrop Formulas, although the level of nesting rarely exceeds 1--2 in practice. The main conclusion \C\ is internally marked as a protected proposition, which is represented explicitly by the notation \#C\ here. This ensures that the decomposition into subgoals and main conclusion is well-defined for arbitrarily structured claims. \<^medskip> Basic goal management is performed via the following Isabelle/Pure rules: \[ \infer[\(init)\]{\C \ #C\}{} \qquad \infer[\(finish)\]{\C\}{\#C\} \] \<^medskip> The following low-level variants admit general reasoning with protected propositions: \[ \infer[\(protect n)\]{\A\<^sub>1 \ \ \ A\<^sub>n \ #C\}{\A\<^sub>1 \ \ \ A\<^sub>n \ C\} \] \[ \infer[\(conclude)\]{\A \ \ \ C\}{\A \ \ \ #C\} \] \ text %mlref \ \begin{mldecls} @{define_ML Goal.init: "cterm -> thm"} \\ @{define_ML Goal.finish: "Proof.context -> thm -> thm"} \\ @{define_ML Goal.protect: "int -> thm -> thm"} \\ @{define_ML Goal.conclude: "thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\Goal.init\~\C\ initializes a tactical goal from the well-formed proposition \C\. \<^descr> \<^ML>\Goal.finish\~\ctxt thm\ checks whether theorem \thm\ is a solved goal (no subgoals), and concludes the result by removing the goal protection. The context is only required for printing error messages. \<^descr> \<^ML>\Goal.protect\~\n thm\ protects the statement of theorem \thm\. The parameter \n\ indicates the number of premises to be retained. \<^descr> \<^ML>\Goal.conclude\~\thm\ removes the goal protection, even if there are pending subgoals. \ section \Tactics\label{sec:tactics}\ text \ A \tactic\ is a function \goal \ goal\<^sup>*\<^sup>*\ that maps a given goal state (represented as a theorem, cf.\ \secref{sec:tactical-goals}) to a lazy sequence of potential successor states. The underlying sequence implementation is lazy both in head and tail, and is purely functional in \<^emph>\not\ supporting memoing.\<^footnote>\The lack of memoing and the strict nature of ML requires some care when working with low-level sequence operations, to avoid duplicate or premature evaluation of results. It also means that modified runtime behavior, such as timeout, is very hard to achieve for general tactics.\ An \<^emph>\empty result sequence\ means that the tactic has failed: in a compound tactic expression other tactics might be tried instead, or the whole refinement step might fail outright, producing a toplevel error message in the end. When implementing tactics from scratch, one should take care to observe the basic protocol of mapping regular error conditions to an empty result; only serious faults should emerge as exceptions. By enumerating \<^emph>\multiple results\, a tactic can easily express the potential outcome of an internal search process. There are also combinators for building proof tools that involve search systematically, see also \secref{sec:tacticals}. \<^medskip> As explained before, a goal state essentially consists of a list of subgoals that imply the main goal (conclusion). Tactics may operate on all subgoals or on a particularly specified subgoal, but must not change the main conclusion (apart from instantiating schematic goal variables). Tactics with explicit \<^emph>\subgoal addressing\ are of the form \int \ tactic\ and may be applied to a particular subgoal (counting from 1). If the subgoal number is out of range, the tactic should fail with an empty result sequence, but must not raise an exception! Operating on a particular subgoal means to replace it by an interval of zero or more subgoals in the same place; other subgoals must not be affected, apart from instantiating schematic variables ranging over the whole goal state. A common pattern of composing tactics with subgoal addressing is to try the first one, and then the second one only if the subgoal has not been solved yet. Special care is required here to avoid bumping into unrelated subgoals that happen to come after the original subgoal. Assuming that there is only a single initial subgoal is a very common error when implementing tactics! Tactics with internal subgoal addressing should expose the subgoal index as \int\ argument in full generality; a hardwired subgoal 1 is not acceptable. \<^medskip> The main well-formedness conditions for proper tactics are summarized as follows. \<^item> General tactic failure is indicated by an empty result, only serious faults may produce an exception. \<^item> The main conclusion must not be changed, apart from instantiating schematic variables. \<^item> A tactic operates either uniformly on all subgoals, or specifically on a selected subgoal (without bumping into unrelated subgoals). \<^item> Range errors in subgoal addressing produce an empty result. Some of these conditions are checked by higher-level goal infrastructure (\secref{sec:struct-goals}); others are not checked explicitly, and violating them merely results in ill-behaved tactics experienced by the user (e.g.\ tactics that insist in being applicable only to singleton goals, or prevent composition via standard tacticals such as \<^ML>\REPEAT\). \ text %mlref \ \begin{mldecls} @{define_ML_type tactic = "thm -> thm Seq.seq"} \\ @{define_ML no_tac: tactic} \\ @{define_ML all_tac: tactic} \\ @{define_ML print_tac: "Proof.context -> string -> tactic"} \\[1ex] @{define_ML PRIMITIVE: "(thm -> thm) -> tactic"} \\[1ex] @{define_ML SUBGOAL: "(term * int -> tactic) -> int -> tactic"} \\ @{define_ML CSUBGOAL: "(cterm * int -> tactic) -> int -> tactic"} \\ @{define_ML SELECT_GOAL: "tactic -> int -> tactic"} \\ @{define_ML PREFER_GOAL: "tactic -> int -> tactic"} \\ \end{mldecls} \<^descr> Type \<^ML_type>\tactic\ represents tactics. The well-formedness conditions described above need to be observed. See also \<^file>\~~/src/Pure/General/seq.ML\ for the underlying implementation of lazy sequences. \<^descr> Type \<^ML_type>\int -> tactic\ represents tactics with explicit subgoal addressing, with well-formedness conditions as described above. \<^descr> \<^ML>\no_tac\ is a tactic that always fails, returning the empty sequence. \<^descr> \<^ML>\all_tac\ is a tactic that always succeeds, returning a singleton sequence with unchanged goal state. \<^descr> \<^ML>\print_tac\~\ctxt message\ is like \<^ML>\all_tac\, but prints a message together with the goal state on the tracing channel. \<^descr> \<^ML>\PRIMITIVE\~\rule\ turns a primitive inference rule into a tactic with unique result. Exception \<^ML>\THM\ is considered a regular tactic failure and produces an empty result; other exceptions are passed through. \<^descr> \<^ML>\SUBGOAL\~\(fn (subgoal, i) => tactic)\ is the most basic form to produce a tactic with subgoal addressing. The given abstraction over the subgoal term and subgoal number allows to peek at the relevant information of the full goal state. The subgoal range is checked as required above. \<^descr> \<^ML>\CSUBGOAL\ is similar to \<^ML>\SUBGOAL\, but passes the subgoal as \<^ML_type>\cterm\ instead of raw \<^ML_type>\term\. This avoids expensive re-certification in situations where the subgoal is used directly for primitive inferences. \<^descr> \<^ML>\SELECT_GOAL\~\tac i\ confines a tactic to the specified subgoal \i\. This rearranges subgoals and the main goal protection (\secref{sec:tactical-goals}), while retaining the syntactic context of the overall goal state (concerning schematic variables etc.). \<^descr> \<^ML>\PREFER_GOAL\~\tac i\ rearranges subgoals to put \i\ in front. This is similar to \<^ML>\SELECT_GOAL\, but without changing the main goal protection. \ subsection \Resolution and assumption tactics \label{sec:resolve-assume-tac}\ text \ \<^emph>\Resolution\ is the most basic mechanism for refining a subgoal using a theorem as object-level rule. \<^emph>\Elim-resolution\ is particularly suited for elimination rules: it resolves with a rule, proves its first premise by assumption, and finally deletes that assumption from any new subgoals. \<^emph>\Destruct-resolution\ is like elim-resolution, but the given destruction rules are first turned into canonical elimination format. \<^emph>\Forward-resolution\ is like destruct-resolution, but without deleting the selected assumption. The \r/e/d/f\ naming convention is maintained for several different kinds of resolution rules and tactics. Assumption tactics close a subgoal by unifying some of its premises against its conclusion. \<^medskip> All the tactics in this section operate on a subgoal designated by a positive integer. Other subgoals might be affected indirectly, due to instantiation of schematic variables. There are various sources of non-determinism, the tactic result sequence enumerates all possibilities of the following choices (if applicable): \<^enum> selecting one of the rules given as argument to the tactic; \<^enum> selecting a subgoal premise to eliminate, unifying it against the first premise of the rule; \<^enum> unifying the conclusion of the subgoal to the conclusion of the rule. Recall that higher-order unification may produce multiple results that are enumerated here. \ text %mlref \ \begin{mldecls} @{define_ML resolve_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML eresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML dresolve_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML forward_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML biresolve_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\[1ex] @{define_ML assume_tac: "Proof.context -> int -> tactic"} \\ @{define_ML eq_assume_tac: "int -> tactic"} \\[1ex] @{define_ML match_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML ematch_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML dmatch_tac: "Proof.context -> thm list -> int -> tactic"} \\ @{define_ML bimatch_tac: "Proof.context -> (bool * thm) list -> int -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\resolve_tac\~\ctxt thms i\ refines the goal state using the given theorems, which should normally be introduction rules. The tactic resolves a rule's conclusion with subgoal \i\, replacing it by the corresponding versions of the rule's premises. \<^descr> \<^ML>\eresolve_tac\~\ctxt thms i\ performs elim-resolution with the given theorems, which are normally be elimination rules. Note that \<^ML_text>\eresolve_tac ctxt [asm_rl]\ is equivalent to \<^ML_text>\assume_tac ctxt\, which facilitates mixing of assumption steps with genuine eliminations. \<^descr> \<^ML>\dresolve_tac\~\ctxt thms i\ performs destruct-resolution with the given theorems, which should normally be destruction rules. This replaces an assumption by the result of applying one of the rules. \<^descr> \<^ML>\forward_tac\ is like \<^ML>\dresolve_tac\ except that the selected assumption is not deleted. It applies a rule to an assumption, adding the result as a new assumption. \<^descr> \<^ML>\biresolve_tac\~\ctxt brls i\ refines the proof state by resolution or elim-resolution on each rule, as indicated by its flag. It affects subgoal \i\ of the proof state. For each pair \(flag, rule)\, it applies resolution if the flag is \false\ and elim-resolution if the flag is \true\. A single tactic call handles a mixture of introduction and elimination rules, which is useful to organize the search process systematically in proof tools. \<^descr> \<^ML>\assume_tac\~\ctxt i\ attempts to solve subgoal \i\ by assumption (modulo higher-order unification). \<^descr> \<^ML>\eq_assume_tac\ is similar to \<^ML>\assume_tac\, but checks only for immediate \\\-convertibility instead of using unification. It succeeds (with a unique next state) if one of the assumptions is equal to the subgoal's conclusion. Since it does not instantiate variables, it cannot make other subgoals unprovable. \<^descr> \<^ML>\match_tac\, \<^ML>\ematch_tac\, \<^ML>\dmatch_tac\, and \<^ML>\bimatch_tac\ are similar to \<^ML>\resolve_tac\, \<^ML>\eresolve_tac\, \<^ML>\dresolve_tac\, and \<^ML>\biresolve_tac\, respectively, but do not instantiate schematic variables in the goal state.\<^footnote>\Strictly speaking, matching means to treat the unknowns in the goal state as constants, but these tactics merely discard unifiers that would update the goal state. In rare situations (where the conclusion and goal state have flexible terms at the same position), the tactic will fail even though an acceptable unifier exists.\ These tactics were written for a specific application within the classical reasoner. Flexible subgoals are not updated at will, but are left alone. \ subsection \Explicit instantiation within a subgoal context\ text \ The main resolution tactics (\secref{sec:resolve-assume-tac}) use higher-order unification, which works well in many practical situations despite its daunting theoretical properties. Nonetheless, there are important problem classes where unguided higher-order unification is not so useful. This typically involves rules like universal elimination, existential introduction, or equational substitution. Here the unification problem involves fully flexible \?P ?x\ schemes, which are hard to manage without further hints. By providing a (small) rigid term for \?x\ explicitly, the remaining unification problem is to assign a (large) term to \?P\, according to the shape of the given subgoal. This is sufficiently well-behaved in most practical situations. \<^medskip> Isabelle provides separate versions of the standard \r/e/d/f\ resolution tactics that allow to provide explicit instantiations of unknowns of the given rule, wrt.\ terms that refer to the implicit context of the selected subgoal. An instantiation consists of a list of pairs of the form \(?x, t)\, where \?x\ is a schematic variable occurring in the given rule, and \t\ is a term from the current proof context, augmented by the local goal parameters of the selected subgoal; cf.\ the \focus\ operation described in \secref{sec:variables}. Entering the syntactic context of a subgoal is a brittle operation, because its exact form is somewhat accidental, and the choice of bound variable names depends on the presence of other local and global names. Explicit renaming of subgoal parameters prior to explicit instantiation might help to achieve a bit more robustness. Type instantiations may be given as well, via pairs like \(?'a, \)\. Type instantiations are distinguished from term instantiations by the syntactic form of the schematic variable. Types are instantiated before terms are. Since term instantiation already performs simple type-inference, so explicit type instantiations are seldom necessary. \ text %mlref \ \begin{mldecls} @{define_ML Rule_Insts.res_inst_tac: "Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic"} \\ @{define_ML Rule_Insts.eres_inst_tac: "Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic"} \\ @{define_ML Rule_Insts.dres_inst_tac: "Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic"} \\ @{define_ML Rule_Insts.forw_inst_tac: "Proof.context -> ((indexname * Position.T) * string) list -> (binding * string option * mixfix) list -> thm -> int -> tactic"} \\ @{define_ML Rule_Insts.subgoal_tac: "Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic"} \\ @{define_ML Rule_Insts.thin_tac: "Proof.context -> string -> (binding * string option * mixfix) list -> int -> tactic"} \\ @{define_ML rename_tac: "string list -> int -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\Rule_Insts.res_inst_tac\~\ctxt insts thm i\ instantiates the rule \thm\ with the instantiations \insts\, as described above, and then performs resolution on subgoal \i\. \<^descr> \<^ML>\Rule_Insts.eres_inst_tac\ is like \<^ML>\Rule_Insts.res_inst_tac\, but performs elim-resolution. \<^descr> \<^ML>\Rule_Insts.dres_inst_tac\ is like \<^ML>\Rule_Insts.res_inst_tac\, but performs destruct-resolution. \<^descr> \<^ML>\Rule_Insts.forw_inst_tac\ is like \<^ML>\Rule_Insts.dres_inst_tac\ except that the selected assumption is not deleted. \<^descr> \<^ML>\Rule_Insts.subgoal_tac\~\ctxt \ i\ adds the proposition \\\ as local premise to subgoal \i\, and poses the same as a new subgoal \i + 1\ (in the original context). \<^descr> \<^ML>\Rule_Insts.thin_tac\~\ctxt \ i\ deletes the specified premise from subgoal \i\. Note that \\\ may contain schematic variables, to abbreviate the intended proposition; the first matching subgoal premise will be deleted. Removing useless premises from a subgoal increases its readability and can make search tactics run faster. \<^descr> \<^ML>\rename_tac\~\names i\ renames the innermost parameters of subgoal \i\ according to the provided \names\ (which need to be distinct identifiers). For historical reasons, the above instantiation tactics take unparsed string arguments, which makes them hard to use in general ML code. The slightly more advanced \<^ML>\Subgoal.FOCUS\ combinator of \secref{sec:struct-goals} allows to refer to internal goal structure with explicit context management. \ subsection \Rearranging goal states\ text \ In rare situations there is a need to rearrange goal states: either the overall collection of subgoals, or the local structure of a subgoal. Various administrative tactics allow to operate on the concrete presentation these conceptual sets of formulae. \ text %mlref \ \begin{mldecls} @{define_ML rotate_tac: "int -> int -> tactic"} \\ @{define_ML distinct_subgoals_tac: tactic} \\ @{define_ML flexflex_tac: "Proof.context -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\rotate_tac\~\n i\ rotates the premises of subgoal \i\ by \n\ positions: from right to left if \n\ is positive, and from left to right if \n\ is negative. \<^descr> \<^ML>\distinct_subgoals_tac\ removes duplicate subgoals from a proof state. This is potentially inefficient. \<^descr> \<^ML>\flexflex_tac\ removes all flex-flex pairs from the proof state by applying the trivial unifier. This drastic step loses information. It is already part of the Isar infrastructure for facts resulting from goals, and rarely needs to be invoked manually. Flex-flex constraints arise from difficult cases of higher-order unification. To prevent this, use \<^ML>\Rule_Insts.res_inst_tac\ to instantiate some variables in a rule. Normally flex-flex constraints can be ignored; they often disappear as unknowns get instantiated. \ subsection \Raw composition: resolution without lifting\ text \ Raw composition of two rules means resolving them without prior lifting or renaming of unknowns. This low-level operation, which underlies the resolution tactics, may occasionally be useful for special effects. Schematic variables are not renamed by default, so beware of clashes! \ text %mlref \ \begin{mldecls} @{define_ML compose_tac: "Proof.context -> (bool * thm * int) -> int -> tactic"} \\ @{define_ML Drule.compose: "thm * int * thm -> thm"} \\ @{define_ML_infix COMP: "thm * thm -> thm"} \\ \end{mldecls} \<^descr> \<^ML>\compose_tac\~\ctxt (flag, rule, m) i\ refines subgoal \i\ using \rule\, without lifting. The \rule\ is taken to have the form \\\<^sub>1 \ \ \\<^sub>m \ \\, where \\\ need not be atomic; thus \m\ determines the number of new subgoals. If \flag\ is \true\ then it performs elim-resolution --- it solves the first premise of \rule\ by assumption and deletes that assumption. \<^descr> \<^ML>\Drule.compose\~\(thm\<^sub>1, i, thm\<^sub>2)\ uses \thm\<^sub>1\, regarded as an atomic formula, to solve premise \i\ of \thm\<^sub>2\. Let \thm\<^sub>1\ and \thm\<^sub>2\ be \\\ and \\\<^sub>1 \ \ \\<^sub>n \ \\. The unique \s\ that unifies \\\ and \\\<^sub>i\ yields the theorem \(\\<^sub>1 \ \ \\<^sub>i\<^sub>-\<^sub>1 \ \\<^sub>i\<^sub>+\<^sub>1 \ \ \\<^sub>n \ \)s\. Multiple results are considered as error (exception \<^ML>\THM\). \<^descr> \thm\<^sub>1 COMP thm\<^sub>2\ is the same as \Drule.compose (thm\<^sub>1, 1, thm\<^sub>2)\. \begin{warn} These low-level operations are stepping outside the structure imposed by regular rule resolution. Used without understanding of the consequences, they may produce results that cause problems with standard rules and tactics later on. \end{warn} \ section \Tacticals \label{sec:tacticals}\ text \ A \<^emph>\tactical\ is a functional combinator for building up complex tactics from simpler ones. Common tacticals perform sequential composition, disjunctive choice, iteration, or goal addressing. Various search strategies may be expressed via tacticals. \ subsection \Combining tactics\ text \ Sequential composition and alternative choices are the most basic ways to combine tactics, similarly to ``\<^verbatim>\,\'' and ``\<^verbatim>\|\'' in Isar method notation. - This corresponds to \<^ML_op>\THEN\ and \<^ML_op>\ORELSE\ in ML, but there + This corresponds to \<^ML_infix>\THEN\ and \<^ML_infix>\ORELSE\ in ML, but there are further possibilities for fine-tuning alternation of tactics such as - \<^ML_op>\APPEND\. Further details become visible in ML due to explicit + \<^ML_infix>\APPEND\. Further details become visible in ML due to explicit subgoal addressing. \ text %mlref \ \begin{mldecls} @{define_ML_infix "THEN": "tactic * tactic -> tactic"} \\ @{define_ML_infix "ORELSE": "tactic * tactic -> tactic"} \\ @{define_ML_infix "APPEND": "tactic * tactic -> tactic"} \\ @{define_ML "EVERY": "tactic list -> tactic"} \\ @{define_ML "FIRST": "tactic list -> tactic"} \\[0.5ex] @{define_ML_infix "THEN'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ @{define_ML_infix "ORELSE'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ @{define_ML_infix "APPEND'": "('a -> tactic) * ('a -> tactic) -> 'a -> tactic"} \\ @{define_ML "EVERY'": "('a -> tactic) list -> 'a -> tactic"} \\ @{define_ML "FIRST'": "('a -> tactic) list -> 'a -> tactic"} \\ \end{mldecls} - \<^descr> \tac\<^sub>1\~\<^ML_op>\THEN\~\tac\<^sub>2\ is the sequential composition of \tac\<^sub>1\ and + \<^descr> \tac\<^sub>1\~\<^ML_infix>\THEN\~\tac\<^sub>2\ is the sequential composition of \tac\<^sub>1\ and \tac\<^sub>2\. Applied to a goal state, it returns all states reachable in two steps by applying \tac\<^sub>1\ followed by \tac\<^sub>2\. First, it applies \tac\<^sub>1\ to the goal state, getting a sequence of possible next states; then, it applies \tac\<^sub>2\ to each of these and concatenates the results to produce again one flat sequence of states. - \<^descr> \tac\<^sub>1\~\<^ML_op>\ORELSE\~\tac\<^sub>2\ makes a choice between \tac\<^sub>1\ and + \<^descr> \tac\<^sub>1\~\<^ML_infix>\ORELSE\~\tac\<^sub>2\ makes a choice between \tac\<^sub>1\ and \tac\<^sub>2\. Applied to a state, it tries \tac\<^sub>1\ and returns the result if non-empty; if \tac\<^sub>1\ fails then it uses \tac\<^sub>2\. This is a deterministic choice: if \tac\<^sub>1\ succeeds then \tac\<^sub>2\ is excluded from the result. - \<^descr> \tac\<^sub>1\~\<^ML_op>\APPEND\~\tac\<^sub>2\ concatenates the possible results of - \tac\<^sub>1\ and \tac\<^sub>2\. Unlike \<^ML_op>\ORELSE\ there is \<^emph>\no commitment\ to - either tactic, so \<^ML_op>\APPEND\ helps to avoid incompleteness during + \<^descr> \tac\<^sub>1\~\<^ML_infix>\APPEND\~\tac\<^sub>2\ concatenates the possible results of + \tac\<^sub>1\ and \tac\<^sub>2\. Unlike \<^ML_infix>\ORELSE\ there is \<^emph>\no commitment\ to + either tactic, so \<^ML_infix>\APPEND\ helps to avoid incompleteness during search, at the cost of potential inefficiencies. - \<^descr> \<^ML>\EVERY\~\[tac\<^sub>1, \, tac\<^sub>n]\ abbreviates \tac\<^sub>1\~\<^ML_op>\THEN\~\\\~\<^ML_op>\THEN\~\tac\<^sub>n\. Note that \<^ML>\EVERY []\ is the same as + \<^descr> \<^ML>\EVERY\~\[tac\<^sub>1, \, tac\<^sub>n]\ abbreviates \tac\<^sub>1\~\<^ML_infix>\THEN\~\\\~\<^ML_infix>\THEN\~\tac\<^sub>n\. Note that \<^ML>\EVERY []\ is the same as \<^ML>\all_tac\: it always succeeds. - \<^descr> \<^ML>\FIRST\~\[tac\<^sub>1, \, tac\<^sub>n]\ abbreviates \tac\<^sub>1\~\<^ML_op>\ORELSE\~\\\~\<^ML_op>\ORELSE\~\tac\<^sub>n\. Note that \<^ML>\FIRST []\ is the + \<^descr> \<^ML>\FIRST\~\[tac\<^sub>1, \, tac\<^sub>n]\ abbreviates \tac\<^sub>1\~\<^ML_infix>\ORELSE\~\\\~\<^ML_infix>\ORELSE\~\tac\<^sub>n\. Note that \<^ML>\FIRST []\ is the same as \<^ML>\no_tac\: it always fails. - \<^descr> \<^ML_op>\THEN'\ is the lifted version of \<^ML_op>\THEN\, for tactics - with explicit subgoal addressing. So \(tac\<^sub>1\~\<^ML_op>\THEN'\~\tac\<^sub>2) i\ is - the same as \(tac\<^sub>1 i\~\<^ML_op>\THEN\~\tac\<^sub>2 i)\. + \<^descr> \<^ML_infix>\THEN'\ is the lifted version of \<^ML_infix>\THEN\, for tactics + with explicit subgoal addressing. So \(tac\<^sub>1\~\<^ML_infix>\THEN'\~\tac\<^sub>2) i\ is + the same as \(tac\<^sub>1 i\~\<^ML_infix>\THEN\~\tac\<^sub>2 i)\. The other primed tacticals work analogously. \ subsection \Repetition tacticals\ text \ These tacticals provide further control over repetition of tactics, beyond the stylized forms of ``\<^verbatim>\?\'' and ``\<^verbatim>\+\'' in Isar method expressions. \ text %mlref \ \begin{mldecls} @{define_ML "TRY": "tactic -> tactic"} \\ @{define_ML "REPEAT": "tactic -> tactic"} \\ @{define_ML "REPEAT1": "tactic -> tactic"} \\ @{define_ML "REPEAT_DETERM": "tactic -> tactic"} \\ @{define_ML "REPEAT_DETERM_N": "int -> tactic -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\TRY\~\tac\ applies \tac\ to the goal state and returns the resulting sequence, if non-empty; otherwise it returns the original state. Thus, it applies \tac\ at most once. Note that for tactics with subgoal addressing, the combinator can be applied - via functional composition: \<^ML>\TRY\~\<^ML_op>\o\~\tac\. There is no need + via functional composition: \<^ML>\TRY\~\<^ML_infix>\o\~\tac\. There is no need for \<^verbatim>\TRY'\. \<^descr> \<^ML>\REPEAT\~\tac\ applies \tac\ to the goal state and, recursively, to each element of the resulting sequence. The resulting sequence consists of those states that make \tac\ fail. Thus, it applies \tac\ as many times as possible (including zero times), and allows backtracking over each invocation of \tac\. \<^ML>\REPEAT\ is more general than \<^ML>\REPEAT_DETERM\, but requires more space. \<^descr> \<^ML>\REPEAT1\~\tac\ is like \<^ML>\REPEAT\~\tac\ but it always applies \tac\ at least once, failing if this is impossible. \<^descr> \<^ML>\REPEAT_DETERM\~\tac\ applies \tac\ to the goal state and, recursively, to the head of the resulting sequence. It returns the first state to make \tac\ fail. It is deterministic, discarding alternative outcomes. \<^descr> \<^ML>\REPEAT_DETERM_N\~\n tac\ is like \<^ML>\REPEAT_DETERM\~\tac\ but the number of repetitions is bound by \n\ (where \<^ML>\~1\ means \\\). \ text %mlex \ The basic tactics and tacticals considered above follow some algebraic laws: - \<^item> \<^ML>\all_tac\ is the identity element of the tactical \<^ML_op>\THEN\. + \<^item> \<^ML>\all_tac\ is the identity element of the tactical \<^ML_infix>\THEN\. - \<^item> \<^ML>\no_tac\ is the identity element of \<^ML_op>\ORELSE\ and \<^ML_op>\APPEND\. Also, it is a zero element for \<^ML_op>\THEN\, which means that - \tac\~\<^ML_op>\THEN\~\<^ML>\no_tac\ is equivalent to \<^ML>\no_tac\. + \<^item> \<^ML>\no_tac\ is the identity element of \<^ML_infix>\ORELSE\ and \<^ML_infix>\APPEND\. Also, it is a zero element for \<^ML_infix>\THEN\, which means that + \tac\~\<^ML_infix>\THEN\~\<^ML>\no_tac\ is equivalent to \<^ML>\no_tac\. \<^item> \<^ML>\TRY\ and \<^ML>\REPEAT\ can be expressed as (recursive) functions over more basic combinators (ignoring some internal implementation tricks): \ ML \ fun TRY tac = tac ORELSE all_tac; fun REPEAT tac st = ((tac THEN REPEAT tac) ORELSE all_tac) st; \ text \ - If \tac\ can return multiple outcomes then so can \<^ML>\REPEAT\~\tac\. \<^ML>\REPEAT\ uses \<^ML_op>\ORELSE\ and not \<^ML_op>\APPEND\, it applies \tac\ + If \tac\ can return multiple outcomes then so can \<^ML>\REPEAT\~\tac\. \<^ML>\REPEAT\ uses \<^ML_infix>\ORELSE\ and not \<^ML_infix>\APPEND\, it applies \tac\ as many times as possible in each outcome. \begin{warn} Note the explicit abstraction over the goal state in the ML definition of \<^ML>\REPEAT\. Recursive tacticals must be coded in this awkward fashion to avoid infinite recursion of eager functional evaluation in Standard ML. The following attempt would make \<^ML>\REPEAT\~\tac\ loop: \end{warn} \ ML_val \ (*BAD -- does not terminate!*) fun REPEAT tac = (tac THEN REPEAT tac) ORELSE all_tac; \ subsection \Applying tactics to subgoal ranges\ text \ Tactics with explicit subgoal addressing \<^ML_type>\int -> tactic\ can be used together with tacticals that act like ``subgoal quantifiers'': guided by success of the body tactic a certain range of subgoals is covered. Thus the body tactic is applied to \<^emph>\all\ subgoals, \<^emph>\some\ subgoal etc. Suppose that the goal state has \n \ 0\ subgoals. Many of these tacticals address subgoal ranges counting downwards from \n\ towards \1\. This has the fortunate effect that newly emerging subgoals are concatenated in the result, without interfering each other. Nonetheless, there might be situations where a different order is desired. \ text %mlref \ \begin{mldecls} @{define_ML ALLGOALS: "(int -> tactic) -> tactic"} \\ @{define_ML SOMEGOAL: "(int -> tactic) -> tactic"} \\ @{define_ML FIRSTGOAL: "(int -> tactic) -> tactic"} \\ @{define_ML HEADGOAL: "(int -> tactic) -> tactic"} \\ @{define_ML REPEAT_SOME: "(int -> tactic) -> tactic"} \\ @{define_ML REPEAT_FIRST: "(int -> tactic) -> tactic"} \\ @{define_ML RANGE: "(int -> tactic) list -> int -> tactic"} \\ \end{mldecls} - \<^descr> \<^ML>\ALLGOALS\~\tac\ is equivalent to \tac n\~\<^ML_op>\THEN\~\\\~\<^ML_op>\THEN\~\tac 1\. It applies the \tac\ to all the subgoals, counting downwards. + \<^descr> \<^ML>\ALLGOALS\~\tac\ is equivalent to \tac n\~\<^ML_infix>\THEN\~\\\~\<^ML_infix>\THEN\~\tac 1\. It applies the \tac\ to all the subgoals, counting downwards. - \<^descr> \<^ML>\SOMEGOAL\~\tac\ is equivalent to \tac n\~\<^ML_op>\ORELSE\~\\\~\<^ML_op>\ORELSE\~\tac 1\. It applies \tac\ to one subgoal, counting downwards. + \<^descr> \<^ML>\SOMEGOAL\~\tac\ is equivalent to \tac n\~\<^ML_infix>\ORELSE\~\\\~\<^ML_infix>\ORELSE\~\tac 1\. It applies \tac\ to one subgoal, counting downwards. - \<^descr> \<^ML>\FIRSTGOAL\~\tac\ is equivalent to \tac 1\~\<^ML_op>\ORELSE\~\\\~\<^ML_op>\ORELSE\~\tac n\. It applies \tac\ to one subgoal, counting upwards. + \<^descr> \<^ML>\FIRSTGOAL\~\tac\ is equivalent to \tac 1\~\<^ML_infix>\ORELSE\~\\\~\<^ML_infix>\ORELSE\~\tac n\. It applies \tac\ to one subgoal, counting upwards. \<^descr> \<^ML>\HEADGOAL\~\tac\ is equivalent to \tac 1\. It applies \tac\ unconditionally to the first subgoal. \<^descr> \<^ML>\REPEAT_SOME\~\tac\ applies \tac\ once or more to a subgoal, counting downwards. \<^descr> \<^ML>\REPEAT_FIRST\~\tac\ applies \tac\ once or more to a subgoal, counting upwards. \<^descr> \<^ML>\RANGE\~\[tac\<^sub>1, \, tac\<^sub>k] i\ is equivalent to \tac\<^sub>k (i + k - - 1)\~\<^ML_op>\THEN\~\\\~\<^ML_op>\THEN\~\tac\<^sub>1 i\. It applies the given list of + 1)\~\<^ML_infix>\THEN\~\\\~\<^ML_infix>\THEN\~\tac\<^sub>1 i\. It applies the given list of tactics to the corresponding range of subgoals, counting downwards. \ subsection \Control and search tacticals\ text \ A predicate on theorems \<^ML_type>\thm -> bool\ can test whether a goal state enjoys some desirable property --- such as having no subgoals. Tactics that search for satisfactory goal states are easy to express. The main search procedures, depth-first, breadth-first and best-first, are provided as tacticals. They generate the search tree by repeatedly applying a given tactic. \ text %mlref "" subsubsection \Filtering a tactic's results\ text \ \begin{mldecls} @{define_ML FILTER: "(thm -> bool) -> tactic -> tactic"} \\ @{define_ML CHANGED: "tactic -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\FILTER\~\sat tac\ applies \tac\ to the goal state and returns a sequence consisting of those result goal states that are satisfactory in the sense of \sat\. \<^descr> \<^ML>\CHANGED\~\tac\ applies \tac\ to the goal state and returns precisely those states that differ from the original state (according to \<^ML>\Thm.eq_thm\). Thus \<^ML>\CHANGED\~\tac\ always has some effect on the state. \ subsubsection \Depth-first search\ text \ \begin{mldecls} @{define_ML DEPTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\ @{define_ML DEPTH_SOLVE: "tactic -> tactic"} \\ @{define_ML DEPTH_SOLVE_1: "tactic -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\DEPTH_FIRST\~\sat tac\ returns the goal state if \sat\ returns true. Otherwise it applies \tac\, then recursively searches from each element of the resulting sequence. The code uses a stack for efficiency, in effect - applying \tac\~\<^ML_op>\THEN\~\<^ML>\DEPTH_FIRST\~\sat tac\ to the state. + applying \tac\~\<^ML_infix>\THEN\~\<^ML>\DEPTH_FIRST\~\sat tac\ to the state. \<^descr> \<^ML>\DEPTH_SOLVE\\tac\ uses \<^ML>\DEPTH_FIRST\ to search for states having no subgoals. \<^descr> \<^ML>\DEPTH_SOLVE_1\~\tac\ uses \<^ML>\DEPTH_FIRST\ to search for states having fewer subgoals than the given state. Thus, it insists upon solving at least one subgoal. \ subsubsection \Other search strategies\ text \ \begin{mldecls} @{define_ML BREADTH_FIRST: "(thm -> bool) -> tactic -> tactic"} \\ @{define_ML BEST_FIRST: "(thm -> bool) * (thm -> int) -> tactic -> tactic"} \\ @{define_ML THEN_BEST_FIRST: "tactic -> (thm -> bool) * (thm -> int) -> tactic -> tactic"} \\ \end{mldecls} These search strategies will find a solution if one exists. However, they do not enumerate all solutions; they terminate after the first satisfactory result from \tac\. \<^descr> \<^ML>\BREADTH_FIRST\~\sat tac\ uses breadth-first search to find states for which \sat\ is true. For most applications, it is too slow. \<^descr> \<^ML>\BEST_FIRST\~\(sat, dist) tac\ does a heuristic search, using \dist\ to estimate the distance from a satisfactory state (in the sense of \sat\). It maintains a list of states ordered by distance. It applies \tac\ to the head of this list; if the result contains any satisfactory states, then it returns them. Otherwise, \<^ML>\BEST_FIRST\ adds the new states to the list, and continues. The distance function is typically \<^ML>\size_of_thm\, which computes the size of the state. The smaller the state, the fewer and simpler subgoals it has. \<^descr> \<^ML>\THEN_BEST_FIRST\~\tac\<^sub>0 (sat, dist) tac\ is like \<^ML>\BEST_FIRST\, except that the priority queue initially contains the result of applying \tac\<^sub>0\ to the goal state. This tactical permits separate tactics for starting the search and continuing the search. \ subsubsection \Auxiliary tacticals for searching\ text \ \begin{mldecls} @{define_ML COND: "(thm -> bool) -> tactic -> tactic -> tactic"} \\ @{define_ML IF_UNSOLVED: "tactic -> tactic"} \\ @{define_ML SOLVE: "tactic -> tactic"} \\ @{define_ML DETERM: "tactic -> tactic"} \\ \end{mldecls} \<^descr> \<^ML>\COND\~\sat tac\<^sub>1 tac\<^sub>2\ applies \tac\<^sub>1\ to the goal state if it satisfies predicate \sat\, and applies \tac\<^sub>2\. It is a conditional tactical in that only one of \tac\<^sub>1\ and \tac\<^sub>2\ is applied to a goal state. However, both \tac\<^sub>1\ and \tac\<^sub>2\ are evaluated because ML uses eager evaluation. \<^descr> \<^ML>\IF_UNSOLVED\~\tac\ applies \tac\ to the goal state if it has any subgoals, and simply returns the goal state otherwise. Many common tactics, such as \<^ML>\resolve_tac\, fail if applied to a goal state that has no subgoals. \<^descr> \<^ML>\SOLVE\~\tac\ applies \tac\ to the goal state and then fails iff there are subgoals left. \<^descr> \<^ML>\DETERM\~\tac\ applies \tac\ to the goal state and returns the head of the resulting sequence. \<^ML>\DETERM\ limits the search space by making its argument deterministic. \ subsubsection \Predicates and functions useful for searching\ text \ \begin{mldecls} @{define_ML has_fewer_prems: "int -> thm -> bool"} \\ @{define_ML Thm.eq_thm: "thm * thm -> bool"} \\ @{define_ML Thm.eq_thm_prop: "thm * thm -> bool"} \\ @{define_ML size_of_thm: "thm -> int"} \\ \end{mldecls} \<^descr> \<^ML>\has_fewer_prems\~\n thm\ reports whether \thm\ has fewer than \n\ premises. \<^descr> \<^ML>\Thm.eq_thm\~\(thm\<^sub>1, thm\<^sub>2)\ reports whether \thm\<^sub>1\ and \thm\<^sub>2\ are equal. Both theorems must have the same conclusions, the same set of hypotheses, and the same set of sort hypotheses. Names of bound variables are ignored as usual. \<^descr> \<^ML>\Thm.eq_thm_prop\~\(thm\<^sub>1, thm\<^sub>2)\ reports whether the propositions of \thm\<^sub>1\ and \thm\<^sub>2\ are equal. Names of bound variables are ignored. \<^descr> \<^ML>\size_of_thm\~\thm\ computes the size of \thm\, namely the number of variables, constants and abstractions in its conclusion. It may serve as a distance function for \<^ML>\BEST_FIRST\. \ end diff --git a/src/Doc/Isar_Ref/Document_Preparation.thy b/src/Doc/Isar_Ref/Document_Preparation.thy --- a/src/Doc/Isar_Ref/Document_Preparation.thy +++ b/src/Doc/Isar_Ref/Document_Preparation.thy @@ -1,700 +1,718 @@ (*:maxLineLen=78:*) theory Document_Preparation imports Main Base begin chapter \Document preparation \label{ch:document-prep}\ text \ Isabelle/Isar provides a simple document preparation system based on {PDF-\LaTeX}, with support for hyperlinks and bookmarks within that format. This allows to produce papers, books, theses etc.\ from Isabelle theory sources. {\LaTeX} output is generated while processing a \<^emph>\session\ in batch mode, as explained in the \<^emph>\The Isabelle System Manual\ @{cite "isabelle-system"}. The main Isabelle tools to get started with document preparation are @{tool_ref mkroot} and @{tool_ref build}. The classic Isabelle/HOL tutorial @{cite "isabelle-hol-book"} also explains some aspects of theory presentation. \ section \Markup commands \label{sec:markup}\ text \ \begin{matharray}{rcl} @{command_def "chapter"} & : & \any \ any\ \\ @{command_def "section"} & : & \any \ any\ \\ @{command_def "subsection"} & : & \any \ any\ \\ @{command_def "subsubsection"} & : & \any \ any\ \\ @{command_def "paragraph"} & : & \any \ any\ \\ @{command_def "subparagraph"} & : & \any \ any\ \\ @{command_def "text"} & : & \any \ any\ \\ @{command_def "txt"} & : & \any \ any\ \\ @{command_def "text_raw"} & : & \any \ any\ \\ \end{matharray} Markup commands provide a structured way to insert text into the document generated from a theory. Each markup command takes a single @{syntax text} argument, which is passed as argument to a corresponding {\LaTeX} macro. The default macros provided by \<^file>\~~/lib/texinputs/isabelle.sty\ can be redefined according to the needs of the underlying document and {\LaTeX} styles. Note that formal comments (\secref{sec:comments}) are similar to markup commands, but have a different status within Isabelle/Isar syntax. \<^rail>\ (@@{command chapter} | @@{command section} | @@{command subsection} | @@{command subsubsection} | @@{command paragraph} | @@{command subparagraph}) @{syntax text} ';'? | (@@{command text} | @@{command txt} | @@{command text_raw}) @{syntax text} \ \<^descr> @{command chapter}, @{command section}, @{command subsection} etc.\ mark section headings within the theory source. This works in any context, even before the initial @{command theory} command. The corresponding {\LaTeX} macros are \<^verbatim>\\isamarkupchapter\, \<^verbatim>\\isamarkupsection\, \<^verbatim>\\isamarkupsubsection\ etc.\ \<^descr> @{command text} and @{command txt} specify paragraphs of plain text. This corresponds to a {\LaTeX} environment \<^verbatim>\\begin{isamarkuptext}\ \\\ \<^verbatim>\\end{isamarkuptext}\ etc. \<^descr> @{command text_raw} is similar to @{command text}, but without any surrounding markup environment. This allows to inject arbitrary {\LaTeX} source into the generated document. All text passed to any of the above markup commands may refer to formal entities via \<^emph>\document antiquotations\, see also \secref{sec:antiq}. These are interpreted in the present theory or proof context. \<^medskip> The proof markup commands closely resemble those for theory specifications, but have a different formal status and produce different {\LaTeX} macros. \ section \Document antiquotations \label{sec:antiq}\ text \ \begin{matharray}{rcl} @{antiquotation_def "theory"} & : & \antiquotation\ \\ @{antiquotation_def "thm"} & : & \antiquotation\ \\ @{antiquotation_def "lemma"} & : & \antiquotation\ \\ @{antiquotation_def "prop"} & : & \antiquotation\ \\ @{antiquotation_def "term"} & : & \antiquotation\ \\ @{antiquotation_def term_type} & : & \antiquotation\ \\ @{antiquotation_def typeof} & : & \antiquotation\ \\ @{antiquotation_def const} & : & \antiquotation\ \\ @{antiquotation_def abbrev} & : & \antiquotation\ \\ @{antiquotation_def typ} & : & \antiquotation\ \\ @{antiquotation_def type} & : & \antiquotation\ \\ @{antiquotation_def class} & : & \antiquotation\ \\ @{antiquotation_def locale} & : & \antiquotation\ \\ @{antiquotation_def "text"} & : & \antiquotation\ \\ @{antiquotation_def goals} & : & \antiquotation\ \\ @{antiquotation_def subgoals} & : & \antiquotation\ \\ @{antiquotation_def prf} & : & \antiquotation\ \\ @{antiquotation_def full_prf} & : & \antiquotation\ \\ + @{antiquotation_def ML_text} & : & \antiquotation\ \\ @{antiquotation_def ML} & : & \antiquotation\ \\ - @{antiquotation_def ML_op} & : & \antiquotation\ \\ + @{antiquotation_def ML_def} & : & \antiquotation\ \\ + @{antiquotation_def ML_ref} & : & \antiquotation\ \\ + @{antiquotation_def ML_infix} & : & \antiquotation\ \\ + @{antiquotation_def ML_infix_def} & : & \antiquotation\ \\ + @{antiquotation_def ML_infix_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_type} & : & \antiquotation\ \\ + @{antiquotation_def ML_type_def} & : & \antiquotation\ \\ + @{antiquotation_def ML_type_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_structure} & : & \antiquotation\ \\ + @{antiquotation_def ML_structure_def} & : & \antiquotation\ \\ + @{antiquotation_def ML_structure_ref} & : & \antiquotation\ \\ @{antiquotation_def ML_functor} & : & \antiquotation\ \\ + @{antiquotation_def ML_functor_def} & : & \antiquotation\ \\ + @{antiquotation_def ML_functor_ref} & : & \antiquotation\ \\ @{antiquotation_def emph} & : & \antiquotation\ \\ @{antiquotation_def bold} & : & \antiquotation\ \\ @{antiquotation_def verbatim} & : & \antiquotation\ \\ @{antiquotation_def bash_function} & : & \antiquotation\ \\ @{antiquotation_def system_option} & : & \antiquotation\ \\ @{antiquotation_def session} & : & \antiquotation\ \\ @{antiquotation_def "file"} & : & \antiquotation\ \\ @{antiquotation_def "url"} & : & \antiquotation\ \\ @{antiquotation_def "cite"} & : & \antiquotation\ \\ @{command_def "print_antiquotations"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} The overall content of an Isabelle/Isar theory may alternate between formal and informal text. The main body consists of formal specification and proof commands, interspersed with markup commands (\secref{sec:markup}) or document comments (\secref{sec:comments}). The argument of markup commands quotes informal text to be printed in the resulting document, but may again refer to formal entities via \<^emph>\document antiquotations\. For example, embedding \<^verbatim>\@{term [show_types] "f x = a + x"}\ within a text block makes \isa{{\isacharparenleft}f{\isasymColon}{\isacharprime}a\ {\isasymRightarrow}\ {\isacharprime}a{\isacharparenright}\ {\isacharparenleft}x{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}a{\isasymColon}{\isacharprime}a{\isacharparenright}\ {\isacharplus}\ x} appear in the final {\LaTeX} document. Antiquotations usually spare the author tedious typing of logical entities in full detail. Even more importantly, some degree of consistency-checking between the main body of formal text and its informal explanation is achieved, since terms and types appearing in antiquotations are checked within the current theory or proof context. \<^medskip> Antiquotations are in general written as \<^verbatim>\@{\\name\~\<^verbatim>\[\\options\\<^verbatim>\]\~\arguments\\<^verbatim>\}\. The short form \<^verbatim>\\\\<^verbatim>\<^\\name\\<^verbatim>\>\\\argument_content\\ (without surrounding \<^verbatim>\@{\\\\\<^verbatim>\}\) works for a single argument that is a cartouche. A cartouche without special decoration is equivalent to \<^verbatim>\\<^cartouche>\\\argument_content\\, which is equivalent to \<^verbatim>\@{cartouche\~\\argument_content\\\<^verbatim>\}\. The special name @{antiquotation_def cartouche} is defined in the context: Isabelle/Pure introduces that as an alias to @{antiquotation_ref text} (see below). Consequently, \\foo_bar + baz \ bazar\\ prints literal quasi-formal text (unchecked). A control symbol \<^verbatim>\\\\<^verbatim>\<^\\name\\<^verbatim>\>\ within the body text, but without a subsequent cartouche, is equivalent to \<^verbatim>\@{\\name\\<^verbatim>\}\. \begingroup \def\isasymcontrolstart{\isatt{\isacharbackslash\isacharless\isacharcircum}} \<^rail>\ @{syntax_def antiquotation}: '@{' antiquotation_body '}' | '\' @{syntax_ref name} '>' @{syntax_ref cartouche} | @{syntax_ref cartouche} ; options: '[' (option * ',') ']' ; option: @{syntax name} | @{syntax name} '=' @{syntax name} ; \ \endgroup Note that the syntax of antiquotations may \<^emph>\not\ include source comments \<^verbatim>\(*\~\\\~\<^verbatim>\*)\ nor verbatim text \<^verbatim>\{*\~\\\~\<^verbatim>\*}\. %% FIXME less monolithic presentation, move to individual sections!? \<^rail>\ @{syntax_def antiquotation_body}: (@@{antiquotation text} | @@{antiquotation cartouche} | @@{antiquotation theory_text}) options @{syntax text} | @@{antiquotation theory} options @{syntax embedded} | @@{antiquotation thm} options styles @{syntax thms} | @@{antiquotation lemma} options @{syntax prop} @'by' @{syntax method} @{syntax method}? | @@{antiquotation prop} options styles @{syntax prop} | @@{antiquotation term} options styles @{syntax term} | @@{antiquotation (HOL) value} options styles @{syntax term} | @@{antiquotation term_type} options styles @{syntax term} | @@{antiquotation typeof} options styles @{syntax term} | @@{antiquotation const} options @{syntax term} | @@{antiquotation abbrev} options @{syntax term} | @@{antiquotation typ} options @{syntax type} | @@{antiquotation type} options @{syntax embedded} | @@{antiquotation class} options @{syntax embedded} | @@{antiquotation locale} options @{syntax embedded} | (@@{antiquotation command} | @@{antiquotation method} | @@{antiquotation attribute}) options @{syntax name} ; @{syntax antiquotation}: @@{antiquotation goals} options | @@{antiquotation subgoals} options | @@{antiquotation prf} options @{syntax thms} | @@{antiquotation full_prf} options @{syntax thms} | + @@{antiquotation ML_text} options @{syntax text} | @@{antiquotation ML} options @{syntax text} | - @@{antiquotation ML_op} options @{syntax text} | + @@{antiquotation ML_infix} options @{syntax text} | @@{antiquotation ML_type} options @{syntax text} | @@{antiquotation ML_structure} options @{syntax text} | @@{antiquotation ML_functor} options @{syntax text} | @@{antiquotation emph} options @{syntax text} | @@{antiquotation bold} options @{syntax text} | @@{antiquotation verbatim} options @{syntax text} | @@{antiquotation bash_function} options @{syntax embedded} | @@{antiquotation system_option} options @{syntax embedded} | @@{antiquotation session} options @{syntax embedded} | @@{antiquotation path} options @{syntax embedded} | @@{antiquotation "file"} options @{syntax name} | @@{antiquotation dir} options @{syntax name} | @@{antiquotation url} options @{syntax embedded} | @@{antiquotation cite} options @{syntax cartouche}? (@{syntax name} + @'and') ; styles: '(' (style + ',') ')' ; style: (@{syntax name} +) ; @@{command print_antiquotations} ('!'?) \ \<^descr> \@{text s}\ prints uninterpreted source text \s\, i.e.\ inner syntax. This is particularly useful to print portions of text according to the Isabelle document style, without demanding well-formedness, e.g.\ small pieces of terms that should not be parsed or type-checked yet. It is also possible to write this in the short form \\s\\ without any further decoration. \<^descr> \@{theory_text s}\ prints uninterpreted theory source text \s\, i.e.\ outer syntax with command keywords and other tokens. \<^descr> \@{theory A}\ prints the session-qualified theory name \A\, which is guaranteed to refer to a valid ancestor theory in the current context. \<^descr> \@{thm a\<^sub>1 \ a\<^sub>n}\ prints theorems \a\<^sub>1 \ a\<^sub>n\. Full fact expressions are allowed here, including attributes (\secref{sec:syn-att}). \<^descr> \@{prop \}\ prints a well-typed proposition \\\. \<^descr> \@{lemma \ by m}\ proves a well-typed proposition \\\ by method \m\ and prints the original \\\. \<^descr> \@{term t}\ prints a well-typed term \t\. \<^descr> \@{value t}\ evaluates a term \t\ and prints its result, see also @{command_ref (HOL) value}. \<^descr> \@{term_type t}\ prints a well-typed term \t\ annotated with its type. \<^descr> \@{typeof t}\ prints the type of a well-typed term \t\. \<^descr> \@{const c}\ prints a logical or syntactic constant \c\. \<^descr> \@{abbrev c x\<^sub>1 \ x\<^sub>n}\ prints a constant abbreviation \c x\<^sub>1 \ x\<^sub>n \ rhs\ as defined in the current context. \<^descr> \@{typ \}\ prints a well-formed type \\\. \<^descr> \@{type \}\ prints a (logical or syntactic) type constructor \\\. \<^descr> \@{class c}\ prints a class \c\. \<^descr> \@{locale c}\ prints a locale \c\. \<^descr> \@{command name}\, \@{method name}\, \@{attribute name}\ print checked entities of the Isar language. \<^descr> \@{goals}\ prints the current \<^emph>\dynamic\ goal state. This is mainly for support of tactic-emulation scripts within Isar. Presentation of goal states does not conform to the idea of human-readable proof documents! When explaining proofs in detail it is usually better to spell out the reasoning via proper Isar proof commands, instead of peeking at the internal machine configuration. \<^descr> \@{subgoals}\ is similar to \@{goals}\, but does not print the main goal. \<^descr> \@{prf a\<^sub>1 \ a\<^sub>n}\ prints the (compact) proof terms corresponding to the theorems \a\<^sub>1 \ a\<^sub>n\. Note that this requires proof terms to be switched on for the current logic session. \<^descr> \@{full_prf a\<^sub>1 \ a\<^sub>n}\ is like \@{prf a\<^sub>1 \ a\<^sub>n}\, but prints the full proof terms, i.e.\ also displays information omitted in the compact proof term, which is denoted by ``\_\'' placeholders there. - - \<^descr> \@{ML s}\, \@{ML_op s}\, \@{ML_type s}\, \@{ML_structure s}\, and + + \<^descr> \@{ML_text s}\ prints ML text verbatim: only the token language is + checked. + + \<^descr> \@{ML s}\, \@{ML_infix s}\, \@{ML_type s}\, \@{ML_structure s}\, and \@{ML_functor s}\ check text \s\ as ML value, infix operator, type, - structure, and functor respectively. The source is printed verbatim. + exception, structure, and functor respectively. The source is printed + verbatim. The variants \@{ML_def s}\ and \@{ML_ref s}\ etc. maintain the + document index: ``def'' means to make a bold entry, ``ref'' means to make a + regular entry. \<^descr> \@{emph s}\ prints document source recursively, with {\LaTeX} markup \<^verbatim>\\emph{\\\\\<^verbatim>\}\. \<^descr> \@{bold s}\ prints document source recursively, with {\LaTeX} markup \<^verbatim>\\textbf{\\\\\<^verbatim>\}\. \<^descr> \@{verbatim s}\ prints uninterpreted source text literally as ASCII characters, using some type-writer font style. \<^descr> \@{bash_function name}\ prints the given GNU bash function verbatim. The name is checked wrt.\ the Isabelle system environment @{cite "isabelle-system"}. \<^descr> \@{system_option name}\ prints the given system option verbatim. The name is checked wrt.\ cumulative \<^verbatim>\etc/options\ of all Isabelle components, notably \<^file>\~~/etc/options\. \<^descr> \@{session name}\ prints given session name verbatim. The name is checked wrt.\ the dependencies of the current session. \<^descr> \@{path name}\ prints the file-system path name verbatim. \<^descr> \@{file name}\ is like \@{path name}\, but ensures that \name\ refers to a plain file. \<^descr> \@{dir name}\ is like \@{path name}\, but ensures that \name\ refers to a directory. \<^descr> \@{url name}\ produces markup for the given URL, which results in an active hyperlink within the text. \<^descr> \@{cite name}\ produces a citation \<^verbatim>\\cite{name}\ in {\LaTeX}, where the name refers to some Bib{\TeX} database entry. This is only checked in batch-mode session builds. The variant \@{cite \opt\ name}\ produces \<^verbatim>\\cite[opt]{name}\ with some free-form optional argument. Multiple names are output with commas, e.g. \@{cite foo \ bar}\ becomes \<^verbatim>\\cite{foo,bar}\. The {\LaTeX} macro name is determined by the antiquotation option @{antiquotation_option_def cite_macro}, or the configuration option @{attribute cite_macro} in the context. For example, \@{cite [cite_macro = nocite] foobar}\ produces \<^verbatim>\\nocite{foobar}\. \<^descr> @{command "print_antiquotations"} prints all document antiquotations that are defined in the current context; the ``\!\'' option indicates extra verbosity. \ subsection \Styled antiquotations\ text \ The antiquotations \thm\, \prop\ and \term\ admit an extra \<^emph>\style\ specification to modify the printed result. A style is specified by a name with a possibly empty number of arguments; multiple styles can be sequenced with commas. The following standard styles are available: \<^descr> \lhs\ extracts the first argument of any application form with at least two arguments --- typically meta-level or object-level equality, or any other binary relation. \<^descr> \rhs\ is like \lhs\, but extracts the second argument. \<^descr> \concl\ extracts the conclusion \C\ from a rule in Horn-clause normal form \A\<^sub>1 \ \ A\<^sub>n \ C\. \<^descr> \prem\ \n\ extract premise number \n\ from from a rule in Horn-clause normal form \A\<^sub>1 \ \ A\<^sub>n \ C\. \ subsection \General options\ text \ The following options are available to tune the printed output of antiquotations. Note that many of these coincide with system and configuration options of the same names. \<^descr> @{antiquotation_option_def show_types}~\= bool\ and @{antiquotation_option_def show_sorts}~\= bool\ control printing of explicit type and sort constraints. \<^descr> @{antiquotation_option_def show_structs}~\= bool\ controls printing of implicit structures. \<^descr> @{antiquotation_option_def show_abbrevs}~\= bool\ controls folding of abbreviations. \<^descr> @{antiquotation_option_def names_long}~\= bool\ forces names of types and constants etc.\ to be printed in their fully qualified internal form. \<^descr> @{antiquotation_option_def names_short}~\= bool\ forces names of types and constants etc.\ to be printed unqualified. Note that internalizing the output again in the current context may well yield a different result. \<^descr> @{antiquotation_option_def names_unique}~\= bool\ determines whether the printed version of qualified names should be made sufficiently long to avoid overlap with names declared further back. Set to \false\ for more concise output. \<^descr> @{antiquotation_option_def eta_contract}~\= bool\ prints terms in \\\-contracted form. \<^descr> @{antiquotation_option_def display}~\= bool\ indicates if the text is to be output as multi-line ``display material'', rather than a small piece of text without line breaks (which is the default). In this mode the embedded entities are printed in the same style as the main theory text. \<^descr> @{antiquotation_option_def break}~\= bool\ controls line breaks in non-display material. \<^descr> @{antiquotation_option_def cartouche}~\= bool\ indicates if the output should be delimited as cartouche. \<^descr> @{antiquotation_option_def quotes}~\= bool\ indicates if the output should be delimited via double quotes (option @{antiquotation_option cartouche} takes precedence). Note that the Isabelle {\LaTeX} styles may suppress quotes on their own account. \<^descr> @{antiquotation_option_def mode}~\= name\ adds \name\ to the print mode to be used for presentation. Note that the standard setup for {\LaTeX} output is already present by default, with mode ``\latex\''. \<^descr> @{antiquotation_option_def margin}~\= nat\ and @{antiquotation_option_def indent}~\= nat\ change the margin or indentation for pretty printing of display material. \<^descr> @{antiquotation_option_def goals_limit}~\= nat\ determines the maximum number of subgoals to be printed (for goal-based antiquotation). \<^descr> @{antiquotation_option_def source}~\= bool\ prints the original source text of the antiquotation arguments, rather than its internal representation. Note that formal checking of @{antiquotation "thm"}, @{antiquotation "term"}, etc. is still enabled; use the @{antiquotation "text"} antiquotation for unchecked output. Regular \term\ and \typ\ antiquotations with \source = false\ involve a full round-trip from the original source to an internalized logical entity back to a source form, according to the syntax of the current context. Thus the printed output is not under direct control of the author, it may even fluctuate a bit as the underlying theory is changed later on. In contrast, @{antiquotation_option source}~\= true\ admits direct printing of the given source text, with the desirable well-formedness check in the background, but without modification of the printed text. Cartouche delimiters of the argument are stripped for antiquotations that are internally categorized as ``embedded''. \<^descr> @{antiquotation_option_def source_cartouche} is like @{antiquotation_option source}, but cartouche delimiters are always preserved in the output. For Boolean flags, ``\name = true\'' may be abbreviated as ``\name\''. All of the above flags are disabled by default, unless changed specifically for a logic session in the corresponding \<^verbatim>\ROOT\ file. \ section \Markdown-like text structure\ text \ The markup commands @{command_ref text}, @{command_ref txt}, @{command_ref text_raw} (\secref{sec:markup}) consist of plain text. Its internal structure consists of paragraphs and (nested) lists, using special Isabelle symbols and some rules for indentation and blank lines. This quasi-visual format resembles \<^emph>\Markdown\\<^footnote>\\<^url>\http://commonmark.org\\, but the full complexity of that notation is avoided. This is a summary of the main principles of minimal Markdown in Isabelle: \<^item> List items start with the following markers \<^descr>[itemize:] \<^verbatim>\\<^item>\ \<^descr>[enumerate:] \<^verbatim>\\<^enum>\ \<^descr>[description:] \<^verbatim>\\<^descr>\ \<^item> Adjacent list items with same indentation and same marker are grouped into a single list. \<^item> Singleton blank lines separate paragraphs. \<^item> Multiple blank lines escape from the current list hierarchy. Notable differences to official Markdown: \<^item> Indentation of list items needs to match exactly. \<^item> Indentation is unlimited (official Markdown interprets four spaces as block quote). \<^item> List items always consist of paragraphs --- there is no notion of ``tight'' list. \<^item> Section headings are expressed via Isar document markup commands (\secref{sec:markup}). \<^item> URLs, font styles, other special content is expressed via antiquotations (\secref{sec:antiq}), usually with proper nesting of sub-languages via text cartouches. \ section \Document markers and command tags \label{sec:document-markers}\ text \ \emph{Document markers} are formal comments of the form \\<^marker>\marker_body\\ (using the control symbol \<^verbatim>\\<^marker>\) and may occur anywhere within the outer syntax of a command: the inner syntax of a marker body resembles that for attributes (\secref{sec:syn-att}). In contrast, \emph{Command tags} may only occur after a command keyword and are treated as special markers as explained below. \<^rail>\ @{syntax_def marker}: '\<^marker>' @{syntax cartouche} ; @{syntax_def marker_body}: (@{syntax name} @{syntax args} * ',') ; @{syntax_def tags}: tag* ; tag: '%' (@{syntax short_ident} | @{syntax string}) \ Document markers are stripped from the document output, but surrounding white-space is preserved: e.g.\ a marker at the end of a line does not affect the subsequent line break. Markers operate within the semantic presentation context of a command, and may modify it to change the overall appearance of a command span (e.g.\ by adding tags). Each document marker has its own syntax defined in the theory context; the following markers are predefined in Isabelle/Pure: \<^rail>\ (@@{document_marker_def title} | @@{document_marker_def creator} | @@{document_marker_def contributor} | @@{document_marker_def date} | @@{document_marker_def license} | @@{document_marker_def description}) @{syntax embedded} ; @@{document_marker_def tag} (scope?) @{syntax name} ; scope: '(' ('proof' | 'command') ')' \ \<^descr> \\<^marker>\title arg\\, \\<^marker>\creator arg\\, \\<^marker>\contributor arg\\, \\<^marker>\date arg\\, \\<^marker>\license arg\\, and \\<^marker>\description arg\\ produce markup in the PIDE document, without any immediate effect on typesetting. This vocabulary is taken from the Dublin Core Metadata Initiative\<^footnote>\\<^url>\https://www.dublincore.org/specifications/dublin-core/dcmi-terms\\. The argument is an uninterpreted string, except for @{document_marker description}, which consists of words that are subject to spell-checking. \<^descr> \\<^marker>\tag name\\ updates the list of command tags in the presentation context: later declarations take precedence, so \\<^marker>\tag a, tag b, tag c\\ produces a reversed list. The default tags are given by the original \<^theory_text>\keywords\ declaration of a command, and the system option @{system_option_ref document_tags}. The optional \scope\ tells how far the tagging is applied to subsequent proof structure: ``\<^theory_text>\("proof")\'' means it applies to the following proof text, and ``\<^theory_text>\(command)\'' means it only applies to the current command. The default within a proof body is ``\<^theory_text>\("proof")\'', but for toplevel goal statements it is ``\<^theory_text>\(command)\''. Thus a \tag\ marker for \<^theory_text>\theorem\, \<^theory_text>\lemma\ etc. does \emph{not} affect its proof by default. An old-style command tag \<^verbatim>\%\\name\ is treated like a document marker \\<^marker>\tag (proof) name\\: the list of command tags precedes the list of document markers. The head of the resulting tags in the presentation context is turned into {\LaTeX} environments to modify the type-setting. The following tags are pre-declared for certain classes of commands, and serve as default markup for certain kinds of commands: \<^medskip> \begin{tabular}{ll} \document\ & document markup commands \\ \theory\ & theory begin/end \\ \proof\ & all proof commands \\ \ML\ & all commands involving ML code \\ \end{tabular} \<^medskip> The Isabelle document preparation system @{cite "isabelle-system"} allows tagged command regions to be presented specifically, e.g.\ to fold proof texts, or drop parts of the text completely. For example ``\<^theory_text>\by auto\~\\<^marker>\tag invisible\\'' causes that piece of proof to be treated as \invisible\ instead of \proof\ (the default), which may be shown or hidden depending on the document setup. In contrast, ``\<^theory_text>\by auto\~\\<^marker>\tag visible\\'' forces this text to be shown invariably. Explicit tag specifications within a proof apply to all subsequent commands of the same level of nesting. For example, ``\<^theory_text>\proof\~\\<^marker>\tag invisible\ \\~\<^theory_text>\qed\'' forces the whole sub-proof to be typeset as \visible\ (unless some of its parts are tagged differently). \<^medskip> Command tags merely produce certain markup environments for type-setting. The meaning of these is determined by {\LaTeX} macros, as defined in \<^file>\~~/lib/texinputs/isabelle.sty\ or by the document author. The Isabelle document preparation tools also provide some high-level options to specify the meaning of arbitrary tags to ``keep'', ``drop'', or ``fold'' the corresponding parts of the text. Logic sessions may also specify ``document versions'', where given tags are interpreted in some particular way. Again see @{cite "isabelle-system"} for further details. \ section \Railroad diagrams\ text \ \begin{matharray}{rcl} @{antiquotation_def "rail"} & : & \antiquotation\ \\ \end{matharray} \<^rail>\ 'rail' @{syntax text} \ The @{antiquotation rail} antiquotation allows to include syntax diagrams into Isabelle documents. {\LaTeX} requires the style file \<^file>\~~/lib/texinputs/railsetup.sty\, which can be used via \<^verbatim>\\usepackage{railsetup}\ in \<^verbatim>\root.tex\, for example. The rail specification language is quoted here as Isabelle @{syntax string} or text @{syntax "cartouche"}; it has its own grammar given below. \begingroup \def\isasymnewline{\isatt{\isacharbackslash\isacharless newline\isachargreater}} \<^rail>\ rule? + ';' ; rule: ((identifier | @{syntax antiquotation}) ':')? body ; body: concatenation + '|' ; concatenation: ((atom '?'?) +) (('*' | '+') atom?)? ; atom: '(' body? ')' | identifier | '@'? (string | @{syntax antiquotation}) | '\' \ \endgroup The lexical syntax of \identifier\ coincides with that of @{syntax short_ident} in regular Isabelle syntax, but \string\ uses single quotes instead of double quotes of the standard @{syntax string} category. Each \rule\ defines a formal language (with optional name), using a notation that is similar to EBNF or regular expressions with recursion. The meaning and visual appearance of these rail language elements is illustrated by the following representative examples. \<^item> Empty \<^verbatim>\()\ \<^rail>\()\ \<^item> Nonterminal \<^verbatim>\A\ \<^rail>\A\ \<^item> Nonterminal via Isabelle antiquotation \<^verbatim>\@{syntax method}\ \<^rail>\@{syntax method}\ \<^item> Terminal \<^verbatim>\'xyz'\ \<^rail>\'xyz'\ \<^item> Terminal in keyword style \<^verbatim>\@'xyz'\ \<^rail>\@'xyz'\ \<^item> Terminal via Isabelle antiquotation \<^verbatim>\@@{method rule}\ \<^rail>\@@{method rule}\ \<^item> Concatenation \<^verbatim>\A B C\ \<^rail>\A B C\ \<^item> Newline inside concatenation \<^verbatim>\A B C \ D E F\ \<^rail>\A B C \ D E F\ \<^item> Variants \<^verbatim>\A | B | C\ \<^rail>\A | B | C\ \<^item> Option \<^verbatim>\A ?\ \<^rail>\A ?\ \<^item> Repetition \<^verbatim>\A *\ \<^rail>\A *\ \<^item> Repetition with separator \<^verbatim>\A * sep\ \<^rail>\A * sep\ \<^item> Strict repetition \<^verbatim>\A +\ \<^rail>\A +\ \<^item> Strict repetition with separator \<^verbatim>\A + sep\ \<^rail>\A + sep\ \ end diff --git a/src/Doc/Isar_Ref/Generic.thy b/src/Doc/Isar_Ref/Generic.thy --- a/src/Doc/Isar_Ref/Generic.thy +++ b/src/Doc/Isar_Ref/Generic.thy @@ -1,1827 +1,1827 @@ (*:maxLineLen=78:*) theory Generic imports Main Base begin chapter \Generic tools and packages \label{ch:gen-tools}\ section \Configuration options \label{sec:config}\ text \ Isabelle/Pure maintains a record of named configuration options within the theory or proof context, with values of type \<^ML_type>\bool\, \<^ML_type>\int\, \<^ML_type>\real\, or \<^ML_type>\string\. Tools may declare options in ML, and then refer to these values (relative to the context). Thus global reference variables are easily avoided. The user may change the value of a configuration option by means of an associated attribute of the same name. This form of context declaration works particularly well with commands such as @{command "declare"} or @{command "using"} like this: \ (*<*)experiment begin(*>*) declare [[show_main_goal = false]] notepad begin note [[show_main_goal = true]] end (*<*)end(*>*) text \ \begin{matharray}{rcll} @{command_def "print_options"} & : & \context \\ \\ \end{matharray} \<^rail>\ @@{command print_options} ('!'?) ; @{syntax name} ('=' ('true' | 'false' | @{syntax int} | @{syntax float} | @{syntax name}))? \ \<^descr> @{command "print_options"} prints the available configuration options, with names, types, and current values; the ``\!\'' option indicates extra verbosity. \<^descr> \name = value\ as an attribute expression modifies the named option, with the syntax of the value depending on the option's type. For \<^ML_type>\bool\ the default value is \true\. Any attempt to change a global option in a local context is ignored. \ section \Basic proof tools\ subsection \Miscellaneous methods and attributes \label{sec:misc-meth-att}\ text \ \begin{matharray}{rcl} @{method_def unfold} & : & \method\ \\ @{method_def fold} & : & \method\ \\ @{method_def insert} & : & \method\ \\[0.5ex] @{method_def erule}\\<^sup>*\ & : & \method\ \\ @{method_def drule}\\<^sup>*\ & : & \method\ \\ @{method_def frule}\\<^sup>*\ & : & \method\ \\ @{method_def intro} & : & \method\ \\ @{method_def elim} & : & \method\ \\ @{method_def fail} & : & \method\ \\ @{method_def succeed} & : & \method\ \\ @{method_def sleep} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method fold} | @@{method unfold} | @@{method insert}) @{syntax thms} ; (@@{method erule} | @@{method drule} | @@{method frule}) ('(' @{syntax nat} ')')? @{syntax thms} ; (@@{method intro} | @@{method elim}) @{syntax thms}? ; @@{method sleep} @{syntax real} \ \<^descr> @{method unfold}~\a\<^sub>1 \ a\<^sub>n\ and @{method fold}~\a\<^sub>1 \ a\<^sub>n\ expand (or fold back) the given definitions throughout all goals; any chained facts provided are inserted into the goal and subject to rewriting as well. Unfolding works in two stages: first, the given equations are used directly for rewriting; second, the equations are passed through the attribute @{attribute_ref abs_def} before rewriting --- to ensure that definitions are fully expanded, regardless of the actual parameters that are provided. \<^descr> @{method insert}~\a\<^sub>1 \ a\<^sub>n\ inserts theorems as facts into all goals of the proof state. Note that current facts indicated for forward chaining are ignored. \<^descr> @{method erule}~\a\<^sub>1 \ a\<^sub>n\, @{method drule}~\a\<^sub>1 \ a\<^sub>n\, and @{method frule}~\a\<^sub>1 \ a\<^sub>n\ are similar to the basic @{method rule} method (see \secref{sec:pure-meth-att}), but apply rules by elim-resolution, destruct-resolution, and forward-resolution, respectively @{cite "isabelle-implementation"}. The optional natural number argument (default 0) specifies additional assumption steps to be performed here. Note that these methods are improper ones, mainly serving for experimentation and tactic script emulation. Different modes of basic rule application are usually expressed in Isar at the proof language level, rather than via implicit proof state manipulations. For example, a proper single-step elimination would be done using the plain @{method rule} method, with forward chaining of current facts. \<^descr> @{method intro} and @{method elim} repeatedly refine some goal by intro- or elim-resolution, after having inserted any chained facts. Exactly the rules given as arguments are taken into account; this allows fine-tuned decomposition of a proof problem, in contrast to common automated tools. \<^descr> @{method fail} yields an empty result sequence; it is the identity of the ``\|\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method succeed} yields a single (unchanged) result; it is the identity of the ``\,\'' method combinator (cf.\ \secref{sec:proof-meth}). \<^descr> @{method sleep}~\s\ succeeds after a real-time delay of \s\ seconds. This is occasionally useful for demonstration and testing purposes. \begin{matharray}{rcl} @{attribute_def tagged} & : & \attribute\ \\ @{attribute_def untagged} & : & \attribute\ \\[0.5ex] @{attribute_def THEN} & : & \attribute\ \\ @{attribute_def unfolded} & : & \attribute\ \\ @{attribute_def folded} & : & \attribute\ \\ @{attribute_def abs_def} & : & \attribute\ \\[0.5ex] @{attribute_def rotated} & : & \attribute\ \\ @{attribute_def (Pure) elim_format} & : & \attribute\ \\ @{attribute_def no_vars}\\<^sup>*\ & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute tagged} @{syntax name} @{syntax name} ; @@{attribute untagged} @{syntax name} ; @@{attribute THEN} ('[' @{syntax nat} ']')? @{syntax thm} ; (@@{attribute unfolded} | @@{attribute folded}) @{syntax thms} ; @@{attribute rotated} @{syntax int}? \ \<^descr> @{attribute tagged}~\name value\ and @{attribute untagged}~\name\ add and remove \<^emph>\tags\ of some theorem. Tags may be any list of string pairs that serve as formal comment. The first string is considered the tag name, the second its value. Note that @{attribute untagged} removes any tags of the same name. \<^descr> @{attribute THEN}~\a\ composes rules by resolution; it resolves with the first premise of \a\ (an alternative position may be also specified). See - also \<^ML_op>\RS\ in @{cite "isabelle-implementation"}. + also \<^ML_infix>\RS\ in @{cite "isabelle-implementation"}. \<^descr> @{attribute unfolded}~\a\<^sub>1 \ a\<^sub>n\ and @{attribute folded}~\a\<^sub>1 \ a\<^sub>n\ expand and fold back again the given definitions throughout a rule. \<^descr> @{attribute abs_def} turns an equation of the form \<^prop>\f x y \ t\ into \<^prop>\f \ \x y. t\, which ensures that @{method simp} steps always expand it. This also works for object-logic equality. \<^descr> @{attribute rotated}~\n\ rotate the premises of a theorem by \n\ (default 1). \<^descr> @{attribute (Pure) elim_format} turns a destruction rule into elimination rule format, by resolving with the rule \<^prop>\PROP A \ (PROP A \ PROP B) \ PROP B\. Note that the Classical Reasoner (\secref{sec:classical}) provides its own version of this operation. \<^descr> @{attribute no_vars} replaces schematic variables by free ones; this is mainly for tuning output of pretty printed theorems. \ subsection \Low-level equational reasoning\ text \ \begin{matharray}{rcl} @{method_def subst} & : & \method\ \\ @{method_def hypsubst} & : & \method\ \\ @{method_def split} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method subst} ('(' 'asm' ')')? \ ('(' (@{syntax nat}+) ')')? @{syntax thm} ; @@{method split} @{syntax thms} \ These methods provide low-level facilities for equational reasoning that are intended for specialized applications only. Normally, single step calculations would be performed in a structured text (see also \secref{sec:calculation}), while the Simplifier methods provide the canonical way for automated normalization (see \secref{sec:simplifier}). \<^descr> @{method subst}~\eq\ performs a single substitution step using rule \eq\, which may be either a meta or object equality. \<^descr> @{method subst}~\(asm) eq\ substitutes in an assumption. \<^descr> @{method subst}~\(i \ j) eq\ performs several substitutions in the conclusion. The numbers \i\ to \j\ indicate the positions to substitute at. Positions are ordered from the top of the term tree moving down from left to right. For example, in \(a + b) + (c + d)\ there are three positions where commutativity of \+\ is applicable: 1 refers to \a + b\, 2 to the whole term, and 3 to \c + d\. If the positions in the list \(i \ j)\ are non-overlapping (e.g.\ \(2 3)\ in \(a + b) + (c + d)\) you may assume all substitutions are performed simultaneously. Otherwise the behaviour of \subst\ is not specified. \<^descr> @{method subst}~\(asm) (i \ j) eq\ performs the substitutions in the assumptions. The positions refer to the assumptions in order from left to right. For example, given in a goal of the form \P (a + b) \ P (c + d) \ \\, position 1 of commutativity of \+\ is the subterm \a + b\ and position 2 is the subterm \c + d\. \<^descr> @{method hypsubst} performs substitution using some assumption; this only works for equations of the form \x = t\ where \x\ is a free or bound variable. \<^descr> @{method split}~\a\<^sub>1 \ a\<^sub>n\ performs single-step case splitting using the given rules. Splitting is performed in the conclusion or some assumption of the subgoal, depending of the structure of the rule. Note that the @{method simp} method already involves repeated application of split rules as declared in the current context, using @{attribute split}, for example. \ section \The Simplifier \label{sec:simplifier}\ text \ The Simplifier performs conditional and unconditional rewriting and uses contextual information: rule declarations in the background theory or local proof context are taken into account, as well as chained facts and subgoal premises (``local assumptions''). There are several general hooks that allow to modify the simplification strategy, or incorporate other proof tools that solve sub-problems, produce rewrite rules on demand etc. The rewriting strategy is always strictly bottom up, except for congruence rules, which are applied while descending into a term. Conditions in conditional rewrite rules are solved recursively before the rewrite rule is applied. The default Simplifier setup of major object logics (HOL, HOLCF, FOL, ZF) makes the Simplifier ready for immediate use, without engaging into the internal structures. Thus it serves as general-purpose proof tool with the main focus on equational reasoning, and a bit more than that. \ subsection \Simplification methods \label{sec:simp-meth}\ text \ \begin{tabular}{rcll} @{method_def simp} & : & \method\ \\ @{method_def simp_all} & : & \method\ \\ \Pure.\@{method_def (Pure) simp} & : & \method\ \\ \Pure.\@{method_def (Pure) simp_all} & : & \method\ \\ @{attribute_def simp_depth_limit} & : & \attribute\ & default \100\ \\ \end{tabular} \<^medskip> \<^rail>\ (@@{method simp} | @@{method simp_all}) opt? (@{syntax simpmod} * ) ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use' | 'asm_lr' ) ')' ; @{syntax_def simpmod}: ('add' | 'del' | 'flip' | 'only' | 'split' (() | '!' | 'del') | 'cong' (() | 'add' | 'del')) ':' @{syntax thms} \ \<^descr> @{method simp} invokes the Simplifier on the first subgoal, after inserting chained facts as additional goal premises; further rule declarations may be included via \(simp add: facts)\. The proof method fails if the subgoal remains unchanged after simplification. Note that the original goal premises and chained facts are subject to simplification themselves, while declarations via \add\/\del\ merely follow the policies of the object-logic to extract rewrite rules from theorems, without further simplification. This may lead to slightly different behavior in either case, which might be required precisely like that in some boundary situations to perform the intended simplification step! \<^medskip> Modifier \flip\ deletes the following theorems from the simpset and adds their symmetric version (i.e.\ lhs and rhs exchanged). No warning is shown if the original theorem was not present. \<^medskip> The \only\ modifier first removes all other rewrite rules, looper tactics (including split rules), congruence rules, and then behaves like \add\. Implicit solvers remain, which means that trivial rules like reflexivity or introduction of \True\ are available to solve the simplified subgoals, but also non-trivial tools like linear arithmetic in HOL. The latter may lead to some surprise of the meaning of ``only'' in Isabelle/HOL compared to English! \<^medskip> The \split\ modifiers add or delete rules for the Splitter (see also \secref{sec:simp-strategies} on the looper). This works only if the Simplifier method has been properly setup to include the Splitter (all major object logics such HOL, HOLCF, FOL, ZF do this already). The \!\ option causes the split rules to be used aggressively: after each application of a split rule in the conclusion, the \safe\ tactic of the classical reasoner (see \secref{sec:classical:partial}) is applied to the new goal. The net effect is that the goal is split into the different cases. This option can speed up simplification of goals with many nested conditional or case expressions significantly. There is also a separate @{method_ref split} method available for single-step case splitting. The effect of repeatedly applying \(split thms)\ can be imitated by ``\(simp only: split: thms)\''. \<^medskip> The \cong\ modifiers add or delete Simplifier congruence rules (see also \secref{sec:simp-rules}); the default is to add. \<^descr> @{method simp_all} is similar to @{method simp}, but acts on all goals, working backwards from the last to the first one as usual in Isabelle.\<^footnote>\The order is irrelevant for goals without schematic variables, so simplification might actually be performed in parallel here.\ Chained facts are inserted into all subgoals, before the simplification process starts. Further rule declarations are the same as for @{method simp}. The proof method fails if all subgoals remain unchanged after simplification. \<^descr> @{attribute simp_depth_limit} limits the number of recursive invocations of the Simplifier during conditional rewriting. By default the Simplifier methods above take local assumptions fully into account, using equational assumptions in the subsequent normalization process, or simplifying assumptions themselves. Further options allow to fine-tune the behavior of the Simplifier in this respect, corresponding to a variety of ML tactics as follows.\<^footnote>\Unlike the corresponding Isar proof methods, the ML tactics do not insist in changing the goal state.\ \begin{center} \small \begin{tabular}{|l|l|p{0.3\textwidth}|} \hline Isar method & ML tactic & behavior \\\hline \(simp (no_asm))\ & \<^ML>\simp_tac\ & assumptions are ignored completely \\\hline \(simp (no_asm_simp))\ & \<^ML>\asm_simp_tac\ & assumptions are used in the simplification of the conclusion but are not themselves simplified \\\hline \(simp (no_asm_use))\ & \<^ML>\full_simp_tac\ & assumptions are simplified but are not used in the simplification of each other or the conclusion \\\hline \(simp)\ & \<^ML>\asm_full_simp_tac\ & assumptions are used in the simplification of the conclusion and to simplify other assumptions \\\hline \(simp (asm_lr))\ & \<^ML>\asm_lr_simp_tac\ & compatibility mode: an assumption is only used for simplifying assumptions which are to the right of it \\\hline \end{tabular} \end{center} \<^medskip> In Isabelle/Pure, proof methods @{method (Pure) simp} and @{method (Pure) simp_all} only know about meta-equality \\\. Any new object-logic needs to re-define these methods via \<^ML>\Simplifier.method_setup\ in ML: Isabelle/FOL or Isabelle/HOL may serve as blue-prints. \ subsubsection \Examples\ text \ We consider basic algebraic simplifications in Isabelle/HOL. The rather trivial goal \<^prop>\0 + (x + 0) = x + 0 + 0\ looks like a good candidate to be solved by a single call of @{method simp}: \ lemma "0 + (x + 0) = x + 0 + 0" apply simp? oops text \ The above attempt \<^emph>\fails\, because \<^term>\0\ and \<^term>\(+)\ in the HOL library are declared as generic type class operations, without stating any algebraic laws yet. More specific types are required to get access to certain standard simplifications of the theory context, e.g.\ like this:\ lemma fixes x :: nat shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: int shows "0 + (x + 0) = x + 0 + 0" by simp lemma fixes x :: "'a :: monoid_add" shows "0 + (x + 0) = x + 0 + 0" by simp text \ \<^medskip> In many cases, assumptions of a subgoal are also needed in the simplification process. For example: \ lemma fixes x :: nat shows "x = 0 \ x + x = 0" by simp lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" apply simp oops lemma fixes x :: nat assumes "x = 0" shows "x + x = 0" using assms by simp text \ As seen above, local assumptions that shall contribute to simplification need to be part of the subgoal already, or indicated explicitly for use by the subsequent method invocation. Both too little or too much information can make simplification fail, for different reasons. In the next example the malicious assumption \<^prop>\\x::nat. f x = g (f (g x))\ does not contribute to solve the problem, but makes the default @{method simp} method loop: the rewrite rule \f ?x \ g (f (g ?x))\ extracted from the assumption does not terminate. The Simplifier notices certain simple forms of nontermination, but not this one. The problem can be solved nonetheless, by ignoring assumptions via special options as explained before: \ lemma "(\x::nat. f x = g (f (g x))) \ f 0 = f 0 + 0" by (simp (no_asm)) text \ The latter form is typical for long unstructured proof scripts, where the control over the goal content is limited. In structured proofs it is usually better to avoid pushing too many facts into the goal state in the first place. Assumptions in the Isar proof context do not intrude the reasoning if not used explicitly. This is illustrated for a toplevel statement and a local proof body as follows: \ lemma assumes "\x::nat. f x = g (f (g x))" shows "f 0 = f 0 + 0" by simp notepad begin assume "\x::nat. f x = g (f (g x))" have "f 0 = f 0 + 0" by simp end text \ \<^medskip> Because assumptions may simplify each other, there can be very subtle cases of nontermination. For example, the regular @{method simp} method applied to \<^prop>\P (f x) \ y = x \ f x = f y \ Q\ gives rise to the infinite reduction sequence \[ \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \P (f y)\ \stackrel{\y \ x\}{\longmapsto} \P (f x)\ \stackrel{\f x \ f y\}{\longmapsto} \cdots \] whereas applying the same to \<^prop>\y = x \ f x = f y \ P (f x) \ Q\ terminates (without solving the goal): \ lemma "y = x \ f x = f y \ P (f x) \ Q" apply simp oops text \ See also \secref{sec:simp-trace} for options to enable Simplifier trace mode, which often helps to diagnose problems with rewrite systems. \ subsection \Declaring rules \label{sec:simp-rules}\ text \ \begin{matharray}{rcl} @{attribute_def simp} & : & \attribute\ \\ @{attribute_def split} & : & \attribute\ \\ @{attribute_def cong} & : & \attribute\ \\ @{command_def "print_simpset"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} \<^rail>\ (@@{attribute simp} | @@{attribute cong}) (() | 'add' | 'del') | @@{attribute split} (() | '!' | 'del') ; @@{command print_simpset} ('!'?) \ \<^descr> @{attribute simp} declares rewrite rules, by adding or deleting them from the simpset within the theory or proof context. Rewrite rules are theorems expressing some form of equality, for example: \Suc ?m + ?n = ?m + Suc ?n\ \\ \?P \ ?P \ ?P\ \\ \?A \ ?B \ {x. x \ ?A \ x \ ?B}\ \<^medskip> Conditional rewrites such as \?m < ?n \ ?m div ?n = 0\ are also permitted; the conditions can be arbitrary formulas. \<^medskip> Internally, all rewrite rules are translated into Pure equalities, theorems with conclusion \lhs \ rhs\. The simpset contains a function for extracting equalities from arbitrary theorems, which is usually installed when the object-logic is configured initially. For example, \\ ?x \ {}\ could be turned into \?x \ {} \ False\. Theorems that are declared as @{attribute simp} and local assumptions within a goal are treated uniformly in this respect. The Simplifier accepts the following formats for the \lhs\ term: \<^enum> First-order patterns, considering the sublanguage of application of constant operators to variable operands, without \\\-abstractions or functional variables. For example: \(?x + ?y) + ?z \ ?x + (?y + ?z)\ \\ \f (f ?x ?y) ?z \ f ?x (f ?y ?z)\ \<^enum> Higher-order patterns in the sense of @{cite "nipkow-patterns"}. These are terms in \\\-normal form (this will always be the case unless you have done something strange) where each occurrence of an unknown is of the form \?F x\<^sub>1 \ x\<^sub>n\, where the \x\<^sub>i\ are distinct bound variables. For example, \(\x. ?P x \ ?Q x) \ (\x. ?P x) \ (\x. ?Q x)\ or its symmetric form, since the \rhs\ is also a higher-order pattern. \<^enum> Physical first-order patterns over raw \\\-term structure without \\\\\-equality; abstractions and bound variables are treated like quasi-constant term material. For example, the rule \?f ?x \ range ?f = True\ rewrites the term \g a \ range g\ to \True\, but will fail to match \g (h b) \ range (\x. g (h x))\. However, offending subterms (in our case \?f ?x\, which is not a pattern) can be replaced by adding new variables and conditions like this: \?y = ?f ?x \ ?y \ range ?f = True\ is acceptable as a conditional rewrite rule of the second category since conditions can be arbitrary terms. \<^descr> @{attribute split} declares case split rules. \<^descr> @{attribute cong} declares congruence rules to the Simplifier context. Congruence rules are equalities of the form @{text [display] "\ \ f ?x\<^sub>1 \ ?x\<^sub>n = f ?y\<^sub>1 \ ?y\<^sub>n"} This controls the simplification of the arguments of \f\. For example, some arguments can be simplified under additional assumptions: @{text [display] "?P\<^sub>1 \ ?Q\<^sub>1 \ (?Q\<^sub>1 \ ?P\<^sub>2 \ ?Q\<^sub>2) \ (?P\<^sub>1 \ ?P\<^sub>2) \ (?Q\<^sub>1 \ ?Q\<^sub>2)"} Given this rule, the Simplifier assumes \?Q\<^sub>1\ and extracts rewrite rules from it when simplifying \?P\<^sub>2\. Such local assumptions are effective for rewriting formulae such as \x = 0 \ y + x = y\. %FIXME %The local assumptions are also provided as theorems to the solver; %see \secref{sec:simp-solver} below. \<^medskip> The following congruence rule for bounded quantifiers also supplies contextual information --- about the bound variable: @{text [display] "(?A = ?B) \ (\x. x \ ?B \ ?P x \ ?Q x) \ (\x \ ?A. ?P x) \ (\x \ ?B. ?Q x)"} \<^medskip> This congruence rule for conditional expressions can supply contextual information for simplifying the arms: @{text [display] "?p = ?q \ (?q \ ?a = ?c) \ (\ ?q \ ?b = ?d) \ (if ?p then ?a else ?b) = (if ?q then ?c else ?d)"} A congruence rule can also \<^emph>\prevent\ simplification of some arguments. Here is an alternative congruence rule for conditional expressions that conforms to non-strict functional evaluation: @{text [display] "?p = ?q \ (if ?p then ?a else ?b) = (if ?q then ?a else ?b)"} Only the first argument is simplified; the others remain unchanged. This can make simplification much faster, but may require an extra case split over the condition \?q\ to prove the goal. \<^descr> @{command "print_simpset"} prints the collection of rules declared to the Simplifier, which is also known as ``simpset'' internally; the ``\!\'' option indicates extra verbosity. The implicit simpset of the theory context is propagated monotonically through the theory hierarchy: forming a new theory, the union of the simpsets of its imports are taken as starting point. Also note that definitional packages like @{command "datatype"}, @{command "primrec"}, @{command "fun"} routinely declare Simplifier rules to the target context, while plain @{command "definition"} is an exception in \<^emph>\not\ declaring anything. \<^medskip> It is up the user to manipulate the current simpset further by explicitly adding or deleting theorems as simplification rules, or installing other tools via simplification procedures (\secref{sec:simproc}). Good simpsets are hard to design. Rules that obviously simplify, like \?n + 0 \ ?n\ are good candidates for the implicit simpset, unless a special non-normalizing behavior of certain operations is intended. More specific rules (such as distributive laws, which duplicate subterms) should be added only for specific proof steps. Conversely, sometimes a rule needs to be deleted just for some part of a proof. The need of frequent additions or deletions may indicate a poorly designed simpset. \begin{warn} The union of simpsets from theory imports (as described above) is not always a good starting point for the new theory. If some ancestors have deleted simplification rules because they are no longer wanted, while others have left those rules in, then the union will contain the unwanted rules, and thus have to be deleted again in the theory body. \end{warn} \ subsection \Ordered rewriting with permutative rules\ text \ A rewrite rule is \<^emph>\permutative\ if the left-hand side and right-hand side are the equal up to renaming of variables. The most common permutative rule is commutativity: \?x + ?y = ?y + ?x\. Other examples include \(?x - ?y) - ?z = (?x - ?z) - ?y\ in arithmetic and \insert ?x (insert ?y ?A) = insert ?y (insert ?x ?A)\ for sets. Such rules are common enough to merit special attention. Because ordinary rewriting loops given such rules, the Simplifier employs a special strategy, called \<^emph>\ordered rewriting\. Permutative rules are detected and only applied if the rewriting step decreases the redex wrt.\ a given term ordering. For example, commutativity rewrites \b + a\ to \a + b\, but then stops, because the redex cannot be decreased further in the sense of the term ordering. The default is lexicographic ordering of term structure, but this could be also changed locally for special applications via @{define_ML Simplifier.set_term_ord} in Isabelle/ML. \<^medskip> Permutative rewrite rules are declared to the Simplifier just like other rewrite rules. Their special status is recognized automatically, and their application is guarded by the term ordering accordingly. \ subsubsection \Rewriting with AC operators\ text \ Ordered rewriting is particularly effective in the case of associative-commutative operators. (Associativity by itself is not permutative.) When dealing with an AC-operator \f\, keep the following points in mind: \<^item> The associative law must always be oriented from left to right, namely \f (f x y) z = f x (f y z)\. The opposite orientation, if used with commutativity, leads to looping in conjunction with the standard term order. \<^item> To complete your set of rewrite rules, you must add not just associativity (A) and commutativity (C) but also a derived rule \<^emph>\left-commutativity\ (LC): \f x (f y z) = f y (f x z)\. Ordered rewriting with the combination of A, C, and LC sorts a term lexicographically --- the rewriting engine imitates bubble-sort. \ experiment fixes f :: "'a \ 'a \ 'a" (infix "\" 60) assumes assoc: "(x \ y) \ z = x \ (y \ z)" assumes commute: "x \ y = y \ x" begin lemma left_commute: "x \ (y \ z) = y \ (x \ z)" proof - have "(x \ y) \ z = (y \ x) \ z" by (simp only: commute) then show ?thesis by (simp only: assoc) qed lemmas AC_rules = assoc commute left_commute text \ Thus the Simplifier is able to establish equalities with arbitrary permutations of subterms, by normalizing to a common standard form. For example: \ lemma "(b \ c) \ a = xxx" apply (simp only: AC_rules) txt \\<^subgoals>\ oops lemma "(b \ c) \ a = a \ (b \ c)" by (simp only: AC_rules) lemma "(b \ c) \ a = c \ (b \ a)" by (simp only: AC_rules) lemma "(b \ c) \ a = (c \ b) \ a" by (simp only: AC_rules) end text \ Martin and Nipkow @{cite "martin-nipkow"} discuss the theory and give many examples; other algebraic structures are amenable to ordered rewriting, such as Boolean rings. The Boyer-Moore theorem prover @{cite bm88book} also employs ordered rewriting. \ subsubsection \Re-orienting equalities\ text \Another application of ordered rewriting uses the derived rule @{thm [source] eq_commute}: @{thm [source = false] eq_commute} to reverse equations. This is occasionally useful to re-orient local assumptions according to the term ordering, when other built-in mechanisms of reorientation and mutual simplification fail to apply.\ subsection \Simplifier tracing and debugging \label{sec:simp-trace}\ text \ \begin{tabular}{rcll} @{attribute_def simp_trace} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_depth_limit} & : & \attribute\ & default \1\ \\ @{attribute_def simp_debug} & : & \attribute\ & default \false\ \\ @{attribute_def simp_trace_new} & : & \attribute\ \\ @{attribute_def simp_break} & : & \attribute\ \\ \end{tabular} \<^medskip> \<^rail>\ @@{attribute simp_trace_new} ('interactive')? \ ('mode' '=' ('full' | 'normal'))? \ ('depth' '=' @{syntax nat})? ; @@{attribute simp_break} (@{syntax term}*) \ These attributes and configurations options control various aspects of Simplifier tracing and debugging. \<^descr> @{attribute simp_trace} makes the Simplifier output internal operations. This includes rewrite steps, but also bookkeeping like modifications of the simpset. \<^descr> @{attribute simp_trace_depth_limit} limits the effect of @{attribute simp_trace} to the given depth of recursive Simplifier invocations (when solving conditions of rewrite rules). \<^descr> @{attribute simp_debug} makes the Simplifier output some extra information about internal operations. This includes any attempted invocation of simplification procedures. \<^descr> @{attribute simp_trace_new} controls Simplifier tracing within Isabelle/PIDE applications, notably Isabelle/jEdit @{cite "isabelle-jedit"}. This provides a hierarchical representation of the rewriting steps performed by the Simplifier. Users can configure the behaviour by specifying breakpoints, verbosity and enabling or disabling the interactive mode. In normal verbosity (the default), only rule applications matching a breakpoint will be shown to the user. In full verbosity, all rule applications will be logged. Interactive mode interrupts the normal flow of the Simplifier and defers the decision how to continue to the user via some GUI dialog. \<^descr> @{attribute simp_break} declares term or theorem breakpoints for @{attribute simp_trace_new} as described above. Term breakpoints are patterns which are checked for matches on the redex of a rule application. Theorem breakpoints trigger when the corresponding theorem is applied in a rewrite step. For example: \ (*<*)experiment begin(*>*) declare conjI [simp_break] declare [[simp_break "?x \ ?y"]] (*<*)end(*>*) subsection \Simplification procedures \label{sec:simproc}\ text \ Simplification procedures are ML functions that produce proven rewrite rules on demand. They are associated with higher-order patterns that approximate the left-hand sides of equations. The Simplifier first matches the current redex against one of the LHS patterns; if this succeeds, the corresponding ML function is invoked, passing the Simplifier context and redex term. Thus rules may be specifically fashioned for particular situations, resulting in a more powerful mechanism than term rewriting by a fixed set of rules. Any successful result needs to be a (possibly conditional) rewrite rule \t \ u\ that is applicable to the current redex. The rule will be applied just as any ordinary rewrite rule. It is expected to be already in \<^emph>\internal form\, bypassing the automatic preprocessing of object-level equivalences. \begin{matharray}{rcl} @{command_def "simproc_setup"} & : & \local_theory \ local_theory\ \\ simproc & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{command simproc_setup} @{syntax name} '(' (@{syntax term} + '|') ')' '=' @{syntax text}; @@{attribute simproc} (('add' ':')? | 'del' ':') (@{syntax name}+) \ \<^descr> @{command "simproc_setup"} defines a named simplification procedure that is invoked by the Simplifier whenever any of the given term patterns match the current redex. The implementation, which is provided as ML source text, needs to be of type \<^ML_type>\morphism -> Proof.context -> cterm -> thm option\, where the \<^ML_type>\cterm\ represents the current redex \r\ and the result is supposed to be some proven rewrite rule \r \ r'\ (or a generalized version), or \<^ML>\NONE\ to indicate failure. The \<^ML_type>\Proof.context\ argument holds the full context of the current Simplifier invocation. The \<^ML_type>\morphism\ informs about the difference of the original compilation context wrt.\ the one of the actual application later on. Morphisms are only relevant for simprocs that are defined within a local target context, e.g.\ in a locale. \<^descr> \simproc add: name\ and \simproc del: name\ add or delete named simprocs to the current Simplifier context. The default is to add a simproc. Note that @{command "simproc_setup"} already adds the new simproc to the subsequent context. \ subsubsection \Example\ text \ The following simplification procedure for @{thm [source = false, show_types] unit_eq} in HOL performs fine-grained control over rule application, beyond higher-order pattern matching. Declaring @{thm unit_eq} as @{attribute simp} directly would make the Simplifier loop! Note that a version of this simplification procedure is already active in Isabelle/HOL. \ (*<*)experiment begin(*>*) simproc_setup unit ("x::unit") = \fn _ => fn _ => fn ct => if HOLogic.is_unit (Thm.term_of ct) then NONE else SOME (mk_meta_eq @{thm unit_eq})\ (*<*)end(*>*) text \ Since the Simplifier applies simplification procedures frequently, it is important to make the failure check in ML reasonably fast.\ subsection \Configurable Simplifier strategies \label{sec:simp-strategies}\ text \ The core term-rewriting engine of the Simplifier is normally used in combination with some add-on components that modify the strategy and allow to integrate other non-Simplifier proof tools. These may be reconfigured in ML as explained below. Even if the default strategies of object-logics like Isabelle/HOL are used unchanged, it helps to understand how the standard Simplifier strategies work.\ subsubsection \The subgoaler\ text \ \begin{mldecls} @{define_ML Simplifier.set_subgoaler: "(Proof.context -> int -> tactic) -> Proof.context -> Proof.context"} \\ @{define_ML Simplifier.prems_of: "Proof.context -> thm list"} \\ \end{mldecls} The subgoaler is the tactic used to solve subgoals arising out of conditional rewrite rules or congruence rules. The default should be simplification itself. In rare situations, this strategy may need to be changed. For example, if the premise of a conditional rule is an instance of its conclusion, as in \Suc ?m < ?n \ ?m < ?n\, the default strategy could loop. % FIXME !?? \<^descr> \<^ML>\Simplifier.set_subgoaler\~\tac ctxt\ sets the subgoaler of the context to \tac\. The tactic will be applied to the context of the running Simplifier instance. \<^descr> \<^ML>\Simplifier.prems_of\~\ctxt\ retrieves the current set of premises from the context. This may be non-empty only if the Simplifier has been told to utilize local assumptions in the first place (cf.\ the options in \secref{sec:simp-meth}). As an example, consider the following alternative subgoaler: \ ML_val \ fun subgoaler_tac ctxt = assume_tac ctxt ORELSE' resolve_tac ctxt (Simplifier.prems_of ctxt) ORELSE' asm_simp_tac ctxt \ text \ This tactic first tries to solve the subgoal by assumption or by resolving with with one of the premises, calling simplification only if that fails.\ subsubsection \The solver\ text \ \begin{mldecls} @{define_ML_type solver} \\ @{define_ML Simplifier.mk_solver: "string -> (Proof.context -> int -> tactic) -> solver"} \\ @{define_ML_infix setSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix setSSolver: "Proof.context * solver -> Proof.context"} \\ @{define_ML_infix addSSolver: "Proof.context * solver -> Proof.context"} \\ \end{mldecls} A solver is a tactic that attempts to solve a subgoal after simplification. Its core functionality is to prove trivial subgoals such as \<^prop>\True\ and \t = t\, but object-logics might be more ambitious. For example, Isabelle/HOL performs a restricted version of linear arithmetic here. Solvers are packaged up in abstract type \<^ML_type>\solver\, with \<^ML>\Simplifier.mk_solver\ as the only operation to create a solver. \<^medskip> Rewriting does not instantiate unknowns. For example, rewriting alone cannot prove \a \ ?A\ since this requires instantiating \?A\. The solver, however, is an arbitrary tactic and may instantiate unknowns as it pleases. This is the only way the Simplifier can handle a conditional rewrite rule whose condition contains extra variables. When a simplification tactic is to be combined with other provers, especially with the Classical Reasoner, it is important whether it can be considered safe or not. For this reason a simpset contains two solvers: safe and unsafe. The standard simplification strategy solely uses the unsafe solver, which is appropriate in most cases. For special applications where the simplification process is not allowed to instantiate unknowns within the goal, simplification starts with the safe solver, but may still apply the ordinary unsafe one in nested simplifications for conditional rules or congruences. Note that in this way the overall tactic is not totally safe: it may instantiate unknowns that appear also in other subgoals. \<^descr> \<^ML>\Simplifier.mk_solver\~\name tac\ turns \tac\ into a solver; the \name\ is only attached as a comment and has no further significance. \<^descr> \ctxt setSSolver solver\ installs \solver\ as the safe solver of \ctxt\. \<^descr> \ctxt addSSolver solver\ adds \solver\ as an additional safe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^descr> \ctxt setSolver solver\ installs \solver\ as the unsafe solver of \ctxt\. \<^descr> \ctxt addSolver solver\ adds \solver\ as an additional unsafe solver; it will be tried after the solvers which had already been present in \ctxt\. \<^medskip> The solver tactic is invoked with the context of the running Simplifier. Further operations may be used to retrieve relevant information, such as the list of local Simplifier premises via \<^ML>\Simplifier.prems_of\ --- this list may be non-empty only if the Simplifier runs in a mode that utilizes local assumptions (see also \secref{sec:simp-meth}). The solver is also presented the full goal including its assumptions in any case. Thus it can use these (e.g.\ by calling \<^ML>\assume_tac\), even if the Simplifier proper happens to ignore local premises at the moment. \<^medskip> As explained before, the subgoaler is also used to solve the premises of congruence rules. These are usually of the form \s = ?x\, where \s\ needs to be simplified and \?x\ needs to be instantiated with the result. Typically, the subgoaler will invoke the Simplifier at some point, which will eventually call the solver. For this reason, solver tactics must be prepared to solve goals of the form \t = ?x\, usually by reflexivity. In particular, reflexivity should be tried before any of the fancy automated proof tools. It may even happen that due to simplification the subgoal is no longer an equality. For example, \False \ ?Q\ could be rewritten to \\ ?Q\. To cover this case, the solver could try resolving with the theorem \\ False\ of the object-logic. \<^medskip> \begin{warn} If a premise of a congruence rule cannot be proved, then the congruence is ignored. This should only happen if the rule is \<^emph>\conditional\ --- that is, contains premises not of the form \t = ?x\. Otherwise it indicates that some congruence rule, or possibly the subgoaler or solver, is faulty. \end{warn} \ subsubsection \The looper\ text \ \begin{mldecls} @{define_ML_infix setloop: "Proof.context * (Proof.context -> int -> tactic) -> Proof.context"} \\ @{define_ML_infix addloop: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delloop: "Proof.context * string -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split: "thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.add_split_bang: " thm -> Proof.context -> Proof.context"} \\ @{define_ML Splitter.del_split: "thm -> Proof.context -> Proof.context"} \\ \end{mldecls} The looper is a list of tactics that are applied after simplification, in case the solver failed to solve the simplified goal. If the looper succeeds, the simplification process is started all over again. Each of the subgoals generated by the looper is attacked in turn, in reverse order. A typical looper is \<^emph>\case splitting\: the expansion of a conditional. Another possibility is to apply an elimination rule on the assumptions. More adventurous loopers could start an induction. \<^descr> \ctxt setloop tac\ installs \tac\ as the only looper tactic of \ctxt\. \<^descr> \ctxt addloop (name, tac)\ adds \tac\ as an additional looper tactic with name \name\, which is significant for managing the collection of loopers. The tactic will be tried after the looper tactics that had already been present in \ctxt\. \<^descr> \ctxt delloop name\ deletes the looper tactic that was associated with \name\ from \ctxt\. \<^descr> \<^ML>\Splitter.add_split\~\thm ctxt\ adds split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.add_split_bang\~\thm ctxt\ adds aggressive (see \S\ref{sec:simp-meth}) split tactic for \thm\ as additional looper tactic of \ctxt\ (overwriting previous split tactic for the same constant). \<^descr> \<^ML>\Splitter.del_split\~\thm ctxt\ deletes the split tactic corresponding to \thm\ from the looper tactics of \ctxt\. The splitter replaces applications of a given function; the right-hand side of the replacement can be anything. For example, here is a splitting rule for conditional expressions: @{text [display] "?P (if ?Q ?x ?y) \ (?Q \ ?P ?x) \ (\ ?Q \ ?P ?y)"} Another example is the elimination operator for Cartesian products (which happens to be called \<^const>\case_prod\ in Isabelle/HOL: @{text [display] "?P (case_prod ?f ?p) \ (\a b. ?p = (a, b) \ ?P (f a b))"} For technical reasons, there is a distinction between case splitting in the conclusion and in the premises of a subgoal. The former is done by \<^ML>\Splitter.split_tac\ with rules like @{thm [source] if_split} or @{thm [source] option.split}, which do not split the subgoal, while the latter is done by \<^ML>\Splitter.split_asm_tac\ with rules like @{thm [source] if_split_asm} or @{thm [source] option.split_asm}, which split the subgoal. The function \<^ML>\Splitter.add_split\ automatically takes care of which tactic to call, analyzing the form of the rules given as argument; it is the same operation behind \split\ attribute or method modifier syntax in the Isar source language. Case splits should be allowed only when necessary; they are expensive and hard to control. Case-splitting on if-expressions in the conclusion is usually beneficial, so it is enabled by default in Isabelle/HOL and Isabelle/FOL/ZF. \begin{warn} With \<^ML>\Splitter.split_asm_tac\ as looper component, the Simplifier may split subgoals! This might cause unexpected problems in tactic expressions that silently assume 0 or 1 subgoals after simplification. \end{warn} \ subsection \Forward simplification \label{sec:simp-forward}\ text \ \begin{matharray}{rcl} @{attribute_def simplified} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute simplified} opt? @{syntax thms}? ; opt: '(' ('no_asm' | 'no_asm_simp' | 'no_asm_use') ')' \ \<^descr> @{attribute simplified}~\a\<^sub>1 \ a\<^sub>n\ causes a theorem to be simplified, either by exactly the specified rules \a\<^sub>1, \, a\<^sub>n\, or the implicit Simplifier context if no arguments are given. The result is fully simplified by default, including assumptions and conclusion; the options \no_asm\ etc.\ tune the Simplifier in the same way as the for the \simp\ method. Note that forward simplification restricts the Simplifier to its most basic operation of term rewriting; solver and looper tactics (\secref{sec:simp-strategies}) are \<^emph>\not\ involved here. The @{attribute simplified} attribute should be only rarely required under normal circumstances. \ section \The Classical Reasoner \label{sec:classical}\ subsection \Basic concepts\ text \Although Isabelle is generic, many users will be working in some extension of classical first-order logic. Isabelle/ZF is built upon theory FOL, while Isabelle/HOL conceptually contains first-order logic as a fragment. Theorem-proving in predicate logic is undecidable, but many automated strategies have been developed to assist in this task. Isabelle's classical reasoner is a generic package that accepts certain information about a logic and delivers a suite of automatic proof tools, based on rules that are classified and declared in the context. These proof procedures are slow and simplistic compared with high-end automated theorem provers, but they can save considerable time and effort in practice. They can prove theorems such as Pelletier's @{cite pelletier86} problems 40 and 41 in a few milliseconds (including full proof reconstruction):\ lemma "(\y. \x. F x y \ F x x) \ \ (\x. \y. \z. F z y \ \ F z x)" by blast lemma "(\z. \y. \x. f x y \ f x z \ \ f x x) \ \ (\z. \x. f x z)" by blast text \The proof tools are generic. They are not restricted to first-order logic, and have been heavily used in the development of the Isabelle/HOL library and applications. The tactics can be traced, and their components can be called directly; in this manner, any proof can be viewed interactively.\ subsubsection \The sequent calculus\ text \Isabelle supports natural deduction, which is easy to use for interactive proof. But natural deduction does not easily lend itself to automation, and has a bias towards intuitionism. For certain proofs in classical logic, it can not be called natural. The \<^emph>\sequent calculus\, a generalization of natural deduction, is easier to automate. A \<^bold>\sequent\ has the form \\ \ \\, where \\\ and \\\ are sets of formulae.\<^footnote>\For first-order logic, sequents can equivalently be made from lists or multisets of formulae.\ The sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ is \<^bold>\valid\ if \P\<^sub>1 \ \ \ P\<^sub>m\ implies \Q\<^sub>1 \ \ \ Q\<^sub>n\. Thus \P\<^sub>1, \, P\<^sub>m\ represent assumptions, each of which is true, while \Q\<^sub>1, \, Q\<^sub>n\ represent alternative goals. A sequent is \<^bold>\basic\ if its left and right sides have a common formula, as in \P, Q \ Q, R\; basic sequents are trivially valid. Sequent rules are classified as \<^bold>\right\ or \<^bold>\left\, indicating which side of the \\\ symbol they operate on. Rules that operate on the right side are analogous to natural deduction's introduction rules, and left rules are analogous to elimination rules. The sequent calculus analogue of \(\I)\ is the rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\P, \ \ \, Q\} \] Applying the rule backwards, this breaks down some implication on the right side of a sequent; \\\ and \\\ stand for the sets of formulae that are unaffected by the inference. The analogue of the pair \(\I1)\ and \(\I2)\ is the single rule \[ \infer[\(\R)\]{\\ \ \, P \ Q\}{\\ \ \, P, Q\} \] This breaks down some disjunction on the right side, replacing it by both disjuncts. Thus, the sequent calculus is a kind of multiple-conclusion logic. To illustrate the use of multiple formulae on the right, let us prove the classical theorem \(P \ Q) \ (Q \ P)\. Working backwards, we reduce this formula to a basic sequent: \[ \infer[\(\R)\]{\\ (P \ Q) \ (Q \ P)\} {\infer[\(\R)\]{\\ (P \ Q), (Q \ P)\} {\infer[\(\R)\]{\P \ Q, (Q \ P)\} {\P, Q \ Q, P\}}} \] This example is typical of the sequent calculus: start with the desired theorem and apply rules backwards in a fairly arbitrary manner. This yields a surprisingly effective proof procedure. Quantifiers add only few complications, since Isabelle handles parameters and schematic variables. See @{cite \Chapter 10\ "paulson-ml2"} for further discussion.\ subsubsection \Simulating sequents by natural deduction\ text \Isabelle can represent sequents directly, as in the object-logic LK. But natural deduction is easier to work with, and most object-logics employ it. Fortunately, we can simulate the sequent \P\<^sub>1, \, P\<^sub>m \ Q\<^sub>1, \, Q\<^sub>n\ by the Isabelle formula \P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>2 \ ... \ \ Q\<^sub>n \ Q\<^sub>1\ where the order of the assumptions and the choice of \Q\<^sub>1\ are arbitrary. Elim-resolution plays a key role in simulating sequent proofs. We can easily handle reasoning on the left. Elim-resolution with the rules \(\E)\, \(\E)\ and \(\E)\ achieves a similar effect as the corresponding sequent rules. For the other connectives, we use sequent-style elimination rules instead of destruction rules such as \(\E1, 2)\ and \(\E)\. But note that the rule \(\L)\ has no effect under our representation of sequents! \[ \infer[\(\L)\]{\\ P, \ \ \\}{\\ \ \, P\} \] What about reasoning on the right? Introduction rules can only affect the formula in the conclusion, namely \Q\<^sub>1\. The other right-side formulae are represented as negated assumptions, \\ Q\<^sub>2, \, \ Q\<^sub>n\. In order to operate on one of these, it must first be exchanged with \Q\<^sub>1\. Elim-resolution with the \swap\ rule has this effect: \\ P \ (\ R \ P) \ R\ To ensure that swaps occur only when necessary, each introduction rule is converted into a swapped form: it is resolved with the second premise of \(swap)\. The swapped form of \(\I)\, which might be called \(\\E)\, is @{text [display] "\ (P \ Q) \ (\ R \ P) \ (\ R \ Q) \ R"} Similarly, the swapped form of \(\I)\ is @{text [display] "\ (P \ Q) \ (\ R \ P \ Q) \ R"} Swapped introduction rules are applied using elim-resolution, which deletes the negated formula. Our representation of sequents also requires the use of ordinary introduction rules. If we had no regard for readability of intermediate goal states, we could treat the right side more uniformly by representing sequents as @{text [display] "P\<^sub>1 \ \ \ P\<^sub>m \ \ Q\<^sub>1 \ \ \ \ Q\<^sub>n \ \"} \ subsubsection \Extra rules for the sequent calculus\ text \As mentioned, destruction rules such as \(\E1, 2)\ and \(\E)\ must be replaced by sequent-style elimination rules. In addition, we need rules to embody the classical equivalence between \P \ Q\ and \\ P \ Q\. The introduction rules \(\I1, 2)\ are replaced by a rule that simulates \(\R)\: @{text [display] "(\ Q \ P) \ P \ Q"} The destruction rule \(\E)\ is replaced by @{text [display] "(P \ Q) \ (\ P \ R) \ (Q \ R) \ R"} Quantifier replication also requires special rules. In classical logic, \\x. P x\ is equivalent to \\ (\x. \ P x)\; the rules \(\R)\ and \(\L)\ are dual: \[ \infer[\(\R)\]{\\ \ \, \x. P x\}{\\ \ \, \x. P x, P t\} \qquad \infer[\(\L)\]{\\x. P x, \ \ \\}{\P t, \x. P x, \ \ \\} \] Thus both kinds of quantifier may be replicated. Theorems requiring multiple uses of a universal formula are easy to invent; consider @{text [display] "(\x. P x \ P (f x)) \ P a \ P (f\<^sup>n a)"} for any \n > 1\. Natural examples of the multiple use of an existential formula are rare; a standard one is \\x. \y. P x \ P y\. Forgoing quantifier replication loses completeness, but gains decidability, since the search space becomes finite. Many useful theorems can be proved without replication, and the search generally delivers its verdict in a reasonable time. To adopt this approach, represent the sequent rules \(\R)\, \(\L)\ and \(\R)\ by \(\I)\, \(\E)\ and \(\I)\, respectively, and put \(\E)\ into elimination form: @{text [display] "\x. P x \ (P t \ Q) \ Q"} Elim-resolution with this rule will delete the universal formula after a single use. To replicate universal quantifiers, replace the rule by @{text [display] "\x. P x \ (P t \ \x. P x \ Q) \ Q"} To replicate existential quantifiers, replace \(\I)\ by @{text [display] "(\ (\x. P x) \ P t) \ \x. P x"} All introduction rules mentioned above are also useful in swapped form. Replication makes the search space infinite; we must apply the rules with care. The classical reasoner distinguishes between safe and unsafe rules, applying the latter only when there is no alternative. Depth-first search may well go down a blind alley; best-first search is better behaved in an infinite search space. However, quantifier replication is too expensive to prove any but the simplest theorems. \ subsection \Rule declarations\ text \The proof tools of the Classical Reasoner depend on collections of rules declared in the context, which are classified as introduction, elimination or destruction and as \<^emph>\safe\ or \<^emph>\unsafe\. In general, safe rules can be attempted blindly, while unsafe rules must be used with care. A safe rule must never reduce a provable goal to an unprovable set of subgoals. The rule \P \ P \ Q\ is unsafe because it reduces \P \ Q\ to \P\, which might turn out as premature choice of an unprovable subgoal. Any rule whose premises contain new unknowns is unsafe. The elimination rule \\x. P x \ (P t \ Q) \ Q\ is unsafe, since it is applied via elim-resolution, which discards the assumption \\x. P x\ and replaces it by the weaker assumption \P t\. The rule \P t \ \x. P x\ is unsafe for similar reasons. The quantifier duplication rule \\x. P x \ (P t \ \x. P x \ Q) \ Q\ is unsafe in a different sense: since it keeps the assumption \\x. P x\, it is prone to looping. In classical first-order logic, all rules are safe except those mentioned above. The safe~/ unsafe distinction is vague, and may be regarded merely as a way of giving some rules priority over others. One could argue that \(\E)\ is unsafe, because repeated application of it could generate exponentially many subgoals. Induction rules are unsafe because inductive proofs are difficult to set up automatically. Any inference that instantiates an unknown in the proof state is unsafe --- thus matching must be used, rather than unification. Even proof by assumption is unsafe if it instantiates unknowns shared with other subgoals. \begin{matharray}{rcl} @{command_def "print_claset"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def intro} & : & \attribute\ \\ @{attribute_def elim} & : & \attribute\ \\ @{attribute_def dest} & : & \attribute\ \\ @{attribute_def rule} & : & \attribute\ \\ @{attribute_def iff} & : & \attribute\ \\ @{attribute_def swapped} & : & \attribute\ \\ \end{matharray} \<^rail>\ (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? ; @@{attribute rule} 'del' ; @@{attribute iff} (((() | 'add') '?'?) | 'del') \ \<^descr> @{command "print_claset"} prints the collection of rules declared to the Classical Reasoner, i.e.\ the \<^ML_type>\claset\ within the context. \<^descr> @{attribute intro}, @{attribute elim}, and @{attribute dest} declare introduction, elimination, and destruction rules, respectively. By default, rules are considered as \<^emph>\unsafe\ (i.e.\ not applied blindly without backtracking), while ``\!\'' classifies as \<^emph>\safe\. Rule declarations marked by ``\?\'' coincide with those of Isabelle/Pure, cf.\ \secref{sec:pure-meth-att} (i.e.\ are only applied in single steps of the @{method rule} method). The optional natural number specifies an explicit weight argument, which is ignored by the automated reasoning tools, but determines the search order of single rule steps. Introduction rules are those that can be applied using ordinary resolution. Their swapped forms are generated internally, which will be applied using elim-resolution. Elimination rules are applied using elim-resolution. Rules are sorted by the number of new subgoals they will yield; rules that generate the fewest subgoals will be tried first. Otherwise, later declarations take precedence over earlier ones. Rules already present in the context with the same classification are ignored. A warning is printed if the rule has already been added with some other classification, but the rule is added anyway as requested. \<^descr> @{attribute rule}~\del\ deletes all occurrences of a rule from the classical context, regardless of its classification as introduction~/ elimination~/ destruction and safe~/ unsafe. \<^descr> @{attribute iff} declares logical equivalences to the Simplifier and the Classical reasoner at the same time. Non-conditional rules result in a safe introduction and elimination pair; conditional ones are considered unsafe. Rules with negative conclusion are automatically inverted (using \\\-elimination internally). The ``\?\'' version of @{attribute iff} declares rules to the Isabelle/Pure context only, and omits the Simplifier declaration. \<^descr> @{attribute swapped} turns an introduction rule into an elimination, by resolving with the classical swap principle \\ P \ (\ R \ P) \ R\ in the second position. This is mainly for illustrative purposes: the Classical Reasoner already swaps rules internally as explained above. \ subsection \Structured methods\ text \ \begin{matharray}{rcl} @{method_def rule} & : & \method\ \\ @{method_def contradiction} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method rule} @{syntax thms}? \ \<^descr> @{method rule} as offered by the Classical Reasoner is a refinement over the Pure one (see \secref{sec:pure-meth-att}). Both versions work the same, but the classical version observes the classical rule context in addition to that of Isabelle/Pure. Common object logics (HOL, ZF, etc.) declare a rich collection of classical rules (even if these would qualify as intuitionistic ones), but only few declarations to the rule context of Isabelle/Pure (\secref{sec:pure-meth-att}). \<^descr> @{method contradiction} solves some goal by contradiction, deriving any result from both \\ A\ and \A\. Chained facts, which are guaranteed to participate, may appear in either order. \ subsection \Fully automated methods\ text \ \begin{matharray}{rcl} @{method_def blast} & : & \method\ \\ @{method_def auto} & : & \method\ \\ @{method_def force} & : & \method\ \\ @{method_def fast} & : & \method\ \\ @{method_def slow} & : & \method\ \\ @{method_def best} & : & \method\ \\ @{method_def fastforce} & : & \method\ \\ @{method_def slowsimp} & : & \method\ \\ @{method_def bestsimp} & : & \method\ \\ @{method_def deepen} & : & \method\ \\ \end{matharray} \<^rail>\ @@{method blast} @{syntax nat}? (@{syntax clamod} * ) ; @@{method auto} (@{syntax nat} @{syntax nat})? (@{syntax clasimpmod} * ) ; @@{method force} (@{syntax clasimpmod} * ) ; (@@{method fast} | @@{method slow} | @@{method best}) (@{syntax clamod} * ) ; (@@{method fastforce} | @@{method slowsimp} | @@{method bestsimp}) (@{syntax clasimpmod} * ) ; @@{method deepen} (@{syntax nat} ?) (@{syntax clamod} * ) ; @{syntax_def clamod}: (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del') ':' @{syntax thms} ; @{syntax_def clasimpmod}: ('simp' (() | 'add' | 'del' | 'only') | 'cong' (() | 'add' | 'del') | 'split' (() | '!' | 'del') | 'iff' (((() | 'add') '?'?) | 'del') | (('intro' | 'elim' | 'dest') ('!' | () | '?') | 'del')) ':' @{syntax thms} \ \<^descr> @{method blast} is a separate classical tableau prover that uses the same classical rule declarations as explained before. Proof search is coded directly in ML using special data structures. A successful proof is then reconstructed using regular Isabelle inferences. It is faster and more powerful than the other classical reasoning tools, but has major limitations too. \<^item> It does not use the classical wrapper tacticals, such as the integration with the Simplifier of @{method fastforce}. \<^item> It does not perform higher-order unification, as needed by the rule @{thm [source=false] rangeI} in HOL. There are often alternatives to such rules, for example @{thm [source=false] range_eqI}. \<^item> Function variables may only be applied to parameters of the subgoal. (This restriction arises because the prover does not use higher-order unification.) If other function variables are present then the prover will fail with the message @{verbatim [display] \Function unknown's argument not a bound variable\} \<^item> Its proof strategy is more general than @{method fast} but can be slower. If @{method blast} fails or seems to be running forever, try @{method fast} and the other proof tools described below. The optional integer argument specifies a bound for the number of unsafe steps used in a proof. By default, @{method blast} starts with a bound of 0 and increases it successively to 20. In contrast, \(blast lim)\ tries to prove the goal using a search bound of \lim\. Sometimes a slow proof using @{method blast} can be made much faster by supplying the successful search bound to this proof method instead. \<^descr> @{method auto} combines classical reasoning with simplification. It is intended for situations where there are a lot of mostly trivial subgoals; it proves all the easy ones, leaving the ones it cannot prove. Occasionally, attempting to prove the hard ones may take a long time. The optional depth arguments in \(auto m n)\ refer to its builtin classical reasoning procedures: \m\ (default 4) is for @{method blast}, which is tried first, and \n\ (default 2) is for a slower but more general alternative that also takes wrappers into account. \<^descr> @{method force} is intended to prove the first subgoal completely, using many fancy proof tools and performing a rather exhaustive search. As a result, proof attempts may take rather long or diverge easily. \<^descr> @{method fast}, @{method best}, @{method slow} attempt to prove the first subgoal using sequent-style reasoning as explained before. Unlike @{method blast}, they construct proofs directly in Isabelle. There is a difference in search strategy and back-tracking: @{method fast} uses depth-first search and @{method best} uses best-first search (guided by a heuristic function: normally the total size of the proof state). Method @{method slow} is like @{method fast}, but conducts a broader search: it may, when backtracking from a failed proof attempt, undo even the step of proving a subgoal by assumption. \<^descr> @{method fastforce}, @{method slowsimp}, @{method bestsimp} are like @{method fast}, @{method slow}, @{method best}, respectively, but use the Simplifier as additional wrapper. The name @{method fastforce}, reflects the behaviour of this popular method better without requiring an understanding of its implementation. \<^descr> @{method deepen} works by exhaustive search up to a certain depth. The start depth is 4 (unless specified explicitly), and the depth is increased iteratively up to 10. Unsafe rules are modified to preserve the formula they act on, so that it be used repeatedly. This method can prove more goals than @{method fast}, but is much slower, for example if the assumptions have many universal quantifiers. Any of the above methods support additional modifiers of the context of classical (and simplifier) rules, but the ones related to the Simplifier are explicitly prefixed by \simp\ here. The semantics of these ad-hoc rule declarations is analogous to the attributes given before. Facts provided by forward chaining are inserted into the goal before commencing proof search. \ subsection \Partially automated methods\label{sec:classical:partial}\ text \These proof methods may help in situations when the fully-automated tools fail. The result is a simpler subgoal that can be tackled by other means, such as by manual instantiation of quantifiers. \begin{matharray}{rcl} @{method_def safe} & : & \method\ \\ @{method_def clarify} & : & \method\ \\ @{method_def clarsimp} & : & \method\ \\ \end{matharray} \<^rail>\ (@@{method safe} | @@{method clarify}) (@{syntax clamod} * ) ; @@{method clarsimp} (@{syntax clasimpmod} * ) \ \<^descr> @{method safe} repeatedly performs safe steps on all subgoals. It is deterministic, with at most one outcome. \<^descr> @{method clarify} performs a series of safe steps without splitting subgoals; see also @{method clarify_step}. \<^descr> @{method clarsimp} acts like @{method clarify}, but also does simplification. Note that if the Simplifier context includes a splitter for the premises, the subgoal may still be split. \ subsection \Single-step tactics\ text \ \begin{matharray}{rcl} @{method_def safe_step} & : & \method\ \\ @{method_def inst_step} & : & \method\ \\ @{method_def step} & : & \method\ \\ @{method_def slow_step} & : & \method\ \\ @{method_def clarify_step} & : & \method\ \\ \end{matharray} These are the primitive tactics behind the automated proof methods of the Classical Reasoner. By calling them yourself, you can execute these procedures one step at a time. \<^descr> @{method safe_step} performs a safe step on the first subgoal. The safe wrapper tacticals are applied to a tactic that may include proof by assumption or Modus Ponens (taking care not to instantiate unknowns), or substitution. \<^descr> @{method inst_step} is like @{method safe_step}, but allows unknowns to be instantiated. \<^descr> @{method step} is the basic step of the proof procedure, it operates on the first subgoal. The unsafe wrapper tacticals are applied to a tactic that tries @{method safe}, @{method inst_step}, or applies an unsafe rule from the context. \<^descr> @{method slow_step} resembles @{method step}, but allows backtracking between using safe rules with instantiation (@{method inst_step}) and using unsafe rules. The resulting search space is larger. \<^descr> @{method clarify_step} performs a safe step on the first subgoal; no splitting step is applied. For example, the subgoal \A \ B\ is left as a conjunction. Proof by assumption, Modus Ponens, etc., may be performed provided they do not instantiate unknowns. Assumptions of the form \x = t\ may be eliminated. The safe wrapper tactical is applied. \ subsection \Modifying the search step\ text \ \begin{mldecls} @{define_ML_type wrapper = "(int -> tactic) -> (int -> tactic)"} \\[0.5ex] @{define_ML_infix addSWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addSbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addSafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delSWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML_infix addWrapper: "Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context"} \\ @{define_ML_infix addbefore: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix addafter: "Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context"} \\ @{define_ML_infix delWrapper: "Proof.context * string -> Proof.context"} \\[0.5ex] @{define_ML addSss: "Proof.context -> Proof.context"} \\ @{define_ML addss: "Proof.context -> Proof.context"} \\ \end{mldecls} The proof strategy of the Classical Reasoner is simple. Perform as many safe inferences as possible; or else, apply certain safe rules, allowing instantiation of unknowns; or else, apply an unsafe rule. The tactics also eliminate assumptions of the form \x = t\ by substitution if they have been set up to do so. They may perform a form of Modus Ponens: if there are assumptions \P \ Q\ and \P\, then replace \P \ Q\ by \Q\. The classical reasoning tools --- except @{method blast} --- allow to modify this basic proof strategy by applying two lists of arbitrary \<^emph>\wrapper tacticals\ to it. The first wrapper list, which is considered to contain safe wrappers only, affects @{method safe_step} and all the tactics that call it. The second one, which may contain unsafe wrappers, affects the unsafe parts of @{method step}, @{method slow_step}, and the tactics that call them. A wrapper transforms each step of the search, for example by attempting other tactics before or after the original step tactic. All members of a wrapper list are applied in turn to the respective step tactic. Initially the two wrapper lists are empty, which means no modification of the step tactics. Safe and unsafe wrappers are added to the context with the functions given below, supplying them with wrapper names. These names may be used to selectively delete wrappers. \<^descr> \ctxt addSWrapper (name, wrapper)\ adds a new wrapper, which should yield a safe tactic, to modify the existing safe step tactic. \<^descr> \ctxt addSbefore (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried \<^emph>\before\ each safe step of the search. \<^descr> \ctxt addSafter (name, tac)\ adds the given tactic as a safe wrapper, such that it is tried when a safe step of the search would fail. \<^descr> \ctxt delSWrapper name\ deletes the safe wrapper with the given name. \<^descr> \ctxt addWrapper (name, wrapper)\ adds a new wrapper to modify the existing (unsafe) step tactic. \<^descr> \ctxt addbefore (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\before\ the result of each unsafe step. \<^descr> \ctxt addafter (name, tac)\ adds the given tactic as an unsafe wrapper, such that it its result is concatenated \<^emph>\after\ the result of each unsafe step. \<^descr> \ctxt delWrapper name\ deletes the unsafe wrapper with the given name. \<^descr> \addSss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, in a rather safe way, after each safe step of the search. \<^descr> \addss\ adds the simpset of the context to its classical set. The assumptions and goal will be simplified, before the each unsafe step of the search. \ section \Object-logic setup \label{sec:object-logic}\ text \ \begin{matharray}{rcl} @{command_def "judgment"} & : & \theory \ theory\ \\ @{method_def atomize} & : & \method\ \\ @{attribute_def atomize} & : & \attribute\ \\ @{attribute_def rule_format} & : & \attribute\ \\ @{attribute_def rulify} & : & \attribute\ \\ \end{matharray} The very starting point for any Isabelle object-logic is a ``truth judgment'' that links object-level statements to the meta-logic (with its minimal language of \prop\ that covers universal quantification \\\ and implication \\\). Common object-logics are sufficiently expressive to internalize rule statements over \\\ and \\\ within their own language. This is useful in certain situations where a rule needs to be viewed as an atomic statement from the meta-level perspective, e.g.\ \\x. x \ A \ P x\ versus \\x \ A. P x\. From the following language elements, only the @{method atomize} method and @{attribute rule_format} attribute are occasionally required by end-users, the rest is for those who need to setup their own object-logic. In the latter case existing formulations of Isabelle/FOL or Isabelle/HOL may be taken as realistic examples. Generic tools may refer to the information provided by object-logic declarations internally. \<^rail>\ @@{command judgment} @{syntax name} '::' @{syntax type} @{syntax mixfix}? ; @@{attribute atomize} ('(' 'full' ')')? ; @@{attribute rule_format} ('(' 'noasm' ')')? \ \<^descr> @{command "judgment"}~\c :: \ (mx)\ declares constant \c\ as the truth judgment of the current object-logic. Its type \\\ should specify a coercion of the category of object-level propositions to \prop\ of the Pure meta-logic; the mixfix annotation \(mx)\ would typically just link the object language (internally of syntactic category \logic\) with that of \prop\. Only one @{command "judgment"} declaration may be given in any theory development. \<^descr> @{method atomize} (as a method) rewrites any non-atomic premises of a sub-goal, using the meta-level equations declared via @{attribute atomize} (as an attribute) beforehand. As a result, heavily nested goals become amenable to fundamental operations such as resolution (cf.\ the @{method (Pure) rule} method). Giving the ``\(full)\'' option here means to turn the whole subgoal into an object-statement (if possible), including the outermost parameters and assumptions as well. A typical collection of @{attribute atomize} rules for a particular object-logic would provide an internalization for each of the connectives of \\\, \\\, and \\\. Meta-level conjunction should be covered as well (this is particularly important for locales, see \secref{sec:locale}). \<^descr> @{attribute rule_format} rewrites a theorem by the equalities declared as @{attribute rulify} rules in the current object-logic. By default, the result is fully normalized, including assumptions and conclusions at any depth. The \(no_asm)\ option restricts the transformation to the conclusion of a rule. In common object-logics (HOL, FOL, ZF), the effect of @{attribute rule_format} is to replace (bounded) universal quantification (\\\) and implication (\\\) by the corresponding rule statements over \\\ and \\\. \ section \Tracing higher-order unification\ text \ \begin{tabular}{rcll} @{attribute_def unify_trace_simp} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_types} & : & \attribute\ & default \false\ \\ @{attribute_def unify_trace_bound} & : & \attribute\ & default \50\ \\ @{attribute_def unify_search_bound} & : & \attribute\ & default \60\ \\ \end{tabular} \<^medskip> Higher-order unification works well in most practical situations, but sometimes needs extra care to identify problems. These tracing options may help. \<^descr> @{attribute unify_trace_simp} controls tracing of the simplification phase of higher-order unification. \<^descr> @{attribute unify_trace_types} controls warnings of incompleteness, when unification is not considering all possible instantiations of schematic type variables. \<^descr> @{attribute unify_trace_bound} determines the depth where unification starts to print tracing information once it reaches depth; 0 for full tracing. At the default value, tracing information is almost never printed in practice. \<^descr> @{attribute unify_search_bound} prevents unification from searching past the given depth. Because of this bound, higher-order unification cannot return an infinite sequence, though it can return an exponentially long one. The search rarely approaches the default value in practice. If the search is cut off, unification prints a warning ``Unification bound exceeded''. \begin{warn} Options for unification cannot be modified in a local context. Only the global theory content is taken into account. \end{warn} \ end diff --git a/src/Doc/Isar_Ref/Proof.thy b/src/Doc/Isar_Ref/Proof.thy --- a/src/Doc/Isar_Ref/Proof.thy +++ b/src/Doc/Isar_Ref/Proof.thy @@ -1,1440 +1,1440 @@ (*:maxLineLen=78:*) theory Proof imports Main Base begin chapter \Proofs \label{ch:proofs}\ text \ Proof commands perform transitions of Isar/VM machine configurations, which are block-structured, consisting of a stack of nodes with three main components: logical proof context, current facts, and open goals. Isar/VM transitions are typed according to the following three different modes of operation: \<^descr> \proof(prove)\ means that a new goal has just been stated that is now to be \<^emph>\proven\; the next command may refine it by some proof method, and enter a sub-proof to establish the actual result. \<^descr> \proof(state)\ is like a nested theory mode: the context may be augmented by \<^emph>\stating\ additional assumptions, intermediate results etc. \<^descr> \proof(chain)\ is intermediate between \proof(state)\ and \proof(prove)\: existing facts (i.e.\ the contents of the special @{fact_ref this} register) have been just picked up in order to be used when refining the goal claimed next. The proof mode indicator may be understood as an instruction to the writer, telling what kind of operation may be performed next. The corresponding typings of proof commands restricts the shape of well-formed proof texts to particular command sequences. So dynamic arrangements of commands eventually turn out as static texts of a certain structure. \Appref{ap:refcard} gives a simplified grammar of the (extensible) language emerging that way from the different types of proof commands. The main ideas of the overall Isar framework are explained in \chref{ch:isar-framework}. \ section \Proof structure\ subsection \Formal notepad\ text \ \begin{matharray}{rcl} @{command_def "notepad"} & : & \local_theory \ proof(state)\ \\ \end{matharray} \<^rail>\ @@{command notepad} @'begin' ; @@{command end} \ \<^descr> @{command "notepad"}~@{keyword "begin"} opens a proof state without any goal statement. This allows to experiment with Isar, without producing any persistent result. The notepad is closed by @{command "end"}. \ subsection \Blocks\ text \ \begin{matharray}{rcl} @{command_def "next"} & : & \proof(state) \ proof(state)\ \\ @{command_def "{"} & : & \proof(state) \ proof(state)\ \\ @{command_def "}"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} While Isar is inherently block-structured, opening and closing blocks is mostly handled rather casually, with little explicit user-intervention. Any local goal statement automatically opens \<^emph>\two\ internal blocks, which are closed again when concluding the sub-proof (by @{command "qed"} etc.). Sections of different context within a sub-proof may be switched via @{command "next"}, which is just a single block-close followed by block-open again. The effect of @{command "next"} is to reset the local proof context; there is no goal focus involved here! For slightly more advanced applications, there are explicit block parentheses as well. These typically achieve a stronger forward style of reasoning. \<^descr> @{command "next"} switches to a fresh block within a sub-proof, resetting the local context to the initial one. \<^descr> @{command "{"} and @{command "}"} explicitly open and close blocks. Any current facts pass through ``@{command "{"}'' unchanged, while ``@{command "}"}'' causes any result to be \<^emph>\exported\ into the enclosing context. Thus fixed variables are generalized, assumptions discharged, and local definitions unfolded (cf.\ \secref{sec:proof-context}). There is no difference of @{command "assume"} and @{command "presume"} in this mode of forward reasoning --- in contrast to plain backward reasoning with the result exported at @{command "show"} time. \ subsection \Omitting proofs\ text \ \begin{matharray}{rcl} @{command_def "oops"} & : & \proof \ local_theory | theory\ \\ \end{matharray} The @{command "oops"} command discontinues the current proof attempt, while considering the partial proof text as properly processed. This is conceptually quite different from ``faking'' actual proofs via @{command_ref "sorry"} (see \secref{sec:proof-steps}): @{command "oops"} does not observe the proof structure at all, but goes back right to the theory level. Furthermore, @{command "oops"} does not produce any result theorem --- there is no intended claim to be able to complete the proof in any way. A typical application of @{command "oops"} is to explain Isar proofs \<^emph>\within\ the system itself, in conjunction with the document preparation tools of Isabelle described in \chref{ch:document-prep}. Thus partial or even wrong proof attempts can be discussed in a logically sound manner. Note that the Isabelle {\LaTeX} macros can be easily adapted to print something like ``\\\'' instead of the keyword ``@{command "oops"}''. \ section \Statements\ subsection \Context elements \label{sec:proof-context}\ text \ \begin{matharray}{rcl} @{command_def "fix"} & : & \proof(state) \ proof(state)\ \\ @{command_def "assume"} & : & \proof(state) \ proof(state)\ \\ @{command_def "presume"} & : & \proof(state) \ proof(state)\ \\ @{command_def "define"} & : & \proof(state) \ proof(state)\ \\ \end{matharray} The logical proof context consists of fixed variables and assumptions. The former closely correspond to Skolem constants, or meta-level universal quantification as provided by the Isabelle/Pure logical framework. Introducing some \<^emph>\arbitrary, but fixed\ variable via ``@{command "fix"}~\x\'' results in a local value that may be used in the subsequent proof as any other variable or constant. Furthermore, any result \\ \[x]\ exported from the context will be universally closed wrt.\ \x\ at the outermost level: \\ \x. \[x]\ (this is expressed in normal form using Isabelle's meta-variables). Similarly, introducing some assumption \\\ has two effects. On the one hand, a local theorem is created that may be used as a fact in subsequent proof steps. On the other hand, any result \\ \ \\ exported from the context becomes conditional wrt.\ the assumption: \\ \ \ \\. Thus, solving an enclosing goal using such a result would basically introduce a new subgoal stemming from the assumption. How this situation is handled depends on the version of assumption command used: while @{command "assume"} insists on solving the subgoal by unification with some premise of the goal, @{command "presume"} leaves the subgoal unchanged in order to be proved later by the user. Local definitions, introduced by ``\<^theory_text>\define x where x = t\'', are achieved by combining ``@{command "fix"}~\x\'' with another version of assumption that causes any hypothetical equation \x \ t\ to be eliminated by the reflexivity rule. Thus, exporting some result \x \ t \ \[x]\ yields \\ \[t]\. \<^rail>\ @@{command fix} @{syntax vars} ; (@@{command assume} | @@{command presume}) concl prems @{syntax for_fixes} ; concl: (@{syntax props} + @'and') ; prems: (@'if' (@{syntax props'} + @'and'))? ; @@{command define} @{syntax vars} @'where' (@{syntax props} + @'and') @{syntax for_fixes} \ \<^descr> @{command "fix"}~\x\ introduces a local variable \x\ that is \<^emph>\arbitrary, but fixed\. \<^descr> @{command "assume"}~\a: \\ and @{command "presume"}~\a: \\ introduce a local fact \\ \ \\ by assumption. Subsequent results applied to an enclosing goal (e.g.\ by @{command_ref "show"}) are handled as follows: @{command "assume"} expects to be able to unify with existing premises in the goal, while @{command "presume"} leaves \\\ as new subgoals. Several lists of assumptions may be given (separated by @{keyword_ref "and"}; the resulting list of current facts consists of all of these concatenated. A structured assumption like \<^theory_text>\assume "B x" if "A x" for x\ is equivalent to \<^theory_text>\assume "\x. A x \ B x"\, but vacuous quantification is avoided: a for-context only effects propositions according to actual use of variables. \<^descr> \<^theory_text>\define x where "x = t"\ introduces a local (non-polymorphic) definition. In results that are exported from the context, \x\ is replaced by \t\. Internally, equational assumptions are added to the context in Pure form, using \x \ t\ instead of \x = t\ or \x \ t\ from the object-logic. When exporting results from the context, \x\ is generalized and the assumption discharged by reflexivity, causing the replacement by \t\. The default name for the definitional fact is \x_def\. Several simultaneous definitions may be given as well, with a collective default name. \<^medskip> It is also possible to abstract over local parameters as follows: \<^theory_text>\define f :: "'a \ 'b" where "f x = t" for x :: 'a\. \ subsection \Term abbreviations \label{sec:term-abbrev}\ text \ \begin{matharray}{rcl} @{command_def "let"} & : & \proof(state) \ proof(state)\ \\ @{keyword_def "is"} & : & syntax \\ \end{matharray} Abbreviations may be either bound by explicit @{command "let"}~\p \ t\ statements, or by annotating assumptions or goal statements with a list of patterns ``\<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\''. In both cases, higher-order matching is invoked to bind extra-logical term variables, which may be either named schematic variables of the form \?x\, or nameless dummies ``@{variable _}'' (underscore). Note that in the @{command "let"} form the patterns occur on the left-hand side, while the @{keyword "is"} patterns are in postfix position. Polymorphism of term bindings is handled in Hindley-Milner style, similar to ML. Type variables referring to local assumptions or open goal statements are \<^emph>\fixed\, while those of finished results or bound by @{command "let"} may occur in \<^emph>\arbitrary\ instances later. Even though actual polymorphism should be rarely used in practice, this mechanism is essential to achieve proper incremental type-inference, as the user proceeds to build up the Isar proof text from left to right. \<^medskip> Term abbreviations are quite different from local definitions as introduced via @{command "define"} (see \secref{sec:proof-context}). The latter are visible within the logic as actual equations, while abbreviations disappear during the input process just after type checking. Also note that @{command "define"} does not support polymorphism. \<^rail>\ @@{command let} ((@{syntax term} + @'and') '=' @{syntax term} + @'and') \ The syntax of @{keyword "is"} patterns follows @{syntax term_pat} or @{syntax prop_pat} (see \secref{sec:term-decls}). \<^descr> \<^theory_text>\let p\<^sub>1 = t\<^sub>1 and \ p\<^sub>n = t\<^sub>n\ binds any text variables in patterns \p\<^sub>1, \, p\<^sub>n\ by simultaneous higher-order matching against terms \t\<^sub>1, \, t\<^sub>n\. \<^descr> \<^theory_text>\(is p\<^sub>1 \ p\<^sub>n)\ resembles @{command "let"}, but matches \p\<^sub>1, \, p\<^sub>n\ against the preceding statement. Also note that @{keyword "is"} is not a separate command, but part of others (such as @{command "assume"}, @{command "have"} etc.). Some \<^emph>\implicit\ term abbreviations\index{term abbreviations} for goals and facts are available as well. For any open goal, @{variable_ref thesis} refers to its object-level statement, abstracted over any meta-level parameters (if present). Likewise, @{variable_ref this} is bound for fact statements resulting from assumptions or finished goals. In case @{variable this} refers to an object-logic statement that is an application \f t\, then \t\ is bound to the special text variable ``@{variable "\"}'' (three dots). The canonical application of this convenience are calculational proofs (see \secref{sec:calculation}). \ subsection \Facts and forward chaining \label{sec:proof-facts}\ text \ \begin{matharray}{rcl} @{command_def "note"} & : & \proof(state) \ proof(state)\ \\ @{command_def "then"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "from"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "with"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "using"} & : & \proof(prove) \ proof(prove)\ \\ @{command_def "unfolding"} & : & \proof(prove) \ proof(prove)\ \\ @{method_def "use"} & : & \method\ \\ @{fact_def "method_facts"} & : & \fact\ \\ \end{matharray} New facts are established either by assumption or proof of local statements. Any fact will usually be involved in further proofs, either as explicit arguments of proof methods, or when forward chaining towards the next goal via @{command "then"} (and variants); @{command "from"} and @{command "with"} are composite forms involving @{command "note"}. The @{command "using"} elements augments the collection of used facts \<^emph>\after\ a goal has been stated. Note that the special theorem name @{fact_ref this} refers to the most recently established facts, but only \<^emph>\before\ issuing a follow-up claim. \<^rail>\ @@{command note} (@{syntax thmdef}? @{syntax thms} + @'and') ; (@@{command from} | @@{command with} | @@{command using} | @@{command unfolding}) (@{syntax thms} + @'and') ; @{method use} @{syntax thms} @'in' @{syntax method} \ \<^descr> @{command "note"}~\a = b\<^sub>1 \ b\<^sub>n\ recalls existing facts \b\<^sub>1, \, b\<^sub>n\, binding the result as \a\. Note that attributes may be involved as well, both on the left and right hand sides. \<^descr> @{command "then"} indicates forward chaining by the current facts in order to establish the goal to be claimed next. The initial proof method invoked to refine that will be offered the facts to do ``anything appropriate'' (see also \secref{sec:proof-steps}). For example, method @{method (Pure) rule} (see \secref{sec:pure-meth-att}) would typically do an elimination rather than an introduction. Automatic methods usually insert the facts into the goal state before operation. This provides a simple scheme to control relevance of facts in automated proof search. \<^descr> @{command "from"}~\b\ abbreviates ``@{command "note"}~\b\~@{command "then"}''; thus @{command "then"} is equivalent to ``@{command "from"}~\this\''. \<^descr> @{command "with"}~\b\<^sub>1 \ b\<^sub>n\ abbreviates ``@{command "from"}~\b\<^sub>1 \ b\<^sub>n \ this\''; thus the forward chaining is from earlier facts together with the current ones. \<^descr> @{command "using"}~\b\<^sub>1 \ b\<^sub>n\ augments the facts to be used by a subsequent refinement step (such as @{command_ref "apply"} or @{command_ref "proof"}). \<^descr> @{command "unfolding"}~\b\<^sub>1 \ b\<^sub>n\ is structurally similar to @{command "using"}, but unfolds definitional equations \b\<^sub>1 \ b\<^sub>n\ throughout the goal state and facts. See also the proof method @{method_ref unfold}. \<^descr> \<^theory_text>\(use b\<^sub>1 \ b\<^sub>n in method)\ uses the facts in the given method expression. The facts provided by the proof state (via @{command "using"} etc.) are ignored, but it is possible to refer to @{fact method_facts} explicitly. \<^descr> @{fact method_facts} is a dynamic fact that refers to the currently used facts of the goal state. Forward chaining with an empty list of theorems is the same as not chaining at all. Thus ``@{command "from"}~\nothing\'' has no effect apart from entering \prove(chain)\ mode, since @{fact_ref nothing} is bound to the empty list of theorems. Basic proof methods (such as @{method_ref (Pure) rule}) expect multiple facts to be given in their proper order, corresponding to a prefix of the premises of the rule involved. Note that positions may be easily skipped using something like @{command "from"}~\_ \ a \ b\, for example. This involves the trivial rule \PROP \ \ PROP \\, which is bound in Isabelle/Pure as ``@{fact_ref "_"}'' (underscore). Automated methods (such as @{method simp} or @{method auto}) just insert any given facts before their usual operation. Depending on the kind of procedure involved, the order of facts is less significant here. \ subsection \Goals \label{sec:goals}\ text \ \begin{matharray}{rcl} @{command_def "lemma"} & : & \local_theory \ proof(prove)\ \\ @{command_def "theorem"} & : & \local_theory \ proof(prove)\ \\ @{command_def "corollary"} & : & \local_theory \ proof(prove)\ \\ @{command_def "proposition"} & : & \local_theory \ proof(prove)\ \\ @{command_def "schematic_goal"} & : & \local_theory \ proof(prove)\ \\ @{command_def "have"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "show"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "hence"} & : & \proof(state) \ proof(prove)\ \\ @{command_def "thus"} & : & \proof(state) \ proof(prove)\ \\ @{command_def "print_statement"}\\<^sup>*\ & : & \context \\ \\ \end{matharray} From a theory context, proof mode is entered by an initial goal command such as @{command "lemma"}. Within a proof context, new claims may be introduced locally; there are variants to interact with the overall proof structure specifically, such as @{command have} or @{command show}. Goals may consist of multiple statements, resulting in a list of facts eventually. A pending multi-goal is internally represented as a meta-level conjunction (\&&&\), which is usually split into the corresponding number of sub-goals prior to an initial method application, via @{command_ref "proof"} (\secref{sec:proof-steps}) or @{command_ref "apply"} (\secref{sec:tactic-commands}). The @{method_ref induct} method covered in \secref{sec:cases-induct} acts on multiple claims simultaneously. Claims at the theory level may be either in short or long form. A short goal merely consists of several simultaneous propositions (often just one). A long goal includes an explicit context specification for the subsequent conclusion, involving local parameters and assumptions. Here the role of each part of the statement is explicitly marked by separate keywords (see also \secref{sec:locale}); the local assumptions being introduced here are available as @{fact_ref assms} in the proof. Moreover, there are two kinds of conclusions: @{element_def "shows"} states several simultaneous propositions (essentially a big conjunction), while @{element_def "obtains"} claims several simultaneous simultaneous contexts of (essentially a big disjunction of eliminated parameters and assumptions, cf.\ \secref{sec:obtain}). \<^rail>\ (@@{command lemma} | @@{command theorem} | @@{command corollary} | @@{command proposition} | @@{command schematic_goal}) (long_statement | short_statement) ; (@@{command have} | @@{command show} | @@{command hence} | @@{command thus}) stmt cond_stmt @{syntax for_fixes} ; @@{command print_statement} @{syntax modes}? @{syntax thms} ; stmt: (@{syntax props} + @'and') ; cond_stmt: ((@'if' | @'when') stmt)? ; short_statement: stmt (@'if' stmt)? @{syntax for_fixes} ; long_statement: @{syntax thmdecl}? context conclusion ; context: (@{syntax_ref "includes"}?) (@{syntax context_elem} *) ; conclusion: @'shows' stmt | @'obtains' @{syntax obtain_clauses} ; @{syntax_def obtain_clauses}: (@{syntax par_name}? obtain_case + '|') ; @{syntax_def obtain_case}: @{syntax vars} @'where' (@{syntax thmdecl}? (@{syntax prop}+) + @'and') \ \<^descr> @{command "lemma"}~\a: \\ enters proof mode with \\\ as main goal, eventually resulting in some fact \\ \\ to be put back into the target context. A @{syntax long_statement} may build up an initial proof context for the subsequent claim, potentially including local definitions and syntax; see also @{syntax "includes"} in \secref{sec:bundle} and @{syntax context_elem} in \secref{sec:locale}. A @{syntax short_statement} consists of propositions as conclusion, with an option context of premises and parameters, via \<^verbatim>\if\/\<^verbatim>\for\ in postfix notation, corresponding to \<^verbatim>\assumes\/\<^verbatim>\fixes\ in the long prefix notation. Local premises (if present) are called ``\assms\'' for @{syntax long_statement}, and ``\that\'' for @{syntax short_statement}. \<^descr> @{command "theorem"}, @{command "corollary"}, and @{command "proposition"} are the same as @{command "lemma"}. The different command names merely serve as a formal comment in the theory source. \<^descr> @{command "schematic_goal"} is similar to @{command "theorem"}, but allows the statement to contain unbound schematic variables. Under normal circumstances, an Isar proof text needs to specify claims explicitly. Schematic goals are more like goals in Prolog, where certain results are synthesized in the course of reasoning. With schematic statements, the inherent compositionality of Isar proofs is lost, which also impacts performance, because proof checking is forced into sequential mode. \<^descr> @{command "have"}~\a: \\ claims a local goal, eventually resulting in a fact within the current logical context. This operation is completely independent of any pending sub-goals of an enclosing goal statements, so @{command "have"} may be freely used for experimental exploration of potential results within a proof body. \<^descr> @{command "show"}~\a: \\ is like @{command "have"}~\a: \\ plus a second stage to refine some pending sub-goal for each one of the finished result, after having been exported into the corresponding context (at the head of the sub-proof of this @{command "show"} command). To accommodate interactive debugging, resulting rules are printed before being applied internally. Even more, interactive execution of @{command "show"} predicts potential failure and displays the resulting error as a warning beforehand. Watch out for the following message: @{verbatim [display] \Local statement fails to refine any pending goal\} \<^descr> @{command "hence"} expands to ``@{command "then"}~@{command "have"}'' and @{command "thus"} expands to ``@{command "then"}~@{command "show"}''. These conflations are left-over from early history of Isar. The expanded syntax is more orthogonal and improves readability and maintainability of proofs. \<^descr> @{command "print_statement"}~\a\ prints facts from the current theory or proof context in long statement form, according to the syntax for @{command "lemma"} given above. Any goal statement causes some term abbreviations (such as @{variable_ref "?thesis"}) to be bound automatically, see also \secref{sec:term-abbrev}. Structured goal statements involving @{keyword_ref "if"} or @{keyword_ref "when"} define the special fact @{fact_ref that} to refer to these assumptions in the proof body. The user may provide separate names according to the syntax of the statement. \ section \Calculational reasoning \label{sec:calculation}\ text \ \begin{matharray}{rcl} @{command_def "also"} & : & \proof(state) \ proof(state)\ \\ @{command_def "finally"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "moreover"} & : & \proof(state) \ proof(state)\ \\ @{command_def "ultimately"} & : & \proof(state) \ proof(chain)\ \\ @{command_def "print_trans_rules"}\\<^sup>*\ & : & \context \\ \\ @{attribute trans} & : & \attribute\ \\ @{attribute sym} & : & \attribute\ \\ @{attribute symmetric} & : & \attribute\ \\ \end{matharray} Calculational proof is forward reasoning with implicit application of transitivity rules (such those of \=\, \\\, \<\). Isabelle/Isar maintains an auxiliary fact register @{fact_ref calculation} for accumulating results obtained by transitivity composed with the current result. Command @{command "also"} updates @{fact calculation} involving @{fact this}, while @{command "finally"} exhibits the final @{fact calculation} by forward chaining towards the next goal statement. Both commands require valid current facts, i.e.\ may occur only after commands that produce theorems such as @{command "assume"}, @{command "note"}, or some finished proof of @{command "have"}, @{command "show"} etc. The @{command "moreover"} and @{command "ultimately"} commands are similar to @{command "also"} and @{command "finally"}, but only collect further results in @{fact calculation} without applying any rules yet. Also note that the implicit term abbreviation ``\\\'' has its canonical application with calculational proofs. It refers to the argument of the preceding statement. (The argument of a curried infix expression happens to be its right-hand side.) Isabelle/Isar calculations are implicitly subject to block structure in the sense that new threads of calculational reasoning are commenced for any new block (as opened by a local goal, for example). This means that, apart from being able to nest calculations, there is no separate \<^emph>\begin-calculation\ command required. \<^medskip> The Isar calculation proof commands may be defined as follows:\<^footnote>\We suppress internal bookkeeping such as proper handling of block-structure.\ \begin{matharray}{rcl} @{command "also"}\\<^sub>0\ & \equiv & @{command "note"}~\calculation = this\ \\ @{command "also"}\\<^sub>n\<^sub>+\<^sub>1\ & \equiv & @{command "note"}~\calculation = trans [OF calculation this]\ \\[0.5ex] @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~\calculation\ \\[0.5ex] @{command "moreover"} & \equiv & @{command "note"}~\calculation = calculation this\ \\ @{command "ultimately"} & \equiv & @{command "moreover"}~@{command "from"}~\calculation\ \\ \end{matharray} \<^rail>\ (@@{command also} | @@{command finally}) ('(' @{syntax thms} ')')? ; @@{attribute trans} (() | 'add' | 'del') \ \<^descr> @{command "also"}~\(a\<^sub>1 \ a\<^sub>n)\ maintains the auxiliary @{fact calculation} register as follows. The first occurrence of @{command "also"} in some calculational thread initializes @{fact calculation} by @{fact this}. Any subsequent @{command "also"} on the same level of block-structure updates @{fact calculation} by some transitivity rule applied to @{fact calculation} and @{fact this} (in that order). Transitivity rules are picked from the current context, unless alternative rules are given as explicit arguments. \<^descr> @{command "finally"}~\(a\<^sub>1 \ a\<^sub>n)\ maintains @{fact calculation} in the same way as @{command "also"} and then concludes the current calculational thread. The final result is exhibited as fact for forward chaining towards the next goal. Basically, @{command "finally"} abbreviates @{command "also"}~@{command "from"}~@{fact calculation}. Typical idioms for concluding calculational proofs are ``@{command "finally"}~@{command "show"}~\?thesis\~@{command "."}'' and ``@{command "finally"}~@{command "have"}~\\\~@{command "."}''. \<^descr> @{command "moreover"} and @{command "ultimately"} are analogous to @{command "also"} and @{command "finally"}, but collect results only, without applying rules. \<^descr> @{command "print_trans_rules"} prints the list of transitivity rules (for calculational commands @{command "also"} and @{command "finally"}) and symmetry rules (for the @{attribute symmetric} operation and single step elimination patters) of the current context. \<^descr> @{attribute trans} declares theorems as transitivity rules. \<^descr> @{attribute sym} declares symmetry rules, as well as @{attribute "Pure.elim"}\?\ rules. \<^descr> @{attribute symmetric} resolves a theorem with some rule declared as @{attribute sym} in the current context. For example, ``@{command "assume"}~\[symmetric]: x = y\'' produces a swapped fact derived from that assumption. In structured proof texts it is often more appropriate to use an explicit single-step elimination proof, such as ``@{command "assume"}~\x = y\~@{command "then"}~@{command "have"}~\y = x\~@{command ".."}''. \ section \Refinement steps\ subsection \Proof method expressions \label{sec:proof-meth}\ text \ Proof methods are either basic ones, or expressions composed of methods via ``\<^verbatim>\,\'' (sequential composition), ``\<^verbatim>\;\'' (structural composition), ``\<^verbatim>\|\'' (alternative choices), ``\<^verbatim>\?\'' (try), ``\<^verbatim>\+\'' (repeat at least once), ``\<^verbatim>\[\\n\\<^verbatim>\]\'' (restriction to first \n\ subgoals). In practice, proof methods are usually just a comma separated list of @{syntax name}~@{syntax args} specifications. Note that parentheses may be dropped for single method specifications (with no arguments). The syntactic precedence of method combinators is \<^verbatim>\|\ \<^verbatim>\;\ \<^verbatim>\,\ \<^verbatim>\[]\ \<^verbatim>\+\ \<^verbatim>\?\ (from low to high). \<^rail>\ @{syntax_def method}: (@{syntax name} | '(' methods ')') (() | '?' | '+' | '[' @{syntax nat}? ']') ; methods: (@{syntax name} @{syntax args} | @{syntax method}) + (',' | ';' | '|') \ Regular Isar proof methods do \<^emph>\not\ admit direct goal addressing, but refer to the first subgoal or to all subgoals uniformly. Nonetheless, the subsequent mechanisms allow to imitate the effect of subgoal addressing that is known from ML tactics. \<^medskip> Goal \<^emph>\restriction\ means the proof state is wrapped-up in a way that certain subgoals are exposed, and other subgoals are ``parked'' elsewhere. Thus a proof method has no other chance than to operate on the subgoals that are presently exposed. Structural composition ``\m\<^sub>1\\<^verbatim>\;\~\m\<^sub>2\'' means that method \m\<^sub>1\ is applied with restriction to the first subgoal, then \m\<^sub>2\ is applied consecutively with restriction to each subgoal that has newly emerged due to - \m\<^sub>1\. This is analogous to the tactic combinator \<^ML_op>\THEN_ALL_NEW\ in + \m\<^sub>1\. This is analogous to the tactic combinator \<^ML_infix>\THEN_ALL_NEW\ in Isabelle/ML, see also @{cite "isabelle-implementation"}. For example, \(rule r; blast)\ applies rule \r\ and then solves all new subgoals by \blast\. Moreover, the explicit goal restriction operator ``\[n]\'' exposes only the first \n\ subgoals (which need to exist), with default \n = 1\. For example, the method expression ``\simp_all[3]\'' simplifies the first three subgoals, while ``\(rule r, simp_all)[]\'' simplifies all new goals that emerge from applying rule \r\ to the originally first one. \<^medskip> Improper methods, notably tactic emulations, offer low-level goal addressing as explicit argument to the individual tactic being involved. Here ``\[!]\'' refers to all goals, and ``\[n-]\'' to all goals starting from \n\. \<^rail>\ @{syntax_def goal_spec}: '[' (@{syntax nat} '-' @{syntax nat} | @{syntax nat} '-' | @{syntax nat} | '!' ) ']' \ \ subsection \Initial and terminal proof steps \label{sec:proof-steps}\ text \ \begin{matharray}{rcl} @{command_def "proof"} & : & \proof(prove) \ proof(state)\ \\ @{command_def "qed"} & : & \proof(state) \ proof(state) | local_theory | theory\ \\ @{command_def "by"} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def ".."} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def "."} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{command_def "sorry"} & : & \proof(prove) \ proof(state) | local_theory | theory\ \\ @{method_def standard} & : & \method\ \\ \end{matharray} Arbitrary goal refinement via tactics is considered harmful. Structured proof composition in Isar admits proof methods to be invoked in two places only. \<^enum> An \<^emph>\initial\ refinement step @{command_ref "proof"}~\m\<^sub>1\ reduces a newly stated goal to a number of sub-goals that are to be solved later. Facts are passed to \m\<^sub>1\ for forward chaining, if so indicated by \proof(chain)\ mode. \<^enum> A \<^emph>\terminal\ conclusion step @{command_ref "qed"}~\m\<^sub>2\ is intended to solve remaining goals. No facts are passed to \m\<^sub>2\. The only other (proper) way to affect pending goals in a proof body is by @{command_ref "show"}, which involves an explicit statement of what is to be solved eventually. Thus we avoid the fundamental problem of unstructured tactic scripts that consist of numerous consecutive goal transformations, with invisible effects. \<^medskip> As a general rule of thumb for good proof style, initial proof methods should either solve the goal completely, or constitute some well-understood reduction to new sub-goals. Arbitrary automatic proof tools that are prone leave a large number of badly structured sub-goals are no help in continuing the proof document in an intelligible manner. Unless given explicitly by the user, the default initial method is @{method standard}, which subsumes at least @{method_ref (Pure) rule} or its classical variant @{method_ref (HOL) rule}. These methods apply a single standard elimination or introduction rule according to the topmost logical connective involved. There is no separate default terminal method. Any remaining goals are always solved by assumption in the very last step. \<^rail>\ @@{command proof} method? ; @@{command qed} method? ; @@{command "by"} method method? ; (@@{command "."} | @@{command ".."} | @@{command sorry}) \ \<^descr> @{command "proof"}~\m\<^sub>1\ refines the goal by proof method \m\<^sub>1\; facts for forward chaining are passed if so indicated by \proof(chain)\ mode. \<^descr> @{command "qed"}~\m\<^sub>2\ refines any remaining goals by proof method \m\<^sub>2\ and concludes the sub-proof by assumption. If the goal had been \show\, some pending sub-goal is solved as well by the rule resulting from the result \<^emph>\exported\ into the enclosing goal context. Thus \qed\ may fail for two reasons: either \m\<^sub>2\ fails, or the resulting rule does not fit to any pending goal\<^footnote>\This includes any additional ``strong'' assumptions as introduced by @{command "assume"}.\ of the enclosing context. Debugging such a situation might involve temporarily changing @{command "show"} into @{command "have"}, or weakening the local context by replacing occurrences of @{command "assume"} by @{command "presume"}. \<^descr> @{command "by"}~\m\<^sub>1 m\<^sub>2\ is a \<^emph>\terminal proof\\index{proof!terminal}; it abbreviates @{command "proof"}~\m\<^sub>1\~@{command "qed"}~\m\<^sub>2\, but with backtracking across both methods. Debugging an unsuccessful @{command "by"}~\m\<^sub>1 m\<^sub>2\ command can be done by expanding its definition; in many cases @{command "proof"}~\m\<^sub>1\ (or even \apply\~\m\<^sub>1\) is already sufficient to see the problem. \<^descr> ``@{command ".."}'' is a \<^emph>\standard proof\\index{proof!standard}; it abbreviates @{command "by"}~\standard\. \<^descr> ``@{command "."}'' is a \<^emph>\trivial proof\\index{proof!trivial}; it abbreviates @{command "by"}~\this\. \<^descr> @{command "sorry"} is a \<^emph>\fake proof\\index{proof!fake} pretending to solve the pending claim without further ado. This only works in interactive development, or if the @{attribute quick_and_dirty} is enabled. Facts emerging from fake proofs are not the real thing. Internally, the derivation object is tainted by an oracle invocation, which may be inspected via the command @{command "thm_oracles"} (\secref{sec:oracles}). The most important application of @{command "sorry"} is to support experimentation and top-down proof development. \<^descr> @{method standard} refers to the default refinement step of some Isar language elements (notably @{command proof} and ``@{command ".."}''). It is \<^emph>\dynamically scoped\, so the behaviour depends on the application environment. In Isabelle/Pure, @{method standard} performs elementary introduction~/ elimination steps (@{method_ref (Pure) rule}), introduction of type classes (@{method_ref intro_classes}) and locales (@{method_ref intro_locales}). In Isabelle/HOL, @{method standard} also takes classical rules into account (cf.\ \secref{sec:classical}). \ subsection \Fundamental methods and attributes \label{sec:pure-meth-att}\ text \ The following proof methods and attributes refer to basic logical operations of Isar. Further methods and attributes are provided by several generic and object-logic specific tools and packages (see \chref{ch:gen-tools} and \partref{part:hol}). \begin{matharray}{rcl} @{command_def "print_rules"}\\<^sup>*\ & : & \context \\ \\[0.5ex] @{method_def "-"} & : & \method\ \\ @{method_def "goal_cases"} & : & \method\ \\ @{method_def "subproofs"} & : & \method\ \\ @{method_def "fact"} & : & \method\ \\ @{method_def "assumption"} & : & \method\ \\ @{method_def "this"} & : & \method\ \\ @{method_def (Pure) "rule"} & : & \method\ \\ @{attribute_def (Pure) "intro"} & : & \attribute\ \\ @{attribute_def (Pure) "elim"} & : & \attribute\ \\ @{attribute_def (Pure) "dest"} & : & \attribute\ \\ @{attribute_def (Pure) "rule"} & : & \attribute\ \\[0.5ex] @{attribute_def "OF"} & : & \attribute\ \\ @{attribute_def "of"} & : & \attribute\ \\ @{attribute_def "where"} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{method goal_cases} (@{syntax name}*) ; @@{method subproofs} @{syntax method} ; @@{method fact} @{syntax thms}? ; @@{method (Pure) rule} @{syntax thms}? ; rulemod: ('intro' | 'elim' | 'dest') ((('!' | () | '?') @{syntax nat}?) | 'del') ':' @{syntax thms} ; (@@{attribute intro} | @@{attribute elim} | @@{attribute dest}) ('!' | () | '?') @{syntax nat}? ; @@{attribute (Pure) rule} 'del' ; @@{attribute OF} @{syntax thms} ; @@{attribute of} @{syntax insts} ('concl' ':' @{syntax insts})? @{syntax for_fixes} ; @@{attribute "where"} @{syntax named_insts} @{syntax for_fixes} \ \<^descr> @{command "print_rules"} prints rules declared via attributes @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure) dest} of Isabelle/Pure. See also the analogous @{command "print_claset"} command for similar rule declarations of the classical reasoner (\secref{sec:classical}). \<^descr> ``@{method "-"}'' (minus) inserts the forward chaining facts as premises into the goal, and nothing else. Note that command @{command_ref "proof"} without any method actually performs a single reduction step using the @{method_ref (Pure) rule} method; thus a plain \<^emph>\do-nothing\ proof step would be ``@{command "proof"}~\-\'' rather than @{command "proof"} alone. \<^descr> @{method "goal_cases"}~\a\<^sub>1 \ a\<^sub>n\ turns the current subgoals into cases within the context (see also \secref{sec:cases-induct}). The specified case names are used if present; otherwise cases are numbered starting from 1. Invoking cases in the subsequent proof body via the @{command_ref case} command will @{command fix} goal parameters, @{command assume} goal premises, and @{command let} variable @{variable_ref ?case} refer to the conclusion. \<^descr> @{method "subproofs"}~\m\ applies the method expression \m\ consecutively to each subgoal, constructing individual subproofs internally (analogous to ``\<^theory_text>\show goal by m\'' for each subgoal of the proof state). Search alternatives of \m\ are truncated: the method is forced to be deterministic. This method combinator impacts the internal construction of proof terms: it makes a cascade of let-expressions within the derivation tree and may thus improve scalability. \<^descr> @{method "fact"}~\a\<^sub>1 \ a\<^sub>n\ composes some fact from \a\<^sub>1, \, a\<^sub>n\ (or implicitly from the current proof context) modulo unification of schematic type and term variables. The rule structure is not taken into account, i.e.\ meta-level implication is considered atomic. This is the same principle underlying literal facts (cf.\ \secref{sec:syn-att}): ``@{command "have"}~\\\~@{command "by"}~\fact\'' is equivalent to ``@{command "note"}~\<^verbatim>\`\\\\\<^verbatim>\`\'' provided that \\ \\ is an instance of some known \\ \\ in the proof context. \<^descr> @{method assumption} solves some goal by a single assumption step. All given facts are guaranteed to participate in the refinement; this means there may be only 0 or 1 in the first place. Recall that @{command "qed"} (\secref{sec:proof-steps}) already concludes any remaining sub-goals by assumption, so structured proofs usually need not quote the @{method assumption} method at all. \<^descr> @{method this} applies all of the current facts directly as rules. Recall that ``@{command "."}'' (dot) abbreviates ``@{command "by"}~\this\''. \<^descr> @{method (Pure) rule}~\a\<^sub>1 \ a\<^sub>n\ applies some rule given as argument in backward manner; facts are used to reduce the rule before applying it to the goal. Thus @{method (Pure) rule} without facts is plain introduction, while with facts it becomes elimination. When no arguments are given, the @{method (Pure) rule} method tries to pick appropriate rules automatically, as declared in the current context using the @{attribute (Pure) intro}, @{attribute (Pure) elim}, @{attribute (Pure) dest} attributes (see below). This is included in the standard behaviour of @{command "proof"} and ``@{command ".."}'' (double-dot) steps (see \secref{sec:proof-steps}). \<^descr> @{attribute (Pure) intro}, @{attribute (Pure) elim}, and @{attribute (Pure) dest} declare introduction, elimination, and destruct rules, to be used with method @{method (Pure) rule}, and similar tools. Note that the latter will ignore rules declared with ``\?\'', while ``\!\'' are used most aggressively. The classical reasoner (see \secref{sec:classical}) introduces its own variants of these attributes; use qualified names to access the present versions of Isabelle/Pure, i.e.\ @{attribute (Pure) "Pure.intro"}. \<^descr> @{attribute (Pure) rule}~\del\ undeclares introduction, elimination, or destruct rules. \<^descr> @{attribute OF}~\a\<^sub>1 \ a\<^sub>n\ applies some theorem to all of the given rules \a\<^sub>1, \, a\<^sub>n\ in canonical right-to-left order, which means that premises stemming from the \a\<^sub>i\ emerge in parallel in the result, without interfering with each other. In many practical situations, the \a\<^sub>i\ do not have premises themselves, so \rule [OF a\<^sub>1 \ a\<^sub>n]\ can be actually read as functional application (modulo unification). Argument positions may be effectively skipped by using ``\_\'' (underscore), which refers to the propositional identity rule in the Pure theory. \<^descr> @{attribute of}~\t\<^sub>1 \ t\<^sub>n\ performs positional instantiation of term variables. The terms \t\<^sub>1, \, t\<^sub>n\ are substituted for any schematic variables occurring in a theorem from left to right; ``\_\'' (underscore) indicates to skip a position. Arguments following a ``\concl:\'' specification refer to positions of the conclusion of a rule. An optional context of local variables \\ x\<^sub>1 \ x\<^sub>m\ may be specified: the instantiated theorem is exported, and these variables become schematic (usually with some shifting of indices). \<^descr> @{attribute "where"}~\x\<^sub>1 = t\<^sub>1 \ \ x\<^sub>n = t\<^sub>n\ performs named instantiation of schematic type and term variables occurring in a theorem. Schematic variables have to be specified on the left-hand side (e.g.\ \?x1.3\). The question mark may be omitted if the variable name is a plain identifier without index. As type instantiations are inferred from term instantiations, explicit type instantiations are seldom necessary. An optional context of local variables \\ x\<^sub>1 \ x\<^sub>m\ may be specified as for @{attribute "of"} above. \ subsection \Defining proof methods\ text \ \begin{matharray}{rcl} @{command_def "method_setup"} & : & \local_theory \ local_theory\ \\ \end{matharray} \<^rail>\ @@{command method_setup} @{syntax name} '=' @{syntax text} @{syntax text}? \ \<^descr> @{command "method_setup"}~\name = text description\ defines a proof method in the current context. The given \text\ has to be an ML expression of type \<^ML_type>\(Proof.context -> Proof.method) context_parser\, cf.\ basic parsers defined in structure \<^ML_structure>\Args\ and \<^ML_structure>\Attrib\. There are also combinators like \<^ML>\METHOD\ and \<^ML>\SIMPLE_METHOD\ to turn certain tactic forms into official proof methods; the primed versions refer to tactics with explicit goal addressing. Here are some example method definitions: \ (*<*)experiment begin(*>*) method_setup my_method1 = \Scan.succeed (K (SIMPLE_METHOD' (fn i: int => no_tac)))\ "my first method (without any arguments)" method_setup my_method2 = \Scan.succeed (fn ctxt: Proof.context => SIMPLE_METHOD' (fn i: int => no_tac))\ "my second method (with context)" method_setup my_method3 = \Attrib.thms >> (fn thms: thm list => fn ctxt: Proof.context => SIMPLE_METHOD' (fn i: int => no_tac))\ "my third method (with theorem arguments and context)" (*<*)end(*>*) section \Proof by cases and induction \label{sec:cases-induct}\ subsection \Rule contexts\ text \ \begin{matharray}{rcl} @{command_def "case"} & : & \proof(state) \ proof(state)\ \\ @{command_def "print_cases"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def case_names} & : & \attribute\ \\ @{attribute_def case_conclusion} & : & \attribute\ \\ @{attribute_def params} & : & \attribute\ \\ @{attribute_def consumes} & : & \attribute\ \\ \end{matharray} The puristic way to build up Isar proof contexts is by explicit language elements like @{command "fix"}, @{command "assume"}, @{command "let"} (see \secref{sec:proof-context}). This is adequate for plain natural deduction, but easily becomes unwieldy in concrete verification tasks, which typically involve big induction rules with several cases. The @{command "case"} command provides a shorthand to refer to a local context symbolically: certain proof methods provide an environment of named ``cases'' of the form \c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>n\; the effect of ``@{command "case"}~\c\'' is then equivalent to ``@{command "fix"}~\x\<^sub>1 \ x\<^sub>m\~@{command "assume"}~\c: \\<^sub>1 \ \\<^sub>n\''. Term bindings may be covered as well, notably @{variable ?case} for the main conclusion. By default, the ``terminology'' \x\<^sub>1, \, x\<^sub>m\ of a case value is marked as hidden, i.e.\ there is no way to refer to such parameters in the subsequent proof text. After all, original rule parameters stem from somewhere outside of the current proof text. By using the explicit form ``@{command "case"}~\(c y\<^sub>1 \ y\<^sub>m)\'' instead, the proof author is able to chose local names that fit nicely into the current context. \<^medskip> It is important to note that proper use of @{command "case"} does not provide means to peek at the current goal state, which is not directly observable in Isar! Nonetheless, goal refinement commands do provide named cases \goal\<^sub>i\ for each subgoal \i = 1, \, n\ of the resulting goal state. Using this extra feature requires great care, because some bits of the internal tactical machinery intrude the proof text. In particular, parameter names stemming from the left-over of automated reasoning tools are usually quite unpredictable. Under normal circumstances, the text of cases emerge from standard elimination or induction rules, which in turn are derived from previous theory specifications in a canonical way (say from @{command "inductive"} definitions). \<^medskip> Proper cases are only available if both the proof method and the rules involved support this. By using appropriate attributes, case names, conclusions, and parameters may be also declared by hand. Thus variant versions of rules that have been derived manually become ready to use in advanced case analysis later. \<^rail>\ @@{command case} @{syntax thmdecl}? (name | '(' name (('_' | @{syntax name}) *) ')') ; @@{attribute case_names} ((@{syntax name} ( '[' (('_' | @{syntax name}) *) ']' ) ? ) +) ; @@{attribute case_conclusion} @{syntax name} (@{syntax name} * ) ; @@{attribute params} ((@{syntax name} * ) + @'and') ; @@{attribute consumes} @{syntax int}? \ \<^descr> @{command "case"}~\a: (c x\<^sub>1 \ x\<^sub>m)\ invokes a named local context \c: x\<^sub>1, \, x\<^sub>m, \\<^sub>1, \, \\<^sub>m\, as provided by an appropriate proof method (such as @{method_ref cases} and @{method_ref induct}). The command ``@{command "case"}~\a: (c x\<^sub>1 \ x\<^sub>m)\'' abbreviates ``@{command "fix"}~\x\<^sub>1 \ x\<^sub>m\~@{command "assume"}~\a.c: \\<^sub>1 \ \\<^sub>n\''. Each local fact is qualified by the prefix \a\, and all such facts are collectively bound to the name \a\. The fact name is specification \a\ is optional, the default is to re-use \c\. So @{command "case"}~\(c x\<^sub>1 \ x\<^sub>m)\ is the same as @{command "case"}~\c: (c x\<^sub>1 \ x\<^sub>m)\. \<^descr> @{command "print_cases"} prints all local contexts of the current state, using Isar proof language notation. \<^descr> @{attribute case_names}~\c\<^sub>1 \ c\<^sub>k\ declares names for the local contexts of premises of a theorem; \c\<^sub>1, \, c\<^sub>k\ refers to the \<^emph>\prefix\ of the list of premises. Each of the cases \c\<^sub>i\ can be of the form \c[h\<^sub>1 \ h\<^sub>n]\ where the \h\<^sub>1 \ h\<^sub>n\ are the names of the hypotheses in case \c\<^sub>i\ from left to right. \<^descr> @{attribute case_conclusion}~\c d\<^sub>1 \ d\<^sub>k\ declares names for the conclusions of a named premise \c\; here \d\<^sub>1, \, d\<^sub>k\ refers to the prefix of arguments of a logical formula built by nesting a binary connective (e.g.\ \\\). Note that proof methods such as @{method induct} and @{method coinduct} already provide a default name for the conclusion as a whole. The need to name subformulas only arises with cases that split into several sub-cases, as in common co-induction rules. \<^descr> @{attribute params}~\p\<^sub>1 \ p\<^sub>m \ \ q\<^sub>1 \ q\<^sub>n\ renames the innermost parameters of premises \1, \, n\ of some theorem. An empty list of names may be given to skip positions, leaving the present parameters unchanged. Note that the default usage of case rules does \<^emph>\not\ directly expose parameters to the proof context. \<^descr> @{attribute consumes}~\n\ declares the number of ``major premises'' of a rule, i.e.\ the number of facts to be consumed when it is applied by an appropriate proof method. The default value of @{attribute consumes} is \n = 1\, which is appropriate for the usual kind of cases and induction rules for inductive sets (cf.\ \secref{sec:hol-inductive}). Rules without any @{attribute consumes} declaration given are treated as if @{attribute consumes}~\0\ had been specified. A negative \n\ is interpreted relatively to the total number of premises of the rule in the target context. Thus its absolute value specifies the remaining number of premises, after subtracting the prefix of major premises as indicated above. This form of declaration has the technical advantage of being stable under more morphisms, notably those that export the result from a nested @{command_ref context} with additional assumptions. Note that explicit @{attribute consumes} declarations are only rarely needed; this is already taken care of automatically by the higher-level @{attribute cases}, @{attribute induct}, and @{attribute coinduct} declarations. \ subsection \Proof methods\ text \ \begin{matharray}{rcl} @{method_def cases} & : & \method\ \\ @{method_def induct} & : & \method\ \\ @{method_def induction} & : & \method\ \\ @{method_def coinduct} & : & \method\ \\ \end{matharray} The @{method cases}, @{method induct}, @{method induction}, and @{method coinduct} methods provide a uniform interface to common proof techniques over datatypes, inductive predicates (or sets), recursive functions etc. The corresponding rules may be specified and instantiated in a casual manner. Furthermore, these methods provide named local contexts that may be invoked via the @{command "case"} proof command within the subsequent proof text. This accommodates compact proof texts even when reasoning about large specifications. The @{method induct} method also provides some infrastructure to work with structured statements (either using explicit meta-level connectives, or including facts and parameters separately). This avoids cumbersome encoding of ``strengthened'' inductive statements within the object-logic. Method @{method induction} differs from @{method induct} only in the names of the facts in the local context invoked by the @{command "case"} command. \<^rail>\ @@{method cases} ('(' 'no_simp' ')')? \ (@{syntax insts} * @'and') rule? ; (@@{method induct} | @@{method induction}) ('(' 'no_simp' ')')? (definsts * @'and') \ arbitrary? taking? rule? ; @@{method coinduct} @{syntax insts} taking rule? ; rule: ('type' | 'pred' | 'set') ':' (@{syntax name} +) | 'rule' ':' (@{syntax thm} +) ; definst: @{syntax name} ('==' | '\') @{syntax term} | '(' @{syntax term} ')' | @{syntax inst} ; definsts: ( definst * ) ; arbitrary: 'arbitrary' ':' ((@{syntax term} * ) @'and' +) ; taking: 'taking' ':' @{syntax insts} \ \<^descr> @{method cases}~\insts R\ applies method @{method rule} with an appropriate case distinction theorem, instantiated to the subjects \insts\. Symbolic case names are bound according to the rule's local contexts. The rule is determined as follows, according to the facts and arguments passed to the @{method cases} method: \<^medskip> \begin{tabular}{llll} facts & & arguments & rule \\\hline \\ R\ & @{method cases} & & implicit rule \R\ \\ & @{method cases} & & classical case split \\ & @{method cases} & \t\ & datatype exhaustion (type of \t\) \\ \\ A t\ & @{method cases} & \\\ & inductive predicate/set elimination (of \A\) \\ \\\ & @{method cases} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Several instantiations may be given, referring to the \<^emph>\suffix\ of premises of the case rule; within each premise, the \<^emph>\prefix\ of variables is instantiated. In most situations, only a single term needs to be specified; this refers to the first variable of the last premise (it is usually the same for all cases). The \(no_simp)\ option can be used to disable pre-simplification of cases (see the description of @{method induct} below for details). \<^descr> @{method induct}~\insts R\ and @{method induction}~\insts R\ are analogous to the @{method cases} method, but refer to induction rules, which are determined as follows: \<^medskip> \begin{tabular}{llll} facts & & arguments & rule \\\hline & @{method induct} & \P x\ & datatype induction (type of \x\) \\ \\ A x\ & @{method induct} & \\\ & predicate/set induction (of \A\) \\ \\\ & @{method induct} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Several instantiations may be given, each referring to some part of a mutual inductive definition or datatype --- only related partial induction rules may be used together, though. Any of the lists of terms \P, x, \\ refers to the \<^emph>\suffix\ of variables present in the induction rule. This enables the writer to specify only induction variables, or both predicates and variables, for example. Instantiations may be definitional: equations \x \ t\ introduce local definitions, which are inserted into the claim and discharged after applying the induction rule. Equalities reappear in the inductive cases, but have been transformed according to the induction principle being involved here. In order to achieve practically useful induction hypotheses, some variables occurring in \t\ need to generalized (see below). Instantiations of the form \t\, where \t\ is not a variable, are taken as a shorthand for \x \ t\, where \x\ is a fresh variable. If this is not intended, \t\ has to be enclosed in parentheses. By default, the equalities generated by definitional instantiations are pre-simplified using a specific set of rules, usually consisting of distinctness and injectivity theorems for datatypes. This pre-simplification may cause some of the parameters of an inductive case to disappear, or may even completely delete some of the inductive cases, if one of the equalities occurring in their premises can be simplified to \False\. The \(no_simp)\ option can be used to disable pre-simplification. Additional rules to be used in pre-simplification can be declared using the @{attribute_def induct_simp} attribute. The optional ``\arbitrary: x\<^sub>1 \ x\<^sub>m\'' specification generalizes variables \x\<^sub>1, \, x\<^sub>m\ of the original goal before applying induction. It is possible to separate variables by ``\and\'' to generalize in goals other than the first. Thus induction hypotheses may become sufficiently general to get the proof through. Together with definitional instantiations, one may effectively perform induction over expressions of a certain structure. The optional ``\taking: t\<^sub>1 \ t\<^sub>n\'' specification provides additional instantiations of a prefix of pending variables in the rule. Such schematic induction rules rarely occur in practice, though. \<^descr> @{method coinduct}~\inst R\ is analogous to the @{method induct} method, but refers to coinduction rules, which are determined as follows: \<^medskip> \begin{tabular}{llll} goal & & arguments & rule \\\hline & @{method coinduct} & \x\ & type coinduction (type of \x\) \\ \A x\ & @{method coinduct} & \\\ & predicate/set coinduction (of \A\) \\ \\\ & @{method coinduct} & \\ rule: R\ & explicit rule \R\ \\ \end{tabular} \<^medskip> Coinduction is the dual of induction. Induction essentially eliminates \A x\ towards a generic result \P x\, while coinduction introduces \A x\ starting with \B x\, for a suitable ``bisimulation'' \B\. The cases of a coinduct rule are typically named after the predicates or sets being covered, while the conclusions consist of several alternatives being named after the individual destructor patterns. The given instantiation refers to the \<^emph>\suffix\ of variables occurring in the rule's major premise, or conclusion if unavailable. An additional ``\taking: t\<^sub>1 \ t\<^sub>n\'' specification may be required in order to specify the bisimulation to be used in the coinduction step. Above methods produce named local contexts, as determined by the instantiated rule as given in the text. Beyond that, the @{method induct} and @{method coinduct} methods guess further instantiations from the goal specification itself. Any persisting unresolved schematic variables of the resulting rule will render the the corresponding case invalid. The term binding @{variable ?case} for the conclusion will be provided with each case, provided that term is fully specified. The @{command "print_cases"} command prints all named cases present in the current proof state. \<^medskip> Despite the additional infrastructure, both @{method cases} and @{method coinduct} merely apply a certain rule, after instantiation, while conforming due to the usual way of monotonic natural deduction: the context of a structured statement \\x\<^sub>1 \ x\<^sub>m. \\<^sub>1 \ \ \\<^sub>n \ \\ reappears unchanged after the case split. The @{method induct} method is fundamentally different in this respect: the meta-level structure is passed through the ``recursive'' course involved in the induction. Thus the original statement is basically replaced by separate copies, corresponding to the induction hypotheses and conclusion; the original goal context is no longer available. Thus local assumptions, fixed parameters and definitions effectively participate in the inductive rephrasing of the original statement. In @{method induct} proofs, local assumptions introduced by cases are split into two different kinds: \hyps\ stemming from the rule and \prems\ from the goal statement. This is reflected in the extracted cases accordingly, so invoking ``@{command "case"}~\c\'' will provide separate facts \c.hyps\ and \c.prems\, as well as fact \c\ to hold the all-inclusive list. In @{method induction} proofs, local assumptions introduced by cases are split into three different kinds: \IH\, the induction hypotheses, \hyps\, the remaining hypotheses stemming from the rule, and \prems\, the assumptions from the goal statement. The names are \c.IH\, \c.hyps\ and \c.prems\, as above. \<^medskip> Facts presented to either method are consumed according to the number of ``major premises'' of the rule involved, which is usually 0 for plain cases and induction rules of datatypes etc.\ and 1 for rules of inductive predicates or sets and the like. The remaining facts are inserted into the goal verbatim before the actual \cases\, \induct\, or \coinduct\ rule is applied. \ subsection \Declaring rules\ text \ \begin{matharray}{rcl} @{command_def "print_induct_rules"}\\<^sup>*\ & : & \context \\ \\ @{attribute_def cases} & : & \attribute\ \\ @{attribute_def induct} & : & \attribute\ \\ @{attribute_def coinduct} & : & \attribute\ \\ \end{matharray} \<^rail>\ @@{attribute cases} spec ; @@{attribute induct} spec ; @@{attribute coinduct} spec ; spec: (('type' | 'pred' | 'set') ':' @{syntax name}) | 'del' \ \<^descr> @{command "print_induct_rules"} prints cases and induct rules for predicates (or sets) and types of the current context. \<^descr> @{attribute cases}, @{attribute induct}, and @{attribute coinduct} (as attributes) declare rules for reasoning about (co)inductive predicates (or sets) and types, using the corresponding methods of the same name. Certain definitional packages of object-logics usually declare emerging cases and induction rules as expected, so users rarely need to intervene. Rules may be deleted via the \del\ specification, which covers all of the \type\/\pred\/\set\ sub-categories simultaneously. For example, @{attribute cases}~\del\ removes any @{attribute cases} rules declared for some type, predicate, or set. Manual rule declarations usually refer to the @{attribute case_names} and @{attribute params} attributes to adjust names of cases and parameters of a rule; the @{attribute consumes} declaration is taken care of automatically: @{attribute consumes}~\0\ is specified for ``type'' rules and @{attribute consumes}~\1\ for ``predicate'' / ``set'' rules. \ section \Generalized elimination and case splitting \label{sec:obtain}\ text \ \begin{matharray}{rcl} @{command_def "consider"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "obtain"} & : & \proof(state) | proof(chain) \ proof(prove)\ \\ @{command_def "guess"}\\<^sup>*\ & : & \proof(state) | proof(chain) \ proof(prove)\ \\ \end{matharray} Generalized elimination means that hypothetical parameters and premises may be introduced in the current context, potentially with a split into cases. This works by virtue of a locally proven rule that establishes the soundness of this temporary context extension. As representative examples, one may think of standard rules from Isabelle/HOL like this: \<^medskip> \begin{tabular}{ll} \\x. B x \ (\x. B x \ thesis) \ thesis\ \\ \A \ B \ (A \ B \ thesis) \ thesis\ \\ \A \ B \ (A \ thesis) \ (B \ thesis) \ thesis\ \\ \end{tabular} \<^medskip> In general, these particular rules and connectives need to get involved at all: this concept works directly in Isabelle/Pure via Isar commands defined below. In particular, the logic of elimination and case splitting is delegated to an Isar proof, which often involves automated tools. \<^rail>\ @@{command consider} @{syntax obtain_clauses} ; @@{command obtain} @{syntax par_name}? @{syntax vars} \ @'where' concl prems @{syntax for_fixes} ; concl: (@{syntax props} + @'and') ; prems: (@'if' (@{syntax props'} + @'and'))? ; @@{command guess} @{syntax vars} \ \<^descr> @{command consider}~\(a) \<^vec>x \ \<^vec>A \<^vec>x | (b) \<^vec>y \ \<^vec>B \<^vec>y | \\ states a rule for case splitting into separate subgoals, such that each case involves new parameters and premises. After the proof is finished, the resulting rule may be used directly with the @{method cases} proof method (\secref{sec:cases-induct}), in order to perform actual case-splitting of the proof text via @{command case} and @{command next} as usual. Optional names in round parentheses refer to case names: in the proof of the rule this is a fact name, in the resulting rule it is used as annotation with the @{attribute_ref case_names} attribute. \<^medskip> Formally, the command @{command consider} is defined as derived Isar language element as follows: \begin{matharray}{l} @{command "consider"}~\(a) \<^vec>x \ \<^vec>A \<^vec>x | (b) \<^vec>y \ \<^vec>B \<^vec>y | \ \\ \\[1ex] \quad @{command "have"}~\[case_names a b \]: thesis\ \\ \qquad \\ a [Pure.intro?]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis\ \\ \qquad \\ b [Pure.intro?]: \\<^vec>y. \<^vec>B \<^vec>y \ thesis\ \\ \qquad \\ \\ \\ \qquad \\ thesis\ \\ \qquad @{command "apply"}~\(insert a b \)\ \\ \end{matharray} See also \secref{sec:goals} for @{keyword "obtains"} in toplevel goal statements, as well as @{command print_statement} to print existing rules in a similar format. \<^descr> @{command obtain}~\\<^vec>x \ \<^vec>A \<^vec>x\ states a generalized elimination rule with exactly one case. After the proof is finished, it is activated for the subsequent proof text: the context is augmented via @{command fix}~\\<^vec>x\ @{command assume}~\\<^vec>A \<^vec>x\, with special provisions to export later results by discharging these assumptions again. Note that according to the parameter scopes within the elimination rule, results \<^emph>\must not\ refer to hypothetical parameters; otherwise the export will fail! This restriction conforms to the usual manner of existential reasoning in Natural Deduction. \<^medskip> Formally, the command @{command obtain} is defined as derived Isar language element as follows, using an instrumented variant of @{command assume}: \begin{matharray}{l} @{command "obtain"}~\\<^vec>x \ a: \<^vec>A \<^vec>x \proof\ \\ \\[1ex] \quad @{command "have"}~\thesis\ \\ \qquad \\ that [Pure.intro?]: \\<^vec>x. \<^vec>A \<^vec>x \ thesis\ \\ \qquad \\ thesis\ \\ \qquad @{command "apply"}~\(insert that)\ \\ \qquad \\proof\\ \\ \quad @{command "fix"}~\\<^vec>x\~@{command "assume"}\\<^sup>* a: \<^vec>A \<^vec>x\ \\ \end{matharray} \<^descr> @{command guess} is similar to @{command obtain}, but it derives the obtained context elements from the course of tactical reasoning in the proof. Thus it can considerably obscure the proof: it is classified as \<^emph>\improper\. A proof with @{command guess} starts with a fixed goal \thesis\. The subsequent refinement steps may turn this to anything of the form \\\<^vec>x. \<^vec>A \<^vec>x \ thesis\, but without splitting into new subgoals. The final goal state is then used as reduction rule for the obtain pattern described above. Obtained parameters \\<^vec>x\ are marked as internal by default, and thus inaccessible in the proof text. The variable names and type constraints given as arguments for @{command "guess"} specify a prefix of accessible parameters. In the proof of @{command consider} and @{command obtain} the local premises are always bound to the fact name @{fact_ref that}, according to structured Isar statements involving @{keyword_ref "if"} (\secref{sec:goals}). Facts that are established by @{command "obtain"} and @{command "guess"} may not be polymorphic: any type-variables occurring here are fixed in the present context. This is a natural consequence of the role of @{command fix} and @{command assume} in these constructs. \ end diff --git a/src/Doc/antiquote_setup.ML b/src/Doc/antiquote_setup.ML --- a/src/Doc/antiquote_setup.ML +++ b/src/Doc/antiquote_setup.ML @@ -1,246 +1,158 @@ (* Title: Doc/antiquote_setup.ML Author: Makarius Auxiliary antiquotations for the Isabelle manuals. *) structure Antiquote_Setup: sig end = struct (* misc utils *) fun translate f = Symbol.explode #> map f #> implode; val clean_string = translate (fn "_" => "\\_" | "#" => "\\#" | "$" => "\\$" | "%" => "\\%" | "<" => "$<$" | ">" => "$>$" | "{" => "\\{" | "|" => "$\\mid$" | "}" => "\\}" | "\" => "-" | c => c); fun clean_name "\" = "dots" | clean_name ".." = "ddot" | clean_name "." = "dot" | clean_name "_" = "underscore" | clean_name "{" = "braceleft" | clean_name "}" = "braceright" | clean_name s = s |> translate (fn "_" => "-" | "\" => "-" | c => c); -(* ML text *) - -local - -fun test_val (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read ");" - | test_val (ml1, ml2) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read ");"; - -fun test_op (ml1, ml2) = test_val (ML_Lex.read "op " @ ml1, ml2); - -fun test_type (ml1, []) = ML_Lex.read "val _ = NONE : (" @ ml1 @ ML_Lex.read ") option;" - | test_type (ml1, ml2) = - ML_Lex.read "val _ = [NONE : (" @ ml1 @ ML_Lex.read ") option, NONE : (" @ - ml2 @ ML_Lex.read ") option];"; - -fun text_exn (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : exn);" - | text_exn (ml1, ml2) = - ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read " -> exn);"; - -fun test_struct (ml, _) = - ML_Lex.read "functor XXX() = struct structure XX = " @ ml @ ML_Lex.read " end;"; - -fun test_functor (Antiquote.Text tok :: _, _) = - ML_Lex.read "ML_Env.check_functor " @ - ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok)) - | test_functor _ = raise Fail "Bad ML functor specification"; - -val is_name = - ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.Long_Ident); - -fun is_ml_identifier ants = - forall Antiquote.is_text ants andalso - (case filter is_name (map (Antiquote.the_text "") ants) of - toks as [_] => Symbol_Pos.is_identifier (Long_Name.base_name (ML_Lex.flatten toks)) - | _ => false); - -val parse_ml = Args.text_input -- Scan.optional (Args.colon |-- Args.text_input) Input.empty; -val parse_type = Args.text_input -- Scan.optional (Args.$$$ "=" |-- Args.text_input) Input.empty; -val parse_exn = Args.text_input -- Scan.optional (Args.$$$ "of" |-- Args.text_input) Input.empty; -val parse_struct = Args.text_input >> rpair Input.empty; - -fun antiquotation_ml parse test kind show_kind binding index = - Document_Output.antiquotation_raw binding (Scan.lift parse) - (fn ctxt => fn (source1, source2) => - let - val (ml1, ml2) = apply2 ML_Lex.read_source (source1, source2); - val pos = Input.pos_of source1; - val _ = - ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos (test (ml1, ml2)) - handle ERROR msg => error (msg ^ Position.here pos); - - val (txt1, txt2) = apply2 (#1 o Input.source_content) (source1, source2); - val sep = if kind = "type" then "=" else if kind = "exception" then "of" else ":"; - val txt = - if txt2 = "" then txt1 - else if sep = ":" andalso is_ml_identifier ml1 then txt1 ^ ": " ^ txt2 - else txt1 ^ " " ^ sep ^ " " ^ txt2; - - val main_text = - Document_Output.verbatim ctxt - (if kind = "" orelse not show_kind then txt else kind ^ " " ^ txt); - - val index_text = - index |> Option.map (fn def => - let - val ctxt' = Config.put Document_Antiquotation.thy_output_display false ctxt; - val kind' = if kind = "" then " (ML)" else " (ML " ^ kind ^ ")"; - val txt' = Latex.block [Document_Output.verbatim ctxt' txt1, Latex.string kind']; - val like = Document_Antiquotation.approx_content ctxt source1; - in Latex.index_entry {items = [{text = txt', like = like}], def = def} end); - in Latex.block (the_list index_text @ [main_text]) end); - -fun antiquotation_ml' parse test kind binding = - antiquotation_ml parse test kind true binding (SOME true); - -in - -val _ = - Theory.setup - (antiquotation_ml' parse_ml test_val "" \<^binding>\define_ML\ #> - antiquotation_ml' parse_ml test_op "infix" \<^binding>\define_ML_infix\ #> - antiquotation_ml' parse_type test_type "type" \<^binding>\define_ML_type\ #> - antiquotation_ml' parse_exn text_exn "exception" \<^binding>\define_ML_exception\ #> - antiquotation_ml' parse_struct test_struct "structure" \<^binding>\define_ML_structure\ #> - antiquotation_ml' parse_struct test_functor "functor" \<^binding>\define_ML_functor\); - -end; - - (* named theorems *) val _ = Theory.setup (Document_Output.antiquotation_raw \<^binding>\named_thms\ (Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name))) (fn ctxt => map (fn (thm, name) => Output.output (Document_Antiquotation.format ctxt (Document_Antiquotation.delimit ctxt (Document_Output.pretty_thm ctxt thm))) ^ enclose "\\rulename{" "}" (Output.output name)) #> space_implode "\\par\\smallskip%\n" #> Latex.string #> single #> Document_Output.isabelle ctxt)); (* Isabelle/Isar entities (with index) *) local fun no_check (_: Proof.context) (name, _: Position.T) = name; fun check_keyword ctxt (name, pos) = if Keyword.is_keyword (Thy_Header.get_keywords' ctxt) name then name else error ("Bad outer syntax keyword " ^ quote name ^ Position.here pos); fun check_system_option ctxt arg = (Completion.check_option (Options.default ()) ctxt arg; true) handle ERROR _ => false; val arg = enclose "{" "}" o clean_string; fun entity check markup binding index = Document_Output.antiquotation_raw (binding |> Binding.map_name (fn name => name ^ (case index of NONE => "" | SOME true => "_def" | SOME false => "_ref"))) (Scan.lift (Scan.optional (Args.parens Args.name) "" -- Args.name_position)) (fn ctxt => fn (logic, (name, pos)) => let val kind = translate (fn "_" => " " | c => c) (Binding.name_of binding); val hyper_name = "{" ^ Long_Name.append kind (Long_Name.append logic (clean_name name)) ^ "}"; val hyper = enclose ("\\hyperlink" ^ hyper_name ^ "{") "}" #> index = SOME true ? enclose ("\\hypertarget" ^ hyper_name ^ "{") "}"; val idx = (case index of NONE => "" | SOME is_def => "\\index" ^ (if is_def then "def" else "ref") ^ arg logic ^ arg kind ^ arg name); val _ = if Context_Position.is_reported ctxt pos then ignore (check ctxt (name, pos)) else (); val latex = idx ^ (Output.output name |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}") |> hyper o enclose "\\mbox{\\isa{" "}}"); in Latex.string latex end); fun entity_antiqs check markup kind = entity check markup kind NONE #> entity check markup kind (SOME true) #> entity check markup kind (SOME false); in val _ = Theory.setup (entity_antiqs no_check "" \<^binding>\syntax\ #> entity_antiqs Outer_Syntax.check_command "isacommand" \<^binding>\command\ #> entity_antiqs check_keyword "isakeyword" \<^binding>\keyword\ #> entity_antiqs check_keyword "isakeyword" \<^binding>\element\ #> entity_antiqs Method.check_name "" \<^binding>\method\ #> entity_antiqs Attrib.check_name "" \<^binding>\attribute\ #> entity_antiqs no_check "" \<^binding>\fact\ #> entity_antiqs no_check "" \<^binding>\variable\ #> entity_antiqs no_check "" \<^binding>\case\ #> entity_antiqs Document_Antiquotation.check "" \<^binding>\antiquotation\ #> entity_antiqs Document_Antiquotation.check_option "" \<^binding>\antiquotation_option\ #> entity_antiqs Document_Marker.check "" \<^binding>\document_marker\ #> entity_antiqs no_check "isasystem" \<^binding>\setting\ #> entity_antiqs check_system_option "isasystem" \<^binding>\system_option\ #> entity_antiqs no_check "" \<^binding>\inference\ #> entity_antiqs no_check "isasystem" \<^binding>\executable\ #> entity_antiqs Isabelle_Tool.check "isatool" \<^binding>\tool\ #> entity_antiqs ML_Context.check_antiquotation "" \<^binding>\ML_antiquotation\ #> entity_antiqs (K JEdit.check_action) "isasystem" \<^binding>\action\); end; (* show symbols *) val _ = Theory.setup (Document_Output.antiquotation_raw \<^binding>\show_symbols\ (Scan.succeed ()) (fn _ => fn _ => let val symbol_name = unprefix "\\newcommand{\\isasym" #> raw_explode #> take_prefix Symbol.is_ascii_letter #> implode; val symbols = File.read \<^file>\~~/lib/texinputs/isabellesym.sty\ |> split_lines |> map_filter (fn line => (case try symbol_name line of NONE => NONE | SOME "" => NONE | SOME name => SOME ("\\verb,\\" ^ "<" ^ name ^ ">, & {\\isasym" ^ name ^ "}"))); val eol = "\\\\\n"; fun table (a :: b :: rest) = a ^ " & " ^ b ^ eol :: table rest | table [a] = [a ^ eol] | table [] = []; in Latex.string ("\\begin{supertabular}{ll@{\\qquad}ll}\n" ^ implode (table symbols) ^ "\\end{supertabular}\n") end)) end; diff --git a/src/Pure/Thy/document_antiquotations.ML b/src/Pure/Thy/document_antiquotations.ML --- a/src/Pure/Thy/document_antiquotations.ML +++ b/src/Pure/Thy/document_antiquotations.ML @@ -1,370 +1,437 @@ (* Title: Pure/Thy/document_antiquotations.ML Author: Makarius Miscellaneous document antiquotations. *) structure Document_Antiquotations: sig end = struct (* basic entities *) local type style = term -> term; fun pretty_term_style ctxt (style: style, t) = Document_Output.pretty_term ctxt (style t); fun pretty_thms_style ctxt (style: style, ths) = map (fn th => Document_Output.pretty_term ctxt (style (Thm.full_prop_of th))) ths; fun pretty_term_typ ctxt (style: style, t) = let val t' = style t in Document_Output.pretty_term ctxt (Type.constraint (Term.fastype_of t') t') end; fun pretty_term_typeof ctxt (style: style, t) = Syntax.pretty_typ ctxt (Term.fastype_of (style t)); fun pretty_const ctxt c = let val t = Const (c, Consts.type_scheme (Proof_Context.consts_of ctxt) c) handle TYPE (msg, _, _) => error msg; val (t', _) = yield_singleton (Variable.import_terms true) t ctxt; in Document_Output.pretty_term ctxt t' end; fun pretty_abbrev ctxt s = let val t = Syntax.read_term (Proof_Context.set_mode Proof_Context.mode_abbrev ctxt) s; fun err () = error ("Abbreviated constant expected: " ^ Syntax.string_of_term ctxt t); val (head, args) = Term.strip_comb t; val (c, T) = Term.dest_Const head handle TERM _ => err (); val (U, u) = Consts.the_abbreviation (Proof_Context.consts_of ctxt) c handle TYPE _ => err (); val t' = Term.betapplys (Envir.expand_atom T (U, u), args); val eq = Logic.mk_equals (t, t'); val ctxt' = Proof_Context.augment eq ctxt; in Proof_Context.pretty_term_abbrev ctxt' eq end; fun pretty_locale ctxt (name, pos) = let val thy = Proof_Context.theory_of ctxt in Pretty.str (Locale.extern thy (Locale.check thy (name, pos))) end; fun pretty_class ctxt s = Pretty.str (Proof_Context.extern_class ctxt (Proof_Context.read_class ctxt s)); fun pretty_type ctxt s = let val Type (name, _) = Proof_Context.read_type_name {proper = true, strict = false} ctxt s in Pretty.str (Proof_Context.extern_type ctxt name) end; fun pretty_prf full ctxt = Proof_Syntax.pretty_standard_proof_of ctxt full; fun pretty_theory ctxt (name, pos) = (Theory.check {long = true} ctxt (name, pos); Pretty.str name); val basic_entity = Document_Output.antiquotation_pretty_source_embedded; fun basic_entities name scan pretty = Document_Antiquotation.setup name scan (fn {context = ctxt, source = src, argument = xs} => Document_Output.pretty_items_source ctxt {embedded = false} src (map (pretty ctxt) xs)); val _ = Theory.setup (basic_entity \<^binding>\prop\ (Term_Style.parse -- Args.prop) pretty_term_style #> basic_entity \<^binding>\term\ (Term_Style.parse -- Args.term) pretty_term_style #> basic_entity \<^binding>\term_type\ (Term_Style.parse -- Args.term) pretty_term_typ #> basic_entity \<^binding>\typeof\ (Term_Style.parse -- Args.term) pretty_term_typeof #> basic_entity \<^binding>\const\ (Args.const {proper = true, strict = false}) pretty_const #> basic_entity \<^binding>\abbrev\ (Scan.lift Args.embedded_inner_syntax) pretty_abbrev #> basic_entity \<^binding>\typ\ Args.typ_abbrev Syntax.pretty_typ #> basic_entity \<^binding>\locale\ (Scan.lift Args.embedded_position) pretty_locale #> basic_entity \<^binding>\class\ (Scan.lift Args.embedded_inner_syntax) pretty_class #> basic_entity \<^binding>\type\ (Scan.lift Args.embedded_inner_syntax) pretty_type #> basic_entity \<^binding>\theory\ (Scan.lift Args.embedded_position) pretty_theory #> basic_entities \<^binding>\prf\ Attrib.thms (pretty_prf false) #> basic_entities \<^binding>\full_prf\ Attrib.thms (pretty_prf true) #> Document_Antiquotation.setup \<^binding>\thm\ (Term_Style.parse -- Attrib.thms) (fn {context = ctxt, source = src, argument = arg} => Document_Output.pretty_items_source ctxt {embedded = false} src (pretty_thms_style ctxt arg))); in end; (* Markdown errors *) local fun markdown_error binding = Document_Antiquotation.setup binding (Scan.succeed ()) (fn {source = src, ...} => error ("Bad Markdown structure: illegal " ^ quote (Binding.name_of binding) ^ Position.here (Position.no_range_position (#1 (Token.range_of src))))) val _ = Theory.setup (markdown_error \<^binding>\item\ #> markdown_error \<^binding>\enum\ #> markdown_error \<^binding>\descr\); in end; (* control spacing *) val _ = Theory.setup (Document_Output.antiquotation_raw \<^binding>\noindent\ (Scan.succeed ()) (fn _ => fn () => Latex.string "\\noindent") #> Document_Output.antiquotation_raw \<^binding>\smallskip\ (Scan.succeed ()) (fn _ => fn () => Latex.string "\\smallskip") #> Document_Output.antiquotation_raw \<^binding>\medskip\ (Scan.succeed ()) (fn _ => fn () => Latex.string "\\medskip") #> Document_Output.antiquotation_raw \<^binding>\bigskip\ (Scan.succeed ()) (fn _ => fn () => Latex.string "\\bigskip")); (* nested document text *) local fun nested_antiquotation name s1 s2 = Document_Output.antiquotation_raw_embedded name (Scan.lift Args.cartouche_input) (fn ctxt => fn txt => (Context_Position.reports ctxt (Document_Output.document_reports txt); Latex.enclose_block s1 s2 (Document_Output.output_document ctxt {markdown = false} txt))); val _ = Theory.setup (nested_antiquotation \<^binding>\footnote\ "\\footnote{" "}" #> nested_antiquotation \<^binding>\emph\ "\\emph{" "}" #> nested_antiquotation \<^binding>\bold\ "\\textbf{" "}"); in end; (* index entries *) local val index_like = Parse.$$$ "(" |-- Parse.!!! (Parse.$$$ "is" |-- Args.name --| Parse.$$$ ")"); val index_args = Parse.enum1 "!" (Args.embedded_input -- Scan.option index_like); fun output_text ctxt = Latex.block o Document_Output.output_document ctxt {markdown = false}; fun index binding def = Document_Output.antiquotation_raw binding (Scan.lift index_args) (fn ctxt => fn args => let val _ = Context_Position.reports ctxt (maps (Document_Output.document_reports o #1) args); fun make_item (txt, opt_like) = let val text = output_text ctxt txt; val like = (case opt_like of SOME s => s | NONE => Document_Antiquotation.approx_content ctxt txt); val _ = if is_none opt_like andalso Context_Position.is_visible ctxt then writeln ("(" ^ Markup.markup Markup.keyword2 "is" ^ " " ^ quote like ^ ")" ^ Position.here (Input.pos_of txt)) else (); in {text = text, like = like} end; in Latex.index_entry {items = map make_item args, def = def} end); val _ = Theory.setup (index \<^binding>\index_ref\ false #> index \<^binding>\index_def\ true); in end; (* quasi-formal text (unchecked) *) local fun report_text ctxt text = let val pos = Input.pos_of text in Context_Position.reports ctxt [(pos, Markup.language_text (Input.is_delimited text)), (pos, Markup.raw_text)] end; fun prepare_text ctxt = Input.source_content #> #1 #> Document_Antiquotation.prepare_lines ctxt; fun text_antiquotation name = Document_Output.antiquotation_raw_embedded name (Scan.lift Args.text_input) (fn ctxt => fn text => let val _ = report_text ctxt text; in prepare_text ctxt text |> Document_Output.output_source ctxt |> Document_Output.isabelle ctxt end); val theory_text_antiquotation = Document_Output.antiquotation_raw_embedded \<^binding>\theory_text\ (Scan.lift Args.text_input) (fn ctxt => fn text => let val keywords = Thy_Header.get_keywords' ctxt; val _ = report_text ctxt text; val _ = Input.source_explode text |> Token.tokenize keywords {strict = true} |> maps (Token.reports keywords) |> Context_Position.reports_text ctxt; in prepare_text ctxt text |> Token.explode0 keywords |> maps (Document_Output.output_token ctxt) |> Document_Output.isabelle ctxt end); val _ = Theory.setup (text_antiquotation \<^binding>\text\ #> text_antiquotation \<^binding>\cartouche\ #> theory_text_antiquotation); in end; (* goal state *) local fun goal_state name main = Document_Output.antiquotation_pretty name (Scan.succeed ()) (fn ctxt => fn () => Goal_Display.pretty_goal (Config.put Goal_Display.show_main_goal main ctxt) (#goal (Proof.goal (Toplevel.proof_of (Toplevel.presentation_state ctxt))))); val _ = Theory.setup (goal_state \<^binding>\goals\ true #> goal_state \<^binding>\subgoals\ false); in end; (* embedded lemma *) val _ = Theory.setup (Document_Antiquotation.setup \<^binding>\lemma\ (Scan.lift (Scan.ahead Parse.not_eof) -- Args.prop -- Scan.lift (Parse.position (Parse.reserved "by") -- Method.parse -- Scan.option Method.parse)) (fn {context = ctxt, source = src, argument = ((prop_tok, prop), (((_, by_pos), m1), m2))} => let val reports = (by_pos, Markup.keyword1 |> Markup.keyword_properties) :: maps Method.reports_of (m1 :: the_list m2); val _ = Context_Position.reports ctxt reports; (* FIXME check proof!? *) val _ = ctxt |> Proof.theorem NONE (K I) [[(prop, [])]] |> Proof.global_terminal_proof (m1, m2); in Document_Output.pretty_source ctxt {embedded = false} [hd src, prop_tok] (Document_Output.pretty_term ctxt prop) end)); (* verbatim text *) val _ = Theory.setup (Document_Output.antiquotation_verbatim_embedded \<^binding>\verbatim\ (Scan.lift Args.text_input) (fn ctxt => fn text => let val pos = Input.pos_of text; val _ = Context_Position.reports ctxt [(pos, Markup.language_verbatim (Input.is_delimited text)), (pos, Markup.raw_text)]; in #1 (Input.source_content text) end)); (* bash functions *) val _ = Theory.setup (Document_Output.antiquotation_verbatim_embedded \<^binding>\bash_function\ (Scan.lift Args.embedded_position) Isabelle_System.check_bash_function); (* system options *) val _ = Theory.setup (Document_Output.antiquotation_verbatim_embedded \<^binding>\system_option\ (Scan.lift Args.embedded_position) (fn ctxt => fn (name, pos) => let val _ = Completion.check_option (Options.default ()) ctxt (name, pos); in name end)); (* ML text *) local -fun ml_text name ml = - Document_Output.antiquotation_verbatim_embedded name (Scan.lift Args.text_input) - (fn ctxt => fn text => - let val _ = ML_Context.eval_in (SOME ctxt) ML_Compiler.flags (Input.pos_of text) (ml text) - in #1 (Input.source_content text) end); - -fun ml_enclose bg en source = - ML_Lex.read bg @ ML_Lex.read_source source @ ML_Lex.read en; +fun test_val (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read ");" + | test_val (ml1, ml2) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read ");"; -val _ = Theory.setup - (ml_text \<^binding>\ML\ (ml_enclose "fn _ => (" ");") #> - ml_text \<^binding>\ML_op\ (ml_enclose "fn _ => (op " ");") #> - ml_text \<^binding>\ML_type\ (ml_enclose "val _ = NONE : (" ") option;") #> - ml_text \<^binding>\ML_structure\ - (ml_enclose "functor XXX() = struct structure XX = " " end;") #> +fun test_op (ml1, ml2) = test_val (ML_Lex.read "op " @ ml1, ml2); - ml_text \<^binding>\ML_functor\ (* FIXME formal treatment of functor name (!?) *) - (fn source => - ML_Lex.read ("ML_Env.check_functor " ^ - ML_Syntax.print_string (#1 (Input.source_content source)))) #> +fun test_type (ml1, []) = ML_Lex.read "val _ = NONE : (" @ ml1 @ ML_Lex.read ") option;" + | test_type (ml1, ml2) = + ML_Lex.read "val _ = [NONE : (" @ ml1 @ ML_Lex.read ") option, NONE : (" @ + ml2 @ ML_Lex.read ") option];"; - ml_text \<^binding>\ML_text\ (K [])); +fun test_exn (ml1, []) = ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : exn);" + | test_exn (ml1, ml2) = + ML_Lex.read "fn _ => (" @ ml1 @ ML_Lex.read " : " @ ml2 @ ML_Lex.read " -> exn);"; -in end; +fun test_struct (ml, _) = + ML_Lex.read "functor XXX() = struct structure XX = " @ ml @ ML_Lex.read " end;"; + +fun test_functor (Antiquote.Text tok :: _, _) = + ML_Lex.read "ML_Env.check_functor " @ + ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok)) + | test_functor _ = raise Fail "Bad ML functor specification"; + +val is_name = + ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.Long_Ident); + +fun is_ml_identifier ants = + forall Antiquote.is_text ants andalso + (case filter is_name (map (Antiquote.the_text "") ants) of + toks as [_] => Symbol_Pos.is_identifier (Long_Name.base_name (ML_Lex.flatten toks)) + | _ => false); + +val parse_ml0 = Args.text_input >> rpair Input.empty; +val parse_ml = Args.text_input -- Scan.optional (Args.colon |-- Args.text_input) Input.empty; +val parse_type = Args.text_input -- Scan.optional (Args.$$$ "=" |-- Args.text_input) Input.empty; +val parse_exn = Args.text_input -- Scan.optional (Args.$$$ "of" |-- Args.text_input) Input.empty; + +fun antiquotation_ml parse test kind show_kind binding index = + Document_Output.antiquotation_raw binding (Scan.lift parse) + (fn ctxt => fn (source1, source2) => + let + val (ml1, ml2) = apply2 ML_Lex.read_source (source1, source2); + val pos = Input.pos_of source1; + val _ = + ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos (test (ml1, ml2)) + handle ERROR msg => error (msg ^ Position.here pos); + + val (txt1, txt2) = apply2 (#1 o Input.source_content) (source1, source2); + val sep = if kind = "type" then "=" else if kind = "exception" then "of" else ":"; + val txt = + if txt2 = "" then txt1 + else if sep = ":" andalso is_ml_identifier ml1 then txt1 ^ ": " ^ txt2 + else txt1 ^ " " ^ sep ^ " " ^ txt2; + + val main_text = + Document_Output.verbatim ctxt + (if kind = "" orelse not show_kind then txt else kind ^ " " ^ txt); + + val index_text = + index |> Option.map (fn def => + let + val ctxt' = Config.put Document_Antiquotation.thy_output_display false ctxt; + val kind' = if kind = "" then " (ML)" else " (ML " ^ kind ^ ")"; + val txt' = Latex.block [Document_Output.verbatim ctxt' txt1, Latex.string kind']; + val like = Document_Antiquotation.approx_content ctxt source1; + in Latex.index_entry {items = [{text = txt', like = like}], def = def} end); + in Latex.block (the_list index_text @ [main_text]) end); + +fun antiquotation_ml0 test kind = + antiquotation_ml parse_ml0 test kind false; + +fun antiquotation_ml1 parse test kind binding = + antiquotation_ml parse test kind true binding (SOME true); + +in + +val _ = + Theory.setup + (Latex.index_variants (antiquotation_ml0 test_val "") \<^binding>\ML\ #> + Latex.index_variants (antiquotation_ml0 test_op "infix") \<^binding>\ML_infix\ #> + Latex.index_variants (antiquotation_ml0 test_type "type") \<^binding>\ML_type\ #> + Latex.index_variants (antiquotation_ml0 test_struct "structure") \<^binding>\ML_structure\ #> + Latex.index_variants (antiquotation_ml0 test_functor "functor") \<^binding>\ML_functor\ #> + antiquotation_ml0 (K []) "text" \<^binding>\ML_text\ NONE #> + antiquotation_ml1 parse_ml test_val "" \<^binding>\define_ML\ #> + antiquotation_ml1 parse_ml test_op "infix" \<^binding>\define_ML_infix\ #> + antiquotation_ml1 parse_type test_type "type" \<^binding>\define_ML_type\ #> + antiquotation_ml1 parse_exn test_exn "exception" \<^binding>\define_ML_exception\ #> + antiquotation_ml1 parse_ml0 test_struct "structure" \<^binding>\define_ML_structure\ #> + antiquotation_ml1 parse_ml0 test_functor "functor" \<^binding>\define_ML_functor\); + +end; (* URLs *) val escape_url = translate_string (fn c => if c = "%" orelse c = "#" orelse c = "^" then "\\" ^ c else c); val _ = Theory.setup (Document_Output.antiquotation_raw_embedded \<^binding>\url\ (Scan.lift Args.embedded_input) (fn ctxt => fn source => let val url = Input.string_of source; val pos = Input.pos_of source; val delimited = Input.is_delimited source; val _ = Context_Position.reports ctxt [(pos, Markup.language_url delimited), (pos, Markup.url url)]; in Latex.enclose_block "\\url{" "}" [Latex.string (escape_url url)] end)); (* formal entities *) local fun entity_antiquotation name check bg en = Document_Output.antiquotation_raw name (Scan.lift Args.name_position) (fn ctxt => fn (name, pos) => let val _ = check ctxt (name, pos) in Latex.enclose_block bg en [Latex.string (Output.output name)] end); val _ = Theory.setup (entity_antiquotation \<^binding>\command\ Outer_Syntax.check_command "\\isacommand{" "}" #> entity_antiquotation \<^binding>\method\ Method.check_name "\\isa{" "}" #> entity_antiquotation \<^binding>\attribute\ Attrib.check_name "\\isa{" "}"); in end; end;