X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fmultiple%2Fmr2_plus.ma;h=455b669bef01513895a68259af88405cec4e98b4;hb=658c000ee2ea2da04cf29efc0acdaf16364fbf5e;hp=99b2d4b17474eb2504a2e85ae79de895855a4b71;hpb=1994fe8e6355243652770f53a02db5fdf26915f0;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/mr2_plus.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/mr2_plus.ma index 99b2d4b17..455b669be 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/mr2_plus.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/mr2_plus.ma @@ -12,11 +12,12 @@ (* *) (**************************************************************************) +include "ground_2/ynat/ynat_plus.ma". include "basic_2/multiple/mr2.ma". (* MULTIPLE RELOCATION WITH PAIRS *******************************************) -let rec pluss (cs:list2 nat nat) (i:nat) on cs ≝ match cs with +let rec pluss (cs:list2 ynat nat) (i:nat) on cs ≝ match cs with [ nil2 ⇒ ◊ | cons2 l m cs ⇒ {l + i, m} @ pluss cs i ]. @@ -24,6 +25,11 @@ let rec pluss (cs:list2 nat nat) (i:nat) on cs ≝ match cs with interpretation "plus (multiple relocation with pairs)" 'plus x y = (pluss x y). +(* Basic properties *********************************************************) + +lemma pluss_SO2: ∀l,m,cs. ({l, m} @ cs) + 1 = {⫯l, m} @ cs + 1. +// qed. + (* Basic inversion lemmas ***************************************************) lemma pluss_inv_nil2: ∀i,cs. cs + i = ◊ → cs = ◊. @@ -33,8 +39,9 @@ qed. lemma pluss_inv_cons2: ∀i,l,m,cs2,cs. cs + i = {l, m} @ cs2 → ∃∃cs1. cs1 + i = cs2 & cs = {l - i, m} @ cs1. -#i #l #m #cs2 * normalize -[ #H destruct -| #l1 #m1 #cs1 #H destruct /2 width=3 by ex2_intro/ +#i #l #m #cs2 * +[ normalize #H destruct +| #l1 #m1 #cs1 whd in ⊢ (??%?→?); #H destruct + >yplus_minus_inj /2 width=3 by ex2_intro/ ] qed-.