src/HOL/ex/Sum_of_Powers.thy
author paulson <lp15@cam.ac.uk>
Sun, 14 Aug 2022 23:51:47 +0100
changeset 75864 3842556b757c
parent 70114 089c17514794
permissions -rw-r--r--
moved some material from Sum_of_Powers
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
     1
(*  Author:     Lukas Bulwahn <lukas.bulwahn-at-gmail.com> *)
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
     2
section \<open>Sum of Powers\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     3
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     4
theory Sum_of_Powers
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     5
imports Complex_Main
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     6
begin
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     7
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
     8
subsection \<open>Preliminaries\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
     9
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    10
lemma integrals_eq:
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
    11
  assumes "f a = g a"
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    12
  assumes "\<And> x. ((\<lambda>x. f x - g x) has_real_derivative 0) (at x)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    13
  shows "f x = g x"
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
    14
  by (metis (no_types, lifting) DERIV_isconst_all assms(1) assms(2) eq_iff_diff_eq_0)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    15
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    16
lemma sum_diff: "((\<Sum>i\<le>n::nat. f (i + 1) - f i)::'a::field) = f (n + 1) - f 0"
75864
3842556b757c moved some material from Sum_of_Powers
paulson <lp15@cam.ac.uk>
parents: 70114
diff changeset
    17
  by (induct n) (auto simp add: field_simps)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    18
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    19
declare One_nat_def [simp del]
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    20
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
    21
subsection \<open>Bernoulli Numbers and Bernoulli Polynomials\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    22
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    23
declare sum.cong [fundef_cong]
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    24
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    25
fun bernoulli :: "nat \<Rightarrow> real"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    26
where
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    27
  "bernoulli 0 = (1::real)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    28
| "bernoulli (Suc n) =  (-1 / (n + 2)) * (\<Sum>k \<le> n. ((n + 2 choose k) * bernoulli k))"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    29
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    30
declare bernoulli.simps[simp del]
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    31
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    32
definition
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    33
  "bernpoly n = (\<lambda>x. \<Sum>k \<le> n. (n choose k) * bernoulli k * x ^ (n - k))"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    34
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
    35
subsection \<open>Basic Observations on Bernoulli Polynomials\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    36
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    37
lemma bernpoly_0: "bernpoly n 0 = bernoulli n"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    38
proof (cases n)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    39
  case 0
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
    40
  then show "bernpoly n 0 = bernoulli n"
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    41
    unfolding bernpoly_def bernoulli.simps by auto
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    42
next
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    43
  case (Suc n')
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    44
  have "(\<Sum>k\<le>n'. real (Suc n' choose k) * bernoulli k * 0 ^ (Suc n' - k)) = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    45
    by (rule sum.neutral) auto
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    46
  with Suc show ?thesis
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    47
    unfolding bernpoly_def by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    48
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    49
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    50
lemma sum_binomial_times_bernoulli:
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    51
  "(\<Sum>k\<le>n. ((Suc n) choose k) * bernoulli k) = (if n = 0 then 1 else 0)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    52
proof (cases n)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    53
  case 0
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
    54
  then show ?thesis by (simp add: bernoulli.simps)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    55
next
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    56
  case Suc
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
    57
  then show ?thesis
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    58
  by (simp add: bernoulli.simps)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    59
    (simp add: field_simps add_2_eq_Suc'[symmetric] del: add_2_eq_Suc add_2_eq_Suc')
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    60
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    61
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
    62
subsection \<open>Sum of Powers with Bernoulli Polynomials\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    63
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    64
lemma bernpoly_derivative [derivative_intros]:
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    65
  "(bernpoly (Suc n) has_real_derivative ((n + 1) * bernpoly n x)) (at x)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    66
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    67
  have "(bernpoly (Suc n) has_real_derivative (\<Sum>k\<le>n. real (Suc n - k) * x ^ (n - k) * (real (Suc n choose k) * bernoulli k))) (at x)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    68
    unfolding bernpoly_def by (rule DERIV_cong) (fast intro!: derivative_intros, simp)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    69
  moreover have "(\<Sum>k\<le>n. real (Suc n - k) * x ^ (n - k) * (real (Suc n choose k) * bernoulli k)) = (n + 1) * bernpoly n x"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    70
    unfolding bernpoly_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    71
    by (auto intro: sum.cong simp add: sum_distrib_left real_binomial_eq_mult_binomial_Suc[of _ n] Suc_eq_plus1 of_nat_diff)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    72
  ultimately show ?thesis by auto
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    73
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    74
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    75
lemma diff_bernpoly:
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    76
  "bernpoly n (x + 1) - bernpoly n x = n * x ^ (n - 1)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    77
proof (induct n arbitrary: x)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    78
  case 0
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    79
  show ?case unfolding bernpoly_def by auto
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    80
next
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    81
  case (Suc n)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    82
  have "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = (Suc n) * 0 ^ n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    83
    unfolding bernpoly_0 unfolding bernpoly_def by (simp add: sum_binomial_times_bernoulli zero_power)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
    84
  then have const: "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = real (Suc n) * 0 ^ n" by (simp add: power_0_left)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    85
  have hyps': "\<And>x. (real n + 1) * bernpoly n (x + 1) - (real n + 1) * bernpoly n x = real n * x ^ (n - Suc 0) * real (Suc n)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    86
    unfolding right_diff_distrib[symmetric] by (simp add: Suc.hyps One_nat_def)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    87
  note [derivative_intros] = DERIV_chain'[where f = "\<lambda>x::real. x + 1" and g = "bernpoly (Suc n)" and s="UNIV"]
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    88
  have derivative: "\<And>x. ((%x. bernpoly (Suc n) (x + 1) - bernpoly (Suc n) x - real (Suc n) * x ^ n) has_real_derivative 0) (at x)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    89
    by (rule DERIV_cong) (fast intro!: derivative_intros, simp add: hyps')
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    90
  from integrals_eq[OF const derivative] show ?case by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    91
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    92
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    93
lemma sum_of_powers: "(\<Sum>k\<le>n::nat. (real k) ^ m) = (bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0) / (m + 1)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    94
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    95
  from diff_bernpoly[of "Suc m", simplified] have "(m + (1::real)) * (\<Sum>k\<le>n. (real k) ^ m) = (\<Sum>k\<le>n. bernpoly (Suc m) (real k + 1) - bernpoly (Suc m) (real k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
    96
    by (auto simp add: sum_distrib_left intro!: sum.cong)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    97
  also have "... = (\<Sum>k\<le>n. bernpoly (Suc m) (real (k + 1)) - bernpoly (Suc m) (real k))"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
    98
    by simp
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
    99
  also have "... = bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   100
    by (simp only: sum_diff[where f="\<lambda>k. bernpoly (Suc m) (real k)"]) simp
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   101
  finally show ?thesis by (auto simp add: field_simps intro!: eq_divide_imp)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   102
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   103
61343
5b5656a63bd6 isabelle update_cartouches;
wenzelm
parents: 60603
diff changeset
   104
subsection \<open>Instances for Square And Cubic Numbers\<close>
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   105
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   106
lemma binomial_unroll:
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   107
  "n > 0 \<Longrightarrow> (n choose k) = (if k = 0 then 1 else (n - 1) choose (k - 1) + ((n - 1) choose k))"
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 61694
diff changeset
   108
  by (auto simp add: gr0_conv_Suc)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   109
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   110
lemma sum_unroll:
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   111
  "(\<Sum>k\<le>n::nat. f k) = (if n = 0 then f 0 else f n + (\<Sum>k\<le>n - 1. f k))"
70097
4005298550a6 The last big tranche of Homology material: invariance of domain; renamings to use generic sum/prod lemmas from their locale
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
   112
by auto (metis One_nat_def Suc_pred add.commute sum.atMost_Suc)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   113
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   114
lemma bernoulli_unroll:
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   115
  "n > 0 \<Longrightarrow> bernoulli n = - 1 / (real n + 1) * (\<Sum>k\<le>n - 1. real (n + 1 choose k) * bernoulli k)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   116
by (cases n) (simp add: bernoulli.simps One_nat_def)+
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   117
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 61694
diff changeset
   118
lemmas unroll = binomial_unroll
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63918
diff changeset
   119
  bernoulli.simps(1) bernoulli_unroll sum_unroll bernpoly_def
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   120
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   121
lemma sum_of_squares: "(\<Sum>k\<le>n::nat. k ^ 2) = (2 * n ^ 3 + 3 * n ^ 2 + n) / 6"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   122
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   123
  have "real (\<Sum>k\<le>n::nat. k ^ 2) = (\<Sum>k\<le>n::nat. (real k) ^ 2)" by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   124
  also have "... = (bernpoly 3 (real (n + 1)) - bernpoly 3 0) / real (3 :: nat)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   125
    by (auto simp add: sum_of_powers)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   126
  also have "... = (2 * n ^ 3 + 3 * n ^ 2 + n) / 6"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   127
    by (simp add: unroll algebra_simps power2_eq_square power3_eq_cube One_nat_def[symmetric])
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   128
  finally show ?thesis by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   129
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   130
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   131
lemma sum_of_squares_nat: "(\<Sum>k\<le>n::nat. k ^ 2) = (2 * n ^ 3 + 3 * n ^ 2 + n) div 6"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   132
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   133
  from sum_of_squares have "real (6 * (\<Sum>k\<le>n. k ^ 2)) = real (2 * n ^ 3 + 3 * n ^ 2 + n)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   134
    by (auto simp add: field_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
   135
  then have "6 * (\<Sum>k\<le>n. k ^ 2) = 2 * n ^ 3 + 3 * n ^ 2 + n"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   136
    using of_nat_eq_iff by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
   137
  then show ?thesis by auto
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   138
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   139
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   140
lemma sum_of_cubes: "(\<Sum>k\<le>n::nat. k ^ 3) = (n ^ 2 + n) ^ 2 / 4"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   141
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   142
  have two_plus_two: "2 + 2 = 4" by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   143
  have power4_eq: "\<And>x::real. x ^ 4 = x * x * x * x"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   144
    by (simp only: two_plus_two[symmetric] power_add power2_eq_square)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   145
  have "real (\<Sum>k\<le>n::nat. k ^ 3) = (\<Sum>k\<le>n::nat. (real k) ^ 3)" by simp
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   146
  also have "... = ((bernpoly 4 (n + 1) - bernpoly 4 0)) / (real (4 :: nat))"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   147
    by (auto simp add: sum_of_powers)
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   148
  also have "... = ((n ^ 2 + n) / 2) ^ 2"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   149
    by (simp add: unroll algebra_simps power2_eq_square power4_eq power3_eq_cube)
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
   150
  finally show ?thesis by (simp add: power_divide)
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   151
qed
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
   152
                       
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   153
lemma sum_of_cubes_nat: "(\<Sum>k\<le>n::nat. k ^ 3) = (n ^ 2 + n) ^ 2 div 4"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   154
proof -
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   155
  from sum_of_cubes have "real (4 * (\<Sum>k\<le>n. k ^ 3)) = real ((n ^ 2 + n) ^ 2)"
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   156
    by (auto simp add: field_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
   157
  then have "4 * (\<Sum>k\<le>n. k ^ 3) = (n ^ 2 + n) ^ 2"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61609
diff changeset
   158
    using of_nat_eq_iff by blast
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61343
diff changeset
   159
  then show ?thesis by auto
60603
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   160
qed
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   161
09ecbd791d4a add examples from Freek's top 100 theorems (thms 30, 73, 77)
bulwahn
parents:
diff changeset
   162
end