X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fmatita%2Flibrary%2Fnat%2Fle_arith.ma;h=a76183063b1e6d211d5c03a75cd1845c844c73aa;hb=4167cea65ca58897d1a3dbb81ff95de5074700cc;hp=a36f96bd451caba6a4091f0e7ecd9ea530957ed1;hpb=7bbce6bc163892cfd99cfcda65db42001b86789f;p=helm.git diff --git a/helm/matita/library/nat/le_arith.ma b/helm/matita/library/nat/le_arith.ma index a36f96bd4..a76183063 100644 --- a/helm/matita/library/nat/le_arith.ma +++ b/helm/matita/library/nat/le_arith.ma @@ -14,47 +14,44 @@ set "baseuri" "cic:/matita/nat/le_arith". -include "higher_order_defs/functions.ma". include "nat/times.ma". include "nat/orders.ma". -include "nat/compare.ma". (* plus *) theorem monotonic_le_plus_r: -\forall n:nat.monotonic nat le (\lambda m.plus n m). +\forall n:nat.monotonic nat le (\lambda m.n + m). simplify.intros.elim n. simplify.assumption. simplify.apply le_S_S.assumption. qed. -theorem le_plus_r: \forall p,n,m:nat. le n m \to le (plus p n) (plus p m) +theorem le_plus_r: \forall p,n,m:nat. n \le m \to p + n \le p + m \def monotonic_le_plus_r. theorem monotonic_le_plus_l: -\forall m:nat.monotonic nat le (\lambda n.plus n m). +\forall m:nat.monotonic nat le (\lambda n.n + m). simplify.intros. -rewrite < sym_plus.rewrite < sym_plus m. +rewrite < sym_plus.rewrite < (sym_plus m). apply le_plus_r.assumption. qed. -theorem le_plus_l: \forall p,n,m:nat. le n m \to le (plus n p) (plus m p) +theorem le_plus_l: \forall p,n,m:nat. n \le m \to n + p \le m + p \def monotonic_le_plus_l. -theorem le_plus: \forall n1,n2,m1,m2:nat. le n1 n2 \to le m1 m2 -\to le (plus n1 m1) (plus n2 m2). +theorem le_plus: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2 +\to n1 + m1 \le n2 + m2. intros. -apply trans_le ? (plus n2 m1). +apply (trans_le ? (n2 + m1)). apply le_plus_l.assumption. apply le_plus_r.assumption. qed. -theorem le_plus_n :\forall n,m:nat. le m (plus n m). -intros.change with le (plus O m) (plus n m). +theorem le_plus_n :\forall n,m:nat. m \le n + m. +intros.change with (O+m \le n+m). apply le_plus_l.apply le_O_n. qed. -theorem eq_plus_to_le: \forall n,m,p:nat.eq nat n (plus m p) -\to le m n. +theorem eq_plus_to_le: \forall n,m,p:nat.n=m+p \to m \le n. intros.rewrite > H. rewrite < sym_plus. apply le_plus_n. @@ -62,7 +59,7 @@ qed. (* times *) theorem monotonic_le_times_r: -\forall n:nat.monotonic nat le (\lambda m.times n m). +\forall n:nat.monotonic nat le (\lambda m. n * m). simplify.intros.elim n. simplify.apply le_O_n. simplify.apply le_plus. @@ -70,30 +67,29 @@ assumption. assumption. qed. -theorem le_times_r: \forall p,n,m:nat. le n m \to le (times p n) (times p m) +theorem le_times_r: \forall p,n,m:nat. n \le m \to p*n \le p*m \def monotonic_le_times_r. theorem monotonic_le_times_l: -\forall m:nat.monotonic nat le (\lambda n.times n m). +\forall m:nat.monotonic nat le (\lambda n.n*m). simplify.intros. -rewrite < sym_times.rewrite < sym_times m. +rewrite < sym_times.rewrite < (sym_times m). apply le_times_r.assumption. qed. -theorem le_times_l: \forall p,n,m:nat. le n m \to le (times n p) (times m p) +theorem le_times_l: \forall p,n,m:nat. n \le m \to n*p \le m*p \def monotonic_le_times_l. -theorem le_times: \forall n1,n2,m1,m2:nat. le n1 n2 \to le m1 m2 -\to le (times n1 m1) (times n2 m2). +theorem le_times: \forall n1,n2,m1,m2:nat. n1 \le n2 \to m1 \le m2 +\to n1*m1 \le n2*m2. intros. -apply trans_le ? (times n2 m1). +apply (trans_le ? (n2*m1)). apply le_times_l.assumption. apply le_times_r.assumption. qed. -theorem le_times_n: \forall n,m:nat.le (S O) n \to le m (times n m). +theorem le_times_n: \forall n,m:nat.(S O) \le n \to m \le n*m. intros.elim H.simplify. elim (plus_n_O ?).apply le_n. simplify.rewrite < sym_plus.apply le_plus_n. qed. -