X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Fcl_restricted_weight.ma;h=879af21b29b4d74ee83f5b0918b5bff2cbe40711;hb=156d974ad89aa04a086fdf9d332c8b04adf279fd;hp=31c3f7adc28ee03ce44cae6fbab1da9c5f05e2a7;hpb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;p=helm.git 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 31c3f7adc..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 @@ -17,36 +17,33 @@ include "static_2/syntax/lenv_weight.ma". (* WEIGHT OF A RESTRICTED CLOSURE *******************************************) -definition rfw: lenv → term → ? ≝ λL,T. ♯{L} + ♯{T}. +definition rfw: lenv → term → ? ≝ λL,T. ♯❨L❩ + ♯❨T❩. 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 monotonic_le_plus_r/ -qed. +lemma rfw_shift: ∀p,I,K,V,T. ♯❨K.ⓑ[I]V,T❩ < ♯❨K,ⓑ[p,I]V.T❩. +/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 monotonic_le_plus_r, le_S_S/ -qed. +lemma rfw_clear: ∀p,I1,I2,K,V,T. ♯❨K.ⓤ[I1],T❩ < ♯❨K,ⓑ[p,I2]V.T❩. +/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. +lemma rfw_tpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L,②[I]V.T❩. +/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. +lemma rfw_tpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L,②[I]V.T❩. +/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 monotonic_lt_plus_l, monotonic_le_plus_r/ -qed. +lemma rfw_lpair_sn: ∀I,L,V,T. ♯❨L,V❩ < ♯❨L.ⓑ[I]V,T❩. +// qed. -lemma rfw_lpair_dx: ∀I,L,V,T. ♯{L, T} < ♯{L.ⓑ{I}V, T}. -normalize /3 width=1 by monotonic_lt_plus_l, monotonic_le_plus_r/ -qed. +lemma rfw_lpair_dx: ∀I,L,V,T. ♯❨L,T❩ < ♯❨L.ⓑ[I]V,T❩. +/2 width=1 by plt_plus_bi_dx/ qed. (* Basic_1: removed theorems 7: flt_thead_sx flt_thead_dx flt_trans