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=5310d8ab791aac77a0b133907090ed8a605fee55;hb=1604f2ee65c57eefb7c6b3122eab2a9f32e0552d;hp=862858c4af0fd07648b0e833509b70cc8460a748;hpb=bc40346f09bcccb9a09560963ccb7157ebfad7ad;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 862858c4a..5310d8ab7 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,8 +12,11 @@ (* *) (**************************************************************************) +include "basic_2/rt_computation/lpxs_lpx.ma". +include "basic_2/rt_computation/lpxs_cpxs.ma". include "basic_2/rt_computation/csx_cpxs.ma". include "basic_2/rt_computation/csx_lsubr.ma". +include "basic_2/rt_computation/lfsx_lpx.ma". include "basic_2/rt_computation/lsubsx_lfsx.ma". (* STRONGLY NORMALIZING LOCAL ENV.S FOR UNCOUNTED PARALLEL RT-TRANSITION ****) @@ -21,47 +24,48 @@ include "basic_2/rt_computation/lsubsx_lfsx.ma". (* Advanced properties ******************************************************) (* 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⦄. +lemma lfsx_pair_lpxs: ∀h,o,G,K1,V. ⦃G, K1⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ → + ∀K2. G ⊢ ⬈*[h, o, V] 𝐒⦃K2⦄ → ⦃G, K1⦄ ⊢ ⬈*[h] 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 +@(lfsx_ind_lpx … H) -K2 #K0 #HK0 #IHK0 #HK10 #I +@lfsx_intro_lpx #Y #HY #HnY +elim (lpx_inv_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/ +[ /5 width=5 by lfsx_lfdeq_trans, lpxs_step_dx, lfdeq_pair/ +| @(IHV0 … HnV02) -IHV0 -HnV02 + [ /2 width=3 by lpxs_cpx_trans/ + | /3 width=3 by lfsx_lpx_trans, lfsx_cpx_trans/ + | /2 width=3 by lpxs_step_dx/ ] ] 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/ +lemma lfsx_lref_pair_drops: ∀h,o,G,K,V. ⦃G, K⦄ ⊢ ⬈*[h, o] 𝐒⦃V⦄ → G ⊢ ⬈*[h, o, V] 𝐒⦃K⦄ → + ∀I,i,L. ⬇*[i] L ≡ K.ⓑ{I}V → G ⊢ ⬈*[h, o, #i] 𝐒⦃L⦄. +#h #o #G #K #V #HV #HK #I #i elim i -i +[ #L #H >(drops_fwd_isid … H) -H /2 width=3 by lfsx_pair_lpxs/ +| #i #IH #L #H + elim (drops_inv_bind2_isuni_next … H) -H // #J #Y #HY #H destruct + @(lfsx_lifts … (𝐔❴1❵)) /3 width=6 by drops_refl, drops_drop/ (**) (* full auto fails *) +] qed. (* Main properties **********************************************************) -theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. +(* Basic_2A1: uses: csx_lsx *) +theorem csx_lfsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬈*[h, o] 𝐒⦃T⦄ → G ⊢ ⬈*[h, o, T] 𝐒⦃L⦄. #h #o #G #L #T @(fqup_wf_ind_eq (Ⓕ) … G L T) -G -L -T #Z #Y #X #IH #G #L * * // [ #i #HG #HL #HT #H destruct elim (csx_inv_lref … H) -H [ |*: * ] [ /2 width=1 by lfsx_lref_atom/ | /2 width=3 by lfsx_lref_unit/ - | /4 width=6 by lfsx_lref_pair, fqup_lref/ + | /4 width=6 by lfsx_lref_pair_drops, fqup_lref/ ] -| #a #I #V #T #HG #HL #HT #H destruct +| #p #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 elim (csx_fwd_flat … H) -H /3 width=1 by lfsx_flat/