X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fgrammar%2Fcl_restricted_weight.ma;h=fb2bcaf94cff0b3f8232069d9f7979a79f0a6acb;hb=e9da8e091898b6e67a2f270581bdc5cdbe80e9b0;hp=e889b05482729a7fc6d560e9101dd39b44e819e8;hpb=54fa4874fc4bfccd061b40d8353cd75a578e99ae;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_restricted_weight.ma b/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_restricted_weight.ma index e889b0548..fb2bcaf94 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_restricted_weight.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/grammar/cl_restricted_weight.ma @@ -24,7 +24,7 @@ interpretation "weight (restricted closure)" 'Weight L T = (rfw L T). (* Basic properties *********************************************************) (* Basic_1: was: flt_shift *) -lemma rfw_shift: ∀a,I,K,V,T. ♯{K.ⓑ{I}V, T} < ♯{K, ⓑ{a,I}V.T}. +lemma rfw_shift: ∀p,I,K,V,T. ♯{K.ⓑ{I}V, T} < ♯{K, ⓑ{p,I}V.T}. normalize // qed. @@ -40,6 +40,10 @@ 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_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. + (* Basic_1: removed theorems 7: flt_thead_sx flt_thead_dx flt_trans flt_arith0 flt_arith1 flt_arith2 flt_wf_ind