]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/grammar/lenv_px.ma
- parallel reduction for local environments: we proved the equivalence
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / grammar / lenv_px.ma
index c3d05949ff289d3c268f4b968260a00d9276ac35..41710626a8cf70bf594baad61604c14f207cb9f7 100644 (file)
@@ -74,6 +74,32 @@ lemma lpx_fwd_length: ∀R,L1,L2. lpx R L1 L2 → |L1| = |L2|.
 #R #L1 #L2 #H elim H -L1 -L2 normalize //
 qed-.
 
+(* Advanced inversion lemmas ************************************************)
+
+lemma lpx_inv_append1: ∀R,L1,K1,L. lpx R (K1 @@ L1) L →
+                       ∃∃K2,L2. lpx R K1 K2 & lpx R L1 L2 & L = K2 @@ L2.
+#R #L1 elim L1 -L1 normalize
+[ #K1 #K2 #HK12
+  @(ex3_2_intro … K2 (⋆)) // (**) (* explicit constructor, /2 width=5/ does not work *) 
+| #L1 #I #V1 #IH #K1 #X #H
+  elim (lpx_inv_pair1 … H) -H #L #V2 #H1 #HV12 #H destruct
+  elim (IH … H1) -IH -H1 #K2 #L2 #HK12 #HL12 #H destruct
+  @(ex3_2_intro … HK12) [2: /2 width=2/ | skip | // ] (* explicit constructor, /3 width=5/ does not work *)
+]
+qed-.
+
+lemma lpx_inv_append2: ∀R,L2,K2,L. lpx R L (K2 @@ L2) →
+                       ∃∃K1,L1. lpx R K1 K2 & lpx R L1 L2 & L = K1 @@ L1.
+#R #L2 elim L2 -L2 normalize
+[ #K2 #K1 #HK12
+  @(ex3_2_intro … K1 (⋆)) // (**) (* explicit constructor, /2 width=5/ does not work *) 
+| #L2 #I #V2 #IH #K2 #X #H
+  elim (lpx_inv_pair2 … H) -H #L #V1 #H1 #HV12 #H destruct
+  elim (IH … H1) -IH -H1 #K1 #L1 #HK12 #HL12 #H destruct
+  @(ex3_2_intro … HK12) [2: /2 width=2/ | skip | // ] (* explicit constructor, /3 width=5/ does not work *)
+]
+qed-.   
+
 (* Basic properties *********************************************************)
 
 lemma lpx_refl: ∀R. reflexive ? R → reflexive … (lpx R).
@@ -125,7 +151,7 @@ lemma TC_lpx_pair_sn: ∀R. reflexive ? R →
 /4 width=5 by lpx_refl, lpx_pair, inj, step/ (**) (* too slow without trace *)
 qed.
 
-lemma lpx_TC: ∀R,L1,L2. TC … (lpx R) L1 L2 → lpx (TC … R) L1 L2. 
+lemma lpx_TC: ∀R,L1,L2. TC … (lpx R) L1 L2 → lpx (TC … R) L1 L2.
 #R #L1 #L2 #H elim H -L2 /2 width=1/ /2 width=3/
 qed.