]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_vector.ma
- commit of the "s_computation" component ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_vector.ma
index 5b4d465cb0aab2c316fa14bea01577ee925ead1b..95269945f490d24cf17d67131b86ea10d2726e8d 100644 (file)
@@ -18,23 +18,23 @@ include "basic_2/computation/csx.ma".
 (* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
 
 definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
-                 λh,g,G,L. all … (csx h g G L).
+                 λh,o,G,L. all … (csx h o G L).
 
 interpretation
    "context-sensitive strong normalization (term vector)"
-   'SN h g G L Ts = (csxv h g G L Ts).
+   'SN h o G L Ts = (csxv h o G L Ts).
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma csxv_inv_cons: ∀h,g,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, g] T @ Ts →
-                     ⦃G, L⦄ ⊢ ⬊*[h, g] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, g] Ts.
+lemma csxv_inv_cons: ∀h,o,G,L,T,Ts. ⦃G, L⦄ ⊢ ⬊*[h, o] T @ Ts →
+                     ⦃G, L⦄ ⊢ ⬊*[h, o] T ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] Ts.
 normalize // qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-lemma csx_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 #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
+lemma csx_fwd_applv: ∀h,o,G,L,T,Vs. ⦃G, L⦄ ⊢ ⬊*[h, o] Ⓐ Vs.T →
+                     ⦃G, L⦄ ⊢ ⬊*[h, o] Vs ∧ ⦃G, L⦄ ⊢ ⬊*[h, o] T.
+#h #o #G #L #T #Vs elim Vs -Vs /2 width=1 by conj/
 #V #Vs #IHVs #HVs
 lapply (csx_fwd_pair_sn … HVs) #HV
 lapply (csx_fwd_flat_dx … HVs) -HVs #HVs