author | paulson <lp15@cam.ac.uk> |
Mon, 09 Mar 2015 16:28:19 +0000 | |
changeset 59658 | 0cc388370041 |
parent 58889 | 5b7a9633cfa8 |
child 59667 | 651ea265d568 |
permissions | -rw-r--r-- |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
1 |
(* Title : Fact.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:
12196
diff
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:
30242
diff
changeset
|
5 |
The integer version of factorial and other additions by Jeremy Avigad. |
12196 | 6 |
*) |
7 |
||
58889 | 8 |
section{*Factorial Function*} |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
9 |
|
15131 | 10 |
theory Fact |
33319 | 11 |
imports Main |
15131 | 12 |
begin |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
13 |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
14 |
class fact = |
41550 | 15 |
fixes fact :: "'a \<Rightarrow> 'a" |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
16 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
17 |
instantiation nat :: fact |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
18 |
begin |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
19 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
20 |
fun |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
21 |
fact_nat :: "nat \<Rightarrow> nat" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
22 |
where |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
23 |
fact_0_nat: "fact_nat 0 = Suc 0" |
32047 | 24 |
| fact_Suc: "fact_nat (Suc x) = Suc x * fact x" |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
25 |
|
41550 | 26 |
instance .. |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
27 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
28 |
end |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
29 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
30 |
(* definitions for the integers *) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
31 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
32 |
instantiation int :: fact |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
33 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
34 |
begin |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
35 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
36 |
definition |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
37 |
fact_int :: "int \<Rightarrow> int" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
38 |
where |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
39 |
"fact_int x = (if x >= 0 then int (fact (nat x)) else 0)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
40 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
41 |
instance proof qed |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
42 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
43 |
end |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
44 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
45 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
46 |
subsection {* Set up Transfer *} |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
47 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
48 |
lemma transfer_nat_int_factorial: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
49 |
"(x::int) >= 0 \<Longrightarrow> fact (nat x) = nat (fact x)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
50 |
unfolding fact_int_def |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
51 |
by auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
52 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
53 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
54 |
lemma transfer_nat_int_factorial_closure: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
55 |
"x >= (0::int) \<Longrightarrow> fact x >= 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
56 |
by (auto simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
57 |
|
35644 | 58 |
declare transfer_morphism_nat_int[transfer add return: |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
59 |
transfer_nat_int_factorial transfer_nat_int_factorial_closure] |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
60 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
61 |
lemma transfer_int_nat_factorial: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
62 |
"fact (int x) = int (fact x)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
63 |
unfolding fact_int_def by auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
64 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
65 |
lemma transfer_int_nat_factorial_closure: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
66 |
"is_nat x \<Longrightarrow> fact x >= 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
67 |
by (auto simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
68 |
|
35644 | 69 |
declare transfer_morphism_int_nat[transfer add return: |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
70 |
transfer_int_nat_factorial transfer_int_nat_factorial_closure] |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
71 |
|
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
72 |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
73 |
subsection {* Factorial *} |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
74 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
75 |
lemma fact_0_int [simp]: "fact (0::int) = 1" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
76 |
by (simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
77 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
78 |
lemma fact_1_nat [simp]: "fact (1::nat) = 1" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
79 |
by simp |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
80 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
81 |
lemma fact_Suc_0_nat [simp]: "fact (Suc 0) = Suc 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
82 |
by simp |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
83 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
84 |
lemma fact_1_int [simp]: "fact (1::int) = 1" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
85 |
by (simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
86 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
87 |
lemma fact_plus_one_nat: "fact ((n::nat) + 1) = (n + 1) * fact n" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
88 |
by simp |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
89 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
90 |
lemma fact_plus_one_int: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
91 |
assumes "n >= 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
92 |
shows "fact ((n::int) + 1) = (n + 1) * fact n" |
41550 | 93 |
using assms unfolding fact_int_def |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
94 |
by (simp add: nat_add_distrib algebra_simps int_mult) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
95 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
96 |
lemma fact_reduce_nat: "(n::nat) > 0 \<Longrightarrow> fact n = n * fact (n - 1)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
97 |
apply (subgoal_tac "n = Suc (n - 1)") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
98 |
apply (erule ssubst) |
32047 | 99 |
apply (subst fact_Suc) |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
100 |
apply simp_all |
41550 | 101 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
102 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
103 |
lemma fact_reduce_int: "(n::int) > 0 \<Longrightarrow> fact n = n * fact (n - 1)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
104 |
apply (subgoal_tac "n = (n - 1) + 1") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
105 |
apply (erule ssubst) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
106 |
apply (subst fact_plus_one_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
107 |
apply simp_all |
41550 | 108 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
109 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
110 |
lemma fact_nonzero_nat [simp]: "fact (n::nat) \<noteq> 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
111 |
apply (induct n) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
112 |
apply (auto simp add: fact_plus_one_nat) |
41550 | 113 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
114 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
115 |
lemma fact_nonzero_int [simp]: "n >= 0 \<Longrightarrow> fact (n::int) ~= 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
116 |
by (simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
117 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
118 |
lemma fact_gt_zero_nat [simp]: "fact (n :: nat) > 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
119 |
by (insert fact_nonzero_nat [of n], arith) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
120 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
121 |
lemma fact_gt_zero_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) > 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
122 |
by (auto simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
123 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
124 |
lemma fact_ge_one_nat [simp]: "fact (n :: nat) >= 1" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
125 |
by (insert fact_nonzero_nat [of n], arith) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
126 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
127 |
lemma fact_ge_Suc_0_nat [simp]: "fact (n :: nat) >= Suc 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
128 |
by (insert fact_nonzero_nat [of n], arith) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
129 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
130 |
lemma fact_ge_one_int [simp]: "n >= 0 \<Longrightarrow> fact (n :: int) >= 1" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
131 |
apply (auto simp add: fact_int_def) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
132 |
apply (subgoal_tac "1 = int 1") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
133 |
apply (erule ssubst) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
134 |
apply (subst zle_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
135 |
apply auto |
41550 | 136 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
137 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
138 |
lemma dvd_fact_nat [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::nat)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
139 |
apply (induct n) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
140 |
apply force |
32047 | 141 |
apply (auto simp only: fact_Suc) |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
142 |
apply (subgoal_tac "m = Suc n") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
143 |
apply (erule ssubst) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
144 |
apply (rule dvd_triv_left) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
145 |
apply auto |
41550 | 146 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
147 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
148 |
lemma dvd_fact_int [rule_format]: "1 <= m \<longrightarrow> m <= n \<longrightarrow> m dvd fact (n::int)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
149 |
apply (case_tac "1 <= n") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
150 |
apply (induct n rule: int_ge_induct) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
151 |
apply (auto simp add: fact_plus_one_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
152 |
apply (subgoal_tac "m = i + 1") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
153 |
apply auto |
41550 | 154 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
155 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
156 |
lemma interval_plus_one_nat: "(i::nat) <= j + 1 \<Longrightarrow> |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
157 |
{i..j+1} = {i..j} Un {j+1}" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
158 |
by auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
159 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
160 |
lemma interval_Suc: "i <= Suc j \<Longrightarrow> {i..Suc j} = {i..j} Un {Suc j}" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
161 |
by auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
162 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
163 |
lemma interval_plus_one_int: "(i::int) <= j + 1 \<Longrightarrow> {i..j+1} = {i..j} Un {j+1}" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
164 |
by auto |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
165 |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
166 |
lemma fact_altdef_nat: "fact (n::nat) = (PROD i:{1..n}. i)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
167 |
apply (induct n) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
168 |
apply force |
32047 | 169 |
apply (subst fact_Suc) |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
170 |
apply (subst interval_Suc) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
171 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
172 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
173 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
174 |
lemma fact_altdef_int: "n >= 0 \<Longrightarrow> fact (n::int) = (PROD i:{1..n}. i)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
175 |
apply (induct n rule: int_ge_induct) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
176 |
apply force |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
177 |
apply (subst fact_plus_one_int, assumption) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
178 |
apply (subst interval_plus_one_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
179 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
180 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
181 |
|
40033
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
182 |
lemma fact_dvd: "n \<le> m \<Longrightarrow> fact n dvd fact (m::nat)" |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
183 |
by (auto simp add: fact_altdef_nat intro!: setprod_dvd_setprod_subset) |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
184 |
|
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
185 |
lemma fact_mod: "m \<le> (n::nat) \<Longrightarrow> fact n mod fact m = 0" |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
186 |
by (auto simp add: dvd_imp_mod_0 fact_dvd) |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
187 |
|
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
188 |
lemma fact_div_fact: |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
189 |
assumes "m \<ge> (n :: nat)" |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
190 |
shows "(fact m) div (fact n) = \<Prod>{n + 1..m}" |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
191 |
proof - |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
192 |
obtain d where "d = m - n" by auto |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
193 |
from assms this have "m = n + d" by auto |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
194 |
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:
35644
diff
changeset
|
195 |
proof (induct d) |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
196 |
case 0 |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
197 |
show ?case by simp |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
198 |
next |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
199 |
case (Suc d') |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
200 |
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:
35644
diff
changeset
|
201 |
by simp |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
202 |
also from Suc.hyps have "... = Suc (n + d') * \<Prod>{n + 1..n + d'}" |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
203 |
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:
35644
diff
changeset
|
204 |
also have "... = \<Prod>{n + 1..n + Suc d'}" |
57418 | 205 |
by (simp add: atLeastAtMostSuc_conv setprod.insert) |
40033
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
206 |
finally show ?case . |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
207 |
qed |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
208 |
from this `m = n + d` show ?thesis by simp |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
209 |
qed |
84200d970bf0
added some facts about factorial and dvd, div and mod
bulwahn
parents:
35644
diff
changeset
|
210 |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
211 |
lemma fact_mono_nat: "(m::nat) \<le> n \<Longrightarrow> fact m \<le> fact n" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
212 |
apply (drule le_imp_less_or_eq) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
213 |
apply (auto dest!: less_imp_Suc_add) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
214 |
apply (induct_tac k, auto) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
215 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
216 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
217 |
lemma fact_neg_int [simp]: "m < (0::int) \<Longrightarrow> fact m = 0" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
218 |
unfolding fact_int_def by auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
219 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
220 |
lemma fact_ge_zero_int [simp]: "fact m >= (0::int)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
221 |
apply (case_tac "m >= 0") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
222 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
223 |
apply (frule fact_gt_zero_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
224 |
apply arith |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
225 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
226 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
227 |
lemma fact_mono_int_aux [rule_format]: "k >= (0::int) \<Longrightarrow> |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
228 |
fact (m + k) >= fact m" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
229 |
apply (case_tac "m < 0") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
230 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
231 |
apply (induct k rule: int_ge_induct) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
232 |
apply auto |
57512
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57418
diff
changeset
|
233 |
apply (subst add.assoc [symmetric]) |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
234 |
apply (subst fact_plus_one_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
235 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
236 |
apply (erule order_trans) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
237 |
apply (subst mult_le_cancel_right1) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
238 |
apply (subgoal_tac "fact (m + i) >= 0") |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
239 |
apply arith |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
240 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
241 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
242 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
243 |
lemma fact_mono_int: "(m::int) <= n \<Longrightarrow> fact m <= fact n" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
244 |
apply (insert fact_mono_int_aux [of "n - m" "m"]) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
245 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
246 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
247 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
248 |
text{*Note that @{term "fact 0 = fact 1"}*} |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
249 |
lemma fact_less_mono_nat: "[| (0::nat) < m; m < n |] ==> fact m < fact n" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
250 |
apply (drule_tac m = m in less_imp_Suc_add, auto) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
251 |
apply (induct_tac k, auto) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
252 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
253 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
254 |
lemma fact_less_mono_int_aux: "k >= 0 \<Longrightarrow> (0::int) < m \<Longrightarrow> |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
255 |
fact m < fact ((m + 1) + k)" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
256 |
apply (induct k rule: int_ge_induct) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
257 |
apply (simp add: fact_plus_one_int) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
258 |
apply (subst (2) fact_reduce_int) |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
259 |
apply (auto simp add: ac_simps) |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
260 |
apply (erule order_less_le_trans) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
261 |
apply auto |
57514
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents:
57512
diff
changeset
|
262 |
done |
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
263 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
264 |
lemma fact_less_mono_int: "(0::int) < m \<Longrightarrow> m < n \<Longrightarrow> fact m < fact n" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
265 |
apply (insert fact_less_mono_int_aux [of "n - (m + 1)" "m"]) |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
266 |
apply auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
267 |
done |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
268 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
269 |
lemma fact_num_eq_if_nat: "fact (m::nat) = |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
270 |
(if m=0 then 1 else m * fact (m - 1))" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
271 |
by (cases m) auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
272 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
273 |
lemma fact_add_num_eq_if_nat: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
274 |
"fact ((m::nat) + n) = (if m + n = 0 then 1 else (m + n) * fact (m + n - 1))" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
275 |
by (cases "m + n") auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
276 |
|
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
277 |
lemma fact_add_num_eq_if2_nat: |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
278 |
"fact ((m::nat) + n) = |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
279 |
(if m = 0 then fact n else (m + n) * fact ((m - 1) + n))" |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
280 |
by (cases m) auto |
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
281 |
|
45930 | 282 |
lemma fact_le_power: "fact n \<le> n^n" |
283 |
proof (induct n) |
|
284 |
case (Suc n) |
|
285 |
then have "fact n \<le> Suc n ^ n" by (rule le_trans) (simp add: power_mono) |
|
286 |
then show ?case by (simp add: add_le_mono) |
|
287 |
qed simp |
|
32036
8a9228872fbd
Moved factorial lemmas from Binomial.thy to Fact.thy and merged.
avigad
parents:
30242
diff
changeset
|
288 |
|
32039
400a519bc888
Use term antiquotation to refer to constant names in subsection title.
berghofe
parents:
32036
diff
changeset
|
289 |
subsection {* @{term fact} and @{term of_nat} *} |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
290 |
|
29693
708dcf7dec9f
moved upwards in thy graph, real related theorems moved to Transcendental.thy
chaieb
parents:
28952
diff
changeset
|
291 |
lemma of_nat_fact_not_zero [simp]: "of_nat (fact n) \<noteq> (0::'a::semiring_char_0)" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
292 |
by auto |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
293 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
33319
diff
changeset
|
294 |
lemma of_nat_fact_gt_zero [simp]: "(0::'a::{linordered_semidom}) < of_nat(fact n)" by auto |
29693
708dcf7dec9f
moved upwards in thy graph, real related theorems moved to Transcendental.thy
chaieb
parents:
28952
diff
changeset
|
295 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
33319
diff
changeset
|
296 |
lemma of_nat_fact_ge_zero [simp]: "(0::'a::linordered_semidom) \<le> of_nat(fact n)" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
297 |
by simp |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
298 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
33319
diff
changeset
|
299 |
lemma inv_of_nat_fact_gt_zero [simp]: "(0::'a::linordered_field) < inverse (of_nat (fact n))" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
300 |
by (auto simp add: positive_imp_inverse_positive) |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
301 |
|
35028
108662d50512
more consistent naming of type classes involving orderings (and lattices) -- c.f. NEWS
haftmann
parents:
33319
diff
changeset
|
302 |
lemma inv_of_nat_fact_ge_zero [simp]: "(0::'a::linordered_field) \<le> inverse (of_nat (fact n))" |
25134
3d4953e88449
Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents:
25112
diff
changeset
|
303 |
by (auto intro: order_less_imp_le) |
15094
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
paulson
parents:
12196
diff
changeset
|
304 |
|
50224 | 305 |
lemma fact_eq_rev_setprod_nat: "fact (k::nat) = (\<Prod>i<k. k - i)" |
306 |
unfolding fact_altdef_nat |
|
57129
7edb7550663e
introduce more powerful reindexing rules for big operators
hoelzl
parents:
57113
diff
changeset
|
307 |
by (rule setprod.reindex_bij_witness[where i="\<lambda>i. k - i" and j="\<lambda>i. k - i"]) auto |
50224 | 308 |
|
50240
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents:
50224
diff
changeset
|
309 |
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:
50224
diff
changeset
|
310 |
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:
50224
diff
changeset
|
311 |
proof - |
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents:
50224
diff
changeset
|
312 |
have "\<And>r. r \<le> n \<Longrightarrow> \<Prod>{n - r..n} = (n - r) * \<Prod>{Suc (n - r)..n}" |
57418 | 313 |
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:
50224
diff
changeset
|
314 |
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:
50224
diff
changeset
|
315 |
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:
50224
diff
changeset
|
316 |
qed |
019d642d422d
add upper bounds for factorial and binomial; add equation for binomial using nat-division (both from AFP/Girth_Chromatic)
hoelzl
parents:
50224
diff
changeset
|
317 |
|
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:
50240
diff
changeset
|
318 |
lemma fact_numeral: --{*Evaluation for specific numerals*} |
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:
50240
diff
changeset
|
319 |
"fact (numeral k) = (numeral k) * (fact (pred_numeral k))" |
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:
50240
diff
changeset
|
320 |
by (simp add: numeral_eq_Suc) |
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:
50240
diff
changeset
|
321 |
|
59658
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
322 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
323 |
text {* This development is based on the work of Andy Gordon and |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
324 |
Florian Kammueller. *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
325 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
326 |
subsection {* Basic definitions and lemmas *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
327 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
328 |
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:
58889
diff
changeset
|
329 |
where |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
330 |
"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:
58889
diff
changeset
|
331 |
| "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:
58889
diff
changeset
|
332 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
333 |
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:
58889
diff
changeset
|
334 |
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:
58889
diff
changeset
|
335 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
336 |
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:
58889
diff
changeset
|
337 |
by simp |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
338 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
339 |
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:
58889
diff
changeset
|
340 |
by simp |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
341 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
342 |
lemma 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:
58889
diff
changeset
|
343 |
"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:
58889
diff
changeset
|
344 |
(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:
58889
diff
changeset
|
345 |
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:
58889
diff
changeset
|
346 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
347 |
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:
58889
diff
changeset
|
348 |
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:
58889
diff
changeset
|
349 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
350 |
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:
58889
diff
changeset
|
351 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
352 |
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:
58889
diff
changeset
|
353 |
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:
58889
diff
changeset
|
354 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
355 |
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:
58889
diff
changeset
|
356 |
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:
58889
diff
changeset
|
357 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
358 |
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:
58889
diff
changeset
|
359 |
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:
58889
diff
changeset
|
360 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
361 |
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:
58889
diff
changeset
|
362 |
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:
58889
diff
changeset
|
363 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
364 |
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:
58889
diff
changeset
|
365 |
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:
58889
diff
changeset
|
366 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
367 |
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:
58889
diff
changeset
|
368 |
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:
58889
diff
changeset
|
369 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
370 |
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:
58889
diff
changeset
|
371 |
"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:
58889
diff
changeset
|
372 |
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:
58889
diff
changeset
|
373 |
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:
58889
diff
changeset
|
374 |
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:
58889
diff
changeset
|
375 |
done |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
376 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
377 |
text{*The absorption property*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
378 |
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:
58889
diff
changeset
|
379 |
"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:
58889
diff
changeset
|
380 |
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:
58889
diff
changeset
|
381 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
382 |
text{*This is the well-known version of absorption, but it's harder to use because of the |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
383 |
need to reason about division.*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
384 |
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:
58889
diff
changeset
|
385 |
"(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:
58889
diff
changeset
|
386 |
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:
58889
diff
changeset
|
387 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
388 |
text{*Another version of absorption, with -1 instead of Suc.*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
389 |
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:
58889
diff
changeset
|
390 |
"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:
58889
diff
changeset
|
391 |
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:
58889
diff
changeset
|
392 |
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:
58889
diff
changeset
|
393 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
394 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
395 |
subsection {* Combinatorial theorems involving @{text "choose"} *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
396 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
397 |
text {*By Florian Kamm\"uller, tidied by LCP.*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
398 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
399 |
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:
58889
diff
changeset
|
400 |
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:
58889
diff
changeset
|
401 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
402 |
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:
58889
diff
changeset
|
403 |
{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:
58889
diff
changeset
|
404 |
{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:
58889
diff
changeset
|
405 |
apply safe |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
406 |
apply (auto intro: finite_subset [THEN card_insert_disjoint]) |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
407 |
by (metis (full_types) Diff_insert_absorb Set.set_insert Zero_neq_Suc card_Diff_singleton_if |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
408 |
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:
58889
diff
changeset
|
409 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
410 |
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:
58889
diff
changeset
|
411 |
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:
58889
diff
changeset
|
412 |
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:
58889
diff
changeset
|
413 |
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:
58889
diff
changeset
|
414 |
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:
58889
diff
changeset
|
415 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
416 |
text{*There are as many subsets of @{term A} having cardinality @{term k} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
417 |
as there are sets obtained from the former by inserting a fixed element |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
418 |
@{term x} into each.*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
419 |
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:
58889
diff
changeset
|
420 |
"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:
58889
diff
changeset
|
421 |
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:
58889
diff
changeset
|
422 |
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:
58889
diff
changeset
|
423 |
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:
58889
diff
changeset
|
424 |
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:
58889
diff
changeset
|
425 |
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:
58889
diff
changeset
|
426 |
done |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
427 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
428 |
text {* |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
429 |
Main theorem: combinatorial statement about number of subsets of a set. |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
430 |
*} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
431 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
432 |
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:
58889
diff
changeset
|
433 |
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:
58889
diff
changeset
|
434 |
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:
58889
diff
changeset
|
435 |
next |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
436 |
case (Suc k) |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
437 |
show ?case using `finite A` |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
438 |
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:
58889
diff
changeset
|
439 |
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:
58889
diff
changeset
|
440 |
next |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
441 |
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:
58889
diff
changeset
|
442 |
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:
58889
diff
changeset
|
443 |
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:
58889
diff
changeset
|
444 |
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:
58889
diff
changeset
|
445 |
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:
58889
diff
changeset
|
446 |
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:
58889
diff
changeset
|
447 |
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:
58889
diff
changeset
|
448 |
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:
58889
diff
changeset
|
449 |
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:
58889
diff
changeset
|
450 |
done |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
451 |
qed |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
452 |
qed |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
453 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
454 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
455 |
subsection {* The binomial theorem (courtesy of Tobias Nipkow): *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
456 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
457 |
text{* Avigad's version, generalized to any commutative ring *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
458 |
theorem binomial_ring: "(a+b::'a::{comm_ring_1,power})^n = |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
459 |
(\<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:
58889
diff
changeset
|
460 |
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:
58889
diff
changeset
|
461 |
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:
58889
diff
changeset
|
462 |
next |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
463 |
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:
58889
diff
changeset
|
464 |
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:
58889
diff
changeset
|
465 |
by auto |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
466 |
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:
58889
diff
changeset
|
467 |
by auto |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
468 |
have "(a+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:
58889
diff
changeset
|
469 |
(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:
58889
diff
changeset
|
470 |
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:
58889
diff
changeset
|
471 |
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:
58889
diff
changeset
|
472 |
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:
58889
diff
changeset
|
473 |
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:
58889
diff
changeset
|
474 |
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:
58889
diff
changeset
|
475 |
(\<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:
58889
diff
changeset
|
476 |
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:
58889
diff
changeset
|
477 |
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:
58889
diff
changeset
|
478 |
(\<Sum>k=1..n+1. 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:
58889
diff
changeset
|
479 |
by (simp add:setsum_shift_bounds_cl_Suc_ivl Suc_diff_le field_simps |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
480 |
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:
58889
diff
changeset
|
481 |
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:
58889
diff
changeset
|
482 |
(\<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:
58889
diff
changeset
|
483 |
(\<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:
58889
diff
changeset
|
484 |
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:
58889
diff
changeset
|
485 |
also have |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
486 |
"\<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:
58889
diff
changeset
|
487 |
(\<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:
58889
diff
changeset
|
488 |
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:
58889
diff
changeset
|
489 |
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:
58889
diff
changeset
|
490 |
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:
58889
diff
changeset
|
491 |
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:
58889
diff
changeset
|
492 |
qed |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
493 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
494 |
text{* Original version for the naturals *} |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
495 |
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:
58889
diff
changeset
|
496 |
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:
58889
diff
changeset
|
497 |
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:
58889
diff
changeset
|
498 |
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:
58889
diff
changeset
|
499 |
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:
58889
diff
changeset
|
500 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
501 |
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:
58889
diff
changeset
|
502 |
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:
58889
diff
changeset
|
503 |
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:
58889
diff
changeset
|
504 |
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:
58889
diff
changeset
|
505 |
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:
58889
diff
changeset
|
506 |
{ 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:
58889
diff
changeset
|
507 |
moreover |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
508 |
{ 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:
58889
diff
changeset
|
509 |
moreover |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
510 |
{ 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:
58889
diff
changeset
|
511 |
moreover |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
512 |
{ 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:
58889
diff
changeset
|
513 |
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:
58889
diff
changeset
|
514 |
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:
58889
diff
changeset
|
515 |
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:
58889
diff
changeset
|
516 |
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:
58889
diff
changeset
|
517 |
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:
58889
diff
changeset
|
518 |
by simp |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
519 |
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:
58889
diff
changeset
|
520 |
have "fact k * fact (n - k) * (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:
58889
diff
changeset
|
521 |
k * (fact h * fact (m - h) * (m choose h)) + |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
522 |
(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:
58889
diff
changeset
|
523 |
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:
58889
diff
changeset
|
524 |
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:
58889
diff
changeset
|
525 |
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:
58889
diff
changeset
|
526 |
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:
58889
diff
changeset
|
527 |
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:
58889
diff
changeset
|
528 |
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:
58889
diff
changeset
|
529 |
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:
58889
diff
changeset
|
530 |
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:
58889
diff
changeset
|
531 |
qed |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
532 |
|
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
533 |
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:
58889
diff
changeset
|
534 |
assumes 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:
58889
diff
changeset
|
535 |
shows "(of_nat (n choose k) :: 'a::{field,ring_char_0}) = |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
536 |
of_nat (fact n) / (of_nat (fact k) * of_nat (fact (n - k)))" |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
537 |
using binomial_fact_lemma[OF kn] |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
538 |
by (simp add: field_simps of_nat_mult [symmetric]) |
0cc388370041
sin, cos generalised from type real to any "'a::{real_normed_field,banach}", including complex
paulson <lp15@cam.ac.uk>
parents:
58889
diff
changeset
|
539 |
|
15131 | 540 |
end |