]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/cprs_cprs.ma
notational change of lift, drop, and gget
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / cprs_cprs.ma
index 0b9d7a021cc0105680e97d6afe60008dda56239d..e86df253ee59842817485ddcc2c669d3f7ab671a 100644 (file)
@@ -59,14 +59,14 @@ theorem cprs_beta: ∀a,G,L,V1,V2,W1,W2,T1,T2.
 qed.
 
 theorem cprs_theta_rc: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
-                       â¦\83G, Lâ¦\84 â\8a¢ V1 â\9e¡ V â\86\92 â\87§[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
+                       â¦\83G, Lâ¦\84 â\8a¢ V1 â\9e¡ V â\86\92 â¬\86[0, 1] V ≡ V2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
                        ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
 #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV1 #HV2 #HT12 #H @(cprs_ind … H) -W2
 /3 width=5 by cprs_trans, cprs_theta_dx, cprs_bind_dx/
 qed.
 
 theorem cprs_theta: ∀a,G,L,V1,V,V2,W1,W2,T1,T2.
-                    â\87§[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
+                    â¬\86[0, 1] V ≡ V2 → ⦃G, L⦄ ⊢ W1 ➡* W2 → ⦃G, L.ⓓW1⦄ ⊢ T1 ➡* T2 →
                     ⦃G, L⦄ ⊢ V1 ➡* V → ⦃G, L⦄ ⊢ ⓐV1.ⓓ{a}W1.T1 ➡* ⓓ{a}W2.ⓐV2.T2.
 #a #G #L #V1 #V #V2 #W1 #W2 #T1 #T2 #HV2 #HW12 #HT12 #H @(cprs_ind_dx … H) -V1
 /3 width=3 by cprs_trans, cprs_theta_rc, cprs_flat_dx/
@@ -80,7 +80,7 @@ lemma cprs_inv_appl1: ∀G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓐV1.T1 ➡* U2 →
                                         U2 = ⓐV2. T2
                        | ∃∃a,W,T.       ⦃G, L⦄ ⊢ T1 ➡* ⓛ{a}W.T &
                                         ⦃G, L⦄ ⊢ ⓓ{a}ⓝW.V1.T ➡* U2
-                       | â\88\83â\88\83a,V0,V2,V,T. â¦\83G, Lâ¦\84 â\8a¢ V1 â\9e¡* V0 & â\87§[0,1] V0 ≡ V2 &
+                       | â\88\83â\88\83a,V0,V2,V,T. â¦\83G, Lâ¦\84 â\8a¢ V1 â\9e¡* V0 & â¬\86[0,1] V0 ≡ V2 &
                                         ⦃G, L⦄ ⊢ T1 ➡* ⓓ{a}V.T &
                                         ⦃G, L⦄ ⊢ ⓓ{a}V.ⓐV2.T ➡* U2.
 #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by or3_intro0, ex3_2_intro/