X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Ftrace_sor.ma;h=65aba02ac4da17fa88f8b24b11aa8e701c892653;hb=fb5c93c9812ea39fb78f1470da2095c80822e158;hp=f4e79a6ce25f44ea770e2c07784f33b91c684aeb;hpb=1e94683f160df35b31f1eee8f4d99a6ec8008b36;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_sor.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_sor.ma index f4e79a6ce..65aba02ac 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_sor.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_sor.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "ground_2/notation/relations/runion_3.ma". -include "ground_2/relocation/trace.ma". +include "ground_2/relocation/trace_isid.ma". (* RELOCATION TRACE *********************************************************) @@ -28,6 +28,35 @@ interpretation (* Basic properties *********************************************************) +lemma sor_length: ∀cs1,cs2. |cs1| = |cs2| → + ∃∃cs. cs2 ⋓ cs1 ≡ cs & |cs| = |cs1| & |cs| = |cs2|. +#cs1 elim cs1 -cs1 +[ #cs2 #H >(length_inv_zero_sn … H) -H /2 width=4 by sor_empty, ex3_intro/ +| #b1 #cs1 #IH #x #H elim (length_inv_succ_sn … H) -H + #cs2 #b2 #H12 #H destruct elim (IH … H12) -IH -H12 + #cs #H12 #H1 #H2 @(ex3_intro … ((b1 ∨ b2) @ cs)) /2 width=1 by sor_inh/ (**) (* explicit constructor *) +] +qed-. + lemma sor_sym: ∀cs1,cs2,cs. cs1 ⋓ cs2 ≡ cs → cs2 ⋓ cs1 ≡ cs. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs /2 width=1 by sor_inh/ qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma sor_inv_length: ∀cs1,cs2,cs. cs2 ⋓ cs1 ≡ cs → + ∧∧ |cs1| = |cs2| & |cs| = |cs1| & |cs| = |cs2|. +#cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs /2 width=1 by and3_intro/ +#cs1 #cs2 #cs #_ #b1 #b2 * /2 width=1 by and3_intro/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma sor_fwd_isid_sn: ∀cs1,cs2,cs. cs1 ⋓ cs2 ≡ cs → 𝐈⦃cs1⦄ → 𝐈⦃cs⦄. +#cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs // +#cs1 #cs2 #cs #_ #b1 #b2 #IH #H elim (isid_inv_cons … H) -H +/3 width=1 by isid_true/ +qed-. + +lemma sor_fwd_isid_dx: ∀cs1,cs2,cs. cs1 ⋓ cs2 ≡ cs → 𝐈⦃cs2⦄ → 𝐈⦃cs⦄. +/3 width=4 by sor_fwd_isid_sn, sor_sym/ qed-.