| author | blanchet | 
| Sat, 04 Feb 2012 12:08:18 +0100 | |
| changeset 46409 | d4754183ccce | 
| parent 42666 | fee67c099d03 | 
| child 48279 | ddf866029eb2 | 
| permissions | -rw-r--r-- | 
| 29716 | 1 | theory Framework | 
| 42651 | 2 | imports Base Main | 
| 29716 | 3 | begin | 
| 4 | ||
| 5 | chapter {* The Isabelle/Isar Framework \label{ch:isar-framework} *}
 | |
| 6 | ||
| 7 | text {*
 | |
| 8 | Isabelle/Isar | |
| 9 |   \cite{Wenzel:1999:TPHOL,Wenzel-PhD,Nipkow-TYPES02,Wenzel-Paulson:2006,Wenzel:2006:Festschrift}
 | |
| 10 | is intended as a generic framework for developing formal | |
| 11 | mathematical documents with full proof checking. Definitions and | |
| 29735 | 12 | proofs are organized as theories. An assembly of theory sources may | 
| 29716 | 13 | be presented as a printed document; see also | 
| 14 |   \chref{ch:document-prep}.
 | |
| 15 | ||
| 16 | The main objective of Isar is the design of a human-readable | |
| 17 | structured proof language, which is called the ``primary proof | |
| 18 | format'' in Isar terminology. Such a primary proof language is | |
| 19 | somewhere in the middle between the extremes of primitive proof | |
| 20 | objects and actual natural language. In this respect, Isar is a bit | |
| 21 | more formalistic than Mizar | |
| 22 |   \cite{Trybulec:1993:MizarFeatures,Rudnicki:1992:MizarOverview,Wiedijk:1999:Mizar},
 | |
| 23 | using logical symbols for certain reasoning schemes where Mizar | |
| 24 |   would prefer English words; see \cite{Wenzel-Wiedijk:2002} for
 | |
| 25 | further comparisons of these systems. | |
| 26 | ||
| 27 | So Isar challenges the traditional way of recording informal proofs | |
| 28 | in mathematical prose, as well as the common tendency to see fully | |
| 29 | formal proofs directly as objects of some logical calculus (e.g.\ | |
| 30 |   @{text "\<lambda>"}-terms in a version of type theory).  In fact, Isar is
 | |
| 31 | better understood as an interpreter of a simple block-structured | |
| 29735 | 32 | language for describing the data flow of local facts and goals, | 
| 29716 | 33 | interspersed with occasional invocations of proof methods. | 
| 34 | Everything is reduced to logical inferences internally, but these | |
| 35 | steps are somewhat marginal compared to the overall bookkeeping of | |
| 36 | the interpretation process. Thanks to careful design of the syntax | |
| 37 | and semantics of Isar language elements, a formal record of Isar | |
| 38 | instructions may later appear as an intelligible text to the | |
| 39 | attentive reader. | |
| 40 | ||
| 41 | The Isar proof language has emerged from careful analysis of some | |
| 42 | inherent virtues of the existing logical framework of Isabelle/Pure | |
| 43 |   \cite{paulson-found,paulson700}, notably composition of higher-order
 | |
| 44 | natural deduction rules, which is a generalization of Gentzen's | |
| 45 |   original calculus \cite{Gentzen:1935}.  The approach of generic
 | |
| 46 | inference systems in Pure is continued by Isar towards actual proof | |
| 47 | texts. | |
| 48 | ||
| 49 | Concrete applications require another intermediate layer: an | |
| 50 |   object-logic.  Isabelle/HOL \cite{isa-tutorial} (simply-typed
 | |
| 51 | set-theory) is being used most of the time; Isabelle/ZF | |
| 52 |   \cite{isabelle-ZF} is less extensively developed, although it would
 | |
| 53 | probably fit better for classical mathematics. | |
| 29721 | 54 | |
| 29735 | 55 | \medskip In order to illustrate natural deduction in Isar, we shall | 
| 56 | refer to the background theory and library of Isabelle/HOL. This | |
| 57 | includes common notions of predicate logic, naive set-theory etc.\ | |
| 58 | using fairly standard mathematical notation. From the perspective | |
| 59 | of generic natural deduction there is nothing special about the | |
| 60 |   logical connectives of HOL (@{text "\<and>"}, @{text "\<or>"}, @{text "\<forall>"},
 | |
| 61 |   @{text "\<exists>"}, etc.), only the resulting reasoning principles are
 | |
| 62 | relevant to the user. There are similar rules available for | |
| 63 |   set-theory operators (@{text "\<inter>"}, @{text "\<union>"}, @{text "\<Inter>"}, @{text
 | |
| 64 | "\<Union>"}, etc.), or any other theory developed in the library (lattice | |
| 65 | theory, topology etc.). | |
| 29721 | 66 | |
| 67 | Subsequently we briefly review fragments of Isar proof texts | |
| 29735 | 68 | corresponding directly to such general deduction schemes. The | 
| 69 | examples shall refer to set-theory, to minimize the danger of | |
| 29721 | 70 | understanding connectives of predicate logic as something special. | 
| 71 | ||
| 72 |   \medskip The following deduction performs @{text "\<inter>"}-introduction,
 | |
| 73 | working forwards from assumptions towards the conclusion. We give | |
| 74 | both the Isar text, and depict the primitive rule involved, as | |
| 29735 | 75 | determined by unification of the problem against rules that are | 
| 76 | declared in the library context. | |
| 29721 | 77 | *} | 
| 78 | ||
| 79 | text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
 | |
| 80 | ||
| 81 | (*<*) | |
| 40964 | 82 | notepad | 
| 83 | begin | |
| 29721 | 84 | (*>*) | 
| 85 | assume "x \<in> A" and "x \<in> B" | |
| 86 | then have "x \<in> A \<inter> B" .. | |
| 87 | (*<*) | |
| 40964 | 88 | end | 
| 29721 | 89 | (*>*) | 
| 90 | ||
| 91 | text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
 | |
| 92 | ||
| 93 | text {*
 | |
| 94 |   \infer{@{prop "x \<in> A \<inter> B"}}{@{prop "x \<in> A"} & @{prop "x \<in> B"}}
 | |
| 95 | *} | |
| 96 | ||
| 97 | text_raw {*\end{minipage}*}
 | |
| 98 | ||
| 99 | text {*
 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 100 |   \medskip\noindent Note that @{command assume} augments the proof
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 101 |   context, @{command then} indicates that the current fact shall be
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 102 |   used in the next step, and @{command have} states an intermediate
 | 
| 29735 | 103 |   goal.  The two dots ``@{command ".."}'' refer to a complete proof of
 | 
| 104 | this claim, using the indicated facts and a canonical rule from the | |
| 29721 | 105 | context. We could have been more explicit here by spelling out the | 
| 106 |   final proof step via the @{command "by"} command:
 | |
| 107 | *} | |
| 108 | ||
| 109 | (*<*) | |
| 40964 | 110 | notepad | 
| 111 | begin | |
| 29721 | 112 | (*>*) | 
| 113 | assume "x \<in> A" and "x \<in> B" | |
| 114 | then have "x \<in> A \<inter> B" by (rule IntI) | |
| 115 | (*<*) | |
| 40964 | 116 | end | 
| 29721 | 117 | (*>*) | 
| 118 | ||
| 119 | text {*
 | |
| 120 |   \noindent The format of the @{text "\<inter>"}-introduction rule represents
 | |
| 121 | the most basic inference, which proceeds from given premises to a | |
| 29735 | 122 | conclusion, without any nested proof context involved. | 
| 29721 | 123 | |
| 29735 | 124 |   The next example performs backwards introduction on @{term "\<Inter>\<A>"},
 | 
| 125 | the intersection of all sets within a given set. This requires a | |
| 126 |   nested proof of set membership within a local context, where @{term
 | |
| 127 | A} is an arbitrary-but-fixed member of the collection: | |
| 29721 | 128 | *} | 
| 129 | ||
| 130 | text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
 | |
| 131 | ||
| 132 | (*<*) | |
| 40964 | 133 | notepad | 
| 134 | begin | |
| 29721 | 135 | (*>*) | 
| 136 | have "x \<in> \<Inter>\<A>" | |
| 137 | proof | |
| 138 | fix A | |
| 139 | assume "A \<in> \<A>" | |
| 29733 | 140 | show "x \<in> A" sorry %noproof | 
| 29721 | 141 | qed | 
| 142 | (*<*) | |
| 40964 | 143 | end | 
| 29721 | 144 | (*>*) | 
| 145 | ||
| 146 | text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
 | |
| 147 | ||
| 148 | text {*
 | |
| 149 |   \infer{@{prop "x \<in> \<Inter>\<A>"}}{\infer*{@{prop "x \<in> A"}}{@{text "[A][A \<in> \<A>]"}}}
 | |
| 150 | *} | |
| 151 | ||
| 152 | text_raw {*\end{minipage}*}
 | |
| 153 | ||
| 154 | text {*
 | |
| 155 | \medskip\noindent This Isar reasoning pattern again refers to the | |
| 156 | primitive rule depicted above. The system determines it in the | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 157 |   ``@{command proof}'' step, which could have been spelt out more
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 158 |   explicitly as ``@{command proof}~@{text "(rule InterI)"}''.  Note
 | 
| 29735 | 159 |   that the rule involves both a local parameter @{term "A"} and an
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 160 |   assumption @{prop "A \<in> \<A>"} in the nested reasoning.  This kind of
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 161 | compound rule typically demands a genuine sub-proof in Isar, working | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 162 | backwards rather than forwards as seen before. In the proof body we | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 163 |   encounter the @{command fix}-@{command assume}-@{command show}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 164 | outline of nested sub-proofs that is typical for Isar. The final | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 165 |   @{command show} is like @{command have} followed by an additional
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 166 | refinement of the enclosing claim, using the rule derived from the | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 167 | proof body. | 
| 29721 | 168 | |
| 169 |   \medskip The next example involves @{term "\<Union>\<A>"}, which can be
 | |
| 170 |   characterized as the set of all @{term "x"} such that @{prop "\<exists>A. x
 | |
| 171 |   \<in> A \<and> A \<in> \<A>"}.  The elimination rule for @{prop "x \<in> \<Union>\<A>"} does
 | |
| 172 |   not mention @{text "\<exists>"} and @{text "\<and>"} at all, but admits to obtain
 | |
| 173 |   directly a local @{term "A"} such that @{prop "x \<in> A"} and @{prop "A
 | |
| 174 | \<in> \<A>"} hold. This corresponds to the following Isar proof and | |
| 175 | inference rule, respectively: | |
| 176 | *} | |
| 177 | ||
| 178 | text_raw {*\medskip\begin{minipage}{0.6\textwidth}*}
 | |
| 179 | ||
| 180 | (*<*) | |
| 40964 | 181 | notepad | 
| 182 | begin | |
| 29721 | 183 | (*>*) | 
| 184 | assume "x \<in> \<Union>\<A>" | |
| 185 | then have C | |
| 186 | proof | |
| 187 | fix A | |
| 188 | assume "x \<in> A" and "A \<in> \<A>" | |
| 29733 | 189 | show C sorry %noproof | 
| 29721 | 190 | qed | 
| 191 | (*<*) | |
| 40964 | 192 | end | 
| 29721 | 193 | (*>*) | 
| 194 | ||
| 195 | text_raw {*\end{minipage}\begin{minipage}{0.4\textwidth}*}
 | |
| 196 | ||
| 197 | text {*
 | |
| 198 |   \infer{@{prop "C"}}{@{prop "x \<in> \<Union>\<A>"} & \infer*{@{prop "C"}~}{@{text "[A][x \<in> A, A \<in> \<A>]"}}}
 | |
| 199 | *} | |
| 200 | ||
| 201 | text_raw {*\end{minipage}*}
 | |
| 202 | ||
| 203 | text {*
 | |
| 204 | \medskip\noindent Although the Isar proof follows the natural | |
| 205 | deduction rule closely, the text reads not as natural as | |
| 206 | anticipated. There is a double occurrence of an arbitrary | |
| 207 |   conclusion @{prop "C"}, which represents the final result, but is
 | |
| 208 | irrelevant for now. This issue arises for any elimination rule | |
| 209 | involving local parameters. Isar provides the derived language | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 210 |   element @{command obtain}, which is able to perform the same
 | 
| 29721 | 211 | elimination proof more conveniently: | 
| 212 | *} | |
| 213 | ||
| 214 | (*<*) | |
| 40964 | 215 | notepad | 
| 216 | begin | |
| 29721 | 217 | (*>*) | 
| 218 | assume "x \<in> \<Union>\<A>" | |
| 219 | then obtain A where "x \<in> A" and "A \<in> \<A>" .. | |
| 220 | (*<*) | |
| 40964 | 221 | end | 
| 29721 | 222 | (*>*) | 
| 223 | ||
| 224 | text {*
 | |
| 225 |   \noindent Here we avoid to mention the final conclusion @{prop "C"}
 | |
| 226 | and return to plain forward reasoning. The rule involved in the | |
| 227 |   ``@{command ".."}'' proof is the same as before.
 | |
| 29716 | 228 | *} | 
| 229 | ||
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 230 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 231 | section {* The Pure framework \label{sec:framework-pure} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 232 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 233 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 234 |   The Pure logic \cite{paulson-found,paulson700} is an intuitionistic
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 235 |   fragment of higher-order logic \cite{church40}.  In type-theoretic
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 236 |   parlance, there are three levels of @{text "\<lambda>"}-calculus with
 | 
| 29735 | 237 |   corresponding arrows @{text "\<Rightarrow>"}/@{text "\<And>"}/@{text "\<Longrightarrow>"}:
 | 
| 238 | ||
| 239 | \medskip | |
| 240 |   \begin{tabular}{ll}
 | |
| 241 |   @{text "\<alpha> \<Rightarrow> \<beta>"} & syntactic function space (terms depending on terms) \\
 | |
| 242 |   @{text "\<And>x. B(x)"} & universal quantification (proofs depending on terms) \\
 | |
| 243 |   @{text "A \<Longrightarrow> B"} & implication (proofs depending on proofs) \\
 | |
| 244 |   \end{tabular}
 | |
| 245 | \medskip | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 246 | |
| 29735 | 247 | \noindent Here only the types of syntactic terms, and the | 
| 248 |   propositions of proof terms have been shown.  The @{text
 | |
| 249 | "\<lambda>"}-structure of proofs can be recorded as an optional feature of | |
| 250 |   the Pure inference kernel \cite{Berghofer-Nipkow:2000:TPHOL}, but
 | |
| 251 |   the formal system can never depend on them due to \emph{proof
 | |
| 252 | irrelevance}. | |
| 253 | ||
| 254 | On top of this most primitive layer of proofs, Pure implements a | |
| 255 | generic calculus for nested natural deduction rules, similar to | |
| 256 |   \cite{Schroeder-Heister:1984}.  Here object-logic inferences are
 | |
| 257 |   internalized as formulae over @{text "\<And>"} and @{text "\<Longrightarrow>"}.
 | |
| 258 | Combining such rule statements may involve higher-order unification | |
| 259 |   \cite{paulson-natural}.
 | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 260 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 261 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 262 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 263 | subsection {* Primitive inferences *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 264 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 265 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 266 |   Term syntax provides explicit notation for abstraction @{text "\<lambda>x ::
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 267 |   \<alpha>. b(x)"} and application @{text "b a"}, while types are usually
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 268 |   implicit thanks to type-inference; terms of type @{text "prop"} are
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 269 |   called propositions.  Logical statements are composed via @{text "\<And>x
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 270 |   :: \<alpha>. B(x)"} and @{text "A \<Longrightarrow> B"}.  Primitive reasoning operates on
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 271 |   judgments of the form @{text "\<Gamma> \<turnstile> \<phi>"}, with standard introduction
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 272 |   and elimination rules for @{text "\<And>"} and @{text "\<Longrightarrow>"} that refer to
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 273 |   fixed parameters @{text "x\<^isub>1, \<dots>, x\<^isub>m"} and hypotheses
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 274 |   @{text "A\<^isub>1, \<dots>, A\<^isub>n"} from the context @{text "\<Gamma>"};
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 275 | the corresponding proof terms are left implicit. The subsequent | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 276 |   inference rules define @{text "\<Gamma> \<turnstile> \<phi>"} inductively, relative to a
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 277 | collection of axioms: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 278 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 279 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 280 |   \infer{@{text "\<turnstile> A"}}{(@{text "A"} \text{~axiom})}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 281 | \qquad | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 282 |   \infer{@{text "A \<turnstile> A"}}{}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 283 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 284 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 285 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 286 |   \infer{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}{@{text "\<Gamma> \<turnstile> B(x)"} & @{text "x \<notin> \<Gamma>"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 287 | \qquad | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 288 |   \infer{@{text "\<Gamma> \<turnstile> B(a)"}}{@{text "\<Gamma> \<turnstile> \<And>x. B(x)"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 289 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 290 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 291 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 292 |   \infer{@{text "\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<Gamma> \<turnstile> B"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 293 | \qquad | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 294 |   \infer{@{text "\<Gamma>\<^sub>1 \<union> \<Gamma>\<^sub>2 \<turnstile> B"}}{@{text "\<Gamma>\<^sub>1 \<turnstile> A \<Longrightarrow> B"} & @{text "\<Gamma>\<^sub>2 \<turnstile> A"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 295 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 296 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 297 |   Furthermore, Pure provides a built-in equality @{text "\<equiv> :: \<alpha> \<Rightarrow> \<alpha> \<Rightarrow>
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 298 | prop"} with axioms for reflexivity, substitution, extensionality, | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 299 |   and @{text "\<alpha>\<beta>\<eta>"}-conversion on @{text "\<lambda>"}-terms.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 300 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 301 | \medskip An object-logic introduces another layer on top of Pure, | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 302 |   e.g.\ with types @{text "i"} for individuals and @{text "o"} for
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 303 |   propositions, term constants @{text "Trueprop :: o \<Rightarrow> prop"} as
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 304 |   (implicit) derivability judgment and connectives like @{text "\<and> :: o
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 305 |   \<Rightarrow> o \<Rightarrow> o"} or @{text "\<forall> :: (i \<Rightarrow> o) \<Rightarrow> o"}, and axioms for object-level
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 306 |   rules such as @{text "conjI: A \<Longrightarrow> B \<Longrightarrow> A \<and> B"} or @{text "allI: (\<And>x. B
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 307 | x) \<Longrightarrow> \<forall>x. B x"}. Derived object rules are represented as theorems of | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 308 | Pure. After the initial object-logic setup, further axiomatizations | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 309 | are usually avoided; plain definitions and derived principles are | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 310 | used exclusively. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 311 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 312 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 313 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 314 | subsection {* Reasoning with rules \label{sec:framework-resolution} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 315 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 316 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 317 | Primitive inferences mostly serve foundational purposes. The main | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 318 | reasoning mechanisms of Pure operate on nested natural deduction | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 319 |   rules expressed as formulae, using @{text "\<And>"} to bind local
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 320 |   parameters and @{text "\<Longrightarrow>"} to express entailment.  Multiple
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 321 | parameters and premises are represented by repeating these | 
| 29735 | 322 | connectives in a right-associative manner. | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 323 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 324 |   Since @{text "\<And>"} and @{text "\<Longrightarrow>"} commute thanks to the theorem
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 325 |   @{prop "(A \<Longrightarrow> (\<And>x. B x)) \<equiv> (\<And>x. A \<Longrightarrow> B x)"}, we may assume w.l.o.g.\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 326 | that rule statements always observe the normal form where | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 327 | quantifiers are pulled in front of implications at each level of | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 328 | nesting. This means that any Pure proposition may be presented as a | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 329 |   \emph{Hereditary Harrop Formula} \cite{Miller:1991} which is of the
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 330 |   form @{text "\<And>x\<^isub>1 \<dots> x\<^isub>m. H\<^isub>1 \<Longrightarrow> \<dots> H\<^isub>n \<Longrightarrow>
 | 
| 29735 | 331 |   A"} for @{text "m, n \<ge> 0"}, and @{text "A"} atomic, and @{text
 | 
| 332 | "H\<^isub>1, \<dots>, H\<^isub>n"} being recursively of the same format. | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 333 | Following the convention that outermost quantifiers are implicit, | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 334 |   Horn clauses @{text "A\<^isub>1 \<Longrightarrow> \<dots> A\<^isub>n \<Longrightarrow> A"} are a special
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 335 | case of this. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 336 | |
| 29735 | 337 |   For example, @{text "\<inter>"}-introduction rule encountered before is
 | 
| 338 | represented as a Pure theorem as follows: | |
| 339 | \[ | |
| 340 |   @{text "IntI:"}~@{prop "x \<in> A \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A \<inter> B"}
 | |
| 341 | \] | |
| 342 | ||
| 343 | \noindent This is a plain Horn clause, since no further nesting on | |
| 344 |   the left is involved.  The general @{text "\<Inter>"}-introduction
 | |
| 345 | corresponds to a Hereditary Harrop Formula with one additional level | |
| 346 | of nesting: | |
| 347 | \[ | |
| 348 |   @{text "InterI:"}~@{prop "(\<And>A. A \<in> \<A> \<Longrightarrow> x \<in> A) \<Longrightarrow> x \<in> \<Inter>\<A>"}
 | |
| 349 | \] | |
| 350 | ||
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 351 |   \medskip Goals are also represented as rules: @{text "A\<^isub>1 \<Longrightarrow>
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 352 |   \<dots> A\<^isub>n \<Longrightarrow> C"} states that the sub-goals @{text "A\<^isub>1, \<dots>,
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 353 |   A\<^isub>n"} entail the result @{text "C"}; for @{text "n = 0"} the
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 354 |   goal is finished.  To allow @{text "C"} being a rule statement
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 355 |   itself, we introduce the protective marker @{text "# :: prop \<Rightarrow>
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 356 | prop"}, which is defined as identity and hidden from the user. We | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 357 | initialize and finish goal states as follows: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 358 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 359 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 360 |   \begin{array}{c@ {\qquad}c}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 361 |   \infer[(@{inference_def init})]{@{text "C \<Longrightarrow> #C"}}{} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 362 |   \infer[(@{inference_def finish})]{@{text C}}{@{text "#C"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 363 |   \end{array}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 364 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 365 | |
| 29735 | 366 | \noindent Goal states are refined in intermediate proof steps until | 
| 367 | a finished form is achieved. Here the two main reasoning principles | |
| 368 |   are @{inference resolution}, for back-chaining a rule against a
 | |
| 369 |   sub-goal (replacing it by zero or more sub-goals), and @{inference
 | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 370 | assumption}, for solving a sub-goal (finding a short-circuit with | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 371 |   local assumptions).  Below @{text "\<^vec>x"} stands for @{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 372 |   "x\<^isub>1, \<dots>, x\<^isub>n"} (@{text "n \<ge> 0"}).
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 373 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 374 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 375 |   \infer[(@{inference_def resolution})]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 376 |   {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>A (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 377 |   {\begin{tabular}{rl}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 378 |     @{text "rule:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 379 |     @{text "\<^vec>A \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 380 |     @{text "goal:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 381 |     @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 382 |     @{text "goal unifier:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 383 |     @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 384 |    \end{tabular}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 385 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 386 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 387 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 388 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 389 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 390 |   \infer[(@{inference_def assumption})]{@{text "C\<vartheta>"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 391 |   {\begin{tabular}{rl}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 392 |     @{text "goal:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 393 |     @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> A \<^vec>x) \<Longrightarrow> C"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 394 |     @{text "assm unifier:"} & @{text "A\<vartheta> = H\<^sub>i\<vartheta>"}~~\text{(for some~@{text "H\<^sub>i"})} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 395 |    \end{tabular}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 396 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 397 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 398 | The following trace illustrates goal-oriented reasoning in | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 399 | Isabelle/Pure: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 400 | |
| 29735 | 401 |   {\footnotesize
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 402 | \medskip | 
| 29735 | 403 |   \begin{tabular}{r@ {\quad}l}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 404 |   @{text "(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #(A \<and> B \<Longrightarrow> B \<and> A)"} & @{text "(init)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 405 |   @{text "(A \<and> B \<Longrightarrow> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution B \<Longrightarrow> A \<Longrightarrow> B \<and> A)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 406 |   @{text "(A \<and> B \<Longrightarrow> A \<and> B) \<Longrightarrow> (A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> B)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 407 |   @{text "(A \<and> B \<Longrightarrow> A) \<Longrightarrow> #\<dots>"} & @{text "(assumption)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 408 |   @{text "(A \<and> B \<Longrightarrow> B \<and> A) \<Longrightarrow> #\<dots>"} & @{text "(resolution A \<and> B \<Longrightarrow> A)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 409 |   @{text "#\<dots>"} & @{text "(assumption)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 410 |   @{text "A \<and> B \<Longrightarrow> B \<and> A"} & @{text "(finish)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 411 |   \end{tabular}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 412 | \medskip | 
| 29735 | 413 | } | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 414 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 415 |   Compositions of @{inference assumption} after @{inference
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 416 | resolution} occurs quite often, typically in elimination steps. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 417 | Traditional Isabelle tactics accommodate this by a combined | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 418 |   @{inference_def elim_resolution} principle.  In contrast, Isar uses
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 419 | a slightly more refined combination, where the assumptions to be | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 420 | closed are marked explicitly, using again the protective marker | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 421 |   @{text "#"}:
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 422 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 423 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 424 |   \infer[(@{inference refinement})]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 425 |   {@{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> \<^vec>G' (\<^vec>a \<^vec>x))\<vartheta> \<Longrightarrow> C\<vartheta>"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 426 |   {\begin{tabular}{rl}
 | 
| 42666 | 427 |     @{text "sub\<hyphen>proof:"} &
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 428 |     @{text "\<^vec>G \<^vec>a \<Longrightarrow> B \<^vec>a"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 429 |     @{text "goal:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 430 |     @{text "(\<And>\<^vec>x. \<^vec>H \<^vec>x \<Longrightarrow> B' \<^vec>x) \<Longrightarrow> C"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 431 |     @{text "goal unifier:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 432 |     @{text "(\<lambda>\<^vec>x. B (\<^vec>a \<^vec>x))\<vartheta> = B'\<vartheta>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 433 |     @{text "assm unifiers:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 434 |     @{text "(\<lambda>\<^vec>x. G\<^sub>j (\<^vec>a \<^vec>x))\<vartheta> = #H\<^sub>i\<vartheta>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 435 |     & \quad (for each marked @{text "G\<^sub>j"} some @{text "#H\<^sub>i"}) \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 436 |    \end{tabular}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 437 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 438 | |
| 42666 | 439 |   \noindent Here the @{text "sub\<hyphen>proof"} rule stems from the
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 440 |   main @{command fix}-@{command assume}-@{command show} outline of
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 441 |   Isar (cf.\ \secref{sec:framework-subproof}): each assumption
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 442 |   indicated in the text results in a marked premise @{text "G"} above.
 | 
| 29735 | 443 | The marking enforces resolution against one of the sub-goal's | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 444 |   premises.  Consequently, @{command fix}-@{command assume}-@{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 445 | show} enables to fit the result of a sub-proof quite robustly into a | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 446 | pending sub-goal, while maintaining a good measure of flexibility. | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 447 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 448 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 449 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 450 | section {* The Isar proof language \label{sec:framework-isar} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 451 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 452 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 453 | Structured proofs are presented as high-level expressions for | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 454 | composing entities of Pure (propositions, facts, and goals). The | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 455 | Isar proof language allows to organize reasoning within the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 456 | underlying rule calculus of Pure, but Isar is not another logical | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 457 | calculus! | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 458 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 459 | Isar is an exercise in sound minimalism. Approximately half of the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 460 | language is introduced as primitive, the rest defined as derived | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 461 | concepts. The following grammar describes the core language | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 462 |   (category @{text "proof"}), which is embedded into theory
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 463 |   specification elements such as @{command theorem}; see also
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 464 |   \secref{sec:framework-stmt} for the separate category @{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 465 | "statement"}. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 466 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 467 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 468 |   \begin{tabular}{rcl}
 | 
| 42666 | 469 |     @{text "theory\<hyphen>stmt"} & = & @{command "theorem"}~@{text "statement proof  |"}~~@{command "definition"}~@{text "\<dots>  |  \<dots>"} \\[1ex]
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 470 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 471 |     @{text "proof"} & = & @{text "prfx\<^sup>*"}~@{command "proof"}~@{text "method\<^sup>? stmt\<^sup>*"}~@{command "qed"}~@{text "method\<^sup>?"} \\[1ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 472 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 473 |     @{text prfx} & = & @{command "using"}~@{text "facts"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 474 |     & @{text "|"} & @{command "unfolding"}~@{text "facts"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 475 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 476 |     @{text stmt} & = & @{command "{"}~@{text "stmt\<^sup>*"}~@{command "}"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 477 |     & @{text "|"} & @{command "next"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 478 |     & @{text "|"} & @{command "note"}~@{text "name = facts"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 479 |     & @{text "|"} & @{command "let"}~@{text "term = term"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 480 |     & @{text "|"} & @{command "fix"}~@{text "var\<^sup>+"} \\
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 481 |     & @{text "|"} & @{command assume}~@{text "\<guillemotleft>inference\<guillemotright> name: props"} \\
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 482 |     & @{text "|"} & @{command "then"}@{text "\<^sup>?"}~@{text goal} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 483 |     @{text goal} & = & @{command "have"}~@{text "name: props proof"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 484 |     & @{text "|"} & @{command "show"}~@{text "name: props proof"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 485 |   \end{tabular}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 486 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 487 | \medskip Simultaneous propositions or facts may be separated by the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 488 |   @{keyword "and"} keyword.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 489 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 490 | \medskip The syntax for terms and propositions is inherited from | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 491 |   Pure (and the object-logic).  A @{text "pattern"} is a @{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 492 | "term"} with schematic variables, to be bound by higher-order | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 493 | matching. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 494 | |
| 29735 | 495 | \medskip Facts may be referenced by name or proposition. For | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 496 |   example, the result of ``@{command have}~@{text "a: A \<langle>proof\<rangle>"}''
 | 
| 29735 | 497 |   becomes available both as @{text "a"} and
 | 
| 498 |   \isacharbackquoteopen@{text "A"}\isacharbackquoteclose.  Moreover,
 | |
| 499 | fact expressions may involve attributes that modify either the | |
| 500 | theorem or the background context. For example, the expression | |
| 501 |   ``@{text "a [OF b]"}'' refers to the composition of two facts
 | |
| 502 |   according to the @{inference resolution} inference of
 | |
| 503 |   \secref{sec:framework-resolution}, while ``@{text "a [intro]"}''
 | |
| 504 | declares a fact as introduction rule in the context. | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 505 | |
| 29735 | 506 |   The special fact called ``@{fact this}'' always refers to the last
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 507 |   result, as produced by @{command note}, @{command assume}, @{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 508 |   have}, or @{command show}.  Since @{command note} occurs
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 509 |   frequently together with @{command then} we provide some
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 510 | abbreviations: | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 511 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 512 | \medskip | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 513 |   \begin{tabular}{rcl}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 514 |     @{command from}~@{text a} & @{text "\<equiv>"} & @{command note}~@{text a}~@{command then} \\
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 515 |     @{command with}~@{text a} & @{text "\<equiv>"} & @{command from}~@{text "a \<AND> this"} \\
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 516 |   \end{tabular}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 517 | \medskip | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 518 | |
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 519 |   The @{text "method"} category is essentially a parameter and may be
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 520 |   populated later.  Methods use the facts indicated by @{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 521 |   "then"} or @{command using}, and then operate on the goal state.
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 522 |   Some basic methods are predefined: ``@{method "-"}'' leaves the goal
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 523 |   unchanged, ``@{method this}'' applies the facts as rules to the
 | 
| 42626 | 524 |   goal, ``@{method (Pure) "rule"}'' applies the facts to another rule and the
 | 
| 525 |   result to the goal (both ``@{method this}'' and ``@{method (Pure) rule}''
 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 526 |   refer to @{inference resolution} of
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 527 |   \secref{sec:framework-resolution}).  The secondary arguments to
 | 
| 42626 | 528 |   ``@{method (Pure) rule}'' may be specified explicitly as in ``@{text "(rule
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 529 | a)"}'', or picked from the context. In the latter case, the system | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 530 |   first tries rules declared as @{attribute (Pure) elim} or
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 531 |   @{attribute (Pure) dest}, followed by those declared as @{attribute
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 532 | (Pure) intro}. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 533 | |
| 42626 | 534 |   The default method for @{command proof} is ``@{method (Pure) rule}''
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 535 |   (arguments picked from the context), for @{command qed} it is
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 536 |   ``@{method "-"}''.  Further abbreviations for terminal proof steps
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 537 |   are ``@{command "by"}~@{text "method\<^sub>1 method\<^sub>2"}'' for
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 538 |   ``@{command proof}~@{text "method\<^sub>1"}~@{command qed}~@{text
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 539 |   "method\<^sub>2"}'', and ``@{command ".."}'' for ``@{command
 | 
| 42626 | 540 |   "by"}~@{method (Pure) rule}, and ``@{command "."}'' for ``@{command
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 541 |   "by"}~@{method this}''.  The @{command unfolding} element operates
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 542 | directly on the current facts and goal by applying equalities. | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 543 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 544 |   \medskip Block structure can be indicated explicitly by ``@{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 545 |   "{"}~@{text "\<dots>"}~@{command "}"}'', although the body of a sub-proof
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 546 |   already involves implicit nesting.  In any case, @{command next}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 547 | jumps into the next section of a block, i.e.\ it acts like closing | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 548 | an implicit block scope and opening another one; there is no direct | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 549 | correspondence to subgoals here. | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 550 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 551 |   The remaining elements @{command fix} and @{command assume} build up
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 552 |   a local context (see \secref{sec:framework-context}), while
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 553 |   @{command show} refines a pending sub-goal by the rule resulting
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 554 |   from a nested sub-proof (see \secref{sec:framework-subproof}).
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 555 | Further derived concepts will support calculational reasoning (see | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 556 |   \secref{sec:framework-calc}).
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 557 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 558 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 559 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 560 | subsection {* Context elements \label{sec:framework-context} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 561 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 562 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 563 |   In judgments @{text "\<Gamma> \<turnstile> \<phi>"} of the primitive framework, @{text "\<Gamma>"}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 564 | essentially acts like a proof context. Isar elaborates this idea | 
| 29735 | 565 | towards a higher-level notion, with additional information for | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 566 | type-inference, term abbreviations, local facts, hypotheses etc. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 567 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 568 |   The element @{command fix}~@{text "x :: \<alpha>"} declares a local
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 569 | parameter, i.e.\ an arbitrary-but-fixed entity of a given type; in | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 570 |   results exported from the context, @{text "x"} may become anything.
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 571 |   The @{command assume}~@{text "\<guillemotleft>inference\<guillemotright>"} element provides a
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 572 |   general interface to hypotheses: ``@{command assume}~@{text
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 573 |   "\<guillemotleft>inference\<guillemotright> A"}'' produces @{text "A \<turnstile> A"} locally, while the
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 574 |   included inference tells how to discharge @{text A} from results
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 575 |   @{text "A \<turnstile> B"} later on.  There is no user-syntax for @{text
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 576 | "\<guillemotleft>inference\<guillemotright>"}, i.e.\ it may only occur internally when derived | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 577 | commands are defined in ML. | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 578 | |
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 579 |   At the user-level, the default inference for @{command assume} is
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 580 |   @{inference discharge} as given below.  The additional variants
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 581 |   @{command presume} and @{command def} are defined as follows:
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 582 | |
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 583 | \medskip | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 584 |   \begin{tabular}{rcl}
 | 
| 42666 | 585 |     @{command presume}~@{text A} & @{text "\<equiv>"} & @{command assume}~@{text "\<guillemotleft>weak\<hyphen>discharge\<guillemotright> A"} \\
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 586 |     @{command def}~@{text "x \<equiv> a"} & @{text "\<equiv>"} & @{command fix}~@{text x}~@{command assume}~@{text "\<guillemotleft>expansion\<guillemotright> x \<equiv> a"} \\
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 587 |   \end{tabular}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 588 | \medskip | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 589 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 590 | \[ | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 591 |   \infer[(@{inference_def discharge})]{@{text "\<strut>\<Gamma> - A \<turnstile> #A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 592 | \] | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 593 | \[ | 
| 42666 | 594 |   \infer[(@{inference_def "weak\<hyphen>discharge"})]{@{text "\<strut>\<Gamma> - A \<turnstile> A \<Longrightarrow> B"}}{@{text "\<strut>\<Gamma> \<turnstile> B"}}
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 595 | \] | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 596 | \[ | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 597 |   \infer[(@{inference_def expansion})]{@{text "\<strut>\<Gamma> - (x \<equiv> a) \<turnstile> B a"}}{@{text "\<strut>\<Gamma> \<turnstile> B x"}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 598 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 599 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 600 |   \medskip Note that @{inference discharge} and @{inference
 | 
| 42666 | 601 |   "weak\<hyphen>discharge"} differ in the marker for @{prop A}, which is
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 602 |   relevant when the result of a @{command fix}-@{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 603 |   assume}-@{command show} outline is composed with a pending goal,
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 604 |   cf.\ \secref{sec:framework-subproof}.
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 605 | |
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 606 |   The most interesting derived context element in Isar is @{command
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 607 |   obtain} \cite[\S5.3]{Wenzel-PhD}, which supports generalized
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 608 |   elimination steps in a purely forward manner.  The @{command obtain}
 | 
| 29739 | 609 |   command takes a specification of parameters @{text "\<^vec>x"} and
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 610 |   assumptions @{text "\<^vec>A"} to be added to the context, together
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 611 | with a proof of a case rule stating that this extension is | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 612 | conservative (i.e.\ may be removed from closed results later on): | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 613 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 614 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 615 |   \begin{tabular}{l}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 616 |   @{text "\<langle>facts\<rangle>"}~~@{command obtain}~@{text "\<^vec>x \<WHERE> \<^vec>A \<^vec>x  \<langle>proof\<rangle> \<equiv>"} \\[0.5ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 617 |   \quad @{command have}~@{text "case: \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis\<rangle>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 618 |   \quad @{command proof}~@{method "-"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 619 |   \qquad @{command fix}~@{text thesis} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 620 |   \qquad @{command assume}~@{text "[intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis"} \\
 | 
| 29733 | 621 |   \qquad @{command show}~@{text thesis}~@{command using}~@{text "\<langle>facts\<rangle> \<langle>proof\<rangle>"} \\
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 622 |   \quad @{command qed} \\
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 623 |   \quad @{command fix}~@{text "\<^vec>x"}~@{command assume}~@{text "\<guillemotleft>elimination case\<guillemotright> \<^vec>A \<^vec>x"} \\
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 624 |   \end{tabular}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 625 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 626 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 627 | \[ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 628 |   \infer[(@{inference elimination})]{@{text "\<Gamma> \<turnstile> B"}}{
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 629 |     \begin{tabular}{rl}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 630 |     @{text "case:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 631 |     @{text "\<Gamma> \<turnstile> \<And>thesis. (\<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis) \<Longrightarrow> thesis"} \\[0.2ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 632 |     @{text "result:"} &
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 633 |     @{text "\<Gamma> \<union> \<^vec>A \<^vec>y \<turnstile> B"} \\[0.2ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 634 |     \end{tabular}}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 635 | \] | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 636 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 637 |   \noindent Here the name ``@{text thesis}'' is a specific convention
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 638 | for an arbitrary-but-fixed proposition; in the primitive natural | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 639 |   deduction rules shown before we have occasionally used @{text C}.
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 640 |   The whole statement of ``@{command obtain}~@{text x}~@{keyword
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 641 |   "where"}~@{text "A x"}'' may be read as a claim that @{text "A x"}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 642 |   may be assumed for some arbitrary-but-fixed @{text "x"}.  Also note
 | 
| 29737 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 643 |   that ``@{command obtain}~@{text "A \<AND> B"}'' without parameters
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 644 |   is similar to ``@{command have}~@{text "A \<AND> B"}'', but the
 | 
| 
866841668584
replaced 'assm' by 'assume' -- regular 'assume' is default, if inference is omitted;
 wenzelm parents: 
29735diff
changeset | 645 | latter involves multiple sub-goals. | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 646 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 647 | \medskip The subsequent Isar proof texts explain all context | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 648 | elements introduced above using the formal proof language itself. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 649 | After finishing a local proof within a block, we indicate the | 
| 29739 | 650 |   exported result via @{command note}.
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 651 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 652 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 653 | (*<*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 654 | theorem True | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 655 | proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 656 | (*>*) | 
| 40476 | 657 |   txt_raw {* \begin{minipage}[t]{0.45\textwidth} *}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 658 |   {
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 659 | fix x | 
| 29735 | 660 | have "B x" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 661 | } | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 662 | note `\<And>x. B x` | 
| 40476 | 663 |   txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
 | 
| 29735 | 664 |   {
 | 
| 665 | assume A | |
| 666 | have B sorry %noproof | |
| 667 | } | |
| 668 | note `A \<Longrightarrow> B` | |
| 40476 | 669 |   txt_raw {* \end{minipage}\\[3ex]\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 670 |   {
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 671 | def x \<equiv> a | 
| 29735 | 672 | have "B x" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 673 | } | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 674 | note `B a` | 
| 40476 | 675 |   txt_raw {* \end{minipage}\quad\begin{minipage}[t]{0.45\textwidth} *}(*<*)next(*>*)
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 676 |   {
 | 
| 29735 | 677 | obtain x where "A x" sorry %noproof | 
| 29733 | 678 | have B sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 679 | } | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 680 | note `B` | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 681 |   txt_raw {* \end{minipage} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 682 | (*<*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 683 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 684 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 685 | |
| 29739 | 686 | text {*
 | 
| 687 | \bigskip\noindent This illustrates the meaning of Isar context | |
| 688 | elements without goals getting in between. | |
| 689 | *} | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 690 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 691 | subsection {* Structured statements \label{sec:framework-stmt} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 692 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 693 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 694 |   The category @{text "statement"} of top-level theorem specifications
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 695 | is defined as follows: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 696 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 697 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 698 |   \begin{tabular}{rcl}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 699 |   @{text "statement"} & @{text "\<equiv>"} & @{text "name: props \<AND> \<dots>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 700 |   & @{text "|"} & @{text "context\<^sup>* conclusion"} \\[0.5ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 701 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 702 |   @{text "context"} & @{text "\<equiv>"} & @{text "\<FIXES> vars \<AND> \<dots>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 703 |   & @{text "|"} & @{text "\<ASSUMES> name: props \<AND> \<dots>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 704 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 705 |   @{text "conclusion"} & @{text "\<equiv>"} & @{text "\<SHOWS> name: props \<AND> \<dots>"} \\
 | 
| 29735 | 706 |   & @{text "|"} & @{text "\<OBTAINS> vars \<AND> \<dots> \<WHERE> name: props \<AND> \<dots>"} \\
 | 
| 707 |   & & \quad @{text "\<BBAR> \<dots>"} \\
 | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 708 |   \end{tabular}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 709 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 710 |   \medskip\noindent A simple @{text "statement"} consists of named
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 711 | propositions. The full form admits local context elements followed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 712 |   by the actual conclusions, such as ``@{keyword "fixes"}~@{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 713 |   x}~@{keyword "assumes"}~@{text "A x"}~@{keyword "shows"}~@{text "B
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 714 | x"}''. The final result emerges as a Pure rule after discharging | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 715 |   the context: @{prop "\<And>x. A x \<Longrightarrow> B x"}.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 716 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 717 |   The @{keyword "obtains"} variant is another abbreviation defined
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 718 |   below; unlike @{command obtain} (cf.\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 719 |   \secref{sec:framework-context}) there may be several ``cases''
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 720 |   separated by ``@{text "\<BBAR>"}'', each consisting of several
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 721 |   parameters (@{text "vars"}) and several premises (@{text "props"}).
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 722 | This specifies multi-branch elimination rules. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 723 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 724 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 725 |   \begin{tabular}{l}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 726 |   @{text "\<OBTAINS> \<^vec>x \<WHERE> \<^vec>A \<^vec>x   \<BBAR>   \<dots>   \<equiv>"} \\[0.5ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 727 |   \quad @{text "\<FIXES> thesis"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 728 |   \quad @{text "\<ASSUMES> [intro]: \<And>\<^vec>x. \<^vec>A \<^vec>x \<Longrightarrow> thesis  \<AND>  \<dots>"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 729 |   \quad @{text "\<SHOWS> thesis"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 730 |   \end{tabular}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 731 | \medskip | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 732 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 733 | Presenting structured statements in such an ``open'' format usually | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 734 | simplifies the subsequent proof, because the outer structure of the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 735 | problem is already laid out directly. E.g.\ consider the following | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 736 |   canonical patterns for @{text "\<SHOWS>"} and @{text "\<OBTAINS>"},
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 737 | respectively: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 738 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 739 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 740 | text_raw {*\begin{minipage}{0.5\textwidth}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 741 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 742 | theorem | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 743 | fixes x and y | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 744 | assumes "A x" and "B y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 745 | shows "C x y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 746 | proof - | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 747 | from `A x` and `B y` | 
| 29733 | 748 | show "C x y" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 749 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 750 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 751 | text_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 752 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 753 | theorem | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 754 | obtains x and y | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 755 | where "A x" and "B y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 756 | proof - | 
| 29733 | 757 | have "A a" and "B b" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 758 | then show thesis .. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 759 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 760 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 761 | text_raw {*\end{minipage}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 762 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 763 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 764 |   \medskip\noindent Here local facts \isacharbackquoteopen@{text "A
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 765 |   x"}\isacharbackquoteclose\ and \isacharbackquoteopen@{text "B
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 766 | y"}\isacharbackquoteclose\ are referenced immediately; there is no | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 767 | need to decompose the logical rule structure again. In the second | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 768 |   proof the final ``@{command then}~@{command show}~@{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 769 |   thesis}~@{command ".."}''  involves the local rule case @{text "\<And>x
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 770 |   y. A x \<Longrightarrow> B y \<Longrightarrow> thesis"} for the particular instance of terms @{text
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 771 |   "a"} and @{text "b"} produced in the body.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 772 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 773 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 774 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 775 | subsection {* Structured proof refinement \label{sec:framework-subproof} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 776 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 777 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 778 | By breaking up the grammar for the Isar proof language, we may | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 779 | understand a proof text as a linear sequence of individual proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 780 | commands. These are interpreted as transitions of the Isar virtual | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 781 | machine (Isar/VM), which operates on a block-structured | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 782 | configuration in single steps. This allows users to write proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 783 | texts in an incremental manner, and inspect intermediate | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 784 | configurations for debugging. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 785 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 786 | The basic idea is analogous to evaluating algebraic expressions on a | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 787 |   stack machine: @{text "(a + b) \<cdot> c"} then corresponds to a sequence
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 788 |   of single transitions for each symbol @{text "(, a, +, b, ), \<cdot>, c"}.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 789 | In Isar the algebraic values are facts or goals, and the operations | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 790 | are inferences. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 791 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 792 | \medskip The Isar/VM state maintains a stack of nodes, each node | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 793 | contains the local proof context, the linguistic mode, and a pending | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 794 | goal (optional). The mode determines the type of transition that | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 795 | may be performed next, it essentially alternates between forward and | 
| 29738 | 796 | backward reasoning, with an intermediate stage for chained facts | 
| 797 |   (see \figref{fig:isar-vm}).
 | |
| 798 | ||
| 799 |   \begin{figure}[htb]
 | |
| 800 |   \begin{center}
 | |
| 801 |   \includegraphics[width=0.8\textwidth]{Thy/document/isar-vm}
 | |
| 802 |   \end{center}
 | |
| 803 |   \caption{Isar/VM modes}\label{fig:isar-vm}
 | |
| 804 |   \end{figure}
 | |
| 805 | ||
| 806 |   For example, in @{text "state"} mode Isar acts like a mathematical
 | |
| 807 |   scratch-pad, accepting declarations like @{command fix}, @{command
 | |
| 808 |   assume}, and claims like @{command have}, @{command show}.  A goal
 | |
| 809 |   statement changes the mode to @{text "prove"}, which means that we
 | |
| 810 |   may now refine the problem via @{command unfolding} or @{command
 | |
| 811 |   proof}.  Then we are again in @{text "state"} mode of a proof body,
 | |
| 812 |   which may issue @{command show} statements to solve pending
 | |
| 813 |   sub-goals.  A concluding @{command qed} will return to the original
 | |
| 814 |   @{text "state"} mode one level upwards.  The subsequent Isar/VM
 | |
| 815 | trace indicates block structure, linguistic mode, goal state, and | |
| 816 | inferences: | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 817 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 818 | |
| 29735 | 819 | text_raw {* \begingroup\footnotesize *}
 | 
| 40964 | 820 | (*<*)notepad begin | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 821 | (*>*) | 
| 29735 | 822 |   txt_raw {* \begin{minipage}[t]{0.18\textwidth} *}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 823 | have "A \<longrightarrow> B" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 824 | proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 825 | assume A | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 826 | show B | 
| 29733 | 827 | sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 828 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 829 |   txt_raw {* \end{minipage}\quad
 | 
| 29735 | 830 | \begin{minipage}[t]{0.06\textwidth}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 831 | @{text "begin"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 832 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 833 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 834 | @{text "begin"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 835 | @{text "end"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 836 | @{text "end"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 837 | \end{minipage}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 838 | \begin{minipage}[t]{0.08\textwidth}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 839 | @{text "prove"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 840 | @{text "state"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 841 | @{text "state"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 842 | @{text "prove"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 843 | @{text "state"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 844 | @{text "state"} \\
 | 
| 29735 | 845 | \end{minipage}\begin{minipage}[t]{0.35\textwidth}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 846 | @{text "(A \<longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 847 | @{text "(A \<Longrightarrow> B) \<Longrightarrow> #(A \<longrightarrow> B)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 848 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 849 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 850 | @{text "#(A \<longrightarrow> B)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 851 | @{text "A \<longrightarrow> B"} \\
 | 
| 29735 | 852 | \end{minipage}\begin{minipage}[t]{0.4\textwidth}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 853 | @{text "(init)"} \\
 | 
| 29735 | 854 | @{text "(resolution impI)"} \\
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 855 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 856 | \\ | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 857 | @{text "(refinement #A \<Longrightarrow> B)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 858 | @{text "(finish)"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 859 | \end{minipage} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 860 | (*<*) | 
| 40964 | 861 | end | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 862 | (*>*) | 
| 29735 | 863 | text_raw {* \endgroup *}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 864 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 865 | text {*
 | 
| 29735 | 866 |   \noindent Here the @{inference refinement} inference from
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 867 |   \secref{sec:framework-resolution} mediates composition of Isar
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 868 | sub-proofs nicely. Observe that this principle incorporates some | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 869 | degree of freedom in proof composition. In particular, the proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 870 | body allows parameters and assumptions to be re-ordered, or commuted | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 871 | according to Hereditary Harrop Form. Moreover, context elements | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 872 | that are not used in a sub-proof may be omitted altogether. For | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 873 | example: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 874 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 875 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 876 | text_raw {*\begin{minipage}{0.5\textwidth}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 877 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 878 | (*<*) | 
| 40964 | 879 | notepad | 
| 880 | begin | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 881 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 882 | have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 883 | proof - | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 884 | fix x and y | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 885 | assume "A x" and "B y" | 
| 29733 | 886 | show "C x y" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 887 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 888 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 889 | txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 890 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 891 | (*<*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 892 | next | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 893 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 894 | have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 895 | proof - | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 896 | fix x assume "A x" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 897 | fix y assume "B y" | 
| 29733 | 898 | show "C x y" sorry %noproof | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 899 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 900 | |
| 29735 | 901 | txt_raw {*\end{minipage}\\[3ex]\begin{minipage}{0.5\textwidth}*}
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 902 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 903 | (*<*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 904 | next | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 905 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 906 | have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 907 | proof - | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 908 | fix y assume "B y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 909 | fix x assume "A x" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 910 | show "C x y" sorry | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 911 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 912 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 913 | txt_raw {*\end{minipage}\begin{minipage}{0.5\textwidth}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 914 | (*<*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 915 | next | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 916 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 917 | have "\<And>x y. A x \<Longrightarrow> B y \<Longrightarrow> C x y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 918 | proof - | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 919 | fix y assume "B y" | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 920 | fix x | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 921 | show "C x y" sorry | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 922 | qed | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 923 | (*<*) | 
| 40964 | 924 | end | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 925 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 926 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 927 | text_raw {*\end{minipage}*}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 928 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 929 | text {*
 | 
| 29735 | 930 | \medskip\noindent Such ``peephole optimizations'' of Isar texts are | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 931 | practically important to improve readability, by rearranging | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 932 | contexts elements according to the natural flow of reasoning in the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 933 | body, while still observing the overall scoping rules. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 934 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 935 | \medskip This illustrates the basic idea of structured proof | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 936 | processing in Isar. The main mechanisms are based on natural | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 937 | deduction rule composition within the Pure framework. In | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 938 | particular, there are no direct operations on goal states within the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 939 | proof body. Moreover, there is no hidden automated reasoning | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 940 | involved, just plain unification. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 941 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 942 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 943 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 944 | subsection {* Calculational reasoning \label{sec:framework-calc} *}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 945 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 946 | text {*
 | 
| 29735 | 947 | The existing Isar infrastructure is sufficiently flexible to support | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 948 | calculational reasoning (chains of transitivity steps) as derived | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 949 | concept. The generic proof elements introduced below depend on | 
| 29732 | 950 |   rules declared as @{attribute trans} in the context.  It is left to
 | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 951 | the object-logic to provide a suitable rule collection for mixed | 
| 29732 | 952 |   relations of @{text "="}, @{text "<"}, @{text "\<le>"}, @{text "\<subset>"},
 | 
| 953 |   @{text "\<subseteq>"} etc.  Due to the flexibility of rule composition
 | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 954 |   (\secref{sec:framework-resolution}), substitution of equals by
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 955 | equals is covered as well, even substitution of inequalities | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 956 |   involving monotonicity conditions; see also \cite[\S6]{Wenzel-PhD}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 957 |   and \cite{Bauer-Wenzel:2001}.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 958 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 959 | The generic calculational mechanism is based on the observation that | 
| 29735 | 960 |   rules such as @{text "trans:"}~@{prop "x = y \<Longrightarrow> y = z \<Longrightarrow> x = z"}
 | 
| 961 | proceed from the premises towards the conclusion in a deterministic | |
| 962 | fashion. Thus we may reason in forward mode, feeding intermediate | |
| 963 | results into rules selected from the context. The course of | |
| 964 | reasoning is organized by maintaining a secondary fact called | |
| 965 |   ``@{fact calculation}'', apart from the primary ``@{fact this}''
 | |
| 966 | already provided by the Isar primitives. In the definitions below, | |
| 967 |   @{attribute OF} refers to @{inference resolution}
 | |
| 968 |   (\secref{sec:framework-resolution}) with multiple rule arguments,
 | |
| 969 |   and @{text "trans"} represents to a suitable rule from the context:
 | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 970 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 971 |   \begin{matharray}{rcl}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 972 |     @{command "also"}@{text "\<^sub>0"} & \equiv & @{command "note"}~@{text "calculation = this"} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 973 |     @{command "also"}@{text "\<^sub>n\<^sub>+\<^sub>1"} & \equiv & @{command "note"}~@{text "calculation = trans [OF calculation this]"} \\[0.5ex]
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 974 |     @{command "finally"} & \equiv & @{command "also"}~@{command "from"}~@{text calculation} \\
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 975 |   \end{matharray}
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 976 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 977 | \noindent The start of a calculation is determined implicitly in the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 978 |   text: here @{command also} sets @{fact calculation} to the current
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 979 |   result; any subsequent occurrence will update @{fact calculation} by
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 980 | combination with the next result and a transitivity rule. The | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 981 |   calculational sequence is concluded via @{command finally}, where
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 982 | the final result is exposed for use in a concluding claim. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 983 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 984 |   Here is a canonical proof pattern, using @{command have} to
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 985 | establish the intermediate results: | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 986 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 987 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 988 | (*<*) | 
| 40964 | 989 | notepad | 
| 990 | begin | |
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 991 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 992 | have "a = b" sorry | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 993 | also have "\<dots> = c" sorry | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 994 | also have "\<dots> = d" sorry | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 995 | finally have "a = d" . | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 996 | (*<*) | 
| 40964 | 997 | end | 
| 29729 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 998 | (*>*) | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 999 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1000 | text {*
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1001 |   \noindent The term ``@{text "\<dots>"}'' above is a special abbreviation
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1002 | provided by the Isabelle/Isar syntax layer: it statically refers to | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1003 | the right-hand side argument of the previous statement given in the | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1004 | text. Thus it happens to coincide with relevant sub-expressions in | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1005 | the calculational chain, but the exact correspondence is dependent | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1006 | on the transitivity rules being involved. | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1007 | |
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1008 |   \medskip Symmetry rules such as @{prop "x = y \<Longrightarrow> y = x"} are like
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1009 | transitivities with only one premise. Isar maintains a separate | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1010 |   rule collection declared via the @{attribute sym} attribute, to be
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1011 |   used in fact expressions ``@{text "a [symmetric]"}'', or single-step
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1012 |   proofs ``@{command assume}~@{text "x = y"}~@{command then}~@{command
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1013 |   have}~@{text "y = x"}~@{command ".."}''.
 | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1014 | *} | 
| 
c2e926455fcc
more on Isar framework -- mostly from Trybulec Festschrift;
 wenzelm parents: 
29721diff
changeset | 1015 | |
| 29742 | 1016 | end |