]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma
update in ground_2 and basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / frees_drops.ma
index 476a459c87c71bb2dd3d9547c088062fa7f5e02f..096683f573a8076cbf70e5b635516d4ab5424979 100644 (file)
@@ -77,20 +77,20 @@ lemma frees_inv_lref_drops: ∀L,i,f. L ⊢ 𝐅*⦃#i⦄ ≡ f →
                             ∨∨ ∃∃g. ⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ & 𝐈⦃g⦄ & f = ↑*[i] ⫯g
                              | ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g &
                                           ⬇*[i] L ≡ K.ⓑ{I}V & f = ↑*[i] ⫯g
-                             | ∃∃g,I,K. ⬇*[i] L ≡ K.ⓤ{I} & f = ↑*[i] ⫯g.
+                             | ∃∃g,I,K. ⬇*[i] L ≡ K.ⓤ{I} & 𝐈⦃g⦄ & f = ↑*[i] ⫯g.
 #L elim L -L
 [ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H
 [ elim (frees_inv_atom … H) -H #f #Hf #H destruct
   /3 width=3 by or3_intro0, ex3_intro/
 | elim (frees_inv_unit … H) -H #f #Hf #H destruct
-  /4 width=3 by drops_refl, or3_intro2, ex2_3_intro/
+  /4 width=3 by drops_refl, or3_intro2, ex3_3_intro/
 | elim (frees_inv_pair … H) -H #f #Hf #H destruct
   /4 width=7 by drops_refl, or3_intro1, ex3_4_intro/
 | elim (frees_inv_lref … H) -H #f #Hf #H destruct
   elim (IH … Hf) -IH -Hf *
   [ /4 width=3 by drops_drop, or3_intro0, ex3_intro/
   | /4 width=7 by drops_drop, or3_intro1, ex3_4_intro/
-  | /4 width=3 by drops_drop, or3_intro2, ex2_3_intro/
+  | /4 width=3 by drops_drop, or3_intro2, ex3_3_intro/
   ]
 ]
 qed-.
@@ -154,6 +154,12 @@ lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≡ f1 →
 ]
 qed-.
 
+lemma frees_lifts_SO: ∀b,L,K. ⬇*[b, 𝐔❴1❵] L ≡ K → ∀T,U. ⬆*[1] T ≡ U →
+                      ∀f. K ⊢ 𝐅*⦃T⦄ ≡ f → L ⊢ 𝐅*⦃U⦄ ≡ ↑f.
+#b #L #K #HLK #T #U #HTU #f #Hf
+@(frees_lifts b … Hf … HTU) //  (**) (* auto fails *)
+qed.
+
 (* Forward lemmas with generic slicing for local environments ***************)
 
 lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
@@ -185,6 +191,7 @@ lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
 /3 width=7 by frees_eq_repl_back, coafter_inj/
 qed-.
 
+(* Note: this is used by lfxs_conf and might be modified *)
 lemma frees_inv_drops_next: ∀f1,L1,T1. L1 ⊢ 𝐅*⦃T1⦄ ≡ f1 →
                             ∀I2,L2,V2,n. ⬇*[n] L1 ≡ L2.ⓑ{I2}V2 →
                             ∀g1. ⫯g1 = ⫱*[n] f1 →