]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/dxprs_dxprs.ma
- lambdadelta: last recursive part of preservation finally proved!
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / dxprs_dxprs.ma
index c0aff4efb56988fc79a7cb03354603d63d2d20a2..9d051adaa4e6de952277bb71aeb75ec564375419 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/unwind/sstas_sstas.ma".
-include "basic_2/computation/cprs_cprs.ma".
+include "basic_2/computation/cprs_lfprs.ma".
 include "basic_2/computation/dxprs.ma".
 
 (* DECOMPOSED EXTENDED PARALLEL COMPUTATION ON TERMS ************************)
@@ -31,3 +31,15 @@ lemma sstas_dxprs_trans: ∀h,g,L,T1,T,T2.
 #h #g #L #T1 #T #T2 #HT1 * #T0 #HT0 #HT02
 lapply (sstas_trans … HT1 … HT0) -T /2 width=3/
 qed-.
+
+(* Advanced inversion lemmas ************************************************)
+
+lemma dxprs_inv_abbr_abst: ∀h,g,a1,a2,L,V1,W2,T1,T2. ⦃h, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[g] ⓛ{a2}W2.T2 →
+                           ∃∃T. ⦃h, L.ⓓV1⦄ ⊢ T1 •*➡*[g] T & ⇧[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
+#h #g #a1 #a2 #L #V1 #W2 #T1 #T2 * #X #H1 #H2
+elim (sstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
+elim (cprs_inv_abbr1 … H2) -H2 *
+[ #V2 #U2 #HV12 #HU12 #H destruct
+| /3 width=3/
+]
+qed-.