X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flex_tc.ma;h=d33476caed973a0c80250407bc2fea1cf6d87c1f;hb=990f97071a9939d47be16b36f6045d3b23f218e0;hp=bb3aec919cbef0bd88949fcbd4e8bd92abcecd18;hpb=42705ef31dd3513a998533e02b5f20fb38dd4fb2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma index bb3aec919..d33476cae 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma @@ -12,14 +12,24 @@ (* *) (**************************************************************************) +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". +alias symbol "subseteq" = "relation inclusion". + (* GENERIC EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS **************) (* Inversion lemmas with transitive closure *********************************) +(* Basic_2A1: was: lpx_sn_LTC_TC_lpx_sn *) +lemma lex_inv_ltc: ∀R. c_reflexive … R → + lex (LTC … R) ⊆ TC … (lex R). +#R #HR #L1 #L2 * +/5 width=11 by lexs_inv_tc_dx, lexs_co, ext2_inv_tc, ext2_refl, monotonic_TC, ex2_intro/ +qed-. + lemma s_rs_transitive_lex_inv_isid: ∀R. s_rs_transitive … R (λ_.lex R) → s_rs_transitive_isid cfull (cext2 R). #R #HR #f #Hf #L2 #T1 #T2 #H #L1 #HL12 @@ -30,3 +40,16 @@ elim (ext2_tc … H) -H @(HR … HV12) -HV12 /2 width=3 by ex2_intro/ (**) (* auto fails *) ] qed-. + +(* Properties with transitive closure ***************************************) + +(* Basic_2A1: was: TC_lpx_sn_inv_lpx_sn_LTC *) +lemma lex_ltc: ∀R. s_rs_transitive … R (λ_. lex R) → + TC … (lex R) ⊆ lex (LTC … R). +#R #HR #L1 #L2 #HL12 +lapply (monotonic_TC … (lexs cfull (cext2 R) 𝐈𝐝) … HL12) -HL12 +[ #L1 #L2 * /3 width=3 by lexs_eq_repl_fwd, eq_id_inv_isid/ +| /5 width=9 by s_rs_transitive_lex_inv_isid, lexs_tc_dx, lexs_co, ext2_tc, ex2_intro/ +] +qed-. +