X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Flt_arith.ma;h=b8339f374e6019dd054719c9928bf735c63aef4b;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=5bef2237103f27308663712a78f78b67d0c3b621;hpb=7bbce6bc163892cfd99cfcda65db42001b86789f;p=helm.git diff --git a/helm/matita/library/nat/lt_arith.ma b/helm/matita/library/nat/lt_arith.ma index 5bef22371..b8339f374 100644 --- a/helm/matita/library/nat/lt_arith.ma +++ b/helm/matita/library/nat/lt_arith.ma @@ -18,145 +18,200 @@ include "nat/div_and_mod.ma". (* plus *) theorem monotonic_lt_plus_r: -\forall n:nat.monotonic nat lt (\lambda m.plus n m). +\forall n:nat.monotonic nat lt (\lambda m.n+m). simplify.intros. elim n.simplify.assumption. -simplify. +simplify.unfold lt. apply le_S_S.assumption. qed. -variant lt_plus_r: \forall n,p,q:nat. lt p q \to lt (plus n p) (plus n q) \def +variant lt_plus_r: \forall n,p,q:nat. p < q \to n + p < n + q \def monotonic_lt_plus_r. theorem monotonic_lt_plus_l: -\forall n:nat.monotonic nat lt (\lambda m.plus m n). -change with \forall n,p,q:nat. lt p q \to lt (plus p n) (plus q n). +\forall n:nat.monotonic nat lt (\lambda m.m+n). +change with (\forall n,p,q:nat. p < q \to p + n < q + n). intros. -rewrite < sym_plus. rewrite < sym_plus n. +rewrite < sym_plus. rewrite < (sym_plus n). apply lt_plus_r.assumption. qed. -variant lt_plus_l: \forall n,p,q:nat. lt p q \to lt (plus p n) (plus q n) \def +variant lt_plus_l: \forall n,p,q:nat. p < q \to p + n < q + n \def monotonic_lt_plus_l. -theorem lt_plus: \forall n,m,p,q:nat. lt n m \to lt p q \to lt (plus n p) (plus m q). +theorem lt_plus: \forall n,m,p,q:nat. n < m \to p < q \to n + p < m + q. intros. -apply trans_lt ? (plus n q). +apply (trans_lt ? (n + q)). apply lt_plus_r.assumption. apply lt_plus_l.assumption. qed. -theorem lt_plus_to_lt_l :\forall n,p,q:nat. lt (plus p n) (plus q n) \to lt p q. +theorem lt_plus_to_lt_l :\forall n,p,q:nat. p+n < q+n \to p plus_n_O. -rewrite > plus_n_O q.assumption. +rewrite > (plus_n_O q).assumption. apply H. -simplify.apply le_S_S_to_le. +unfold lt.apply le_S_S_to_le. rewrite > plus_n_Sm. -rewrite > plus_n_Sm q. +rewrite > (plus_n_Sm q). exact H1. qed. -theorem lt_plus_to_lt_r :\forall n,p,q:nat. lt (plus n p) (plus n q) \to lt p q. -intros.apply lt_plus_to_lt_l n. +theorem lt_plus_to_lt_r :\forall n,p,q:nat. n+p < n+q \to p sym_plus. -rewrite > sym_plus q.assumption. +rewrite > (sym_plus q).assumption. qed. (* times and zero *) -theorem lt_O_times_S_S: \forall n,m:nat.lt O (times (S n) (S m)). -intros.simplify.apply le_S_S.apply le_O_n. +theorem lt_O_times_S_S: \forall n,m:nat.O < (S n)*(S m). +intros.simplify.unfold lt.apply le_S_S.apply le_O_n. qed. (* times *) theorem monotonic_lt_times_r: -\forall n:nat.monotonic nat lt (\lambda m.times (S n) m). -change with \forall n,p,q:nat. lt p q \to lt (times (S n) p) (times (S n) q). +\forall n:nat.monotonic nat lt (\lambda m.(S n)*m). +change with (\forall n,p,q:nat. p < q \to (S n) * p < (S n) * q). intros.elim n. simplify.rewrite < plus_n_O.rewrite < plus_n_O.assumption. -change with lt (plus p (times (S n1) p)) (plus q (times (S n1) q)). +change with (p + (S n1) * p < q + (S n1) * q). apply lt_plus.assumption.assumption. qed. -theorem lt_times_r: \forall n,p,q:nat. lt p q \to lt (times (S n) p) (times (S n) q) +theorem lt_times_r: \forall n,p,q:nat. p < q \to (S n) * p < (S n) * q \def monotonic_lt_times_r. theorem monotonic_lt_times_l: -\forall m:nat.monotonic nat lt (\lambda n.times n (S m)). +\forall m:nat.monotonic nat lt (\lambda n.n * (S m)). change with -\forall n,p,q:nat. lt p q \to lt (times p (S n)) (times q (S n)). +(\forall n,p,q:nat. p < q \to p*(S n) < q*(S n)). intros. -rewrite < sym_times.rewrite < sym_times (S n). +rewrite < sym_times.rewrite < (sym_times (S n)). apply lt_times_r.assumption. qed. -variant lt_times_l: \forall n,p,q:nat. lt p q \to lt (times p (S n)) (times q (S n)) +variant lt_times_l: \forall n,p,q:nat. p nat_compare_n_n.reflexivity. intro.apply nat_compare_elim.intro. -absurd (lt p q). -apply lt_times_to_lt_r n.assumption. +absurd (p (plus_n_O ((S m1)*(n / (S m1)))). +rewrite < H2. +rewrite < sym_times. +rewrite < div_mod. +rewrite > H2. +assumption. +unfold lt.apply le_S_S.apply le_O_n. +qed. + +theorem lt_div_n_m_n: \forall n,m:nat. (S O) < m \to O < n \to n / m \lt n. +intros. +apply (nat_case1 (n / m)).intro. +assumption.intros.rewrite < H2. +rewrite > (div_mod n m) in \vdash (? ? %). +apply (lt_to_le_to_lt ? ((n / m)*m)). +apply (lt_to_le_to_lt ? ((n / m)*(S (S O)))). +rewrite < sym_times. +rewrite > H2. +simplify.unfold lt. +rewrite < plus_n_O. +rewrite < plus_n_Sm. +apply le_S_S. +apply le_S_S. +apply le_plus_n. +apply le_times_r. +assumption. +rewrite < sym_plus. +apply le_plus_n. +apply (trans_lt ? (S O)). +unfold lt. apply le_n.assumption. +qed. + +(* general properties of functions *) +theorem monotonic_to_injective: \forall f:nat\to nat. +monotonic nat lt f \to injective nat nat f. +unfold injective.intros. +apply (nat_compare_elim x y). +intro.apply False_ind.apply (not_le_Sn_n (f x)). +rewrite > H1 in \vdash (? ? %).apply H.apply H2. +intros.assumption. +intro.apply False_ind.apply (not_le_Sn_n (f y)). +rewrite < H1 in \vdash (? ? %).apply H.apply H2. +qed. + +theorem increasing_to_injective: \forall f:nat\to nat. +increasing f \to injective nat nat f. +intros.apply monotonic_to_injective. +apply increasing_to_monotonic.assumption. +qed.