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