author | huffman |
Tue, 21 Aug 2007 02:30:14 +0200 | |
changeset 24376 | e403ab5c9415 |
parent 24375 | 4aa80fadc071 |
child 24377 | 223622422d7b |
src/HOL/Power.thy | file | annotate | diff | comparison | revisions |
--- a/src/HOL/Power.thy Tue Aug 21 02:15:13 2007 +0200 +++ b/src/HOL/Power.thy Tue Aug 21 02:30:14 2007 +0200 @@ -77,6 +77,10 @@ finally show ?thesis by simp qed +lemma one_less_power: + "\<lbrakk>1 < (a::'a::{ordered_semidom,recpower}); 0 < n\<rbrakk> \<Longrightarrow> 1 < a ^ n" +by (cases n, simp_all add: power_gt1_lemma power_Suc) + lemma power_gt1: "1 < (a::'a::{ordered_semidom,recpower}) ==> 1 < a ^ (Suc n)" by (simp add: power_gt1_lemma power_Suc)