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=d9569d4d57ffd3a236d7d591c72688a435b6160b;hb=6c86c70b005e3f3efd375868b27f3cff84febfad;hp=f5f7feb7fc6f3bd95e32153641d2ba3c58d4c200;hpb=a2144f09d1bd7022c1f2dfd4909a1fb9772c8d56;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 f5f7feb7f..d9569d4d5 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,44 +17,50 @@ include "basic_2/grammar/cl_shift.ma". (* WEIGHT OF A CLOSURE ******************************************************) -definition cw: lenv → term → ? ≝ λL,T. #[L] + #[T]. +definition fw: lenv → term → ? ≝ λL,T. #{L} + #{T}. -interpretation "weight (closure)" 'Weight L T = (cw L T). +interpretation "weight (closure)" 'Weight L T = (fw L T). (* Basic properties *********************************************************) (* Basic_1: was: flt_wf__q_ind *) (* 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) → - ∀L,T. R L T. +axiom fw_ind: ∀R:relation2 lenv term. + (∀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: ∀a,K,I,V,T. #[K. ⓑ{I} V, T] < #[K, ⓑ{a,I} V. T]. +lemma fw_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]. -#I #L #V #T normalize in ⊢ (? % %); // +lemma fw_tpair_sn: ∀I,L,V,T. #{L, V} < #{L, ②{I}V.T}. +normalize in ⊢ (?→?→?→?→?%%); // qed. -lemma cw_tpair_dx: ∀I,L,V,T. #[L, T] < #[L, ②{I}V.T]. -#I #L #V #T normalize in ⊢ (? % %); // +lemma fw_tpair_dx: ∀I,L,V,T. #{L, T} < #{L, ②{I}V.T}. +normalize in ⊢ (?→?→?→?→?%%); // qed. -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/ +lemma fw_tpair_dx_sn: ∀I1,I2,L,V1,V2,T. #{L, V2} < #{L, ②{I1}V1.②{I2}V2.T}. +normalize in ⊢ (?→?→?→?→?→?→?%%); /2 width=1/ qed. -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/ +lemma fw_tpair_sn_sn_shift: ∀I,I1,I2,L,V1,V2,T. + #{L.ⓑ{I}V1, T} < #{L, ②{I1}V1.②{I2}V2.T}. +normalize in ⊢ (?→?→?→?→?→?→?→?%%); /3 width=1/ +qed. + +lemma fw_tpair_sn_dx_shift: ∀I,I1,I2,L,V1,V2,T. + #{L.ⓑ{I}V2, T} < #{L, ②{I1}V1.②{I2}V2.T}. +normalize in ⊢ (?→?→?→?→?→?→?→?%%); /2 width=1/ qed. (* Basic_1: removed theorems 6: