X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fgrammar%2Fcl_weight.ma;h=40a364a44124ccf03c9f38d01074936252ec3465;hb=439b6ec33d749ba4e6ae0938e973a85bc23e306e;hp=e3dee1e0bac03c0764bf5854c1b1a9fdc3d25160;hpb=84713c9446ff13dd26533590985a0df67cb5ec7e;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 e3dee1e0b..40a364a44 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 @@ -17,7 +17,7 @@ include "basic_2/grammar/cl_shift.ma". (* WEIGHT OF A CLOSURE ******************************************************) -definition cw: lenv → term → ? ≝ λL,T. #[L] + #[T]. +definition cw: lenv → term → ? ≝ λL,T. #{L} + #{T}. interpretation "weight (closure)" 'Weight L T = (cw L T). @@ -27,34 +27,35 @@ interpretation "weight (closure)" 'Weight L T = (cw L T). (* Basic_1: was: flt_wf_ind *) axiom cw_wf_ind: ∀R:lenv→predicate term. - (∀L2,T2. (∀L1,T1. #[L1,T1] < #[L2,T2] → R L1 T1) → R L2 T2) → + (∀L2,T2. (∀L1,T1. #{L1,T1} < #{L2,T2} → R L1 T1) → R L2 T2) → ∀L,T. R L T. (* Basic_1: was: flt_shift *) -lemma cw_shift: ∀K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ②{I} V. T]. +lemma cw_shift: ∀a,K,I,V,T. #{K. ⓑ{I} V, T} < #{K, ⓑ{a,I} V. T}. normalize // qed. -lemma tw_shift: ∀L,T. #[L, T] ≤ #[L @ T]. +lemma tw_shift: ∀L,T. #{L, T} ≤ #{L @@ T}. #L elim L // #K #I #V #IHL #T @transitive_le [3: @IHL |2: /2 width=2/ | skip ] qed. -lemma cw_tpair_sn: ∀I,L,V,T. #[L, V] < #[L, ②{I}V.T]. +lemma cw_tpair_sn: ∀I,L,V,T. #{L, V} < #{L, ②{I}V.T}. #I #L #V #T normalize in ⊢ (? % %); // qed. -lemma cw_tpair_dx: ∀I,L,V,T. #[L, T] < #[L, ②{I}V.T]. +lemma cw_tpair_dx: ∀I,L,V,T. #{L, T} < #{L, ②{I}V.T}. #I #L #V #T normalize in ⊢ (? % %); // qed. -lemma cw_tpair_dx_sn: ∀I1,I2,L,V1,V2,T. #[L, V2] < #[L, ②{I1}V1.②{I2}V2.T]. +lemma cw_tpair_dx_sn: ∀I1,I2,L,V1,V2,T. #{L, V2} < #{L, ②{I1}V1.②{I2}V2.T}. #I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/ qed. -lemma cw_tpair_dx_sn_shift: ∀I1,I2,L,V1,V2,T. #[L.ⓑ{I2}V2, T] < #[L, ②{I1}V1.②{I2}V2.T]. -#I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/ +lemma cw_tpair_dx_sn_shift: ∀a2,I1,I2,L,V1,V2,T. + #{L.ⓑ{I2}V2, T} < #{L, ②{I1}V1.ⓑ{a2,I2}V2.T}. +#a2 #I1 #I2 #L #V1 #V2 #T normalize in ⊢ (? % %); /2 width=1/ qed. (* Basic_1: removed theorems 6: