X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fynat%2Fynat_plus.ma;h=2e1d05077a37e9fd25e5d328ba893097629c97e9;hb=d8d00d6f6694155be5be486a8239f5953efe28b7;hp=16a933acf7472d41f4d6269153c6333f919a48ef;hpb=5102e7f780e83c7fef1d3826f81dfd37ee4028bc;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 16a933acf..2e1d05077 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_plus.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "ground_2/xoa/ex_3_2.ma". include "ground_2/ynat/ynat_lt.ma". (* NATURAL NUMBERS WITH INFINITY ********************************************) @@ -27,7 +28,7 @@ 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). +lemma yplus_S2: ∀m:ynat. ∀n. m + S n = ↑(m + n). // qed. lemma yplus_Y2: ∀m:ynat. m + (∞) = ∞. @@ -35,18 +36,18 @@ lemma yplus_Y2: ∀m:ynat. m + (∞) = ∞. (* Properties on successor **************************************************) -lemma yplus_succ2: ∀m,n. m + ⫯n = ⫯(m + n). +lemma yplus_succ2: ∀m,n. m + ↑n = ↑(m + n). #m * // qed. -lemma yplus_succ1: ∀m,n. ⫯m + n = ⫯(m + n). +lemma yplus_succ1: ∀m,n. ↑m + n = ↑(m + n). #m * // #n elim n -n // qed. -lemma yplus_succ_swap: ∀m,n. m + ⫯n = ⫯m + n. +lemma yplus_succ_swap: ∀m,n. m + ↑n = ↑m + n. // qed. -lemma yplus_SO2: ∀m. m + 1 = ⫯m. +lemma yplus_SO2: ∀m. m + 1 = ↑m. * // qed. @@ -86,19 +87,19 @@ lemma yplus_comm_24: ∀x1,x2,x3,x4. x1 + x4 + x3 + x2 = x1 + x2 + x3 + x4. /2 width=1 by eq_f2/ qed. -lemma yplus_assoc_23: ∀x1,x2,x3,x4. x1 + x2 + (x3 + x4) = x1 + (x2 + x3) + x4. +lemma yplus_assoc_23: ∀x1,x2,x3,x4. x1 + x2 + (x3 + x4) = x1 + (x2 + x3) + x4. #x1 #x2 #x3 #x4 >yplus_assoc >yplus_assoc /2 width=1 by eq_f2/ qed. (* Inversion lemmas on successor *********************************************) -lemma yplus_inv_succ_lt_dx: ∀x,y,z:ynat. 0 < y → x + y = ⫯z → x + ⫰y = z. +lemma yplus_inv_succ_lt_dx: ∀x,y,z:ynat. 0 < y → x + y = ↑z → x + ↓y = z. #x #y #z #H <(ylt_inv_O1 y) // >yplus_succ2 /2 width=1 by ysucc_inv_inj/ qed-. -lemma yplus_inv_succ_lt_sn: ∀x,y,z:ynat. 0 < x → x + y = ⫯z → ⫰x + y = z. +lemma yplus_inv_succ_lt_sn: ∀x,y,z:ynat. 0 < x → x + y = ↑z → ↓x + y = z. #x #y #z #H <(ylt_inv_O1 x) // >yplus_succ1 /2 width=1 by ysucc_inv_inj/ @@ -148,15 +149,15 @@ lemma yplus_inv_monotonic_dx_inj: ∀z,x,y. x + yinj z = y + yinj z → x = y. qed-. lemma yplus_inv_monotonic_dx: ∀z,x,y. z < ∞ → x + z = y + z → x = y. -#z #x #y #H elim (ylt_inv_Y2 … H) -H /2 width=2 by yplus_inv_monotonic_dx_inj/ +#z #x #y #H elim (ylt_inv_Y2 … H) -H /2 width=2 by yplus_inv_monotonic_dx_inj/ qed-. lemma yplus_inv_Y2: ∀x,y. x + y = ∞ → x = ∞ ∨ y = ∞. -* /2 width=1 by or_introl/ #x * // #y >yplus_inj #H destruct +* /2 width=1 by or_introl/ #x * // #y >yplus_inj #H destruct qed-. lemma yplus_inv_monotonic_23: ∀z,x1,x2,y1,y2. z < ∞ → - x1 + z + y1 = x2 + z + y2 → x1 + y1 = x2 + y2. + x1 + z + y1 = x2 + z + y2 → x1 + y1 = x2 + y2. #z #x1 #x2 #y1 #y2 #Hz #H @(yplus_inv_monotonic_dx z) // >yplus_comm_23 >H -H // qed-. @@ -168,13 +169,13 @@ lemma ylt_inv_plus_Y: ∀x,y. x + y < ∞ → x < ∞ ∧ y < ∞. #z #H elim (yplus_inv_inj … H) -H /2 width=1 by conj/ qed-. -lemma ylt_inv_plus_sn: ∀x,y. x < y → ∃∃z. ⫯z + x = y & x < ∞. +lemma ylt_inv_plus_sn: ∀x,y. x < y → ∃∃z. ↑z + x = y & x < ∞. #x #y #H elim (ylt_inv_le … H) -H #Hx #H elim (yle_inv_plus_sn … H) -H /2 width=2 by ex2_intro/ qed-. -lemma ylt_inv_plus_dx: ∀x,y. x < y → ∃∃z. x + ⫯z = y & x < ∞. +lemma ylt_inv_plus_dx: ∀x,y. x < y → ∃∃z. x + ↑z = y & x < ∞. #x #y #H elim (ylt_inv_plus_sn … H) -H #z >yplus_comm /2 width=2 by ex2_intro/ qed-. @@ -269,7 +270,7 @@ lemma monotonic_ylt_plus_sn_inj: ∀x,y. x < y → ∀z:nat. yinj z + x < yinj z lemma monotonic_ylt_plus_dx: ∀x,y. x < y → ∀z. z < ∞ → x + z < y + z. #x #y #Hxy #z #Hz elim (ylt_inv_Y2 … Hz) -Hz #m #H destruct /2 width=1 by monotonic_ylt_plus_dx_inj/ -qed. +qed. lemma monotonic_ylt_plus_sn: ∀x,y. x < y → ∀z. z < ∞ → z + x < z + y. /2 width=1 by monotonic_ylt_plus_dx/ qed. @@ -302,11 +303,11 @@ qed-. (* Properties on predeccessor ***********************************************) -lemma yplus_pred1: ∀x,y:ynat. 0 < x → ⫰x + y = ⫰(x+y). +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). +lemma yplus_pred2: ∀x,y:ynat. 0 < y → x + ↓y = ↓(x+y). /2 width=1 by yplus_pred1/ qed-.