51 qed |
49 qed |
52 qed |
50 qed |
53 qed |
51 qed |
54 qed |
52 qed |
55 |
53 |
56 text {* |
54 text {* Isar provides several ways to fine-tune the reasoning, |
57 Isar provides several ways to fine-tune the reasoning, avoiding |
55 avoiding excessive detail. Several abbreviated language elements |
58 excessive detail. Several abbreviated language elements are |
56 are available, enabling the writer to express proofs in a more |
59 available, enabling the writer to express proofs in a more concise |
57 concise way, even without referring to any automated proof tools |
60 way, even without referring to any automated proof tools yet. |
58 yet. |
61 |
59 |
62 First of all, proof by assumption may be abbreviated as a single |
60 First of all, proof by assumption may be abbreviated as a single |
63 dot. |
61 dot. *} |
64 *} |
|
65 |
62 |
66 lemma "A --> A" |
63 lemma "A --> A" |
67 proof |
64 proof |
68 assume A |
65 assume A |
69 show A by fact+ |
66 show A by fact+ |
70 qed |
67 qed |
71 |
68 |
72 text {* |
69 text {* In fact, concluding any (sub-)proof already involves solving |
73 In fact, concluding any (sub-)proof already involves solving any |
70 any remaining goals by assumption\footnote{This is not a completely |
74 remaining goals by assumption\footnote{This is not a completely |
|
75 trivial operation, as proof by assumption may involve full |
71 trivial operation, as proof by assumption may involve full |
76 higher-order unification.}. Thus we may skip the rather vacuous |
72 higher-order unification.}. Thus we may skip the rather vacuous |
77 body of the above proof as well. |
73 body of the above proof as well. *} |
78 *} |
|
79 |
74 |
80 lemma "A --> A" |
75 lemma "A --> A" |
81 proof |
76 proof |
82 qed |
77 qed |
83 |
78 |
84 text {* |
79 text {* Note that the \isacommand{proof} command refers to the @{text |
85 Note that the \isacommand{proof} command refers to the @{text rule} |
80 rule} method (without arguments) by default. Thus it implicitly |
86 method (without arguments) by default. Thus it implicitly applies a |
81 applies a single rule, as determined from the syntactic form of the |
87 single rule, as determined from the syntactic form of the statements |
82 statements involved. The \isacommand{by} command abbreviates any |
88 involved. The \isacommand{by} command abbreviates any proof with |
83 proof with empty body, so the proof may be further pruned. *} |
89 empty body, so the proof may be further pruned. |
|
90 *} |
|
91 |
84 |
92 lemma "A --> A" |
85 lemma "A --> A" |
93 by rule |
86 by rule |
94 |
87 |
95 text {* |
88 text {* Proof by a single rule may be abbreviated as double-dot. *} |
96 Proof by a single rule may be abbreviated as double-dot. |
|
97 *} |
|
98 |
89 |
99 lemma "A --> A" .. |
90 lemma "A --> A" .. |
100 |
91 |
101 text {* |
92 text {* Thus we have arrived at an adequate representation of the |
102 Thus we have arrived at an adequate representation of the proof of a |
93 proof of a tautology that holds by a single standard |
103 tautology that holds by a single standard rule.\footnote{Apparently, |
94 rule.\footnote{Apparently, the rule here is implication |
104 the rule here is implication introduction.} |
95 introduction.} *} |
105 *} |
96 |
106 |
97 text {* Let us also reconsider @{text K}. Its statement is composed |
107 text {* |
98 of iterated connectives. Basic decomposition is by a single rule at |
108 Let us also reconsider @{text K}. Its statement is composed of |
99 a time, which is why our first version above was by nesting two |
109 iterated connectives. Basic decomposition is by a single rule at a |
|
110 time, which is why our first version above was by nesting two |
|
111 proofs. |
100 proofs. |
112 |
101 |
113 The @{text intro} proof method repeatedly decomposes a goal's |
102 The @{text intro} proof method repeatedly decomposes a goal's |
114 conclusion.\footnote{The dual method is @{text elim}, acting on a |
103 conclusion.\footnote{The dual method is @{text elim}, acting on a |
115 goal's premises.} |
104 goal's premises.} *} |
116 *} |
|
117 |
105 |
118 lemma "A --> B --> A" |
106 lemma "A --> B --> A" |
119 proof (intro impI) |
107 proof (intro impI) |
120 assume A |
108 assume A |
121 show A by fact |
109 show A by fact |
122 qed |
110 qed |
123 |
111 |
124 text {* |
112 text {* Again, the body may be collapsed. *} |
125 Again, the body may be collapsed. |
|
126 *} |
|
127 |
113 |
128 lemma "A --> B --> A" |
114 lemma "A --> B --> A" |
129 by (intro impI) |
115 by (intro impI) |
130 |
116 |
131 text {* |
117 text {* Just like @{text rule}, the @{text intro} and @{text elim} |
132 Just like @{text rule}, the @{text intro} and @{text elim} proof |
118 proof methods pick standard structural rules, in case no explicit |
133 methods pick standard structural rules, in case no explicit |
|
134 arguments are given. While implicit rules are usually just fine for |
119 arguments are given. While implicit rules are usually just fine for |
135 single rule application, this may go too far with iteration. Thus |
120 single rule application, this may go too far with iteration. Thus |
136 in practice, @{text intro} and @{text elim} would be typically |
121 in practice, @{text intro} and @{text elim} would be typically |
137 restricted to certain structures by giving a few rules only, e.g.\ |
122 restricted to certain structures by giving a few rules only, e.g.\ |
138 \isacommand{proof}~@{text "(intro impI allI)"} to strip implications |
123 \isacommand{proof}~@{text "(intro impI allI)"} to strip implications |
140 |
125 |
141 Such well-tuned iterated decomposition of certain structures is the |
126 Such well-tuned iterated decomposition of certain structures is the |
142 prime application of @{text intro} and @{text elim}. In contrast, |
127 prime application of @{text intro} and @{text elim}. In contrast, |
143 terminal steps that solve a goal completely are usually performed by |
128 terminal steps that solve a goal completely are usually performed by |
144 actual automated proof methods (such as \isacommand{by}~@{text |
129 actual automated proof methods (such as \isacommand{by}~@{text |
145 blast}. |
130 blast}. *} |
146 *} |
|
147 |
131 |
148 |
132 |
149 subsection {* Variations of backward vs.\ forward reasoning *} |
133 subsection {* Variations of backward vs.\ forward reasoning *} |
150 |
134 |
151 text {* |
135 text {* Certainly, any proof may be performed in backward-style only. |
152 Certainly, any proof may be performed in backward-style only. On |
136 On the other hand, small steps of reasoning are often more naturally |
153 the other hand, small steps of reasoning are often more naturally |
|
154 expressed in forward-style. Isar supports both backward and forward |
137 expressed in forward-style. Isar supports both backward and forward |
155 reasoning as a first-class concept. In order to demonstrate the |
138 reasoning as a first-class concept. In order to demonstrate the |
156 difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}. |
139 difference, we consider several proofs of @{text "A \<and> B \<longrightarrow> B \<and> A"}. |
157 |
140 |
158 The first version is purely backward. |
141 The first version is purely backward. *} |
159 *} |
|
160 |
142 |
161 lemma "A & B --> B & A" |
143 lemma "A & B --> B & A" |
162 proof |
144 proof |
163 assume "A & B" |
145 assume "A & B" |
164 show "B & A" |
146 show "B & A" |
166 show B by (rule conjunct2) fact |
148 show B by (rule conjunct2) fact |
167 show A by (rule conjunct1) fact |
149 show A by (rule conjunct1) fact |
168 qed |
150 qed |
169 qed |
151 qed |
170 |
152 |
171 text {* |
153 text {* Above, the @{text "conjunct_1/2"} projection rules had to be |
172 Above, the @{text "conjunct_1/2"} projection rules had to be named |
154 named explicitly, since the goals @{text B} and @{text A} did not |
173 explicitly, since the goals @{text B} and @{text A} did not provide |
155 provide any structural clue. This may be avoided using |
174 any structural clue. This may be avoided using \isacommand{from} to |
156 \isacommand{from} to focus on the @{text "A \<and> B"} assumption as the |
175 focus on the @{text "A \<and> B"} assumption as the current facts, |
157 current facts, enabling the use of double-dot proofs. Note that |
176 enabling the use of double-dot proofs. Note that \isacommand{from} |
158 \isacommand{from} already does forward-chaining, involving the |
177 already does forward-chaining, involving the \name{conjE} rule here. |
159 @{text conjE} rule here. *} |
178 *} |
|
179 |
160 |
180 lemma "A & B --> B & A" |
161 lemma "A & B --> B & A" |
181 proof |
162 proof |
182 assume "A & B" |
163 assume "A & B" |
183 show "B & A" |
164 show "B & A" |
185 from `A & B` show B .. |
166 from `A & B` show B .. |
186 from `A & B` show A .. |
167 from `A & B` show A .. |
187 qed |
168 qed |
188 qed |
169 qed |
189 |
170 |
190 text {* |
171 text {* In the next version, we move the forward step one level |
191 In the next version, we move the forward step one level upwards. |
172 upwards. Forward-chaining from the most recent facts is indicated |
192 Forward-chaining from the most recent facts is indicated by the |
173 by the \isacommand{then} command. Thus the proof of @{text "B \<and> A"} |
193 \isacommand{then} command. Thus the proof of @{text "B \<and> A"} from |
174 from @{text "A \<and> B"} actually becomes an elimination, rather than an |
194 @{text "A \<and> B"} actually becomes an elimination, rather than an |
|
195 introduction. The resulting proof structure directly corresponds to |
175 introduction. The resulting proof structure directly corresponds to |
196 that of the @{text conjE} rule, including the repeated goal |
176 that of the @{text conjE} rule, including the repeated goal |
197 proposition that is abbreviated as @{text ?thesis} below. |
177 proposition that is abbreviated as @{text ?thesis} below. *} |
198 *} |
|
199 |
178 |
200 lemma "A & B --> B & A" |
179 lemma "A & B --> B & A" |
201 proof |
180 proof |
202 assume "A & B" |
181 assume "A & B" |
203 then show "B & A" |
182 then show "B & A" |
205 assume B A |
184 assume B A |
206 then show ?thesis .. -- {* rule @{text conjI} of @{text "B \<and> A"} *} |
185 then show ?thesis .. -- {* rule @{text conjI} of @{text "B \<and> A"} *} |
207 qed |
186 qed |
208 qed |
187 qed |
209 |
188 |
210 text {* |
189 text {* In the subsequent version we flatten the structure of the main |
211 In the subsequent version we flatten the structure of the main body |
190 body by doing forward reasoning all the time. Only the outermost |
212 by doing forward reasoning all the time. Only the outermost |
191 decomposition step is left as backward. *} |
213 decomposition step is left as backward. |
|
214 *} |
|
215 |
192 |
216 lemma "A & B --> B & A" |
193 lemma "A & B --> B & A" |
217 proof |
194 proof |
218 assume "A & B" |
195 assume "A & B" |
219 from `A & B` have A .. |
196 from `A & B` have A .. |
220 from `A & B` have B .. |
197 from `A & B` have B .. |
221 from `B` `A` show "B & A" .. |
198 from `B` `A` show "B & A" .. |
222 qed |
199 qed |
223 |
200 |
224 text {* |
201 text {* We can still push forward-reasoning a bit further, even at the |
225 We can still push forward-reasoning a bit further, even at the risk |
202 risk of getting ridiculous. Note that we force the initial proof |
226 of getting ridiculous. Note that we force the initial proof step to |
203 step to do nothing here, by referring to the ``-'' proof method. *} |
227 do nothing here, by referring to the ``-'' proof method. |
|
228 *} |
|
229 |
204 |
230 lemma "A & B --> B & A" |
205 lemma "A & B --> B & A" |
231 proof - |
206 proof - |
232 { |
207 { |
233 assume "A & B" |
208 assume "A & B" |
234 from `A & B` have A .. |
209 from `A & B` have A .. |
235 from `A & B` have B .. |
210 from `A & B` have B .. |
236 from `B` `A` have "B & A" .. |
211 from `B` `A` have "B & A" .. |
237 } |
212 } |
238 then show ?thesis .. -- {* rule \name{impI} *} |
213 then show ?thesis .. -- {* rule @{text impI} *} |
239 qed |
214 qed |
240 |
215 |
241 text {* |
216 text {* \medskip With these examples we have shifted through a whole |
242 \medskip With these examples we have shifted through a whole range |
217 range from purely backward to purely forward reasoning. Apparently, |
243 from purely backward to purely forward reasoning. Apparently, in |
218 in the extreme ends we get slightly ill-structured proofs, which |
244 the extreme ends we get slightly ill-structured proofs, which also |
219 also require much explicit naming of either rules (backward) or |
245 require much explicit naming of either rules (backward) or local |
220 local facts (forward). |
246 facts (forward). |
|
247 |
221 |
248 The general lesson learned here is that good proof style would |
222 The general lesson learned here is that good proof style would |
249 achieve just the \emph{right} balance of top-down backward |
223 achieve just the \emph{right} balance of top-down backward |
250 decomposition, and bottom-up forward composition. In general, there |
224 decomposition, and bottom-up forward composition. In general, there |
251 is no single best way to arrange some pieces of formal reasoning, of |
225 is no single best way to arrange some pieces of formal reasoning, of |
252 course. Depending on the actual applications, the intended audience |
226 course. Depending on the actual applications, the intended audience |
253 etc., rules (and methods) on the one hand vs.\ facts on the other |
227 etc., rules (and methods) on the one hand vs.\ facts on the other |
254 hand have to be emphasized in an appropriate way. This requires the |
228 hand have to be emphasized in an appropriate way. This requires the |
255 proof writer to develop good taste, and some practice, of course. |
229 proof writer to develop good taste, and some practice, of course. *} |
256 *} |
230 |
257 |
231 text {* For our example the most appropriate way of reasoning is |
258 text {* |
232 probably the middle one, with conjunction introduction done after |
259 For our example the most appropriate way of reasoning is probably |
233 elimination. *} |
260 the middle one, with conjunction introduction done after |
|
261 elimination. |
|
262 *} |
|
263 |
234 |
264 lemma "A & B --> B & A" |
235 lemma "A & B --> B & A" |
265 proof |
236 proof |
266 assume "A & B" |
237 assume "A & B" |
267 then show "B & A" |
238 then show "B & A" |
299 next |
267 next |
300 assume P show P by fact |
268 assume P show P by fact |
301 qed |
269 qed |
302 qed |
270 qed |
303 |
271 |
304 text {* |
272 text {* Case splits are \emph{not} hardwired into the Isar language as |
305 Case splits are \emph{not} hardwired into the Isar language as a |
273 a special feature. The \isacommand{next} command used to separate |
306 special feature. The \isacommand{next} command used to separate the |
274 the cases above is just a short form of managing block structure. |
307 cases above is just a short form of managing block structure. |
|
308 |
275 |
309 \medskip In general, applying proof methods may split up a goal into |
276 \medskip In general, applying proof methods may split up a goal into |
310 separate ``cases'', i.e.\ new subgoals with individual local |
277 separate ``cases'', i.e.\ new subgoals with individual local |
311 assumptions. The corresponding proof text typically mimics this by |
278 assumptions. The corresponding proof text typically mimics this by |
312 establishing results in appropriate contexts, separated by blocks. |
279 establishing results in appropriate contexts, separated by blocks. |
334 show P by fact |
300 show P by fact |
335 show P by fact |
301 show P by fact |
336 qed |
302 qed |
337 qed |
303 qed |
338 |
304 |
339 text {* |
305 text {* Again, the rather vacuous body of the proof may be collapsed. |
340 Again, the rather vacuous body of the proof may be collapsed. Thus |
306 Thus the case analysis degenerates into two assumption steps, which |
341 the case analysis degenerates into two assumption steps, which are |
307 are implicitly performed when concluding the single rule step of the |
342 implicitly performed when concluding the single rule step of the |
308 double-dot proof as follows. *} |
343 double-dot proof as follows. |
|
344 *} |
|
345 |
309 |
346 lemma "P | P --> P" |
310 lemma "P | P --> P" |
347 proof |
311 proof |
348 assume "P | P" |
312 assume "P | P" |
349 then show P .. |
313 then show P .. |
350 qed |
314 qed |
351 |
315 |
352 |
316 |
353 subsubsection {* A quantifier proof *} |
317 subsubsection {* A quantifier proof *} |
354 |
318 |
355 text {* |
319 text {* To illustrate quantifier reasoning, let us prove @{text |
356 To illustrate quantifier reasoning, let us prove @{text "(\<exists>x. P (f |
320 "(\<exists>x. P (f x)) \<longrightarrow> (\<exists>y. P y)"}. Informally, this holds because any |
357 x)) \<longrightarrow> (\<exists>y. P y)"}. Informally, this holds because any @{text a} |
321 @{text a} with @{text "P (f a)"} may be taken as a witness for the |
358 with @{text "P (f a)"} may be taken as a witness for the second |
322 second existential statement. |
359 existential statement. |
|
360 |
323 |
361 The first proof is rather verbose, exhibiting quite a lot of |
324 The first proof is rather verbose, exhibiting quite a lot of |
362 (redundant) detail. It gives explicit rules, even with some |
325 (redundant) detail. It gives explicit rules, even with some |
363 instantiation. Furthermore, we encounter two new language elements: |
326 instantiation. Furthermore, we encounter two new language elements: |
364 the \isacommand{fix} command augments the context by some new |
327 the \isacommand{fix} command augments the context by some new |
365 ``arbitrary, but fixed'' element; the \isacommand{is} annotation |
328 ``arbitrary, but fixed'' element; the \isacommand{is} annotation |
366 binds term abbreviations by higher-order pattern matching. |
329 binds term abbreviations by higher-order pattern matching. *} |
367 *} |
|
368 |
330 |
369 lemma "(EX x. P (f x)) --> (EX y. P y)" |
331 lemma "(EX x. P (f x)) --> (EX y. P y)" |
370 proof |
332 proof |
371 assume "EX x. P (f x)" |
333 assume "EX x. P (f x)" |
372 then show "EX y. P y" |
334 then show "EX y. P y" |
373 proof (rule exE) -- {* |
335 proof (rule exE) -- {* |
374 rule \name{exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$} |
336 rule @{text exE}: \smash{$\infer{B}{\ex x A(x) & \infer*{B}{[A(x)]_x}}$} |
375 *} |
337 *} |
376 fix a |
338 fix a |
377 assume "P (f a)" (is "P ?witness") |
339 assume "P (f a)" (is "P ?witness") |
378 then show ?thesis by (rule exI [of P ?witness]) |
340 then show ?thesis by (rule exI [of P ?witness]) |
379 qed |
341 qed |
380 qed |
342 qed |
381 |
343 |
382 text {* |
344 text {* While explicit rule instantiation may occasionally improve |
383 While explicit rule instantiation may occasionally improve |
|
384 readability of certain aspects of reasoning, it is usually quite |
345 readability of certain aspects of reasoning, it is usually quite |
385 redundant. Above, the basic proof outline gives already enough |
346 redundant. Above, the basic proof outline gives already enough |
386 structural clues for the system to infer both the rules and their |
347 structural clues for the system to infer both the rules and their |
387 instances (by higher-order unification). Thus we may as well prune |
348 instances (by higher-order unification). Thus we may as well prune |
388 the text as follows. |
349 the text as follows. *} |
389 *} |
|
390 |
350 |
391 lemma "(EX x. P (f x)) --> (EX y. P y)" |
351 lemma "(EX x. P (f x)) --> (EX y. P y)" |
392 proof |
352 proof |
393 assume "EX x. P (f x)" |
353 assume "EX x. P (f x)" |
394 then show "EX y. P y" |
354 then show "EX y. P y" |
397 assume "P (f a)" |
357 assume "P (f a)" |
398 then show ?thesis .. |
358 then show ?thesis .. |
399 qed |
359 qed |
400 qed |
360 qed |
401 |
361 |
402 text {* |
362 text {* Explicit @{text \<exists>}-elimination as seen above can become quite |
403 Explicit @{text \<exists>}-elimination as seen above can become quite |
|
404 cumbersome in practice. The derived Isar language element |
363 cumbersome in practice. The derived Isar language element |
405 ``\isakeyword{obtain}'' provides a more handsome way to do |
364 ``\isakeyword{obtain}'' provides a more handsome way to do |
406 generalized existence reasoning. |
365 generalized existence reasoning. *} |
407 *} |
|
408 |
366 |
409 lemma "(EX x. P (f x)) --> (EX y. P y)" |
367 lemma "(EX x. P (f x)) --> (EX y. P y)" |
410 proof |
368 proof |
411 assume "EX x. P (f x)" |
369 assume "EX x. P (f x)" |
412 then obtain a where "P (f a)" .. |
370 then obtain a where "P (f a)" .. |
413 then show "EX y. P y" .. |
371 then show "EX y. P y" .. |
414 qed |
372 qed |
415 |
373 |
416 text {* |
374 text {* Technically, \isakeyword{obtain} is similar to |
417 Technically, \isakeyword{obtain} is similar to \isakeyword{fix} and |
375 \isakeyword{fix} and \isakeyword{assume} together with a soundness |
418 \isakeyword{assume} together with a soundness proof of the |
376 proof of the elimination involved. Thus it behaves similar to any |
419 elimination involved. Thus it behaves similar to any other forward |
377 other forward proof element. Also note that due to the nature of |
420 proof element. Also note that due to the nature of general |
378 general existence reasoning involved here, any result exported from |
421 existence reasoning involved here, any result exported from the |
379 the context of an \isakeyword{obtain} statement may \emph{not} refer |
422 context of an \isakeyword{obtain} statement may \emph{not} refer to |
380 to the parameters introduced there. *} |
423 the parameters introduced there. |
|
424 *} |
|
425 |
|
426 |
381 |
427 |
382 |
428 subsubsection {* Deriving rules in Isabelle *} |
383 subsubsection {* Deriving rules in Isabelle *} |
429 |
384 |
430 text {* |
385 text {* We derive the conjunction elimination rule from the |
431 We derive the conjunction elimination rule from the corresponding |
386 corresponding projections. The proof is quite straight-forward, |
432 projections. The proof is quite straight-forward, since |
387 since Isabelle/Isar supports non-atomic goals and assumptions fully |
433 Isabelle/Isar supports non-atomic goals and assumptions fully |
388 transparently. *} |
434 transparently. |
|
435 *} |
|
436 |
389 |
437 theorem conjE: "A & B ==> (A ==> B ==> C) ==> C" |
390 theorem conjE: "A & B ==> (A ==> B ==> C) ==> C" |
438 proof - |
391 proof - |
439 assume "A & B" |
392 assume "A & B" |
440 assume r: "A ==> B ==> C" |
393 assume r: "A ==> B ==> C" |