X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flex_tc.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Flex_tc.ma;h=bb3aec919cbef0bd88949fcbd4e8bd92abcecd18;hb=c44a7c4d35c1bb9651c3596519d8262e52e90ff4;hp=0000000000000000000000000000000000000000;hpb=d7c5846e4a362a366f5600d079e08f8a75b9d566;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 new file mode 100644 index 000000000..bb3aec919 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lex_tc.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/syntax/ext2_tc.ma". +include "basic_2/relocation/lexs_tc.ma". +include "basic_2/relocation/lex.ma". + +(* GENERIC EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS **************) + +(* Inversion lemmas with transitive closure *********************************) + +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 +elim (ext2_tc … H) -H +[ /3 width=1 by ext2_inv_tc, ext2_unit/ +| #I #V1 #V2 #HV12 + @ext2_inv_tc @ext2_pair + @(HR … HV12) -HV12 /2 width=3 by ex2_intro/ (**) (* auto fails *) +] +qed-.