]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/gcp_cr.ma
notational change of lift, drop, and gget
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / gcp_cr.ma
index a596e246d814b5a55926f5048679665be7d9679c..fa040b9155cb8ecd6b816f69252741258c292dab 100644 (file)
@@ -39,11 +39,11 @@ definition S4 ≝ λRP,C:candidate.
                 ∀G,L,Vs. all … (RP G L) Vs → ∀k. C G L (ⒶVs.⋆k).
 
 definition S5 ≝ λC:candidate. ∀I,G,L,K,Vs,V1,V2,i.
-                C G L (â\92¶Vs.V2) â\86\92 â\87§[0, i+1] V1 ≡ V2 →
-                â\87©[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
+                C G L (â\92¶Vs.V2) â\86\92 â¬\86[0, i+1] V1 ≡ V2 →
+                â¬\87[i] L ≡ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
 
 definition S6 ≝ λRP,C:candidate.
-                â\88\80G,L,V1s,V2s. â\87§[0, 1] V1s ≡ V2s →
+                â\88\80G,L,V1s,V2s. â¬\86[0, 1] V1s ≡ V2s →
                 ∀a,V,T. C G (L.ⓓV) (ⒶV2s.T) → RP G L V → C G L (ⒶV1s.ⓓ{a}V.T).
 
 definition S7 ≝ λC:candidate.
@@ -63,7 +63,7 @@ record gcr (RR:relation4 genv lenv term term) (RS:relation term) (RP,C:candidate
 (* the functional construction for candidates *)
 definition cfun: candidate → candidate → candidate ≝
                  λC1,C2,G,K,T. ∀L,W,U,des.
-                 â\87©*[â\92», des] L â\89¡ K â\86\92 â\87§*[des] T ≡ U → C1 G L W → C2 G L (ⓐW.U).
+                 â¬\87*[â\92», des] L â\89¡ K â\86\92 â¬\86*[des] T ≡ U → C1 G L W → C2 G L (ⓐW.U).
 
 (* the reducibility candidate associated to an atomic arity *)
 let rec acr (RP:candidate) (A:aarity) on A: candidate ≝
@@ -148,7 +148,7 @@ qed.
 
 lemma acr_abst: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
                 ∀a,G,L,W,T,A,B. ⦃G, L, W⦄ ϵ[RP] 〚B〛 → (
-                   â\88\80L0,V0,W0,T0,des. â\87©*[â\92», des] L0 â\89¡ L â\86\92 â\87§*[des] W â\89¡ W0 â\86\92 â\87§*[des + 1] T ≡ T0 →
+                   â\88\80L0,V0,W0,T0,des. â¬\87*[â\92», des] L0 â\89¡ L â\86\92 â¬\86*[des] W â\89¡ W0 â\86\92 â¬\86*[des + 1] T ≡ T0 →
                                    ⦃G, L0, V0⦄ ϵ[RP] 〚B〛 → ⦃G, L0, W0⦄ ϵ[RP] 〚B〛 → ⦃G, L0.ⓓⓝW0.V0, T0⦄ ϵ[RP] 〚A〛
                 ) →
                 ⦃G, L, ⓛ{a}W.T⦄ ϵ[RP] 〚②B.A〛.