]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda-delta/Basic-2/reduction/ltpr.ma
- the theory of parallel substitution of local environments (ltps) is ready
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / reduction / ltpr.ma
index 902bc96dabb5dc02bd3a7ec6a1cc8d939ca92fb2..97ef4901ac2ff98ebef5241c3c61dae21db31020 100644 (file)
@@ -26,9 +26,26 @@ interpretation
   "context-free parallel reduction (environment)"
   'PRed L1 L2 = (ltpr L1 L2).
 
+(* Basic properties *********************************************************)
+
+lemma ltpr_refl: ∀L:lenv. L ⇒ L.
+#L elim L -L /2/
+qed.
+
 (* Basic inversion lemmas ***************************************************)
 
-fact ltpr_inv_item1_aux: ∀L1,L2. L1 ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1 →
+fact ltpr_inv_atom1_aux: ∀L1,L2. L1 ⇒ L2 → L1 = ⋆ → L2 = ⋆.
+#L1 #L2 * -L1 L2
+[ //
+| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
+]
+qed.
+
+(* Basic-1: was: wcpr0_gen_sort *)
+lemma ltpr_inv_atom1: ∀L2. ⋆ ⇒ L2 → L2 = ⋆.
+/2/ qed.
+
+fact ltpr_inv_pair1_aux: ∀L1,L2. L1 ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1 →
                          ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
 #L1 #L2 * -L1 L2
 [ #K1 #I #V1 #H destruct
@@ -36,6 +53,31 @@ fact ltpr_inv_item1_aux: ∀L1,L2. L1 ⇒ L2 → ∀K1,I,V1. L1 = K1. 𝕓{I} V1
 ]
 qed.
 
-lemma ltpr_inv_item1: ∀K1,I,V1,L2. K1. 𝕓{I} V1 ⇒ L2 →
+(* Basic-1: was: wcpr0_gen_head *)
+lemma ltpr_inv_pair1: ∀K1,I,V1,L2. K1. 𝕓{I} V1 ⇒ L2 →
                       ∃∃K2,V2. K1 ⇒ K2 & V1 ⇒ V2 & L2 = K2. 𝕓{I} V2.
 /2/ qed.
+
+fact ltpr_inv_atom2_aux: ∀L1,L2. L1 ⇒ L2 → L2 = ⋆ → L1 = ⋆.
+#L1 #L2 * -L1 L2
+[ //
+| #K1 #K2 #I #V1 #V2 #_ #_ #H destruct
+]
+qed.
+
+lemma ltpr_inv_atom2: ∀L1. L1 ⇒ ⋆ → L1 = ⋆.
+/2/ qed.
+
+fact ltpr_inv_pair2_aux: ∀L1,L2. L1 ⇒ L2 → ∀K2,I,V2. L2 = K2. 𝕓{I} V2 →
+                         ∃∃K1,V1. K1 ⇒ K2 & V1 ⇒ V2 & L1 = K1. 𝕓{I} V1.
+#L1 #L2 * -L1 L2
+[ #K2 #I #V2 #H destruct
+| #K1 #K2 #I #V1 #V2 #HK12 #HV12 #K #J #W #H destruct -K2 I V2 /2 width=5/
+]
+qed.
+
+lemma ltpr_inv_pair2: ∀L1,K2,I,V2. L1 ⇒ K2. 𝕓{I} V2 →
+                      ∃∃K1,V1. K1 ⇒ K2 & V1 ⇒ V2 & L1 = K1. 𝕓{I} V1.
+/2/ qed.
+
+(* Basic-1: removed theorems 2: wcpr0_getl wcpr0_getl_back *)