]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_lpx_sn.ma
- theory of llor now includes (long awaited) non-recursive alternative definition
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / llpx_sn_lpx_sn.ma
index 7486958ae806c6830e6924c800287f30c2b34f49..6a0a210ca523b23b41f207c87c8f158c5db4708d 100644 (file)
@@ -19,7 +19,7 @@ include "basic_2/substitution/llpx_sn.ma".
 
 (* Properties on pointwise extensions ***************************************)
 
-lemma lpx_sn_llpx_sn: ∀R. (∀I,L. reflexive … (R I L)) →
+lemma lpx_sn_llpx_sn: ∀R. (∀L. reflexive … (R L)) →
                       ∀T,L1,L2,d. lpx_sn R L1 L2 → llpx_sn R d T L1 L2.
 #R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T
 #n #IH #L1 * *
@@ -28,7 +28,7 @@ lemma lpx_sn_llpx_sn: ∀R. (∀I,L. reflexive … (R I L)) →
   [2: -IH /4 width=4 by lpx_sn_fwd_length, llpx_sn_free, le_repl_sn_conf_aux/ ]
   #Hi #Hn #L2 #d elim (ylt_split i d) 
   [ -n /3 width=2 by llpx_sn_skip, lpx_sn_fwd_length/ ]
-  #Hdi #HL12 elim (ldrop_O1_lt L1 i) //
+  #Hdi #HL12 elim (ldrop_O1_lt (Ⓕ) L1 i) //
   #I #K1 #V1 #HLK1 elim (lpx_sn_ldrop_conf … HL12 … HLK1) -HL12
   /4 width=9 by llpx_sn_lref, ldrop_fwd_rfw/
 | -HR -IH /4 width=2 by lpx_sn_fwd_length, llpx_sn_gref/