]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/multiple/llpx_sn_llor.ma
- some renaming according to the written version of basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / multiple / llpx_sn_llor.ma
index 91d9e649e4e1acadaa15af9adc43616418b69483..235eedbdacb2c5987f445cd0d02d685ec6fbc17d 100644 (file)
@@ -21,9 +21,9 @@ include "basic_2/multiple/lleq_alt.ma".
 (* Inversion lemmas on pointwise union for local environments ****************)
 
 lemma llpx_sn_llor_fwd_sn: ∀R. (∀L. reflexive … (R L)) →
-                           ∀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
+                           ∀L1,L2,T,l. llpx_sn R l T L1 L2 →
+                           ∀L. L1 ⋓[T, l] L2 ≡ L → lpx_sn R L1 L.
+#R #HR #L1 #L2 #T #l #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