diff -r fb7756087101 -r 38b049cd3aad src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy Wed Dec 16 16:31:36 2015 +0100 +++ b/src/Doc/Implementation/Logic.thy Wed Dec 16 17:28:49 2015 +0100 @@ -7,110 +7,97 @@ 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. + 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). + 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!\ + 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. + 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>\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. + 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\. + 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. + 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\. + 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)\\. + 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 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 \\ :: + 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 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"}. + 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 \ @@ -135,48 +122,42 @@ \<^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 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 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> 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.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 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.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.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}~\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.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_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_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\. + \<^descr> @{ML Sign.primitive_arity}~\(\, \<^vec>s, s)\ declares the arity \\ :: + (\<^vec>s)s\. \ text %mlantiq \ @@ -201,92 +182,84 @@ @@{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} + \<^descr> \@{class c}\ inlines the internalized class \c\ --- as @{ML_type string} literal. - \<^descr> \@{typ \}\ inlines the internalized type \\\ - --- as constructor term for datatype @{ML_type typ}. + \<^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. + 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>\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>\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>\\. + 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. + 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 + 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. + 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. + 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: + 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 @@ -296,47 +269,46 @@ \] 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. + 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 + 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. + 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. + 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\. + 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. + 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 \ @@ -361,56 +333,52 @@ @{index_ML Sign.const_instance: "theory -> string * typ list -> typ"} \\ \end{mldecls} - \<^descr> Type @{ML_type term} represents de-Bruijn terms, with comments - in abstractions, and explicitly named free variables and constants; - this is a datatype with constructors @{index_ML Bound}, @{index_ML - Free}, @{index_ML Var}, @{index_ML Const}, @{index_ML Abs}, - @{index_ML_op "$"}. + \<^descr> Type @{ML_type term} represents de-Bruijn terms, with comments in + abstractions, and explicitly named free variables and constants; this is a + datatype with constructors @{index_ML Bound}, @{index_ML Free}, @{index_ML + Var}, @{index_ML Const}, @{index_ML Abs}, @{index_ML_op "$"}. - \<^descr> \t\~@{ML_text aconv}~\u\ checks \\\-equivalence of two terms. This is the basic equality relation - on type @{ML_type term}; raw datatype equality should only be used - for operations related to parsing or printing! - - \<^descr> @{ML Term.map_types}~\f t\ applies the mapping \f\ to all types occurring in \t\. + \<^descr> \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.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_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.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 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 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 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 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 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.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.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. + \<^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 \ @@ -433,33 +401,31 @@ @@{ML_antiquotation prop} prop \} - \<^descr> \@{const_name c}\ inlines the internalized logical - constant name \c\ --- as @{ML_type string} literal. + \<^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_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> \@{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 + \<^descr> \@{term t}\ inlines the internalized term \t\ --- as constructor term for datatype @{ML_type term}. - \<^descr> \@{term t}\ inlines the internalized term \t\ - --- as constructor term for datatype @{ML_type term}. - - \<^descr> \@{prop \}\ inlines the internalized proposition - \\\ --- as constructor term for datatype @{ML_type term}. + \<^descr> \@{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 + 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 \\\. \ @@ -467,16 +433,14 @@ 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. + 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} @@ -523,26 +487,29 @@ \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 + 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"}). + \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.\ + 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}. + 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} @@ -560,40 +527,39 @@ \end{center} \end{figure} - Note that \instantiate\ does not require an explicit - side-condition, because \\\ may never contain schematic - variables. + 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. + 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 system always records oracle invocations within derivations of - theorems by a unique tag. + 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 system always records oracle invocations + within derivations of theorems by a unique tag. - 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. + 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\. + 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"}. + 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 \ @@ -635,101 +601,89 @@ Defs.entry -> Defs.entry list -> theory -> theory"} \\ \end{mldecls} - \<^descr> @{ML Thm.peek_status}~\thm\ informs about the current - status of the derivation object behind the given theorem. This is a - snapshot of a potentially ongoing (parallel) evaluation of proofs. - The three Boolean values indicate the following: \<^verbatim>\oracle\ - if the finished part contains some oracle invocation; \<^verbatim>\unfinished\ - if some future proofs are still pending; \<^verbatim>\failed\ if some future - proof has failed, rendering the theorem invalid! + \<^descr> @{ML Thm.peek_status}~\thm\ informs about the current status of the + derivation object behind the given theorem. This is a snapshot of a + potentially ongoing (parallel) evaluation of proofs. The three Boolean + values indicate the following: \<^verbatim>\oracle\ if the finished part contains some + oracle invocation; \<^verbatim>\unfinished\ if some future proofs are still pending; + \<^verbatim>\failed\ if some future proof has failed, rendering the theorem invalid! - \<^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.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> @{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> 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.ctyp_of}~\ctxt \\ and @{ML Thm.cterm_of}~\ctxt t\ explicitly + check types and terms, respectively. This also involves some basic + normalizations, such expansion of type and term abbreviations from the + underlying theory context. Full re-certification is relatively slow and + should be avoided in tight reasoning loops. - \<^descr> @{ML Thm.apply}, @{ML Thm.lambda}, @{ML Thm.all}, @{ML - Drule.mk_implies} etc.\ compose certified terms (or propositions) - incrementally. This is equivalent to @{ML Thm.cterm_of} after - unchecked @{ML_op "$"}, @{ML lambda}, @{ML Logic.all}, @{ML - Logic.mk_implies} etc., but there can be a big difference in - performance when large existing entities are composed by a few extra - constructions on top. There are separate operations to decompose + \<^descr> @{ML Thm.apply}, @{ML Thm.lambda}, @{ML Thm.all}, @{ML Drule.mk_implies} + etc.\ compose certified terms (or propositions) incrementally. This is + equivalent to @{ML Thm.cterm_of} after unchecked @{ML_op "$"}, @{ML lambda}, + @{ML Logic.all}, @{ML Logic.mk_implies} etc., but there can be a big + difference in performance when large existing entities are composed by a few + extra constructions on top. There are separate operations to decompose certified terms and theorems to produce certified terms again. - \<^descr> Type @{ML_type thm} represents proven propositions. This is - an abstract datatype that guarantees that its values have been - constructed by basic principles of the @{ML_structure Thm} module. - Every @{ML_type thm} value refers its background theory, - cf.\ \secref{sec:context-theory}. + \<^descr> 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.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.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.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.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_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_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 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 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. \ - text %mlantiq \ \begin{matharray}{rcl} @{ML_antiquotation_def "ctyp"} & : & \ML_antiquotation\ \\ @@ -755,46 +709,42 @@ @'by' method method? \} - \<^descr> \@{ctyp \}\ produces a certified type wrt.\ the - current background theory --- as abstract value of type @{ML_type - ctyp}. + \<^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> \@{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> \@{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> \@{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. + \<^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). + 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. + 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. \ 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. +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} @@ -812,37 +762,32 @@ \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 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 \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 \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. + 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 \ @@ -857,8 +802,7 @@ \<^descr> @{ML Conjunction.intr} derives \A &&& B\ from \A\ and \B\. - \<^descr> @{ML Conjunction.elim} derives \A\ and \B\ - from \A &&& B\. + \<^descr> @{ML Conjunction.elim} derives \A\ and \B\ from \A &&& B\. \<^descr> @{ML Drule.mk_term} derives \TERM t\. @@ -866,35 +810,35 @@ \<^descr> @{ML Logic.mk_type}~\\\ produces the term \TYPE(\)\. - \<^descr> @{ML Logic.dest_type}~\TYPE(\)\ recovers the 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 \\ \ \\. +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 \ \\. + 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\.\ + 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} @@ -902,17 +846,18 @@ @{index_ML Thm.strip_shyps: "thm -> thm"} \\ \end{mldecls} - \<^descr> @{ML Thm.extra_shyps}~\thm\ determines the extraneous - sort hypotheses of the given theorem, i.e.\ the sorts that are not - present within type variables of the statement. + \<^descr> @{ML Thm.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. + \<^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.\ +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" @@ -922,55 +867,54 @@ 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!\ +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. + 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 extensions91}. The most basic - rule format is that of a \<^emph>\Horn Clause\: + 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 extensions91}. 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: + 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: + 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: + 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} @@ -980,29 +924,26 @@ \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. + 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: + 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> 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. + \<^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 \ @@ -1010,43 +951,42 @@ @{index_ML Simplifier.norm_hhf: "Proof.context -> thm -> thm"} \\ \end{mldecls} - \<^descr> @{ML Simplifier.norm_hhf}~\ctxt thm\ normalizes the given - theorem according to the canonical form specified above. This is - occasionally helpful to repair some low-level tools that do not - handle Hereditary Harrop Formulae properly. + \<^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}. + 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 \\\): + 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.). + 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: + 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\} \] @@ -1065,10 +1005,10 @@ \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): + 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\\~~\text{(for some~\i\)}} @@ -1089,133 +1029,125 @@ @{index_ML_op "OF": "thm * thm list -> thm"} \\ \end{mldecls} - \<^descr> \rule\<^sub>1 RSN (i, rule\<^sub>2)\ resolves the conclusion of - \rule\<^sub>1\ with the \i\-th premise of \rule\<^sub>2\, - according to the @{inference resolution} principle explained above. - Unless there is precisely one resolvent it raises exception @{ML - THM}. + \<^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. + 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> \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 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> \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\<^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. + \<^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. + 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. +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. + 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. + \<^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. + 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>\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. + 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}). + \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 + 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>\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. + 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). +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"}. + 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.\ + 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: +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} @@ -1240,13 +1172,14 @@ \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). + 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. + There are separate read and print operations for proof terms, in order to + avoid conflicts with the regular term language. \ text %mlref \ @@ -1263,65 +1196,60 @@ @{index_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ \end{mldecls} - \<^descr> Type @{ML_type proof} represents proof terms; this is a - datatype with constructors @{index_ML Abst}, @{index_ML AbsP}, - @{index_ML_op "%"}, @{index_ML_op "%%"}, @{index_ML PBound}, - @{index_ML MinProof}, @{index_ML Hyp}, @{index_ML PAxm}, @{index_ML - Oracle}, @{index_ML Promise}, @{index_ML PThm} as explained above. - %FIXME OfClass (!?) + \<^descr> Type @{ML_type proof} represents proof terms; this is a datatype with + constructors @{index_ML Abst}, @{index_ML AbsP}, @{index_ML_op "%"}, + @{index_ML_op "%%"}, @{index_ML PBound}, @{index_ML MinProof}, @{index_ML + Hyp}, @{index_ML PAxm}, @{index_ML Oracle}, @{index_ML Promise}, @{index_ML + PThm} as explained above. %FIXME OfClass (!?) + + \<^descr> Type @{ML_type proof_body} represents the nested proof information of a + named theorem, consisting of a digest of oracles and named theorem over some + proof term. The digest only covers the directly visible part of the proof: + in order to get the full information, the implicit graph of nested theorems + needs to be traversed (e.g.\ using @{ML Proofterm.fold_body_thms}). - \<^descr> 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 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 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 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 Reconstruct.reconstruct_proof}~\thy prop prf\ turns the implicit + proof term \prf\ into a full proof of the given proposition. - \<^descr> @{ML Reconstruct.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. + 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 Reconstruct.expand_proof}~\thy [thm\<^sub>1, \, thm\<^sub>n] - prf\ expands and reconstructs the proofs of all specified theorems, - with the given (full) proof. Theorems that are not unique specified - via their name may be disambiguated by giving their proposition. + \<^descr> @{ML Reconstruct.expand_proof}~\thy [thm\<^sub>1, \, thm\<^sub>n] prf\ expands and + reconstructs the proofs of all specified theorems, with the given (full) + proof. Theorems that are not unique specified via their name may be + disambiguated by giving their proposition. - \<^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_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.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. + \<^descr> @{ML Proof_Syntax.pretty_proof}~\ctxt prf\ pretty-prints the given proof + term. \ -text %mlex \Detailed proof information of a theorem may be retrieved - as follows:\ +text %mlex \ + Detailed proof information of a theorem may be retrieved as follows: +\ lemma ex: "A \ B \ B \ A" proof @@ -1344,15 +1272,16 @@ (fn (name, _, _) => insert (op =) name) [body] []; \ -text \The result refers to various basic facts of Isabelle/HOL: - @{thm [source] HOL.impI}, @{thm [source] HOL.conjE}, @{thm [source] - HOL.conjI} etc. The combinator @{ML Proofterm.fold_body_thms} - recursively explores the graph of the proofs of all theorems being - used here. +text \ + The result refers to various basic facts of Isabelle/HOL: @{thm [source] + HOL.impI}, @{thm [source] HOL.conjE}, @{thm [source] HOL.conjI} etc. The + combinator @{ML Proofterm.fold_body_thms} recursively explores the graph of + the proofs of all theorems being used here. \<^medskip> - Alternatively, we may produce a proof term manually, and - turn it into a theorem as follows:\ + Alternatively, we may produce a proof term manually, and turn it into a + theorem as follows: +\ ML_val \ val thy = @{theory}; @@ -1371,9 +1300,8 @@ text \ \<^medskip> - See also @{file "~~/src/HOL/Proofs/ex/XML_Data.thy"} - for further examples, with export and import of proof terms via - XML/ML data representation. + See also @{file "~~/src/HOL/Proofs/ex/XML_Data.thy"} for further examples, + with export and import of proof terms via XML/ML data representation. \ end