]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / cpr.ma
index 072d951babf91b355a7005f522192e6f60afb2b2..22fbfc148440b49c722985a6ae9a518598942845 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/cl_shift.ma".
 include "basic_2/unfold/tpss.ma".
 include "basic_2/reducibility/tpr.ma".
 
@@ -76,26 +75,6 @@ lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k.
 >(tpss_inv_sort1 … H) -H //
 qed-.
 
-(* Basic_1: was pr2_gen_abbr *)
-lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
-                     (∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 &
-                                 L. ⓓV ⊢ T1 ➡ T2 &
-                                 U2 = ⓓV2. T2
-                      ) ∨
-                      ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2.
-#L #V1 #T1 #Y * #X #H1 #H2
-elim (tpr_inv_abbr1 … H1) -H1 *
-[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct
-  elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
-  lapply (tps_lsubs_trans … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0
-  lapply (tps_weak_all … HT0) -HT0 #HT0
-  lapply (tpss_lsubs_trans … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2
-  lapply (tpss_weak_all … HT2) -HT2 #HT2
-  lapply (tpss_strap … HT0 HT2) -T /4 width=7/
-| /4 width=5/
-]
-qed-.
-
 (* Basic_1: was: pr2_gen_cast *)
 lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓝV1. T1 ➡ U2 → (
                         ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
@@ -107,6 +86,15 @@ elim (tpr_inv_cast1 … H) -H /3 width=3/
 elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/
 qed-.
 
+(* Basic forward lemmas *****************************************************)
+
+lemma cpr_fwd_shift1: ∀L,L1,T1,T. L ⊢ L1 @@ T1 ➡ T →
+                      ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2.
+#L #L1 #T1 #T * #X #H1 #H2
+elim (tpr_fwd_shift1 … H1) -H1 #L0 #T0 #HL10 #H destruct
+elim (tpss_fwd_shift1 … H2) -H2 #L2 #T2 #HL02 #H destruct /2 width=4/
+qed-.
+
 (* Basic_1: removed theorems 6: 
             pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
             pr2_gen_ctail pr2_ctail