tuned proofs
authorhuffman
Sun, 01 Apr 2012 09:12:03 +0200
changeset 48115a7f85074c169
parent 48114 403b363c1387
child 48116 30a1692557b0
child 48118 880e587eee9f
child 48121 ff1770df59b8
tuned proofs
src/HOL/Ln.thy
     1.1 --- a/src/HOL/Ln.thy	Sat Mar 31 22:45:46 2012 +0200
     1.2 +++ b/src/HOL/Ln.thy	Sun Apr 01 09:12:03 2012 +0200
     1.3 @@ -22,73 +22,48 @@
     1.4    finally show ?thesis .
     1.5  qed
     1.6  
     1.7 -lemma exp_tail_after_first_two_terms_summable: 
     1.8 -  "summable (%n. inverse(fact (n+2)) * (x ^ (n+2)))"
     1.9 -proof -
    1.10 -  note summable_exp
    1.11 -  thus ?thesis
    1.12 -    by (frule summable_ignore_initial_segment)
    1.13 -qed
    1.14 -
    1.15 -lemma aux1: assumes a: "0 <= x" and b: "x <= 1"
    1.16 -    shows "inverse (fact ((n::nat) + 2)) * x ^ (n + 2) <= (x^2/2) * ((1/2)^n)"
    1.17 -proof -
    1.18 -  have "2 * 2 ^ n \<le> fact (n + 2)"
    1.19 -    by (induct n, simp, simp)
    1.20 -  hence "real ((2::nat) * 2 ^ n) \<le> real (fact (n + 2))"
    1.21 -    by (simp only: real_of_nat_le_iff)
    1.22 -  hence "2 * 2 ^ n \<le> real (fact (n + 2))"
    1.23 -    by simp
    1.24 -  hence "inverse (fact (n + 2)) \<le> inverse (2 * 2 ^ n)"
    1.25 -    by (rule le_imp_inverse_le) simp
    1.26 -  hence "inverse (fact (n + 2)) \<le> 1/2 * (1/2)^n"
    1.27 -    by (simp add: inverse_mult_distrib power_inverse)
    1.28 -  hence "inverse (fact (n + 2)) * (x^n * x\<twosuperior>) \<le> 1/2 * (1/2)^n * (1 * x\<twosuperior>)"
    1.29 -    by (rule mult_mono)
    1.30 -      (rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg)
    1.31 -  thus ?thesis
    1.32 -    unfolding power_add by (simp add: mult_ac del: fact_Suc)
    1.33 -qed
    1.34 -
    1.35 -lemma aux2: "(%n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2"
    1.36 -proof -
    1.37 -  have "(%n. (1 / 2::real)^n) sums (1 / (1 - (1/2)))"
    1.38 -    apply (rule geometric_sums)
    1.39 -    by (simp add: abs_less_iff)
    1.40 -  also have "(1::real) / (1 - 1/2) = 2"
    1.41 -    by simp
    1.42 -  finally have "(%n. (1 / 2::real)^n) sums 2" .
    1.43 -  then have "(%n. x ^ 2 / 2 * (1 / 2) ^ n) sums (x^2 / 2 * 2)"
    1.44 -    by (rule sums_mult)
    1.45 -  also have "x^2 / 2 * 2 = x^2"
    1.46 -    by simp
    1.47 -  finally show ?thesis .
    1.48 -qed
    1.49 -
    1.50  lemma exp_bound: "0 <= (x::real) ==> x <= 1 ==> exp x <= 1 + x + x^2"
    1.51  proof -
    1.52    assume a: "0 <= x"
    1.53    assume b: "x <= 1"
    1.54 -  have c: "exp x = 1 + x + suminf (%n. inverse(fact (n+2)) * 
    1.55 -      (x ^ (n+2)))"
    1.56 -    by (rule exp_first_two_terms)
    1.57 -  moreover have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
    1.58 +  { fix n :: nat
    1.59 +    have "2 * 2 ^ n \<le> fact (n + 2)"
    1.60 +      by (induct n, simp, simp)
    1.61 +    hence "real ((2::nat) * 2 ^ n) \<le> real (fact (n + 2))"
    1.62 +      by (simp only: real_of_nat_le_iff)
    1.63 +    hence "2 * 2 ^ n \<le> real (fact (n + 2))"
    1.64 +      by simp
    1.65 +    hence "inverse (fact (n + 2)) \<le> inverse (2 * 2 ^ n)"
    1.66 +      by (rule le_imp_inverse_le) simp
    1.67 +    hence "inverse (fact (n + 2)) \<le> 1/2 * (1/2)^n"
    1.68 +      by (simp add: inverse_mult_distrib power_inverse)
    1.69 +    hence "inverse (fact (n + 2)) * (x^n * x\<twosuperior>) \<le> 1/2 * (1/2)^n * (1 * x\<twosuperior>)"
    1.70 +      by (rule mult_mono)
    1.71 +        (rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg)
    1.72 +    hence "inverse (fact (n + 2)) * x ^ (n + 2) \<le> (x\<twosuperior>/2) * ((1/2)^n)"
    1.73 +      unfolding power_add by (simp add: mult_ac del: fact_Suc) }
    1.74 +  note aux1 = this
    1.75 +  have "(\<lambda>n. x\<twosuperior> / 2 * (1 / 2) ^ n) sums (x\<twosuperior> / 2 * (1 / (1 - 1 / 2)))"
    1.76 +    by (intro sums_mult geometric_sums, simp)
    1.77 +  hence aux2: "(\<lambda>n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2"
    1.78 +    by simp
    1.79 +  have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
    1.80    proof -
    1.81      have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <=
    1.82          suminf (%n. (x^2/2) * ((1/2)^n))"
    1.83        apply (rule summable_le)
    1.84 -      apply (auto simp only: aux1 a b)
    1.85 -      apply (rule exp_tail_after_first_two_terms_summable)
    1.86 -      by (rule sums_summable, rule aux2)  
    1.87 +      apply (rule allI, rule aux1)
    1.88 +      apply (rule summable_exp [THEN summable_ignore_initial_segment])
    1.89 +      by (rule sums_summable, rule aux2)
    1.90      also have "... = x^2"
    1.91        by (rule sums_unique [THEN sym], rule aux2)
    1.92      finally show ?thesis .
    1.93    qed
    1.94 -  ultimately show ?thesis
    1.95 -    by auto
    1.96 +  thus ?thesis unfolding exp_first_two_terms by auto
    1.97  qed
    1.98  
    1.99 -lemma aux4: "0 <= (x::real) ==> x <= 1 ==> exp (x - x^2) <= 1 + x" 
   1.100 +lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
   1.101 +    x - x^2 <= ln (1 + x)"
   1.102  proof -
   1.103    assume a: "0 <= x" and b: "x <= 1"
   1.104    have "exp (x - x^2) = exp x / exp (x^2)"
   1.105 @@ -101,25 +76,13 @@
   1.106      done
   1.107    also have "... <= (1 + x + x^2) / (1 + x^2)"
   1.108      apply (rule divide_left_mono)
   1.109 -    apply (auto simp add: exp_ge_add_one_self_aux)
   1.110 -    apply (rule add_nonneg_nonneg)
   1.111 -    using a apply auto
   1.112 -    apply (rule mult_pos_pos)
   1.113 -    apply auto
   1.114 -    apply (rule add_pos_nonneg)
   1.115 -    apply auto
   1.116 +    apply (simp add: exp_ge_add_one_self_aux)
   1.117 +    apply (simp add: a)
   1.118 +    apply (simp add: mult_pos_pos add_pos_nonneg)
   1.119      done
   1.120    also from a have "... <= 1 + x"
   1.121      by (simp add: field_simps add_strict_increasing zero_le_mult_iff)
   1.122 -  finally show ?thesis .
   1.123 -qed
   1.124 -
   1.125 -lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
   1.126 -    x - x^2 <= ln (1 + x)"
   1.127 -proof -
   1.128 -  assume a: "0 <= x" and b: "x <= 1"
   1.129 -  then have "exp (x - x^2) <= 1 + x"
   1.130 -    by (rule aux4)
   1.131 +  finally have "exp (x - x^2) <= 1 + x" .
   1.132    also have "... = exp (ln (1 + x))"
   1.133    proof -
   1.134      from a have "0 < 1 + x" by auto
   1.135 @@ -138,19 +101,18 @@
   1.136    also have "... <= 1"
   1.137      by (auto simp add: a)
   1.138    finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
   1.139 -  moreover have "0 < 1 + x + x^2"
   1.140 -    apply (rule add_pos_nonneg)
   1.141 -    using a apply auto
   1.142 -    done
   1.143 +  moreover have c: "0 < 1 + x + x\<twosuperior>"
   1.144 +    by (simp add: add_pos_nonneg a)
   1.145    ultimately have "1 - x <= 1 / (1 + x + x^2)"
   1.146      by (elim mult_imp_le_div_pos)
   1.147    also have "... <= 1 / exp x"
   1.148      apply (rule divide_left_mono)
   1.149      apply (rule exp_bound, rule a)
   1.150 -    using a b apply auto
   1.151 +    apply (rule b [THEN less_imp_le])
   1.152 +    apply simp
   1.153      apply (rule mult_pos_pos)
   1.154 -    apply (rule add_pos_nonneg)
   1.155 -    apply auto
   1.156 +    apply (rule c)
   1.157 +    apply simp
   1.158      done
   1.159    also have "... = exp (-x)"
   1.160      by (auto simp add: exp_minus divide_inverse)
   1.161 @@ -230,8 +192,7 @@
   1.162  done
   1.163  
   1.164  lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
   1.165 -  apply (subgoal_tac "x = ln (exp x)")
   1.166 -  apply (erule ssubst)back
   1.167 +  apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
   1.168    apply (subst ln_le_cancel_iff)
   1.169    apply auto
   1.170  done