]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_theq_vector.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_theq_vector.ma
index bab5c447ede214fb299c5481c2bf5eac8aa1291e..147867a8b216bf9234ba673684fe05d252ca3b69 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/syntax/theq_simple_vector.ma".
-include "basic_2/relocation/lifts_vector.ma".
+include "static_2/syntax/theq_simple_vector.ma".
+include "static_2/relocation/lifts_vector.ma".
 include "basic_2/rt_computation/cpxs_theq.ma".
 
-(* UNCOUNTED CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS ************)
+(* UNBOUND CONTEXT-SENSITIVE PARALLEL RT-COMPUTATION FOR TERMS **************)
 
 (* Vector form of forward lemmas with head equivalence for terms ************)
 
@@ -44,8 +44,8 @@ elim (cpxs_inv_appl1 … H) -H *
 qed-.
 
 (* Basic_2A1: was: cpxs_fwd_delta_vector *)
-lemma cpxs_fwd_delta_drops_vector: â\88\80h,o,I,G,L,K,V1,i. â¬\87*[i] L â\89¡ K.ⓑ{I}V1 →
-                                   â\88\80V2. â¬\86*[⫯i] V1 â\89¡ V2 →
+lemma cpxs_fwd_delta_drops_vector: â\88\80h,o,I,G,L,K,V1,i. â¬\87*[i] L â\89\98 K.ⓑ{I}V1 →
+                                   â\88\80V2. â¬\86*[â\86\91i] V1 â\89\98 V2 →
                                    ∀Vs,U. ⦃G, L⦄ ⊢ ⒶVs.#i ⬈*[h] U →
                                    ⒶVs.#i ⩳[h, o] U ∨ ⦃G, L⦄ ⊢ ⒶVs.V2 ⬈*[h] U.
 #h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #Vs elim Vs -Vs /2 width=5 by cpxs_fwd_delta_drops/
@@ -94,7 +94,7 @@ elim (cpxs_inv_appl1 … H) -H *
 qed-.
 
 (* Basic_1: was just: pr3_iso_appls_abbr *)
-lemma cpxs_fwd_theta_vector: â\88\80h,o,G,L,V1b,V2b. â¬\86*[1] V1b â\89¡ V2b →
+lemma cpxs_fwd_theta_vector: â\88\80h,o,G,L,V1b,V2b. â¬\86*[1] V1b â\89\98 V2b →
                              ∀p,V,T,U. ⦃G, L⦄ ⊢ ⒶV1b.ⓓ{p}V.T ⬈*[h] U →
                              ⒶV1b.ⓓ{p}V.T ⩳[h, o] U ∨ ⦃G, L⦄ ⊢ ⓓ{p}V.ⒶV2b.T ⬈*[h] U.
 #h #o #G #L #V1b #V2b * -V1b -V2b /3 width=1 by or_intror/