]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma
previous lemma proved ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / frees_drops.ma
index c5762eacbbe202d8c246ec8f26d0bb63782efc5f..f153ad53cab0704ed87278833552481d6705f9ff 100644 (file)
@@ -14,6 +14,7 @@
 
 include "ground_2/relocation/nstream_coafter.ma".
 include "basic_2/relocation/drops_drops.ma".
+include "basic_2/static/frees_fqup.ma".
 include "basic_2/static/frees_frees.ma".
 
 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
@@ -153,61 +154,26 @@ lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
 
 (* Inversion lemmas with generic slicing for local environments *************)
 
+lemma frees_inv_lifts_ex: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
+                          ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U →
+                          ∃∃f1. f ~⊚ f1 ≡ f2 & K ⊢ 𝐅*⦃T⦄ ≡ f1.
+#b #f2 #L #U #Hf2 #f #K #HLK #T elim (frees_total K T)
+/3 width=9 by frees_fwd_coafter, ex2_intro/
+qed-.
+
+lemma frees_inv_lifts_SO: ∀b,f,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f →
+                          ∀K. ⬇*[b, 𝐔❴1❵] L ≡ K → ∀T. ⬆*[1] T ≡ U →
+                          K ⊢ 𝐅*⦃T⦄ ≡ ⫱f.
+#b #f #L #U #H #K #HLK #T #HTU elim(frees_inv_lifts_ex … H … HLK … HTU) -b -L -U
+#f1 #Hf #Hf1 elim (coafter_inv_nxx … Hf) -Hf
+/3 width=5 by frees_eq_repl_back, coafter_isid_inv_sn/
+qed-.
+
 lemma frees_inv_lifts: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →
                        ∀f,K. ⬇*[b, f] L ≡ K → ∀T. ⬆*[f] T ≡ U →
                        ∀f1. f ~⊚ f1 ≡ f2 → K ⊢ 𝐅*⦃T⦄ ≡ f1.
-#b #f2 #L #U #H lapply (frees_fwd_isfin … H) elim H -f2 -L -U
-[ #f2 #I #Hf2 #_ #f #K #H1 #T #H2 #f1 #H3
-  lapply (coafter_fwd_isid2 … H3 … Hf2) -H3 // -Hf2 #Hf1
-  elim (drops_inv_atom1 … H1) -H1 #H #_ destruct
-  elim (lifts_inv_atom2 … H2) -H2 * /2 width=3 by frees_atom/
-| #f2 #I #L #W #s #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3
-  lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
-  lapply (lifts_inv_sort2 … H2) -H2 #H destruct
-  elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ]
-  [ #g #g1 #Hf2 #H #H0 destruct
-    elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct
-  | #g #Hf2 #H destruct
-    lapply (drops_inv_drop1 … H1) -H1
-  ] /3 width=4 by frees_sort/
-| #f2 #I #L #W #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3
-  lapply (isfin_inv_next … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
-  elim (lifts_inv_lref2 … H2) -H2 #i #H2 #H destruct
-  lapply (at_inv_xxp … H2 ?) -H2 // * #g #H #H0 destruct
-  elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct
-  elim (coafter_inv_pxn … H3) -H3 [ |*: // ] #g1 #Hf2 #H destruct
-  /3 width=4 by frees_zero/
-| #f2 #I #L #W #j #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3
-  lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
-  elim (lifts_inv_lref2 … H2) -H2 #x #H2 #H destruct
-  elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ]
-  [ #g #g1 #Hf2 #H #H0 destruct
-    elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #HVW #H destruct
-    elim (at_inv_xpn … H2) -H2 [ |*: // ] #j #Hg #H destruct
-  | #g #Hf2 #H destruct
-    lapply (drops_inv_drop1 … H1) -H1
-    lapply (at_inv_xnn … H2 ????) -H2 [5: |*: // ]
-  ] /4 width=4 by lifts_lref, frees_lref/
-| #f2 #I #L #W #l #_ #IH #Hf2 #f #Y #H1 #T #H2 #f1 #H3
-  lapply (isfin_fwd_push … Hf2 ??) -Hf2 [3: |*: // ] #Hf2
-  lapply (lifts_inv_gref2 … H2) -H2 #H destruct
-  elim (coafter_inv_xxp … H3) -H3 [1,3: * |*: // ]
-  [ #g #g1 #Hf2 #H #H0 destruct
-    elim (drops_inv_skip1 … H1) -H1 #K #V #HLK #_ #H destruct
-  | #g #Hf2 #H destruct
-    lapply (drops_inv_drop1 … H1) -H1
-  ] /3 width=4 by frees_gref/
-| #f2W #f2U #f2 #p #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3
-  elim (sor_inv_isfin3 … H1f2) // #H1f2W #H
-  lapply (isfin_inv_tl … H) -H
-  elim (lifts_inv_bind2 … H2) -H2 #V #T #HVW #HTU #H destruct
-  elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 // #f1W #f1U #H2f2W #H
-  elim (coafter_inv_tl0 … H) -H /4 width=5 by frees_bind, drops_skip/
-| #f2W #f2U #f2 #I #L #W #U #_ #_ #H1f2 #IHW #IHU #H2f2 #f #K #H1 #X #H2 #f1 #H3
-  elim (sor_inv_isfin3 … H1f2) //
-  elim (lifts_inv_flat2 … H2) -H2 #V #T #HVW #HTU #H destruct
-  elim (coafter_inv_sor … H3 … H1f2) -H3 -H1f2 /3 width=5 by frees_flat/
-]
+#b #f2 #L #U #H #f #K #HLK #T #HTU #f1 #Hf2 elim (frees_inv_lifts_ex … H … HLK … HTU) -b -L -U
+/3 width=7 by frees_eq_repl_back, coafter_inj/
 qed-.
 
 lemma frees_inv_drops: ∀f2,L,U. L ⊢ 𝐅*⦃U⦄ ≡ f2 →