X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Fterm_weight.ma;h=5b7853c572cac973b5a104e21e3a03be8766bdd4;hp=bcb826bb41bf7477067c229b2116c9d48105f31a;hb=156d974ad89aa04a086fdf9d332c8b04adf279fd;hpb=8fe4dc148d50a0352313633bea61441bc817afbf diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/term_weight.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/term_weight.ma index bcb826bb4..5b7853c57 100644 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/term_weight.ma +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/term_weight.ma @@ -12,13 +12,14 @@ (* *) (**************************************************************************) +include "ground/arith/pnat_lt_plus.ma". include "static_2/notation/functions/weight_1.ma". include "static_2/syntax/term.ma". (* WEIGHT OF A TERM *********************************************************) rec definition tw T ≝ match T with -[ TAtom _ ⇒ 1 +[ TAtom _ ⇒ 𝟏 | TPair _ V T ⇒ ↑(tw V + tw T) ]. @@ -26,16 +27,18 @@ interpretation "weight (term)" 'Weight T = (tw T). (* Basic properties *********************************************************) -(* Basic_1: was: tweight_lt *) -lemma tw_pos: ∀T. 1 ≤ ♯❨T❩. -#T elim T -T // -qed. +lemma tw_atom_unfold (I): 𝟏 = ♯❨⓪[I]❩. +// qed. + +lemma tw_pair_unfold (I) (V) (T): ↑(♯❨V❩ + ♯❨T❩) = ♯❨②[I]V.T❩. +// qed. lemma tw_le_pair_dx (I): ∀V,T. ♯❨T❩ < ♯❨②[I]V.T❩. -#I #V #T /2 width=1 by nle_succ_bi/ +/2 width=1 by plt_succ_dx_trans/ qed. -(* Basic_1: removed theorems 11: +(* Basic_1: removed theorems 12: + tweight_lt wadd_le wadd_lt wadd_O weight_le weight_eq weight_add_O weight_add_S tlt_trans tlt_head_sx tlt_head_dx tlt_wf_ind *)