]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma
- advances on free variables allow to reduce lleq_lpx_trans to llor_total :)
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / multiple / llpx_sn_llor.ma
index 4828b7b755cc18ed79c43d722fb5b3bdfcca1e8c..6f27ec9b088cf2ee37aa7f14027713307a754b5f 100644 (file)
@@ -21,14 +21,14 @@ include "basic_2/multiple/lleq_alt.ma".
 (* Inversion lemmas on poinwise union for local environments ****************)
 
 lemma llpx_sn_llor_fwd_sn: ∀R. (∀L. reflexive … (R L)) →
-                           ∀L1,L2,T. llpx_sn R 0 T L1 L2 →
-                           ∀L. L1 ⩖[T] L2 ≡ L → lpx_sn R L1 L.
-#R #HR #L1 #L2 #T #H1 #L #H2
+                           ∀L1,L2,T,d. llpx_sn R d T L1 L2 →
+                           ∀L. L1 ⩖[T, d] L2 ≡ L → lpx_sn R L1 L.
+#R #HR #L1 #L2 #T #d #H1 #L #H2
 elim (llpx_sn_llpx_sn_alt … H1) -H1 #HL12 #IH1
 elim H2 -H2 #_ #HL1 #IH2
 @lpx_sn_intro_alt // #I1 #I #K1 #K #V1 #V #i #HLK1 #HLK
 lapply (ldrop_fwd_length_lt2 … HLK) #HiL
 elim (ldrop_O1_lt (Ⓕ) L2 i) // -HiL -HL1 -HL12 #I2 #K2 #V2 #HLK2
-elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK * [ /2 width=1 by conj/ ]
+elim (IH2 … HLK1 HLK2 HLK) -IH2 -HLK * /2 width=1 by conj/
 #HnT #H1 #H2 elim (IH1 … HnT … HLK1 HLK2) -IH1 -HnT -HLK1 -HLK2 /2 width=1 by conj/
 qed-.