| author | paulson | 
| Tue, 18 Aug 1998 10:24:54 +0200 | |
| changeset 5330 | 8c9fadda81f4 | 
| parent 5143 | b94cd208f073 | 
| child 6865 | 5577ffe4c2f1 | 
| 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 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 11 | open Power; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 12 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 13 | (*** Simple laws about Power ***) | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 14 | |
| 5069 | 15 | Goal "!!i::nat. i ^ (j+k) = (i^j) * (i^k)"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 16 | by (induct_tac "k" 1); | 
| 4089 | 17 | by (ALLGOALS (asm_simp_tac (simpset() addsimps mult_ac))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 18 | qed "power_add"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 19 | |
| 5069 | 20 | Goal "!!i::nat. i ^ (j*k) = (i^j) ^ k"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 21 | by (induct_tac "k" 1); | 
| 4089 | 22 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [power_add]))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 23 | qed "power_mult"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 24 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 25 | Goal "0 < i ==> 0 < i^n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 26 | by (induct_tac "n" 1); | 
| 4089 | 27 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [zero_less_mult_iff]))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 28 | qed "zero_less_power"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 29 | |
| 5069 | 30 | Goalw [dvd_def] "!!m::nat. m<=n ==> i^m dvd i^n"; | 
| 3457 | 31 | by (etac (not_less_iff_le RS iffD2 RS add_diff_inverse RS subst) 1); | 
| 4089 | 32 | by (asm_simp_tac (simpset() addsimps [power_add]) 1); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 33 | by (Blast_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 34 | qed "le_imp_power_dvd"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 35 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 36 | Goal "[| 0 < i; i^m < i^n |] ==> m<n"; | 
| 3457 | 37 | by (rtac ccontr 1); | 
| 38 | by (dtac (leI RS le_imp_power_dvd RS dvd_imp_le RS leD) 1); | |
| 39 | by (etac zero_less_power 1); | |
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 40 | by (contr_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 41 | qed "power_less_imp_less"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 42 | |
| 5143 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 paulson parents: 
5069diff
changeset | 43 | Goal "k^j dvd n --> i<j --> k^i dvd n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 44 | by (induct_tac "j" 1); | 
| 4089 | 45 | by (ALLGOALS (simp_tac (simpset() addsimps [less_Suc_eq]))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 46 | by (stac mult_commute 1); | 
| 4089 | 47 | by (blast_tac (claset() addSDs [dvd_mult_left]) 1); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 48 | qed_spec_mp "power_less_dvd"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 49 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 50 | |
| 3396 | 51 | (*** Binomial Coefficients, following Andy Gordon and Florian Kammueller ***) | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 52 | |
| 5069 | 53 | Goal "(n choose 0) = 1"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 54 | by (exhaust_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 55 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 56 | qed "binomial_n_0"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 57 | |
| 5069 | 58 | Goal "(0 choose Suc k) = 0"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 59 | by (Simp_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 60 | qed "binomial_0_Suc"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 61 | |
| 5069 | 62 | 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 | 63 | by (Simp_tac 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 64 | qed "binomial_Suc_Suc"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 65 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 66 | Addsimps [binomial_n_0, binomial_0_Suc, binomial_Suc_Suc]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 67 | Delsimps [binomial_0, binomial_Suc]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 68 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 69 | |
| 5069 | 70 | Goal "!k. n<k --> (n choose k) = 0"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 71 | by (induct_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 72 | by (ALLGOALS (rtac allI THEN' exhaust_tac "k")); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 73 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 74 | qed_spec_mp "less_imp_binomial_eq_0"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 75 | |
| 5069 | 76 | Goal "(n choose n) = 1"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 77 | by (induct_tac "n" 1); | 
| 4089 | 78 | by (ALLGOALS (asm_simp_tac (simpset() addsimps [less_imp_binomial_eq_0]))); | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 79 | qed "binomial_n_n"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 80 | Addsimps [binomial_n_n]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 81 | |
| 5069 | 82 | Goal "(Suc n choose n) = Suc n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 83 | by (induct_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 84 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 85 | qed "binomial_Suc_n"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 86 | Addsimps [binomial_Suc_n]; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 87 | |
| 5069 | 88 | Goal "(n choose 1) = n"; | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 89 | by (induct_tac "n" 1); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 90 | by (ALLGOALS Asm_simp_tac); | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 91 | qed "binomial_1"; | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 92 | Addsimps [binomial_1]; |