# HG changeset patch # User haftmann # Date 1240813357 -7200 # Node ID c2524d1235281a3097e1dd0e62b5262602713a0e # Parent e46639644fcd2ab7f9a05747d0d07c088eb86ba7# Parent a1efb13fc5d8b8cf356676a7d9cdd60eacb3224f merged diff -r e46639644fcd -r c2524d123528 src/HOL/Power.thy --- a/src/HOL/Power.thy Sun Apr 26 23:41:18 2009 +0100 +++ b/src/HOL/Power.thy Mon Apr 27 08:22:37 2009 +0200 @@ -11,85 +11,169 @@ subsection {* Powers for Arbitrary Monoids *} -class recpower = monoid_mult +class power = one + times begin primrec power :: "'a \ nat \ 'a" (infixr "^" 80) where power_0: "a ^ 0 = 1" | power_Suc: "a ^ Suc n = a * a ^ n" +notation (latex output) + power ("(_\<^bsup>_\<^esup>)" [1000] 1000) + +notation (HTML output) + power ("(_\<^bsup>_\<^esup>)" [1000] 1000) + end -lemma power_0_Suc [simp]: "(0::'a::{recpower,semiring_0}) ^ (Suc n) = 0" - by simp +context monoid_mult +begin -text{*It looks plausible as a simprule, but its effect can be strange.*} -lemma power_0_left: "0^n = (if n=0 then 1 else (0::'a::{recpower,semiring_0}))" - by (induct n) simp_all +subclass power .. -lemma power_one [simp]: "1^n = (1::'a::recpower)" +lemma power_one [simp]: + "1 ^ n = 1" by (induct n) simp_all -lemma power_one_right [simp]: "(a::'a::recpower) ^ 1 = a" - unfolding One_nat_def by simp +lemma power_one_right [simp]: + "a ^ 1 = a * 1" + by simp -lemma power_commutes: "(a::'a::recpower) ^ n * a = a * a ^ n" +lemma power_commutes: + "a ^ n * a = a * a ^ n" by (induct n) (simp_all add: mult_assoc) -lemma power_Suc2: "(a::'a::recpower) ^ Suc n = a ^ n * a" +lemma power_Suc2: + "a ^ Suc n = a ^ n * a" by (simp add: power_commutes) -lemma power_add: "(a::'a::recpower) ^ (m+n) = (a^m) * (a^n)" - by (induct m) (simp_all add: mult_ac) +lemma power_add: + "a ^ (m + n) = a ^ m * a ^ n" + by (induct m) (simp_all add: algebra_simps) -lemma power_mult: "(a::'a::recpower) ^ (m*n) = (a^m) ^ n" +lemma power_mult: + "a ^ (m * n) = (a ^ m) ^ n" by (induct n) (simp_all add: power_add) -lemma power_mult_distrib: "((a::'a::{recpower,comm_monoid_mult}) * b) ^ n = (a^n) * (b^n)" +end + +context comm_monoid_mult +begin + +lemma power_mult_distrib: + "(a * b) ^ n = (a ^ n) * (b ^ n)" by (induct n) (simp_all add: mult_ac) -lemma zero_less_power[simp]: - "0 < (a::'a::{ordered_semidom,recpower}) ==> 0 < a^n" -by (induct n) (simp_all add: mult_pos_pos) +end + +context semiring_1 +begin + +lemma of_nat_power: + "of_nat (m ^ n) = of_nat m ^ n" + by (induct n) (simp_all add: of_nat_mult) + +end + +context comm_semiring_1 +begin + +text {* The divides relation *} + +lemma le_imp_power_dvd: + 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: + "a ^ n dvd b \ m \ n \ a ^ m dvd b" + by (rule dvd_trans [OF le_imp_power_dvd]) + +lemma dvd_power_same: + "x dvd y \ x ^ n dvd y ^ n" + by (induct n) (auto simp add: mult_dvd_mono) + +lemma dvd_power_le: + "x dvd y \ m \ n \ x ^ n dvd y ^ m" + by (rule power_le_dvd [OF dvd_power_same]) -lemma zero_le_power[simp]: - "0 \ (a::'a::{ordered_semidom,recpower}) ==> 0 \ a^n" -by (induct n) (simp_all add: mult_nonneg_nonneg) +lemma dvd_power [simp]: + assumes "n > (0::nat) \ x = 1" + shows "x dvd (x ^ n)" +using assms proof + assume "0 < n" + then have "x ^ n = x ^ Suc (n - 1)" by simp + then show "x dvd (x ^ n)" by simp +next + assume "x = 1" + then show "x dvd (x ^ n)" by simp +qed + +end + +context ring_1 +begin + +lemma power_minus: + "(- a) ^ n = (- 1) ^ n * a ^ n" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) then show ?case + by (simp del: power_Suc add: power_Suc2 mult_assoc) +qed + +end + +context ordered_semidom +begin + +lemma zero_less_power [simp]: + "0 < a \ 0 < a ^ n" + by (induct n) (simp_all add: mult_pos_pos) + +lemma zero_le_power [simp]: + "0 \ a \ 0 \ a ^ n" + by (induct n) (simp_all add: mult_nonneg_nonneg) lemma one_le_power[simp]: - "1 \ (a::'a::{ordered_semidom,recpower}) ==> 1 \ a^n" -apply (induct "n") -apply simp_all -apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) -apply (simp_all add: order_trans [OF zero_le_one]) -done - -lemma gt1_imp_ge0: "1 < a ==> 0 \ (a::'a::ordered_semidom)" - by (simp add: order_trans [OF zero_le_one order_less_imp_le]) + "1 \ a \ 1 \ a ^ n" + apply (induct n) + apply simp_all + apply (rule order_trans [OF _ mult_mono [of 1 _ 1]]) + apply (simp_all add: order_trans [OF zero_le_one]) + done lemma power_gt1_lemma: - assumes gt1: "1 < (a::'a::{ordered_semidom,recpower})" - shows "1 < a * a^n" + assumes gt1: "1 < a" + shows "1 < a * a ^ n" proof - - have "1*1 < a*1" using gt1 by simp - also have "\ \ a * a^n" using gt1 - by (simp only: mult_mono gt1_imp_ge0 one_le_power order_less_imp_le + from gt1 have "0 \ a" + by (fact order_trans [OF zero_le_one less_imp_le]) + have "1 * 1 < a * 1" using gt1 by simp + also have "\ \ a * a ^ n" using gt1 + by (simp only: mult_mono `0 \ a` one_le_power order_less_imp_le zero_le_one order_refl) finally show ?thesis by simp qed -lemma one_less_power[simp]: - "\1 < (a::'a::{ordered_semidom,recpower}); 0 < n\ \ 1 < a ^ n" -by (cases n, simp_all add: power_gt1_lemma) +lemma power_gt1: + "1 < a \ 1 < a ^ Suc n" + by (simp add: power_gt1_lemma) -lemma power_gt1: - "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)" -by (simp add: power_gt1_lemma) +lemma one_less_power [simp]: + "1 < a \ 0 < n \ 1 < a ^ n" + by (cases n) (simp_all add: power_gt1_lemma) lemma power_le_imp_le_exp: - assumes gt1: "(1::'a::{recpower,ordered_semidom}) < a" - shows "!!n. a^m \ a^n ==> m \ n" -proof (induct m) + assumes gt1: "1 < a" + shows "a ^ m \ a ^ n \ m \ n" +proof (induct m arbitrary: n) case 0 show ?case by simp next @@ -97,212 +181,128 @@ show ?case proof (cases n) case 0 - from prems have "a * a^m \ 1" by simp + with Suc.prems Suc.hyps have "a * a ^ m \ 1" by simp with gt1 show ?thesis by (force simp only: power_gt1_lemma - linorder_not_less [symmetric]) + not_less [symmetric]) next case (Suc n) - from prems show ?thesis + with Suc.prems Suc.hyps show ?thesis by (force dest: mult_left_le_imp_le - simp add: order_less_trans [OF zero_less_one gt1]) + simp add: less_trans [OF zero_less_one gt1]) qed qed text{*Surely we can strengthen this? It holds for @{text "0 (a^m = a^n) = (m=n)" + "1 < a \ a ^ m = a ^ n \ m = n" by (force simp add: order_antisym power_le_imp_le_exp) text{*Can relax the first premise to @{term "0 m < n" -by (simp add: order_less_le [of m n] order_less_le [of "a^m" "a^n"] - power_le_imp_le_exp) - + "1 < a \ a ^ m < a ^ n \ m < n" + by (simp add: order_less_le [of m n] less_le [of "a^m" "a^n"] + power_le_imp_le_exp) lemma power_mono: - "[|a \ b; (0::'a::{recpower,ordered_semidom}) \ a|] ==> a^n \ b^n" -apply (induct "n") -apply simp_all -apply (auto intro: mult_mono order_trans [of 0 a b]) -done + "a \ b \ 0 \ a \ a ^ n \ b ^ n" + by (induct n) + (auto intro: mult_mono order_trans [of 0 a b]) lemma power_strict_mono [rule_format]: - "[|a < b; (0::'a::{recpower,ordered_semidom}) \ a|] - ==> 0 < n --> a^n < b^n" -apply (induct "n") -apply (auto simp add: mult_strict_mono order_le_less_trans [of 0 a b]) -done - -lemma power_eq_0_iff [simp]: - "(a^n = 0) \ - (a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,recpower}) & n\0)" -apply (induct "n") -apply (auto simp add: no_zero_divisors) -done - - -lemma field_power_not_zero: - "a \ (0::'a::{ring_1_no_zero_divisors,recpower}) ==> a^n \ 0" -by force - -lemma nonzero_power_inverse: - fixes a :: "'a::{division_ring,recpower}" - shows "a \ 0 ==> inverse (a ^ n) = (inverse a) ^ n" -apply (induct "n") -apply (auto simp add: nonzero_inverse_mult_distrib power_commutes) -done (* TODO: reorient or rename to nonzero_inverse_power *) - -text{*Perhaps these should be simprules.*} -lemma power_inverse: - fixes a :: "'a::{division_ring,division_by_zero,recpower}" - shows "inverse (a ^ n) = (inverse a) ^ n" -apply (cases "a = 0") -apply (simp add: power_0_left) -apply (simp add: nonzero_power_inverse) -done (* TODO: reorient or rename to inverse_power *) - -lemma power_one_over: "1 / (a::'a::{field,division_by_zero,recpower})^n = - (1 / a)^n" -apply (simp add: divide_inverse) -apply (rule power_inverse) -done - -lemma nonzero_power_divide: - "b \ 0 ==> (a/b) ^ n = ((a::'a::{field,recpower}) ^ n) / (b ^ n)" -by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) - -lemma power_divide: - "(a/b) ^ n = ((a::'a::{field,division_by_zero,recpower}) ^ n / b ^ n)" -apply (case_tac "b=0", simp add: power_0_left) -apply (rule nonzero_power_divide) -apply assumption -done - -lemma power_abs: "abs(a ^ n) = abs(a::'a::{ordered_idom,recpower}) ^ n" -apply (induct "n") -apply (auto simp add: abs_mult) -done - -lemma abs_power_minus [simp]: - fixes a:: "'a::{ordered_idom,recpower}" shows "abs((-a) ^ n) = abs(a ^ n)" - by (simp add: abs_minus_cancel power_abs) - -lemma zero_less_power_abs_iff [simp,noatp]: - "(0 < (abs a)^n) = (a \ (0::'a::{ordered_idom,recpower}) | n=0)" -proof (induct "n") - case 0 - show ?case by simp -next - case (Suc n) - show ?case by (auto simp add: prems zero_less_mult_iff) -qed - -lemma zero_le_power_abs [simp]: - "(0::'a::{ordered_idom,recpower}) \ (abs a)^n" -by (rule zero_le_power [OF abs_ge_zero]) - -lemma power_minus: "(-a) ^ n = (- 1)^n * (a::'a::{ring_1,recpower}) ^ n" -proof (induct n) - case 0 show ?case by simp -next - case (Suc n) then show ?case - by (simp del: power_Suc add: power_Suc2 mult_assoc) -qed + "a < b \ 0 \ a \ 0 < n \ a ^ n < b ^ n" + by (induct n) + (auto simp add: mult_strict_mono le_less_trans [of 0 a b]) text{*Lemma for @{text power_strict_decreasing}*} lemma power_Suc_less: - "[|(0::'a::{ordered_semidom,recpower}) < a; a < 1|] - ==> a * a^n < a^n" -apply (induct n) -apply (auto simp add: mult_strict_left_mono) -done + "0 < a \ a < 1 \ a * a ^ n < a ^ n" + by (induct n) + (auto simp add: mult_strict_left_mono) -lemma power_strict_decreasing: - "[|n < N; 0 < a; a < (1::'a::{ordered_semidom,recpower})|] - ==> a^N < a^n" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: power_Suc_less less_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "a * a^m < 1 * a^n", simp) -apply (rule mult_strict_mono) -apply (auto simp add: order_less_imp_le) -done +lemma power_strict_decreasing [rule_format]: + "n < N \ 0 < a \ a < 1 \ a ^ N < a ^ n" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: power_Suc_less less_Suc_eq) + apply (subgoal_tac "a * a^N < 1 * a^n") + apply simp + apply (rule mult_strict_mono) apply auto + done +qed text{*Proof resembles that of @{text power_strict_decreasing}*} -lemma power_decreasing: - "[|n \ N; 0 \ a; a \ (1::'a::{ordered_semidom,recpower})|] - ==> a^N \ a^n" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: le_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "a * a^m \ 1 * a^n", simp) -apply (rule mult_mono) -apply auto -done +lemma power_decreasing [rule_format]: + "n \ N \ 0 \ a \ a \ 1 \ a ^ N \ a ^ n" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "a * a^N \ 1 * a^n", simp) + apply (rule mult_mono) apply auto + done +qed lemma power_Suc_less_one: - "[| 0 < a; a < (1::'a::{ordered_semidom,recpower}) |] ==> a ^ Suc n < 1" -apply (insert power_strict_decreasing [of 0 "Suc n" a], simp) -done + "0 < a \ a < 1 \ a ^ Suc n < 1" + using power_strict_decreasing [of 0 "Suc n" a] by simp text{*Proof again resembles that of @{text power_strict_decreasing}*} -lemma power_increasing: - "[|n \ N; (1::'a::{ordered_semidom,recpower}) \ a|] ==> a^n \ a^N" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: le_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "1 * a^n \ a * a^m", simp) -apply (rule mult_mono) -apply (auto simp add: order_trans [OF zero_le_one]) -done +lemma power_increasing [rule_format]: + "n \ N \ 1 \ a \ a ^ n \ a ^ N" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: le_Suc_eq) + apply (subgoal_tac "1 * a^n \ a * a^N", simp) + apply (rule mult_mono) apply (auto simp add: order_trans [OF zero_le_one]) + done +qed text{*Lemma for @{text power_strict_increasing}*} lemma power_less_power_Suc: - "(1::'a::{ordered_semidom,recpower}) < a ==> a^n < a * a^n" -apply (induct n) -apply (auto simp add: mult_strict_left_mono order_less_trans [OF zero_less_one]) -done + "1 < a \ a ^ n < a * a ^ n" + by (induct n) (auto simp add: mult_strict_left_mono less_trans [OF zero_less_one]) -lemma power_strict_increasing: - "[|n < N; (1::'a::{ordered_semidom,recpower}) < a|] ==> a^n < a^N" -apply (erule rev_mp) -apply (induct "N") -apply (auto simp add: power_less_power_Suc less_Suc_eq) -apply (rename_tac m) -apply (subgoal_tac "1 * a^n < a * a^m", simp) -apply (rule mult_strict_mono) -apply (auto simp add: order_less_trans [OF zero_less_one] order_less_imp_le) -done +lemma power_strict_increasing [rule_format]: + "n < N \ 1 < a \ a ^ n < a ^ N" +proof (induct N) + case 0 then show ?case by simp +next + case (Suc N) then show ?case + apply (auto simp add: power_less_power_Suc less_Suc_eq) + apply (subgoal_tac "1 * a^n < a * a^N", simp) + apply (rule mult_strict_mono) apply (auto simp add: less_trans [OF zero_less_one] less_imp_le) + done +qed lemma power_increasing_iff [simp]: - "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x \ b ^ y) = (x \ y)" -by (blast intro: power_le_imp_le_exp power_increasing order_less_imp_le) + "1 < b \ b ^ x \ b ^ y \ x \ y" + by (blast intro: power_le_imp_le_exp power_increasing less_imp_le) lemma power_strict_increasing_iff [simp]: - "1 < (b::'a::{ordered_semidom,recpower}) ==> (b ^ x < b ^ y) = (x < y)" + "1 < b \ b ^ x < b ^ y \ x < y" by (blast intro: power_less_imp_less_exp power_strict_increasing) lemma power_le_imp_le_base: -assumes le: "a ^ Suc n \ b ^ Suc n" - and ynonneg: "(0::'a::{ordered_semidom,recpower}) \ b" -shows "a \ b" + assumes le: "a ^ Suc n \ b ^ Suc n" + and ynonneg: "0 \ b" + shows "a \ b" proof (rule ccontr) assume "~ a \ b" then have "b < a" by (simp only: linorder_not_le) then have "b ^ Suc n < a ^ Suc n" by (simp only: prems power_strict_mono) - from le and this show "False" + from le and this show False by (simp add: linorder_not_less [symmetric]) qed lemma power_less_imp_less_base: - fixes a b :: "'a::{ordered_semidom,recpower}" assumes less: "a ^ n < b ^ n" assumes nonneg: "0 \ b" shows "a < b" @@ -310,83 +310,144 @@ assume "~ a < b" hence "b \ a" by (simp only: linorder_not_less) hence "b ^ n \ a ^ n" using nonneg by (rule power_mono) - thus "~ a ^ n < b ^ n" by (simp only: linorder_not_less) + thus "\ a ^ n < b ^ n" by (simp only: linorder_not_less) qed lemma power_inject_base: - "[| a ^ Suc n = b ^ Suc n; 0 \ a; 0 \ b |] - ==> a = (b::'a::{ordered_semidom,recpower})" -by (blast intro: power_le_imp_le_base order_antisym order_eq_refl sym) + "a ^ Suc n = b ^ Suc n \ 0 \ a \ 0 \ b \ a = b" +by (blast intro: power_le_imp_le_base antisym eq_refl sym) lemma power_eq_imp_eq_base: - fixes a b :: "'a::{ordered_semidom,recpower}" - shows "\a ^ n = b ^ n; 0 \ a; 0 \ b; 0 < n\ \ a = b" -by (cases n, simp_all del: power_Suc, rule power_inject_base) + "a ^ n = b ^ n \ 0 \ a \ 0 \ b \ 0 < n \ a = b" + by (cases n) (simp_all del: power_Suc, rule power_inject_base) -text {* The divides relation *} +end + +context ordered_idom +begin -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)" . +lemma power_abs: + "abs (a ^ n) = abs a ^ n" + by (induct n) (auto simp add: abs_mult) + +lemma abs_power_minus [simp]: + "abs ((-a) ^ n) = abs (a ^ n)" + by (simp add: abs_minus_cancel power_abs) + +lemma zero_less_power_abs_iff [simp, noatp]: + "0 < abs a ^ n \ a \ 0 \ n = 0" +proof (induct n) + case 0 show ?case by simp +next + case (Suc n) show ?case by (auto simp add: Suc zero_less_mult_iff) 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]) +lemma zero_le_power_abs [simp]: + "0 \ abs a ^ n" + by (rule zero_le_power [OF abs_ge_zero]) + +end + +context ring_1_no_zero_divisors +begin + +lemma field_power_not_zero: + "a \ 0 \ a ^ n \ 0" + by (induct n) auto + +end + +context division_ring +begin +text {* FIXME reorient or rename to @{text nonzero_inverse_power} *} +lemma nonzero_power_inverse: + "a \ 0 \ inverse (a ^ n) = (inverse a) ^ n" + by (induct n) + (simp_all add: nonzero_inverse_mult_distrib power_commutes field_power_not_zero) -lemma dvd_power_same: - "(x::'a::{comm_semiring_1,recpower}) dvd y \ x^n dvd y^n" -by (induct n) (auto simp add: mult_dvd_mono) +end + +context field +begin + +lemma nonzero_power_divide: + "b \ 0 \ (a / b) ^ n = a ^ n / b ^ n" + by (simp add: divide_inverse power_mult_distrib nonzero_power_inverse) + +end + +lemma power_0_Suc [simp]: + "(0::'a::{power, semiring_0}) ^ Suc n = 0" + by simp -lemma dvd_power_le: - "(x::'a::{comm_semiring_1,recpower}) dvd y \ m >= n \ x^n dvd y^m" -by(rule power_le_dvd[OF dvd_power_same]) +text{*It looks plausible as a simprule, but its effect can be strange.*} +lemma power_0_left: + "0 ^ n = (if n = 0 then 1 else (0::'a::{power, semiring_0}))" + by (induct n) simp_all + +lemma power_eq_0_iff [simp]: + "a ^ n = 0 \ + a = (0::'a::{mult_zero,zero_neq_one,no_zero_divisors,power}) \ n \ 0" + by (induct n) + (auto simp add: no_zero_divisors elim: contrapos_pp) + +lemma power_diff: + fixes a :: "'a::field" + assumes nz: "a \ 0" + shows "n \ m \ a ^ (m - n) = a ^ m / a ^ n" + by (induct m n rule: diff_induct) (simp_all add: nz) -lemma dvd_power [simp]: - "n > 0 | (x::'a::{comm_semiring_1,recpower}) = 1 \ x dvd x^n" -apply (erule disjE) - apply (subgoal_tac "x ^ n = x^(Suc (n - 1))") - apply (erule ssubst) - apply (subst power_Suc) - apply auto +text{*Perhaps these should be simprules.*} +lemma power_inverse: + fixes a :: "'a::{division_ring,division_by_zero,power}" + shows "inverse (a ^ n) = (inverse a) ^ n" +apply (cases "a = 0") +apply (simp add: power_0_left) +apply (simp add: nonzero_power_inverse) +done (* TODO: reorient or rename to inverse_power *) + +lemma power_one_over: + "1 / (a::'a::{field,division_by_zero, power}) ^ n = (1 / a) ^ n" + by (simp add: divide_inverse) (rule power_inverse) + +lemma power_divide: + "(a / b) ^ n = (a::'a::{field,division_by_zero}) ^ n / b ^ n" +apply (cases "b = 0") +apply (simp add: power_0_left) +apply (rule nonzero_power_divide) +apply assumption done +class recpower = monoid_mult + subsection {* Exponentiation for the Natural Numbers *} instance nat :: recpower .. -lemma of_nat_power: - "of_nat (m ^ n) = (of_nat m::'a::{semiring_1,recpower}) ^ n" -by (induct n, simp_all add: of_nat_mult) +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_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_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 m, auto) + "x ^ m = Suc 0 \ m = 0 \ x = Suc 0" + by (induct m) auto -lemma power_Suc_0[simp]: "(Suc 0)^n = Suc 0" -by simp +lemma power_Suc_0 [simp]: + "Suc 0 ^ n = Suc 0" + by simp text{*Valid for the naturals, but what if @{text"0nat)" - assumes less: "i^m < i^n" + assumes less: "i ^ m < i ^ n" shows "m < n" proof (cases "i = 1") case True with less power_one [where 'a = nat] show ?thesis by simp @@ -395,10 +456,4 @@ from power_strict_increasing_iff [OF this] less show ?thesis .. qed -lemma power_diff: - assumes nz: "a ~= 0" - shows "n <= m ==> (a::'a::{recpower, field}) ^ (m-n) = (a^m) / (a^n)" - by (induct m n rule: diff_induct) - (simp_all add: nonzero_mult_divide_cancel_left nz) - end