| author | wenzelm | 
| Sun, 14 Oct 2001 20:06:13 +0200 | |
| changeset 11757 | 122be3f5b4b7 | 
| parent 11701 | 3d51fbf81c17 | 
| child 12196 | a3be6b3a9c0b | 
| permissions | -rw-r--r-- | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Power.ML  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
4  | 
Copyright 1997 University of Cambridge  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
5  | 
|
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
6  | 
The (overloaded) exponentiation operator, ^ :: [nat,nat]=>nat  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
7  | 
Also binomial coefficents  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
10  | 
(*** Simple laws about Power ***)  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
11  | 
|
| 5069 | 12  | 
Goal "!!i::nat. i ^ (j+k) = (i^j) * (i^k)";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
13  | 
by (induct_tac "k" 1);  | 
| 4089 | 14  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps mult_ac)));  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
15  | 
qed "power_add";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
16  | 
|
| 5069 | 17  | 
Goal "!!i::nat. i ^ (j*k) = (i^j) ^ k";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
18  | 
by (induct_tac "k" 1);  | 
| 4089 | 19  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [power_add])));  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
20  | 
qed "power_mult";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 
8935
 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 
paulson 
parents: 
8861 
diff
changeset
 | 
22  | 
Goal "!!i::nat. 0 < i ==> 0 < i^n";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
23  | 
by (induct_tac "n" 1);  | 
| 
8935
 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 
paulson 
parents: 
8861 
diff
changeset
 | 
24  | 
by (ALLGOALS Asm_simp_tac);  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
25  | 
qed "zero_less_power";  | 
| 8861 | 26  | 
Addsimps [zero_less_power];  | 
27  | 
||
| 11365 | 28  | 
Goal "i^n = 0 ==> i = (0::nat)";  | 
29  | 
by (etac contrapos_pp 1);  | 
|
30  | 
by Auto_tac;  | 
|
31  | 
qed "power_eq_0D";  | 
|
32  | 
||
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
33  | 
Goal "!!i::nat. 1 <= i ==> Suc 0 <= i^n";  | 
| 8861 | 34  | 
by (induct_tac "n" 1);  | 
35  | 
by Auto_tac;  | 
|
36  | 
qed "one_le_power";  | 
|
37  | 
Addsimps [one_le_power];  | 
|
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
38  | 
|
| 11331 | 39  | 
Goal "!!i::nat. 1 < i ==> !m. (i^n = i^m) = (n=m)";  | 
40  | 
by (induct_tac "n" 1);  | 
|
41  | 
by Auto_tac;  | 
|
42  | 
by (ALLGOALS (case_tac "m"));  | 
|
43  | 
by Auto_tac;  | 
|
44  | 
qed_spec_mp "power_inject";  | 
|
45  | 
Addsimps [power_inject];  | 
|
46  | 
||
| 
8935
 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 
paulson 
parents: 
8861 
diff
changeset
 | 
47  | 
Goalw [dvd_def] "!!i::nat. m<=n ==> i^m dvd i^n";  | 
| 3457 | 48  | 
by (etac (not_less_iff_le RS iffD2 RS add_diff_inverse RS subst) 1);  | 
| 4089 | 49  | 
by (asm_simp_tac (simpset() addsimps [power_add]) 1);  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
50  | 
qed "le_imp_power_dvd";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
51  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
52  | 
Goal "(1::nat) < i ==> \\<forall>n. i ^ m <= i ^ n --> m <= n";  | 
| 11365 | 53  | 
by (induct_tac "m" 1);  | 
54  | 
by Auto_tac;  | 
|
55  | 
by (case_tac "na" 1);  | 
|
56  | 
by Auto_tac;  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
57  | 
by (subgoal_tac "Suc 1 * 1 <= i * i^n" 1);  | 
| 11365 | 58  | 
by (Asm_full_simp_tac 1);  | 
59  | 
by (rtac mult_le_mono 1);  | 
|
60  | 
by Auto_tac;  | 
|
61  | 
qed_spec_mp "power_le_imp_le";  | 
|
62  | 
||
| 
8935
 
548901d05a0e
added type constraint ::nat because 0 is now overloaded
 
paulson 
parents: 
8861 
diff
changeset
 | 
63  | 
Goal "!!i::nat. [| 0 < i; i^m < i^n |] ==> m < n";  | 
| 3457 | 64  | 
by (rtac ccontr 1);  | 
65  | 
by (dtac (leI RS le_imp_power_dvd RS dvd_imp_le RS leD) 1);  | 
|
66  | 
by (etac zero_less_power 1);  | 
|
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
67  | 
by (contr_tac 1);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
68  | 
qed "power_less_imp_less";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
69  | 
|
| 11311 | 70  | 
Goal "k^j dvd n --> i<=j --> k^i dvd (n::nat)";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
71  | 
by (induct_tac "j" 1);  | 
| 11311 | 72  | 
by (ALLGOALS (simp_tac (simpset() addsimps [le_Suc_eq])));  | 
73  | 
by (blast_tac (claset() addSDs [dvd_mult_right]) 1);  | 
|
74  | 
qed_spec_mp "power_le_dvd";  | 
|
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
75  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
76  | 
Goal "[|i^m dvd i^n; (1::nat) < i|] ==> m <= n";  | 
| 11365 | 77  | 
by (rtac power_le_imp_le 1);  | 
78  | 
by (assume_tac 1);  | 
|
79  | 
by (etac dvd_imp_le 1);  | 
|
80  | 
by (Asm_full_simp_tac 1);  | 
|
81  | 
qed "power_dvd_imp_le";  | 
|
82  | 
||
83  | 
||
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
84  | 
|
| 3396 | 85  | 
(*** Binomial Coefficients, following Andy Gordon and Florian Kammueller ***)  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
86  | 
|
| 5069 | 87  | 
Goal "(n choose 0) = 1";  | 
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
88  | 
by (case_tac "n" 1);  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
89  | 
by (ALLGOALS Asm_simp_tac);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
90  | 
qed "binomial_n_0";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
91  | 
|
| 5069 | 92  | 
Goal "(0 choose Suc k) = 0";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
93  | 
by (Simp_tac 1);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
94  | 
qed "binomial_0_Suc";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
95  | 
|
| 5069 | 96  | 
Goal "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
97  | 
by (Simp_tac 1);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
98  | 
qed "binomial_Suc_Suc";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
99  | 
|
| 7084 | 100  | 
Goal "ALL k. n < k --> (n choose k) = 0";  | 
101  | 
by (induct_tac "n" 1);  | 
|
102  | 
by Auto_tac;  | 
|
103  | 
by (etac allE 1);  | 
|
104  | 
by (etac mp 1);  | 
|
105  | 
by (arith_tac 1);  | 
|
106  | 
qed_spec_mp "binomial_eq_0";  | 
|
107  | 
||
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
108  | 
Addsimps [binomial_n_0, binomial_0_Suc, binomial_Suc_Suc];  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
109  | 
Delsimps [binomial_0, binomial_Suc];  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
110  | 
|
| 5069 | 111  | 
Goal "(n choose n) = 1";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
112  | 
by (induct_tac "n" 1);  | 
| 7844 | 113  | 
by (ALLGOALS (asm_simp_tac (simpset() addsimps [binomial_eq_0])));  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
114  | 
qed "binomial_n_n";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
115  | 
Addsimps [binomial_n_n];  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
116  | 
|
| 5069 | 117  | 
Goal "(Suc n choose n) = Suc n";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
118  | 
by (induct_tac "n" 1);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
119  | 
by (ALLGOALS Asm_simp_tac);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
120  | 
qed "binomial_Suc_n";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
121  | 
Addsimps [binomial_Suc_n];  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
122  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
123  | 
Goal "(n choose Suc 0) = n";  | 
| 
3390
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
124  | 
by (induct_tac "n" 1);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
125  | 
by (ALLGOALS Asm_simp_tac);  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
126  | 
qed "binomial_1";  | 
| 
 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 
paulson 
parents:  
diff
changeset
 | 
127  | 
Addsimps [binomial_1];  | 
| 7084 | 128  | 
|
129  | 
Goal "k <= n --> 0 < (n choose k)";  | 
|
130  | 
by (res_inst_tac [("m","n"),("n","k")] diff_induct 1);
 | 
|
131  | 
by (ALLGOALS Asm_simp_tac);  | 
|
132  | 
qed_spec_mp "zero_less_binomial";  | 
|
133  | 
||
| 11365 | 134  | 
Goal "(n choose k = 0) = (n<k)";  | 
135  | 
by (safe_tac (claset() addSIs [binomial_eq_0]));  | 
|
136  | 
by (etac contrapos_pp 1);  | 
|
137  | 
by (asm_full_simp_tac (simpset() addsimps [zero_less_binomial]) 1);  | 
|
138  | 
qed "binomial_eq_0_iff";  | 
|
139  | 
||
140  | 
Goal "(0 < n choose k) = (k<=n)";  | 
|
141  | 
by (simp_tac (simpset() addsimps [linorder_not_less RS sym,  | 
|
142  | 
binomial_eq_0_iff RS sym]) 1);  | 
|
143  | 
qed "zero_less_binomial_iff";  | 
|
144  | 
||
| 7084 | 145  | 
(*Might be more useful if re-oriented*)  | 
146  | 
Goal "ALL k. k <= n --> Suc n * (n choose k) = (Suc n choose Suc k) * Suc k";  | 
|
147  | 
by (induct_tac "n" 1);  | 
|
148  | 
by (simp_tac (simpset() addsimps [binomial_0]) 1);  | 
|
149  | 
by (Clarify_tac 1);  | 
|
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
150  | 
by (case_tac "k" 1);  | 
| 7084 | 151  | 
by (auto_tac (claset(),  | 
152  | 
simpset() addsimps [add_mult_distrib, add_mult_distrib2,  | 
|
153  | 
le_Suc_eq, binomial_eq_0]));  | 
|
154  | 
qed_spec_mp "Suc_times_binomial_eq";  | 
|
155  | 
||
| 11365 | 156  | 
(*This is the well-known version, but it's harder to use because of the  | 
157  | 
need to reason about division.*)  | 
|
| 7084 | 158  | 
Goal "k <= n ==> (Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k";  | 
159  | 
by (asm_simp_tac  | 
|
| 9424 | 160  | 
(simpset_of NatDef.thy addsimps [Suc_times_binomial_eq,  | 
| 7084 | 161  | 
div_mult_self_is_m]) 1);  | 
162  | 
qed "binomial_Suc_Suc_eq_times";  | 
|
163  | 
||
| 11365 | 164  | 
(*Another version, with -1 instead of Suc.*)  | 
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
165  | 
Goal "[|k <= n; 0<k|] ==> (n choose k) * k = n * ((n - 1) choose (k - 1))";  | 
| 
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11464 
diff
changeset
 | 
166  | 
by (cut_inst_tac [("n","n - 1"),("k","k - 1")] Suc_times_binomial_eq 1);
 | 
| 11365 | 167  | 
by (asm_full_simp_tac (simpset() addsplits [nat_diff_split]) 1);  | 
168  | 
by Auto_tac;  | 
|
169  | 
qed "times_binomial_minus1_eq";  | 
|
170  |