X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfsx_csx.ma;h=862858c4af0fd07648b0e833509b70cc8460a748;hb=bc40346f09bcccb9a09560963ccb7157ebfad7ad;hp=ff005164fb89a0650b3aa1a9d3df536e3e65a5a5;hpb=c879284b576409cec07e96c1f08510d9d9ac14f3;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_csx.ma index ff005164f..862858c4a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfsx_csx.ma @@ -12,34 +12,44 @@ (* *) (**************************************************************************) +include "basic_2/rt_computation/csx_cpxs.ma". include "basic_2/rt_computation/csx_lsubr.ma". -include "basic_2/rt_computation/lfsx_lfpxs.ma". +include "basic_2/rt_computation/lsubsx_lfsx.ma". (* STRONGLY NORMALIZING LOCAL ENV.S FOR UNCOUNTED PARALLEL RT-TRANSITION ****) -(* + (* Advanced properties ******************************************************) -lemma lsx_lref_be_lpxs: ∀h,o,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, o] V → - ∀K2. G ⊢ ⬊*[h, o, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, o] K2 → - ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L2. -#h #o #I #G #K1 #V #i #l #Hli #H @(csx_ind_alt … H) -V -#V0 #_ #IHV0 #K2 #H @(lsx_ind … H) -K2 -#K0 #HK0 #IHK0 #HK10 #L0 #HLK0 @lsx_intro -#L2 #HL02 #HnL02 elim (lpx_drop_conf … HLK0 … HL02) -HL02 -#Y #H #HLK2 elim (lpx_inv_pair1 … H) -H -#K2 #V2 #HK02 #HV02 #H destruct -elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0 ] -[ /4 width=8 by lpxs_strap1, lleq_lref/ -| @(IHV0 … HnV02 … HLK2) -IHV0 -HnV02 -HLK2 - /3 width=4 by lsx_cpx_trans_O, lsx_lpx_trans, lpxs_cpx_trans, lpxs_strap1/ (**) (* full auto too slow *) +(* Basic_2A1: uses: lsx_lref_be_lpxs *) +lemma lfsx_pair_lfpxs: ∀h,o,G,K1,V. ⦃G, K1⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ → + ∀K2. G ⊢ ⬈*[h, o, V] 𝐒⦃K2⦄ → ⦃G, K1⦄ ⊢ ⬈*[h, V] K2 → + ∀I. G ⊢ ⬈*[h, o, #0] 𝐒⦃K2.ⓑ{I}V⦄. +#h #o #G #K1 #V #H +@(csx_ind_cpxs … H) -V #V0 #_ #IHV0 #K2 #H +@(lfsx_ind … H) -K2 #K0 #HK0 #IHK0 #HK10 #I +@lfsx_intro #Y #HY #HnY +elim (lfpx_inv_zero_pair_sn … HY) -HY #K2 #V2 #HK02 #HV02 #H destruct +elim (tdeq_dec h o V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnY ] +[ /5 width=5 by lfsx_lfdeq_trans, lfpxs_step_dx, lfdeq_pair/ +| @lfsx_lfpx_trans + [2: @(IHV0 … HnV02 K0 … I) -IHV0 -HnV02 + [ /2 width=3 by lfpxs_cpx_trans/ + | /2 width=3 by lfsx_cpx_trans/ + | + ] + |1: skip + |3: @lfpx_pair /2 width=3 by lfpx_cpx_conf/ + ] ] qed. -lemma lsx_lref_be: ∀h,o,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, o] V → - G ⊢ ⬊*[h, o, V, 0] K → - ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L. -/2 width=8 by lsx_lref_be_lpxs/ qed. -*) +(* Basic_2A1: uses: lsx_lref_be *) +lemma lfsx_lref_pair: ∀h,o,G,K,V. ⦃G, K⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ → G ⊢ ⬈*[h, o, V] 𝐒⦃K⦄ → + ∀I,L,i. ⬇*[i] L ≡ K.ⓑ{I}V → G ⊢ ⬈*[h, o, #i] 𝐒⦃L⦄. +#h #o #G #K #V #HV #HK #I #L #i #HLK +@(lfsx_lifts … (#0) … HLK) -L /2 width=3 by lfsx_pair_lpxs/ +qed. + (* Main properties **********************************************************) theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. @@ -47,17 +57,10 @@ theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → G ⊢ #Z #Y #X #IH #G #L * * // [ #i #HG #HL #HT #H destruct elim (csx_inv_lref … H) -H [ |*: * ] - [ #HL - | #I #K #HLK - | #I #K #V #HLK #HV + [ /2 width=1 by lfsx_lref_atom/ + | /2 width=3 by lfsx_lref_unit/ + | /4 width=6 by lfsx_lref_pair, fqup_lref/ ] -(* - elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/ - elim (ylt_split i l) /2 width=1 by lsx_lref_skip/ - #Hli #Hi elim (drop_O1_lt (Ⓕ) … Hi) -Hi - #I #K #V #HLK lapply (csx_inv_lref_bind … HLK … H) -H - /4 width=6 by lsx_lref_be, fqup_lref/ -*) | #a #I #V #T #HG #HL #HT #H destruct elim (csx_fwd_bind_unit … H Void) -H /3 width=1 by lfsx_bind_void/ | #I #V #T #HG #HL #HT #H destruct