X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Fcl_weight.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Fgrammar%2Fcl_weight.ma;h=355b8a7973a46bbcd67c4093df77939cbc2a9b0f;hb=78f21d7d9014e5c7655f58239e4f1a128ea2c558;hp=35bf32a9dd22cbdd66787e9300d61f7c3e5800eb;hpb=0c547d79d0edc850787ac27edc25565aa271ec2a;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma b/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma index 35bf32a9d..355b8a797 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/grammar/cl_weight.ma @@ -12,8 +12,8 @@ (* *) (**************************************************************************) -include "Basic-2/grammar/lenv_weight.ma". -include "Basic-2/grammar/cl_shift.ma". +include "Basic_2/grammar/lenv_weight.ma". +include "Basic_2/grammar/cl_shift.ma". (* WEIGHT OF A CLOSURE ******************************************************) @@ -23,14 +23,14 @@ interpretation "weight (closure)" 'Weight L T = (cw L T). (* Basic properties *********************************************************) -(* Basic-1: was: flt_wf__q_ind *) +(* Basic_1: was: flt_wf__q_ind *) -(* Basic-1: was: flt_wf_ind *) +(* Basic_1: was: flt_wf_ind *) axiom cw_wf_ind: ∀R:lenv→term→Prop. (∀L2,T2. (∀L1,T1. #[L1,T1] < #[L2,T2] → R L1 T1) → R L2 T2) → ∀L,T. R L T. -(* Basic-1: was: flt_shift *) +(* Basic_1: was: flt_shift *) lemma cw_shift: ∀K,I,V,T. #[K. 𝕓{I} V, T] < #[K, 𝕔{I} V. T]. normalize // qed. @@ -41,6 +41,6 @@ lemma tw_shift: ∀L,T. #[L, T] ≤ #[L @ T]. @transitive_le [3: @IHL |2: /2/ | skip ] qed. -(* Basic-1: removed theorems 6: +(* Basic_1: removed theorems 6: flt_thead_sx flt_thead_dx flt_arith0 flt_arith1 flt_arith2 flt_trans *)