]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/computation/acp.ma
- main lemmas about abstract reducibility candidates closed
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / computation / acp.ma
index 437a51c55da4ef136aaf74c7d4bf072d0422fd12..ca1ba468640db0e0577b3165d9f82bc1aee367b3 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic_2/substitution/ldrop.ma".
+include "Basic_2/unfold/ldrops.ma".
 
 (* ABSTRACT COMPUTATION PROPERTIES ******************************************)
 
@@ -20,14 +20,35 @@ 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.
-                 â\88\80L,K,W,i. â\87\93[0,i] L â\89¡ K. ð\9d\95\93{Abst} W → NF … (RR L) RS (#i).
+                 â\88\80L,K,W,i. â\87©[0,i] L â\89¡ K. â\93\9bW → 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.
+
+definition CP4s ≝ λRR:lenv→relation term. λRS:relation term.
+                  ∀L0,L,des. ⇩*[des] L0 ≡ L →
+                  ∀T,T0. ⇧*[des] T ≡ T0 →
+                  NF … (RR L) RS T → NF … (RR L0) RS T0.
 
 (* requirements for abstract computation properties *)
 record acp (RR:lenv->relation term) (RS:relation term) (RP:lenv→predicate term) : Prop ≝
 { cp1: CP1 RR RS;
   cp2: CP2 RR RS;
-  cp3: CP3 RR RP
+  cp3: CP3 RR RP;
+  cp4: CP4 RR RS
 }.
+
+(* Basic properties *********************************************************)
+
+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
+  <(lifts_inv_nil … H) -H //
+| #L1 #L #L2 #des #d #e #_ #HL2 #IHL #T2 #T1 #H #HLT2
+  elim (lifts_inv_cons … H) -H /3 width=9/
+]
+qed.