]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrop.ma
- we polarized binders to control zeta reduction
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / lsubc_ldrop.ma
index 03ea7b6289922c57c645eee94c05369545162883..a7c7c7a99be5922b2decdaf5d116f9f0eb6c7072 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic_2/static/aaa_lift.ma".
-include "Basic_2/computation/acp_cr.ma".
-include "Basic_2/computation/lsubc.ma".
+include "basic_2/static/aaa_lift.ma".
+include "basic_2/computation/lsubc.ma".
 
 (* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****)
 
 (* Properties concerning basic local environment slicing ********************)
 
 (* Basic_1: was: csubc_drop_conf_O *)
-(* Note: the constant (0) can not be generalized *)
-lemma lsubc_ldrop_O1_trans: ∀RP,L1,L2. L1 [RP] ⊑ L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 →
-                            ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 [RP] ⊑ K2.
+(* Note: the constant 0 can not be generalized *)
+lemma lsubc_ldrop_O1_trans: ∀RP,L1,L2. L1 ⊑[RP] L2 → ∀K2,e. ⇩[0, e] L2 ≡ K2 →
+                            ∃∃K1. ⇩[0, e] L1 ≡ K1 & K1 ⊑[RP] K2.
 #RP #L1 #L2 #H elim H -L1 -L2
 [ #X #e #H
   >(ldrop_inv_atom1 … H) -H /2 width=3/
@@ -42,8 +41,8 @@ qed-.
 (* Basic_1: was: csubc_drop_conf_rev *)
 lemma ldrop_lsubc_trans: ∀RR,RS,RP.
                          acp RR RS RP → acr RR RS RP (λL,T. RP L T) →
-                         ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 →
-                         ∃∃L2. L1 [RP] ⊑ L2 & ⇩[d, e] L2 ≡ K2.
+                         ∀L1,K1,d,e. ⇩[d, e] L1 ≡ K1 → ∀K2. K1 ⊑[RP] K2 →
+                         ∃∃L2. L1 ⊑[RP] L2 & ⇩[d, e] L2 ≡ K2.
 #RR #RS #RP #Hacp #Hacr #L1 #K1 #d #e #H elim H -L1 -K1 -d -e
 [ #d #e #X #H
   >(lsubc_inv_atom1 … H) -H /2 width=3/