| author | haftmann | 
| Sat, 29 Sep 2007 08:58:56 +0200 | |
| changeset 24751 | dbb34a03af5a | 
| parent 19469 | 958d2f2dd8d4 | 
| permissions | -rw-r--r-- | 
| 19358 | 1 | (* Title: HOL/Complex/ex/Arithmetic_Series_Complex | 
| 2 | ID: $Id$ | |
| 3 | Author: Benjamin Porter, 2006 | |
| 4 | *) | |
| 5 | ||
| 6 | ||
| 7 | header {* Arithmetic Series for Reals *}
 | |
| 8 | ||
| 9 | theory Arithmetic_Series_Complex | |
| 19469 
958d2f2dd8d4
moved arithmetic series to geometric series in SetInterval
 kleing parents: 
19358diff
changeset | 10 | imports Complex_Main | 
| 19358 | 11 | begin | 
| 12 | ||
| 13 | lemma arith_series_real: | |
| 14 |   "(2::real) * (\<Sum>i\<in>{..<n}. a + of_nat i * d) =
 | |
| 15 | of_nat n * (a + (a + of_nat(n - 1)*d))" | |
| 16 | proof - | |
| 17 | have | |
| 18 |     "((1::real) + 1) * (\<Sum>i\<in>{..<n}. a + of_nat(i)*d) =
 | |
| 19 | of_nat(n) * (a + (a + of_nat(n - 1)*d))" | |
| 20 | by (rule arith_series_general) | |
| 21 | thus ?thesis by simp | |
| 22 | qed | |
| 23 | ||
| 24 | end |