X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpg_drops.ma;h=482a4b12630589d574299b4d83406455da91f572;hb=d8d00d6f6694155be5be486a8239f5953efe28b7;hp=8b33836a250f77f43660632e261970e7b2d8e192;hpb=397413c4196f84c81d61ba7dd79b54ab1c428ebb;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma index 8b33836a2..482a4b126 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma @@ -12,17 +12,18 @@ (* *) (**************************************************************************) -include "basic_2/relocation/drops_drops.ma". -include "basic_2/s_computation/fqup_weight.ma". -include "basic_2/s_computation/fqup_drops.ma". +include "ground_2/xoa/ex_5_5.ma". +include "static_2/relocation/drops_drops.ma". +include "static_2/s_computation/fqup_weight.ma". +include "static_2/s_computation/fqup_drops.ma". include "basic_2/rt_transition/cpg.ma". -(* COUNTED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS ***************) +(* BOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION FOR TERMS *****************) (* Advanced properties ******************************************************) -lemma cpg_delta_drops: ∀Rt,c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≘ K.ⓓV → ⦃G, K⦄ ⊢ V ⬈[Rt, c, h] V2 → - ⬆*[↑i] V2 ≘ T2 → ⦃G, L⦄ ⊢ #i ⬈[Rt, c, h] T2. +lemma cpg_delta_drops: ∀Rt,c,h,G,K,V,V2,i,L,T2. ⇩*[i] L ≘ K.ⓓV → ⦃G,K⦄ ⊢ V ⬈[Rt,c,h] V2 → + ⇧*[↑i] V2 ≘ T2 → ⦃G,L⦄ ⊢ #i ⬈[Rt,c,h] T2. #Rt #c #h #G #K #V #V2 #i elim i -i [ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_delta/ | #i #IH #L0 #T0 #H0 #HV2 #HVT2 @@ -31,8 +32,8 @@ lemma cpg_delta_drops: ∀Rt,c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≘ K.ⓓV → ⦃G, ] qed. -lemma cpg_ell_drops: ∀Rt,c,h,G,K,V,V2,i,L,T2. ⬇*[i] L ≘ K.ⓛV → ⦃G, K⦄ ⊢ V ⬈[Rt,c, h] V2 → - ⬆*[↑i] V2 ≘ T2 → ⦃G, L⦄ ⊢ #i ⬈[Rt, c+𝟘𝟙, h] T2. +lemma cpg_ell_drops: ∀Rt,c,h,G,K,V,V2,i,L,T2. ⇩*[i] L ≘ K.ⓛV → ⦃G,K⦄ ⊢ V ⬈[Rt,c,h] V2 → + ⇧*[↑i] V2 ≘ T2 → ⦃G,L⦄ ⊢ #i ⬈[Rt,c+𝟘𝟙,h] T2. #Rt #c #h #G #K #V #V2 #i elim i -i [ #L #T2 #HLK lapply (drops_fwd_isid … HLK ?) // #H destruct /3 width=3 by cpg_ell/ | #i #IH #L0 #T0 #H0 #HV2 #HVT2 @@ -43,12 +44,12 @@ qed. (* Advanced inversion lemmas ************************************************) -lemma cpg_inv_lref1_drops: ∀Rt,c,h,G,i,L,T2. ⦃G, L⦄ ⊢ #i ⬈[Rt,c, h] T2 → +lemma cpg_inv_lref1_drops: ∀Rt,c,h,G,i,L,T2. ⦃G,L⦄ ⊢ #i ⬈[Rt,c,h] T2 → ∨∨ T2 = #i ∧ c = 𝟘𝟘 - | ∃∃cV,K,V,V2. ⬇*[i] L ≘ K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 & - ⬆*[↑i] V2 ≘ T2 & c = cV - | ∃∃cV,K,V,V2. ⬇*[i] L ≘ K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 & - ⬆*[↑i] V2 ≘ T2 & c = cV + 𝟘𝟙. + | ∃∃cV,K,V,V2. ⇩*[i] L ≘ K.ⓓV & ⦃G,K⦄ ⊢ V ⬈[Rt,cV,h] V2 & + ⇧*[↑i] V2 ≘ T2 & c = cV + | ∃∃cV,K,V,V2. ⇩*[i] L ≘ K.ⓛV & ⦃G,K⦄ ⊢ V ⬈[Rt,cV,h] V2 & + ⇧*[↑i] V2 ≘ T2 & c = cV + 𝟘𝟙. #Rt #c #h #G #i elim i -i [ #L #T2 #H elim (cpg_inv_zero1 … H) -H * /3 width=1 by or3_intro0, conj/ /4 width=8 by drops_refl, ex4_4_intro, or3_intro2, or3_intro1/ @@ -61,13 +62,13 @@ lemma cpg_inv_lref1_drops: ∀Rt,c,h,G,i,L,T2. ⦃G, L⦄ ⊢ #i ⬈[Rt,c, h] T2 ] qed-. -lemma cpg_inv_atom1_drops: ∀Rt,c,h,I,G,L,T2. ⦃G, L⦄ ⊢ ⓪{I} ⬈[Rt, c, h] T2 → +lemma cpg_inv_atom1_drops: ∀Rt,c,h,I,G,L,T2. ⦃G,L⦄ ⊢ ⓪{I} ⬈[Rt,c,h] T2 → ∨∨ T2 = ⓪{I} ∧ c = 𝟘𝟘 - | ∃∃s. T2 = ⋆(next h s) & I = Sort s & c = 𝟘𝟙 - | ∃∃cV,i,K,V,V2. ⬇*[i] L ≘ K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 & - ⬆*[↑i] V2 ≘ T2 & I = LRef i & c = cV - | ∃∃cV,i,K,V,V2. ⬇*[i] L ≘ K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 & - ⬆*[↑i] V2 ≘ T2 & I = LRef i & c = cV + 𝟘𝟙. + | ∃∃s. T2 = ⋆(⫯[h]s) & I = Sort s & c = 𝟘𝟙 + | ∃∃cV,i,K,V,V2. ⇩*[i] L ≘ K.ⓓV & ⦃G,K⦄ ⊢ V ⬈[Rt,cV,h] V2 & + ⇧*[↑i] V2 ≘ T2 & I = LRef i & c = cV + | ∃∃cV,i,K,V,V2. ⇩*[i] L ≘ K.ⓛV & ⦃G,K⦄ ⊢ V ⬈[Rt,cV,h] V2 & + ⇧*[↑i] V2 ≘ T2 & I = LRef i & c = cV + 𝟘𝟙. #Rt #c #h * #n #G #L #T2 #H [ elim (cpg_inv_sort1 … H) -H * /3 width=3 by or4_intro0, or4_intro1, ex3_intro, conj/ @@ -115,10 +116,9 @@ lemma cpg_lifts_sn: ∀Rt. reflexive … Rt → elim (IH … HV12 … HLK … HVW1) -HV12 // elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ] /3 width=5 by cpg_bind, lifts_bind, ex2_intro/ - | #cT #T2 #HT12 #HXT2 #H1 #H2 #H3 destruct - elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ] #U2 #HTU2 #HU12 - lapply (lifts_trans … HXT2 … HTU2 ??) -T2 [3: |*: // ] #HXU2 - elim (lifts_split_trans … HXU2 f (𝐔❴↑O❵)) [2: /2 width=1 by after_uni_one_dx/ ] + | #cT #T2 #HT21 #HTX2 #H1 #H2 #H3 destruct + elim (lifts_trans4_one … HT21 … HTU1) -HTU1 #U2 #HTU2 #HU21 + elim (IH … HTX2 … HLK … HTU2) [| /3 width=1 by fqup_zeta/ ] -K -V1 -T1 -T2 /3 width=5 by cpg_zeta, ex2_intro/ ] | * #V1 #T1 #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct @@ -199,9 +199,10 @@ lemma cpg_inv_lifts_sn: ∀Rt. reflexive … Rt → elim (IH … HW12 … HLK … HVW1) -HW12 // elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ] /3 width=5 by cpg_bind, lifts_bind, ex2_intro/ - | #cU #U2 #HU12 #HXU2 #H1 #H2 #H3 destruct - elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ] #T2 #HTU2 #HT12 - elim (lifts_div4_one … HTU2 … HXU2) -U2 /3 width=5 by cpg_zeta, ex2_intro/ + | #cU #U2 #HU21 #HUX2 #H1 #H2 #H3 destruct + elim (lifts_div4_one … HTU1 … HU21) -HTU1 #T2 #HT21 #HTU2 + elim (IH … HUX2 … HLK … HTU2) [| /3 width=1 by fqup_zeta/ ] -L -W1 -U1 -U2 + /3 width=5 by cpg_zeta, ex2_intro/ ] | * #W1 #U1 #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct elim (lifts_inv_flat2 … H1) -H1 #V1 #T1 #HVW1 #HTU1 #H destruct