X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Ftrace_isid.ma;h=d87487c897fa098b1b59e56c107ce1c2fda70c75;hb=952ec5aa2e9a54787acb63a5c8d6fdbf9011ab60;hp=0aa9a4c038364b05edbfc46d3a204f381d58ceb5;hpb=a5548736278a0b63f6f25c2721934ed8a7d2eef8;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma index 0aa9a4c03..d87487c89 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma @@ -58,24 +58,43 @@ qed-. (* Properties on composition ************************************************) -lemma isid_after_sn: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs2 → 𝐈⦃cs1⦄ . +lemma isid_after_sn: ∀cs2. ∃∃cs1. 𝐈⦃cs1⦄ & cs1 ⊚ cs2 ≡ cs2. +#cs2 elim cs2 -cs2 /2 width=3 by after_empty, ex2_intro/ +#b #cs2 * /3 width=3 by isid_true, after_true, ex2_intro/ +qed-. + +lemma isid_after_dx: ∀cs1. ∃∃cs2. 𝐈⦃cs2⦄ & cs1 ⊚ cs2 ≡ cs1. +#cs1 elim cs1 -cs1 /2 width=3 by after_empty, ex2_intro/ +* #cs1 * /3 width=3 by isid_true, after_true, after_false, ex2_intro/ +qed-. + +lemma after_isid_sn: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs2 → 𝐈⦃cs1⦄ . #cs1 #cs2 #H elim (after_inv_length … H) -H // qed. -lemma isid_after_dx: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs1 → 𝐈⦃cs2⦄ . +lemma after_isid_dx: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs1 → 𝐈⦃cs2⦄ . #cs1 #cs2 #H elim (after_inv_length … H) -H // qed. (* Inversion lemmas on composition ******************************************) -lemma isid_inv_after_sn: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs1⦄ → cs = cs2. +lemma after_isid_inv_sn: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs1⦄ → cs = cs2. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs // #cs1 #cs2 #cs #_ [ #b ] #IH #H [ >IH -IH // | elim (isid_inv_false … H) ] qed-. -lemma isid_inv_after_dx: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs2⦄ → cs = cs1. +lemma after_isid_inv_dx: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs2⦄ → cs = cs1. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs // #cs1 #cs2 #cs #_ [ #b ] #IH #H [ elim (isid_inv_cons … H) -H #H >IH -IH // | >IH -IH // ] qed-. + +lemma after_inv_isid3: ∀t1,t2,t. t1 ⊚ t2 ≡ t → 𝐈⦃t⦄ → 𝐈⦃t1⦄ ∧ 𝐈⦃t2⦄. +#t1 #t2 #t #H elim H -t1 -t2 -t +[ /2 width=1 by conj/ +| #t1 #t2 #t #_ #b #IHt #H elim (isid_inv_cons … H) -H + #Ht #H elim (IHt Ht) -t /2 width=1 by isid_true, conj/ +| #t1 #t2 #t #_ #_ #H elim (isid_inv_false … H) +] +qed-.