]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpms_drops.ma
update in basic_2 and apps_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpms_drops.ma
index b8917b38c453f57b1f9102716c2526abc3611949..fbab9de5544e0cec5bb3833e3967d6fc048a4e96 100644 (file)
@@ -20,14 +20,14 @@ include "basic_2/rt_computation/cpms.ma".
 
 (* Properties with generic slicing for local environments *******************)
 
-lemma cpms_lifts_sn: ∀n,h,G. d_liftable2_sn … lifts (λL. cpms h G L n).
+lemma cpms_lifts_sn: ∀h,n,G. d_liftable2_sn … lifts (λL. cpms h G L n).
 /3 width=6 by d2_liftable_sn_ltc, cpm_lifts_sn/ qed-.
 
 (* Basic_2A1: uses: scpds_lift *)
 (* Basic_2A1: includes: cprs_lift *)
 (* Basic_1: includes: pr3_lift *)
-lemma cpms_lifts_bi: ∀n,h,G. d_liftable2_bi … lifts (λL. cpms h G L n).
-#n #h #G @d_liftable2_sn_bi
+lemma cpms_lifts_bi: ∀h,n,G. d_liftable2_bi … lifts (λL. cpms h G L n).
+#h #n #G @d_liftable2_sn_bi
 /2 width=6 by cpms_lifts_sn, lifts_mono/
 qed-.
 
@@ -36,19 +36,19 @@ qed-.
 (* Basic_2A1: uses: scpds_inv_lift1 *)
 (* Basic_2A1: includes: cprs_inv_lift1 *)
 (* Basic_1: includes: pr3_gen_lift *)
-lemma cpms_inv_lifts_sn: ∀n,h,G. d_deliftable2_sn … lifts (λL. cpms h G L n).
+lemma cpms_inv_lifts_sn: ∀h,n,G. d_deliftable2_sn … lifts (λL. cpms h G L n).
 /3 width=6 by d2_deliftable_sn_ltc, cpm_inv_lifts_sn/ qed-.
 
-lemma cpms_inv_lifts_bi: ∀n,h,G. d_deliftable2_bi … lifts (λL. cpms h G L n).
-#n #h #G @d_deliftable2_sn_bi
+lemma cpms_inv_lifts_bi: ∀h,n,G. d_deliftable2_bi … lifts (λL. cpms h G L n).
+#h #n #G @d_deliftable2_sn_bi
 /2 width=6 by cpms_inv_lifts_sn, lifts_inj/
 qed-.
 
 (* Advanced properties ******************************************************)
 
-lemma cpms_delta (n) (h) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[n,h] V2 →
-                              ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓓV1❫ ⊢ #0 ➡*[n,h] W2.
-#n #h #G #K #V1 #V2 #H @(cpms_ind_dx … H) -V2
+lemma cpms_delta (h) (n) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[h,n] V2 →
+                              ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓓV1❫ ⊢ #0 ➡*[h,n] W2.
+#h #n #G #K #V1 #V2 #H @(cpms_ind_dx … H) -V2
 [ /3 width=3 by cpm_cpms, cpm_delta/
 | #n1 #n2 #V #V2 #_ #IH #HV2 #W2 #HVW2
   elim (lifts_total V (𝐔❨1❩)) #W #HVW
@@ -56,9 +56,9 @@ lemma cpms_delta (n) (h) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[n,h] V2 →
 ]
 qed.
 
-lemma cpms_ell (n) (h) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[n,h] V2 →
-                            ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓛV1❫ ⊢ #0 ➡*[↑n,h] W2.
-#n #h #G #K #V1 #V2 #H @(cpms_ind_dx … H) -V2
+lemma cpms_ell (h) (n) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[h,n] V2 →
+                            ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓛV1❫ ⊢ #0 ➡*[h,↑n] W2.
+#h #n #G #K #V1 #V2 #H @(cpms_ind_dx … H) -V2
 [ /3 width=3 by cpm_cpms, cpm_ell/
 | #n1 #n2 #V #V2 #_ #IH #HV2 #W2 #HVW2
   elim (lifts_total V (𝐔❨1❩)) #W #HVW >plus_S1
@@ -66,9 +66,9 @@ lemma cpms_ell (n) (h) (G): ∀K,V1,V2. ❪G,K❫ ⊢ V1 ➡*[n,h] V2 →
 ]
 qed.
 
-lemma cpms_lref (n) (h) (I) (G): ∀K,T,i. ❪G,K❫ ⊢ #i ➡*[n,h] T →
-                                 ∀U. ⇧[1] T ≘ U → ❪G,K.ⓘ[I]❫ ⊢ #↑i ➡*[n,h] U.
-#n #h #I #G #K #T #i #H @(cpms_ind_dx … H) -T
+lemma cpms_lref (h) (n) (I) (G): ∀K,T,i. ❪G,K❫ ⊢ #i ➡*[h,n] T →
+                                 ∀U. ⇧[1] T ≘ U → ❪G,K.ⓘ[I]❫ ⊢ #↑i ➡*[h,n] U.
+#h #n #I #G #K #T #i #H @(cpms_ind_dx … H) -T
 [ /3 width=3 by cpm_cpms, cpm_lref/
 | #n1 #n2 #T #T2 #_ #IH #HT2 #U2 #HTU2
   elim (lifts_total T (𝐔❨1❩)) #U #TU
@@ -76,11 +76,11 @@ lemma cpms_lref (n) (h) (I) (G): ∀K,T,i. ❪G,K❫ ⊢ #i ➡*[n,h] T →
 ]
 qed.
 
-lemma cpms_cast_sn (n) (h) (G) (L):
-                   ∀U1,U2. ❪G,L❫ ⊢ U1 ➡*[n,h] U2 →
-                   ∀T1,T2. ❪G,L❫ ⊢ T1 ➡[n,h] T2 →
-                   ❪G,L❫ ⊢ ⓝU1.T1 ➡*[n,h] ⓝU2.T2.
-#n #h #G #L #U1 #U2 #H @(cpms_ind_sn … H) -U1 -n
+lemma cpms_cast_sn (h) (n) (G) (L):
+                   ∀U1,U2. ❪G,L❫ ⊢ U1 ➡*[h,n] U2 →
+                   ∀T1,T2. ❪G,L❫ ⊢ T1 ➡[h,n] T2 →
+                   ❪G,L❫ ⊢ ⓝU1.T1 ➡*[h,n] ⓝU2.T2.
+#h #n #G #L #U1 #U2 #H @(cpms_ind_sn … H) -U1 -n
 [ /3 width=3 by cpm_cpms, cpm_cast/
 | #n1 #n2 #U1 #U #HU1 #_ #IH #T1 #T2 #H
   elim (cpm_fwd_plus … H) -H #T #HT1 #HT2
@@ -90,11 +90,11 @@ qed.
 
 (* Note: apparently this was missing in basic_1 *)
 (* Basic_2A1: uses: cprs_delta *)
-lemma cpms_delta_drops (n) (h) (G):
+lemma cpms_delta_drops (h) (n) (G):
                        ∀L,K,V,i. ⇩[i] L ≘ K.ⓓV →
-                       ∀V2. ❪G,K❫ ⊢ V ➡*[n,h] V2 →
-                       ∀W2. ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡*[n,h] W2.
-#n #h #G #L #K #V #i #HLK #V2 #H @(cpms_ind_dx … H) -V2
+                       ∀V2. ❪G,K❫ ⊢ V ➡*[h,n] V2 →
+                       ∀W2. ⇧[↑i] V2 ≘ W2 → ❪G,L❫ ⊢ #i ➡*[h,n] W2.
+#h #n #G #L #K #V #i #HLK #V2 #H @(cpms_ind_dx … H) -V2
 [ /3 width=6 by cpm_cpms, cpm_delta_drops/
 | #n1 #n2 #V1 #V2 #_ #IH #HV12 #W2 #HVW2
   lapply (drops_isuni_fwd_drop2 … HLK) -HLK // #HLK
@@ -103,11 +103,11 @@ lemma cpms_delta_drops (n) (h) (G):
 ]
 qed.
 
-lemma cpms_ell_drops (n) (h) (G):
+lemma cpms_ell_drops (h) (n) (G):
                      ∀L,K,W,i. ⇩[i] L ≘ K.ⓛW →
-                     ∀W2. ❪G,K❫ ⊢ W ➡*[n,h] W2 →
-                     ∀V2. ⇧[↑i] W2 ≘ V2 → ❪G,L❫ ⊢ #i ➡*[↑n,h] V2.
-#n #h #G #L #K #W #i #HLK #W2 #H @(cpms_ind_dx … H) -W2
+                     ∀W2. ❪G,K❫ ⊢ W ➡*[h,n] W2 →
+                     ∀V2. ⇧[↑i] W2 ≘ V2 → ❪G,L❫ ⊢ #i ➡*[h,↑n] V2.
+#h #n #G #L #K #W #i #HLK #W2 #H @(cpms_ind_dx … H) -W2
 [ /3 width=6 by cpm_cpms, cpm_ell_drops/
 | #n1 #n2 #W1 #W2 #_ #IH #HW12 #V2 #HWV2
   lapply (drops_isuni_fwd_drop2 … HLK) -HLK // #HLK
@@ -118,14 +118,14 @@ qed.
 
 (* Advanced inversion lemmas ************************************************)
 
-lemma cpms_inv_lref1_drops (n) (h) (G):
-                           ∀L,T2,i. ❪G,L❫ ⊢ #i ➡*[n,h] T2 →
+lemma cpms_inv_lref1_drops (h) (n) (G):
+                           ∀L,T2,i. ❪G,L❫ ⊢ #i ➡*[h,n] T2 →
                            ∨∨ ∧∧ T2 = #i & n = 0
-                            | ∃∃K,V,V2. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡*[n,h] V2 &
+                            | ∃∃K,V,V2. ⇩[i] L ≘ K.ⓓV & ❪G,K❫ ⊢ V ➡*[h,n] V2 &
                                         ⇧[↑i] V2 ≘ T2
-                            | ∃∃m,K,V,V2. ⇩[i] L ≘ K.ⓛV & ❪G,K❫ ⊢ V ➡*[m,h] V2 &
+                            | ∃∃m,K,V,V2. ⇩[i] L ≘ K.ⓛV & ❪G,K❫ ⊢ V ➡*[h,m] V2 &
                                           ⇧[↑i] V2 ≘ T2 & n = ↑m.
-#n #h #G #L #T2 #i #H @(cpms_ind_dx … H) -T2
+#h #n #G #L #T2 #i #H @(cpms_ind_dx … H) -T2
 [ /3 width=1 by or3_intro0, conj/
 | #n1 #n2 #T #T2 #_ #IH #HT2 cases IH -IH *
   [ #H1 #H2 destruct
@@ -146,11 +146,11 @@ lemma cpms_inv_lref1_drops (n) (h) (G):
 ]
 qed-.
 
-lemma cpms_inv_delta_sn (n) (h) (G) (K) (V):
-      ∀T2. ❪G,K.ⓓV❫ ⊢ #0 ➡*[n,h] T2 →
+lemma cpms_inv_delta_sn (h) (n) (G) (K) (V):
+      ∀T2. ❪G,K.ⓓV❫ ⊢ #0 ➡*[h,n] T2 →
       ∨∨ ∧∧ T2 = #0 & n = 0
-       | ∃∃V2. ❪G,K❫ ⊢ V ➡*[n,h] V2 & ⇧[1] V2 ≘ T2.
-#n #h #G #K #V #T2 #H
+       | ∃∃V2. ❪G,K❫ ⊢ V ➡*[h,n] V2 & ⇧[1] V2 ≘ T2.
+#h #n #G #K #V #T2 #H
 elim (cpms_inv_lref1_drops … H) -H *
 [ /3 width=1 by or_introl, conj/
 | #Y #X #V2 #H #HV2 #HVT2
@@ -161,11 +161,11 @@ elim (cpms_inv_lref1_drops … H) -H *
 ]
 qed-.
 
-lemma cpms_inv_ell_sn (n) (h) (G) (K) (V):
-      ∀T2. ❪G,K.ⓛV❫ ⊢ #0 ➡*[n,h] T2 →
+lemma cpms_inv_ell_sn (h) (n) (G) (K) (V):
+      ∀T2. ❪G,K.ⓛV❫ ⊢ #0 ➡*[h,n] T2 →
       ∨∨ ∧∧ T2 = #0 & n = 0
-       | ∃∃m,V2. ❪G,K❫ ⊢ V ➡*[m,h] V2 & ⇧[1] V2 ≘ T2 & n = ↑m.
-#n #h #G #K #V #T2 #H
+       | ∃∃m,V2. ❪G,K❫ ⊢ V ➡*[h,m] V2 & ⇧[1] V2 ≘ T2 & n = ↑m.
+#h #n #G #K #V #T2 #H
 elim (cpms_inv_lref1_drops … H) -H *
 [ /3 width=1 by or_introl, conj/
 | #Y #X #V2 #H #HV2 #HVT2
@@ -176,11 +176,11 @@ elim (cpms_inv_lref1_drops … H) -H *
 ]
 qed-.
 
-lemma cpms_inv_lref_sn (n) (h) (G) (I) (K):
-      ∀U2,i. ❪G,K.ⓘ[I]❫ ⊢ #↑i ➡*[n,h] U2 →
+lemma cpms_inv_lref_sn (h) (n) (G) (I) (K):
+      ∀U2,i. ❪G,K.ⓘ[I]❫ ⊢ #↑i ➡*[h,n] U2 →
       ∨∨ ∧∧ U2 = #↑i & n = 0
-       | ∃∃T2. ❪G,K❫ ⊢ #i ➡*[n,h] T2 & ⇧[1] T2 ≘ U2.
-#n #h #G #I #K #U2 #i #H
+       | ∃∃T2. ❪G,K❫ ⊢ #i ➡*[h,n] T2 & ⇧[1] T2 ≘ U2.
+#h #n #G #I #K #U2 #i #H
 elim (cpms_inv_lref1_drops … H) -H *
 [ /3 width=1 by or_introl, conj/
 | #L #V #V2 #H #HV2 #HVU2
@@ -196,10 +196,10 @@ elim (cpms_inv_lref1_drops … H) -H *
 ]
 qed-.
 
-fact cpms_inv_succ_sn (n) (h) (G) (L):
-                      ∀T1,T2. ❪G,L❫ ⊢ T1 ➡*[↑n,h] T2 →
-                      ∃∃T. ❪G,L❫ ⊢ T1 ➡*[1,h] T & ❪G,L❫ ⊢ T ➡*[n,h] T2.
-#n #h #G #L #T1 #T2
+fact cpms_inv_succ_sn (h) (n) (G) (L):
+                      ∀T1,T2. ❪G,L❫ ⊢ T1 ➡*[h,↑n] T2 →
+                      ∃∃T. ❪G,L❫ ⊢ T1 ➡*[h,1] T & ❪G,L❫ ⊢ T ➡*[h,n] T2.
+#h #n #G #L #T1 #T2
 @(insert_eq_0 … (↑n)) #m #H
 @(cpms_ind_sn … H) -T1 -m
 [ #H destruct