X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fynat%2Fynat_plus.ma;h=f7a15c310424605c3f37d7a9a8c86013c590e6b7;hb=658c000ee2ea2da04cf29efc0acdaf16364fbf5e;hp=e8390c8dd2526da1907854a5339e5fb09c9acf88;hpb=1994fe8e6355243652770f53a02db5fdf26915f0;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma index e8390c8dd..f7a15c310 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma @@ -24,6 +24,15 @@ definition yplus: ynat → ynat → ynat ≝ λx,y. match y with interpretation "ynat plus" 'plus x y = (yplus x y). +lemma yplus_O2: ∀m:ynat. m + 0 = m. +// qed. + +lemma yplus_S2: ∀m:ynat. ∀n. m + S n = ⫯(m + n). +// qed. + +lemma yplus_Y2: ∀m:ynat. m + (∞) = ∞. +// qed. + (* Properties on successor **************************************************) lemma yplus_succ2: ∀m,n. m + ⫯n = ⫯(m + n). @@ -31,7 +40,7 @@ lemma yplus_succ2: ∀m,n. m + ⫯n = ⫯(m + n). qed. lemma yplus_succ1: ∀m,n. ⫯m + n = ⫯(m + n). -#m * normalize // +#m * // #n elim n -n // qed. lemma yplus_succ_swap: ∀m,n. m + ⫯n = ⫯m + n. @@ -44,18 +53,17 @@ qed. (* Basic properties *********************************************************) lemma yplus_inj: ∀n,m. yinj m + yinj n = yinj (m + n). -#n elim n -n [ normalize // ] +#n elim n -n // #n #IHn #m >(yplus_succ2 ? n) >IHn -IHn ysucc_iter_Y // qed. lemma yplus_assoc: associative … yplus. @@ -68,6 +76,10 @@ qed. lemma yplus_O1: ∀n:ynat. 0 + n = n. #n >yplus_comm // qed. +lemma yplus_comm_23: ∀x,y,z. x + z + y = x + y + z. +#x #y #z >yplus_assoc // +qed. + (* Basic inversion lemmas ***************************************************) lemma yplus_inv_inj: ∀z,y,x. x + y = yinj z → @@ -104,7 +116,7 @@ lemma monotonic_yle_plus_sn: ∀x,y. x ≤ y → ∀z. z + x ≤ z + y. lemma monotonic_yle_plus: ∀x1,y1. x1 ≤ y1 → ∀x2,y2. x2 ≤ y2 → x1 + x2 ≤ y1 + y2. -/3 width=3 by monotonic_yle_plus_dx, yle_trans/ qed. +/3 width=3 by monotonic_yle_plus_dx, monotonic_yle_plus_sn, yle_trans/ qed. (* Forward lemmas on order **************************************************) @@ -132,6 +144,10 @@ lemma yle_fwd_plus_ge_inj: ∀m1:nat. ∀m2,n1,n2:ynat. m2 ≤ m1 → m1 + n1 #x #H0 #H destruct /3 width=4 by yle_fwd_plus_ge, yle_inj/ qed-. +lemma yle_fwd_plus_yge: ∀n2,m1:ynat. ∀n1,m2:nat. m2 ≤ m1 → m1 + n1 ≤ m2 + n2 → n1 ≤ n2. +* // #n2 * /2 width=4 by yle_fwd_plus_ge_inj/ +qed-. + (* Forward lemmas on strict order *******************************************) lemma ylt_inv_monotonic_plus_dx: ∀x,y,z. x + z < y + z → x < y. @@ -155,6 +171,17 @@ qed. lemma monotonic_ylt_plus_sn: ∀x,y. x < y → ∀z:nat. yinj z + x < yinj z + y. /2 width=1 by monotonic_ylt_plus_dx/ qed. +(* Properties on predeccessor ***********************************************) + +lemma yplus_pred1: ∀x,y:ynat. 0 < x → ⫰x + y = ⫰(x+y). +#x * // #y elim y -y // #y #IH #Hx +>yplus_S2 >yplus_S2 >IH -IH // >ylt_inv_O1 +/2 width=1 by ylt_plus_dx1_trans/ +qed-. + +lemma yplus_pred2: ∀x,y:ynat. 0 < y → x + ⫰y = ⫰(x+y). +/2 width=1 by yplus_pred1/ qed-. + (* Properties on minus ******************************************************) lemma yplus_minus_inj: ∀m:ynat. ∀n:nat. m + n - n = m. @@ -166,10 +193,18 @@ lemma yplus_minus: ∀m,n. m + n - n ≤ m. #m * // qed. +lemma yminus_plus2: ∀z,y,x:ynat. x - (y + z) = x - y - z. +* // #z * [2: >yplus_Y1 >yminus_O1 // ] #y * +[ #x >yplus_inj >yminus_inj >yminus_inj >yminus_inj /2 width=1 by eq_f/ +| >yplus_inj >yminus_Y_inj // +] +qed. + (* Forward lemmas on minus **************************************************) lemma yle_plus1_to_minus_inj2: ∀x,z:ynat. ∀y:nat. x + y ≤ z → x ≤ z - y. -/2 width=1 by monotonic_yle_minus_dx/ qed-. +#x #z #y #H lapply (monotonic_yle_minus_dx … H y) -H // +qed-. lemma yle_plus1_to_minus_inj1: ∀x,z:ynat. ∀y:nat. y + x ≤ z → x ≤ z - y. /2 width=1 by yle_plus1_to_minus_inj2/ qed-. @@ -188,12 +223,36 @@ lemma yplus_minus_assoc_inj: ∀x:nat. ∀y,z:ynat. x ≤ y → z + (y - x) = z ] qed-. +lemma yplus_minus_assoc_comm_inj: ∀x:nat. ∀y,z:ynat. x ≤ y → z - (y - x) = z + x - y. +#x * +[ #y * + [ #z >yminus_inj >yminus_inj >yplus_inj >yminus_inj + /4 width=1 by yle_inv_inj, minus_le_minus_minus_comm, eq_f/ + | >yminus_inj >yminus_Y_inj // + ] +| >yminus_Y_inj // +] +qed-. + lemma yplus_minus_comm_inj: ∀y:nat. ∀x,z:ynat. y ≤ x → x + z - y = x - y + z. #y * // #x * // #z #Hxy >yplus_inj >yminus_inj