X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Fmr2_plus.ma;h=e1ef5d2875cfd78e4eadc2cfd50b1e881d2483a2;hp=7d9e5e60332d3628c044a31ff08ae7e8e489fdb2;hb=3c8da07d7a5d7cf0432a83732a6d103f527afaef;hpb=2976c347e18717e691825ebdf73a5ce941c57d1b diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_plus.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_plus.ma index 7d9e5e603..e1ef5d287 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_plus.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/mr2_plus.ma @@ -18,7 +18,7 @@ include "ground_2/relocation/mr2.ma". rec definition pluss (cs:mr2) (i:nat) on cs ≝ match cs with [ nil2 ⇒ ◊ -| cons2 l m cs ⇒ {l + i, m} ⨮ pluss cs i +| cons2 l m cs ⇒ {l + i, m};pluss cs i ]. interpretation "plus (multiple relocation with pairs)" @@ -26,7 +26,7 @@ interpretation "plus (multiple relocation with pairs)" (* Basic properties *********************************************************) -lemma pluss_SO2: ∀l,m,cs. ({l, m} ⨮ cs) + 1 = {↑l, m} ⨮ cs + 1. +lemma pluss_SO2: ∀l,m,cs. ({l, m};cs) + 1 = {↑l, m};cs + 1. normalize // qed. (* Basic inversion lemmas ***************************************************) @@ -36,8 +36,8 @@ lemma pluss_inv_nil2: ∀i,cs. cs + i = ◊ → cs = ◊. #l #m #cs #H destruct qed. -lemma pluss_inv_cons2: ∀i,l,m,cs2,cs. cs + i = {l, m} ⨮ cs2 → - ∃∃cs1. cs1 + i = cs2 & cs = {l - i, m} ⨮ cs1. +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 whd in ⊢ (??%?→?); #H destruct