X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfpxs_etc.ma;h=d8989b799d2c7783797a1db008cd1c692d741ec1;hb=990f97071a9939d47be16b36f6045d3b23f218e0;hp=0fda037b18022749f623cb821d72d0f7316f7e8a;hpb=bc40346f09bcccb9a09560963ccb7157ebfad7ad;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma index 0fda037b1..d8989b799 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_etc.ma @@ -1,20 +1,18 @@ -include "basic_2/static/lfxs_lfxs.ma". -include "basic_2/rt_transition/lfpx_frees.ma". -include "basic_2/rt_computation/lfpxs_fqup.ma". +include "basic_2/static/lfxs_lex.ma". +include "basic_2/rt_transition/cpx_etc.ma". +include "basic_2/rt_computation/lfpxs_lpxs.ma". -axiom cpx_frees_conf_lfpxs: ∀h,G,L1,T1,T2. ⦃G, L1⦄ ⊢ T1 ⬈[h] T2 → - ∀f1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 → - ∀L2. ⦃G, L1⦄ ⊢ ⬈*[h, T1] L2 → - ∀g1. L2 ⊢ 𝐅*⦃T1⦄ ≡ g1 → - ∃∃g2. L2 ⊢ 𝐅*⦃T2⦄ ≡ g2 & g2 ⊆ g1 & g1 ⊆ f1. +lemma R_fle_comp_LTC: ∀R. R_fle_compatible R → R_fle_compatible (LTC … R). +#R #HR #L #T1 #T2 #H elim H -T2 +/3 width=3 by fle_trans_dx/ +qed-. lemma lfpxs_cpx_conf: ∀h,G. s_r_confluent1 … (cpx h G) (lfpxs h G). #h #G #L1 #T1 #T2 #HT12 #L2 #H -lapply (cpx_frees_conf_lfpxs … HT12) -HT12 #HT12 -@(lfpxs_ind_sn … H) -L2 // -#L #L2 #HL1 * #g1 #Hg1 #HL2 #IH -elim (frees_total L1 T1) #f1 #Hf1 -elim (HT12 … Hf1 … HL1 … Hg1) -T1 #g2 #Hg2 #Hg21 #_ -f1 -/4 width=7 by lfpxs_step_dx, sle_lexs_trans, ex2_intro/ +elim (tc_lfxs_inv_lex_lfeq … H) -H #L #HL1 #HL2 +lapply (lfxs_lex … HL1 T1) #H +elim (cpx_lfxs_conf_fle … HT12 … H) -HT12 -H // #_ #HT21 #_ +@(lfpxs_lpxs_lfeq … HL1) -HL1 +@(fle_lfxs_trans … HL2) // qed-.