]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr_lift.ma
component "reducibility" updated to new syntax!
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reducibility / cpr_lift.ma
index 4c2101504a9a3eab9bc4c7e1d62a6035861747ac..0775cefd22684717b9b4969dddc0f9eaff60f3fa 100644 (file)
@@ -24,7 +24,8 @@ lemma cpr_cdelta: ∀L,K,V1,W1,W2,i.
                   ↓[0, i] L ≡ K. 𝕓{Abbr} V1 → K ⊢ V1 [0, |L| - i - 1] ≫* W1 →
                   ↑[0, i + 1] W1 ≡ W2 → L ⊢ #i ⇒ W2.
 #L #K #V1 #W1 #W2 #i #HLK #HVW1 #HW12
-@ex2_1_intro [2: // | skip | @tpss_subst /2 width=6/ ] (**) (* /4 width=6/ is too slow *)
+lapply (ldrop_fwd_ldrop2_length … HLK) #Hi
+@ex2_1_intro [2: // | skip | @tpss_subst /width=6/ ] (**) (* /3 width=6/ is too slow *)
 qed.
 
 (* Advanced inversion lemmas ************************************************)
@@ -38,14 +39,14 @@ lemma cpr_inv_lref1: ∀L,T2,i. L ⊢ #i ⇒ T2 →
                                 i < |L|.
 #L #T2 #i * #X #H
 >(tpr_inv_atom1 … H) -H #H
-elim (tpss_inv_lref1 … H) -H /2/
+elim (tpss_inv_lref1 … H) -H /2 width=1/
 * /3 width=6/
-qed.
+qed-.
 
 (* Basic_1: was: pr2_gen_abst *)
 lemma cpr_inv_abst1: ∀V1,T1,U2. 𝕔{Abst} V1. T1 ⇒ U2 →
                      ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕔{Abst} V2. T2.
-/2/ qed.
+/2 width=3/ qed-.
 
 (* Relocation properties ****************************************************)
 
@@ -57,8 +58,8 @@ lemma cpr_lift: ∀L,K,d,e. ↓[d, e] L ≡ K →
 elim (lift_total T d e) #U #HTU 
 lapply (tpr_lift … HT1 … HTU1 … HTU) -T1 #HU1
 elim (lt_or_ge (|K|) d) #HKd
-[ lapply (tpss_lift_le … HT2 … HLK HTU … HTU2) -T2 T HLK [ /2/ | /3/ ]
-| lapply (tpss_lift_be … HT2 … HLK HTU … HTU2) -T2 T HLK // /3/
+[ lapply (tpss_lift_le … HT2 … HLK HTU … HTU2) -T2 -T -HLK [ /2 width=1/ | /3 width=4/ ]
+| lapply (tpss_lift_be … HT2 … HLK HTU … HTU2) -T2 -T -HLK // /3 width=4/
 ]
 qed.
 
@@ -69,10 +70,10 @@ lemma cpr_inv_lift: ∀L,K,d,e. ↓[d, e] L ≡ K →
 #L #K #d #e #HLK #T1 #U1 #HTU1 #U2 * #U #HU1 #HU2
 elim (tpr_inv_lift … HU1 … HTU1) -U1 #T #HTU #T1
 elim (lt_or_ge (|L|) d) #HLd
-[ elim (tpss_inv_lift1_le … HU2 … HLK … HTU ?) -U HLK [ /5/ | /2/ ]
+[ elim (tpss_inv_lift1_le … HU2 … HLK … HTU ?) -U -HLK [ /5 width=4/ | /2 width=1/ ]
 | elim (lt_or_ge (|L|) (d + e)) #HLde
-  [ elim (tpss_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -U HLK // [ /5/ | /2/ ] 
-  | elim (tpss_inv_lift1_be … HU2 … HLK … HTU ? ?) -U HLK // /5/
+  [ elim (tpss_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -U -HLK // [ /5 width=4/ | /2 width=1/ ] 
+  | elim (tpss_inv_lift1_be … HU2 … HLK … HTU ? ?) -U -HLK // /5 width=4/
   ]
 ]
 qed.