X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fetc%2Fcpr%2Fltpss_sn_alt.etc;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Fetc%2Fcpr%2Fltpss_sn_alt.etc;h=d2cdba023c7baa581f5df8f3866ea4a4fe4a4bdd;hb=1fd63df4c77f5c24024769432ea8492748b4ac79;hp=0000000000000000000000000000000000000000;hpb=277fc8ff21ce3dbd6893b1994c55cf5c06a98355;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/etc/cpr/ltpss_sn_alt.etc b/matita/matita/contribs/lambdadelta/basic_2A/etc/cpr/ltpss_sn_alt.etc new file mode 100644 index 000000000..d2cdba023 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/etc/cpr/ltpss_sn_alt.etc @@ -0,0 +1,160 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +notation "hvbox( T1 break ⊢ ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStarSnAlt $T1 $d $e $T2 }. + +include "basic_2/unfold/ltpss_dx_ltpss_dx.ma". +include "basic_2/unfold/ltpss_sn_ltpss_sn.ma". + +(* SN PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) + +(* alternative definition of ltpss_sn *) +definition ltpssa: nat → nat → relation lenv ≝ + λd,e. TC … (ltpss_dx d e). + +interpretation "parallel unfold (local environment, sn variant) alternative" + 'PSubstStarSnAlt L1 d e L2 = (ltpssa d e L1 L2). + +(* Basic eliminators ********************************************************) + +lemma ltpssa_ind: ∀d,e,L1. ∀R:predicate lenv. R L1 → + (∀L,L2. L1 ⊢ ▶▶* [d, e] L → L ▶* [d, e] L2 → R L → R L2) → + ∀L2. L1 ⊢ ▶▶* [d, e] L2 → R L2. +#d #e #L1 #R #HL1 #IHL1 #L2 #HL12 @(TC_star_ind … HL1 IHL1 … HL12) // +qed-. + +lemma ltpssa_ind_dx: ∀d,e,L2. ∀R:predicate lenv. R L2 → + (∀L1,L. L1 ▶* [d, e] L → L ⊢ ▶▶* [d, e] L2 → R L → R L1) → + ∀L1. L1 ⊢ ▶▶* [d, e] L2 → R L1. +#d #e #L2 #R #HL2 #IHL2 #L1 #HL12 @(TC_star_ind_dx … HL2 IHL2 … HL12) // +qed-. + +(* Basic properties *********************************************************) + +lemma ltpssa_refl: ∀L,d,e. L ⊢ ▶▶* [d, e] L. +/2 width=1/ qed. + +lemma ltpssa_tpss2: ∀I,L1,V1,V2,e. L1 ⊢ V1 ▶*[0, e] V2 → + ∀L2. L1 ⊢ ▶▶* [0, e] L2 → + L1.ⓑ{I}V1 ⊢ ▶▶* [O, e + 1] L2.ⓑ{I}V2. +#I #L1 #V1 #V2 #e #HV12 #L2 #H @(ltpssa_ind … H) -L2 +[ /3 width=1/ | /3 width=5/ ] +qed. + +lemma ltpssa_tpss1: ∀I,L1,V1,V2,d,e. L1 ⊢ V1 ▶*[d, e] V2 → + ∀L2. L1 ⊢ ▶▶* [d, e] L2 → + L1.ⓑ{I}V1 ⊢ ▶▶* [d + 1, e] L2.ⓑ{I}V2. +#I #L1 #V1 #V2 #d #e #HV12 #L2 #H @(ltpssa_ind … H) -L2 +[ /3 width=1/ | /3 width=5/ ] +qed. + +lemma ltpss_sn_ltpssa: ∀L1,L2,d,e. L1 ⊢ ▶* [d, e] L2 → L1 ⊢ ▶▶* [d, e] L2. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // /2 width=1/ +qed. + +lemma ltpss_sn_dx_trans_eq: ∀L1,L,d,e. L1 ⊢ ▶* [d, e] L → + ∀L2. L ▶* [d, e] L2 → L1 ⊢ ▶* [d, e] L2. +#L1 #L #d #e #H elim H -L1 -L -d -e +[ #d #e #X #H + lapply (ltpss_dx_inv_atom1 … H) -H #H destruct // +| #L #I #V #X #H + lapply (ltpss_dx_inv_refl_O2 … H) -H #H destruct // +| #L1 #L #I #V1 #V #e #_ #HV1 #IHL1 #X #H + elim (ltpss_dx_inv_tpss21 … H ?) -H //