src/HOL/Ln.thy
changeset 41550 efa734d9b221
parent 40864 4abaaadfdaf2
child 41959 b460124855b8
equal deleted inserted replaced
41549:2c65ad10bec8 41550:efa734d9b221
    69         done
    69         done
    70       finally show ?thesis .
    70       finally show ?thesis .
    71     qed
    71     qed
    72     moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)"
    72     moreover have "x ^ (Suc n + 2) <= x ^ (n + 2)"
    73       apply (simp add: mult_compare_simps)
    73       apply (simp add: mult_compare_simps)
    74       apply (simp add: prems)
    74       apply (simp add: assms)
    75       apply (subgoal_tac "0 <= x * (x * x^n)")
    75       apply (subgoal_tac "0 <= x * (x * x^n)")
    76       apply force
    76       apply force
    77       apply (rule mult_nonneg_nonneg, rule a)+
    77       apply (rule mult_nonneg_nonneg, rule a)+
    78       apply (rule zero_le_power, rule a)
    78       apply (rule zero_le_power, rule a)
    79       done
    79       done
    89       done
    89       done
    90     also have "... = 1 / 2 * (inverse (fact (n + 2)) * x ^ (n + 2))"
    90     also have "... = 1 / 2 * (inverse (fact (n + 2)) * x ^ (n + 2))"
    91       by simp
    91       by simp
    92     also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)"
    92     also have "... <= 1 / 2 * (x ^ 2 / 2 * (1 / 2) ^ n)"
    93       apply (rule mult_left_mono)
    93       apply (rule mult_left_mono)
    94       apply (rule prems)
    94       apply (rule c)
    95       apply simp
    95       apply simp
    96       done
    96       done
    97     also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)"
    97     also have "... = x ^ 2 / 2 * (1 / 2 * (1 / 2) ^ n)"
    98       by auto
    98       by auto
    99     also have "(1::real) / 2 * (1 / 2) ^ n = (1 / 2) ^ (Suc n)"
    99     also have "(1::real) / 2 * (1 / 2) ^ n = (1 / 2) ^ (Suc n)"
   127   moreover have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
   127   moreover have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <= x^2"
   128   proof -
   128   proof -
   129     have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <=
   129     have "suminf (%n. inverse(fact (n+2)) * (x ^ (n+2))) <=
   130         suminf (%n. (x^2/2) * ((1/2)^n))"
   130         suminf (%n. (x^2/2) * ((1/2)^n))"
   131       apply (rule summable_le)
   131       apply (rule summable_le)
   132       apply (auto simp only: aux1 prems)
   132       apply (auto simp only: aux1 a b)
   133       apply (rule exp_tail_after_first_two_terms_summable)
   133       apply (rule exp_tail_after_first_two_terms_summable)
   134       by (rule sums_summable, rule aux2)  
   134       by (rule sums_summable, rule aux2)  
   135     also have "... = x^2"
   135     also have "... = x^2"
   136       by (rule sums_unique [THEN sym], rule aux2)
   136       by (rule sums_unique [THEN sym], rule aux2)
   137     finally show ?thesis .
   137     finally show ?thesis .
   153     done
   153     done
   154   also have "... <= (1 + x + x^2) / (1 + x^2)"
   154   also have "... <= (1 + x + x^2) / (1 + x^2)"
   155     apply (rule divide_left_mono)
   155     apply (rule divide_left_mono)
   156     apply (auto simp add: exp_ge_add_one_self_aux)
   156     apply (auto simp add: exp_ge_add_one_self_aux)
   157     apply (rule add_nonneg_nonneg)
   157     apply (rule add_nonneg_nonneg)
   158     apply (insert prems, auto)
   158     using a apply auto
   159     apply (rule mult_pos_pos)
   159     apply (rule mult_pos_pos)
   160     apply auto
   160     apply auto
   161     apply (rule add_pos_nonneg)
   161     apply (rule add_pos_nonneg)
   162     apply auto
   162     apply auto
   163     done
   163     done
   164   also from a have "... <= 1 + x"
   164   also from a have "... <= 1 + x"
   165     by(simp add:field_simps zero_compare_simps)
   165     by (simp add: field_simps zero_compare_simps)
   166   finally show ?thesis .
   166   finally show ?thesis .
   167 qed
   167 qed
   168 
   168 
   169 lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
   169 lemma ln_one_plus_pos_lower_bound: "0 <= x ==> x <= 1 ==> 
   170     x - x^2 <= ln (1 + x)"
   170     x - x^2 <= ln (1 + x)"
   190   also have "... <= 1"
   190   also have "... <= 1"
   191     by (auto simp add: a)
   191     by (auto simp add: a)
   192   finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
   192   finally have "(1 - x) * (1 + x + x ^ 2) <= 1" .
   193   moreover have "0 < 1 + x + x^2"
   193   moreover have "0 < 1 + x + x^2"
   194     apply (rule add_pos_nonneg)
   194     apply (rule add_pos_nonneg)
   195     apply (insert a, auto)
   195     using a apply auto
   196     done
   196     done
   197   ultimately have "1 - x <= 1 / (1 + x + x^2)"
   197   ultimately have "1 - x <= 1 / (1 + x + x^2)"
   198     by (elim mult_imp_le_div_pos)
   198     by (elim mult_imp_le_div_pos)
   199   also have "... <= 1 / exp x"
   199   also have "... <= 1 / exp x"
   200     apply (rule divide_left_mono)
   200     apply (rule divide_left_mono)
   201     apply (rule exp_bound, rule a)
   201     apply (rule exp_bound, rule a)
   202     apply (insert prems, auto)
   202     using a b apply auto
   203     apply (rule mult_pos_pos)
   203     apply (rule mult_pos_pos)
   204     apply (rule add_pos_nonneg)
   204     apply (rule add_pos_nonneg)
   205     apply auto
   205     apply auto
   206     done
   206     done
   207   also have "... = exp (-x)"
   207   also have "... = exp (-x)"
   254       by auto
   254       by auto
   255   qed
   255   qed
   256   also have "- (x / (1 - x)) = -x / (1 - x)"
   256   also have "- (x / (1 - x)) = -x / (1 - x)"
   257     by auto
   257     by auto
   258   finally have d: "- x / (1 - x) <= ln (1 - x)" .
   258   finally have d: "- x / (1 - x) <= ln (1 - x)" .
   259   have "0 < 1 - x" using prems by simp
   259   have "0 < 1 - x" using a b by simp
   260   hence e: "-x - 2 * x^2 <= - x / (1 - x)"
   260   hence e: "-x - 2 * x^2 <= - x / (1 - x)"
   261     using mult_right_le_one_le[of "x*x" "2*x"] prems
   261     using mult_right_le_one_le[of "x*x" "2*x"] a b
   262     by(simp add:field_simps power2_eq_square)
   262     by (simp add:field_simps power2_eq_square)
   263   from e d show "- x - 2 * x^2 <= ln (1 - x)"
   263   from e d show "- x - 2 * x^2 <= ln (1 - x)"
   264     by (rule order_trans)
   264     by (rule order_trans)
   265 qed
   265 qed
   266 
   266 
   267 lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x"
   267 lemma exp_ge_add_one_self [simp]: "1 + (x::real) <= exp x"
   290 
   290 
   291 lemma abs_ln_one_plus_x_minus_x_bound_nonneg:
   291 lemma abs_ln_one_plus_x_minus_x_bound_nonneg:
   292     "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2"
   292     "0 <= x ==> x <= 1 ==> abs(ln (1 + x) - x) <= x^2"
   293 proof -
   293 proof -
   294   assume x: "0 <= x"
   294   assume x: "0 <= x"
   295   assume "x <= 1"
   295   assume x1: "x <= 1"
   296   from x have "ln (1 + x) <= x"
   296   from x have "ln (1 + x) <= x"
   297     by (rule ln_add_one_self_le_self)
   297     by (rule ln_add_one_self_le_self)
   298   then have "ln (1 + x) - x <= 0" 
   298   then have "ln (1 + x) - x <= 0" 
   299     by simp
   299     by simp
   300   then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)"
   300   then have "abs(ln(1 + x) - x) = - (ln(1 + x) - x)"
   301     by (rule abs_of_nonpos)
   301     by (rule abs_of_nonpos)
   302   also have "... = x - ln (1 + x)" 
   302   also have "... = x - ln (1 + x)" 
   303     by simp
   303     by simp
   304   also have "... <= x^2"
   304   also have "... <= x^2"
   305   proof -
   305   proof -
   306     from prems have "x - x^2 <= ln (1 + x)"
   306     from x x1 have "x - x^2 <= ln (1 + x)"
   307       by (intro ln_one_plus_pos_lower_bound)
   307       by (intro ln_one_plus_pos_lower_bound)
   308     thus ?thesis
   308     thus ?thesis
   309       by simp
   309       by simp
   310   qed
   310   qed
   311   finally show ?thesis .
   311   finally show ?thesis .
   312 qed
   312 qed
   313 
   313 
   314 lemma abs_ln_one_plus_x_minus_x_bound_nonpos:
   314 lemma abs_ln_one_plus_x_minus_x_bound_nonpos:
   315     "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2"
   315     "-(1 / 2) <= x ==> x <= 0 ==> abs(ln (1 + x) - x) <= 2 * x^2"
   316 proof -
   316 proof -
   317   assume "-(1 / 2) <= x"
   317   assume a: "-(1 / 2) <= x"
   318   assume "x <= 0"
   318   assume b: "x <= 0"
   319   have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" 
   319   have "abs(ln (1 + x) - x) = x - ln(1 - (-x))" 
   320     apply (subst abs_of_nonpos)
   320     apply (subst abs_of_nonpos)
   321     apply simp
   321     apply simp
   322     apply (rule ln_add_one_self_le_self2)
   322     apply (rule ln_add_one_self_le_self2)
   323     apply (insert prems, auto)
   323     using a apply auto
   324     done
   324     done
   325   also have "... <= 2 * x^2"
   325   also have "... <= 2 * x^2"
   326     apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
   326     apply (subgoal_tac "- (-x) - 2 * (-x)^2 <= ln (1 - (-x))")
   327     apply (simp add: algebra_simps)
   327     apply (simp add: algebra_simps)
   328     apply (rule ln_one_minus_pos_lower_bound)
   328     apply (rule ln_one_minus_pos_lower_bound)
   329     apply (insert prems, auto)
   329     using a b apply auto
   330     done
   330     done
   331   finally show ?thesis .
   331   finally show ?thesis .
   332 qed
   332 qed
   333 
   333 
   334 lemma abs_ln_one_plus_x_minus_x_bound:
   334 lemma abs_ln_one_plus_x_minus_x_bound:
   341   apply auto
   341   apply auto
   342 done
   342 done
   343 
   343 
   344 lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"  
   344 lemma ln_x_over_x_mono: "exp 1 <= x ==> x <= y ==> (ln y / y) <= (ln x / x)"  
   345 proof -
   345 proof -
   346   assume "exp 1 <= x" and "x <= y"
   346   assume x: "exp 1 <= x" "x <= y"
   347   have a: "0 < x" and b: "0 < y"
   347   have a: "0 < x" and b: "0 < y"
   348     apply (insert prems)
   348     apply (insert x)
   349     apply (subgoal_tac "0 < exp (1::real)")
   349     apply (subgoal_tac "0 < exp (1::real)")
   350     apply arith
   350     apply arith
   351     apply auto
   351     apply auto
   352     apply (subgoal_tac "0 < exp (1::real)")
   352     apply (subgoal_tac "0 < exp (1::real)")
   353     apply arith
   353     apply arith
   359     apply (subst ln_div)
   359     apply (subst ln_div)
   360     apply (rule b, rule a, rule refl)
   360     apply (rule b, rule a, rule refl)
   361     done
   361     done
   362   also have "y / x = (x + (y - x)) / x"
   362   also have "y / x = (x + (y - x)) / x"
   363     by simp
   363     by simp
   364   also have "... = 1 + (y - x) / x" using a prems by(simp add:field_simps)
   364   also have "... = 1 + (y - x) / x" using x a by (simp add: field_simps)
   365   also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)"
   365   also have "x * ln(1 + (y - x) / x) <= x * ((y - x) / x)"
   366     apply (rule mult_left_mono)
   366     apply (rule mult_left_mono)
   367     apply (rule ln_add_one_self_le_self)
   367     apply (rule ln_add_one_self_le_self)
   368     apply (rule divide_nonneg_pos)
   368     apply (rule divide_nonneg_pos)
   369     apply (insert prems a, simp_all) 
   369     using x a apply simp_all
   370     done
   370     done
   371   also have "... = y - x" using a by simp
   371   also have "... = y - x" using a by simp
   372   also have "... = (y - x) * ln (exp 1)" by simp
   372   also have "... = (y - x) * ln (exp 1)" by simp
   373   also have "... <= (y - x) * ln x"
   373   also have "... <= (y - x) * ln x"
   374     apply (rule mult_left_mono)
   374     apply (rule mult_left_mono)
   375     apply (subst ln_le_cancel_iff)
   375     apply (subst ln_le_cancel_iff)
   376     apply force
   376     apply force
   377     apply (rule a)
   377     apply (rule a)
   378     apply (rule prems)
   378     apply (rule x)
   379     apply (insert prems, simp)
   379     using x apply simp
   380     done
   380     done
   381   also have "... = y * ln x - x * ln x"
   381   also have "... = y * ln x - x * ln x"
   382     by (rule left_diff_distrib)
   382     by (rule left_diff_distrib)
   383   finally have "x * ln y <= y * ln x"
   383   finally have "x * ln y <= y * ln x"
   384     by arith
   384     by arith
   385   then have "ln y <= (y * ln x) / x" using a by(simp add:field_simps)
   385   then have "ln y <= (y * ln x) / x" using a by (simp add: field_simps)
   386   also have "... = y * (ln x / x)"  by simp
   386   also have "... = y * (ln x / x)" by simp
   387   finally show ?thesis using b by(simp add:field_simps)
   387   finally show ?thesis using b by (simp add: field_simps)
   388 qed
   388 qed
   389 
   389 
   390 end
   390 end