src/ZF/ex/NatSum.thy
author wenzelm
Wed, 30 Aug 2023 21:34:53 +0200
changeset 78622 c42f316f0a01
parent 76219 cf7db6353322
permissions -rw-r--r--
tuned NEWS;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41777
1f7cbe39d425 more precise headers;
wenzelm
parents: 35762
diff changeset
     1
(*  Title:      ZF/ex/NatSum.thy
76219
cf7db6353322 recover informal "&" from 0c18df79b1c8;
wenzelm
parents: 76215
diff changeset
     2
    Author:     Tobias Nipkow & Lawrence C Paulson
9647
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
     3
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
     4
A summation operator. sum(f,n+1) is the sum of all f(i), i=0...n.
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
     5
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
     6
Note: n is a natural number but the sum is an integer,
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
     7
                            and f maps integers to integers
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
     8
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
     9
Summing natural numbers, squares, cubes, etc.
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    10
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    11
Originally demonstrated permutative rewriting, but add_ac is no longer needed
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    12
  thanks to new simprocs.
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    13
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    14
Thanks to Sloane's On-Line Encyclopedia of Integer Sequences,
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 65449
diff changeset
    15
  http://www.research.att.com/\<not>njas/sequences/
9647
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    16
*)
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    17
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    18
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 41777
diff changeset
    19
theory NatSum imports ZF begin
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    20
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    21
consts sum :: "[i\<Rightarrow>i, i] \<Rightarrow> i"
9647
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    22
primrec 
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    23
  "sum (f,0) = #0"
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    24
  "sum (f, succ(n)) = f($#n) $+ sum(f,n)"
e9623f47275b new example ZF/ex/NatSum
paulson
parents:
diff changeset
    25
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    26
declare zadd_zmult_distrib [simp]  zadd_zmult_distrib2 [simp]
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    27
declare zdiff_zmult_distrib [simp] zdiff_zmult_distrib2 [simp]
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    28
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    29
(*The sum of the first n odd numbers equals n squared.*)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    30
lemma sum_of_odds: "n \<in> nat \<Longrightarrow> sum (\<lambda>i. i $+ i $+ #1, n) = $#n $* $#n"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    31
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    32
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    33
(*The sum of the first n odd squares*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    34
lemma sum_of_odd_squares:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    35
     "n \<in> nat \<Longrightarrow> #3 $* sum (\<lambda>i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    36
      $#n $* (#4 $* $#n $* $#n $- #1)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    37
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    38
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    39
(*The sum of the first n odd cubes*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    40
lemma sum_of_odd_cubes:
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    41
     "n \<in> nat  
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    42
      \<Longrightarrow> sum (\<lambda>i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    43
          $#n $* $#n $* (#2 $* $#n $* $#n $- #1)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    44
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    45
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    46
(*The sum of the first n positive integers equals n(n+1)/2.*)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    47
lemma sum_of_naturals:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    48
     "n \<in> nat \<Longrightarrow> #2 $* sum(\<lambda>i. i, succ(n)) = $#n $* $#succ(n)"
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    49
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    50
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    51
lemma sum_of_squares:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    52
     "n \<in> nat \<Longrightarrow> #6 $* sum (\<lambda>i. i$*i, succ(n)) =  
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    53
                  $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    54
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    55
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    56
lemma sum_of_cubes:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    57
     "n \<in> nat \<Longrightarrow> #4 $* sum (\<lambda>i. i$*i$*i, succ(n)) =  
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    58
                  $#n $* $#n $* ($#n $+ #1) $* ($#n $+ #1)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    59
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    60
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    61
(** Sum of fourth powers **)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    62
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    63
lemma sum_of_fourth_powers:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    64
     "n \<in> nat \<Longrightarrow> #30 $* sum (\<lambda>i. i$*i$*i$*i, succ(n)) =  
12867
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    65
                    $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1) $*  
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    66
                    (#3 $* $#n $* $#n $+ #3 $* $#n $- #1)"
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    67
by (induct_tac "n", auto)
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    68
5c900a821a7c New-style versions of these old examples
paulson
parents: 9647
diff changeset
    69
end