author | haftmann |
Tue, 12 May 2009 21:17:47 +0200 | |
changeset 31129 | d2cead76fca2 |
parent 18193 | 54419506df9e |
child 31758 | 3edd5f813f01 |
permissions | -rw-r--r-- |
7443 | 1 |
(* Title: HOL/Isar_examples/Summation.thy |
2 |
ID: $Id$ |
|
3 |
Author: Markus Wenzel |
|
4 |
*) |
|
5 |
||
10007 | 6 |
header {* Summing natural numbers *} |
7443 | 7 |
|
15561 | 8 |
theory Summation |
9 |
imports Main |
|
10 |
begin |
|
11 |
||
7968 | 12 |
text_raw {* |
13 |
\footnote{This example is somewhat reminiscent of the |
|
14 |
\url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, which is |
|
15 |
discussed in \cite{isabelle-ref} in the context of permutative |
|
16 |
rewrite rules and ordered rewriting.} |
|
10007 | 17 |
*} |
7968 | 18 |
|
19 |
text {* |
|
20 |
Subsequently, we prove some summation laws of natural numbers |
|
7982 | 21 |
(including odds, squares, and cubes). These examples demonstrate how |
7968 | 22 |
plain natural deduction (including induction) may be combined with |
23 |
calculational proof. |
|
10007 | 24 |
*} |
7968 | 25 |
|
7761 | 26 |
|
10007 | 27 |
subsection {* Summation laws *} |
7443 | 28 |
|
7968 | 29 |
text {* |
30 |
The sum of natural numbers $0 + \cdots + n$ equals $n \times (n + |
|
7982 | 31 |
1)/2$. Avoiding formal reasoning about division we prove this |
32 |
equation multiplied by $2$. |
|
10007 | 33 |
*} |
7800 | 34 |
|
35 |
theorem sum_of_naturals: |
|
15561 | 36 |
"2 * (\<Sum>i::nat=0..n. i) = n * (n + 1)" |
10007 | 37 |
(is "?P n" is "?S n = _") |
38 |
proof (induct n) |
|
39 |
show "?P 0" by simp |
|
10146 | 40 |
next |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
41 |
fix n have "?S (n + 1) = ?S n + 2 * (n + 1)" by simp |
10007 | 42 |
also assume "?S n = n * (n + 1)" |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
43 |
also have "... + 2 * (n + 1) = (n + 1) * (n + 2)" by simp |
10007 | 44 |
finally show "?P (Suc n)" by simp |
45 |
qed |
|
7443 | 46 |
|
7968 | 47 |
text {* |
48 |
The above proof is a typical instance of mathematical induction. The |
|
49 |
main statement is viewed as some $\var{P} \ap n$ that is split by the |
|
50 |
induction method into base case $\var{P} \ap 0$, and step case |
|
7982 | 51 |
$\var{P} \ap n \Impl \var{P} \ap (\idt{Suc} \ap n)$ for arbitrary $n$. |
7968 | 52 |
|
53 |
The step case is established by a short calculation in forward |
|
54 |
manner. Starting from the left-hand side $\var{S} \ap (n + 1)$ of |
|
7982 | 55 |
the thesis, the final result is achieved by transformations involving |
56 |
basic arithmetic reasoning (using the Simplifier). The main point is |
|
57 |
where the induction hypothesis $\var{S} \ap n = n \times (n + 1)$ is |
|
58 |
introduced in order to replace a certain subterm. So the |
|
7968 | 59 |
``transitivity'' rule involved here is actual \emph{substitution}. |
60 |
Also note how the occurrence of ``\dots'' in the subsequent step |
|
7982 | 61 |
documents the position where the right-hand side of the hypothesis |
7968 | 62 |
got filled in. |
63 |
||
64 |
\medskip A further notable point here is integration of calculations |
|
7982 | 65 |
with plain natural deduction. This works so well in Isar for two |
66 |
reasons. |
|
7968 | 67 |
\begin{enumerate} |
68 |
||
69 |
\item Facts involved in \isakeyword{also}~/ \isakeyword{finally} |
|
70 |
calculational chains may be just anything. There is nothing special |
|
71 |
about \isakeyword{have}, so the natural deduction element |
|
72 |
\isakeyword{assume} works just as well. |
|
73 |
||
74 |
\item There are two \emph{separate} primitives for building natural |
|
75 |
deduction contexts: \isakeyword{fix}~$x$ and \isakeyword{assume}~$A$. |
|
7982 | 76 |
Thus it is possible to start reasoning with some new ``arbitrary, but |
77 |
fixed'' elements before bringing in the actual assumption. In |
|
78 |
contrast, natural deduction is occasionally formalized with basic |
|
79 |
context elements of the form $x:A$ instead. |
|
7968 | 80 |
|
81 |
\end{enumerate} |
|
10007 | 82 |
*} |
7968 | 83 |
|
84 |
text {* |
|
7982 | 85 |
\medskip We derive further summation laws for odds, squares, and |
86 |
cubes as follows. The basic technique of induction plus calculation |
|
87 |
is the same as before. |
|
10007 | 88 |
*} |
7968 | 89 |
|
7800 | 90 |
theorem sum_of_odds: |
15561 | 91 |
"(\<Sum>i::nat=0..<n. 2 * i + 1) = n^Suc (Suc 0)" |
10007 | 92 |
(is "?P n" is "?S n = _") |
93 |
proof (induct n) |
|
94 |
show "?P 0" by simp |
|
10146 | 95 |
next |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
96 |
fix n have "?S (n + 1) = ?S n + 2 * n + 1" by simp |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
10672
diff
changeset
|
97 |
also assume "?S n = n^Suc (Suc 0)" |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
98 |
also have "... + 2 * n + 1 = (n + 1)^Suc (Suc 0)" by simp |
10007 | 99 |
finally show "?P (Suc n)" by simp |
100 |
qed |
|
7443 | 101 |
|
8814 | 102 |
text {* |
103 |
Subsequently we require some additional tweaking of Isabelle built-in |
|
104 |
arithmetic simplifications, such as bringing in distributivity by |
|
105 |
hand. |
|
10007 | 106 |
*} |
8814 | 107 |
|
10007 | 108 |
lemmas distrib = add_mult_distrib add_mult_distrib2 |
8814 | 109 |
|
7761 | 110 |
theorem sum_of_squares: |
15561 | 111 |
"6 * (\<Sum>i::nat=0..n. i^Suc (Suc 0)) = n * (n + 1) * (2 * n + 1)" |
10007 | 112 |
(is "?P n" is "?S n = _") |
113 |
proof (induct n) |
|
114 |
show "?P 0" by simp |
|
10146 | 115 |
next |
18193 | 116 |
fix n have "?S (n + 1) = ?S n + 6 * (n + 1)^Suc (Suc 0)" |
117 |
by (simp add: distrib) |
|
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
118 |
also assume "?S n = n * (n + 1) * (2 * n + 1)" |
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
119 |
also have "... + 6 * (n + 1)^Suc (Suc 0) = |
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
120 |
(n + 1) * (n + 2) * (2 * (n + 1) + 1)" by (simp add: distrib) |
10007 | 121 |
finally show "?P (Suc n)" by simp |
122 |
qed |
|
7443 | 123 |
|
7800 | 124 |
theorem sum_of_cubes: |
15561 | 125 |
"4 * (\<Sum>i::nat=0..n. i^3) = (n * (n + 1))^Suc (Suc 0)" |
10007 | 126 |
(is "?P n" is "?S n = _") |
127 |
proof (induct n) |
|
128 |
show "?P 0" by (simp add: power_eq_if) |
|
10146 | 129 |
next |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
130 |
fix n have "?S (n + 1) = ?S n + 4 * (n + 1)^3" |
10007 | 131 |
by (simp add: power_eq_if distrib) |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
10672
diff
changeset
|
132 |
also assume "?S n = (n * (n + 1))^Suc (Suc 0)" |
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
133 |
also have "... + 4 * (n + 1)^3 = ((n + 1) * ((n + 1) + 1))^Suc (Suc 0)" |
10007 | 134 |
by (simp add: power_eq_if distrib) |
135 |
finally show "?P (Suc n)" by simp |
|
136 |
qed |
|
7443 | 137 |
|
7968 | 138 |
text {* |
139 |
Comparing these examples with the tactic script version |
|
140 |
\url{http://isabelle.in.tum.de/library/HOL/ex/NatSum.html}, we note |
|
7982 | 141 |
an important difference of how induction vs.\ simplification is |
7968 | 142 |
applied. While \cite[\S10]{isabelle-ref} advises for these examples |
143 |
that ``induction should not be applied until the goal is in the |
|
144 |
simplest form'' this would be a very bad idea in our setting. |
|
145 |
||
146 |
Simplification normalizes all arithmetic expressions involved, |
|
7982 | 147 |
producing huge intermediate goals. With applying induction |
148 |
afterwards, the Isar proof text would have to reflect the emerging |
|
149 |
configuration by appropriate sub-proofs. This would result in badly |
|
150 |
structured, low-level technical reasoning, without any good idea of |
|
151 |
the actual point. |
|
7968 | 152 |
|
153 |
\medskip As a general rule of good proof style, automatic methods |
|
7982 | 154 |
such as $\idt{simp}$ or $\idt{auto}$ should normally be never used as |
7968 | 155 |
initial proof methods, but only as terminal ones, solving certain |
156 |
goals completely. |
|
10007 | 157 |
*} |
7968 | 158 |
|
10007 | 159 |
end |