]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/csx_vector.ma
milestone update in basic_2, update in ground and static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / csx_vector.ma
index 6fdda2bbfb0ba80334d3f068698209206b1509b4..faad2323aacbb97c37ef93c852bf34cc6b4041e1 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/syntax/term_vector.ma".
+include "static_2/syntax/term_vector.ma".
 include "basic_2/rt_computation/csx.ma".
 
-(* STRONGLY NORMALIZING TERMS VECTORS FOR UNBOUND PARALLEL RT-TRANSITION ****)
+(* STRONGLY NORMALIZING TERMS VECTORS FOR EXTENDED PARALLEL RT-TRANSITION ***)
 
-definition csxv: ∀h. sd h → relation3 genv lenv (list term) ≝
-                 λh,o,G,L. all … (csx h o G L).
+definition csxv (G) (L): predicate (list term) ≝
+           all … (csx G L).
 
 interpretation
-   "strong normalization for unbound context-sensitive parallel rt-transition (term vector)"
-   'PRedTyStrong h o G L Ts = (csxv h o G L Ts).
+  "strong normalization for extended context-sensitive parallel rt-transition (term vector)"
+  'PRedTyStrong G L Ts = (csxv G L Ts).
 
 (* Basic inversion lemmas ***************************************************)
 
-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⦄.
+lemma csxv_inv_cons (G) (L):
+      ∀T,Ts. ❪G,L❫ ⊢ ⬈*𝐒 T⨮Ts →
+      ∧∧ ❪G,L❫ ⊢ ⬈*𝐒 T & ❪G,L❫ ⊢ ⬈*𝐒 Ts.
 normalize // qed-.
 
 (* Basic forward lemmas *****************************************************)
 
-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/
+lemma csx_fwd_applv (G) (L):
+      ∀T,Vs. ❪G,L❫ ⊢ ⬈*𝐒 ⒶVs.T →
+      ∧∧ ❪G,L❫ ⊢ ⬈*𝐒 Vs & ❪G,L❫ ⊢ ⬈*𝐒 T.
+#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