diff -r 179ff9cb160b -r 5b25fee0362c src/HOL/Power.thy --- a/src/HOL/Power.thy Wed Mar 04 10:43:39 2009 +0100 +++ b/src/HOL/Power.thy Wed Mar 04 10:45:52 2009 +0100 @@ -31,7 +31,7 @@ by (induct n) (simp_all add: power_Suc) lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a" - by (simp add: power_Suc) + unfolding One_nat_def by (simp add: power_Suc) lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n" by (induct n) (simp_all add: power_Suc mult_assoc) @@ -143,11 +143,13 @@ done lemma power_eq_0_iff [simp]: - "(a^n = 0) = (a = (0::'a::{ring_1_no_zero_divisors,recpower}) & n>0)" + "(a^n = 0) \ + (a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,recpower}) & n\0)" apply (induct "n") -apply (auto simp add: power_Suc zero_neq_one [THEN not_sym]) +apply (auto simp add: power_Suc zero_neq_one [THEN not_sym] no_zero_divisors) done + lemma field_power_not_zero: "a \ (0::'a::{ring_1_no_zero_divisors,recpower}) ==> a^n \ 0" by force @@ -324,6 +326,24 @@ shows "\a ^ n = b ^ n; 0 \ a; 0 \ b; 0 < n\ \ a = b" by (cases n, simp_all, rule power_inject_base) +text {* The divides relation *} + +lemma le_imp_power_dvd: + fixes a :: "'a::{comm_semiring_1,recpower}" + assumes "m \ n" shows "a^m dvd a^n" +proof + have "a^n = a^(m + (n - m))" + using `m \ n` by simp + also have "\ = a^m * a^(n - m)" + by (rule power_add) + finally show "a^n = a^m * a^(n - m)" . +qed + +lemma power_le_dvd: + fixes a b :: "'a::{comm_semiring_1,recpower}" + shows "a^n dvd b \ m \ n \ a^m dvd b" + by (rule dvd_trans [OF le_imp_power_dvd]) + subsection{*Exponentiation for the Natural Numbers*} @@ -346,12 +366,19 @@ "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n" by (induct n, simp_all add: power_Suc of_nat_mult) -lemma nat_one_le_power [simp]: "1 \ i ==> Suc 0 \ i^n" -by (insert one_le_power [of i n], simp) +lemma nat_one_le_power [simp]: "Suc 0 \ i ==> Suc 0 \ i^n" +by (rule one_le_power [of i n, unfolded One_nat_def]) lemma nat_zero_less_power_iff [simp]: "(x^n > 0) = (x > (0::nat) | n=0)" by (induct "n", auto) +lemma nat_power_eq_Suc_0_iff [simp]: + "((x::nat)^m = Suc 0) = (m = 0 | x = Suc 0)" +by (induct_tac m, auto) + +lemma power_Suc_0[simp]: "(Suc 0)^n = Suc 0" +by simp + text{*Valid for the naturals, but what if @{text"0