|
1 (***************************************************************** |
|
2 |
|
3 Isabelle Tutorial |
|
4 ----------------- |
|
5 |
|
6 27th May 2009, Beijing |
|
7 |
|
8 This file contains most of the material that will be covered in the |
|
9 tutorial. The file can be "stepped through"; though it contains much |
|
10 commented code (purple text). |
|
11 |
|
12 *) |
|
13 |
|
14 (***************************************************************** |
|
15 |
|
16 Proof General |
|
17 ------------- |
|
18 |
|
19 Proof General is the front end for Isabelle. |
|
20 |
|
21 Proof General "paints" blue the part of the proof-script that has already |
|
22 been processed by Isabelle. You can advance or retract the "frontier" of |
|
23 this processed part using the "Next" and "Undo" buttons in the |
|
24 menubar. "Goto" will process everything up to the current cursor position, |
|
25 or retract if the cursor is inside the blue part. The key-combination |
|
26 control-c control-return is a short-cut for the "Goto"-operation. |
|
27 |
|
28 Proof General gives feedback inside the "Response" and "Goals" buffers. |
|
29 The response buffer will contain warning messages (in yellow) and |
|
30 error messages (in red). Warning messages can generally be ignored. Error |
|
31 messages must be dealt with in order to advance the proof script. The goal |
|
32 buffer shows which goals need to be proved. The sole idea of interactive |
|
33 theorem proving is to get the message "No subgoals." ;o) |
|
34 |
|
35 *) |
|
36 |
|
37 (***************************************************************** |
|
38 |
|
39 X-Symbols |
|
40 --------- |
|
41 |
|
42 X-symbols provide a nice way to input non-ascii characters. For example: |
|
43 |
|
44 \<forall>, \<exists>, \<Down>, \<sharp>, \<And>, \<Gamma>, \<times>, \<noteq>, \<in>, \<subseteq>, \<dots> |
|
45 |
|
46 They need to be input via the combination \<name-of-x-symbol> |
|
47 where name-of-x-symbol coincides with the usual latex name of |
|
48 that symbol. |
|
49 |
|
50 However, there are some handy short-cuts for frequently used X-symbols. |
|
51 For example |
|
52 |
|
53 [| \<dots> \<lbrakk> |
|
54 |] \<dots> \<rbrakk> |
|
55 ==> \<dots> \<Longrightarrow> |
|
56 => \<dots> \<Rightarrow> |
|
57 --> \<dots> \<longrightarrow> |
|
58 /\ \<dots> \<and> |
|
59 \/ \<dots> \<or> |
|
60 |-> \<dots> \<mapsto> |
|
61 =_ \<dots> \<equiv> |
|
62 |
|
63 *) |
|
64 |
|
65 (***************************************************************** |
|
66 |
|
67 Theories |
|
68 -------- |
|
69 |
|
70 Every Isabelle proof-script needs to have a name and must import |
|
71 some pre-existing theory. This will normally be the theory Main. |
|
72 |
|
73 *) |
|
74 |
|
75 theory Lec1 |
|
76 imports "Main" |
|
77 begin |
|
78 |
|
79 text {***************************************************************** |
|
80 |
|
81 Types |
|
82 ----- |
|
83 |
|
84 Isabelle is based on types including some polymorphism. It also includes |
|
85 some overloading, which means that sometimes explicit type annotations |
|
86 need to be given. |
|
87 |
|
88 - Base types include: nat, bool, string |
|
89 |
|
90 - Type formers include: 'a list, ('a\<times>'b), 'c set |
|
91 |
|
92 - Type variables are written like in ML with an apostrophe: 'a, 'b, \<dots> |
|
93 |
|
94 Types known to Isabelle can be queried using: |
|
95 |
|
96 *} |
|
97 |
|
98 typ nat |
|
99 typ bool |
|
100 typ string (* the type for alpha-equated lambda-terms *) |
|
101 typ "('a \<times> 'b)" (* pair type *) |
|
102 typ "'c set" (* set type *) |
|
103 typ "'a list" (* list type *) |
|
104 typ "nat \<Rightarrow> bool" (* the type of functions from nat to bool *) |
|
105 |
|
106 (* These give errors: *) |
|
107 (*typ boolean *) |
|
108 (*typ set*) |
|
109 |
|
110 text {***************************************************************** |
|
111 |
|
112 Terms |
|
113 ----- |
|
114 |
|
115 Every term in Isabelle needs to be well-typed (however they can have |
|
116 polymorphic type). Whether a term is accepted can be queried using: |
|
117 |
|
118 *} |
|
119 |
|
120 term c (* a variable of polymorphic type *) |
|
121 term "1::nat" (* the constant 1 in natural numbers *) |
|
122 term 1 (* the constant 1 with polymorphic type *) |
|
123 term "{1, 2, 3::nat}" (* the set containing natural numbers 1, 2 and 3 *) |
|
124 term "[1, 2, 3]" (* the list containing the polymorphic numbers 1, 2 and 3 *) |
|
125 term "(True, ''c'')" (* a pair consisting of true and the string "c" *) |
|
126 term "Suc 0" (* successor of 0, in other words 1 *) |
|
127 |
|
128 text {* |
|
129 Isabelle provides some useful colour feedback about what are constants (in black), |
|
130 free variables (in blue) and bound variables (in green). *} |
|
131 |
|
132 term "True" (* a constant that is defined in HOL *) |
|
133 term "true" (* not recognised as a constant, therefore it is interpreted as a variable *) |
|
134 term "\<forall>x. P x" (* x is bound, P is free *) |
|
135 |
|
136 text {* Every formula in Isabelle needs to have type bool *} |
|
137 |
|
138 term "True" |
|
139 term "True \<and> False" |
|
140 term "True \<or> B" |
|
141 term "{1,2,3} = {3,2,1}" |
|
142 term "\<forall>x. P x" |
|
143 term "A \<longrightarrow> B" |
|
144 |
|
145 text {* |
|
146 When working with Isabelle, one is confronted with an object logic (that is HOL) |
|
147 and Isabelle's meta-logic (called Pure). Sometimes one has to pay attention |
|
148 to this fact. *} |
|
149 |
|
150 term "A \<longrightarrow> B" (* versus *) |
|
151 term "A \<Longrightarrow> B" |
|
152 |
|
153 term "\<forall>x. P x" (* versus *) |
|
154 term "\<And>x. P x" |
|
155 |
|
156 term "A \<Longrightarrow> B \<Longrightarrow> C" (* is synonymous with *) |
|
157 term "\<lbrakk>A; B\<rbrakk> \<Longrightarrow> C" |
|
158 |
|
159 text {***************************************************************** |
|
160 |
|
161 Inductive Definitions: The even predicate |
|
162 ----------------------------------------- |
|
163 |
|
164 Inductive definitions start with the keyword "inductive" and a predicate name. |
|
165 One also needs to provide a type for the predicate. Clauses of the inductive |
|
166 predicate are introduced by "where" and more than two clauses need to be |
|
167 separated by "|". Optionally one can give a name to each clause and indicate |
|
168 that it should be added to the hints database ("[intro]"). A typical clause |
|
169 has some premises and a conclusion. This is written in Isabelle as: |
|
170 |
|
171 "premise \<Longrightarrow> conclusion" |
|
172 "\<lbrakk>premise1; premise2; \<dots> \<rbrakk> \<Longrightarrow> conclusion" |
|
173 |
|
174 If no premise is present, then one just writes |
|
175 |
|
176 "conclusion" |
|
177 |
|
178 Below we define the even predicate for natural numbers. |
|
179 |
|
180 *} |
|
181 |
|
182 inductive |
|
183 even :: "nat \<Rightarrow> bool" |
|
184 where |
|
185 eZ[intro]: "even 0" |
|
186 | eSS[intro]: "even n \<Longrightarrow> even (Suc (Suc n))" |
|
187 |
|
188 text {***************************************************************** |
|
189 |
|
190 Theorems |
|
191 -------- |
|
192 |
|
193 A central concept in Isabelle is that of theorems. Isabelle's theorem |
|
194 database can be queried using |
|
195 |
|
196 *} |
|
197 |
|
198 thm eZ |
|
199 thm eSS |
|
200 thm conjI |
|
201 thm conjunct1 |
|
202 |
|
203 text {* |
|
204 Notice that theorems usually contain schematic variables (e.g. ?P, ?Q, \<dots>). |
|
205 These schematic variables can be substituted with any term (of the right type |
|
206 of course). It is sometimes useful to suppress the "?" from the schematic |
|
207 variables. This can be achieved by using the attribute "[no_vars]". *} |
|
208 |
|
209 thm eZ[no_vars] |
|
210 thm eSS[no_vars] |
|
211 thm conjI[no_vars] |
|
212 thm conjunct1[no_vars] |
|
213 |
|
214 |
|
215 text {* |
|
216 When defining the predicate eval, Isabelle provides us automatically with |
|
217 the following theorems that state how the even predicate can be introduced |
|
218 and what constitutes an induction over the predicate even. *} |
|
219 |
|
220 thm even.intros[no_vars] |
|
221 thm even.induct[no_vars] |
|
222 |
|
223 text {***************************************************************** |
|
224 |
|
225 Lemma / Theorem / Corollary Statements |
|
226 -------------------------------------- |
|
227 |
|
228 Whether to use lemma, theorem or corollary makes no semantic difference |
|
229 in Isabelle. A lemma starts with "lemma" and consists of a statement |
|
230 ("shows \<dots>") and optionally a lemma name, some type-information for |
|
231 variables ("fixes \<dots>") and some assumptions ("assumes \<dots>"). Lemmas |
|
232 also need to have a proof, but ignore this 'detail' for the moment. |
|
233 |
|
234 *} |
|
235 |
|
236 lemma even_double: |
|
237 shows "even (2 * n)" |
|
238 by (induct n) (auto) |
|
239 |
|
240 lemma even_add_n_m: |
|
241 assumes a: "even n" |
|
242 and b: "even m" |
|
243 shows "even (n + m)" |
|
244 using a b by (induct) (auto) |
|
245 |
|
246 lemma neutral_element: |
|
247 fixes x::"nat" |
|
248 shows "x + 0 = x" |
|
249 by simp |
|
250 |
|
251 text {***************************************************************** |
|
252 |
|
253 Isar Proofs |
|
254 ----------- |
|
255 |
|
256 Isar is a language for writing down proofs that can be understood by humans |
|
257 and by Isabelle. An Isar proof can be thought of as a sequence of 'stepping stones' |
|
258 that start with the assumptions and lead to the goal to be established. Every such |
|
259 stepping stone is introduced by "have" followed by the statement of the stepping |
|
260 stone. An exception is the goal to be proved, which need to be introduced with "show". |
|
261 |
|
262 Since proofs usually do not proceed in a linear fashion, a label can be given |
|
263 to each stepping stone and then used later to refer to this stepping stone |
|
264 ("using"). |
|
265 |
|
266 Each stepping stone (or have-statement) needs to have a justification. The |
|
267 simplest justification is "sorry" which admits any stepping stone, even false |
|
268 ones (this is good during the development of proofs). Assumption can be |
|
269 "justified" using "by fact". Derived facts can be justified using |
|
270 |
|
271 - by simp (* simplification *) |
|
272 - by auto (* proof search and simplification *) |
|
273 - by blast (* only proof search *) |
|
274 |
|
275 If facts or lemmas are needed in order to justify a have-statement, then |
|
276 one can feed these facts into the proof by using "using label \<dots>" or |
|
277 "using theorem-name \<dots>". More than one label at the time is allowed. |
|
278 |
|
279 Induction proofs in Isar are set up by indicating over which predicate(s) |
|
280 the induction proceeds ("using a b") followed by the command "proof (induct)". |
|
281 In this way, Isabelle uses default settings for which induction should |
|
282 be performed. These default settings can be overridden by giving more |
|
283 information, like the variable over which a structural induction should |
|
284 proceed, or a specific induction principle such as well-founded inductions. |
|
285 |
|
286 After the induction is set up, the proof proceeds by cases. In Isar these |
|
287 cases can be given in any order, but must be started with "case" and the |
|
288 name of the case, and optionally some legible names for the variables |
|
289 referred to inside the case. |
|
290 |
|
291 The possible case-names can be found by looking inside the menu "Isabelle -> |
|
292 Show me -> cases". In each "case", we need to establish a statement introduced |
|
293 by "show". Once this has been done, the next case can be started using "next". |
|
294 When all cases are completed, the proof can be finished using "qed". |
|
295 |
|
296 This means a typical induction proof has the following pattern |
|
297 |
|
298 proof (induct) |
|
299 case \<dots> |
|
300 \<dots> |
|
301 show \<dots> |
|
302 next |
|
303 case \<dots> |
|
304 \<dots> |
|
305 show \<dots> |
|
306 \<dots> |
|
307 qed |
|
308 |
|
309 The four lemmas are by induction on the predicate machines. All proofs establish |
|
310 the same property, namely a transitivity rule for machines. The complete Isar |
|
311 proofs are given for the first three proofs. The point of these three proofs is |
|
312 that each proof increases the readability for humans. |
|
313 |
|
314 *} |
|
315 |
|
316 text {***************************************************************** |
|
317 |
|
318 1.) EXERCISE |
|
319 ------------ |
|
320 |
|
321 Remove the sorries in the proof below and fill in the correct |
|
322 justifications. |
|
323 *} |
|
324 |
|
325 lemma |
|
326 shows "even (2 * n)" |
|
327 proof (induct n) |
|
328 case 0 |
|
329 show "even (2 * 0)" sorry |
|
330 next |
|
331 case (Suc n) |
|
332 have ih: "even (2 * n)" by fact |
|
333 have eq: "2 * (Suc n) = Suc (Suc (2 * n))" sorry |
|
334 have h1: "even (Suc (Suc (2 * n)))" sorry |
|
335 show "even (2 * (Suc n))" sorry |
|
336 qed |
|
337 |
|
338 |
|
339 text {***************************************************************** |
|
340 |
|
341 2.) EXERCISE |
|
342 ------------ |
|
343 |
|
344 Remove the sorries in the proof below and fill in the correct |
|
345 justifications. |
|
346 *} |
|
347 |
|
348 lemma |
|
349 shows "even (n + n)" |
|
350 proof (induct n) |
|
351 case 0 |
|
352 show "even (0 + 0)" sorry |
|
353 next |
|
354 case (Suc n) |
|
355 have ih: "even (n + n)" by fact |
|
356 have eq: "(Suc n) + (Suc n) = Suc (Suc (n + n))" sorry |
|
357 have a: "even (Suc (Suc (n + n)))" sorry |
|
358 show "even ((Suc n) + (Suc n))" sorry |
|
359 qed |
|
360 |
|
361 text {* |
|
362 Just like gotos in the Basic programming language, labels can reduce |
|
363 the readability of proofs. Therefore one can use in Isar the notation |
|
364 "then have" in order to feed a have-statement to the proof of |
|
365 the next have-statement. In the proof below this has been used |
|
366 in order to get rid of the labels a. |
|
367 *} |
|
368 |
|
369 lemma even_twice: |
|
370 shows "even (n + n)" |
|
371 proof (induct n) |
|
372 case 0 |
|
373 show "even (0 + 0)" by auto |
|
374 next |
|
375 case (Suc n) |
|
376 have ih: "even (n + n)" by fact |
|
377 have eq: "(Suc n) + (Suc n) = Suc (Suc (n + n))" by simp |
|
378 have "even (Suc (Suc (n + n)))" using ih by auto |
|
379 then show "even ((Suc n) + (Suc n))" using eq by simp |
|
380 qed |
|
381 |
|
382 text {* |
|
383 The label eq cannot be got rid of in this way, because both |
|
384 facts are needed to prove the show-statement. We can still avoid the |
|
385 labels by feeding a sequence of facts into a proof using the chaining |
|
386 mechanism: |
|
387 |
|
388 have "statement1" \<dots> |
|
389 moreover |
|
390 have "statement2" \<dots> |
|
391 \<dots> |
|
392 moreover |
|
393 have "statementn" \<dots> |
|
394 ultimately have "statement" \<dots> |
|
395 |
|
396 In this chain, all "statementi" can be used in the proof of the final |
|
397 "statement". With this we can simplify our proof further to: |
|
398 *} |
|
399 |
|
400 lemma |
|
401 shows "even (n + n)" |
|
402 proof (induct n) |
|
403 case 0 |
|
404 show "even (0 + 0)" by auto |
|
405 next |
|
406 case (Suc n) |
|
407 have ih: "even (n + n)" by fact |
|
408 have "(Suc n) + (Suc n) = Suc (Suc (n + n))" by simp |
|
409 moreover |
|
410 have "even (Suc (Suc (n + n)))" using ih by auto |
|
411 ultimately show "even ((Suc n) + (Suc n))" by simp |
|
412 qed |
|
413 |
|
414 text {* |
|
415 While automatic proof procedures in Isabelle are not able to prove statements |
|
416 like "P = NP" assuming usual definitions for P and NP, they can automatically |
|
417 discharge the lemmas we just proved. For this we only have to set up the induction |
|
418 and auto will take care of the rest. This means we can write: |
|
419 *} |
|
420 |
|
421 lemma |
|
422 shows "even (2 * n)" |
|
423 by (induct n) (auto) |
|
424 |
|
425 lemma |
|
426 shows "even (n + n)" |
|
427 by (induct n) (auto) |
|
428 |
|
429 text {***************************************************************** |
|
430 |
|
431 Rule Inductions |
|
432 --------------- |
|
433 |
|
434 Inductive predicates are defined in terms of rules of the form |
|
435 |
|
436 prem1 \<dots> premn |
|
437 -------------- |
|
438 concl |
|
439 |
|
440 In Isabelle we write such rules as |
|
441 |
|
442 \<lbrakk>prem1; \<dots>; premn\<rbrakk> \<Longrightarrow> concl |
|
443 |
|
444 You can do induction over such rules according to the following |
|
445 scheme: |
|
446 |
|
447 1.) Assume the property for the premises. Assume the side-conditions. |
|
448 2.) Show the property for the conclusion. |
|
449 |
|
450 In Isabelle the corresponding theorem is called, for example |
|
451 *} |
|
452 |
|
453 thm even.induct |
|
454 |
|
455 text {***************************************************************** |
|
456 |
|
457 3.) EXERCISE |
|
458 ------------ |
|
459 |
|
460 Remove the sorries in the proof below and fill in the correct |
|
461 justifications. |
|
462 *} |
|
463 |
|
464 lemma even_add: |
|
465 assumes a: "even n" |
|
466 and b: "even m" |
|
467 shows "even (n + m)" |
|
468 using a b |
|
469 proof (induct) |
|
470 case eZ |
|
471 have as: "even m" by fact |
|
472 show "even (0 + m)" sorry |
|
473 next |
|
474 case (eSS n) |
|
475 have ih: "even m \<Longrightarrow> even (n + m)" by fact |
|
476 have as: "even m" by fact |
|
477 |
|
478 show "even (Suc (Suc n) + m)" sorry |
|
479 qed |
|
480 |
|
481 text {* |
|
482 |
|
483 Whenever a lemma is of the form |
|
484 |
|
485 lemma |
|
486 assumes "pred" |
|
487 and \<dots> |
|
488 shows "something" |
|
489 |
|
490 then a rule induction is generally appropriate (but |
|
491 not always). |
|
492 |
|
493 In case of even_add it is definitely *not* appropriate. |
|
494 Compare for example what you have to prove when you attempt |
|
495 an induction over natural numbers. |
|
496 |
|
497 *} |
|
498 |
|
499 lemma even_add_does_not_work: |
|
500 assumes a: "even n" |
|
501 and b: "even m" |
|
502 shows "even (n + m)" |
|
503 using a b |
|
504 proof (induct n rule: nat_induct) |
|
505 case 0 |
|
506 have "even m" by fact |
|
507 then show "even (0 + m)" by simp |
|
508 next |
|
509 case (Suc n) |
|
510 have ih: "\<lbrakk>even n; even m\<rbrakk> \<Longrightarrow> even (n + m)" by fact |
|
511 have as1: "even (Suc n)" by fact |
|
512 have as2: "even m" by fact |
|
513 |
|
514 show "even ((Suc n) + m)" oops |
|
515 |
|
516 text {***************************************************************** |
|
517 |
|
518 4.) EXERCISE (Last fact about even) |
|
519 ----------------------------------- |
|
520 |
|
521 Remove the sorries in the proof below and fill in the correct |
|
522 justifications. The following two lemmas should be useful |
|
523 *} |
|
524 |
|
525 thm even_twice[no_vars] |
|
526 thm even_add[no_vars] |
|
527 |
|
528 text {* |
|
529 Remember you can include lemmas in the "using" statement, |
|
530 just like other facts. For example: |
|
531 |
|
532 have "something" using even_twice by simp |
|
533 |
|
534 *} |
|
535 |
|
536 lemma even_mul: |
|
537 assumes a: "even n" |
|
538 shows "even (n * m)" |
|
539 using a |
|
540 proof (induct) |
|
541 case eZ |
|
542 show "even (0 * m)" by auto |
|
543 next |
|
544 case (eSS n) |
|
545 have as: "even n" by fact |
|
546 have ih: "even (n * m)" by fact |
|
547 |
|
548 show "even (Suc (Suc n) * m)" sorry |
|
549 qed |
|
550 |
|
551 text {***************************************************************** |
|
552 |
|
553 Definitions |
|
554 ----------- |
|
555 |
|
556 Often it is useful to define new concepts in terms of existsing |
|
557 concepts. In Isabelle you introduce definitions with the key- |
|
558 word "definition". For example |
|
559 |
|
560 *} |
|
561 |
|
562 definition |
|
563 divide :: "nat \<Rightarrow> nat \<Rightarrow> bool" ("_ DVD _" [100,100] 100) |
|
564 where |
|
565 "m DVD n = (\<exists>k. n = m * k)" |
|
566 |
|
567 text {* |
|
568 |
|
569 The annotation ("_ DVD _" [100,100] 100) introduces some fancier |
|
570 syntax. In this case we define this predicate as infix. The underscores |
|
571 correspond to the to arguments. The numbers stand for precedences. |
|
572 |
|
573 Once this definition is stated, it can be accessed as a normal |
|
574 theorem. For example |
|
575 *} |
|
576 |
|
577 thm divide_def |
|
578 |
|
579 text {* |
|
580 Below we prove the really last fact about even. Note |
|
581 how we deal with existential quantified formulas, where |
|
582 we have to use |
|
583 |
|
584 obtain \<dots> where \<dots> |
|
585 |
|
586 in order to obtain a witness with which we can reason further. |
|
587 *} |
|
588 |
|
589 lemma even_divide: |
|
590 assumes a: "even n" |
|
591 shows "2 DVD n" |
|
592 using a |
|
593 proof (induct) |
|
594 case eZ |
|
595 have "0 = 2 * (0::nat)" by simp |
|
596 then show "2 DVD 0" by (auto simp add: divide_def) |
|
597 next |
|
598 case (eSS n) |
|
599 have "2 DVD n" by fact |
|
600 then have "\<exists>k. n = 2 * k" by (simp add: divide_def) |
|
601 then obtain k where eq: "n = 2 * k" by (auto) |
|
602 have "Suc (Suc n) = 2 * (Suc k)" using eq by simp |
|
603 then have "\<exists>k. Suc (Suc n) = 2 * k" by blast |
|
604 then show "2 DVD (Suc (Suc n))" by (simp add: divide_def) |
|
605 qed |
|
606 |
|
607 text {***************************************************************** |
|
608 |
|
609 Function Definitions |
|
610 -------------------- |
|
611 |
|
612 Many functions over datatypes can be defined by recursion on the |
|
613 structure. For this purpose, Isabelle provides "fun". To use it one needs |
|
614 to give a name for the function, its type, optionally some pretty-syntax |
|
615 and then some equations defining the function. Like in "inductive", |
|
616 "fun" expects that more than one such equation is separated by "|". |
|
617 |
|
618 Below we define function iteration using function composition |
|
619 defined as "o". |
|
620 *} |
|
621 |
|
622 fun |
|
623 iter :: "('a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> ('a \<Rightarrow> 'a)" ("_ !! _") |
|
624 where |
|
625 "f !! 0 = (\<lambda>x. x)" |
|
626 | "f !! (Suc n) = (f !! n) o f" |
|
627 |
|
628 text {* |
|
629 5.) EXERCISE |
|
630 |
|
631 Prove the following property by induction over n. |
|
632 |
|
633 *} |
|
634 |
|
635 lemma |
|
636 shows "f !! (m + n) = (f !! m) o (f !! n)" |
|
637 sorry |
|
638 |
|
639 text {* |
|
640 A textbook proof of this lemma usually goes: |
|
641 |
|
642 Case 0: Trivial |
|
643 |
|
644 Case (Suc n): We have to prove that |
|
645 |
|
646 "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" |
|
647 |
|
648 holds. The induction hypothesis is |
|
649 |
|
650 "f !! (m + n) = (f !! m) o (f !! n)" |
|
651 |
|
652 The proof is as follows |
|
653 |
|
654 "f !! (m + (Suc n)) = f !! (Suc (m + n))" |
|
655 = f !! (m + n) o f |
|
656 = (f !! m) o (f !! n) o f (by ih) |
|
657 = (f !! m) o ((f !! n) o f) (by o_assoc) |
|
658 = (f !! m) o (f !! (Suc n)) |
|
659 Done. |
|
660 |
|
661 *} |
|
662 |
|
663 lemma |
|
664 shows "f !! (m + n) = (f !! m) o (f !! n)" |
|
665 proof (induct n) |
|
666 case 0 |
|
667 show "f !! (m + 0) = (f !! m) o (f !! 0)" sorry |
|
668 next |
|
669 case (Suc n) |
|
670 have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact |
|
671 |
|
672 show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" sorry |
|
673 qed |
|
674 |
|
675 |
|
676 text {* |
|
677 The following proof involves several steps of equational reasoning. |
|
678 Isar provides some convenient means to express such equational |
|
679 reasoning in a much cleaner fashion using the "also have" |
|
680 construction. *} |
|
681 |
|
682 lemma |
|
683 shows "f !! (m + n) = (f !! m) o (f !! n)" |
|
684 proof (induct n) |
|
685 case 0 |
|
686 show "f !! (m + 0) = (f !! m) o (f !! 0)" by (simp add: comp_def) |
|
687 next |
|
688 case (Suc n) |
|
689 have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact |
|
690 have eq1: "f !! (m + (Suc n)) = f !! (Suc (m + n))" by simp |
|
691 have eq2: "f !! (Suc (m + n)) = f !! (m + n) o f" by simp |
|
692 have eq3: "f !! (m + n) o f = (f !! m) o (f !! n) o f" using ih by simp |
|
693 have eq4: "(f !! m) o (f !! n) o f = (f !! m) o ((f !! n) o f)" by (simp add: o_assoc) |
|
694 have eq5: "(f !! m) o ((f !! n) o f) = (f !! m) o (f !! (Suc n))" by simp |
|
695 show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" |
|
696 using eq1 eq2 eq3 eq4 eq5 by (simp only:) |
|
697 qed |
|
698 |
|
699 text {* |
|
700 The equations can be stringed together with "also have" and "\<dots>", |
|
701 as shown below. *} |
|
702 |
|
703 lemma |
|
704 shows "f !! (m + n) = (f !! m) o (f !! n)" |
|
705 proof (induct n) |
|
706 case 0 |
|
707 show "f !! (m + 0) = (f !! m) o (f !! 0)" by (simp add: comp_def) |
|
708 next |
|
709 case (Suc n) |
|
710 have ih: "f !! (m + n) = (f !! m) o (f !! n)" by fact |
|
711 have "f !! (m + (Suc n)) = f !! (Suc (m + n))" by simp |
|
712 also have "\<dots> = f !! (m + n) o f" by simp |
|
713 also have "\<dots> = (f !! m) o (f !! n) o f" using ih by simp |
|
714 also have "\<dots> = (f !! m) o ((f !! n) o f)" by (simp add: o_assoc) |
|
715 also have "\<dots> = (f !! m) o (f !! (Suc n))" by simp |
|
716 finally show "f !! (m + (Suc n)) = f !! m o (f !! (Suc n))" by simp |
|
717 qed |
|
718 |
|
719 text {* |
|
720 This type of equational reasoning also extends to relations. |
|
721 For this consider the following power function and the auxiliary |
|
722 fact about less-equal and multiplication. *} |
|
723 |
|
724 fun |
|
725 pow :: "nat \<Rightarrow> nat \<Rightarrow> nat" ("_ \<up> _") |
|
726 where |
|
727 "m \<up> 0 = 1" |
|
728 | "m \<up> (Suc n) = m * (m \<up> n)" |
|
729 |
|
730 lemma aux: |
|
731 fixes a b c::"nat" |
|
732 assumes a: "a \<le> b" |
|
733 shows " (c * a) \<le> (c * b)" |
|
734 using a by (auto) |
|
735 |
|
736 text {* |
|
737 With this you can easily implement the following proof |
|
738 which is straight taken out of Velleman's "How To Prove It". |
|
739 *} |
|
740 |
|
741 lemma |
|
742 shows "1 + n * x \<le> (1 + x) \<up> n" |
|
743 proof (induct n) |
|
744 case 0 |
|
745 show "1 + 0 * x \<le> (1 + x) \<up> 0" by simp |
|
746 next |
|
747 case (Suc n) |
|
748 have ih: "1 + n * x \<le> (1 + x) \<up> n" by fact |
|
749 have "1 + (Suc n) * x \<le> 1 + x + (n * x) + (n * x * x)" by (simp) |
|
750 also have "\<dots> = (1 + x) * (1 + n * x)" by simp |
|
751 also have "\<dots> \<le> (1 + x) * ((1 + x) \<up> n)" using ih aux by blast |
|
752 also have "\<dots> = (1 + x) \<up> (Suc n)" by simp |
|
753 finally show "1 + (Suc n) * x \<le> (1 + x) \<up> (Suc n)" by simp |
|
754 qed |
|
755 |
|
756 text {* |
|
757 What Velleman actually wants to prove is a proper |
|
758 inequality. For this we can nest proofs. |
|
759 *} |
|
760 |
|
761 lemma |
|
762 shows "n * x < (1 + x) \<up> n" |
|
763 proof - |
|
764 have "1 + n * x \<le> (1 + x) \<up> n" |
|
765 proof (induct n) |
|
766 case 0 |
|
767 show "1 + 0 * x \<le> (1 + x) \<up> 0" by simp |
|
768 next |
|
769 case (Suc n) |
|
770 have ih: "1 + n * x \<le> (1 + x) \<up> n" by fact |
|
771 have "1 + (Suc n) * x \<le> 1 + x + (n * x) + (n * x * x)" by (simp) |
|
772 also have "\<dots> = (1 + x) * (1 + n * x)" by simp |
|
773 also have "\<dots> \<le> (1 + x) * ((1 + x) \<up> n)" using ih aux by blast |
|
774 also have "\<dots> = (1 + x) \<up> (Suc n)" by simp |
|
775 finally show "1 + (Suc n) * x \<le> (1 + x) \<up> (Suc n)" by simp |
|
776 qed |
|
777 then show "n * x < (1 + x) \<up> n" by simp |
|
778 qed |
|
779 |
|
780 |
|
781 end |
|
782 |
|
783 |
|
784 |
|
785 |
|
786 |
|
787 |
|
788 |
|
789 |
|
790 |
|
791 |