]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/computation/acp.ma
- first properties of strongly normalizing terms
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / acp.ma
index 3e13d9ac3bcc36145c7292a5489d02e53b7f926d..8be75a31f23398cb8a22ca3db4b82f708ad66199 100644 (file)
@@ -20,14 +20,14 @@ definition CP1 ≝ λRR:lenv→relation term. λRS:relation term.
                  ∀L,k. NF … (RR L) RS (⋆k).
 
 definition CP2 ≝ λRR:lenv→relation term. λRS:relation term.
-                 ∀L,K,W,i. ⇩[0,i] L ≡ K. 𝕓{Abst} W → NF … (RR L) RS (#i).
+                 ∀L,K,W,i. ⇩[0,i] L ≡ K. W → NF … (RR L) RS (#i).
 
 definition CP3 ≝ λRR:lenv→relation term. λRP:lenv→predicate term.
-                 ∀L,V,k. RP L (𝕔{Appl}⋆k.V) → RP L V.
+                 ∀L,V,k. RP L (⋆k.V) → RP L V.
 
 definition CP4 ≝ λRR:lenv→relation term. λRS:relation term.
                  ∀L0,L,T,T0,d,e. NF … (RR L) RS T → 
-                 ⇩[d,e] L0 ≡ L → ⇧[d, e] T ≡ T0 → NF … (RR L0) RS T0.
+                 ⇩[d, e] L0 ≡ L → ⇧[d, e] T ≡ T0 → NF … (RR L0) RS T0.
 
 definition CP4s ≝ λRR:lenv→relation term. λRS:relation term.
                   ∀L0,L,des. ⇩*[des] L0 ≡ L →
@@ -44,6 +44,7 @@ record acp (RR:lenv->relation term) (RS:relation term) (RP:lenv→predicate term
 
 (* Basic properties *********************************************************)
 
+(* Basic_1: was: nf2_lift1 *)
 lemma acp_lifts: ∀RR,RS. CP4 RR RS → CP4s RR RS.
 #RR #RS #HRR #L1 #L2 #des #H elim H -L1 -L2 -des
 [ #L #T1 #T2 #H #HT1