| author | wenzelm | 
| Tue, 10 Nov 2015 23:41:20 +0100 | |
| changeset 61626 | c304402cc3df | 
| parent 61541 | 846c72206207 | 
| child 61932 | 2e48182cc82c | 
| permissions | -rw-r--r-- | 
| 33026 | 1 | (* Title: HOL/Isar_Examples/Summation.thy | 
| 7443 | 2 | Author: Markus Wenzel | 
| 3 | *) | |
| 4 | ||
| 58882 | 5 | section \<open>Summing natural numbers\<close> | 
| 7443 | 6 | |
| 15561 | 7 | theory Summation | 
| 8 | imports Main | |
| 9 | begin | |
| 10 | ||
| 58614 | 11 | text \<open>Subsequently, we prove some summation laws of natural numbers | 
| 61541 | 12 | (including odds, squares, and cubes). These examples demonstrate how plain | 
| 13 | natural deduction (including induction) may be combined with calculational | |
| 14 | proof.\<close> | |
| 7968 | 15 | |
| 7761 | 16 | |
| 58614 | 17 | subsection \<open>Summation laws\<close> | 
| 7443 | 18 | |
| 61541 | 19 | text \<open>The sum of natural numbers \<open>0 + \<cdots> + n\<close> equals \<open>n \<times> (n + 1)/2\<close>. | 
| 20 | Avoiding formal reasoning about division we prove this equation multiplied | |
| 21 | by \<open>2\<close>.\<close> | |
| 7800 | 22 | |
| 23 | theorem sum_of_naturals: | |
| 15561 | 24 | "2 * (\<Sum>i::nat=0..n. i) = n * (n + 1)" | 
| 10007 | 25 | (is "?P n" is "?S n = _") | 
| 26 | proof (induct n) | |
| 27 | show "?P 0" by simp | |
| 10146 | 28 | next | 
| 55640 | 29 | fix n have "?S (n + 1) = ?S n + 2 * (n + 1)" | 
| 30 | by simp | |
| 10007 | 31 | also assume "?S n = n * (n + 1)" | 
| 55640 | 32 | also have "\<dots> + 2 * (n + 1) = (n + 1) * (n + 2)" | 
| 33 | by simp | |
| 34 | finally show "?P (Suc n)" | |
| 35 | by simp | |
| 10007 | 36 | qed | 
| 7443 | 37 | |
| 61541 | 38 | text \<open>The above proof is a typical instance of mathematical induction. The | 
| 39 | main statement is viewed as some \<open>?P n\<close> that is split by the induction | |
| 40 | method into base case \<open>?P 0\<close>, and step case \<open>?P n \<Longrightarrow> ?P (Suc n)\<close> for | |
| 41 | arbitrary \<open>n\<close>. | |
| 7968 | 42 | |
| 61541 | 43 | The step case is established by a short calculation in forward manner. | 
| 44 | Starting from the left-hand side \<open>?S (n + 1)\<close> of the thesis, the final | |
| 45 | result is achieved by transformations involving basic arithmetic reasoning | |
| 46 | (using the Simplifier). The main point is where the induction hypothesis | |
| 47 | \<open>?S n = n \<times> (n + 1)\<close> is introduced in order to replace a certain subterm. | |
| 48 | So the ``transitivity'' rule involved here is actual \<^emph>\<open>substitution\<close>. Also | |
| 49 | note how the occurrence of ``\dots'' in the subsequent step documents the | |
| 50 | position where the right-hand side of the hypothesis got filled in. | |
| 7968 | 51 | |
| 61541 | 52 | \<^medskip> | 
| 53 | A further notable point here is integration of calculations with plain | |
| 54 | natural deduction. This works so well in Isar for two reasons. | |
| 7968 | 55 | |
| 61541 | 56 |     \<^enum> Facts involved in \isakeyword{also}~/ \isakeyword{finally}
 | 
| 57 | calculational chains may be just anything. There is nothing special | |
| 58 |     about \isakeyword{have}, so the natural deduction element
 | |
| 59 |     \isakeyword{assume} works just as well.
 | |
| 60 | ||
| 61 | \<^enum> There are two \<^emph>\<open>separate\<close> primitives for building natural deduction | |
| 62 |     contexts: \isakeyword{fix}~\<open>x\<close> and \isakeyword{assume}~\<open>A\<close>. Thus it is
 | |
| 63 | possible to start reasoning with some new ``arbitrary, but fixed'' | |
| 64 | elements before bringing in the actual assumption. In contrast, natural | |
| 65 | deduction is occasionally formalized with basic context elements of the | |
| 66 | form \<open>x:A\<close> instead. | |
| 7968 | 67 | |
| 61541 | 68 | \<^medskip> | 
| 69 | We derive further summation laws for odds, squares, and cubes as follows. | |
| 70 | The basic technique of induction plus calculation is the same as before.\<close> | |
| 7968 | 71 | |
| 7800 | 72 | theorem sum_of_odds: | 
| 15561 | 73 | "(\<Sum>i::nat=0..<n. 2 * i + 1) = n^Suc (Suc 0)" | 
| 10007 | 74 | (is "?P n" is "?S n = _") | 
| 75 | proof (induct n) | |
| 76 | show "?P 0" by simp | |
| 10146 | 77 | next | 
| 55640 | 78 | fix n | 
| 79 | have "?S (n + 1) = ?S n + 2 * n + 1" | |
| 80 | by simp | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
10672diff
changeset | 81 | also assume "?S n = n^Suc (Suc 0)" | 
| 55640 | 82 | also have "\<dots> + 2 * n + 1 = (n + 1)^Suc (Suc 0)" | 
| 83 | by simp | |
| 84 | finally show "?P (Suc n)" | |
| 85 | by simp | |
| 10007 | 86 | qed | 
| 7443 | 87 | |
| 61541 | 88 | text \<open>Subsequently we require some additional tweaking of Isabelle built-in | 
| 89 | arithmetic simplifications, such as bringing in distributivity by hand.\<close> | |
| 8814 | 90 | |
| 10007 | 91 | lemmas distrib = add_mult_distrib add_mult_distrib2 | 
| 8814 | 92 | |
| 7761 | 93 | theorem sum_of_squares: | 
| 15561 | 94 | "6 * (\<Sum>i::nat=0..n. i^Suc (Suc 0)) = n * (n + 1) * (2 * n + 1)" | 
| 10007 | 95 | (is "?P n" is "?S n = _") | 
| 96 | proof (induct n) | |
| 97 | show "?P 0" by simp | |
| 10146 | 98 | next | 
| 55640 | 99 | fix n | 
| 100 | have "?S (n + 1) = ?S n + 6 * (n + 1)^Suc (Suc 0)" | |
| 18193 | 101 | by (simp add: distrib) | 
| 11704 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 wenzelm parents: 
11701diff
changeset | 102 | also assume "?S n = n * (n + 1) * (2 * n + 1)" | 
| 55640 | 103 | also have "\<dots> + 6 * (n + 1)^Suc (Suc 0) = | 
| 104 | (n + 1) * (n + 2) * (2 * (n + 1) + 1)" | |
| 105 | by (simp add: distrib) | |
| 106 | finally show "?P (Suc n)" | |
| 107 | by simp | |
| 10007 | 108 | qed | 
| 7443 | 109 | |
| 7800 | 110 | theorem sum_of_cubes: | 
| 15561 | 111 | "4 * (\<Sum>i::nat=0..n. i^3) = (n * (n + 1))^Suc (Suc 0)" | 
| 10007 | 112 | (is "?P n" is "?S n = _") | 
| 113 | proof (induct n) | |
| 114 | show "?P 0" by (simp add: power_eq_if) | |
| 10146 | 115 | next | 
| 55640 | 116 | fix n | 
| 117 | have "?S (n + 1) = ?S n + 4 * (n + 1)^3" | |
| 10007 | 118 | by (simp add: power_eq_if distrib) | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
10672diff
changeset | 119 | also assume "?S n = (n * (n + 1))^Suc (Suc 0)" | 
| 55640 | 120 | also have "\<dots> + 4 * (n + 1)^3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)" | 
| 10007 | 121 | by (simp add: power_eq_if distrib) | 
| 55640 | 122 | finally show "?P (Suc n)" | 
| 123 | by simp | |
| 10007 | 124 | qed | 
| 7443 | 125 | |
| 58614 | 126 | text \<open>Note that in contrast to older traditions of tactical proof | 
| 50086 
ecffea78d381
tuned -- eliminated obsolete citation of isabelle-ref;
 wenzelm parents: 
40880diff
changeset | 127 | scripts, the structured proof applies induction on the original, | 
| 61541 | 128 | unsimplified statement. This allows to state the induction cases robustly | 
| 129 | and conveniently. Simplification (or other automated) methods are then | |
| 130 | applied in terminal position to solve certain sub-problems completely. | |
| 7968 | 131 | |
| 61541 | 132 | As a general rule of good proof style, automatic methods such as \<open>simp\<close> or | 
| 133 | \<open>auto\<close> should normally be never used as initial proof methods with a | |
| 134 | nested sub-proof to address the automatically produced situation, but only | |
| 135 | as terminal ones to solve sub-problems.\<close> | |
| 7968 | 136 | |
| 10007 | 137 | end |