| author | wenzelm | 
| Thu, 20 Aug 2015 17:39:07 +0200 | |
| changeset 60986 | 077f663b6c24 | 
| parent 60758 | d8d85a8172b5 | 
| child 61076 | bdc1e2f0a86a | 
| permissions | -rw-r--r-- | 
| 59669 
de7792ea4090
renaming HOL/Fact.thy -> Binomial.thy
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 1 | (* Title : Binomial.thy | 
| 12196 | 2 | Author : Jacques D. Fleuriot | 
| 3 | Copyright : 1998 University of Cambridge | |
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
12196diff
changeset | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 5 | The integer version of factorial and other additions by Jeremy Avigad. | 
| 12196 | 6 | *) | 
| 7 | ||
| 60758 | 8 | section\<open>Factorial Function, Binomial Coefficients and Binomial Theorem\<close> | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
12196diff
changeset | 9 | |
| 59669 
de7792ea4090
renaming HOL/Fact.thy -> Binomial.thy
 paulson <lp15@cam.ac.uk> parents: 
59667diff
changeset | 10 | theory Binomial | 
| 33319 | 11 | imports Main | 
| 15131 | 12 | begin | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
12196diff
changeset | 13 | |
| 60758 | 14 | subsection \<open>Factorial\<close> | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 15 | |
| 59733 
cd945dc13bec
more general type class for factorial. Now allows code generation (?)
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 16 | fun fact :: "nat \<Rightarrow> ('a::semiring_char_0)"
 | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 17 | where "fact 0 = 1" | "fact (Suc n) = of_nat (Suc n) * fact n" | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 18 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 19 | lemmas fact_Suc = fact.simps(2) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 20 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 21 | lemma fact_1 [simp]: "fact 1 = 1" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 22 | by simp | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 23 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 24 | lemma fact_Suc_0 [simp]: "fact (Suc 0) = Suc 0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 25 | by simp | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 26 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 27 | lemma of_nat_fact [simp]: "of_nat (fact n) = fact n" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 28 | by (induct n) (auto simp add: algebra_simps of_nat_mult) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 29 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 30 | lemma fact_reduce: "n > 0 \<Longrightarrow> fact n = of_nat n * fact (n - 1)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 31 | by (cases n) auto | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 32 | |
| 59733 
cd945dc13bec
more general type class for factorial. Now allows code generation (?)
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 33 | lemma fact_nonzero [simp]: "fact n \<noteq> (0::'a::{semiring_char_0,semiring_no_zero_divisors})"
 | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 34 | apply (induct n) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 35 | apply auto | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 36 | using of_nat_eq_0_iff by fastforce | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 37 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 38 | lemma fact_mono_nat: "m \<le> n \<Longrightarrow> fact m \<le> (fact n :: nat)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 39 | by (induct n) (auto simp: le_Suc_eq) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 40 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 41 | context | 
| 60241 | 42 |   assumes "SORT_CONSTRAINT('a::linordered_semidom)"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 43 | begin | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 44 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 45 | lemma fact_mono: "m \<le> n \<Longrightarrow> fact m \<le> (fact n :: 'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 46 | by (metis of_nat_fact of_nat_le_iff fact_mono_nat) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 47 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 48 | lemma fact_ge_1 [simp]: "fact n \<ge> (1 :: 'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 49 | by (metis le0 fact.simps(1) fact_mono) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 50 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 51 | lemma fact_gt_zero [simp]: "fact n > (0 :: 'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 52 | using fact_ge_1 less_le_trans zero_less_one by blast | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 53 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 54 | lemma fact_ge_zero [simp]: "fact n \<ge> (0 :: 'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 55 | by (simp add: less_imp_le) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 56 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 57 | lemma fact_not_neg [simp]: "~ (fact n < (0 :: 'a))" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 58 | by (simp add: not_less_iff_gr_or_eq) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 59 | |
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 60 | lemma fact_le_power: | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 61 | "fact n \<le> (of_nat (n^n) ::'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 62 | proof (induct n) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 63 | case (Suc n) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 64 | then have *: "fact n \<le> (of_nat (Suc n ^ n) ::'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 65 | by (rule order_trans) (simp add: power_mono) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 66 | have "fact (Suc n) = (of_nat (Suc n) * fact n ::'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 67 | by (simp add: algebra_simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 68 | also have "... \<le> (of_nat (Suc n) * of_nat (Suc n ^ n) ::'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 69 | by (simp add: "*" ordered_comm_semiring_class.comm_mult_left_mono) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 70 | also have "... \<le> (of_nat (Suc n ^ Suc n) ::'a)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 71 | by (metis of_nat_mult order_refl power_Suc) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 72 | finally show ?case . | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 73 | qed simp | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 74 | |
| 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 75 | end | 
| 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 76 | |
| 60758 | 77 | text\<open>Note that @{term "fact 0 = fact 1"}\<close>
 | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 78 | lemma fact_less_mono_nat: "\<lbrakk>0 < m; m < n\<rbrakk> \<Longrightarrow> fact m < (fact n :: nat)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 79 | by (induct n) (auto simp: less_Suc_eq) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 80 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 81 | lemma fact_less_mono: | 
| 60241 | 82 | "\<lbrakk>0 < m; m < n\<rbrakk> \<Longrightarrow> fact m < (fact n :: 'a::linordered_semidom)" | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 83 | by (metis of_nat_fact of_nat_less_iff fact_less_mono_nat) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 84 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 85 | lemma fact_ge_Suc_0_nat [simp]: "fact n \<ge> Suc 0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 86 | by (metis One_nat_def fact_ge_1) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 87 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 88 | lemma dvd_fact: | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 89 | shows "1 \<le> m \<Longrightarrow> m \<le> n \<Longrightarrow> m dvd fact n" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 90 | by (induct n) (auto simp: dvdI le_Suc_eq) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 91 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 92 | lemma fact_altdef_nat: "fact n = \<Prod>{1..n}"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 93 | by (induct n) (auto simp: atLeastAtMostSuc_conv) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 94 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 95 | lemma fact_altdef: "fact n = setprod of_nat {1..n}"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 96 | by (induct n) (auto simp: atLeastAtMostSuc_conv) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
12196diff
changeset | 97 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 98 | lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd (fact m :: 'a :: {semiring_div,linordered_semidom})"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 99 | by (induct m) (auto simp: le_Suc_eq) | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 100 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 101 | lemma fact_mod: "m \<le> n \<Longrightarrow> fact n mod (fact m :: 'a :: {semiring_div,linordered_semidom}) = 0"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 102 | by (auto simp add: fact_dvd) | 
| 40033 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 103 | |
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 104 | lemma fact_div_fact: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 105 | assumes "m \<ge> n" | 
| 40033 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 106 |   shows "(fact m) div (fact n) = \<Prod>{n + 1..m}"
 | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 107 | proof - | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 108 | obtain d where "d = m - n" by auto | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 109 | from assms this have "m = n + d" by auto | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 110 |   have "fact (n + d) div (fact n) = \<Prod>{n + 1..n + d}"
 | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 111 | proof (induct d) | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 112 | case 0 | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 113 | show ?case by simp | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 114 | next | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 115 | case (Suc d') | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 116 | have "fact (n + Suc d') div fact n = Suc (n + d') * fact (n + d') div fact n" | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 117 | by simp | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 118 |     also from Suc.hyps have "... = Suc (n + d') * \<Prod>{n + 1..n + d'}"
 | 
| 40033 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 119 | unfolding div_mult1_eq[of _ "fact (n + d')"] by (simp add: fact_mod) | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 120 |     also have "... = \<Prod>{n + 1..n + Suc d'}"
 | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 121 | by (simp add: atLeastAtMostSuc_conv) | 
| 40033 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 122 | finally show ?case . | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 123 | qed | 
| 60758 | 124 | from this \<open>m = n + d\<close> show ?thesis by simp | 
| 40033 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 125 | qed | 
| 
84200d970bf0
added some facts about factorial and dvd, div and mod
 bulwahn parents: 
35644diff
changeset | 126 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 127 | lemma fact_num_eq_if: | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 128 | "fact m = (if m=0 then 1 else of_nat m * fact (m - 1))" | 
| 32036 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 129 | by (cases m) auto | 
| 
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
 avigad parents: 
30242diff
changeset | 130 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 131 | lemma fact_eq_rev_setprod_nat: "fact k = (\<Prod>i<k. k - i)" | 
| 50224 | 132 | unfolding fact_altdef_nat | 
| 57129 
7edb7550663e
introduce more powerful reindexing rules for big operators
 hoelzl parents: 
57113diff
changeset | 133 | by (rule setprod.reindex_bij_witness[where i="\<lambda>i. k - i" and j="\<lambda>i. k - i"]) auto | 
| 50224 | 134 | |
| 50240 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 135 | lemma fact_div_fact_le_pow: | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 136 | assumes "r \<le> n" shows "fact n div fact (n - r) \<le> n ^ r" | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 137 | proof - | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 138 |   have "\<And>r. r \<le> n \<Longrightarrow> \<Prod>{n - r..n} = (n - r) * \<Prod>{Suc (n - r)..n}"
 | 
| 57418 | 139 | by (subst setprod.insert[symmetric]) (auto simp: atLeastAtMost_insertL) | 
| 50240 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 140 | with assms show ?thesis | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 141 | by (induct r rule: nat.induct) (auto simp add: fact_div_fact Suc_diff_Suc mult_le_mono) | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 142 | qed | 
| 
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
 hoelzl parents: 
50224diff
changeset | 143 | |
| 60758 | 144 | lemma fact_numeral: --\<open>Evaluation for specific numerals\<close> | 
| 57113 
7e95523302e6
New theorems to enable the simplification of certain functions when applied to specific natural number constants (such as 4)
 paulson <lp15@cam.ac.uk> parents: 
50240diff
changeset | 145 | "fact (numeral k) = (numeral k) * (fact (pred_numeral k))" | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 146 | by (metis fact.simps(2) numeral_eq_Suc of_nat_numeral) | 
| 57113 
7e95523302e6
New theorems to enable the simplification of certain functions when applied to specific natural number constants (such as 4)
 paulson <lp15@cam.ac.uk> parents: 
50240diff
changeset | 147 | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 148 | |
| 60758 | 149 | text \<open>This development is based on the work of Andy Gordon and | 
| 150 | Florian Kammueller.\<close> | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 151 | |
| 60758 | 152 | subsection \<open>Basic definitions and lemmas\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 153 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 154 | primrec binomial :: "nat \<Rightarrow> nat \<Rightarrow> nat" (infixl "choose" 65) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 155 | where | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 156 | "0 choose k = (if k = 0 then 1 else 0)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 157 | | "Suc n choose k = (if k = 0 then 1 else (n choose (k - 1)) + (n choose k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 158 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 159 | lemma binomial_n_0 [simp]: "(n choose 0) = 1" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 160 | by (cases n) simp_all | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 161 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 162 | lemma binomial_0_Suc [simp]: "(0 choose Suc k) = 0" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 163 | by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 164 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 165 | lemma binomial_Suc_Suc [simp]: "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 166 | by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 167 | |
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 168 | lemma choose_reduce_nat: | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 169 | "0 < (n::nat) \<Longrightarrow> 0 < k \<Longrightarrow> | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 170 | (n choose k) = ((n - 1) choose (k - 1)) + ((n - 1) choose k)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 171 | by (metis Suc_diff_1 binomial.simps(2) neq0_conv) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 172 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 173 | lemma binomial_eq_0: "n < k \<Longrightarrow> n choose k = 0" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 174 | by (induct n arbitrary: k) auto | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 175 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 176 | declare binomial.simps [simp del] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 177 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 178 | lemma binomial_n_n [simp]: "n choose n = 1" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 179 | by (induct n) (simp_all add: binomial_eq_0) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 180 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 181 | lemma binomial_Suc_n [simp]: "Suc n choose n = Suc n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 182 | by (induct n) simp_all | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 183 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 184 | lemma binomial_1 [simp]: "n choose Suc 0 = n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 185 | by (induct n) simp_all | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 186 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 187 | lemma zero_less_binomial: "k \<le> n \<Longrightarrow> n choose k > 0" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 188 | by (induct n k rule: diff_induct) simp_all | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 189 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 190 | lemma binomial_eq_0_iff [simp]: "n choose k = 0 \<longleftrightarrow> n < k" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 191 | by (metis binomial_eq_0 less_numeral_extra(3) not_less zero_less_binomial) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 192 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 193 | lemma zero_less_binomial_iff [simp]: "n choose k > 0 \<longleftrightarrow> k \<le> n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 194 | by (metis binomial_eq_0_iff not_less0 not_less zero_less_binomial) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 195 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 196 | lemma Suc_times_binomial_eq: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 197 | "Suc n * (n choose k) = (Suc n choose Suc k) * Suc k" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 198 | apply (induct n arbitrary: k, simp add: binomial.simps) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 199 | apply (case_tac k) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 200 | apply (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq binomial_eq_0) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 201 | done | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 202 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 203 | lemma binomial_le_pow2: "n choose k \<le> 2^n" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 204 | apply (induction n arbitrary: k) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 205 | apply (simp add: binomial.simps) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 206 | apply (case_tac k) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 207 | apply (auto simp: power_Suc) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 208 | by (simp add: add_le_mono mult_2) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 209 | |
| 60758 | 210 | text\<open>The absorption property\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 211 | lemma Suc_times_binomial: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 212 | "Suc k * (Suc n choose Suc k) = Suc n * (n choose k)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 213 | using Suc_times_binomial_eq by auto | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 214 | |
| 60758 | 215 | text\<open>This is the well-known version of absorption, but it's harder to use because of the | 
| 216 | need to reason about division.\<close> | |
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 217 | lemma binomial_Suc_Suc_eq_times: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 218 | "(Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 219 | by (simp add: Suc_times_binomial_eq del: mult_Suc mult_Suc_right) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 220 | |
| 60758 | 221 | text\<open>Another version of absorption, with -1 instead of Suc.\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 222 | lemma times_binomial_minus1_eq: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 223 | "0 < k \<Longrightarrow> k * (n choose k) = n * ((n - 1) choose (k - 1))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 224 | using Suc_times_binomial_eq [where n = "n - 1" and k = "k - 1"] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 225 | by (auto split add: nat_diff_split) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 226 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 227 | |
| 60758 | 228 | subsection \<open>Combinatorial theorems involving @{text "choose"}\<close>
 | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 229 | |
| 60758 | 230 | text \<open>By Florian Kamm\"uller, tidied by LCP.\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 231 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 232 | lemma card_s_0_eq_empty: "finite A \<Longrightarrow> card {B. B \<subseteq> A & card B = 0} = 1"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 233 | by (simp cong add: conj_cong add: finite_subset [THEN card_0_eq]) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 234 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 235 | lemma choose_deconstruct: "finite M \<Longrightarrow> x \<notin> M \<Longrightarrow> | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 236 |     {s. s \<subseteq> insert x M \<and> card s = Suc k} =
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 237 |     {s. s \<subseteq> M \<and> card s = Suc k} \<union> {s. \<exists>t. t \<subseteq> M \<and> card t = k \<and> s = insert x t}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 238 | apply safe | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 239 | apply (auto intro: finite_subset [THEN card_insert_disjoint]) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 240 | by (metis (full_types) Diff_insert_absorb Set.set_insert Zero_neq_Suc card_Diff_singleton_if | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 241 | card_eq_0_iff diff_Suc_1 in_mono subset_insert_iff) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 242 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 243 | lemma finite_bex_subset [simp]: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 244 | assumes "finite B" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 245 |     and "\<And>A. A \<subseteq> B \<Longrightarrow> finite {x. P x A}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 246 |   shows "finite {x. \<exists>A \<subseteq> B. P x A}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 247 | by (metis (no_types) assms finite_Collect_bounded_ex finite_Collect_subsets) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 248 | |
| 60758 | 249 | text\<open>There are as many subsets of @{term A} having cardinality @{term k}
 | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 250 | as there are sets obtained from the former by inserting a fixed element | 
| 60758 | 251 |  @{term x} into each.\<close>
 | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 252 | lemma constr_bij: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 253 | "finite A \<Longrightarrow> x \<notin> A \<Longrightarrow> | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 254 |     card {B. \<exists>C. C \<subseteq> A \<and> card C = k \<and> B = insert x C} =
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 255 |     card {B. B \<subseteq> A & card(B) = k}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 256 |   apply (rule card_bij_eq [where f = "\<lambda>s. s - {x}" and g = "insert x"])
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 257 | apply (auto elim!: equalityE simp add: inj_on_def) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 258 | apply (metis card_Diff_singleton_if finite_subset in_mono) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 259 | done | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 260 | |
| 60758 | 261 | text \<open> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 262 | Main theorem: combinatorial statement about number of subsets of a set. | 
| 60758 | 263 | \<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 264 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 265 | theorem n_subsets: "finite A \<Longrightarrow> card {B. B \<subseteq> A \<and> card B = k} = (card A choose k)"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 266 | proof (induct k arbitrary: A) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 267 | case 0 then show ?case by (simp add: card_s_0_eq_empty) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 268 | next | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 269 | case (Suc k) | 
| 60758 | 270 | show ?case using \<open>finite A\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 271 | proof (induct A) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 272 | case empty show ?case by (simp add: card_s_0_eq_empty) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 273 | next | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 274 | case (insert x A) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 275 | then show ?case using Suc.hyps | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 276 | apply (simp add: card_s_0_eq_empty choose_deconstruct) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 277 | apply (subst card_Un_disjoint) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 278 | prefer 4 apply (force simp add: constr_bij) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 279 | prefer 3 apply force | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 280 | prefer 2 apply (blast intro: finite_Pow_iff [THEN iffD2] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 281 | finite_subset [of _ "Pow (insert x F)" for F]) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 282 | apply (blast intro: finite_Pow_iff [THEN iffD2, THEN [2] finite_subset]) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 283 | done | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 284 | qed | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 285 | qed | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 286 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 287 | |
| 60758 | 288 | subsection \<open>The binomial theorem (courtesy of Tobias Nipkow):\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 289 | |
| 60758 | 290 | text\<open>Avigad's version, generalized to any commutative ring\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 291 | theorem binomial_ring: "(a+b::'a::{comm_ring_1,power})^n =
 | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 292 | (\<Sum>k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" (is "?P n") | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 293 | proof (induct n) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 294 | case 0 then show "?P 0" by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 295 | next | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 296 | case (Suc n) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 297 |   have decomp: "{0..n+1} = {0} Un {n+1} Un {1..n}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 298 | by auto | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 299 |   have decomp2: "{0..n} = {0} Un {1..n}"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 300 | by auto | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 301 | have "(a+b)^(n+1) = | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 302 | (a+b) * (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 303 | using Suc.hyps by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 304 | also have "\<dots> = a*(\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k)) + | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 305 | b*(\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 306 | by (rule distrib_right) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 307 | also have "\<dots> = (\<Sum>k=0..n. of_nat (n choose k) * a^(k+1) * b^(n-k)) + | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 308 | (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n-k+1))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 309 | by (auto simp add: setsum_right_distrib ac_simps) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 310 | also have "\<dots> = (\<Sum>k=0..n. of_nat (n choose k) * a^k * b^(n+1-k)) + | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 311 | (\<Sum>k=1..n+1. of_nat (n choose (k - 1)) * a^k * b^(n+1-k))" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 312 | by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 313 | del:setsum_cl_ivl_Suc) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 314 | also have "\<dots> = a^(n+1) + b^(n+1) + | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 315 | (\<Sum>k=1..n. of_nat (n choose (k - 1)) * a^k * b^(n+1-k)) + | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 316 | (\<Sum>k=1..n. of_nat (n choose k) * a^k * b^(n+1-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 317 | by (simp add: decomp2) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 318 | also have | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 319 | "\<dots> = a^(n+1) + b^(n+1) + | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 320 | (\<Sum>k=1..n. of_nat(n+1 choose k) * a^k * b^(n+1-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 321 | by (auto simp add: field_simps setsum.distrib [symmetric] choose_reduce_nat) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 322 | also have "\<dots> = (\<Sum>k=0..n+1. of_nat (n+1 choose k) * a^k * b^(n+1-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 323 | using decomp by (simp add: field_simps) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 324 | finally show "?P (Suc n)" by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 325 | qed | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 326 | |
| 60758 | 327 | text\<open>Original version for the naturals\<close> | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 328 | corollary binomial: "(a+b::nat)^n = (\<Sum>k=0..n. (of_nat (n choose k)) * a^k * b^(n-k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 329 | using binomial_ring [of "int a" "int b" n] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 330 | by (simp only: of_nat_add [symmetric] of_nat_mult [symmetric] of_nat_power [symmetric] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 331 | of_nat_setsum [symmetric] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 332 | of_nat_eq_iff of_nat_id) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 333 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 334 | lemma binomial_fact_lemma: "k \<le> n \<Longrightarrow> fact k * fact (n - k) * (n choose k) = fact n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 335 | proof (induct n arbitrary: k rule: nat_less_induct) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 336 | fix n k assume H: "\<forall>m<n. \<forall>x\<le>m. fact x * fact (m - x) * (m choose x) = | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 337 | fact m" and kn: "k \<le> n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 338 | let ?ths = "fact k * fact (n - k) * (n choose k) = fact n" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 339 |   { assume "n=0" then have ?ths using kn by simp }
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 340 | moreover | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 341 |   { assume "k=0" then have ?ths using kn by simp }
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 342 | moreover | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 343 |   { assume nk: "n=k" then have ?ths by simp }
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 344 | moreover | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 345 |   { fix m h assume n: "n = Suc m" and h: "k = Suc h" and hm: "h < m"
 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 346 | from n have mn: "m < n" by arith | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 347 | from hm have hm': "h \<le> m" by arith | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 348 | from hm h n kn have km: "k \<le> m" by arith | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 349 | have "m - h = Suc (m - Suc h)" using h km hm by arith | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 350 | with km h have th0: "fact (m - h) = (m - h) * fact (m - k)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 351 | by simp | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 352 | from n h th0 | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 353 | have "fact k * fact (n - k) * (n choose k) = | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 354 | k * (fact h * fact (m - h) * (m choose h)) + | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 355 | (m - h) * (fact k * fact (m - k) * (m choose k))" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 356 | by (simp add: field_simps) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 357 | also have "\<dots> = (k + (m - h)) * fact m" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 358 | using H[rule_format, OF mn hm'] H[rule_format, OF mn km] | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 359 | by (simp add: field_simps) | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 360 | finally have ?ths using h n km by simp } | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 361 | moreover have "n=0 \<or> k = 0 \<or> k = n \<or> (\<exists>m h. n = Suc m \<and> k = Suc h \<and> h < m)" | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 362 | using kn by presburger | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 363 | ultimately show ?ths by blast | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 364 | qed | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 365 | |
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 366 | lemma binomial_fact: | 
| 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 367 | assumes kn: "k \<le> n" | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 368 | shows "(of_nat (n choose k) :: 'a::field_char_0) = | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 369 | (fact n) / (fact k * fact(n - k))" | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 370 | using binomial_fact_lemma[OF kn] | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 371 | apply (simp add: field_simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 372 | by (metis mult.commute of_nat_fact of_nat_mult) | 
| 59658 
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
 paulson <lp15@cam.ac.uk> parents: 
58889diff
changeset | 373 | |
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 374 | lemma choose_row_sum: "(\<Sum>k=0..n. n choose k) = 2^n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 375 | using binomial [of 1 "1" n] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 376 | by (simp add: numeral_2_eq_2) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 377 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 378 | lemma sum_choose_lower: "(\<Sum>k=0..n. (r+k) choose k) = Suc (r+n) choose n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 379 | by (induct n) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 380 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 381 | lemma sum_choose_upper: "(\<Sum>k=0..n. k choose m) = Suc n choose Suc m" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 382 | by (induct n) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 383 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 384 | lemma natsum_reverse_index: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 385 | fixes m::nat | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 386 | shows "(\<And>k. m \<le> k \<Longrightarrow> k \<le> n \<Longrightarrow> g k = f (m + n - k)) \<Longrightarrow> (\<Sum>k=m..n. f k) = (\<Sum>k=m..n. g k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 387 | by (rule setsum.reindex_bij_witness[where i="\<lambda>k. m+n-k" and j="\<lambda>k. m+n-k"]) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 388 | |
| 60758 | 389 | text\<open>NW diagonal sum property\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 390 | lemma sum_choose_diagonal: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 391 | assumes "m\<le>n" shows "(\<Sum>k=0..m. (n-k) choose (m-k)) = Suc n choose m" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 392 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 393 | have "(\<Sum>k=0..m. (n-k) choose (m-k)) = (\<Sum>k=0..m. (n-m+k) choose k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 394 | by (rule natsum_reverse_index) (simp add: assms) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 395 | also have "... = Suc (n-m+m) choose m" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 396 | by (rule sum_choose_lower) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 397 | also have "... = Suc n choose m" using assms | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 398 | by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 399 | finally show ?thesis . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 400 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 401 | |
| 60758 | 402 | subsection\<open>Pochhammer's symbol : generalized rising factorial\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 403 | |
| 60758 | 404 | text \<open>See @{url "http://en.wikipedia.org/wiki/Pochhammer_symbol"}\<close>
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 405 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 406 | definition "pochhammer (a::'a::comm_semiring_1) n = | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 407 |   (if n = 0 then 1 else setprod (\<lambda>n. a + of_nat n) {0 .. n - 1})"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 408 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 409 | lemma pochhammer_0 [simp]: "pochhammer a 0 = 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 410 | by (simp add: pochhammer_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 411 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 412 | lemma pochhammer_1 [simp]: "pochhammer a 1 = a" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 413 | by (simp add: pochhammer_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 414 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 415 | lemma pochhammer_Suc0 [simp]: "pochhammer a (Suc 0) = a" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 416 | by (simp add: pochhammer_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 417 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 418 | lemma pochhammer_Suc_setprod: "pochhammer a (Suc n) = setprod (\<lambda>n. a + of_nat n) {0 .. n}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 419 | by (simp add: pochhammer_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 420 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 421 | lemma setprod_nat_ivl_Suc: "setprod f {0 .. Suc n} = setprod f {0..n} * f (Suc n)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 422 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 423 |   have "{0..Suc n} = {0..n} \<union> {Suc n}" by auto
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 424 | then show ?thesis by (simp add: field_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 425 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 426 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 427 | lemma setprod_nat_ivl_1_Suc: "setprod f {0 .. Suc n} = f 0 * setprod f {1.. Suc n}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 428 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 429 |   have "{0..Suc n} = {0} \<union> {1 .. Suc n}" by auto
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 430 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 431 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 432 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 433 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 434 | lemma pochhammer_Suc: "pochhammer a (Suc n) = pochhammer a n * (a + of_nat n)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 435 | proof (cases n) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 436 | case 0 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 437 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 438 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 439 | case (Suc n) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 440 | show ?thesis unfolding Suc pochhammer_Suc_setprod setprod_nat_ivl_Suc .. | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 441 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 442 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 443 | lemma pochhammer_rec: "pochhammer a (Suc n) = a * pochhammer (a + 1) n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 444 | proof (cases "n = 0") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 445 | case True | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 446 | then show ?thesis by (simp add: pochhammer_Suc_setprod) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 447 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 448 | case False | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 449 |   have *: "finite {1 .. n}" "0 \<notin> {1 .. n}" by auto
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 450 |   have eq: "insert 0 {1 .. n} = {0..n}" by auto
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 451 |   have **: "(\<Prod>n\<in>{1\<Colon>nat..n}. a + of_nat n) = (\<Prod>n\<in>{0\<Colon>nat..n - 1}. a + 1 + of_nat n)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 452 | apply (rule setprod.reindex_cong [where l = Suc]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 453 | using False | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 454 | apply (auto simp add: fun_eq_iff field_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 455 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 456 | show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 457 | apply (simp add: pochhammer_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 458 | unfolding setprod.insert [OF *, unfolded eq] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 459 | using ** apply (simp add: field_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 460 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 461 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 462 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 463 | lemma pochhammer_fact: "fact n = pochhammer 1 n" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 464 | unfolding fact_altdef | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 465 | apply (cases n) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 466 | apply (simp_all add: of_nat_setprod pochhammer_Suc_setprod) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 467 | apply (rule setprod.reindex_cong [where l = Suc]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 468 | apply (auto simp add: fun_eq_iff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 469 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 470 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 471 | lemma pochhammer_of_nat_eq_0_lemma: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 472 | assumes "k > n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 473 | shows "pochhammer (- (of_nat n :: 'a:: idom)) k = 0" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 474 | proof (cases "n = 0") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 475 | case True | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 476 | then show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 477 | using assms by (cases k) (simp_all add: pochhammer_rec) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 478 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 479 | case False | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 480 | from assms obtain h where "k = Suc h" by (cases k) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 481 | then show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 482 | by (simp add: pochhammer_Suc_setprod) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 483 | (metis Suc_leI Suc_le_mono assms atLeastAtMost_iff less_eq_nat.simps(1)) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 484 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 485 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 486 | lemma pochhammer_of_nat_eq_0_lemma': | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 487 | assumes kn: "k \<le> n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 488 |   shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k \<noteq> 0"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 489 | proof (cases k) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 490 | case 0 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 491 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 492 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 493 | case (Suc h) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 494 | then show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 495 | apply (simp add: pochhammer_Suc_setprod) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 496 | using Suc kn apply (auto simp add: algebra_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 497 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 498 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 499 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 500 | lemma pochhammer_of_nat_eq_0_iff: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 501 |   shows "pochhammer (- (of_nat n :: 'a:: {idom,ring_char_0})) k = 0 \<longleftrightarrow> k > n"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 502 | (is "?l = ?r") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 503 | using pochhammer_of_nat_eq_0_lemma[of n k, where ?'a='a] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 504 | pochhammer_of_nat_eq_0_lemma'[of k n, where ?'a = 'a] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 505 | by (auto simp add: not_le[symmetric]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 506 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 507 | lemma pochhammer_eq_0_iff: "pochhammer a n = (0::'a::field_char_0) \<longleftrightarrow> (\<exists>k < n. a = - of_nat k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 508 | apply (auto simp add: pochhammer_of_nat_eq_0_iff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 509 | apply (cases n) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 510 | apply (auto simp add: pochhammer_def algebra_simps group_add_class.eq_neg_iff_add_eq_0) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 511 | apply (metis leD not_less_eq) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 512 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 513 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 514 | lemma pochhammer_eq_0_mono: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 515 | "pochhammer a n = (0::'a::field_char_0) \<Longrightarrow> m \<ge> n \<Longrightarrow> pochhammer a m = 0" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 516 | unfolding pochhammer_eq_0_iff by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 517 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 518 | lemma pochhammer_neq_0_mono: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 519 | "pochhammer a m \<noteq> (0::'a::field_char_0) \<Longrightarrow> m \<ge> n \<Longrightarrow> pochhammer a n \<noteq> 0" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 520 | unfolding pochhammer_eq_0_iff by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 521 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 522 | lemma pochhammer_minus: | 
| 59862 | 523 | "pochhammer (- b) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (b - of_nat k + 1) k" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 524 | proof (cases k) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 525 | case 0 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 526 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 527 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 528 | case (Suc h) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 529 | have eq: "((- 1) ^ Suc h :: 'a) = (\<Prod>i=0..h. - 1)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 530 |     using setprod_constant[where A="{0 .. h}" and y="- 1 :: 'a"]
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 531 | by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 532 | show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 533 | unfolding Suc pochhammer_Suc_setprod eq setprod.distrib[symmetric] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 534 | by (rule setprod.reindex_bij_witness[where i="op - h" and j="op - h"]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 535 | (auto simp: of_nat_diff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 536 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 537 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 538 | lemma pochhammer_minus': | 
| 59862 | 539 | "pochhammer (b - of_nat k + 1) k = ((- 1) ^ k :: 'a::comm_ring_1) * pochhammer (- b) k" | 
| 540 | unfolding pochhammer_minus[where b=b] | |
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 541 | unfolding mult.assoc[symmetric] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 542 | unfolding power_add[symmetric] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 543 | by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 544 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 545 | lemma pochhammer_same: "pochhammer (- of_nat n) n = | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 546 |     ((- 1) ^ n :: 'a::{semiring_char_0,comm_ring_1,semiring_no_zero_divisors}) * (fact n)"
 | 
| 59862 | 547 | unfolding pochhammer_minus | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 548 | by (simp add: of_nat_diff pochhammer_fact) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 549 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 550 | |
| 60758 | 551 | subsection\<open>Generalized binomial coefficients\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 552 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 553 | definition gbinomial :: "'a::field_char_0 \<Rightarrow> nat \<Rightarrow> 'a" (infixl "gchoose" 65) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 554 | where "a gchoose n = | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 555 |     (if n = 0 then 1 else (setprod (\<lambda>i. a - of_nat i) {0 .. n - 1}) / (fact n))"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 556 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 557 | lemma gbinomial_0 [simp]: "a gchoose 0 = 1" "0 gchoose (Suc n) = 0" | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 558 | by (simp_all add: gbinomial_def) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 559 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 560 | lemma gbinomial_pochhammer: "a gchoose n = (- 1) ^ n * pochhammer (- a) n / (fact n)" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 561 | proof (cases "n = 0") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 562 | case True | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 563 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 564 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 565 | case False | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 566 |   from this setprod_constant[of "{0 .. n - 1}" "- (1:: 'a)"]
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 567 |   have eq: "(- (1\<Colon>'a)) ^ n = setprod (\<lambda>i. - 1) {0 .. n - 1}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 568 | by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 569 | from False show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 570 | by (simp add: pochhammer_def gbinomial_def field_simps | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 571 | eq setprod.distrib[symmetric]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 572 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 573 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 574 | lemma binomial_gbinomial: | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 575 | "of_nat (n choose k) = (of_nat n gchoose k :: 'a::field_char_0)" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 576 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 577 |   { assume kn: "k > n"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 578 | then have ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 579 | by (subst binomial_eq_0[OF kn]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 580 | (simp add: gbinomial_pochhammer field_simps pochhammer_of_nat_eq_0_iff) } | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 581 | moreover | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 582 |   { assume "k=0" then have ?thesis by simp }
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 583 | moreover | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 584 |   { assume kn: "k \<le> n" and k0: "k\<noteq> 0"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 585 | from k0 obtain h where h: "k = Suc h" by (cases k) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 586 | from h | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 587 |     have eq:"(- 1 :: 'a) ^ k = setprod (\<lambda>i. - 1) {0..h}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 588 | by (subst setprod_constant) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 589 |     have eq': "(\<Prod>i\<in>{0..h}. of_nat n + - (of_nat i :: 'a)) = (\<Prod>i\<in>{n - h..n}. of_nat i)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 590 | using h kn | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 591 | by (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 592 | (auto simp: of_nat_diff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 593 |     have th0: "finite {1..n - Suc h}" "finite {n - h .. n}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 594 |         "{1..n - Suc h} \<inter> {n - h .. n} = {}" and
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 595 |         eq3: "{1..n - Suc h} \<union> {n - h .. n} = {1..n}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 596 | using h kn by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 597 | from eq[symmetric] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 598 | have ?thesis using kn | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 599 | apply (simp add: binomial_fact[OF kn, where ?'a = 'a] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 600 | gbinomial_pochhammer field_simps pochhammer_Suc_setprod) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 601 | apply (simp add: pochhammer_Suc_setprod fact_altdef h | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 602 | of_nat_setprod setprod.distrib[symmetric] eq' del: One_nat_def power_Suc) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 603 | unfolding setprod.union_disjoint[OF th0, unfolded eq3, of "of_nat:: nat \<Rightarrow> 'a"] eq[unfolded h] | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 604 | unfolding mult.assoc | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 605 | unfolding setprod.distrib[symmetric] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 606 | apply simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 607 | apply (intro setprod.reindex_bij_witness[where i="op - n" and j="op - n"]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 608 | apply (auto simp: of_nat_diff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 609 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 610 | } | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 611 | moreover | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 612 | have "k > n \<or> k = 0 \<or> (k \<le> n \<and> k \<noteq> 0)" by arith | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 613 | ultimately show ?thesis by blast | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 614 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 615 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 616 | lemma gbinomial_1[simp]: "a gchoose 1 = a" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 617 | by (simp add: gbinomial_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 618 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 619 | lemma gbinomial_Suc0[simp]: "a gchoose (Suc 0) = a" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 620 | by (simp add: gbinomial_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 621 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 622 | lemma gbinomial_mult_1: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 623 | fixes a :: "'a :: field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 624 | shows "a * (a gchoose n) = | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 625 | of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" (is "?l = ?r") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 626 | proof - | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 627 | have "?r = ((- 1) ^n * pochhammer (- a) n / (fact n)) * (of_nat n - (- a + of_nat n))" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 628 | unfolding gbinomial_pochhammer | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 629 | pochhammer_Suc of_nat_mult right_diff_distrib power_Suc | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 630 | apply (simp del: of_nat_Suc fact.simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 631 | apply (auto simp add: field_simps simp del: of_nat_Suc) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 632 | done | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 633 | also have "\<dots> = ?l" unfolding gbinomial_pochhammer | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 634 | by (simp add: field_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 635 | finally show ?thesis .. | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 636 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 637 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 638 | lemma gbinomial_mult_1': | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 639 | fixes a :: "'a :: field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 640 | shows "(a gchoose n) * a = of_nat n * (a gchoose n) + of_nat (Suc n) * (a gchoose (Suc n))" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 641 | by (simp add: mult.commute gbinomial_mult_1) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 642 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 643 | lemma gbinomial_Suc: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 644 |     "a gchoose (Suc k) = (setprod (\<lambda>i. a - of_nat i) {0 .. k}) / (fact (Suc k))"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 645 | by (simp add: gbinomial_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 646 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 647 | lemma gbinomial_mult_fact: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 648 | fixes a :: "'a::field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 649 | shows | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 650 | "fact (Suc k) * (a gchoose (Suc k)) = | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 651 |     (setprod (\<lambda>i. a - of_nat i) {0 .. k})"
 | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 652 | by (simp_all add: gbinomial_Suc field_simps del: fact.simps) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 653 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 654 | lemma gbinomial_mult_fact': | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 655 | fixes a :: "'a::field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 656 |   shows "(a gchoose (Suc k)) * fact (Suc k) = (setprod (\<lambda>i. a - of_nat i) {0 .. k})"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 657 | using gbinomial_mult_fact[of k a] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 658 | by (subst mult.commute) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 659 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 660 | lemma gbinomial_Suc_Suc: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 661 | fixes a :: "'a :: field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 662 | shows "(a + 1) gchoose (Suc k) = a gchoose k + (a gchoose (Suc k))" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 663 | proof (cases k) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 664 | case 0 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 665 | then show ?thesis by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 666 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 667 | case (Suc h) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 668 |   have eq0: "(\<Prod>i\<in>{1..k}. (a + 1) - of_nat i) = (\<Prod>i\<in>{0..h}. a - of_nat i)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 669 | apply (rule setprod.reindex_cong [where l = Suc]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 670 | using Suc | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 671 | apply auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 672 | done | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 673 | have "fact (Suc k) * (a gchoose k + (a gchoose (Suc k))) = | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 674 | (a gchoose Suc h) * (fact (Suc (Suc h))) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 675 | (a gchoose Suc (Suc h)) * (fact (Suc (Suc h)))" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 676 | by (simp add: Suc field_simps del: fact.simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 677 | also have "... = (a gchoose Suc h) * of_nat (Suc (Suc h) * fact (Suc h)) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 678 | (\<Prod>i = 0..Suc h. a - of_nat i)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 679 | by (metis fact.simps(2) gbinomial_mult_fact' of_nat_fact of_nat_id) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 680 | also have "... = (fact (Suc h) * (a gchoose Suc h)) * of_nat (Suc (Suc h)) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 681 | (\<Prod>i = 0..Suc h. a - of_nat i)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 682 | by (simp only: fact.simps(2) mult.commute mult.left_commute of_nat_fact of_nat_id of_nat_mult) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 683 | also have "... = of_nat (Suc (Suc h)) * (\<Prod>i = 0..h. a - of_nat i) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 684 | (\<Prod>i = 0..Suc h. a - of_nat i)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 685 | by (metis gbinomial_mult_fact mult.commute) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 686 | also have "... = (\<Prod>i = 0..Suc h. a - of_nat i) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 687 | (of_nat h * (\<Prod>i = 0..h. a - of_nat i) + 2 * (\<Prod>i = 0..h. a - of_nat i))" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 688 | by (simp add: field_simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 689 | also have "... = | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 690 |     ((a gchoose Suc h) * (fact (Suc h)) * of_nat (Suc k)) + (\<Prod>i\<in>{0\<Colon>nat..Suc h}. a - of_nat i)"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 691 | unfolding gbinomial_mult_fact' | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 692 | by (simp add: comm_semiring_class.distrib field_simps Suc) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 693 |   also have "\<dots> = (\<Prod>i\<in>{0..h}. a - of_nat i) * (a + 1)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 694 | unfolding gbinomial_mult_fact' setprod_nat_ivl_Suc | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 695 | by (simp add: field_simps Suc) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 696 |   also have "\<dots> = (\<Prod>i\<in>{0..k}. (a + 1) - of_nat i)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 697 | using eq0 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 698 | by (simp add: Suc setprod_nat_ivl_1_Suc) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 699 | also have "\<dots> = (fact (Suc k)) * ((a + 1) gchoose (Suc k))" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 700 | unfolding gbinomial_mult_fact .. | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 701 | finally show ?thesis | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 702 | by (metis fact_nonzero mult_cancel_left) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 703 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 704 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 705 | lemma gbinomial_reduce_nat: | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 706 | fixes a :: "'a :: field_char_0" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 707 | shows "0 < k \<Longrightarrow> a gchoose k = (a - 1) gchoose (k - 1) + ((a - 1) gchoose k)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 708 | by (metis Suc_pred' diff_add_cancel gbinomial_Suc_Suc) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 709 | |
| 60141 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 710 | lemma gchoose_row_sum_weighted: | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 711 | fixes r :: "'a::field_char_0" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 712 | shows "(\<Sum>k = 0..m. (r gchoose k) * (r/2 - of_nat k)) = of_nat(Suc m) / 2 * (r gchoose (Suc m))" | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 713 | proof (induct m) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 714 | case 0 show ?case by simp | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 715 | next | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 716 | case (Suc m) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 717 | from Suc show ?case | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 718 | by (simp add: algebra_simps distrib gbinomial_mult_1) | 
| 
833adf7db7d8
New material, mostly about limits. Consolidation.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 719 | qed | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 720 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 721 | lemma binomial_symmetric: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 722 | assumes kn: "k \<le> n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 723 | shows "n choose k = n choose (n - k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 724 | proof- | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 725 | from kn have kn': "n - k \<le> n" by arith | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 726 | from binomial_fact_lemma[OF kn] binomial_fact_lemma[OF kn'] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 727 | have "fact k * fact (n - k) * (n choose k) = | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 728 | fact (n - k) * fact (n - (n - k)) * (n choose (n - k))" by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 729 | then show ?thesis using kn by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 730 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 731 | |
| 60758 | 732 | text\<open>Contributed by Manuel Eberl, generalised by LCP. | 
| 733 |   Alternative definition of the binomial coefficient as @{term "\<Prod>i<k. (n - i) / (k - i)"}\<close>
 | |
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 734 | lemma gbinomial_altdef_of_nat: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 735 | fixes k :: nat | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 736 |     and x :: "'a :: {field_char_0,field}"
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 737 | shows "x gchoose k = (\<Prod>i<k. (x - of_nat i) / of_nat (k - i) :: 'a)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 738 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 739 | have "(x gchoose k) = (\<Prod>i<k. x - of_nat i) / of_nat (fact k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 740 | unfolding gbinomial_def | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 741 | by (auto simp: gr0_conv_Suc lessThan_Suc_atMost atLeast0AtMost) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 742 | also have "\<dots> = (\<Prod>i<k. (x - of_nat i) / of_nat (k - i) :: 'a)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 743 | unfolding fact_eq_rev_setprod_nat of_nat_setprod | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 744 | by (auto simp add: setprod_dividef intro!: setprod.cong of_nat_diff[symmetric]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 745 | finally show ?thesis . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 746 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 747 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 748 | lemma gbinomial_ge_n_over_k_pow_k: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 749 | fixes k :: nat | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 750 | and x :: "'a :: linordered_field" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 751 | assumes "of_nat k \<le> x" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 752 | shows "(x / of_nat k :: 'a) ^ k \<le> x gchoose k" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 753 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 754 | have x: "0 \<le> x" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 755 | using assms of_nat_0_le_iff order_trans by blast | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 756 | have "(x / of_nat k :: 'a) ^ k = (\<Prod>i<k. x / of_nat k :: 'a)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 757 | by (simp add: setprod_constant) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 758 | also have "\<dots> \<le> x gchoose k" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 759 | unfolding gbinomial_altdef_of_nat | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 760 | proof (safe intro!: setprod_mono) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 761 | fix i :: nat | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 762 | assume ik: "i < k" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 763 | from assms have "x * of_nat i \<ge> of_nat (i * k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 764 | by (metis mult.commute mult_le_cancel_right of_nat_less_0_iff of_nat_mult) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 765 | then have "x * of_nat k - x * of_nat i \<le> x * of_nat k - of_nat (i * k)" by arith | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 766 | then have "x * of_nat (k - i) \<le> (x - of_nat i) * of_nat k" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 767 | using ik | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 768 | by (simp add: algebra_simps zero_less_mult_iff of_nat_diff of_nat_mult) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 769 | then have "x * of_nat (k - i) \<le> (x - of_nat i) * (of_nat k :: 'a)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 770 | unfolding of_nat_mult[symmetric] of_nat_le_iff . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 771 | with assms show "x / of_nat k \<le> (x - of_nat i) / (of_nat (k - i) :: 'a)" | 
| 60758 | 772 | using \<open>i < k\<close> by (simp add: field_simps) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 773 | qed (simp add: x zero_le_divide_iff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 774 | finally show ?thesis . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 775 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 776 | |
| 60758 | 777 | text\<open>Versions of the theorems above for the natural-number version of "choose"\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 778 | lemma binomial_altdef_of_nat: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 779 | fixes n k :: nat | 
| 60758 | 780 |     and x :: "'a :: {field_char_0,field}"  --\<open>the point is to constrain @{typ 'a}\<close>
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 781 | assumes "k \<le> n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 782 | shows "of_nat (n choose k) = (\<Prod>i<k. of_nat (n - i) / of_nat (k - i) :: 'a)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 783 | using assms | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 784 | by (simp add: gbinomial_altdef_of_nat binomial_gbinomial of_nat_diff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 785 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 786 | lemma binomial_ge_n_over_k_pow_k: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 787 | fixes k n :: nat | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59862diff
changeset | 788 | and x :: "'a :: linordered_field" | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 789 | assumes "k \<le> n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 790 | shows "(of_nat n / of_nat k :: 'a) ^ k \<le> of_nat (n choose k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 791 | by (simp add: assms gbinomial_ge_n_over_k_pow_k binomial_gbinomial of_nat_diff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 792 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 793 | lemma binomial_le_pow: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 794 | assumes "r \<le> n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 795 | shows "n choose r \<le> n ^ r" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 796 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 797 | have "n choose r \<le> fact n div fact (n - r)" | 
| 60758 | 798 | using \<open>r \<le> n\<close> by (subst binomial_fact_lemma[symmetric]) auto | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 799 | with fact_div_fact_le_pow [OF assms] show ?thesis by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 800 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 801 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 802 | lemma binomial_altdef_nat: "(k::nat) \<le> n \<Longrightarrow> | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 803 | n choose k = fact n div (fact k * fact (n - k))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 804 | by (subst binomial_fact_lemma [symmetric]) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 805 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 806 | lemma choose_dvd: "k \<le> n \<Longrightarrow> fact k * fact (n - k) dvd (fact n :: 'a :: {semiring_div,linordered_semidom})"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 807 | unfolding dvd_def | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 808 | apply (rule exI [where x="of_nat (n choose k)"]) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 809 | using binomial_fact_lemma [of k n, THEN arg_cong [where f = of_nat and 'b='a]] | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 810 | apply (auto simp: of_nat_mult) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 811 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 812 | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 813 | lemma fact_fact_dvd_fact: | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 814 |     "fact k * fact n dvd (fact (k+n) :: 'a :: {semiring_div,linordered_semidom})"
 | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 815 | by (metis add.commute add_diff_cancel_left' choose_dvd le_add2) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 816 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 817 | lemma choose_mult_lemma: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 818 | "((m+r+k) choose (m+k)) * ((m+k) choose k) = ((m+r+k) choose k) * ((m+r) choose m)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 819 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 820 | have "((m+r+k) choose (m+k)) * ((m+k) choose k) = | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 821 | fact (m+r + k) div (fact (m + k) * fact (m+r - m)) * (fact (m + k) div (fact k * fact m))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 822 | by (simp add: assms binomial_altdef_nat) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 823 | also have "... = fact (m+r+k) div (fact r * (fact k * fact m))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 824 | apply (subst div_mult_div_if_dvd) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 825 | apply (auto simp: algebra_simps fact_fact_dvd_fact) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 826 | apply (metis add.assoc add.commute fact_fact_dvd_fact) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 827 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 828 | also have "... = (fact (m+r+k) * fact (m+r)) div (fact r * (fact k * fact m) * fact (m+r))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 829 | apply (subst div_mult_div_if_dvd [symmetric]) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 830 | apply (auto simp add: algebra_simps) | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 831 | apply (metis fact_fact_dvd_fact dvd.order.trans nat_mult_dvd_cancel_disj) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 832 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 833 | also have "... = (fact (m+r+k) div (fact k * fact (m+r)) * (fact (m+r) div (fact r * fact m)))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 834 | apply (subst div_mult_div_if_dvd) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 835 | apply (auto simp: fact_fact_dvd_fact algebra_simps) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 836 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 837 | finally show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 838 | by (simp add: binomial_altdef_nat mult.commute) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 839 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 840 | |
| 60758 | 841 | text\<open>The "Subset of a Subset" identity\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 842 | lemma choose_mult: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 843 | assumes "k\<le>m" "m\<le>n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 844 | shows "(n choose m) * (m choose k) = (n choose k) * ((n-k) choose (m-k))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 845 | using assms choose_mult_lemma [of "m-k" "n-m" k] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 846 | by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 847 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 848 | |
| 60758 | 849 | subsection \<open>Binomial coefficients\<close> | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 850 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 851 | lemma choose_one: "(n::nat) choose 1 = n" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 852 | by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 853 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 854 | (*FIXME: messy and apparently unused*) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 855 | lemma binomial_induct [rule_format]: "(ALL (n::nat). P n n) \<longrightarrow> | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 856 | (ALL n. P (Suc n) 0) \<longrightarrow> (ALL n. (ALL k < n. P n k \<longrightarrow> P n (Suc k) \<longrightarrow> | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 857 | P (Suc n) (Suc k))) \<longrightarrow> (ALL k <= n. P n k)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 858 | apply (induct n) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 859 | apply auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 860 | apply (case_tac "k = 0") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 861 | apply auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 862 | apply (case_tac "k = Suc n") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 863 | apply auto | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59669diff
changeset | 864 | apply (metis Suc_le_eq fact.cases le_Suc_eq le_eq_less_or_eq) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 865 | done | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 866 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 867 | lemma card_UNION: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 868 | assumes "finite A" and "\<forall>k \<in> A. finite k" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 869 |   shows "card (\<Union>A) = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (- 1) ^ (card I + 1) * int (card (\<Inter>I)))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 870 | (is "?lhs = ?rhs") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 871 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 872 |   have "?rhs = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (- 1) ^ (card I + 1) * (\<Sum>_\<in>\<Inter>I. 1))" by simp
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 873 |   also have "\<dots> = nat (\<Sum>I | I \<subseteq> A \<and> I \<noteq> {}. (\<Sum>_\<in>\<Inter>I. (- 1) ^ (card I + 1)))" (is "_ = nat ?rhs")
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 874 | by(subst setsum_right_distrib) simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 875 |   also have "?rhs = (\<Sum>(I, _)\<in>Sigma {I. I \<subseteq> A \<and> I \<noteq> {}} Inter. (- 1) ^ (card I + 1))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 876 | using assms by(subst setsum.Sigma)(auto) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 877 |   also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:UNIV. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). (- 1) ^ (card I + 1))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 878 | by (rule setsum.reindex_cong [where l = "\<lambda>(x, y). (y, x)"]) (auto intro: inj_onI simp add: split_beta) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 879 |   also have "\<dots> = (\<Sum>(x, I)\<in>(SIGMA x:\<Union>A. {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}). (- 1) ^ (card I + 1))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 880 | using assms by(auto intro!: setsum.mono_neutral_cong_right finite_SigmaI2 intro: finite_subset[where B="\<Union>A"]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 881 |   also have "\<dots> = (\<Sum>x\<in>\<Union>A. (\<Sum>I|I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I. (- 1) ^ (card I + 1)))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 882 | using assms by(subst setsum.Sigma) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 883 | also have "\<dots> = (\<Sum>_\<in>\<Union>A. 1)" (is "setsum ?lhs _ = _") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 884 | proof(rule setsum.cong[OF refl]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 885 | fix x | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 886 | assume x: "x \<in> \<Union>A" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 887 |     def K \<equiv> "{X \<in> A. x \<in> X}"
 | 
| 60758 | 888 | with \<open>finite A\<close> have K: "finite K" by auto | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 889 |     let ?I = "\<lambda>i. {I. I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 890 |     have "inj_on snd (SIGMA i:{1..card A}. ?I i)"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 891 | using assms by(auto intro!: inj_onI) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 892 |     moreover have [symmetric]: "snd ` (SIGMA i:{1..card A}. ?I i) = {I. I \<subseteq> A \<and> I \<noteq> {} \<and> x \<in> \<Inter>I}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 893 | using assms by(auto intro!: rev_image_eqI[where x="(card a, a)" for a] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 894 | simp add: card_gt_0_iff[folded Suc_le_eq] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 895 | dest: finite_subset intro: card_mono) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 896 |     ultimately have "?lhs x = (\<Sum>(i, I)\<in>(SIGMA i:{1..card A}. ?I i). (- 1) ^ (i + 1))"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 897 | by (rule setsum.reindex_cong [where l = snd]) fastforce | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 898 | also have "\<dots> = (\<Sum>i=1..card A. (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. (- 1) ^ (i + 1)))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 899 | using assms by(subst setsum.Sigma) auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 900 | also have "\<dots> = (\<Sum>i=1..card A. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1))" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 901 | by(subst setsum_right_distrib) simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 902 | also have "\<dots> = (\<Sum>i=1..card K. (- 1) ^ (i + 1) * (\<Sum>I|I \<subseteq> K \<and> card I = i. 1))" (is "_ = ?rhs") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 903 | proof(rule setsum.mono_neutral_cong_right[rule_format]) | 
| 60758 | 904 |       show "{1..card K} \<subseteq> {1..card A}" using \<open>finite A\<close>
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 905 | by(auto simp add: K_def intro: card_mono) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 906 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 907 | fix i | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 908 |       assume "i \<in> {1..card A} - {1..card K}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 909 | hence i: "i \<le> card A" "card K < i" by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 910 |       have "{I. I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I} = {I. I \<subseteq> K \<and> card I = i}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 911 | by(auto simp add: K_def) | 
| 60758 | 912 |       also have "\<dots> = {}" using \<open>finite A\<close> i
 | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 913 | by(auto simp add: K_def dest: card_mono[rotated 1]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 914 | finally show "(- 1) ^ (i + 1) * (\<Sum>I | I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1 :: int) = 0" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 915 | by(simp only:) simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 916 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 917 | fix i | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 918 | have "(\<Sum>I | I \<subseteq> A \<and> card I = i \<and> x \<in> \<Inter>I. 1) = (\<Sum>I | I \<subseteq> K \<and> card I = i. 1 :: int)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 919 | (is "?lhs = ?rhs") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 920 | by(rule setsum.cong)(auto simp add: K_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 921 | thus "(- 1) ^ (i + 1) * ?lhs = (- 1) ^ (i + 1) * ?rhs" by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 922 | qed simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 923 |     also have "{I. I \<subseteq> K \<and> card I = 0} = {{}}" using assms
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 924 | by(auto simp add: card_eq_0_iff K_def dest: finite_subset) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 925 | hence "?rhs = (\<Sum>i = 0..card K. (- 1) ^ (i + 1) * (\<Sum>I | I \<subseteq> K \<and> card I = i. 1 :: int)) + 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 926 | by(subst (2) setsum_head_Suc)(simp_all ) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 927 | also have "\<dots> = (\<Sum>i = 0..card K. (- 1) * ((- 1) ^ i * int (card K choose i))) + 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 928 | using K by(subst n_subsets[symmetric]) simp_all | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 929 | also have "\<dots> = - (\<Sum>i = 0..card K. (- 1) ^ i * int (card K choose i)) + 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 930 | by(subst setsum_right_distrib[symmetric]) simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 931 | also have "\<dots> = - ((-1 + 1) ^ card K) + 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 932 | by(subst binomial_ring)(simp add: ac_simps) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 933 | also have "\<dots> = 1" using x K by(auto simp add: K_def card_gt_0_iff) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 934 | finally show "?lhs x = 1" . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 935 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 936 | also have "nat \<dots> = card (\<Union>A)" by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 937 | finally show ?thesis .. | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 938 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 939 | |
| 60758 | 940 | text\<open>The number of nat lists of length @{text m} summing to @{text N} is
 | 
| 941 | @{term "(N + m - 1) choose N"}:\<close>
 | |
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 942 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 943 | lemma card_length_listsum_rec: | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 944 | assumes "m\<ge>1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 945 |   shows "card {l::nat list. length l = m \<and> listsum l = N} =
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 946 |     (card {l. length l = (m - 1) \<and> listsum l = N} +
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 947 |     card {l. length l = m \<and> listsum l + 1 =  N})"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 948 | (is "card ?C = (card ?A + card ?B)") | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 949 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 950 |   let ?A'="{l. length l = m \<and> listsum l = N \<and> hd l = 0}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 951 |   let ?B'="{l. length l = m \<and> listsum l = N \<and> hd l \<noteq> 0}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 952 | let ?f ="\<lambda> l. 0#l" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 953 | let ?g ="\<lambda> l. (hd l + 1) # tl l" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 954 | have 1: "\<And>xs x. xs \<noteq> [] \<Longrightarrow> x = hd xs \<Longrightarrow> x # tl xs = xs" by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 955 | have 2: "\<And>xs. (xs::nat list) \<noteq> [] \<Longrightarrow> listsum(tl xs) = listsum xs - hd xs" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 956 | by(auto simp add: neq_Nil_conv) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 957 | have f: "bij_betw ?f ?A ?A'" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 958 | apply(rule bij_betw_byWitness[where f' = tl]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 959 | using assms | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 960 | by (auto simp: 2 length_0_conv[symmetric] 1 simp del: length_0_conv) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 961 | have 3: "\<And>xs:: nat list. xs \<noteq> [] \<Longrightarrow> hd xs + (listsum xs - hd xs) = listsum xs" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 962 | by (metis 1 listsum_simps(2) 2) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 963 | have g: "bij_betw ?g ?B ?B'" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 964 | apply(rule bij_betw_byWitness[where f' = "\<lambda> l. (hd l - 1) # tl l"]) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 965 | using assms | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 966 | by (auto simp: 2 length_0_conv[symmetric] intro!: 3 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 967 | simp del: length_greater_0_conv length_0_conv) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 968 |   { fix M N :: nat have "finite {xs. size xs = M \<and> set xs \<subseteq> {0..<N}}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 969 | using finite_lists_length_eq[OF finite_atLeastLessThan] conj_commute by auto } | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 970 | note fin = this | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 971 | have fin_A: "finite ?A" using fin[of _ "N+1"] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 972 |     by (intro finite_subset[where ?A = "?A" and ?B = "{xs. size xs = m - 1 \<and> set xs \<subseteq> {0..<N+1}}"],
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 973 | auto simp: member_le_listsum_nat less_Suc_eq_le) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 974 | have fin_B: "finite ?B" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 975 |     by (intro finite_subset[where ?A = "?B" and ?B = "{xs. size xs = m \<and> set xs \<subseteq> {0..<N}}"],
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 976 | auto simp: member_le_listsum_nat less_Suc_eq_le fin) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 977 | have uni: "?C = ?A' \<union> ?B'" by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 978 |   have disj: "?A' \<inter> ?B' = {}" by auto
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 979 | have "card ?C = card(?A' \<union> ?B')" using uni by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 980 | also have "\<dots> = card ?A + card ?B" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 981 | using card_Un_disjoint[OF _ _ disj] bij_betw_finite[OF f] bij_betw_finite[OF g] | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 982 | bij_betw_same_card[OF f] bij_betw_same_card[OF g] fin_A fin_B | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 983 | by presburger | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 984 | finally show ?thesis . | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 985 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 986 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 987 | lemma card_length_listsum: --"By Holden Lee, tidied by Tobias Nipkow" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 988 |   "card {l::nat list. size l = m \<and> listsum l = N} = (N + m - 1) choose N"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 989 | proof (cases m) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 990 | case 0 then show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 991 | by (cases N) (auto simp: cong: conj_cong) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 992 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 993 | case (Suc m') | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 994 | have m: "m\<ge>1" by (simp add: Suc) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 995 | then show ?thesis | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 996 | proof (induct "N + m - 1" arbitrary: N m) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 997 | case 0 -- "In the base case, the only solution is [0]." | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 998 |       have [simp]: "{l::nat list. length l = Suc 0 \<and> (\<forall>n\<in>set l. n = 0)} = {[0]}"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 999 | by (auto simp: length_Suc_conv) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1000 | have "m=1 \<and> N=0" using 0 by linarith | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1001 | then show ?case by simp | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1002 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1003 | case (Suc k) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1004 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1005 |       have c1: "card {l::nat list. size l = (m - 1) \<and> listsum l =  N} =
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1006 | (N + (m - 1) - 1) choose N" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1007 | proof cases | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1008 | assume "m = 1" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1009 | with Suc.hyps have "N\<ge>1" by auto | 
| 60758 | 1010 | with \<open>m = 1\<close> show ?thesis by (simp add: binomial_eq_0) | 
| 59667 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1011 | next | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1012 | assume "m \<noteq> 1" thus ?thesis using Suc by fastforce | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1013 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1014 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1015 |       from Suc have c2: "card {l::nat list. size l = m \<and> listsum l + 1 = N} =
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1016 | (if N>0 then ((N - 1) + m - 1) choose (N - 1) else 0)" | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1017 | proof - | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1018 | have aux: "\<And>m n. n > 0 \<Longrightarrow> Suc m = n \<longleftrightarrow> m = n - 1" by arith | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1019 | from Suc have "N>0 \<Longrightarrow> | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1020 |           card {l::nat list. size l = m \<and> listsum l + 1 = N} =
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1021 | ((N - 1) + m - 1) choose (N - 1)" by (simp add: aux) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1022 | thus ?thesis by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1023 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1024 | |
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1025 |       from Suc.prems have "(card {l::nat list. size l = (m - 1) \<and> listsum l = N} +
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1026 |           card {l::nat list. size l = m \<and> listsum l + 1 = N}) = (N + m - 1) choose N"
 | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1027 | by (auto simp: c1 c2 choose_reduce_nat[of "N + m - 1" N] simp del: One_nat_def) | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1028 | thus ?case using card_length_listsum_rec[OF Suc.prems] by auto | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1029 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1030 | qed | 
| 
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
 paulson <lp15@cam.ac.uk> parents: 
59658diff
changeset | 1031 | |
| 60604 | 1032 | |
| 1033 | lemma Suc_times_binomial_add: -- \<open>by Lukas Bulwahn\<close> | |
| 1034 | "Suc a * (Suc (a + b) choose Suc a) = Suc b * (Suc (a + b) choose a)" | |
| 1035 | proof - | |
| 1036 | have dvd: "Suc a * (fact a * fact b) dvd fact (Suc (a + b))" for a b | |
| 1037 | using fact_fact_dvd_fact[of "Suc a" "b", where 'a=nat] | |
| 1038 | by (simp only: fact_Suc add_Suc[symmetric] of_nat_id mult.assoc) | |
| 1039 | ||
| 1040 | have "Suc a * (fact (Suc (a + b)) div (Suc a * fact a * fact b)) = | |
| 1041 | Suc a * fact (Suc (a + b)) div (Suc a * (fact a * fact b))" | |
| 1042 | by (subst div_mult_swap[symmetric]; simp only: mult.assoc dvd) | |
| 1043 | also have "\<dots> = Suc b * fact (Suc (a + b)) div (Suc b * (fact a * fact b))" | |
| 1044 | by (simp only: div_mult_mult1) | |
| 1045 | also have "\<dots> = Suc b * (fact (Suc (a + b)) div (Suc b * (fact a * fact b)))" | |
| 1046 | using dvd[of b a] by (subst div_mult_swap[symmetric]; simp only: ac_simps dvd) | |
| 1047 | finally show ?thesis | |
| 1048 | by (subst (1 2) binomial_altdef_nat) | |
| 1049 | (simp_all only: ac_simps diff_Suc_Suc Suc_diff_le diff_add_inverse fact_Suc of_nat_id) | |
| 1050 | qed | |
| 1051 | ||
| 15131 | 1052 | end |