| author | wenzelm | 
| Wed, 05 Dec 2001 03:07:44 +0100 | |
| changeset 12372 | cd3a09c7dac9 | 
| parent 12196 | a3be6b3a9c0b | 
| child 13451 | 467bccacc051 | 
| 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: 
8861diff
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: 
8861diff
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: 
11464diff
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: 
8861diff
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: 
11464diff
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: 
11464diff
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: 
8861diff
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: 
11464diff
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 | ||
| 12196 | 84 | (*** Logical equivalences for inequalities ***) | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 85 | |
| 12196 | 86 | Goal "(x^n = 0) = (x = (0::nat) & 0<n)"; | 
| 87 | by (induct_tac "n" 1); | |
| 88 | by Auto_tac; | |
| 89 | qed "power_eq_0_iff"; | |
| 90 | Addsimps [power_eq_0_iff]; | |
| 91 | ||
| 92 | Goal "(0 < x^n) = (x ~= (0::nat) | n=0)"; | |
| 93 | by (induct_tac "n" 1); | |
| 94 | by Auto_tac; | |
| 95 | qed "zero_less_power_iff"; | |
| 96 | Addsimps [zero_less_power_iff]; | |
| 97 | ||
| 98 | Goal "(0::nat) <= x^n"; | |
| 99 | by (induct_tac "n" 1); | |
| 100 | by Auto_tac; | |
| 101 | qed "zero_le_power"; | |
| 102 | Addsimps [zero_le_power]; | |
| 103 | ||
| 104 | ||
| 105 | (**** Binomial Coefficients, following Andy Gordon and Florian Kammueller ****) | |
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 106 | |
| 5069 | 107 | Goal "(n choose 0) = 1"; | 
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 108 | by (case_tac "n" 1); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 109 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 110 | qed "binomial_n_0"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 111 | |
| 5069 | 112 | Goal "(0 choose Suc k) = 0"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 113 | by (Simp_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 114 | qed "binomial_0_Suc"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 115 | |
| 5069 | 116 | 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 | 117 | by (Simp_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 118 | qed "binomial_Suc_Suc"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 119 | |
| 7084 | 120 | Goal "ALL k. n < k --> (n choose k) = 0"; | 
| 121 | by (induct_tac "n" 1); | |
| 122 | by Auto_tac; | |
| 123 | by (etac allE 1); | |
| 124 | by (etac mp 1); | |
| 125 | by (arith_tac 1); | |
| 126 | qed_spec_mp "binomial_eq_0"; | |
| 127 | ||
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 128 | Addsimps [binomial_n_0, binomial_0_Suc, binomial_Suc_Suc]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 129 | Delsimps [binomial_0, binomial_Suc]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 130 | |
| 5069 | 131 | Goal "(n choose n) = 1"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 132 | by (induct_tac "n" 1); | 
| 7844 | 133 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [binomial_eq_0]))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 134 | qed "binomial_n_n"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 135 | Addsimps [binomial_n_n]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 136 | |
| 5069 | 137 | Goal "(Suc n choose n) = Suc n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 138 | by (induct_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 139 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 140 | qed "binomial_Suc_n"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 141 | Addsimps [binomial_Suc_n]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 142 | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11464diff
changeset | 143 | Goal "(n choose Suc 0) = n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 144 | by (induct_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 145 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 146 | qed "binomial_1"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 147 | Addsimps [binomial_1]; | 
| 7084 | 148 | |
| 149 | Goal "k <= n --> 0 < (n choose k)"; | |
| 150 | by (res_inst_tac [("m","n"),("n","k")] diff_induct 1);
 | |
| 151 | by (ALLGOALS Asm_simp_tac); | |
| 152 | qed_spec_mp "zero_less_binomial"; | |
| 153 | ||
| 11365 | 154 | Goal "(n choose k = 0) = (n<k)"; | 
| 155 | by (safe_tac (claset() addSIs [binomial_eq_0])); | |
| 156 | by (etac contrapos_pp 1); | |
| 157 | by (asm_full_simp_tac (simpset() addsimps [zero_less_binomial]) 1); | |
| 158 | qed "binomial_eq_0_iff"; | |
| 159 | ||
| 160 | Goal "(0 < n choose k) = (k<=n)"; | |
| 161 | by (simp_tac (simpset() addsimps [linorder_not_less RS sym, | |
| 162 | binomial_eq_0_iff RS sym]) 1); | |
| 163 | qed "zero_less_binomial_iff"; | |
| 164 | ||
| 7084 | 165 | (*Might be more useful if re-oriented*) | 
| 166 | Goal "ALL k. k <= n --> Suc n * (n choose k) = (Suc n choose Suc k) * Suc k"; | |
| 167 | by (induct_tac "n" 1); | |
| 168 | by (simp_tac (simpset() addsimps [binomial_0]) 1); | |
| 169 | by (Clarify_tac 1); | |
| 8442 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 wenzelm parents: 
8423diff
changeset | 170 | by (case_tac "k" 1); | 
| 7084 | 171 | by (auto_tac (claset(), | 
| 172 | simpset() addsimps [add_mult_distrib, add_mult_distrib2, | |
| 173 | le_Suc_eq, binomial_eq_0])); | |
| 174 | qed_spec_mp "Suc_times_binomial_eq"; | |
| 175 | ||
| 11365 | 176 | (*This is the well-known version, but it's harder to use because of the | 
| 177 | need to reason about division.*) | |
| 7084 | 178 | Goal "k <= n ==> (Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k"; | 
| 179 | by (asm_simp_tac | |
| 9424 | 180 | (simpset_of NatDef.thy addsimps [Suc_times_binomial_eq, | 
| 7084 | 181 | div_mult_self_is_m]) 1); | 
| 182 | qed "binomial_Suc_Suc_eq_times"; | |
| 183 | ||
| 11365 | 184 | (*Another version, with -1 instead of Suc.*) | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11464diff
changeset | 185 | 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: 
11464diff
changeset | 186 | by (cut_inst_tac [("n","n - 1"),("k","k - 1")] Suc_times_binomial_eq 1);
 | 
| 11365 | 187 | by (asm_full_simp_tac (simpset() addsplits [nat_diff_split]) 1); | 
| 188 | by Auto_tac; | |
| 189 | qed "times_binomial_minus1_eq"; | |
| 190 |