]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpms_drops.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpms_drops.ma
index 568100b547bb6e75d569aa357547ea0f057c8675..894ece04c1f3e937f03201044a5465133b84d33c 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/drops_ltc.ma".
+include "static_2/relocation/drops_ltc.ma".
 include "basic_2/rt_transition/cpm_drops.ma".
 include "basic_2/rt_computation/cpms.ma".
 
 (* T-BOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************)
 
+(* Properties with generic slicing for local environments *******************)
+
+lemma cpms_lifts_sn: ∀n,h,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
+/2 width=6 by cpms_lifts_sn, lifts_mono/
+qed-.
+
+(* Inversion lemmas with generic slicing for local environments *************)
+
+(* 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).
+/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
+/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
+[ /3 width=3 by cpm_cpms, cpm_delta/
+| #n1 #n2 #V #V2 #_ #IH #HV2 #W2 #HVW2
+  elim (lifts_total V (𝐔❨1❩)) #W #HVW
+  /5 width=11 by cpms_step_dx, cpm_lifts_bi, drops_refl, drops_drop/
+]
+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
+[ /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
+  /5 width=11 by cpms_step_dx, cpm_lifts_bi, drops_refl, drops_drop/
+]
+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
+[ /3 width=3 by cpm_cpms, cpm_lref/
+| #n1 #n2 #T #T2 #_ #IH #HT2 #U2 #HTU2
+  elim (lifts_total T (𝐔❨1❩)) #U #TU
+  /5 width=11 by cpms_step_dx, cpm_lifts_bi, drops_refl, drops_drop/
+]
+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
+[ /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
+  /3 width=3 by cpms_step_sn, cpm_cast/
+]
+qed.
+
 (* Note: apparently this was missing in basic_1 *)
 (* Basic_2A1: uses: cprs_delta *)
 lemma cpms_delta_drops (n) (h) (G):
-                       â\88\80L,K,V,i. â¬\87*[i] L ≘ K.ⓓV →
-                       â\88\80V2. â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡*[n, h] V2 →
-                       â\88\80W2. â¬\86*[â\86\91i] V2 â\89\98 W2 â\86\92 â¦\83G, Lâ¦\84 â\8a¢ #i â\9e¡*[n, h] W2.
+                       â\88\80L,K,V,i. â\87©*[i] L ≘ K.ⓓV →
+                       â\88\80V2. â\9dªG,Kâ\9d« â\8a¢ V â\9e¡*[n,h] V2 →
+                       â\88\80W2. â\87§*[â\86\91i] V2 â\89\98 W2 â\86\92 â\9dªG,Lâ\9d« â\8a¢ #i â\9e¡*[n,h] W2.
 #n #h #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
-  elim (lifts_total V1 (𝐔❴↑i❵)) #W1 #HVW1
+  elim (lifts_total V1 (𝐔❨↑i❩)) #W1 #HVW1
+  /3 width=11 by cpm_lifts_bi, cpms_step_dx/
+]
+qed.
+
+lemma cpms_ell_drops (n) (h) (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
+[ /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
+  elim (lifts_total W1 (𝐔❨↑i❩)) #V1 #HWV1 >plus_S1
   /3 width=11 by cpm_lifts_bi, cpms_step_dx/
 ]
 qed.
@@ -38,12 +119,12 @@ qed.
 (* Advanced inversion lemmas ************************************************)
 
 lemma cpms_inv_lref1_drops (n) (h) (G):
-                           â\88\80L,T2,i. â¦\83G, Lâ¦\84 â\8a¢ #i â\9e¡*[n, h] T2 →
+                           â\88\80L,T2,i. â\9dªG,Lâ\9d« â\8a¢ #i â\9e¡*[n,h] T2 →
                            ∨∨ ∧∧ T2 = #i & n = 0
-                            | â\88\83â\88\83K,V,V2. â¬\87*[i] L â\89\98 K.â\93\93V & â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡*[n, h] V2 &
-                                        â¬\86*[↑i] V2 ≘ T2
-                            | â\88\83â\88\83m,K,V,V2. â¬\87*[i] L â\89\98 K.â\93\9bV & â¦\83G, Kâ¦\84 â\8a¢ V â\9e¡*[m, h] V2 &
-                                          â¬\86*[↑i] V2 ≘ T2 & n = ↑m.
+                            | â\88\83â\88\83K,V,V2. â\87©*[i] L â\89\98 K.â\93\93V & â\9dªG,Kâ\9d« â\8a¢ V â\9e¡*[n,h] V2 &
+                                        â\87§*[↑i] V2 ≘ T2
+                            | â\88\83â\88\83m,K,V,V2. â\87©*[i] L â\89\98 K.â\93\9bV & â\9dªG,Kâ\9d« â\8a¢ V â\9e¡*[m,h] V2 &
+                                          â\87§*[↑i] V2 ≘ T2 & n = ↑m.
 #n #h #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 *
@@ -65,28 +146,71 @@ lemma cpms_inv_lref1_drops (n) (h) (G):
 ]
 qed-.
 
-(* Properties with generic slicing for local environments *******************)
-
-lemma cpms_lifts_sn: ∀n,h,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
-/2 width=6 by cpms_lifts_sn, lifts_mono/
+lemma cpms_inv_delta_sn (n) (h) (G) (K) (V):
+      ∀T2. ❪G,K.ⓓV❫ ⊢ #0 ➡*[n,h] T2 →
+      ∨∨ ∧∧ T2 = #0 & n = 0
+       | ∃∃V2. ❪G,K❫ ⊢ V ➡*[n,h] V2 & ⇧*[1] V2 ≘ T2.
+#n #h #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
+  lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+  /3 width=3 by ex2_intro, or_intror/
+| #m #Y #X #V2 #H #HV2 #HVT2
+  lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+]
 qed-.
 
-(* Inversion lemmas with generic slicing for local environments *************)
+lemma cpms_inv_ell_sn (n) (h) (G) (K) (V):
+      ∀T2. ❪G,K.ⓛV❫ ⊢ #0 ➡*[n,h] 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
+elim (cpms_inv_lref1_drops … H) -H *
+[ /3 width=1 by or_introl, conj/
+| #Y #X #V2 #H #HV2 #HVT2
+  lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+| #m #Y #X #V2 #H #HV2 #HVT2 #H0 destruct
+  lapply (drops_fwd_isid … H ?) -H [ // ] #H destruct
+  /3 width=5 by ex3_2_intro, or_intror/
+]
+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).
-/3 width=6 by d2_deliftable_sn_ltc, cpm_inv_lifts_sn/ qed-.
+lemma cpms_inv_lref_sn (n) (h) (G) (I) (K):
+      ∀U2,i. ❪G,K.ⓘ[I]❫ ⊢ #↑i ➡*[n,h] U2 →
+      ∨∨ ∧∧ U2 = #↑i & n = 0
+       | ∃∃T2. ❪G,K❫ ⊢ #i ➡*[n,h] T2 & ⇧*[1] T2 ≘ U2.
+#n #h #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
+  lapply (drops_inv_drop1 … H) -H #HLK
+  elim (lifts_split_trans … HVU2 (𝐔❨↑i❩) (𝐔❨1❩)) -HVU2
+  [| // ] #T2 #HVT2 #HTU2
+  /4 width=6 by cpms_delta_drops, ex2_intro, or_intror/
+| #m #L #V #V2 #H #HV2 #HVU2 #H0 destruct
+  lapply (drops_inv_drop1 … H) -H #HLK
+  elim (lifts_split_trans … HVU2 (𝐔❨↑i❩) (𝐔❨1❩)) -HVU2
+  [| // ] #T2 #HVT2 #HTU2
+  /4 width=6 by cpms_ell_drops, ex2_intro, or_intror/
+]
+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
-/2 width=6 by cpms_inv_lifts_sn, lifts_inj/
+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
+@(insert_eq_0 … (↑n)) #m #H
+@(cpms_ind_sn … H) -T1 -m
+[ #H destruct
+| #x1 #n2 #T1 #T #HT1 #HT2 #IH #H
+  elim (plus_inv_S3_sn x1 n2) [1,2: * |*: // ] -H
+  [ #H1 #H2 destruct -HT2
+    elim IH -IH // #T0 #HT0 #HT02
+    /3 width=3 by cpms_step_sn, ex2_intro/
+  | #n1 #H1 #H2 destruct -IH
+    elim (cpm_fwd_plus … 1 n1 T1 T) [|*: // ] -HT1 #T0 #HT10 #HT0
+    /3 width=5 by cpms_step_sn, cpm_cpms, ex2_intro/
+  ]
+]
 qed-.