# HG changeset patch # User huffman # Date 1333226746 -7200 # Node ID 403b363c138749ba123c2247df1711c76c3591d4 # Parent 1caeecc72aea02838c759dd6803334ff12820bd0# Parent 72ab1fbf2f41097cd52df9084705a2d54108fc07 merged diff -r 72ab1fbf2f41 -r 403b363c1387 src/HOL/Ln.thy --- a/src/HOL/Ln.thy Sat Mar 31 19:38:41 2012 +0200 +++ b/src/HOL/Ln.thy Sat Mar 31 22:45:46 2012 +0200 @@ -32,68 +32,22 @@ lemma aux1: assumes a: "0 <= x" and b: "x <= 1" shows "inverse (fact ((n::nat) + 2)) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)" -proof (induct n) - show "inverse (fact ((0::nat) + 2)) * x ^ (0 + 2) <= - x ^ 2 / 2 * (1 / 2) ^ 0" - by (simp add: real_of_nat_Suc power2_eq_square) -next - fix n :: nat - assume c: "inverse (fact (n + 2)) * x ^ (n + 2) - <= x ^ 2 / 2 * (1 / 2) ^ n" - show "inverse (fact (Suc n + 2)) * x ^ (Suc n + 2) - <= x ^ 2 / 2 * (1 / 2) ^ Suc n" - proof - - have "inverse(fact (Suc n + 2)) <= (1/2) * inverse (fact (n+2))" - proof - - have "Suc n + 2 = Suc (n + 2)" by simp - then have "fact (Suc n + 2) = Suc (n + 2) * fact (n + 2)" - by simp - then have "real(fact (Suc n + 2)) = real(Suc (n + 2) * fact (n + 2))" - apply (rule subst) - apply (rule refl) - done - also have "... = real(Suc (n + 2)) * real(fact (n + 2))" - by (rule real_of_nat_mult) - finally have "real (fact (Suc n + 2)) = - real (Suc (n + 2)) * real (fact (n + 2))" . - then have "inverse(fact (Suc n + 2)) = - inverse(Suc (n + 2)) * inverse(fact (n + 2))" - apply (rule ssubst) - apply (rule inverse_mult_distrib) - done - also have "... <= (1/2) * inverse(fact (n + 2))" - apply (rule mult_right_mono) - apply (subst inverse_eq_divide) - apply simp - apply (simp del: fact_Suc) - done - finally show ?thesis . - qed - moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)" - by (simp add: mult_left_le_one_le mult_nonneg_nonneg a b) - ultimately have "inverse (fact (Suc n + 2)) * x ^ (Suc n + 2) <= - (1 / 2 * inverse (fact (n + 2))) * x ^ (n + 2)" - apply (rule mult_mono) - apply (rule mult_nonneg_nonneg) - apply simp - apply (subst inverse_nonnegative_iff_nonnegative) - apply (rule real_of_nat_ge_zero) - apply (rule zero_le_power) - apply (rule a) - done - also have "... = 1 / 2 * (inverse (fact (n + 2)) * x ^ (n + 2))" - by simp - also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)" - apply (rule mult_left_mono) - apply (rule c) - apply simp - done - also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)" - by auto - also have "(1::real) / 2 * (1 / 2) ^ n = (1 / 2) ^ (Suc n)" - by (rule power_Suc [THEN sym]) - finally show ?thesis . - qed +proof - + have "2 * 2 ^ n \ fact (n + 2)" + by (induct n, simp, simp) + hence "real ((2::nat) * 2 ^ n) \ real (fact (n + 2))" + by (simp only: real_of_nat_le_iff) + hence "2 * 2 ^ n \ real (fact (n + 2))" + by simp + hence "inverse (fact (n + 2)) \ inverse (2 * 2 ^ n)" + by (rule le_imp_inverse_le) simp + hence "inverse (fact (n + 2)) \ 1/2 * (1/2)^n" + by (simp add: inverse_mult_distrib power_inverse) + hence "inverse (fact (n + 2)) * (x^n * x\) \ 1/2 * (1/2)^n * (1 * x\)" + by (rule mult_mono) + (rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg) + thus ?thesis + unfolding power_add by (simp add: mult_ac del: fact_Suc) qed lemma aux2: "(%n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2" diff -r 72ab1fbf2f41 -r 403b363c1387 src/HOL/Power.thy --- a/src/HOL/Power.thy Sat Mar 31 19:38:41 2012 +0200 +++ b/src/HOL/Power.thy Sat Mar 31 22:45:46 2012 +0200 @@ -289,13 +289,15 @@ "0 \ a \ 0 \ a ^ n" by (induct n) (simp_all add: mult_nonneg_nonneg) -lemma one_le_power[simp]: - "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_mono: + "a \ b \ 0 \ a \ a ^ n \ b ^ n" + by (induct n) (auto intro: mult_mono order_trans [of 0 a b]) + +lemma one_le_power [simp]: "1 \ a \ 1 \ a ^ n" + using power_mono [of 1 a n] by simp + +lemma power_le_one: "\0 \ a; a \ 1\ \ a ^ n \ 1" + using power_mono [of a 1 n] by simp lemma power_gt1_lemma: assumes gt1: "1 < a" @@ -353,11 +355,6 @@ 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 \ 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 \ 0 < n \ a ^ n < b ^ n" by (induct n)