]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr_lift.ma
- nDestructTac: Sys.break handled in two places
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / cpr_lift.ma
index fd4940b7cd38257687f92a7a30152ca8154130c3..dedb174e6e4b72c3dafa91e0fea2c14c039f3407 100644 (file)
@@ -21,20 +21,29 @@ include "basic_2/reducibility/cpr.ma".
 (* Advanced properties ******************************************************)
 
 lemma cpr_cdelta: ∀L,K,V1,W1,W2,i.
-                  ⇩[0, i] L ≡ K. ⓓV1 → K ⊢ V1 [0, |L| - i - 1] ▶* W1 →
+                  ⇩[0, i] L ≡ K. ⓓ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
 lapply (ldrop_fwd_ldrop2_length … HLK) #Hi
 @ex2_1_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 → L ⊢ ⓛV1. T1 ➡ ⓛV2. T2.
+#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02
+lapply (tpss_inv_S2 … HT02 L V ?) -HT02 // #HT02
+@(ex2_1_intro … (ⓛV0.T0)) /2 width=1/ -V1 -T1 (**) (* explicit constructors *)
+@tpss_bind // -V0
+@(tpss_lsubs_conf (L.ⓛV)) // -T0 -T2 /2 width=1/
+qed.
+
 (* Advanced inversion lemmas ************************************************)
 
 (* Basic_1: was: pr2_gen_lref *)
 lemma cpr_inv_lref1: ∀L,T2,i. L ⊢ #i ➡ T2 →
                      T2 = #i ∨
                      ∃∃K,V1,T1. ⇩[0, i] L ≡ K. ⓓV1 &
-                                K ⊢ V1 [0, |L| - i - 1] ▶* T1 &
+                                K ⊢ V1 ▶* [0, |L| - i - 1] T1 &
                                 ⇧[0, i + 1] T1 ≡ T2 &
                                 i < |L|.
 #L #T2 #i * #X #H
@@ -79,7 +88,7 @@ elim (tpr_inv_appl1 … H1) -H1 *
 qed-.
 
 (* Note: the main property of simple terms *)
-lemma cpr_inv_appl1_simple: ∀L,V1,T1,U. L ⊢ ⓐV1. T1 ➡ U → 𝐒[T1] →
+lemma cpr_inv_appl1_simple: ∀L,V1,T1,U. L ⊢ ⓐV1. T1 ➡ U → 𝐒⦃T1⦄ →
                             ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
                                      U = ⓐV2. T2.
 #L #V1 #T1 #U #H #HT1