]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/unfold/lpqs_ldrop.ma
- basic_2: induction for preservation results now uses supclosure
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / unfold / lpqs_ldrop.ma
index bea4a91341d61652546faed7a80ecf1f566ddf01..854420d60aa390cb4c3ea1ed24a97ddff917bdab 100644 (file)
@@ -12,6 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
+include "basic_2/relocation/fsup.ma".
 include "basic_2/relocation/ldrop_lpx_sn.ma".
 include "basic_2/unfold/cpqs_lift.ma".
 include "basic_2/unfold/lpqs.ma".
@@ -28,3 +29,15 @@ lemma ldrop_lpqs_trans: dedropable_sn lpqs.
 
 lemma lpqs_ldrop_trans_O1: dropable_dx lpqs.
 /2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on context-sensitive rest. parallel computation for terms *****)
+
+lemma fsup_cpqs_trans: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ➤* U2 →
+                       ∃∃L,U1. L1 ⊢ ➤* L & L ⊢ T1 ➤* U1 & ⦃L, U1⦄ ⊃ ⦃L2, U2⦄.
+#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [2: * ] [1,2,3,4,5: /3 width=5/ ]
+#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2
+elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2
+elim (lift_total T d e) #U #HTU
+elim (ldrop_lpqs_trans … HLK1 … HK1) -HLK1 -HK1 #L2 #HL12 #HL2K
+lapply (cpqs_lift … HT1 … HL2K … HTU1 … HTU) -HT1 -HTU1 /3 width=11/
+qed-.