]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpg_drops.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpg_drops.ma
index 6c9fdfd22190e36447009acb8ebb9abfd6e4b55c..a98e061f39d5a61c872c36ca97cf219d573453f3 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/drops_drops.ma".
-include "basic_2/s_computation/fqup_weight.ma".
-include "basic_2/s_computation/fqup_drops.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: â\88\80Rt,c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.ⓓV → ⦃G, K⦄ ⊢ V ⬈[Rt, c, h] V2 →
-                       â¬\86*[⫯i] V2 â\89¡ T2 →  ⦃G, L⦄ ⊢ #i ⬈[Rt, c, h] T2.
+lemma cpg_delta_drops: â\88\80Rt,c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89\98 K.ⓓV → ⦃G, K⦄ ⊢ V ⬈[Rt, c, h] V2 →
+                       â¬\86*[â\86\91i] V2 â\89\98 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 #V0 #HLK #H destruct
-  elim (lifts_split_trans â\80¦ HVT2 (ð\9d\90\94â\9d´â«¯i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
+  elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct
+  elim (lifts_split_trans â\80¦ HVT2 (ð\9d\90\94â\9d´â\86\91i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
 ]
 qed.
 
-lemma cpg_ell_drops: â\88\80Rt,c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89¡ K.ⓛV → ⦃G, K⦄ ⊢ V ⬈[Rt,c, h] V2 →
-                     â¬\86*[⫯i] V2 â\89¡ T2 →  ⦃G, L⦄ ⊢ #i ⬈[Rt, c+𝟘𝟙, h] T2.
+lemma cpg_ell_drops: â\88\80Rt,c,h,G,K,V,V2,i,L,T2. â¬\87*[i] L â\89\98 K.ⓛV → ⦃G, K⦄ ⊢ V ⬈[Rt,c, h] V2 →
+                     â¬\86*[â\86\91i] V2 â\89\98 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 #V0 #HLK #H destruct
-  elim (lifts_split_trans â\80¦ HVT2 (ð\9d\90\94â\9d´â«¯i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
+  elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct
+  elim (lifts_split_trans â\80¦ HVT2 (ð\9d\90\94â\9d´â\86\91i❵) (𝐔❴1❵) ?) -HVT2 /3 width=3 by cpg_lref/
 ]
 qed.
 
@@ -45,15 +45,15 @@ qed.
 
 lemma cpg_inv_lref1_drops: ∀Rt,c,h,G,i,L,T2. ⦃G, L⦄ ⊢ #i ⬈[Rt,c, h] T2 →
                            ∨∨ T2 = #i ∧ c = 𝟘𝟘
-                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
-                                           â¬\86*[⫯i] V2 â\89¡ T2 & c = cV
-                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89¡ K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
-                                           â¬\86*[⫯i] V2 â\89¡ T2 & c = cV + 𝟘𝟙.
+                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89\98 K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
+                                           â¬\86*[â\86\91i] V2 â\89\98 T2 & c = cV
+                            | â\88\83â\88\83cV,K,V,V2. â¬\87*[i] L â\89\98 K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
+                                           â¬\86*[â\86\91i] V2 â\89\98 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/
 | #i #IH #L #T2 #H elim (cpg_inv_lref1 … H) -H * /3 width=1 by or3_intro0, conj/
-  #I #K #V #V2 #H #HVT2 #H0 destruct elim (IH … H) -IH -H
+  #I #K #V2 #H #HVT2 #H0 destruct elim (IH … H) -IH -H
   [ * #H1 #H2 destruct lapply (lifts_inv_lref1_uni … HVT2) -HVT2 #H destruct /3 width=1 by or3_intro0, conj/ ] *
   #cV #L #W #W2 #HKL #HW2 #HWV2 #H destruct
   lapply (lifts_trans … HWV2 … HVT2 ??) -V2
@@ -64,10 +64,10 @@ qed-.
 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 = 𝟘𝟙
-                            | â\88\83â\88\83cV,i,K,V,V2. â¬\87*[i] L â\89¡ K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
-                                             â¬\86*[⫯i] V2 â\89¡ T2 & I = LRef i & c = cV
-                            | â\88\83â\88\83cV,i,K,V,V2. â¬\87*[i] L â\89¡ K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
-                                             â¬\86*[⫯i] V2 â\89¡ T2 & I = LRef i & c = cV + 𝟘𝟙.
+                            | â\88\83â\88\83cV,i,K,V,V2. â¬\87*[i] L â\89\98 K.ⓓV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
+                                             â¬\86*[â\86\91i] V2 â\89\98 T2 & I = LRef i & c = cV
+                            | â\88\83â\88\83cV,i,K,V,V2. â¬\87*[i] L â\89\98 K.ⓛV & ⦃G, K⦄ ⊢ V ⬈[Rt, cV, h] V2 &
+                                             â¬\86*[â\86\91i] V2 â\89\98 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/
@@ -81,9 +81,10 @@ qed-.
 (* Properties with generic slicing for local environments *******************)
 
 (* Note: it should use drops_split_trans_pair2 *)
-lemma cpg_lifts: ∀Rt. reflexive … Rt → ∀c,h,G. d_liftable2 (cpg Rt h c G).
+lemma cpg_lifts_sn: ∀Rt. reflexive … Rt →
+                    ∀c,h,G. d_liftable2_sn … lifts (cpg Rt h c G).
 #Rt #HRt #c #h #G #K #T generalize in match c; -c
-@(fqup_wf_ind_eq … G K T) -G -K -T #G0 #K0 #T0 #IH #G #K * *
+@(fqup_wf_ind_eq (Ⓣ) … G K T) -G -K -T #G0 #K0 #T0 #IH #G #K * *
 [ #s #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct -IH
   lapply (lifts_inv_sort1 … H1) -H1 #H destruct
   elim (cpg_inv_sort1 … H2) -H2 * #H1 #H2 destruct
@@ -96,9 +97,10 @@ lemma cpg_lifts: ∀Rt. reflexive … Rt → ∀c,h,G. d_liftable2 (cpg Rt h c G
   lapply (drops_trans … HLK … HK0 ??) -HLK [3,6: |*: // ] #H
   elim (drops_split_trans … H) -H [1,6: |*: /2 width=6 by after_uni_dx/ ] #Y #HL0 #HY
   lapply (drops_tls_at … Hf … HY) -HY #HY
-  elim (drops_inv_skip2 … HY) -HY #L0 #W #HLK0 #HVW #H destruct
+  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 (ð\9d\90\94â\9d´â«¯i2❵)) #U2 #HWU2
+  elim (lifts_total W2 (ð\9d\90\94â\9d´â\86\91i2❵)) #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/
@@ -111,12 +113,12 @@ lemma cpg_lifts: ∀Rt. reflexive … Rt → ∀c,h,G. d_liftable2 (cpg Rt h c G
   elim (cpg_inv_bind1 … H2) -H2 *
   [ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
     elim (IH … HV12 … HLK … HVW1) -HV12 //
-    elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    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/ ] #U2 #HTU2 #HU12
+    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 â\80¦ HXU2 f (ð\9d\90\94â\9d´â«¯O❵)) [2: /2 width=1 by after_uni_one_dx/ ]
+    elim (lifts_split_trans â\80¦ HXU2 f (ð\9d\90\94â\9d´â\86\91O❵)) [2: /2 width=1 by after_uni_one_dx/ ]
     /3 width=5 by cpg_zeta, ex2_intro/
   ]
 | * #V1 #T1 #HG #HK #HT #c #X2 #H2 #b #f #L #HLK #X1 #H1 destruct
@@ -130,16 +132,16 @@ lemma cpg_lifts: ∀Rt. reflexive … Rt → ∀c,h,G. d_liftable2 (cpg Rt h c G
       elim (lifts_inv_bind1 … HTU1) -HTU1 #Z1 #U0 #HYZ1 #HTU1 #H destruct
       elim (IH … HV12 … HLK … HVW1) -HV12 -HVW1 //
       elim (IH … HY12 … HLK … HYZ1) -HY12 //
-      elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+      elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ]
       /4 width=7 by cpg_beta, lifts_bind, lifts_flat, ex2_intro/
     | #cV #cY #cT #a #V2 #V20 #Y1 #Y2 #T0 #T2 #HV12 #HV20 #HY12 #HT12 #H1 #H2 #H3 destruct
       elim (lifts_inv_bind1 … HTU1) -HTU1 #Z1 #U0 #HYZ1 #HTU1 #H destruct
       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/ ]
+      elim (IH … HT12 … HTU1) -IH -HT12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ]
       elim (lifts_total W2 (𝐔❴1❵)) #W20 #HW20
       lapply (lifts_trans … HVW2 … HW20 ??) -HVW2 [3: |*: // ] #H
-      lapply (lifts_conf â\80¦ HV20 â\80¦ H (â\86\91f) ?) -V2 /2 width=3 by after_uni_one_sn/
+      lapply (lifts_conf â\80¦ HV20 â\80¦ H (⫯f) ?) -V2 /2 width=3 by after_uni_one_sn/
       /4 width=9 by cpg_theta, lifts_bind, lifts_flat, ex2_intro/
     ]
   | elim (cpg_inv_cast1 … H2) -H2 *
@@ -158,11 +160,16 @@ lemma cpg_lifts: ∀Rt. reflexive … Rt → ∀c,h,G. d_liftable2 (cpg Rt h c G
 ]
 qed-.
 
+lemma cpg_lifts_bi: ∀Rt. reflexive … Rt →
+                    ∀c,h,G. d_liftable2_bi … lifts (cpg Rt h c G).
+/3 width=12 by cpg_lifts_sn, d_liftable2_sn_bi, lifts_mono/ qed-.
+
 (* Inversion lemmas with generic slicing for local environments *************)
 
-lemma cpg_inv_lifts1: ∀Rt. reflexive … Rt → ∀c,h,G. d_deliftable2_sn (cpg Rt h c G).
+lemma cpg_inv_lifts_sn: ∀Rt. reflexive … Rt →
+                        ∀c,h,G. d_deliftable2_sn … lifts (cpg Rt h c G).
 #Rt #HRt #c #h #G #L #U generalize in match c; -c
-@(fqup_wf_ind_eq … G L U) -G -L -U #G0 #L0 #U0 #IH #G #L * *
+@(fqup_wf_ind_eq (Ⓣ) … G L U) -G -L -U #G0 #L0 #U0 #IH #G #L * *
 [ #s #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct -IH
   lapply (lifts_inv_sort2 … H1) -H1 #H destruct
   elim (cpg_inv_sort1 … H2) -H2 * #H1 #H2 destruct
@@ -175,7 +182,8 @@ lemma cpg_inv_lifts1: ∀Rt. reflexive … Rt → ∀c,h,G. d_deliftable2_sn (cp
   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 #K0 #V #HLK0 #HVW #H destruct
+  elim (drops_inv_skip1 … HLY0) -HLY0 #Z #K0 #HLK0 #HZ #H destruct
+  elim (liftsb_inv_pair_dx … HZ) -HZ #V #HVW #H destruct
   elim (IH … HW2 … HLK0 … HVW) -IH /2 width=2 by fqup_lref/ -L -L0 -W #V2 #HVW2 #HV2
   lapply (lifts_trans … HVW2 … HWU2 ??) -W2 [3,6: |*: // ] #HVU2
   elim (lifts_split_trans … HVU2 ? f) -HVU2 [1,4: |*: /2 width=4 by after_uni_succ_sn/ ]
@@ -189,10 +197,10 @@ lemma cpg_inv_lifts1: ∀Rt. reflexive … Rt → ∀c,h,G. d_deliftable2_sn (cp
   elim (cpg_inv_bind1 … H2) -H2 *
   [ #cW #cU #W2 #U2 #HW12 #HU12 #H1 #H2 destruct
     elim (IH … HW12 … HLK … HVW1) -HW12 //
-    elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+    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/ ] #T2 #HTU2 #HT12
+    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/
   ]
 | * #W1 #U1 #HG #HL #HU #c #X2 #H2 #b #f #K #HLK #X1 #H1 destruct
@@ -206,15 +214,15 @@ lemma cpg_inv_lifts1: ∀Rt. reflexive … Rt → ∀c,h,G. d_deliftable2_sn (cp
       elim (lifts_inv_bind2 … HTU1) -HTU1 #Y1 #T0 #HYZ1 #HTU1 #H destruct
       elim (IH … HW12 … HLK … HVW1) -HW12 -HVW1 //
       elim (IH … HZ12 … HLK … HYZ1) -HZ12 //
-      elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+      elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ]
       /4 width=7 by cpg_beta, lifts_bind, lifts_flat, ex2_intro/
     | #cW #cZ #cU #a #W2 #W20 #Z1 #Z2 #U0 #U2 #HW12 #HW20 #HZ12 #HU12 #H1 #H2 #H3 destruct
       elim (lifts_inv_bind2 … HTU1) -HTU1 #Y1 #T0 #HYZ1 #HTU1 #H destruct
       elim (IH … HW12 … HLK … HVW1) -HW12 -HVW1 // #V2 #HVW2 #HV12
       elim (IH … HZ12 … HLK … HYZ1) -HZ12 //
-      elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip/ ]
+      elim (IH … HU12 … HTU1) -IH -HU12 -HTU1 [ |*: /3 width=3 by drops_skip, ext2_pair/ ]
       lapply (lifts_trans … HVW2 … HW20 ??) -W2 [3: |*: // ] #H
-      elim (lifts_split_trans â\80¦ H ? (â\86\91f)) -H [ |*: /2 width=3 by after_uni_one_sn/ ]
+      elim (lifts_split_trans â\80¦ H ? (⫯f)) -H [ |*: /2 width=3 by after_uni_one_sn/ ]
       /4 width=9 by cpg_theta, lifts_bind, lifts_flat, ex2_intro/
     ]
   | elim (cpg_inv_cast1 … H2) -H2 *
@@ -232,3 +240,7 @@ lemma cpg_inv_lifts1: ∀Rt. reflexive … Rt → ∀c,h,G. d_deliftable2_sn (cp
   ]
 ]
 qed-.
+
+lemma cpg_inv_lifts_bi: ∀Rt. reflexive … Rt →
+                        ∀c,h,G. d_deliftable2_bi … lifts (cpg Rt h c G).
+/3 width=12 by cpg_inv_lifts_sn, d_deliftable2_sn_bi, lifts_inj/ qed-.