X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=79bcc2548bca92b905b524c6adfda5ae919040cc;hb=d8d00d6f6694155be5be486a8239f5953efe28b7;hp=6ce878b44bf0d73d1ba13dfd8e23a19cf008441e;hpb=bac74b5cff042d37e1abc9c961a6c41094b8a294;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma index 6ce878b44..79bcc2548 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -12,9 +12,12 @@ (* *) (**************************************************************************) +include "arithmetics/nat.ma". +include "ground_2/xoa/ex_3_1.ma". +include "ground_2/xoa/or_3.ma". include "ground_2/notation/functions/uparrow_1.ma". include "ground_2/notation/functions/downarrow_1.ma". -include "arithmetics/nat.ma". +include "ground_2/pull/pull_2.ma". include "ground_2/lib/relations.ma". (* ARITHMETICAL PROPERTIES **************************************************) @@ -51,7 +54,10 @@ qed. (* Equalities ***************************************************************) -lemma plus_SO: ∀n. n + 1 = ↑n. +lemma plus_SO_sn (n): 1 + n = ↑n. +// qed-. + +lemma plus_SO_dx (n): n + 1 = ↑n. // qed. lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m. @@ -150,6 +156,14 @@ lemma lt_S_S: ∀x,y. x < y → ↑x < ↑y. lemma lt_S: ∀n,m. n < m → n < ↑m. /2 width=1 by le_S/ qed. +lemma monotonic_lt_minus_r: +∀p,q,n. q < n -> q < p → n-p < n-q. +#p #q #n #Hn #H +lapply (monotonic_le_minus_r … n H) -H #H +@(le_to_lt_to_lt … H) -H +/2 width=1 by lt_plus_to_minus/ +qed. + lemma max_S1_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (↑n1 ∨ n2) ≤ ↑n. /4 width=2 by to_max, le_maxr, le_S_S, le_S/ qed-. @@ -291,6 +305,21 @@ lemma le_elim: ∀R:relation nat. #n1 #H elim (lt_le_false … H) -H // qed-. +lemma nat_elim_le_sn (Q:relation …): + (∀m1,m2. (∀m. m < m2-m1 → Q (m2-m) m2) → m1 ≤ m2 → Q m1 m2) → + ∀n1,n2. n1 ≤ n2 → Q n1 n2. +#Q #IH #n1 #n2 #Hn +<(minus_minus_m_m … Hn) -Hn +lapply (minus_le n2 n1) +let d ≝ (n2-n1) +@(nat_elim1 … d) -d -n1 #d +@pull_2 #Hd +<(minus_minus_m_m … Hd) in ⊢ (%→?); -Hd +let n1 ≝ (n2-d) #IHd +@IH -IH [| // ] #m #Hn +/4 width=3 by lt_to_le, lt_to_le_to_lt/ +qed-. + (* Iterators ****************************************************************) (* Note: see also: lib/arithemetics/bigops.ma *)