# HG changeset patch # User wenzelm # Date 1455093136 -3600 # Node ID 374d1b6c473d62803e6b32df4de82045bfe713ec # Parent 199f4d6dab0a96387d5f3095a7f3b3d0eb7792de tuned whitespace; diff -r 199f4d6dab0a -r 374d1b6c473d src/Doc/Isar_Ref/Framework.thy --- a/src/Doc/Isar_Ref/Framework.thy Sun Feb 07 21:39:10 2016 +0100 +++ b/src/Doc/Isar_Ref/Framework.thy Wed Feb 10 09:32:16 2016 +0100 @@ -8,75 +8,67 @@ text \ Isabelle/Isar @{cite "Wenzel:1999:TPHOL" and "Wenzel-PhD" and - "Nipkow-TYPES02" and "Wenzel-Paulson:2006" and "Wenzel:2006:Festschrift"} - is intended as a generic framework for developing formal - mathematical documents with full proof checking. Definitions and - proofs are organized as theories. An assembly of theory sources may - be presented as a printed document; see also - \chref{ch:document-prep}. + "Nipkow-TYPES02" and "Wenzel-Paulson:2006" and "Wenzel:2006:Festschrift"} is + intended as a generic framework for developing formal mathematical documents + with full proof checking. Definitions and proofs are organized as theories. + An assembly of theory sources may be presented as a printed document; see + also \chref{ch:document-prep}. - The main objective of Isar is the design of a human-readable - structured proof language, which is called the ``primary proof - format'' in Isar terminology. Such a primary proof language is - somewhere in the middle between the extremes of primitive proof - objects and actual natural language. In this respect, Isar is a bit - more formalistic than Mizar @{cite "Trybulec:1993:MizarFeatures" and - "Rudnicki:1992:MizarOverview" and "Wiedijk:1999:Mizar"}, - using logical symbols for certain reasoning schemes where Mizar - would prefer English words; see @{cite "Wenzel-Wiedijk:2002"} for - further comparisons of these systems. + The main objective of Isar is the design of a human-readable structured + proof language, which is called the ``primary proof format'' in Isar + terminology. Such a primary proof language is somewhere in the middle + between the extremes of primitive proof objects and actual natural language. + In this respect, Isar is a bit more formalistic than Mizar @{cite + "Trybulec:1993:MizarFeatures" and "Rudnicki:1992:MizarOverview" and + "Wiedijk:1999:Mizar"}, using logical symbols for certain reasoning schemes + where Mizar would prefer English words; see @{cite "Wenzel-Wiedijk:2002"} + for further comparisons of these systems. - So Isar challenges the traditional way of recording informal proofs - in mathematical prose, as well as the common tendency to see fully - formal proofs directly as objects of some logical calculus (e.g.\ - \\\-terms in a version of type theory). In fact, Isar is - better understood as an interpreter of a simple block-structured - language for describing the data flow of local facts and goals, - interspersed with occasional invocations of proof methods. - Everything is reduced to logical inferences internally, but these - steps are somewhat marginal compared to the overall bookkeeping of - the interpretation process. Thanks to careful design of the syntax - and semantics of Isar language elements, a formal record of Isar - instructions may later appear as an intelligible text to the - attentive reader. + So Isar challenges the traditional way of recording informal proofs in + mathematical prose, as well as the common tendency to see fully formal + proofs directly as objects of some logical calculus (e.g.\ \\\-terms in a + version of type theory). In fact, Isar is better understood as an + interpreter of a simple block-structured language for describing the data + flow of local facts and goals, interspersed with occasional invocations of + proof methods. Everything is reduced to logical inferences internally, but + these steps are somewhat marginal compared to the overall bookkeeping of the + interpretation process. Thanks to careful design of the syntax and semantics + of Isar language elements, a formal record of Isar instructions may later + appear as an intelligible text to the attentive reader. - The Isar proof language has emerged from careful analysis of some - inherent virtues of the existing logical framework of Isabelle/Pure - @{cite "paulson-found" and "paulson700"}, notably composition of higher-order - natural deduction rules, which is a generalization of Gentzen's - original calculus @{cite "Gentzen:1935"}. The approach of generic - inference systems in Pure is continued by Isar towards actual proof - texts. + The Isar proof language has emerged from careful analysis of some inherent + virtues of the existing logical framework of Isabelle/Pure @{cite + "paulson-found" and "paulson700"}, notably composition of higher-order + natural deduction rules, which is a generalization of Gentzen's original + calculus @{cite "Gentzen:1935"}. The approach of generic inference systems + in Pure is continued by Isar towards actual proof texts. - Concrete applications require another intermediate layer: an - object-logic. Isabelle/HOL @{cite "isa-tutorial"} (simply-typed - set-theory) is being used most of the time; Isabelle/ZF - @{cite "isabelle-ZF"} is less extensively developed, although it would - probably fit better for classical mathematics. + Concrete applications require another intermediate layer: an object-logic. + Isabelle/HOL @{cite "isa-tutorial"} (simply-typed set-theory) is being used + most of the time; Isabelle/ZF @{cite "isabelle-ZF"} is less extensively + developed, although it would probably fit better for classical mathematics. \<^medskip> - In order to illustrate natural deduction in Isar, we shall - refer to the background theory and library of Isabelle/HOL. This - includes common notions of predicate logic, naive set-theory etc.\ - using fairly standard mathematical notation. From the perspective - of generic natural deduction there is nothing special about the - logical connectives of HOL (\\\, \\\, \\\, - \\\, etc.), only the resulting reasoning principles are - relevant to the user. There are similar rules available for - set-theory operators (\\\, \\\, \\\, \\\, etc.), or any other theory developed in the library (lattice + In order to illustrate natural deduction in Isar, we shall refer to the + background theory and library of Isabelle/HOL. This includes common notions + of predicate logic, naive set-theory etc.\ using fairly standard + mathematical notation. From the perspective of generic natural deduction + there is nothing special about the logical connectives of HOL (\\\, \\\, + \\\, \\\, etc.), only the resulting reasoning principles are relevant to the + user. There are similar rules available for set-theory operators (\\\, \\\, + \\\, \\\, etc.), or any other theory developed in the library (lattice theory, topology etc.). - Subsequently we briefly review fragments of Isar proof texts - corresponding directly to such general deduction schemes. The - examples shall refer to set-theory, to minimize the danger of - understanding connectives of predicate logic as something special. + Subsequently we briefly review fragments of Isar proof texts corresponding + directly to such general deduction schemes. The examples shall refer to + set-theory, to minimize the danger of understanding connectives of predicate + logic as something special. \<^medskip> - The following deduction performs \\\-introduction, - working forwards from assumptions towards the conclusion. We give - both the Isar text, and depict the primitive rule involved, as - determined by unification of the problem against rules that are - declared in the library context. + The following deduction performs \\\-introduction, working forwards from + assumptions towards the conclusion. We give both the Isar text, and depict + the primitive rule involved, as determined by unification of the problem + against rules that are declared in the library context. \ text_raw \\medskip\begin{minipage}{0.6\textwidth}\ @@ -101,13 +93,12 @@ text \ \<^medskip> - Note that @{command assume} augments the proof - context, @{command then} indicates that the current fact shall be - used in the next step, and @{command have} states an intermediate - goal. The two dots ``@{command ".."}'' refer to a complete proof of - this claim, using the indicated facts and a canonical rule from the - context. We could have been more explicit here by spelling out the - final proof step via the @{command "by"} command: + Note that @{command assume} augments the proof context, @{command then} + indicates that the current fact shall be used in the next step, and + @{command have} states an intermediate goal. The two dots ``@{command + ".."}'' refer to a complete proof of this claim, using the indicated facts + and a canonical rule from the context. We could have been more explicit here + by spelling out the final proof step via the @{command "by"} command: \ (*<*) @@ -121,14 +112,14 @@ (*>*) text \ - The format of the \\\-introduction rule represents - the most basic inference, which proceeds from given premises to a - conclusion, without any nested proof context involved. + The format of the \\\-introduction rule represents the most basic inference, + which proceeds from given premises to a conclusion, without any nested proof + context involved. - The next example performs backwards introduction on @{term "\\"}, - the intersection of all sets within a given set. This requires a - nested proof of set membership within a local context, where @{term - A} is an arbitrary-but-fixed member of the collection: + The next example performs backwards introduction on @{term "\\"}, the + intersection of all sets within a given set. This requires a nested proof of + set membership within a local context, where @{term A} is an + arbitrary-but-fixed member of the collection: \ text_raw \\medskip\begin{minipage}{0.6\textwidth}\ @@ -157,27 +148,24 @@ text \ \<^medskip> - This Isar reasoning pattern again refers to the - primitive rule depicted above. The system determines it in the - ``@{command proof}'' step, which could have been spelled out more - explicitly as ``@{command proof}~\(rule InterI)\''. Note - that the rule involves both a local parameter @{term "A"} and an - assumption @{prop "A \ \"} in the nested reasoning. This kind of + This Isar reasoning pattern again refers to the primitive rule depicted + above. The system determines it in the ``@{command proof}'' step, which + could have been spelled out more explicitly as ``@{command proof}~\(rule + InterI)\''. Note that the rule involves both a local parameter @{term "A"} + and an assumption @{prop "A \ \"} in the nested reasoning. This kind of compound rule typically demands a genuine sub-proof in Isar, working - backwards rather than forwards as seen before. In the proof body we - encounter the @{command fix}-@{command assume}-@{command show} - outline of nested sub-proofs that is typical for Isar. The final - @{command show} is like @{command have} followed by an additional - refinement of the enclosing claim, using the rule derived from the - proof body. + backwards rather than forwards as seen before. In the proof body we + encounter the @{command fix}-@{command assume}-@{command show} outline of + nested sub-proofs that is typical for Isar. The final @{command show} is + like @{command have} followed by an additional refinement of the enclosing + claim, using the rule derived from the proof body. \<^medskip> - The next example involves @{term "\\"}, which can be - characterized as the set of all @{term "x"} such that @{prop "\A. x - \ A \ A \ \"}. The elimination rule for @{prop "x \ \\"} does - not mention \\\ and \\\ at all, but admits to obtain - directly a local @{term "A"} such that @{prop "x \ A"} and @{prop "A - \ \"} hold. This corresponds to the following Isar proof and + The next example involves @{term "\\"}, which can be characterized as the + set of all @{term "x"} such that @{prop "\A. x \ A \ A \ \"}. The + elimination rule for @{prop "x \ \\"} does not mention \\\ and \\\ at all, + but admits to obtain directly a local @{term "A"} such that @{prop "x \ A"} + and @{prop "A \ \"} hold. This corresponds to the following Isar proof and inference rule, respectively: \ @@ -208,14 +196,13 @@ text \ \<^medskip> - Although the Isar proof follows the natural - deduction rule closely, the text reads not as natural as - anticipated. There is a double occurrence of an arbitrary - conclusion @{prop "C"}, which represents the final result, but is - irrelevant for now. This issue arises for any elimination rule - involving local parameters. Isar provides the derived language - element @{command obtain}, which is able to perform the same - elimination proof more conveniently: + Although the Isar proof follows the natural deduction rule closely, the text + reads not as natural as anticipated. There is a double occurrence of an + arbitrary conclusion @{prop "C"}, which represents the final result, but is + irrelevant for now. This issue arises for any elimination rule involving + local parameters. Isar provides the derived language element @{command + obtain}, which is able to perform the same elimination proof more + conveniently: \ (*<*) @@ -229,9 +216,9 @@ (*>*) text \ - Here we avoid to mention the final conclusion @{prop "C"} - and return to plain forward reasoning. The rule involved in the - ``@{command ".."}'' proof is the same as before. + Here we avoid to mention the final conclusion @{prop "C"} and return to + plain forward reasoning. The rule involved in the ``@{command ".."}'' proof + is the same as before. \ @@ -239,9 +226,9 @@ text \ The Pure logic @{cite "paulson-found" and "paulson700"} is an intuitionistic - fragment of higher-order logic @{cite "church40"}. In type-theoretic - parlance, there are three levels of \\\-calculus with - corresponding arrows \\\/\\\/\\\: + fragment of higher-order logic @{cite "church40"}. In type-theoretic + parlance, there are three levels of \\\-calculus with corresponding arrows + \\\/\\\/\\\: \<^medskip> \begin{tabular}{ll} @@ -251,36 +238,31 @@ \end{tabular} \<^medskip> - Here only the types of syntactic terms, and the - propositions of proof terms have been shown. The \\\-structure of proofs can be recorded as an optional feature of - the Pure inference kernel @{cite "Berghofer-Nipkow:2000:TPHOL"}, but - the formal system can never depend on them due to \<^emph>\proof - irrelevance\. + Here only the types of syntactic terms, and the propositions of proof terms + have been shown. The \\\-structure of proofs can be recorded as an optional + feature of the Pure inference kernel @{cite "Berghofer-Nipkow:2000:TPHOL"}, + but the formal system can never depend on them due to \<^emph>\proof irrelevance\. - On top of this most primitive layer of proofs, Pure implements a - generic calculus for nested natural deduction rules, similar to - @{cite "Schroeder-Heister:1984"}. Here object-logic inferences are - internalized as formulae over \\\ and \\\. - Combining such rule statements may involve higher-order unification - @{cite "paulson-natural"}. + On top of this most primitive layer of proofs, Pure implements a generic + calculus for nested natural deduction rules, similar to @{cite + "Schroeder-Heister:1984"}. Here object-logic inferences are internalized as + formulae over \\\ and \\\. Combining such rule statements may involve + higher-order unification @{cite "paulson-natural"}. \ subsection \Primitive inferences\ text \ - Term syntax provides explicit notation for abstraction \\x :: - \. b(x)\ and application \b a\, while types are usually - implicit thanks to type-inference; terms of type \prop\ are - called propositions. Logical statements are composed via \\x - :: \. B(x)\ and \A \ B\. Primitive reasoning operates on - judgments of the form \\ \ \\, with standard introduction - and elimination rules for \\\ and \\\ that refer to - fixed parameters \x\<^sub>1, \, x\<^sub>m\ and hypotheses - \A\<^sub>1, \, A\<^sub>n\ from the context \\\; - the corresponding proof terms are left implicit. The subsequent - inference rules define \\ \ \\ inductively, relative to a - collection of axioms: + Term syntax provides explicit notation for abstraction \\x :: \. b(x)\ and + application \b a\, while types are usually implicit thanks to + type-inference; terms of type \prop\ are called propositions. Logical + statements are composed via \\x :: \. B(x)\ and \A \ B\. Primitive reasoning + operates on judgments of the form \\ \ \\, with standard introduction and + elimination rules for \\\ and \\\ that refer to fixed parameters \x\<^sub>1, \, + x\<^sub>m\ and hypotheses \A\<^sub>1, \, A\<^sub>n\ from the context \\\; the corresponding + proof terms are left implicit. The subsequent inference rules define \\ \ \\ + inductively, relative to a collection of axioms: \[ \infer{\\ A\}{(\A\ \mbox{~axiom})} @@ -300,68 +282,60 @@ \infer{\\\<^sub>1 \ \\<^sub>2 \ B\}{\\\<^sub>1 \ A \ B\ & \\\<^sub>2 \ A\} \] - Furthermore, Pure provides a built-in equality \\ :: \ \ \ \ - prop\ with axioms for reflexivity, substitution, extensionality, - and \\\\\-conversion on \\\-terms. + Furthermore, Pure provides a built-in equality \\ :: \ \ \ \ prop\ with + axioms for reflexivity, substitution, extensionality, and \\\\\-conversion + on \\\-terms. \<^medskip> - An object-logic introduces another layer on top of Pure, - e.g.\ with types \i\ for individuals and \o\ for - propositions, term constants \Trueprop :: o \ prop\ as - (implicit) derivability judgment and connectives like \\ :: o - \ o \ o\ or \\ :: (i \ o) \ o\, and axioms for object-level - rules such as \conjI: A \ B \ A \ B\ or \allI: (\x. B - x) \ \x. B x\. Derived object rules are represented as theorems of - Pure. After the initial object-logic setup, further axiomatizations - are usually avoided; plain definitions and derived principles are - used exclusively. + An object-logic introduces another layer on top of Pure, e.g.\ with types + \i\ for individuals and \o\ for propositions, term constants \Trueprop :: o + \ prop\ as (implicit) derivability judgment and connectives like \\ :: o \ o + \ o\ or \\ :: (i \ o) \ o\, and axioms for object-level rules such as + \conjI: A \ B \ A \ B\ or \allI: (\x. B x) \ \x. B x\. Derived object rules + are represented as theorems of Pure. After the initial object-logic setup, + further axiomatizations are usually avoided; plain definitions and derived + principles are used exclusively. \ subsection \Reasoning with rules \label{sec:framework-resolution}\ text \ - Primitive inferences mostly serve foundational purposes. The main - reasoning mechanisms of Pure operate on nested natural deduction - rules expressed as formulae, using \\\ to bind local - parameters and \\\ to express entailment. Multiple - parameters and premises are represented by repeating these + Primitive inferences mostly serve foundational purposes. The main reasoning + mechanisms of Pure operate on nested natural deduction rules expressed as + formulae, using \\\ to bind local parameters and \\\ to express entailment. + Multiple parameters and premises are represented by repeating these connectives in a right-associative manner. - Since \\\ and \\\ commute thanks to the theorem - @{prop "(A \ (\x. B x)) \ (\x. A \ B x)"}, we may assume w.l.o.g.\ - that rule statements always observe the normal form where - quantifiers are pulled in front of implications at each level of - nesting. This means that any Pure proposition may be presented as a - \<^emph>\Hereditary Harrop Formula\ @{cite "Miller:1991"} which is of the - form \\x\<^sub>1 \ x\<^sub>m. H\<^sub>1 \ \ H\<^sub>n \ - A\ for \m, n \ 0\, and \A\ atomic, and \H\<^sub>1, \, H\<^sub>n\ being recursively of the same format. - Following the convention that outermost quantifiers are implicit, - Horn clauses \A\<^sub>1 \ \ A\<^sub>n \ A\ are a special - case of this. + Since \\\ and \\\ commute thanks to the theorem @{prop "(A \ (\x. B x)) \ + (\x. A \ B x)"}, we may assume w.l.o.g.\ that rule statements always observe + the normal form where quantifiers are pulled in front of implications at + each level of nesting. This means that any Pure proposition may be presented + as a \<^emph>\Hereditary Harrop Formula\ @{cite "Miller:1991"} which is of the form + \\x\<^sub>1 \ x\<^sub>m. H\<^sub>1 \ \ H\<^sub>n \ A\ for \m, n \ 0\, and \A\ atomic, and \H\<^sub>1, \, + H\<^sub>n\ being recursively of the same format. Following the convention that + outermost quantifiers are implicit, Horn clauses \A\<^sub>1 \ \ A\<^sub>n \ A\ are a + special case of this. - For example, \\\-introduction rule encountered before is - represented as a Pure theorem as follows: + For example, \\\-introduction rule encountered before is represented as a + Pure theorem as follows: \[ \IntI:\~@{prop "x \ A \ x \ B \ x \ A \ B"} \] - This is a plain Horn clause, since no further nesting on - the left is involved. The general \\\-introduction - corresponds to a Hereditary Harrop Formula with one additional level - of nesting: + This is a plain Horn clause, since no further nesting on the left is + involved. The general \\\-introduction corresponds to a Hereditary Harrop + Formula with one additional level of nesting: \[ \InterI:\~@{prop "(\A. A \ \ \ x \ A) \ x \ \\"} \] \<^medskip> - Goals are also represented as rules: \A\<^sub>1 \ - \ A\<^sub>n \ C\ states that the sub-goals \A\<^sub>1, \, - A\<^sub>n\ entail the result \C\; for \n = 0\ the - goal is finished. To allow \C\ being a rule statement - itself, we introduce the protective marker \# :: prop \ - prop\, which is defined as identity and hidden from the user. We - initialize and finish goal states as follows: + Goals are also represented as rules: \A\<^sub>1 \ \ A\<^sub>n \ C\ states that the + sub-goals \A\<^sub>1, \, A\<^sub>n\ entail the result \C\; for \n = 0\ the goal is + finished. To allow \C\ being a rule statement itself, we introduce the + protective marker \# :: prop \ prop\, which is defined as identity and + hidden from the user. We initialize and finish goal states as follows: \[ \begin{array}{c@ {\qquad}c} @@ -370,12 +344,12 @@ \end{array} \] - Goal states are refined in intermediate proof steps until - a finished form is achieved. Here the two main reasoning principles - are @{inference resolution}, for back-chaining a rule against a - sub-goal (replacing it by zero or more sub-goals), and @{inference - assumption}, for solving a sub-goal (finding a short-circuit with - local assumptions). Below \\<^vec>x\ stands for \x\<^sub>1, \, x\<^sub>n\ (\n \ 0\). + Goal states are refined in intermediate proof steps until a finished form is + achieved. Here the two main reasoning principles are @{inference + resolution}, for back-chaining a rule against a sub-goal (replacing it by + zero or more sub-goals), and @{inference assumption}, for solving a sub-goal + (finding a short-circuit with local assumptions). Below \\<^vec>x\ stands + for \x\<^sub>1, \, x\<^sub>n\ (\n \ 0\). \[ \infer[(@{inference_def resolution})] @@ -418,13 +392,12 @@ \<^medskip> } - Compositions of @{inference assumption} after @{inference - resolution} occurs quite often, typically in elimination steps. - Traditional Isabelle tactics accommodate this by a combined - @{inference_def elim_resolution} principle. In contrast, Isar uses - a slightly more refined combination, where the assumptions to be - closed are marked explicitly, using again the protective marker - \#\: + Compositions of @{inference assumption} after @{inference resolution} occurs + quite often, typically in elimination steps. Traditional Isabelle tactics + accommodate this by a combined @{inference_def elim_resolution} principle. + In contrast, Isar uses a slightly more refined combination, where the + assumptions to be closed are marked explicitly, using again the protective + marker \#\: \[ \infer[(@{inference refinement})] @@ -442,32 +415,30 @@ \end{tabular}} \] - Here the \sub\proof\ rule stems from the - main @{command fix}-@{command assume}-@{command show} outline of - Isar (cf.\ \secref{sec:framework-subproof}): each assumption - indicated in the text results in a marked premise \G\ above. - The marking enforces resolution against one of the sub-goal's - premises. Consequently, @{command fix}-@{command assume}-@{command - show} enables to fit the result of a sub-proof quite robustly into a - pending sub-goal, while maintaining a good measure of flexibility. + Here the \sub\proof\ rule stems from the main @{command fix}-@{command + assume}-@{command show} outline of Isar (cf.\ + \secref{sec:framework-subproof}): each assumption indicated in the text + results in a marked premise \G\ above. The marking enforces resolution + against one of the sub-goal's premises. Consequently, @{command + fix}-@{command assume}-@{command show} enables to fit the result of a + sub-proof quite robustly into a pending sub-goal, while maintaining a good + measure of flexibility. \ section \The Isar proof language \label{sec:framework-isar}\ text \ - Structured proofs are presented as high-level expressions for - composing entities of Pure (propositions, facts, and goals). The - Isar proof language allows to organize reasoning within the - underlying rule calculus of Pure, but Isar is not another logical - calculus! + Structured proofs are presented as high-level expressions for composing + entities of Pure (propositions, facts, and goals). The Isar proof language + allows to organize reasoning within the underlying rule calculus of Pure, + but Isar is not another logical calculus! - Isar is an exercise in sound minimalism. Approximately half of the - language is introduced as primitive, the rest defined as derived - concepts. The following grammar describes the core language - (category \proof\), which is embedded into theory - specification elements such as @{command theorem}; see also - \secref{sec:framework-stmt} for the separate category \statement\. + Isar is an exercise in sound minimalism. Approximately half of the language + is introduced as primitive, the rest defined as derived concepts. The + following grammar describes the core language (category \proof\), which is + embedded into theory specification elements such as @{command theorem}; see + also \secref{sec:framework-stmt} for the separate category \statement\. \<^medskip> \begin{tabular}{rcl} @@ -490,31 +461,28 @@ \end{tabular} \<^medskip> - Simultaneous propositions or facts may be separated by the - @{keyword "and"} keyword. + Simultaneous propositions or facts may be separated by the @{keyword "and"} + keyword. \<^medskip> - The syntax for terms and propositions is inherited from - Pure (and the object-logic). A \pattern\ is a \term\ with schematic variables, to be bound by higher-order - matching. + The syntax for terms and propositions is inherited from Pure (and the + object-logic). A \pattern\ is a \term\ with schematic variables, to be bound + by higher-order matching. \<^medskip> - Facts may be referenced by name or proposition. For - example, the result of ``@{command have}~\a: A \proof\\'' - becomes available both as \a\ and - \isacharbackquoteopen\A\\isacharbackquoteclose. Moreover, - fact expressions may involve attributes that modify either the - theorem or the background context. For example, the expression - ``\a [OF b]\'' refers to the composition of two facts - according to the @{inference resolution} inference of - \secref{sec:framework-resolution}, while ``\a [intro]\'' - declares a fact as introduction rule in the context. + Facts may be referenced by name or proposition. For example, the result of + ``@{command have}~\a: A \proof\\'' becomes available both as \a\ and + \isacharbackquoteopen\A\\isacharbackquoteclose. Moreover, fact expressions + may involve attributes that modify either the theorem or the background + context. For example, the expression ``\a [OF b]\'' refers to the + composition of two facts according to the @{inference resolution} inference + of \secref{sec:framework-resolution}, while ``\a [intro]\'' declares a fact + as introduction rule in the context. - The special fact called ``@{fact this}'' always refers to the last - result, as produced by @{command note}, @{command assume}, @{command - have}, or @{command show}. Since @{command note} occurs - frequently together with @{command then} we provide some - abbreviations: + The special fact called ``@{fact this}'' always refers to the last result, + as produced by @{command note}, @{command assume}, @{command have}, or + @{command show}. Since @{command note} occurs frequently together with + @{command then} we provide some abbreviations: \<^medskip> \begin{tabular}{rcl} @@ -523,67 +491,63 @@ \end{tabular} \<^medskip> - The \method\ category is essentially a parameter and may be - populated later. Methods use the facts indicated by @{command - "then"} or @{command using}, and then operate on the goal state. - Some basic methods are predefined: ``@{method "-"}'' leaves the goal - unchanged, ``@{method this}'' applies the facts as rules to the - goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the - result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}'' - refer to @{inference resolution} of - \secref{sec:framework-resolution}). The secondary arguments to - ``@{method (Pure) rule}'' may be specified explicitly as in ``\(rule - a)\'', or picked from the context. In the latter case, the system - first tries rules declared as @{attribute (Pure) elim} or - @{attribute (Pure) dest}, followed by those declared as @{attribute - (Pure) intro}. + The \method\ category is essentially a parameter and may be populated later. + Methods use the facts indicated by @{command "then"} or @{command using}, + and then operate on the goal state. Some basic methods are predefined: + ``@{method "-"}'' leaves the goal unchanged, ``@{method this}'' applies the + facts as rules to the goal, ``@{method (Pure) "rule"}'' applies the facts to + another rule and the result to the goal (both ``@{method this}'' and + ``@{method (Pure) rule}'' refer to @{inference resolution} of + \secref{sec:framework-resolution}). The secondary arguments to ``@{method + (Pure) rule}'' may be specified explicitly as in ``\(rule a)\'', or picked + from the context. In the latter case, the system first tries rules declared + as @{attribute (Pure) elim} or @{attribute (Pure) dest}, followed by those + declared as @{attribute (Pure) intro}. - The default method for @{command proof} is ``@{method standard}'' - (arguments picked from the context), for @{command qed} it is - ``@{method "succeed"}''. Further abbreviations for terminal proof steps - are ``@{command "by"}~\method\<^sub>1 method\<^sub>2\'' for - ``@{command proof}~\method\<^sub>1\~@{command qed}~\method\<^sub>2\'', and ``@{command ".."}'' for ``@{command - "by"}~@{method standard}, and ``@{command "."}'' for ``@{command - "by"}~@{method this}''. The @{command unfolding} element operates - directly on the current facts and goal by applying equalities. + The default method for @{command proof} is ``@{method standard}'' (arguments + picked from the context), for @{command qed} it is ``@{method "succeed"}''. + Further abbreviations for terminal proof steps are ``@{command + "by"}~\method\<^sub>1 method\<^sub>2\'' for ``@{command proof}~\method\<^sub>1\~@{command + qed}~\method\<^sub>2\'', and ``@{command ".."}'' for ``@{command "by"}~@{method + standard}, and ``@{command "."}'' for ``@{command "by"}~@{method this}''. + The @{command unfolding} element operates directly on the current facts and + goal by applying equalities. \<^medskip> Block structure can be indicated explicitly by ``@{command - "{"}~\\\~@{command "}"}'', although the body of a sub-proof - already involves implicit nesting. In any case, @{command next} - jumps into the next section of a block, i.e.\ it acts like closing - an implicit block scope and opening another one; there is no direct - correspondence to subgoals here. + "{"}~\\\~@{command "}"}'', although the body of a sub-proof already involves + implicit nesting. In any case, @{command next} jumps into the next section + of a block, i.e.\ it acts like closing an implicit block scope and opening + another one; there is no direct correspondence to subgoals here. - The remaining elements @{command fix} and @{command assume} build up - a local context (see \secref{sec:framework-context}), while - @{command show} refines a pending sub-goal by the rule resulting - from a nested sub-proof (see \secref{sec:framework-subproof}). - Further derived concepts will support calculational reasoning (see - \secref{sec:framework-calc}). + The remaining elements @{command fix} and @{command assume} build up a local + context (see \secref{sec:framework-context}), while @{command show} refines + a pending sub-goal by the rule resulting from a nested sub-proof (see + \secref{sec:framework-subproof}). Further derived concepts will support + calculational reasoning (see \secref{sec:framework-calc}). \ subsection \Context elements \label{sec:framework-context}\ text \ - In judgments \\ \ \\ of the primitive framework, \\\ - essentially acts like a proof context. Isar elaborates this idea - towards a higher-level notion, with additional information for - type-inference, term abbreviations, local facts, hypotheses etc. + In judgments \\ \ \\ of the primitive framework, \\\ essentially acts like a + proof context. Isar elaborates this idea towards a higher-level notion, with + additional information for type-inference, term abbreviations, local facts, + hypotheses etc. - The element @{command fix}~\x :: \\ declares a local - parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in - results exported from the context, \x\ may become anything. - The @{command assume}~\\inference\\ element provides a - general interface to hypotheses: ``@{command assume}~\\inference\ A\'' produces \A \ A\ locally, while the - included inference tells how to discharge \A\ from results - \A \ B\ later on. There is no user-syntax for \\inference\\, i.e.\ it may only occur internally when derived - commands are defined in ML. + The element @{command fix}~\x :: \\ declares a local parameter, i.e.\ an + arbitrary-but-fixed entity of a given type; in results exported from the + context, \x\ may become anything. The @{command assume}~\\inference\\ + element provides a general interface to hypotheses: ``@{command + assume}~\\inference\ A\'' produces \A \ A\ locally, while the included + inference tells how to discharge \A\ from results \A \ B\ later on. There is + no user-syntax for \\inference\\, i.e.\ it may only occur internally when + derived commands are defined in ML. At the user-level, the default inference for @{command assume} is - @{inference discharge} as given below. The additional variants - @{command presume} and @{command def} are defined as follows: + @{inference discharge} as given below. The additional variants @{command + presume} and @{command def} are defined as follows: \<^medskip> \begin{tabular}{rcl} @@ -603,19 +567,18 @@ \] \<^medskip> - Note that @{inference discharge} and @{inference - "weak\discharge"} differ in the marker for @{prop A}, which is - relevant when the result of a @{command fix}-@{command - assume}-@{command show} outline is composed with a pending goal, - cf.\ \secref{sec:framework-subproof}. + Note that @{inference discharge} and @{inference "weak\discharge"} differ in + the marker for @{prop A}, which is relevant when the result of a @{command + fix}-@{command assume}-@{command show} outline is composed with a pending + goal, cf.\ \secref{sec:framework-subproof}. - The most interesting derived context element in Isar is @{command - obtain} @{cite \\S5.3\ "Wenzel-PhD"}, which supports generalized - elimination steps in a purely forward manner. The @{command obtain} - command takes a specification of parameters \\<^vec>x\ and - assumptions \\<^vec>A\ to be added to the context, together - with a proof of a case rule stating that this extension is - conservative (i.e.\ may be removed from closed results later on): + The most interesting derived context element in Isar is @{command obtain} + @{cite \\S5.3\ "Wenzel-PhD"}, which supports generalized elimination steps + in a purely forward manner. The @{command obtain} command takes a + specification of parameters \\<^vec>x\ and assumptions \\<^vec>A\ to be + added to the context, together with a proof of a case rule stating that this + extension is conservative (i.e.\ may be removed from closed results later + on): \<^medskip> \begin{tabular}{l} @@ -640,21 +603,19 @@ \end{tabular}} \] - Here the name ``\thesis\'' is a specific convention - for an arbitrary-but-fixed proposition; in the primitive natural - deduction rules shown before we have occasionally used \C\. - The whole statement of ``@{command obtain}~\x\~@{keyword - "where"}~\A x\'' may be read as a claim that \A x\ - may be assumed for some arbitrary-but-fixed \x\. Also note - that ``@{command obtain}~\A \ B\'' without parameters - is similar to ``@{command have}~\A \ B\'', but the - latter involves multiple sub-goals. + Here the name ``\thesis\'' is a specific convention for an + arbitrary-but-fixed proposition; in the primitive natural deduction rules + shown before we have occasionally used \C\. The whole statement of + ``@{command obtain}~\x\~@{keyword "where"}~\A x\'' may be read as a claim + that \A x\ may be assumed for some arbitrary-but-fixed \x\. Also note that + ``@{command obtain}~\A \ B\'' without parameters is similar to + ``@{command have}~\A \ B\'', but the latter involves multiple + sub-goals. \<^medskip> - The subsequent Isar proof texts explain all context - elements introduced above using the formal proof language itself. - After finishing a local proof within a block, we indicate the - exported result via @{command note}. + The subsequent Isar proof texts explain all context elements introduced + above using the formal proof language itself. After finishing a local proof + within a block, we indicate the exported result via @{command note}. \ (*<*) @@ -692,8 +653,8 @@ text \ \<^bigskip> - This illustrates the meaning of Isar context - elements without goals getting in between. + This illustrates the meaning of Isar context elements without goals getting + in between. \ @@ -717,18 +678,17 @@ \end{tabular} \<^medskip> - A simple \statement\ consists of named - propositions. The full form admits local context elements followed - by the actual conclusions, such as ``@{keyword "fixes"}~\x\~@{keyword "assumes"}~\A x\~@{keyword "shows"}~\B - x\''. The final result emerges as a Pure rule after discharging - the context: @{prop "\x. A x \ B x"}. + A simple \statement\ consists of named propositions. The full form admits + local context elements followed by the actual conclusions, such as + ``@{keyword "fixes"}~\x\~@{keyword "assumes"}~\A x\~@{keyword "shows"}~\B + x\''. The final result emerges as a Pure rule after discharging the context: + @{prop "\x. A x \ B x"}. - The @{keyword "obtains"} variant is another abbreviation defined - below; unlike @{command obtain} (cf.\ - \secref{sec:framework-context}) there may be several ``cases'' - separated by ``\\\'', each consisting of several - parameters (\vars\) and several premises (\props\). - This specifies multi-branch elimination rules. + The @{keyword "obtains"} variant is another abbreviation defined below; + unlike @{command obtain} (cf.\ \secref{sec:framework-context}) there may be + several ``cases'' separated by ``\\\'', each consisting of several + parameters (\vars\) and several premises (\props\). This specifies + multi-branch elimination rules. \<^medskip> \begin{tabular}{l} @@ -740,10 +700,9 @@ \<^medskip> Presenting structured statements in such an ``open'' format usually - simplifies the subsequent proof, because the outer structure of the - problem is already laid out directly. E.g.\ consider the following - canonical patterns for \\\ and \\\, - respectively: + simplifies the subsequent proof, because the outer structure of the problem + is already laid out directly. E.g.\ consider the following canonical + patterns for \\\ and \\\, respectively: \ text_raw \\begin{minipage}{0.5\textwidth}\ @@ -771,39 +730,36 @@ text \ \<^medskip> - Here local facts \isacharbackquoteopen\A - x\\isacharbackquoteclose\ and \isacharbackquoteopen\B - y\\isacharbackquoteclose\ are referenced immediately; there is no - need to decompose the logical rule structure again. In the second - proof the final ``@{command then}~@{command show}~\thesis\~@{command ".."}'' involves the local rule case \\x - y. A x \ B y \ thesis\ for the particular instance of terms \a\ and \b\ produced in the body. -\ + Here local facts \isacharbackquoteopen\A x\\isacharbackquoteclose\ and + \isacharbackquoteopen\B y\\isacharbackquoteclose\ are referenced + immediately; there is no need to decompose the logical rule structure again. + In the second proof the final ``@{command then}~@{command + show}~\thesis\~@{command ".."}'' involves the local rule case \\x y. A x \ B + y \ thesis\ for the particular instance of terms \a\ and \b\ produced in the + body. \ subsection \Structured proof refinement \label{sec:framework-subproof}\ text \ - By breaking up the grammar for the Isar proof language, we may - understand a proof text as a linear sequence of individual proof - commands. These are interpreted as transitions of the Isar virtual - machine (Isar/VM), which operates on a block-structured - configuration in single steps. This allows users to write proof - texts in an incremental manner, and inspect intermediate - configurations for debugging. + By breaking up the grammar for the Isar proof language, we may understand a + proof text as a linear sequence of individual proof commands. These are + interpreted as transitions of the Isar virtual machine (Isar/VM), which + operates on a block-structured configuration in single steps. This allows + users to write proof texts in an incremental manner, and inspect + intermediate configurations for debugging. - The basic idea is analogous to evaluating algebraic expressions on a - stack machine: \(a + b) \ c\ then corresponds to a sequence - of single transitions for each symbol \(, a, +, b, ), \, c\. - In Isar the algebraic values are facts or goals, and the operations - are inferences. + The basic idea is analogous to evaluating algebraic expressions on a stack + machine: \(a + b) \ c\ then corresponds to a sequence of single transitions + for each symbol \(, a, +, b, ), \, c\. In Isar the algebraic values are + facts or goals, and the operations are inferences. \<^medskip> - The Isar/VM state maintains a stack of nodes, each node - contains the local proof context, the linguistic mode, and a pending - goal (optional). The mode determines the type of transition that - may be performed next, it essentially alternates between forward and - backward reasoning, with an intermediate stage for chained facts - (see \figref{fig:isar-vm}). + The Isar/VM state maintains a stack of nodes, each node contains the local + proof context, the linguistic mode, and a pending goal (optional). The mode + determines the type of transition that may be performed next, it essentially + alternates between forward and backward reasoning, with an intermediate + stage for chained facts (see \figref{fig:isar-vm}). \begin{figure}[htb] \begin{center} @@ -812,17 +768,15 @@ \caption{Isar/VM modes}\label{fig:isar-vm} \end{figure} - For example, in \state\ mode Isar acts like a mathematical - scratch-pad, accepting declarations like @{command fix}, @{command - assume}, and claims like @{command have}, @{command show}. A goal - statement changes the mode to \prove\, which means that we - may now refine the problem via @{command unfolding} or @{command - proof}. Then we are again in \state\ mode of a proof body, - which may issue @{command show} statements to solve pending - sub-goals. A concluding @{command qed} will return to the original - \state\ mode one level upwards. The subsequent Isar/VM - trace indicates block structure, linguistic mode, goal state, and - inferences: + For example, in \state\ mode Isar acts like a mathematical scratch-pad, + accepting declarations like @{command fix}, @{command assume}, and claims + like @{command have}, @{command show}. A goal statement changes the mode to + \prove\, which means that we may now refine the problem via @{command + unfolding} or @{command proof}. Then we are again in \state\ mode of a proof + body, which may issue @{command show} statements to solve pending sub-goals. + A concluding @{command qed} will return to the original \state\ mode one + level upwards. The subsequent Isar/VM trace indicates block structure, + linguistic mode, goal state, and inferences: \ text_raw \\begingroup\footnotesize\ @@ -873,13 +827,12 @@ text \ Here the @{inference refinement} inference from - \secref{sec:framework-resolution} mediates composition of Isar - sub-proofs nicely. Observe that this principle incorporates some - degree of freedom in proof composition. In particular, the proof - body allows parameters and assumptions to be re-ordered, or commuted - according to Hereditary Harrop Form. Moreover, context elements - that are not used in a sub-proof may be omitted altogether. For - example: + \secref{sec:framework-resolution} mediates composition of Isar sub-proofs + nicely. Observe that this principle incorporates some degree of freedom in + proof composition. In particular, the proof body allows parameters and + assumptions to be re-ordered, or commuted according to Hereditary Harrop + Form. Moreover, context elements that are not used in a sub-proof may be + omitted altogether. For example: \ text_raw \\begin{minipage}{0.5\textwidth}\ @@ -937,17 +890,16 @@ text \ \<^medskip> - Such ``peephole optimizations'' of Isar texts are - practically important to improve readability, by rearranging - contexts elements according to the natural flow of reasoning in the - body, while still observing the overall scoping rules. + Such ``peephole optimizations'' of Isar texts are practically important to + improve readability, by rearranging contexts elements according to the + natural flow of reasoning in the body, while still observing the overall + scoping rules. \<^medskip> - This illustrates the basic idea of structured proof - processing in Isar. The main mechanisms are based on natural - deduction rule composition within the Pure framework. In - particular, there are no direct operations on goal states within the - proof body. Moreover, there is no hidden automated reasoning + This illustrates the basic idea of structured proof processing in Isar. The + main mechanisms are based on natural deduction rule composition within the + Pure framework. In particular, there are no direct operations on goal states + within the proof body. Moreover, there is no hidden automated reasoning involved, just plain unification. \ @@ -956,28 +908,26 @@ text \ The existing Isar infrastructure is sufficiently flexible to support - calculational reasoning (chains of transitivity steps) as derived - concept. The generic proof elements introduced below depend on - rules declared as @{attribute trans} in the context. It is left to - the object-logic to provide a suitable rule collection for mixed - relations of \=\, \<\, \\\, \\\, - \\\ etc. Due to the flexibility of rule composition - (\secref{sec:framework-resolution}), substitution of equals by - equals is covered as well, even substitution of inequalities - involving monotonicity conditions; see also @{cite \\S6\ "Wenzel-PhD"} - and @{cite "Bauer-Wenzel:2001"}. + calculational reasoning (chains of transitivity steps) as derived concept. + The generic proof elements introduced below depend on rules declared as + @{attribute trans} in the context. It is left to the object-logic to provide + a suitable rule collection for mixed relations of \=\, \<\, \\\, \\\, \\\ + etc. Due to the flexibility of rule composition + (\secref{sec:framework-resolution}), substitution of equals by equals is + covered as well, even substitution of inequalities involving monotonicity + conditions; see also @{cite \\S6\ "Wenzel-PhD"} and @{cite + "Bauer-Wenzel:2001"}. - The generic calculational mechanism is based on the observation that - rules such as \trans:\~@{prop "x = y \ y = z \ x = z"} - proceed from the premises towards the conclusion in a deterministic - fashion. Thus we may reason in forward mode, feeding intermediate - results into rules selected from the context. The course of - reasoning is organized by maintaining a secondary fact called - ``@{fact calculation}'', apart from the primary ``@{fact this}'' - already provided by the Isar primitives. In the definitions below, + The generic calculational mechanism is based on the observation that rules + such as \trans:\~@{prop "x = y \ y = z \ x = z"} proceed from the premises + towards the conclusion in a deterministic fashion. Thus we may reason in + forward mode, feeding intermediate results into rules selected from the + context. The course of reasoning is organized by maintaining a secondary + fact called ``@{fact calculation}'', apart from the primary ``@{fact this}'' + already provided by the Isar primitives. In the definitions below, @{attribute OF} refers to @{inference resolution} - (\secref{sec:framework-resolution}) with multiple rule arguments, - and \trans\ represents to a suitable rule from the context: + (\secref{sec:framework-resolution}) with multiple rule arguments, and + \trans\ represents to a suitable rule from the context: \begin{matharray}{rcl} @{command "also"}\\<^sub>0\ & \equiv & @{command "note"}~\calculation = this\ \\ @@ -985,15 +935,15 @@ @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~\calculation\ \\ \end{matharray} - The start of a calculation is determined implicitly in the - text: here @{command also} sets @{fact calculation} to the current - result; any subsequent occurrence will update @{fact calculation} by - combination with the next result and a transitivity rule. The - calculational sequence is concluded via @{command finally}, where - the final result is exposed for use in a concluding claim. + The start of a calculation is determined implicitly in the text: here + @{command also} sets @{fact calculation} to the current result; any + subsequent occurrence will update @{fact calculation} by combination with + the next result and a transitivity rule. The calculational sequence is + concluded via @{command finally}, where the final result is exposed for use + in a concluding claim. - Here is a canonical proof pattern, using @{command have} to - establish the intermediate results: + Here is a canonical proof pattern, using @{command have} to establish the + intermediate results: \ (*<*) @@ -1009,20 +959,18 @@ (*>*) text \ - The term ``\\\'' above is a special abbreviation - provided by the Isabelle/Isar syntax layer: it statically refers to - the right-hand side argument of the previous statement given in the - text. Thus it happens to coincide with relevant sub-expressions in - the calculational chain, but the exact correspondence is dependent - on the transitivity rules being involved. + The term ``\\\'' above is a special abbreviation provided by the + Isabelle/Isar syntax layer: it statically refers to the right-hand side + argument of the previous statement given in the text. Thus it happens to + coincide with relevant sub-expressions in the calculational chain, but the + exact correspondence is dependent on the transitivity rules being involved. \<^medskip> - Symmetry rules such as @{prop "x = y \ y = x"} are like - transitivities with only one premise. Isar maintains a separate - rule collection declared via the @{attribute sym} attribute, to be - used in fact expressions ``\a [symmetric]\'', or single-step - proofs ``@{command assume}~\x = y\~@{command then}~@{command - have}~\y = x\~@{command ".."}''. + Symmetry rules such as @{prop "x = y \ y = x"} are like transitivities with + only one premise. Isar maintains a separate rule collection declared via the + @{attribute sym} attribute, to be used in fact expressions ``\a + [symmetric]\'', or single-step proofs ``@{command assume}~\x = y\~@{command + then}~@{command have}~\y = x\~@{command ".."}''. \ end \ No newline at end of file