X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fnat%2Fexp.ma;h=49efe8525504d4f431e81de36ea702f7c9d46f76;hb=d6c41ecd6b4d6944becbef2f7f6a625c51d8867c;hp=c69be6b5fda43ece413ac03cbfab7103a08aa0c1;hpb=6423f1b6e3056883016598e454c55cab1004dfd2;p=helm.git diff --git a/helm/software/matita/library/nat/exp.ma b/helm/software/matita/library/nat/exp.ma index c69be6b5f..49efe8525 100644 --- a/helm/software/matita/library/nat/exp.ma +++ b/helm/software/matita/library/nat/exp.ma @@ -40,6 +40,12 @@ theorem exp_n_SO : \forall n:nat. n = n \sup (S O). intro.simplify.rewrite < times_n_SO.reflexivity. qed. +theorem exp_SSO: \forall n. exp n (S(S O)) = n*n. +intro.simplify. +rewrite < times_n_SO. +reflexivity. +qed. + theorem exp_exp_times : \forall n,p,q:nat. (n \sup p) \sup q = n \sup (p * q). intros. @@ -113,7 +119,7 @@ apply nat_elim2 ] ] 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. @@ -140,6 +146,46 @@ apply nat_elim2 ] qed. +theorem le_exp_to_le: +\forall a,n,m. S O < a \to exp a n \le exp a m \to n \le m. +intro. +apply nat_elim2;intros + [apply le_O_n + |apply False_ind. + apply (le_to_not_lt ? ? H1). + simplify. + rewrite > times_n_SO. + apply lt_to_le_to_lt_times + [assumption + |apply lt_O_exp.apply lt_to_le.assumption + |apply lt_O_exp.apply lt_to_le.assumption + ] + |simplify in H2. + apply le_S_S. + apply H + [assumption + |apply (le_times_to_le a) + [apply lt_to_le.assumption|assumption] + ] + ] +qed. + +theorem lt_exp_to_lt: +\forall a,n,m. S O < a \to exp a n < exp a m \to n < m. +intros. +elim (le_to_or_lt_eq n m) + [assumption + |apply False_ind. + apply (lt_to_not_eq ? ? H1). + rewrite < H2. + reflexivity + |apply (le_exp_to_le a) + [assumption + |apply lt_to_le. + assumption + ] + ] +qed.