]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/fpr.ma
- we introduced the pointer_step rc in the perspective of proving
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / fpr.ma
index 46c201147ed4c67f3646b1160ae2f55d8b1cf71c..3910615656c51e32cd6174241a6e87aa52f60a02 100644 (file)
 
 include "basic_2/reducibility/tpr.ma".
 
-(* FOCALIZED PARALLEL REDUCTION ON CLOSURES *********************************)
+(* CONTEXT-FREE PARALLEL REDUCTION ON CLOSURES ******************************)
 
 definition fpr: bi_relation lenv term ≝
                 λL1,T1,L2,T2. |L1| = |L2| ∧ L1 @@ T1 ➡ L2 @@ T2.
 
 interpretation
-   "focalized parallel reduction (closure)"
+   "context-free parallel reduction (closure)"
    'FocalizedPRed L1 T1 L2 T2 = (fpr L1 T1 L2 T2).
 
 (* Basic properties *********************************************************)
@@ -28,6 +28,13 @@ interpretation
 lemma fpr_refl: bi_reflexive … fpr.
 /2 width=1/ qed.
 
+lemma fpr_shift: ∀I1,I2,L1,L2,V1,V2,T1,T2.
+                 ⦃L1, -ⓑ{I1}V1.T1⦄ ➡ ⦃L2, -ⓑ{I2}V2.T2⦄ →
+                 ⦃L1.ⓑ{I1}V1, T1⦄ ➡ ⦃L2.ⓑ{I2}V2, T2⦄.
+#I1 #I2 #L1 #L2 #V1 #V2 #T1 #T2 * #HL12 #HT12
+@conj // normalize // (**) (* explicit constructor *)
+qed.
+
 (* Basic inversion lemmas ***************************************************)
 
 lemma fpr_inv_atom1: ∀L2,T1,T2. ⦃⋆, T1⦄ ➡ ⦃L2, T2⦄ → T1 ➡ T2 ∧ L2 = ⋆.
@@ -35,23 +42,23 @@ lemma fpr_inv_atom1: ∀L2,T1,T2. ⦃⋆, T1⦄ ➡ ⦃L2, T2⦄ → T1 ➡ T2 
 lapply (length_inv_zero_sn … H) -H #H destruct /2 width=1/
 qed-.
 
-lemma fpr_inv_pair1: ∀I,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I}V1, T1⦄ ➡ ⦃L2, T2⦄ →
-                     ∃∃K2,V2. ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄  &
-                              L2 = K2.ⓑ{I}V2.
-#I1 #K1 #L2 #V1 #T1 #T2 * #H
-elim (length_inv_pos_sn … H) -H #I2 #K2 #V2 #HK12 #H destruct #H
-elim (tpr_fwd_shift_bind_minus … H) // #_ #H0 destruct /3 width=4/
-qed-.
-
 lemma fpr_inv_atom3: ∀L1,T1,T2. ⦃L1,T1⦄ ➡ ⦃⋆,T2⦄ → T1 ➡ T2 ∧ L1 = ⋆.
 #L1 #T1 #T2 * #H
 lapply (length_inv_zero_dx … H) -H #H destruct /2 width=1/
 qed-.
 
-lemma fpr_inv_pair3: ∀I,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I}V2, T2⦄ →
-                     ∃∃K1,V1. ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄  &
-                              L1 = K1.ⓑ{I}V1.
+(* Basic forward lemmas *****************************************************)
+
+lemma fpr_fwd_pair1: ∀I1,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I1}V1, T1⦄ ➡ ⦃L2, T2⦄ →
+                     ∃∃I2,K2,V2. ⦃K1, -ⓑ{I1}V1.T1⦄ ➡ ⦃K2, -ⓑ{I2}V2.T2⦄  &
+                                 L2 = K2.ⓑ{I2}V2.
+#I1 #K1 #L2 #V1 #T1 #T2 * #H
+elim (length_inv_pos_sn … H) -H #I2 #K2 #V2 #HK12 #H destruct /3 width=5/
+qed-.
+
+lemma fpr_fwd_pair3: ∀I2,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I2}V2, T2⦄ →
+                     ∃∃I1,K1,V1. ⦃K1, -ⓑ{I1}V1.T1⦄ ➡ ⦃K2, -ⓑ{I2}V2.T2⦄  &
+                                 L1 = K1.ⓑ{I1}V1.
 #I2 #L1 #K2 #V2 #T1 #T2 * #H
-elim (length_inv_pos_dx … H) -H #I1 #K1 #V1 #HK12 #H destruct #H
-elim (tpr_fwd_shift_bind_minus … H) // #_ #H0 destruct /3 width=4/
+elim (length_inv_pos_dx … H) -H #I1 #K1 #V1 #HK12 #H destruct /3 width=5/
 qed-.