]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_vector.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_vector.ma
index 95269945f490d24cf17d67131b86ea10d2726e8d..32afe79a4a87cf73d6d69943f70b3261fd8d655e 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/grammar/term_vector.ma".
-include "basic_2/computation/csx.ma".
+include "static_2/syntax/term_vector.ma".
+include "basic_2/rt_computation/csx.ma".
 
-(* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERM VECTORS *************)
+(* STRONGLY NORMALIZING TERMS VECTORS FOR UNBOUND PARALLEL RT-TRANSITION ****)
 
 definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
                  λh,o,G,L. all … (csx h o G L).
 
 interpretation
-   "context-sensitive strong normalization (term vector)"
-   'SN h o G L Ts = (csxv h o G L Ts).
+   "strong normalization for unbound context-sensitive parallel rt-transition (term vector)"
+   'PRedTyStrong h o G L Ts = (csxv h o G L Ts).
 
 (* Basic inversion lemmas ***************************************************)
 
-lemma csxv_inv_cons: â\88\80h,o,G,L,T,Ts. â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] T @ Ts →
-                     â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] T â\88§ â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] Ts.
+lemma csxv_inv_cons: â\88\80h,o,G,L,T,Ts. â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83T⨮Tsâ¦\84 →
+                     â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83Tâ¦\84 â\88§ â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83Tsâ¦\84.
 normalize // qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-lemma csx_fwd_applv: â\88\80h,o,G,L,T,Vs. â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] â\92¶ Vs.T →
-                     â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] Vs â\88§ â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] T.
+lemma csx_fwd_applv: â\88\80h,o,G,L,T,Vs. â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83â\92¶Vs.Tâ¦\84 →
+                     â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83Vsâ¦\84 â\88§ â¦\83G, Lâ¦\84 â\8a¢ â¬\88*[h, o] ð\9d\90\92â¦\83Tâ¦\84.
 #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