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