| author | wenzelm | 
| Sat, 08 Sep 2018 11:44:47 +0200 | |
| changeset 68941 | c192c8f9f19b | 
| parent 68484 | 59793df7f853 | 
| child 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 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: 
61343diff
changeset | 1 | (* Author: Lukas Bulwahn <lukas.bulwahn-at-gmail.com> *) | 
| 61343 | 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 | |
| 68484 | 8 | subsection \<open>Additions to @{theory HOL.Binomial} Theory\<close>
 | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 9 | |
| 63417 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 10 | lemma (in field_char_0) one_plus_of_nat_neq_zero [simp]: | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 11 | "1 + of_nat n \<noteq> 0" | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 12 | proof - | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 13 | have "of_nat (Suc n) \<noteq> of_nat 0" | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 14 | unfolding of_nat_eq_iff by simp | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 15 | then show ?thesis by simp | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 16 | qed | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 17 | |
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 18 | lemma of_nat_binomial_eq_mult_binomial_Suc: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 19 | assumes "k \<le> n" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 20 |   shows "(of_nat :: (nat \<Rightarrow> ('a :: field_char_0))) (n choose k) = of_nat (n + 1 - k) / of_nat (n + 1) * of_nat (Suc n choose k)"
 | 
| 63417 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 21 | proof (cases k) | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 22 | case 0 then show ?thesis by simp | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 23 | next | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 24 | case (Suc l) | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 25 | have "of_nat (n + 1) * (\<Prod>i=0..<k. of_nat (n - i)) = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))" | 
| 64272 | 26 | using prod.atLeast0_lessThan_Suc [where ?'a = 'a, symmetric, of "\<lambda>i. of_nat (Suc n - i)" k] | 
| 27 | by (simp add: ac_simps prod.atLeast0_lessThan_Suc_shift) | |
| 63417 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 28 | also have "... = (of_nat :: (nat \<Rightarrow> 'a)) (Suc n - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))" | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 29 | by (simp add: Suc atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 30 | also have "... = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) * (\<Prod>i=0..<k. of_nat (Suc n - i))" | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 31 | by (simp only: Suc_eq_plus1) | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 32 | finally have "(\<Prod>i=0..<k. of_nat (n - i)) = (of_nat :: (nat \<Rightarrow> 'a)) (n + 1 - k) / of_nat (n + 1) * (\<Prod>i=0..<k. of_nat (Suc n - i))" | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 33 | by (simp add: field_simps) | 
| 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 haftmann parents: 
63367diff
changeset | 34 | with assms show ?thesis | 
| 64272 | 35 | by (simp add: binomial_altdef_of_nat prod_dividef) | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 36 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 37 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 38 | lemma real_binomial_eq_mult_binomial_Suc: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 39 | assumes "k \<le> n" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 40 | shows "(n choose k) = (n + 1 - k) / (n + 1) * (Suc n choose 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: 
61343diff
changeset | 41 | by (metis Suc_eq_plus1 add.commute assms le_SucI of_nat_Suc of_nat_binomial_eq_mult_binomial_Suc of_nat_diff) | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 42 | |
| 61343 | 43 | subsection \<open>Preliminaries\<close> | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 44 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 45 | lemma integrals_eq: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 46 | assumes "f 0 = g 0" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 47 | 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 | 48 | shows "f x = g x" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 49 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 50 | show "f x = g x" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 51 | proof (cases "x \<noteq> 0") | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 52 | case True | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 53 | from assms DERIV_const_ratio_const[OF this, of "\<lambda>x. f x - g x" 0] | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 54 | show ?thesis by auto | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 55 | qed (simp add: assms) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 56 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 57 | |
| 64267 | 58 | lemma sum_diff: "((\<Sum>i\<le>n::nat. f (i + 1) - f i)::'a::field) = f (n + 1) - f 0" | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 59 | by (induct n) (auto simp add: field_simps) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 60 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 61 | declare One_nat_def [simp del] | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 62 | |
| 61343 | 63 | 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 | 64 | |
| 64267 | 65 | declare sum.cong [fundef_cong] | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 66 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 67 | fun bernoulli :: "nat \<Rightarrow> real" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 68 | where | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 69 | "bernoulli 0 = (1::real)" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 70 | | "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 | 71 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 72 | declare bernoulli.simps[simp del] | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 73 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 74 | definition | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 75 | "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 | 76 | |
| 61343 | 77 | 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 | 78 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 79 | 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 | 80 | proof (cases n) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 81 | 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: 
61343diff
changeset | 82 | 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 | 83 | unfolding bernpoly_def bernoulli.simps by auto | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 84 | next | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 85 | case (Suc n') | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 86 | have "(\<Sum>k\<le>n'. real (Suc n' choose k) * bernoulli k * 0 ^ (Suc n' - k)) = 0" | 
| 64267 | 87 | by (rule sum.neutral) auto | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 88 | with Suc show ?thesis | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 89 | unfolding bernpoly_def by simp | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 90 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 91 | |
| 64267 | 92 | lemma sum_binomial_times_bernoulli: | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 93 | "(\<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 | 94 | proof (cases n) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 95 | 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: 
61343diff
changeset | 96 | 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 | 97 | next | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 98 | 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: 
61343diff
changeset | 99 | then show ?thesis | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 100 | by (simp add: bernoulli.simps) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 101 | (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 | 102 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 103 | |
| 61343 | 104 | 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 | 105 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 106 | lemma bernpoly_derivative [derivative_intros]: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 107 | "(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 | 108 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 109 | 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 | 110 | 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 | 111 | 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 | 112 | unfolding bernpoly_def | 
| 64267 | 113 | 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 | 114 | ultimately show ?thesis by auto | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 115 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 116 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 117 | lemma diff_bernpoly: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 118 | "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 | 119 | proof (induct n arbitrary: x) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 120 | case 0 | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 121 | show ?case unfolding bernpoly_def by auto | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 122 | next | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 123 | case (Suc n) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 124 | have "bernpoly (Suc n) (0 + 1) - bernpoly (Suc n) 0 = (Suc n) * 0 ^ n" | 
| 64267 | 125 | 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: 
61343diff
changeset | 126 | 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 | 127 | 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 | 128 | 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 | 129 | 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 | 130 | 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 | 131 | 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 | 132 | 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 | 133 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 134 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 135 | 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 | 136 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 137 | 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 | 138 | 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 | 139 | 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: 
61343diff
changeset | 140 | by simp | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 141 | also have "... = bernpoly (Suc m) (n + 1) - bernpoly (Suc m) 0" | 
| 64267 | 142 | 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 | 143 | 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 | 144 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 145 | |
| 61343 | 146 | 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 | 147 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 148 | lemma binomial_unroll: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 149 | "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: 
61694diff
changeset | 150 | 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 | 151 | |
| 64267 | 152 | lemma sum_unroll: | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 153 | "(\<Sum>k\<le>n::nat. f k) = (if n = 0 then f 0 else f n + (\<Sum>k\<le>n - 1. f k))" | 
| 64267 | 154 | 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 | 155 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 156 | lemma bernoulli_unroll: | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 157 | "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 | 158 | 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 | 159 | |
| 63367 
6c731c8b7f03
simplified definitions of combinatorial functions
 haftmann parents: 
61694diff
changeset | 160 | lemmas unroll = binomial_unroll | 
| 64267 | 161 | 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 | 162 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 163 | 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 | 164 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 165 | 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 | 166 | 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 | 167 | by (auto simp add: sum_of_powers) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 168 | 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 | 169 | 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 | 170 | finally show ?thesis by simp | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 171 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 172 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 173 | 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 | 174 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 175 | 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 | 176 | 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: 
61343diff
changeset | 177 | 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: 
61609diff
changeset | 178 | 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: 
61343diff
changeset | 179 | then show ?thesis by auto | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 180 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 181 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 182 | 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 | 183 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 184 | 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 | 185 | 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 | 186 | 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 | 187 | 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 | 188 | 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 | 189 | by (auto simp add: sum_of_powers) | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 190 | also have "... = ((n ^ 2 + n) / 2) ^ 2" | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 191 | 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: 
61649diff
changeset | 192 | 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 | 193 | 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: 
61343diff
changeset | 194 | |
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 195 | 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 | 196 | proof - | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 197 | 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 | 198 | 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: 
61343diff
changeset | 199 | 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: 
61609diff
changeset | 200 | 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: 
61343diff
changeset | 201 | then show ?thesis by auto | 
| 60603 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 202 | qed | 
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 203 | |
| 
09ecbd791d4a
add examples from Freek's top 100 theorems (thms 30, 73, 77)
 bulwahn parents: diff
changeset | 204 | end |