]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxs_theq.ma
update in basic_2 and ground_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxs_theq.ma
index 8a52882a2b5ef4ca58dedb110a99088cdbfe294f..580f90ab2bd05b3454a730799f63efdf9c2a0588 100644 (file)
@@ -32,8 +32,8 @@ qed-.
 
 (* Note: probably this is an inversion lemma *)
 (* Basic_2A1: was: cpxs_fwd_delta *)
-lemma cpxs_fwd_delta_drops: â\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: â\88\80h,o,I,G,L,K,V1,i. â¬\87*[i] L â\89\98 K.ⓑ{I}V1 →
+                            â\88\80V2. â¬\86*[⫯i] V1 â\89\98 V2 →
                             ∀U. ⦃G, L⦄ ⊢ #i ⬈*[h] U →
                             #i ⩳[h, o] U ∨ ⦃G, L⦄ ⊢ V2 ⬈*[h] U.
 #h #o #I #G #L #K #V1 #i #HLK #V2 #HV12 #U #H
@@ -58,7 +58,7 @@ lemma cpxs_fwd_beta: ∀h,o,p,G,L,V,W,T,U. ⦃G, L⦄ ⊢ ⓐV.ⓛ{p}W.T ⬈*[h]
 qed-.
 
 lemma cpxs_fwd_theta: ∀h,o,p,G,L,V1,V,T,U. ⦃G, L⦄ ⊢ ⓐV1.ⓓ{p}V.T ⬈*[h] U →
-                      â\88\80V2. â¬\86*[1] V1 â\89¡ V2 → ⓐV1.ⓓ{p}V.T ⩳[h, o] U ∨
+                      â\88\80V2. â¬\86*[1] V1 â\89\98 V2 → ⓐV1.ⓓ{p}V.T ⩳[h, o] U ∨
                       ⦃G, L⦄ ⊢ ⓓ{p}V.ⓐV2.T ⬈*[h] U.
 #h #o #p #G #L #V1 #V #T #U #H #V2 #HV12
 elim (cpxs_inv_appl1 … H) -H *