]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reducibility/cpr_lift.ma
- lambdadelta: third recursive part of preservation finally proved!
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reducibility / cpr_lift.ma
index 9ff04b705c19e1889ce205560684160418915796..b59a3e6ed02628a85c99f82ed8c23e0e3cea5709 100644 (file)
@@ -28,12 +28,12 @@ lapply (ldrop_fwd_ldrop2_length … HLK) #Hi
 @ex2_intro [2: // | skip | @tpss_subst /width=6/ ] (**) (* /3 width=6/ is too slow *)
 qed.
 
-lemma cpr_abst: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀V,T1,T2.
-                L.ⓛV ⊢ T1 ➡ T2 → ∀a. L ⊢ ⓛ{a}V1. T1 ➡ ⓛ{a}V2. T2.
-#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02 #a
+lemma cpr_abst: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀V,T1,T2. L.ⓛV ⊢ T1 ➡ T2 →
+                          ∀a,I. L ⊢ ⓑ{a,I}V1. T1 ➡ ⓑ{a,I}V2. T2.
+#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02 #a #I
 lapply (tpss_inv_S2 … HT02 L V ?) -HT02 // #HT02
-lapply (tpss_lsubs_trans â\80¦ HT02 (L.â\93\9bV2) ?) -HT02 /2 width=1/ #HT02
-@(ex2_intro â\80¦ (â\93\9b{a}V0.T0)) /2 width=1/ (* explicit constructors *)
+lapply (tpss_lsubs_trans â\80¦ HT02 (L.â\93\91{I}V2) ?) -HT02 /2 width=1/ #HT02
+@(ex2_intro â\80¦ (â\93\91{a,I}V0.T0)) /2 width=1/ (* explicit constructors *)
 qed.
 
 lemma cpr_beta: ∀a,L,V1,V2,W,T1,T2.
@@ -134,7 +134,18 @@ elim (cpr_inv_appl1 … H) -H *
   elim (simple_inv_bind … HT1)
 ]
 qed-.
-     
+
+(* Advanced forward lemmas **************************************************)
+
+lemma cpr_fwd_abst1: ∀a,L,V1,T1,U2. L ⊢ ⓛ{a}V1.T1 ➡ U2 → ∀b,I,W.
+                     ∃∃V2,T2. L ⊢ ⓑ{b,I}W.T1 ➡ ⓑ{b,I}W.T2 &
+                              U2 = ⓛ{a}V2.T2.
+#a #L #V1 #T1 #U2 * #U #H #HU2 #b #I #W
+elim (tpr_fwd_abst1 … H b I W) -H #V #T #HT1 #H destruct
+elim (tpss_inv_bind1 … HU2) -HU2 #V2 #T2 #_ #HT2
+lapply (tpss_lsubs_trans … HT2 (L.ⓑ{I}W) ?) -HT2 /2 width=1/ /4 width=5/
+qed-.
+
 (* Relocation properties ****************************************************)
 
 (* Basic_1: was: pr2_lift *)
@@ -157,10 +168,13 @@ lemma cpr_inv_lift1: ∀L,K,d,e. ⇩[d, e] L ≡ K →
 #L #K #d #e #HLK #T1 #U1 #HTU1 #U2 * #U #HU1 #HU2
 elim (tpr_inv_lift1 … HU1 … HTU1) -U1 #T #HTU #T1
 elim (lt_or_ge (|L|) d) #HLd
-[ elim (tpss_inv_lift1_le … HU2 … HLK … HTU ?) -U -HLK [ /5 width=4/ | /2 width=2/ ]
+[ elim (tpss_inv_lift1_le … HU2 … HLK … HTU ?) -U -HLK /2 width=2/
+  /3 width=7 by ex2_intro, cpr_intro/
 | elim (lt_or_ge (|L|) (d + e)) #HLde
-  [ elim (tpss_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -U -HLK // [ /5 width=4/ | /2 width=2/ ] 
-  | elim (tpss_inv_lift1_be … HU2 … HLK … HTU ? ?) -U -HLK // /5 width=4/
+  [ elim (tpss_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -U -HLK // /2 width=2/
+    /3 width=7 by ex2_intro, cpr_intro/
+  | elim (tpss_inv_lift1_be … HU2 … HLK … HTU ? ?) -U -HLK //
+    /3 width=7 by ex2_intro, cpr_intro/
   ]
 ]
 qed.