]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma
long awaited update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lex_tc.ma
index d33476caed973a0c80250407bc2fea1cf6d87c1f..d9112a12e4461c3981a6ef0c44419b739bb52234 100644 (file)
@@ -12,7 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/relocation/rtmap_id.ma".
 include "basic_2/syntax/ext2_tc.ma".
 include "basic_2/relocation/lexs_tc.ma".
 include "basic_2/relocation/lex.ma".
@@ -53,3 +52,12 @@ lapply (monotonic_TC … (lexs cfull (cext2 R) 𝐈𝐝) … HL12) -HL12
 ]
 qed-.
 
+lemma lex_ltc_step_dx: ∀R. c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                       ∀L1,L. lex (LTC … R) L1 L →
+                       ∀L2. lex R L L2 → lex (LTC … R) L1 L2.
+/4 width=3 by lex_ltc, lex_inv_ltc, step/ qed-.
+
+lemma lex_ltc_step_sn: ∀R. c_reflexive … R → s_rs_transitive … R (λ_. lex R) →
+                       ∀L1,L. lex R L1 L →
+                       ∀L2. lex (LTC … R) L L2 → lex (LTC … R) L1 L2.
+/4 width=3 by lex_ltc, lex_inv_ltc, TC_strap/ qed-.