X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Fcl_restricted_weight.ma;h=879af21b29b4d74ee83f5b0918b5bff2cbe40711;hp=62ce9d6fca80fa26e839b01a0fb9b16dc3bba84c;hb=156d974ad89aa04a086fdf9d332c8b04adf279fd;hpb=8fe4dc148d50a0352313633bea61441bc817afbf diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/cl_restricted_weight.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/cl_restricted_weight.ma index 62ce9d6fc..879af21b2 100644 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/cl_restricted_weight.ma +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/cl_restricted_weight.ma @@ -23,30 +23,27 @@ interpretation "weight (restricted closure)" 'Weight L T = (rfw L T). (* Basic properties *********************************************************) +lemma rfw_unfold (L) (T): ♯❨L❩ + ♯❨T❩ = ♯❨L,T❩. +// qed. + (* Basic_1: was: flt_shift *) lemma rfw_shift: ∀p,I,K,V,T. ♯❨K.ⓑ[I]V,T❩ < ♯❨K,ⓑ[p,I]V.T❩. -normalize /2 width=1 by nle_plus_bi_sn/ -qed. +/2 width=1 by plt_plus_bi_sn/ qed. lemma rfw_clear: ∀p,I1,I2,K,V,T. ♯❨K.ⓤ[I1],T❩ < ♯❨K,ⓑ[p,I2]V.T❩. -normalize /4 width=1 by nle_plus_bi_sn, nle_succ_bi/ -qed. +/2 width=1 by plt_plus_bi_sn/ qed. lemma rfw_tpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L,②[I]V.T❩. -normalize in ⊢ (?→?→?→?→?%%); // -qed. +/2 width=1 by plt_plus_bi_sn/ qed. lemma rfw_tpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L,②[I]V.T❩. -normalize in ⊢ (?→?→?→?→?%%); // -qed. +/2 width=1 by plt_plus_bi_sn/ qed. lemma rfw_lpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L.ⓑ[I]V,T❩. -normalize /3 width=1 by nlt_plus_bi_dx, nle_plus_bi_sn/ -qed. +// qed. lemma rfw_lpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L.ⓑ[I]V,T❩. -normalize /3 width=1 by nlt_plus_bi_dx, nle_plus_bi_sn/ -qed. +/2 width=1 by plt_plus_bi_dx/ qed. (* Basic_1: removed theorems 7: flt_thead_sx flt_thead_dx flt_trans