X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Fcsn_vector.ma;h=73019d0533cdf6b83d334a6a924a884bfc93c866;hb=82500a9ceb53e1af0263c22afbd5954fa3a83190;hp=f29cedc04d6aca787dd9aa0ae2fe22977d9c581b;hpb=8ed01fd6a38bea715ceb449bb7b72a46bad87851;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma index f29cedc04..73019d053 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/csn_vector.ma @@ -17,24 +17,24 @@ include "basic_2/computation/csn.ma". (* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************) -definition csnv: ∀h. sd h → lenv → predicate (list term) ≝ - λh,g,L. all … (csn h g L). +definition csnv: ∀h. sd h → relation3 genv lenv (list term) ≝ + λh,g,G,L. all … (csn h g G L). interpretation "context-sensitive strong normalization (term vector)" - 'SN h g L Ts = (csnv h g L Ts). + 'SN h g G L Ts = (csnv h g G L Ts). (* Basic inversion lemmas ***************************************************) -lemma csnv_inv_cons: ∀h,g,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, g] T @ Ts → +lemma csnv_inv_cons: ∀h,g,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, g] T @ Ts → ⦃G, L⦄ ⊢ ⬊*[h, g] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] Ts. normalize // qed-. (* Basic forward lemmas *****************************************************) -lemma csn_fwd_applv: ∀h,g,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Ⓐ Vs.T → +lemma csn_fwd_applv: ∀h,g,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, g] Ⓐ Vs.T → ⦃G, L⦄ ⊢ ⬊*[h, g] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] T. -#h #g #L #T #Vs elim Vs -Vs /2 width=1/ +#h #g #G #L #T #Vs elim Vs -Vs /2 width=1/ #V #Vs #IHVs #HVs lapply (csn_fwd_pair_sn … HVs) #HV lapply (csn_fwd_flat_dx … HVs) -HVs #HVs