author | hoelzl |
Wed, 18 Apr 2012 14:29:18 +0200 | |
changeset 47595 | 836b4c4d7c86 |
parent 47594 | be2ac449488c |
child 47596 | c031e65c8ddc |
src/HOL/Log.thy | file | annotate | diff | comparison | revisions |
--- a/src/HOL/Log.thy Wed Apr 18 14:29:17 2012 +0200 +++ b/src/HOL/Log.thy Wed Apr 18 14:29:18 2012 +0200 @@ -285,6 +285,10 @@ apply (rule powr_less_mono2, auto) done +lemma powr_inj: + "0 < a \<Longrightarrow> a \<noteq> 1 \<Longrightarrow> a powr x = a powr y \<longleftrightarrow> x = y" + unfolding powr_def exp_inj_iff by simp + lemma ln_powr_bound: "1 <= x ==> 0 < a ==> ln x <= (x powr a) / a" apply (rule mult_imp_le_div_pos) apply (assumption)