src/HOL/Ln.thy
changeset 50326 b5afeccab2db
parent 47244 a7f85074c169
equal deleted inserted replaced
50325:5e40ad9f212a 50326:b5afeccab2db
     5 header {* Properties of ln *}
     5 header {* Properties of ln *}
     6 
     6 
     7 theory Ln
     7 theory Ln
     8 imports Transcendental
     8 imports Transcendental
     9 begin
     9 begin
    10 
       
    11 lemma exp_first_two_terms: "exp x = 1 + x + suminf (%n. 
       
    12   inverse(fact (n+2)) * (x ^ (n+2)))"
       
    13 proof -
       
    14   have "exp x = suminf (%n. inverse(fact n) * (x ^ n))"
       
    15     by (simp add: exp_def)
       
    16   also from summable_exp have "... = (SUM n::nat : {0..<2}. 
       
    17       inverse(fact n) * (x ^ n)) + suminf (%n.
       
    18       inverse(fact(n+2)) * (x ^ (n+2)))" (is "_ = ?a + _")
       
    19     by (rule suminf_split_initial_segment)
       
    20   also have "?a = 1 + x"
       
    21     by (simp add: numeral_2_eq_2)
       
    22   finally show ?thesis .
       
    23 qed
       
    24 
       
    25 lemma exp_bound: "0 <= (x::real) ==> x <= 1 ==> exp x <= 1 + x + x^2"
       
    26 proof -
       
    27   assume a: "0 <= x"
       
    28   assume b: "x <= 1"
       
    29   { fix n :: nat
       
    30     have "2 * 2 ^ n \<le> fact (n + 2)"
       
    31       by (induct n, simp, simp)
       
    32     hence "real ((2::nat) * 2 ^ n) \<le> real (fact (n + 2))"
       
    33       by (simp only: real_of_nat_le_iff)
       
    34     hence "2 * 2 ^ n \<le> real (fact (n + 2))"
       
    35       by simp
       
    36     hence "inverse (fact (n + 2)) \<le> inverse (2 * 2 ^ n)"
       
    37       by (rule le_imp_inverse_le) simp
       
    38     hence "inverse (fact (n + 2)) \<le> 1/2 * (1/2)^n"
       
    39       by (simp add: inverse_mult_distrib power_inverse)
       
    40     hence "inverse (fact (n + 2)) * (x^n * x\<twosuperior>) \<le> 1/2 * (1/2)^n * (1 * x\<twosuperior>)"
       
    41       by (rule mult_mono)
       
    42         (rule mult_mono, simp_all add: power_le_one a b mult_nonneg_nonneg)
       
    43     hence "inverse (fact (n + 2)) * x ^ (n + 2) \<le> (x\<twosuperior>/2) * ((1/2)^n)"
       
    44       unfolding power_add by (simp add: mult_ac del: fact_Suc) }
       
    45   note aux1 = this
       
    46   have "(\<lambda>n. x\<twosuperior> / 2 * (1 / 2) ^ n) sums (x\<twosuperior> / 2 * (1 / (1 - 1 / 2)))"
       
    47     by (intro sums_mult geometric_sums, simp)
       
    48   hence aux2: "(\<lambda>n. (x::real) ^ 2 / 2 * (1 / 2) ^ n) sums x^2"
       
    49     by simp
       
    50   have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
       
    51   proof -
       
    52     have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <=
       
    53         suminf (%n. (x^2/2) * ((1/2)^n))"
       
    54       apply (rule summable_le)
       
    55       apply (rule allI, rule aux1)
       
    56       apply (rule summable_exp [THEN summable_ignore_initial_segment])
       
    57       by (rule sums_summable, rule aux2)
       
    58     also have "... = x^2"
       
    59       by (rule sums_unique [THEN sym], rule aux2)
       
    60     finally show ?thesis .
       
    61   qed
       
    62   thus ?thesis unfolding exp_first_two_terms by auto
       
    63 qed
       
    64 
    10 
    65 lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
    11 lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
    66     x - x^2 <= ln (1 + x)"
    12     x - x^2 <= ln (1 + x)"
    67 proof -
    13 proof -
    68   assume a: "0 <= x" and b: "x <= 1"
    14   assume a: "0 <= x" and b: "x <= 1"
    88     from a have "0 < 1 + x" by auto
    34     from a have "0 < 1 + x" by auto
    89     thus ?thesis
    35     thus ?thesis
    90       by (auto simp only: exp_ln_iff [THEN sym])
    36       by (auto simp only: exp_ln_iff [THEN sym])
    91   qed
    37   qed
    92   finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" .
    38   finally have "exp (x - x ^ 2) <= exp (ln (1 + x))" .
    93   thus ?thesis by (auto simp only: exp_le_cancel_iff)
       
    94 qed
       
    95 
       
    96 lemma ln_one_minus_pos_upper_bound: "0 <= x ==> x < 1 ==> ln (1 - x) <= - x"
       
    97 proof -
       
    98   assume a: "0 <= (x::real)" and b: "x < 1"
       
    99   have "(1 - x) * (1 + x + x^2) = (1 - x^3)"
       
   100     by (simp add: algebra_simps power2_eq_square power3_eq_cube)
       
   101   also have "... <= 1"
       
   102     by (auto simp add: a)
       
   103   finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
       
   104   moreover have c: "0 < 1 + x + x\<twosuperior>"
       
   105     by (simp add: add_pos_nonneg a)
       
   106   ultimately have "1 - x <= 1 / (1 + x + x^2)"
       
   107     by (elim mult_imp_le_div_pos)
       
   108   also have "... <= 1 / exp x"
       
   109     apply (rule divide_left_mono)
       
   110     apply (rule exp_bound, rule a)
       
   111     apply (rule b [THEN less_imp_le])
       
   112     apply simp
       
   113     apply (rule mult_pos_pos)
       
   114     apply (rule c)
       
   115     apply simp
       
   116     done
       
   117   also have "... = exp (-x)"
       
   118     by (auto simp add: exp_minus divide_inverse)
       
   119   finally have "1 - x <= exp (- x)" .
       
   120   also have "1 - x = exp (ln (1 - x))"
       
   121   proof -
       
   122     have "0 < 1 - x"
       
   123       by (insert b, auto)
       
   124     thus ?thesis
       
   125       by (auto simp only: exp_ln_iff [THEN sym])
       
   126   qed
       
   127   finally have "exp (ln (1 - x)) <= exp (- x)" .
       
   128   thus ?thesis by (auto simp only: exp_le_cancel_iff)
    39   thus ?thesis by (auto simp only: exp_le_cancel_iff)
   129 qed
    40 qed
   130 
    41 
   131 lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))"
    42 lemma aux5: "x < 1 ==> ln(1 - x) = - ln(1 + x / (1 - x))"
   132 proof -
    43 proof -
   171     using mult_right_le_one_le[of "x*x" "2*x"] a b
    82     using mult_right_le_one_le[of "x*x" "2*x"] a b
   172     by (simp add:field_simps power2_eq_square)
    83     by (simp add:field_simps power2_eq_square)
   173   from e d show "- x - 2 * x^2 <= ln (1 - x)"
    84   from e d show "- x - 2 * x^2 <= ln (1 - x)"
   174     by (rule order_trans)
    85     by (rule order_trans)
   175 qed
    86 qed
   176 
       
   177 lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x"
       
   178   apply (case_tac "0 <= x")
       
   179   apply (erule exp_ge_add_one_self_aux)
       
   180   apply (case_tac "x <= -1")
       
   181   apply (subgoal_tac "1 + x <= 0")
       
   182   apply (erule order_trans)
       
   183   apply simp
       
   184   apply simp
       
   185   apply (subgoal_tac "1 + x = exp(ln (1 + x))")
       
   186   apply (erule ssubst)
       
   187   apply (subst exp_le_cancel_iff)
       
   188   apply (subgoal_tac "ln (1 - (- x)) <= - (- x)")
       
   189   apply simp
       
   190   apply (rule ln_one_minus_pos_upper_bound) 
       
   191   apply auto
       
   192 done
       
   193 
    87 
   194 lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
    88 lemma ln_add_one_self_le_self2: "-1 < x ==> ln(1 + x) <= x"
   195   apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
    89   apply (subgoal_tac "ln (1 + x) \<le> ln (exp x)", simp)
   196   apply (subst ln_le_cancel_iff)
    90   apply (subst ln_le_cancel_iff)
   197   apply auto
    91   apply auto