| author | haftmann | 
| Sun, 05 Jun 2005 14:41:23 +0200 | |
| changeset 16276 | 3a50bf1f04d0 | 
| parent 15251 | bb6f072c8d10 | 
| child 16733 | 236dfafbeb63 | 
| permissions | -rw-r--r-- | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 1 | (* Title: HOL/Power.thy | 
| 
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 | *) | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 7 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 8 | header{*Exponentiation and Binomial Coefficients*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 9 | |
| 15131 | 10 | theory Power | 
| 15140 | 11 | imports Divides | 
| 15131 | 12 | begin | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 13 | |
| 15066 | 14 | subsection{*Powers for Arbitrary Semirings*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 15 | |
| 15004 | 16 | axclass recpower \<subseteq> comm_semiring_1_cancel, power | 
| 17 | power_0 [simp]: "a ^ 0 = 1" | |
| 18 | power_Suc: "a ^ (Suc n) = a * (a ^ n)" | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 19 | |
| 15004 | 20 | lemma power_0_Suc [simp]: "(0::'a::recpower) ^ (Suc n) = 0" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 21 | by (simp add: power_Suc) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 22 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 23 | text{*It looks plausible as a simprule, but its effect can be strange.*}
 | 
| 15004 | 24 | lemma power_0_left: "0^n = (if n=0 then 1 else (0::'a::recpower))" | 
| 15251 | 25 | by (induct "n", auto) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 26 | |
| 15004 | 27 | lemma power_one [simp]: "1^n = (1::'a::recpower)" | 
| 15251 | 28 | apply (induct "n") | 
| 14577 | 29 | apply (auto simp add: power_Suc) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 30 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 31 | |
| 15004 | 32 | lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 33 | by (simp add: power_Suc) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 34 | |
| 15004 | 35 | lemma power_add: "(a::'a::recpower) ^ (m+n) = (a^m) * (a^n)" | 
| 15251 | 36 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 37 | apply (simp_all add: power_Suc mult_ac) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 38 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 39 | |
| 15004 | 40 | lemma power_mult: "(a::'a::recpower) ^ (m*n) = (a^m) ^ n" | 
| 15251 | 41 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 42 | apply (simp_all add: power_Suc power_add) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 43 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 44 | |
| 15004 | 45 | lemma power_mult_distrib: "((a::'a::recpower) * b) ^ n = (a^n) * (b^n)" | 
| 15251 | 46 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 47 | apply (auto simp add: power_Suc mult_ac) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 48 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 49 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 50 | lemma zero_less_power: | 
| 15004 | 51 |      "0 < (a::'a::{ordered_semidom,recpower}) ==> 0 < a^n"
 | 
| 15251 | 52 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 53 | apply (simp_all add: power_Suc zero_less_one mult_pos) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 54 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 55 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 56 | lemma zero_le_power: | 
| 15004 | 57 |      "0 \<le> (a::'a::{ordered_semidom,recpower}) ==> 0 \<le> a^n"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 58 | apply (simp add: order_le_less) | 
| 14577 | 59 | apply (erule disjE) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 60 | apply (simp_all add: zero_less_power zero_less_one power_0_left) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 61 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 62 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 63 | lemma one_le_power: | 
| 15004 | 64 |      "1 \<le> (a::'a::{ordered_semidom,recpower}) ==> 1 \<le> a^n"
 | 
| 15251 | 65 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 66 | apply (simp_all add: power_Suc) | 
| 14577 | 67 | apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) | 
| 68 | apply (simp_all add: zero_le_one order_trans [OF zero_le_one]) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 69 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 70 | |
| 14738 | 71 | lemma gt1_imp_ge0: "1 < a ==> 0 \<le> (a::'a::ordered_semidom)" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 72 | by (simp add: order_trans [OF zero_le_one order_less_imp_le]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 73 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 74 | lemma power_gt1_lemma: | 
| 15004 | 75 |   assumes gt1: "1 < (a::'a::{ordered_semidom,recpower})"
 | 
| 14577 | 76 | shows "1 < a * a^n" | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 77 | proof - | 
| 14577 | 78 | have "1*1 < a*1" using gt1 by simp | 
| 79 | also have "\<dots> \<le> a * a^n" using gt1 | |
| 80 | by (simp only: mult_mono gt1_imp_ge0 one_le_power order_less_imp_le | |
| 81 | zero_le_one order_refl) | |
| 82 | finally show ?thesis by simp | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 83 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 84 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 85 | lemma power_gt1: | 
| 15004 | 86 |      "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 87 | by (simp add: power_gt1_lemma power_Suc) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 88 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 89 | lemma power_le_imp_le_exp: | 
| 15004 | 90 |   assumes gt1: "(1::'a::{recpower,ordered_semidom}) < a"
 | 
| 14577 | 91 | shows "!!n. a^m \<le> a^n ==> m \<le> n" | 
| 92 | proof (induct m) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 93 | case 0 | 
| 14577 | 94 | show ?case by simp | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 95 | next | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 96 | case (Suc m) | 
| 14577 | 97 | show ?case | 
| 98 | proof (cases n) | |
| 99 | case 0 | |
| 100 | from prems have "a * a^m \<le> 1" by (simp add: power_Suc) | |
| 101 | with gt1 show ?thesis | |
| 102 | by (force simp only: power_gt1_lemma | |
| 103 | linorder_not_less [symmetric]) | |
| 104 | next | |
| 105 | case (Suc n) | |
| 106 | from prems show ?thesis | |
| 107 | by (force dest: mult_left_le_imp_le | |
| 108 | simp add: power_Suc order_less_trans [OF zero_less_one gt1]) | |
| 109 | qed | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 110 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 111 | |
| 14577 | 112 | text{*Surely we can strengthen this? It holds for @{text "0<a<1"} too.*}
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 113 | lemma power_inject_exp [simp]: | 
| 15004 | 114 |      "1 < (a::'a::{ordered_semidom,recpower}) ==> (a^m = a^n) = (m=n)"
 | 
| 14577 | 115 | by (force simp add: order_antisym power_le_imp_le_exp) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 116 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 117 | text{*Can relax the first premise to @{term "0<a"} in the case of the
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 118 | natural numbers.*} | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 119 | lemma power_less_imp_less_exp: | 
| 15004 | 120 |      "[| (1::'a::{recpower,ordered_semidom}) < a; a^m < a^n |] ==> m < n"
 | 
| 14577 | 121 | by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"] | 
| 122 | power_le_imp_le_exp) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 123 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 124 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 125 | lemma power_mono: | 
| 15004 | 126 |      "[|a \<le> b; (0::'a::{recpower,ordered_semidom}) \<le> a|] ==> a^n \<le> b^n"
 | 
| 15251 | 127 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 128 | apply (simp_all add: power_Suc) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 129 | apply (auto intro: mult_mono zero_le_power order_trans [of 0 a b]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 130 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 131 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 132 | lemma power_strict_mono [rule_format]: | 
| 15004 | 133 |      "[|a < b; (0::'a::{recpower,ordered_semidom}) \<le> a|]
 | 
| 14577 | 134 | ==> 0 < n --> a^n < b^n" | 
| 15251 | 135 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 136 | apply (auto simp add: mult_strict_mono zero_le_power power_Suc | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 137 | order_le_less_trans [of 0 a b]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 138 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 139 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 140 | lemma power_eq_0_iff [simp]: | 
| 15004 | 141 |      "(a^n = 0) = (a = (0::'a::{ordered_idom,recpower}) & 0<n)"
 | 
| 15251 | 142 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 143 | apply (auto simp add: power_Suc zero_neq_one [THEN not_sym]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 144 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 145 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 146 | lemma field_power_eq_0_iff [simp]: | 
| 15004 | 147 |      "(a^n = 0) = (a = (0::'a::{field,recpower}) & 0<n)"
 | 
| 15251 | 148 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 149 | apply (auto simp add: power_Suc field_mult_eq_0_iff zero_neq_one[THEN not_sym]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 150 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 151 | |
| 15004 | 152 | lemma field_power_not_zero: "a \<noteq> (0::'a::{field,recpower}) ==> a^n \<noteq> 0"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 153 | by force | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 154 | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 155 | lemma nonzero_power_inverse: | 
| 15004 | 156 |   "a \<noteq> 0 ==> inverse ((a::'a::{field,recpower}) ^ n) = (inverse a) ^ n"
 | 
| 15251 | 157 | apply (induct "n") | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 158 | apply (auto simp add: power_Suc nonzero_inverse_mult_distrib mult_commute) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 159 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 160 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 161 | text{*Perhaps these should be simprules.*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 162 | lemma power_inverse: | 
| 15004 | 163 |   "inverse ((a::'a::{field,division_by_zero,recpower}) ^ n) = (inverse a) ^ n"
 | 
| 15251 | 164 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 165 | apply (auto simp add: power_Suc inverse_mult_distrib) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 166 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 167 | |
| 14577 | 168 | lemma nonzero_power_divide: | 
| 15004 | 169 |     "b \<noteq> 0 ==> (a/b) ^ n = ((a::'a::{field,recpower}) ^ n) / (b ^ n)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 170 | by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 171 | |
| 14577 | 172 | lemma power_divide: | 
| 15004 | 173 |     "(a/b) ^ n = ((a::'a::{field,division_by_zero,recpower}) ^ n / b ^ n)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 174 | apply (case_tac "b=0", simp add: power_0_left) | 
| 14577 | 175 | apply (rule nonzero_power_divide) | 
| 176 | apply assumption | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 177 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 178 | |
| 15004 | 179 | lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,recpower}) ^ n"
 | 
| 15251 | 180 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 181 | apply (auto simp add: power_Suc abs_mult) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 182 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 183 | |
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 184 | lemma zero_less_power_abs_iff [simp]: | 
| 15004 | 185 |      "(0 < (abs a)^n) = (a \<noteq> (0::'a::{ordered_idom,recpower}) | n=0)"
 | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 186 | proof (induct "n") | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 187 | case 0 | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 188 | show ?case by (simp add: zero_less_one) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 189 | next | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 190 | case (Suc n) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 191 | show ?case by (force simp add: prems power_Suc zero_less_mult_iff) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 192 | qed | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 193 | |
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 194 | lemma zero_le_power_abs [simp]: | 
| 15004 | 195 |      "(0::'a::{ordered_idom,recpower}) \<le> (abs a)^n"
 | 
| 15251 | 196 | apply (induct "n") | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 197 | apply (auto simp add: zero_le_one zero_le_power) | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 198 | done | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 199 | |
| 15004 | 200 | lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{comm_ring_1,recpower}) ^ n"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 201 | proof - | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 202 | have "-a = (- 1) * a" by (simp add: minus_mult_left [symmetric]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 203 | thus ?thesis by (simp only: power_mult_distrib) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 204 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 205 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 206 | text{*Lemma for @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 207 | lemma power_Suc_less: | 
| 15004 | 208 |      "[|(0::'a::{ordered_semidom,recpower}) < a; a < 1|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 209 | ==> a * a^n < a^n" | 
| 15251 | 210 | apply (induct n) | 
| 14577 | 211 | apply (auto simp add: power_Suc mult_strict_left_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 212 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 213 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 214 | lemma power_strict_decreasing: | 
| 15004 | 215 |      "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,recpower})|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 216 | ==> a^N < a^n" | 
| 14577 | 217 | apply (erule rev_mp) | 
| 15251 | 218 | apply (induct "N") | 
| 14577 | 219 | apply (auto simp add: power_Suc power_Suc_less less_Suc_eq) | 
| 220 | apply (rename_tac m) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 221 | apply (subgoal_tac "a * a^m < 1 * a^n", simp) | 
| 14577 | 222 | apply (rule mult_strict_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 223 | apply (auto simp add: zero_le_power zero_less_one order_less_imp_le) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 224 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 225 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 226 | text{*Proof resembles that of @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 227 | lemma power_decreasing: | 
| 15004 | 228 |      "[|n \<le> N; 0 \<le> a; a \<le> (1::'a::{ordered_semidom,recpower})|]
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 229 | ==> a^N \<le> a^n" | 
| 14577 | 230 | apply (erule rev_mp) | 
| 15251 | 231 | apply (induct "N") | 
| 14577 | 232 | apply (auto simp add: power_Suc le_Suc_eq) | 
| 233 | apply (rename_tac m) | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 234 | apply (subgoal_tac "a * a^m \<le> 1 * a^n", simp) | 
| 14577 | 235 | apply (rule mult_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 236 | apply (auto simp add: zero_le_power zero_le_one) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 237 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 238 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 239 | lemma power_Suc_less_one: | 
| 15004 | 240 |      "[| 0 < a; a < (1::'a::{ordered_semidom,recpower}) |] ==> a ^ Suc n < 1"
 | 
| 14577 | 241 | apply (insert power_strict_decreasing [of 0 "Suc n" a], simp) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 242 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 243 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 244 | text{*Proof again resembles that of @{text power_strict_decreasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 245 | lemma power_increasing: | 
| 15004 | 246 |      "[|n \<le> N; (1::'a::{ordered_semidom,recpower}) \<le> a|] ==> a^n \<le> a^N"
 | 
| 14577 | 247 | apply (erule rev_mp) | 
| 15251 | 248 | apply (induct "N") | 
| 14577 | 249 | apply (auto simp add: power_Suc le_Suc_eq) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 250 | apply (rename_tac m) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 251 | apply (subgoal_tac "1 * a^n \<le> a * a^m", simp) | 
| 14577 | 252 | apply (rule mult_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 253 | apply (auto simp add: order_trans [OF zero_le_one] zero_le_power) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 254 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 255 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 256 | text{*Lemma for @{text power_strict_increasing}*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 257 | lemma power_less_power_Suc: | 
| 15004 | 258 |      "(1::'a::{ordered_semidom,recpower}) < a ==> a^n < a * a^n"
 | 
| 15251 | 259 | apply (induct n) | 
| 14577 | 260 | apply (auto simp add: power_Suc mult_strict_left_mono order_less_trans [OF zero_less_one]) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 261 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 262 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 263 | lemma power_strict_increasing: | 
| 15004 | 264 |      "[|n < N; (1::'a::{ordered_semidom,recpower}) < a|] ==> a^n < a^N"
 | 
| 14577 | 265 | apply (erule rev_mp) | 
| 15251 | 266 | apply (induct "N") | 
| 14577 | 267 | apply (auto simp add: power_less_power_Suc power_Suc less_Suc_eq) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 268 | apply (rename_tac m) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 269 | apply (subgoal_tac "1 * a^n < a * a^m", simp) | 
| 14577 | 270 | apply (rule mult_strict_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 271 | apply (auto simp add: order_less_trans [OF zero_less_one] zero_le_power | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 272 | order_less_imp_le) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 273 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 274 | |
| 15066 | 275 | lemma power_increasing_iff [simp]: | 
| 276 |      "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x \<le> b ^ y) = (x \<le> y)"
 | |
| 277 | by (blast intro: power_le_imp_le_exp power_increasing order_less_imp_le) | |
| 278 | ||
| 279 | lemma power_strict_increasing_iff [simp]: | |
| 280 |      "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x < b ^ y) = (x < y)"
 | |
| 281 | by (blast intro: power_less_imp_less_exp power_strict_increasing) | |
| 282 | ||
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 283 | lemma power_le_imp_le_base: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 284 | assumes le: "a ^ Suc n \<le> b ^ Suc n" | 
| 15004 | 285 |       and xnonneg: "(0::'a::{ordered_semidom,recpower}) \<le> a"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 286 | and ynonneg: "0 \<le> b" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 287 | shows "a \<le> b" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 288 | proof (rule ccontr) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 289 | assume "~ a \<le> b" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 290 | then have "b < a" by (simp only: linorder_not_le) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 291 | then have "b ^ Suc n < a ^ Suc n" | 
| 14577 | 292 | by (simp only: prems power_strict_mono) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 293 | from le and this show "False" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 294 | by (simp add: linorder_not_less [symmetric]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 295 | qed | 
| 14577 | 296 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 297 | lemma power_inject_base: | 
| 14577 | 298 | "[| a ^ Suc n = b ^ Suc n; 0 \<le> a; 0 \<le> b |] | 
| 15004 | 299 |       ==> a = (b::'a::{ordered_semidom,recpower})"
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 300 | by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 301 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 302 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 303 | subsection{*Exponentiation for the Natural Numbers*}
 | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 304 | |
| 8844 | 305 | primrec (power) | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 306 | "p ^ 0 = 1" | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 307 | "p ^ (Suc n) = (p::nat) * (p ^ n)" | 
| 14577 | 308 | |
| 15004 | 309 | instance nat :: recpower | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 310 | proof | 
| 14438 | 311 | fix z n :: nat | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 312 | show "z^0 = 1" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 313 | show "z^(Suc n) = z * (z^n)" by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 314 | qed | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 315 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 316 | lemma nat_one_le_power [simp]: "1 \<le> i ==> Suc 0 \<le> i^n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 317 | by (insert one_le_power [of i n], simp) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 318 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 319 | lemma le_imp_power_dvd: "!!i::nat. m \<le> n ==> i^m dvd i^n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 320 | apply (unfold dvd_def) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 321 | apply (erule not_less_iff_le [THEN iffD2, THEN add_diff_inverse, THEN subst]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 322 | apply (simp add: power_add) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 323 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 324 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 325 | text{*Valid for the naturals, but what if @{text"0<i<1"}?
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 326 | Premises cannot be weakened: consider the case where @{term "i=0"},
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 327 | @{term "m=1"} and @{term "n=0"}.*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 328 | lemma nat_power_less_imp_less: "!!i::nat. [| 0 < i; i^m < i^n |] ==> m < n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 329 | apply (rule ccontr) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 330 | apply (drule leI [THEN le_imp_power_dvd, THEN dvd_imp_le, THEN leD]) | 
| 14577 | 331 | apply (erule zero_less_power, auto) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 332 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 333 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 334 | lemma nat_zero_less_power_iff [simp]: "(0 < x^n) = (x \<noteq> (0::nat) | n=0)" | 
| 15251 | 335 | by (induct "n", auto) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 336 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 337 | lemma power_le_dvd [rule_format]: "k^j dvd n --> i\<le>j --> k^i dvd (n::nat)" | 
| 15251 | 338 | apply (induct "j") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 339 | apply (simp_all add: le_Suc_eq) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 340 | apply (blast dest!: dvd_mult_right) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 341 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 342 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 343 | lemma power_dvd_imp_le: "[|i^m dvd i^n; (1::nat) < i|] ==> m \<le> n" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 344 | apply (rule power_le_imp_le_exp, assumption) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 345 | apply (erule dvd_imp_le, simp) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 346 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 347 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 348 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 349 | subsection{*Binomial Coefficients*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 350 | |
| 14577 | 351 | text{*This development is based on the work of Andy Gordon and
 | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 352 | Florian Kammueller*} | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 353 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 354 | consts | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 355 | binomial :: "[nat,nat] => nat" (infixl "choose" 65) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 356 | |
| 5183 | 357 | primrec | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 358 | binomial_0: "(0 choose k) = (if k = 0 then 1 else 0)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 359 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 360 | binomial_Suc: "(Suc n choose k) = | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 361 | (if k = 0 then 1 else (n choose (k - 1)) + (n choose k))" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 362 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 363 | lemma binomial_n_0 [simp]: "(n choose 0) = 1" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 364 | by (case_tac "n", simp_all) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 365 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 366 | lemma binomial_0_Suc [simp]: "(0 choose Suc k) = 0" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 367 | by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 368 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 369 | lemma binomial_Suc_Suc [simp]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 370 | "(Suc n choose Suc k) = (n choose k) + (n choose Suc k)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 371 | by simp | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 372 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 373 | lemma binomial_eq_0 [rule_format]: "\<forall>k. n < k --> (n choose k) = 0" | 
| 15251 | 374 | apply (induct "n", auto) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 375 | apply (erule allE) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 376 | apply (erule mp, arith) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 377 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 378 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 379 | declare binomial_0 [simp del] binomial_Suc [simp del] | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 380 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 381 | lemma binomial_n_n [simp]: "(n choose n) = 1" | 
| 15251 | 382 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 383 | apply (simp_all add: binomial_eq_0) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 384 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 385 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 386 | lemma binomial_Suc_n [simp]: "(Suc n choose n) = Suc n" | 
| 15251 | 387 | by (induct "n", simp_all) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 388 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 389 | lemma binomial_1 [simp]: "(n choose Suc 0) = n" | 
| 15251 | 390 | by (induct "n", simp_all) | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 391 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 392 | lemma zero_less_binomial [rule_format]: "k \<le> n --> 0 < (n choose k)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 393 | by (rule_tac m = n and n = k in diff_induct, simp_all) | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 394 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 395 | lemma binomial_eq_0_iff: "(n choose k = 0) = (n<k)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 396 | apply (safe intro!: binomial_eq_0) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 397 | apply (erule contrapos_pp) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 398 | apply (simp add: zero_less_binomial) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 399 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 400 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 401 | lemma zero_less_binomial_iff: "(0 < n choose k) = (k\<le>n)" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 402 | by (simp add: linorder_not_less [symmetric] binomial_eq_0_iff [symmetric]) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 403 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 404 | (*Might be more useful if re-oriented*) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 405 | lemma Suc_times_binomial_eq [rule_format]: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 406 | "\<forall>k. k \<le> n --> Suc n * (n choose k) = (Suc n choose Suc k) * Suc k" | 
| 15251 | 407 | apply (induct "n") | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 408 | apply (simp add: binomial_0, clarify) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 409 | apply (case_tac "k") | 
| 14577 | 410 | apply (auto simp add: add_mult_distrib add_mult_distrib2 le_Suc_eq | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 411 | binomial_eq_0) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 412 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 413 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 414 | text{*This is the well-known version, but it's harder to use because of the
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 415 | need to reason about division.*} | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 416 | lemma binomial_Suc_Suc_eq_times: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 417 | "k \<le> n ==> (Suc n choose Suc k) = (Suc n * (n choose k)) div Suc k" | 
| 14577 | 418 | by (simp add: Suc_times_binomial_eq div_mult_self_is_m zero_less_Suc | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 419 | del: mult_Suc mult_Suc_right) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 420 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 421 | text{*Another version, with -1 instead of Suc.*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 422 | lemma times_binomial_minus1_eq: | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 423 | "[|k \<le> n; 0<k|] ==> (n choose k) * k = n * ((n - 1) choose (k - 1))" | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 424 | apply (cut_tac n = "n - 1" and k = "k - 1" in Suc_times_binomial_eq) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 425 | apply (simp split add: nat_diff_split, auto) | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 426 | done | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 427 | |
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 428 | text{*ML bindings for the general exponentiation theorems*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 429 | ML | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 430 | {*
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 431 | val power_0 = thm"power_0"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 432 | val power_Suc = thm"power_Suc"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 433 | val power_0_Suc = thm"power_0_Suc"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 434 | val power_0_left = thm"power_0_left"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 435 | val power_one = thm"power_one"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 436 | val power_one_right = thm"power_one_right"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 437 | val power_add = thm"power_add"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 438 | val power_mult = thm"power_mult"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 439 | val power_mult_distrib = thm"power_mult_distrib"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 440 | val zero_less_power = thm"zero_less_power"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 441 | val zero_le_power = thm"zero_le_power"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 442 | val one_le_power = thm"one_le_power"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 443 | val gt1_imp_ge0 = thm"gt1_imp_ge0"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 444 | val power_gt1_lemma = thm"power_gt1_lemma"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 445 | val power_gt1 = thm"power_gt1"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 446 | val power_le_imp_le_exp = thm"power_le_imp_le_exp"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 447 | val power_inject_exp = thm"power_inject_exp"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 448 | val power_less_imp_less_exp = thm"power_less_imp_less_exp"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 449 | val power_mono = thm"power_mono"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 450 | val power_strict_mono = thm"power_strict_mono"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 451 | val power_eq_0_iff = thm"power_eq_0_iff"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 452 | val field_power_eq_0_iff = thm"field_power_eq_0_iff"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 453 | val field_power_not_zero = thm"field_power_not_zero"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 454 | val power_inverse = thm"power_inverse"; | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 455 | val nonzero_power_divide = thm"nonzero_power_divide"; | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 456 | val power_divide = thm"power_divide"; | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 457 | val power_abs = thm"power_abs"; | 
| 14353 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 458 | val zero_less_power_abs_iff = thm"zero_less_power_abs_iff"; | 
| 
79f9fbef9106
Added lemmas to Ring_and_Field with slightly modified simplification rules
 paulson parents: 
14348diff
changeset | 459 | val zero_le_power_abs = thm "zero_le_power_abs"; | 
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 460 | val power_minus = thm"power_minus"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 461 | val power_Suc_less = thm"power_Suc_less"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 462 | val power_strict_decreasing = thm"power_strict_decreasing"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 463 | val power_decreasing = thm"power_decreasing"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 464 | val power_Suc_less_one = thm"power_Suc_less_one"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 465 | val power_increasing = thm"power_increasing"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 466 | val power_strict_increasing = thm"power_strict_increasing"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 467 | val power_le_imp_le_base = thm"power_le_imp_le_base"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 468 | val power_inject_base = thm"power_inject_base"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 469 | *} | 
| 14577 | 470 | |
| 14348 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 471 | text{*ML bindings for the remaining theorems*}
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 472 | ML | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 473 | {*
 | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 474 | val nat_one_le_power = thm"nat_one_le_power"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 475 | val le_imp_power_dvd = thm"le_imp_power_dvd"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 476 | val nat_power_less_imp_less = thm"nat_power_less_imp_less"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 477 | val nat_zero_less_power_iff = thm"nat_zero_less_power_iff"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 478 | val power_le_dvd = thm"power_le_dvd"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 479 | val power_dvd_imp_le = thm"power_dvd_imp_le"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 480 | val binomial_n_0 = thm"binomial_n_0"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 481 | val binomial_0_Suc = thm"binomial_0_Suc"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 482 | val binomial_Suc_Suc = thm"binomial_Suc_Suc"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 483 | val binomial_eq_0 = thm"binomial_eq_0"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 484 | val binomial_n_n = thm"binomial_n_n"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 485 | val binomial_Suc_n = thm"binomial_Suc_n"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 486 | val binomial_1 = thm"binomial_1"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 487 | val zero_less_binomial = thm"zero_less_binomial"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 488 | val binomial_eq_0_iff = thm"binomial_eq_0_iff"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 489 | val zero_less_binomial_iff = thm"zero_less_binomial_iff"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 490 | val Suc_times_binomial_eq = thm"Suc_times_binomial_eq"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 491 | val binomial_Suc_Suc_eq_times = thm"binomial_Suc_Suc_eq_times"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 492 | val times_binomial_minus1_eq = thm"times_binomial_minus1_eq"; | 
| 
744c868ee0b7
Defining the type class "ringpower" and deleting superseded theorems for
 paulson parents: 
8844diff
changeset | 493 | *} | 
| 3390 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 494 | |
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 495 | end | 
| 
0c7625196d95
New theory "Power" of exponentiation (and binomial coefficients)
 paulson parents: diff
changeset | 496 |