]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_ssta.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / dynamic / lsubsv_ssta.ma
index 94db2292b6336b85310675172967670aa364f2d3..594c72662507e83af8f5efa25158f4d495ef1090 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/equivalence/lsubse_ssta.ma".
+include "basic_2/equivalence/lsubss_ssta.ma".
 include "basic_2/dynamic/lsubsv.ma".
 
 (* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************)
 
 (* Properties on stratified static type assignment **************************)
 
-lemma lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g,l] U2 →
-                         ∀L1. h ⊢ L1 ⊩:⊑[g] L2 →
-                         ∃∃U1. ⦃h, L1⦄ ⊢ T •[g,l] U1 & L1 ⊢ U1 ⬌* U2.
-/3 width=3 by lsubsv_fwd_lsubse, lsubse_ssta_trans/
+lemma lsubsv_ssta_trans: ∀h,g,L2,T,U2,l. ⦃h, L2⦄ ⊢ T •[g] ⦃l, U2⦄ →
+                         ∀L1. h ⊢ L1 ¡⊑[g] L2 →
+                         ∃∃U1. ⦃h, L1⦄ ⊢ T •[g] ⦃l, U1⦄ & L1 ⊢ U1 ⬌* U2.
+/3 width=3 by lsubsv_fwd_lsubss, lsubss_ssta_trans/
 qed-.