]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/gcp_cr.ma
update in static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / gcp_cr.ma
index d13387c3f743120c850910b296d0f984b1dc380d..e14ca2a6f9c18847776888afb83ef9aa4314a9bc 100644 (file)
@@ -36,11 +36,11 @@ definition S3 ≝ λC:candidate.
                 C G L (ⒶVs.ⓓ{a}ⓝW.V.T) → C G L (ⒶVs.ⓐV.ⓛ{a}W.T).
 
 definition S5 ≝ λC:candidate. ∀I,G,L,K,Vs,V1,V2,i.
-                C G L (â\92¶Vs.V2) â\86\92 â¬\86*[↑i] V1 ≘ V2 →
-                â¬\87*[i] L ≘ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
+                C G L (â\92¶Vs.V2) â\86\92 â\87§*[↑i] V1 ≘ V2 →
+                â\87©*[i] L ≘ K.ⓑ{I}V1 → C G L (ⒶVs.#i).
 
 definition S6 ≝ λRP,C:candidate.
-                â\88\80G,L,V1b,V2b. â¬\86*[1] V1b ≘ V2b →
+                â\88\80G,L,V1b,V2b. â\87§*[1] V1b ≘ V2b →
                 ∀a,V,T. C G (L.ⓓV) (ⒶV2b.T) → RP G L V → C G L (ⒶV1b.ⓓ{a}V.T).
 
 definition S7 ≝ λC:candidate.
@@ -59,7 +59,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. ∀f,L,W,U.
-                 â¬\87*[â\92»,f] L â\89\98 K â\86\92 â¬\86*[f] T ≘ U → C1 G L W → C2 G L (ⓐW.U).
+                 â\87©*[â\92»,f] L â\89\98 K â\86\92 â\87§*[f] T ≘ U → C1 G L W → C2 G L (ⓐW.U).
 
 (* the reducibility candidate associated to an atomic arity *)
 rec definition acr (RP:candidate) (A:aarity) on A: candidate ≝
@@ -149,7 +149,7 @@ qed.
 
 lemma acr_abst: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP →
                 ∀p,G,L,W,T,A,B. ⦃G,L,W⦄ ϵ[RP] 〚B〛 → (
-                   â\88\80b,f,L0,V0,W0,T0. â¬\87*[b,f] L0 â\89\98 L â\86\92 â¬\86*[f] W â\89\98 W0 â\86\92 â¬\86*[⫯f] T ≘ T0 →
+                   â\88\80b,f,L0,V0,W0,T0. â\87©*[b,f] L0 â\89\98 L â\86\92 â\87§*[f] W â\89\98 W0 â\86\92 â\87§*[⫯f] T ≘ T0 →
                                    ⦃G,L0,V0⦄ ϵ[RP] 〚B〛 → ⦃G,L0,W0⦄ ϵ[RP] 〚B〛 → ⦃G,L0.ⓓⓝW0.V0,T0⦄ ϵ[RP] 〚A〛
                 ) →
                 ⦃G,L,ⓛ{p}W.T⦄ ϵ[RP] 〚②B.A〛.