X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcpg_drops.ma;h=6bfc84c4e1088071ca8d7c7eea4a1eb14e546aae;hp=482a4b12630589d574299b4d83406455da91f572;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 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 482a4b126..6bfc84c4e 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 @@ -22,33 +22,33 @@ include "basic_2/rt_transition/cpg.ma". (* 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 elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct - elim (lifts_split_trans … HVT2 (𝐔❴↑i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/ + elim (lifts_split_trans … HVT2 (𝐔❨↑i❩) (𝐔❨1❩) ?) -HVT2 /3 width=3 by cpg_lref/ ] 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 elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct - elim (lifts_split_trans … HVT2 (𝐔❴↑i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/ + elim (lifts_split_trans … HVT2 (𝐔❨↑i❩) (𝐔❨1❩) ?) -HVT2 /3 width=3 by cpg_lref/ ] 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 & + | ∃∃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 & + | ∃∃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/ @@ -62,12 +62,12 @@ 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 → - ∨∨ T2 = ⓪{I} ∧ c = 𝟘𝟘 +lemma cpg_inv_atom1_drops: ∀Rt,c,h,I,G,L,T2. ❪G,L❫ ⊢ ⓪[I] ⬈[Rt,c,h] T2 → + ∨∨ T2 = ⓪[I] ∧ c = 𝟘𝟘 | ∃∃s. T2 = ⋆(⫯[h]s) & I = Sort s & c = 𝟘𝟙 - | ∃∃cV,i,K,V,V2. ⇩*[i] L ≘ K.ⓓV & ⦃G,K⦄ ⊢ V ⬈[Rt,cV,h] V2 & + | ∃∃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 & + | ∃∃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 * @@ -101,7 +101,7 @@ lemma cpg_lifts_sn: ∀Rt. reflexive … Rt → elim (drops_inv_skip2 … HY) -HY #Z #L0 #HLK0 #HZ #H destruct elim (liftsb_inv_pair_sn … HZ) -HZ #W #HVW #H destruct elim (IH … HV2 … HLK0 … HVW) -IH /2 width=2 by fqup_lref/ -K -K0 -V #W2 #HVW2 #HW2 - elim (lifts_total W2 (𝐔❴↑i2❵)) #U2 #HWU2 + elim (lifts_total W2 (𝐔❨↑i2❩)) #U2 #HWU2 lapply (lifts_trans … HVW2 … HWU2 ??) -HVW2 [3,6: |*: // ] #HVU2 lapply (lifts_conf … HVT2 … HVU2 f ?) -V2 [1,3: /2 width=3 by after_uni_succ_sn/ ] /4 width=8 by cpg_ell_drops, cpg_delta_drops, drops_inv_gen, ex2_intro/ @@ -139,7 +139,7 @@ lemma cpg_lifts_sn: ∀Rt. reflexive … Rt → elim (IH … HV12 … HLK … HVW1) -HV12 -HVW1 // #W2 #HVW2 #HW12 elim (IH … HY12 … HLK … HYZ1) -HY12 // elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ] - elim (lifts_total W2 (𝐔❴1❵)) #W20 #HW20 + elim (lifts_total W2 (𝐔❨1❩)) #W20 #HW20 lapply (lifts_trans … HVW2 … HW20 ??) -HVW2 [3: |*: // ] #H lapply (lifts_conf … HV20 … H (⫯f) ?) -V2 /2 width=3 by after_uni_one_sn/ /4 width=9 by cpg_theta, lifts_bind, lifts_flat, ex2_intro/ @@ -179,7 +179,7 @@ lemma cpg_inv_lifts_sn: ∀Rt. reflexive … Rt → [ #H1 #H2 destruct /3 width=3 by cpg_refl, ex2_intro/ ] #cW #L0 #W #W2 #HL0 #HW2 #HWU2 #H destruct elim (lifts_inv_lref2 … H1) -H1 #i1 #Hf #H destruct - lapply (drops_split_div … HLK (𝐔❴i1❵) ???) -HLK [4,8: * |*: // ] #Y0 #HK0 #HLY0 + lapply (drops_split_div … HLK (𝐔❨i1❩) ???) -HLK [4,8: * |*: // ] #Y0 #HK0 #HLY0 lapply (drops_conf … HL0 … HLY0 ??) -HLY0 [3,6: |*: /2 width=6 by after_uni_dx/ ] #HLY0 lapply (drops_tls_at … Hf … HLY0) -HLY0 #HLY0 elim (drops_inv_skip1 … HLY0) -HLY0 #Z #K0 #HLK0 #HZ #H destruct