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