src/HOL/ex/NatSum.thy
author wenzelm
Sat, 03 Nov 2001 01:33:54 +0100
changeset 12023 d982f98e0f0d
parent 11786 51ce34ef5113
child 12196 a3be6b3a9c0b
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     1
(*  Title:      HOL/ex/NatSum.ML
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     2
    ID:         $Id$
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     3
    Author:     Tobias Nipkow
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     4
*)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     5
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     6
header {* Summing natural numbers *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     7
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     8
theory NatSum = Main:
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
     9
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    10
text {*
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    11
  Summing natural numbers, squares, cubes, etc.
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    12
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    13
  Originally demonstrated permutative rewriting, but @{thm [source]
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    14
  add_ac} is no longer needed thanks to new simprocs.
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    15
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    16
  Thanks to Sloane's On-Line Encyclopedia of Integer Sequences,
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    17
  \url{http://www.research.att.com/~njas/sequences/}.
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    18
*}
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    19
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    20
declare lessThan_Suc [simp] atMost_Suc [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    21
declare add_mult_distrib [simp] add_mult_distrib2 [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    22
declare diff_mult_distrib [simp] diff_mult_distrib2 [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    23
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    24
text {*
12023
wenzelm
parents: 11786
diff changeset
    25
  \medskip The sum of the first @{text n} odd numbers equals @{text n}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    26
  squared.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    27
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    28
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    29
lemma sum_of_odds: "(\<Sum>i \<in> {..n(}. Suc (i + i)) = n * n"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    30
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    31
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    32
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    33
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    34
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    35
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    36
  \medskip The sum of the first @{text n} odd squares.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    37
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    38
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    39
lemma sum_of_odd_squares:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    40
  "3 * (\<Sum>i \<in> {..n(}. Suc (i + i) * Suc (i + i)) =
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    41
    n * (4 * n * n - 1)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    42
  apply (induct n)
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    43
   apply (case_tac [2] n)  -- {* eliminates the subtraction *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    44
    apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    45
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    46
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    47
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    48
text {*
12023
wenzelm
parents: 11786
diff changeset
    49
  \medskip The sum of the first @{text n} odd cubes
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    50
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    51
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    52
lemma sum_of_odd_cubes:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    53
  "(\<Sum>i \<in> {..n(}. Suc (i + i) * Suc (i + i) * Suc (i + i)) =
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    54
    n * n * (2 * n * n - 1)"
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    55
  apply (induct n)
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    56
   apply (case_tac [2] n)  -- {* eliminates the subtraction *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    57
    apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    58
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    59
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    60
text {*
12023
wenzelm
parents: 11786
diff changeset
    61
  \medskip The sum of the first @{text n} positive integers equals
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    62
  @{text "n (n + 1) / 2"}.*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    63
11586
wenzelm
parents: 11377
diff changeset
    64
lemma sum_of_naturals:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    65
    "2 * (\<Sum>i \<in> {..n}. i) = n * Suc n"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    66
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    67
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    68
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    69
11586
wenzelm
parents: 11377
diff changeset
    70
lemma sum_of_squares:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    71
    "6 * (\<Sum>i \<in> {..n}. i * i) = n * Suc n * Suc (2 * n)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    72
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    73
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    74
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    75
11586
wenzelm
parents: 11377
diff changeset
    76
lemma sum_of_cubes:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    77
    "4 * (\<Sum>i \<in> {..n}. i * i * i) = n * n * Suc n * Suc n"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    78
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    79
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    80
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    81
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    82
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    83
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    84
  \medskip Sum of fourth powers: two versions.
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    85
*}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    86
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    87
lemma sum_of_fourth_powers:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    88
  "30 * (\<Sum>i \<in> {..n}. i * i * i * i) =
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    89
    n * Suc n * Suc (2 * n) * (3 * n * n + 3 * n - 1)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    90
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    91
   apply auto
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
    92
  apply (case_tac n)  -- {* eliminates the subtraction *}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    93
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    94
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    95
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    96
text {*
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    97
  Alternative proof, with a change of variables and much more
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    98
  subtraction, performed using the integers. *}
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
    99
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   100
declare
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   101
  zmult_int [symmetric, simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   102
  zadd_zmult_distrib [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   103
  zadd_zmult_distrib2 [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   104
  zdiff_zmult_distrib [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   105
  zdiff_zmult_distrib2 [simp]
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   106
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   107
lemma int_sum_of_fourth_powers:
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   108
  "30 * int (\<Sum>i \<in> {..m(}. i * i * i * i) =
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   109
    int m * (int m - 1) * (int (2 * m) - 1) *
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   110
    (int (3 * m * m) - int (3 * m) - 1)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   111
  apply (induct m)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   112
   apply simp_all
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   113
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   114
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   115
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   116
text {*
12023
wenzelm
parents: 11786
diff changeset
   117
  \medskip Sums of geometric series: @{text 2}, @{text 3} and the
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   118
  general case.
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   119
*}
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   120
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   121
lemma sum_of_2_powers: "(\<Sum>i \<in> {..n(}. 2^i) = 2^n - (1::nat)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   122
  apply (induct n)
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   123
   apply (auto split: nat_diff_split)
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   124
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   125
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   126
lemma sum_of_3_powers: "2 * (\<Sum>i \<in> {..n(}. 3^i) = 3^n - (1::nat)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   127
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   128
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   129
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   130
11786
51ce34ef5113 setsum syntax;
wenzelm
parents: 11704
diff changeset
   131
lemma sum_of_powers: "0 < k ==> (k - 1) * (\<Sum>i \<in> {..n(}. k^i) = k^n - (1::nat)"
11024
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   132
  apply (induct n)
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   133
   apply auto
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   134
  done
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   135
23bf8d787b04 converted to new-style theories;
wenzelm
parents: 8944
diff changeset
   136
end