X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fstatic%2Frex_lex.ma;h=9329d7f84c7704c87bcc840b16f430ccbd3d9d8c;hb=3c7b4071a9ac096b02334c1d47468776b948e2de;hp=d031d93c40aefb5dc4d595e70cc67f99c3a764e6;hpb=f308429a0fde273605a2330efc63268b4ac36c99;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/static_2/static/rex_lex.ma b/matita/matita/contribs/lambdadelta/static_2/static/rex_lex.ma index d031d93c4..9329d7f84 100644 --- a/matita/matita/contribs/lambdadelta/static_2/static/rex_lex.ma +++ b/matita/matita/contribs/lambdadelta/static_2/static/rex_lex.ma @@ -20,7 +20,8 @@ include "static_2/static/req.ma". (* Properties with generic extension of a context-sensitive relation ********) -lemma rex_lex: ∀R,L1,L2. L1 ⪤[R] L2 → ∀T. L1 ⪤[R,T] L2. +lemma rex_lex (R): + ∀L1,L2. L1 ⪤[R] L2 → ∀T. L1 ⪤[R,T] L2. #R #L1 #L2 * #f #Hf #HL12 #T elim (frees_total L1 T) #g #Hg /4 width=5 by sex_sdj, sdj_isid_sn, ex2_intro/ @@ -28,14 +29,14 @@ qed. (* Inversion lemmas with generic extension of a context sensitive relation **) -lemma rex_inv_lex_req: ∀R. c_reflexive … R → - rex_fsge_compatible R → - ∀L1,L2,T. L1 ⪤[R,T] L2 → - ∃∃L. L1 ⪤[R] L & L ≡[T] L2. +lemma rex_inv_lex_req (R): + c_reflexive … R → rex_fsge_compatible R → + ∀L1,L2,T. L1 ⪤[R,T] L2 → + ∃∃L. L1 ⪤[R] L & L ≡[T] L2. #R #H1R #H2R #L1 #L2 #T * #f1 #Hf1 #HL elim (sex_sdj_split … ceq_ext … HL 𝐈𝐝 ?) -HL [ #L0 #HL10 #HL02 |*: /2 width=1 by ext2_refl, sdj_isid_dx/ ] -H1R lapply (sex_sdj … HL10 f1 ?) /2 width=1 by sdj_isid_sn/ #H -elim (frees_sex_conf … Hf1 … H) // -H2R -H #f0 #Hf0 #Hf01 +elim (frees_sex_conf_fsge … Hf1 … H) // -H2R -H #f0 #Hf0 #Hf01 /4 width=7 by sle_sex_trans, (* 2x *) ex2_intro/ qed-.