X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Fcpxs_tsts.ma;h=6434ff37fd9e36a344c4ebd527bd74cd6b244821;hb=88dd0e28758c693660a93ee0a9a5202c61ca09a0;hp=ace2532d977db9b6c33ce4b77466f96d66084bef;hpb=e9f96fa56226dfd74de214c89d827de0c5018ac7;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tsts.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tsts.ma index ace2532d9..6434ff37f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tsts.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_tsts.ma @@ -19,13 +19,13 @@ include "basic_2/computation/lpxs_cpxs.ma". (* Forward lemmas involving same top term structure *************************) -lemma cpxs_fwd_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ➡[h, o] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ➡*[h, o] U → T ≂ U. +lemma cpxs_fwd_cnx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄ → ∀U. ⦃G, L⦄ ⊢ T ⬈*[h, o] U → T ≂ U. #h #o #G #L #T #HT #U #H >(cpxs_inv_cnx1 … H HT) -G -L -T // qed-. -lemma cpxs_fwd_sort: ∀h,o,G,L,U,s. ⦃G, L⦄ ⊢ ⋆s ➡*[h, o] U → - ⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h s) ➡*[h, o] U. +lemma cpxs_fwd_sort: ∀h,o,G,L,U,s. ⦃G, L⦄ ⊢ ⋆s ⬈*[h, o] U → + ⋆s ≂ U ∨ ⦃G, L⦄ ⊢ ⋆(next h s) ⬈*[h, o] U. #h #o #G #L #U #s #H elim (cpxs_inv_sort1 … H) -H #n #d generalize in match s; -s @(nat_ind_plus … n) -n [ #s #_ #H -d destruct /2 width=1 by or_introl/ @@ -41,8 +41,8 @@ elim (cpxs_inv_sort1 … H) -H #n #d generalize in match s; -s @(nat_ind_plus qed-. (* Basic_1: was just: pr3_iso_beta *) -lemma cpxs_fwd_beta: ∀h,o,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ➡*[h, o] U → - ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ➡*[h, o] U. +lemma cpxs_fwd_beta: ∀h,o,a,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{a}W.T ⬈*[h, o] U → + ⓐV.ⓛ{a}W.T ≂ U ∨ ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V.T ⬈*[h, o] U. #h #o #a #G #L #V #W #T #U #H elim (cpxs_inv_appl1 … H) -H * [ #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/ @@ -58,8 +58,8 @@ qed-. (* Note: probably this is an inversion lemma *) lemma cpxs_fwd_delta: ∀h,o,I,G,L,K,V1,i. ⬇[i] L ≡ K.ⓑ{I}V1 → ∀V2. ⬆[0, i + 1] V1 ≡ V2 → - ∀U. ⦃G, L⦄ ⊢ #i ➡*[h, o] U → - #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ➡*[h, o] U. + ∀U. ⦃G, L⦄ ⊢ #i ⬈*[h, o] U → + #i ≂ U ∨ ⦃G, L⦄ ⊢ V2 ⬈*[h, o] U. #h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H elim (cpxs_inv_lref1 … H) -H /2 width=1 by or_introl/ * #I0 #K0 #V0 #U0 #HLK0 #HVU0 #HU0 @@ -67,9 +67,9 @@ lapply (drop_mono … HLK0 … HLK) -HLK0 #H destruct /4 width=10 by cpxs_lift, drop_fwd_drop2, or_intror/ qed-. -lemma cpxs_fwd_theta: ∀h,o,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ➡*[h, o] U → +lemma cpxs_fwd_theta: ∀h,o,a,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}V.T ⬈*[h, o] U → ∀V2. ⬆[0, 1] V1 ≡ V2 → ⓐV1.ⓓ{a}V.T ≂ U ∨ - ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡*[h, o] U. + ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ⬈*[h, o] U. #h #o #a #G #L #V1 #V #T #U #H #V2 #HV12 elim (cpxs_inv_appl1 … H) -H * [ -HV12 #V0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or_introl/ @@ -99,8 +99,8 @@ elim (cpxs_inv_appl1 … H) -H * ] qed-. -lemma cpxs_fwd_cast: ∀h,o,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ➡*[h, o] U → - ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ➡*[h, o] U | ⦃G, L⦄ ⊢ W ➡*[h, o] U. +lemma cpxs_fwd_cast: ∀h,o,G,L,W,T,U. ⦃G, L⦄ ⊢ ⓝW.T ⬈*[h, o] U → + ∨∨ ⓝW. T ≂ U | ⦃G, L⦄ ⊢ T ⬈*[h, o] U | ⦃G, L⦄ ⊢ W ⬈*[h, o] U. #h #o #G #L #W #T #U #H elim (cpxs_inv_cast1 … H) -H /2 width=1 by or3_intro1, or3_intro2/ * #W0 #T0 #_ #_ #H destruct /2 width=1 by tsts_pair, or3_intro0/