From: Ferruccio Guidi Date: Thu, 9 Nov 2017 21:59:39 +0000 (+0000) Subject: advances on lfsx_csx ... X-Git-Tag: make_still_working~411 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=9dd07546d24cfdcf62276ce4518a69e8bbf0f9d3 advances on lfsx_csx ... --- 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..29ed16c91 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,16 +12,20 @@ (* *) (**************************************************************************) +include "basic_2/rt_computation/lfpxs.ma". include "basic_2/rt_computation/csx_lsubr.ma". +include "basic_2/rt_computation/lfsx_drops.ma". include "basic_2/rt_computation/lfsx_lfpxs.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. +(* Basic_2A1: uses: lsx_lref_be_lpxs *) +axiom lfsx_pair_lpxs: ∀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 #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 @@ -34,12 +38,14 @@ elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0 /3 width=4 by lsx_cpx_trans_O, lsx_lpx_trans, lpxs_cpx_trans, lpxs_strap1/ (**) (* full auto too slow *) ] 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 +53,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