]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/cpxe.ma
- notational change for cpg and cpx
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / cpxe.ma
index 00fd04247f90aefe5d1ee428c42c2cbc421e3566..e71eebdc4fc1a4231fd14d594bace8f776b916c6 100644 (file)
@@ -19,14 +19,14 @@ include "basic_2/computation/csx.ma".
 (* EVALUATION FOR CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION ON TERMS ****)
 
 definition cpxe: ∀h. sd h → relation4 genv lenv term term ≝
-                 Î»h,o,G,L,T1,T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â\9e¡*[h, o] T2 â\88§ â¦\83G, Lâ¦\84 â\8a¢ â\9e¡[h, o] 𝐍⦃T2⦄.
+                 Î»h,o,G,L,T1,T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â¬\88*[h, o] T2 â\88§ â¦\83G, Lâ¦\84 â\8a¢ â¬\88[h, o] 𝐍⦃T2⦄.
 
 interpretation "evaluation for context-sensitive extended parallel reduction (term)"
    'PRedEval h o G L T1 T2 = (cpxe h o G L T1 T2).
 
 (* Basic properties *********************************************************)
 
-lemma csx_cpxe: â\88\80h,o,G,L,T1. â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] T1 â\86\92 â\88\83T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â\9e¡*[h, o] 𝐍⦃T2⦄.
+lemma csx_cpxe: â\88\80h,o,G,L,T1. â¦\83G, Lâ¦\84 â\8a¢ â¬\8a*[h, o] T1 â\86\92 â\88\83T2. â¦\83G, Lâ¦\84 â\8a¢ T1 â¬\88*[h, o] 𝐍⦃T2⦄.
 #h #o #G #L #T1 #H @(csx_ind … H) -T1
 #T1 #_ #IHT1 elim (cnx_dec h o G L T1) /3 width=3 by ex_intro, conj/
 * #T #H1T1 #H2T1 elim (IHT1 … H1T1 H2T1) -IHT1 -H2T1