]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/tpr.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / tpr.ma
index fbe5f3dcc2c849ac6276ac12851a3621cb0d860f..957b5f3a89926458ddf3feaa3371b7bc68766e0d 100644 (file)
@@ -43,7 +43,7 @@ lemma tpr_bind: ∀a,I,V1,V2,T1,T2. V1 ➡ V2 → T1 ➡ T2 → ⓑ{a,I} V1. T1
 /2 width=3/ qed.
 
 (* Basic_1: was by definition: pr0_refl *)
-lemma tpr_refl: ∀T. T ➡ T.
+lemma tpr_refl: reflexive … tpr.
 #T elim T -T //
 #I elim I -I /2 width=1/
 qed.
@@ -203,6 +203,26 @@ lemma tpr_inv_lref2: ∀T1,i. T1 ➡ #i →
                       | ∃∃V,T. T ➡ #i & T1 = ⓝV. T.
 /2 width=3/ qed-.
 
+(* Basic forward lemmas *****************************************************)
+
+lemma tpr_fwd_shift1: ∀L1,T1,T. L1 @@ T1 ➡ T →
+                      ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2.
+#L1 @(lenv_ind_dx … L1) -L1 normalize
+[ #T1 #T #HT1
+  @(ex2_2_intro … (⋆)) // (**) (* explicit constructor *)
+| #I #L1 #V1 #IH #T1 #X
+  >shift_append_assoc normalize #H
+  elim (tpr_inv_bind1 … H) -H *
+  [ #V0 #T0 #X0 #_ #HT10 #H0 #H destruct
+    elim (IH … HT10) -IH -T1 #L #T #HL1 #H destruct
+    elim (tps_fwd_shift1 … H0) -T #L2 #T2 #HL2 #H destruct
+    >append_length >HL1 >HL2 -L1 -L
+    @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *)
+  | #T #_ #_ #H destruct
+  ]
+]
+qed-.
+
 (* Basic_1: removed theorems 3:
             pr0_subst0_back pr0_subst0_fwd pr0_subst0
    Basic_1: removed local theorems: 1: pr0_delta_tau