X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Flibrary%2Fnat%2Fexp.ma;h=c69be6b5fda43ece413ac03cbfab7103a08aa0c1;hb=dd6568cd572eb236f2c3803ce18914731423eae5;hp=11d84f74ca7deb3b676007693f72b585c0547435;hpb=7f2444c2670cadafddd8785b687ef312158376b0;p=helm.git diff --git a/matita/library/nat/exp.ma b/matita/library/nat/exp.ma index 11d84f74c..c69be6b5f 100644 --- a/matita/library/nat/exp.ma +++ b/matita/library/nat/exp.ma @@ -15,6 +15,7 @@ set "baseuri" "cic:/matita/nat/exp". include "nat/div_and_mod.ma". +include "nat/lt_arith.ma". let rec exp n m on m\def match m with @@ -54,7 +55,7 @@ apply le_times.assumption.assumption. qed. theorem lt_m_exp_nm: \forall n,m:nat. (S O) < n \to m < n \sup m. -intros.elim m.simplify.unfold lt.reflexivity. +intros.elim m.simplify.unfold lt.apply le_n. simplify.unfold lt. apply (trans_le ? ((S(S O))*(S n1))). simplify. @@ -95,3 +96,51 @@ qed. variant inj_exp_r: \forall p:nat. (S O) < p \to \forall n,m:nat. p \sup n = p \sup m \to n = m \def injective_exp_r. + +theorem le_exp: \forall n,m,p:nat. O < p \to n \le m \to exp p n \le exp p m. +apply nat_elim2 + [intros. + apply lt_O_exp.assumption + |intros. + apply False_ind. + apply (le_to_not_lt ? ? ? H1). + apply le_O_n + |intros. + simplify. + apply le_times + [apply le_n + |apply H[assumption|apply le_S_S_to_le.assumption] + ] + ] +qed. + +theorem lt_exp: \forall n,m,p:nat. S O < p \to n < m \to exp p n < exp p m. +apply nat_elim2 + [intros. + apply (lt_O_n_elim ? H1).intro. + simplify.unfold lt. + rewrite > times_n_SO. + apply le_times + [assumption + |apply lt_O_exp. + apply (trans_lt ? (S O))[apply le_n|assumption] + ] + |intros. + apply False_ind. + apply (le_to_not_lt ? ? ? H1). + apply le_O_n + |intros.simplify. + apply lt_times_r1 + [apply (trans_lt ? (S O))[apply le_n|assumption] + |apply H + [apply H1 + |apply le_S_S_to_le.assumption + ] + ] + ] +qed. + + + + + \ No newline at end of file