]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma
lfsx_drops completed
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / drops.ma
index 7702a5a94ff6bd9e5778a6825e58df57b5002a9c..d469f0ac178cb7c9c66d767fd22d392a46de8c08 100644 (file)
@@ -42,10 +42,18 @@ definition d_liftable1: predicate (relation2 lenv term) ≝
                         λR. ∀K,T. R K T → ∀b,f,L. ⬇*[b, f] L ≡ K →
                         ∀U. ⬆*[f] T ≡ U → R L U.
 
+definition d_liftable1_isuni: predicate (relation2 lenv term) ≝
+                              λR. ∀K,T. R K T → ∀b,f,L. ⬇*[b, f] L ≡ K → 𝐔⦃f⦄ →
+                              ∀U. ⬆*[f] T ≡ U → R L U.
+
 definition d_deliftable1: predicate (relation2 lenv term) ≝
                           λR. ∀L,U. R L U → ∀b,f,K. ⬇*[b, f] L ≡ K →
                           ∀T. ⬆*[f] T ≡ U → R K T.
 
+definition d_deliftable1_isuni: predicate (relation2 lenv term) ≝
+                                λR. ∀L,U. R L U → ∀b,f,K. ⬇*[b, f] L ≡ K → 𝐔⦃f⦄ →
+                                ∀T. ⬆*[f] T ≡ U → R K T.
+
 definition d_liftable2_sn: ∀C:Type[0]. ∀S:rtmap → relation C.
                            predicate (lenv → relation C) ≝
                            λC,S,R. ∀K,T1,T2. R K T1 T2 → ∀b,f,L. ⬇*[b, f] L ≡ K →
@@ -235,7 +243,6 @@ lemma drops_fwd_isid: ∀b,f,L1,L2. ⬇*[b, f] L1 ≡ L2 → 𝐈⦃f⦄ → L1
 ]
 qed-.
 
-
 lemma drops_after_fwd_drop2: ∀b,f2,I,X,K. ⬇*[b, f2] X ≡ K.ⓘ{I} →
                              ∀f1,f. 𝐈⦃f1⦄ → f2 ⊚ ⫯f1 ≡ f → ⬇*[b, f] X ≡ K.
 #b #f2 #I #X #K #H #f1 #f #Hf1 #Hf elim (drops_fwd_drop2 … H) -H