]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/static/frees_drops.ma
renaming in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / static / frees_drops.ma
index 8edb998d2cb9c5258ee2b07a791a5e0a6745d390..fe24eadf5d0496b0dce3a9145e2e441f55e4174b 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "ground_2/relocation/rtmap_pushs.ma".
-include "basic_2/relocation/drops.ma".
-include "basic_2/static/frees.ma".
+include "ground_2/relocation/nstream_coafter.ma".
+include "basic_2/relocation/drops_drops.ma".
+include "basic_2/static/frees_fqup.ma".
 
 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
 
 (* Advanced properties ******************************************************)
 
-lemma drops_atom_F: ∀f. ⬇*[Ⓕ, f] ⋆ ≡ ⋆.
-#f @drops_atom #H destruct
+lemma frees_atom_drops: ∀b,L,i. ⬇*[b, 𝐔❴i❵] L ≘ ⋆ →
+                        ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i]↑f.
+#b #L elim L -L /2 width=1 by frees_atom/
+#L #I #IH *
+[ #H lapply (drops_fwd_isid … H ?) -H // #H destruct
+| /4 width=3 by frees_lref, drops_inv_drop1/
+]
 qed.
 
-lemma frees_inv_lref_drops: ∀i,f,L. L ⊢ 𝐅*⦃#i⦄ ≡ f →
-                            (⬇*[Ⓕ, 𝐔❴i❵] L ≡ ⋆ ∧ 𝐈⦃f⦄) ∨
-                            ∃∃g,I,K,V. K ⊢ 𝐅*⦃V⦄ ≡ g &
-                                       ⬇*[i] L ≡ K.ⓑ{I}V & f = ↑*[i] ⫯g.
-#i elim i -i
-[ #f #L #H elim (frees_inv_zero … H) -H *
-  /4 width=7 by ex3_4_intro, or_introl, or_intror, conj, drops_refl/
-| #i #IH #f #L #H elim (frees_inv_lref … H) -H * /3 width=1 by or_introl, conj/
-  #g #I #K #V #Hg #H1 #H2 destruct
-  elim (IH … Hg) -IH -Hg *
-  [ /4 width=3 by or_introl, conj, isid_push, drops_drop/
-  | /4 width=7 by drops_drop, ex3_4_intro, or_intror/
-  ]
+lemma frees_pair_drops: ∀f,K,V. K ⊢ 𝐅*⦃V⦄ ≘ f → 
+                        ∀i,I,L. ⬇*[i] L ≘ K.ⓑ{I}V → L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i] ↑f.
+#f #K #V #Hf #i elim i -i
+[ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_pair/
+| #i #IH #I #L #H elim (drops_inv_succ … H) -H /3 width=2 by frees_lref/
 ]
-qed-.
-
+qed.
 
+lemma frees_unit_drops: ∀f.  𝐈⦃f⦄ → ∀I,K,i,L. ⬇*[i] L ≘ K.ⓤ{I} →
+                       L ⊢ 𝐅*⦃#i⦄ ≘ ⫯*[i] ↑f.
+#f #Hf #I #K #i elim i -i
+[ #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_unit/
+| #i #IH #Y #H elim (drops_inv_succ … H) -H
+  #J #L #HLK #H destruct /3 width=1 by frees_lref/
+]
+qed.
+(*
+lemma frees_sort_pushs: ∀f,K,s. K ⊢ 𝐅*⦃⋆s⦄ ≘ f →
+                        ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃⋆s⦄ ≘ ⫯*[i] f.
+#f #K #s #Hf #i elim i -i
+[ #L #H lapply (drops_fwd_isid … H ?) -H //
+| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/
+]
+qed.
+*)
+lemma frees_lref_pushs: ∀f,K,j. K ⊢ 𝐅*⦃#j⦄ ≘ f →
+                        ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃#(i+j)⦄ ≘ ⫯*[i] f.
+#f #K #j #Hf #i elim i -i
+[ #L #H lapply (drops_fwd_isid … H ?) -H //
+| #i #IH #L #H elim (drops_inv_succ … H) -H
+  #I #Y #HYK #H destruct /3 width=1 by frees_lref/
+]
+qed.
+(*
+lemma frees_gref_pushs: ∀f,K,l. K ⊢ 𝐅*⦃§l⦄ ≘ f →
+                        ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅*⦃§l⦄ ≘ ⫯*[i] f.
+#f #K #l #Hf #i elim i -i
+[ #L #H lapply (drops_fwd_isid … H ?) -H //
+| #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/
+]
+qed.
+*)
+(* Advanced inversion lemmas ************************************************)
 
-lemma frees_dec: ∀L,U,l,i. Decidable (frees l L U i).
-#L #U @(f2_ind … rfw … L U) -L -U
-#x #IH #L * *
-[ -IH /3 width=5 by frees_inv_sort, or_intror/
-| #j #Hx #l #i elim (ylt_split_eq i j) #Hji
-  [ -x @or_intror #H elim (ylt_yle_false … Hji)
-    lapply (frees_inv_lref_ge … H ?) -L -l /2 width=1 by ylt_fwd_le/
-  | -x /2 width=1 by or_introl/
-  | elim (ylt_split j l) #Hli
-    [ -x @or_intror #H elim (ylt_yle_false … Hji)
-      lapply (frees_inv_lref_skip … H ?) -L //
-    | elim (lt_or_ge j (|L|)) #Hj
-      [ elim (drop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
-        elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, drop_fwd_rfw, or_introl/ ] #HnW
-        @or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -l
-        lapply (drop_mono … HLY … HLK) -L #H destruct /2 width=1 by/  
-      | -x @or_intror #H elim (ylt_yle_false … Hji)
-        lapply (frees_inv_lref_free … H ?) -l //
-      ]
-    ]
+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} & 𝐈⦃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, 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, ex3_3_intro/
   ]
-| -IH /3 width=5 by frees_inv_gref, or_intror/
-| #a #I #W #U #Hx #l #i destruct
-  elim (IH L W … l i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
-  elim (IH (L.ⓑ{I}W) U … (⫯l) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
-  @or_intror #H elim (frees_inv_bind … H) -H /2 width=1 by/
-| #I #W #U #Hx #l #i destruct
-  elim (IH L W … l i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
-  elim (IH L U … l i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
-  @or_intror #H elim (frees_inv_flat … H) -H /2 width=1 by/
 ]
 qed-.
 
-lemma frees_S: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[yinj l]⦃U⦄ → ∀I,K,W. ⬇[l] L ≡ K.ⓑ{I}W →
-               (K ⊢ ⫰(i-l) ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯l]⦃U⦄.
-#L #U #l #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
-* #I #K #W #j #Hlj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
-lapply (yle_inv_inj … Hlj) -Hlj #Hlj
-elim (le_to_or_lt_eq … Hlj) -Hlj
-[ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
-| -Hji -HnU #H destruct
-  lapply (drop_mono … HLK0 … HLK) #H destruct -I
-  elim HnW0 -L -U -HnW0 //
+(* Properties with generic slicing for local environments *******************)
+
+lemma frees_lifts: ∀b,f1,K,T. K ⊢ 𝐅*⦃T⦄ ≘ f1 →
+                   ∀f,L. ⬇*[b, f] L ≘ K → ∀U. ⬆*[f] T ≘ U →
+                   ∀f2. f ~⊚ f1 ≘ f2 → L ⊢ 𝐅*⦃U⦄ ≘ f2.
+#b #f1 #K #T #H lapply (frees_fwd_isfin … H) elim H -f1 -K -T
+[ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
+  lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
+  >(lifts_inv_sort1 … H2) -U /2 width=1 by frees_sort/
+| #f1 #i #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
+  elim (lifts_inv_lref1 … H2) -H2 #j #Hij #H destruct
+  elim (coafter_fwd_xnx_pushs … Hij H3) -H3 #g2 #Hg2 #H2 destruct
+  lapply (coafter_isid_inv_dx … Hg2 … Hf1) -f1 #Hf2
+  elim (drops_inv_atom2 … H1) -H1 #n #g #H1 #Hf
+  elim (after_at_fwd … Hij … Hf) -f #x #_ #Hj -g -i
+  lapply (at_inv_uni … Hj) -Hj #H destruct
+  /3 width=8 by frees_atom_drops, drops_trans/
+| #f1 #I #K #V #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
+  lapply (isfin_inv_next … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
+  lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
+  elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #H
+  elim (liftsb_inv_pair_sn … H) -H #W #HVW #H destruct
+  elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
+  lapply (IH … HYK … HVW … H3) -IH -H3 -HYK -HVW //
+  /2 width=5 by frees_pair_drops/
+| #f1 #I #K #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
+  lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
+  elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
+  lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hg2
+  elim (drops_split_trans_bind2 … H1 … Hf) -H1 -Hf #Z #Y #HLY #_ #H
+  lapply (liftsb_inv_unit_sn … H) -H #H destruct
+  /2 width=3 by frees_unit_drops/
+| #f1 #I #K #i #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
+  lapply (isfin_inv_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
+  lapply (lifts_inv_lref1 … H2) -H2 * #x #Hf #H destruct
+  elim (at_inv_nxx … Hf) -Hf [ |*: // ] #j #Hf #H destruct
+  elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #_
+  elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H3 #H2 destruct
+  lapply (drops_isuni_fwd_drop2 … HLY) -HLY // #HLY
+  lapply (IH … HYK … H3) -IH -H3 -HYK [4: |*: /2 width=2 by lifts_lref/ ]
+  >plus_S1 /2 width=3 by frees_lref_pushs/ (**) (* full auto fails *)
+| #f1 #K #l #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
+  lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
+  >(lifts_inv_gref1 … H2) -U /2 width=1 by frees_gref/
+| #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
+  elim (sor_inv_isfin3 … H1f1) // #Hf1V #H
+  lapply (isfin_inv_tl … H) -H
+  elim (lifts_inv_bind1 … H2) -H2 #W #U #HVW #HTU #H destruct
+  elim (coafter_sor … H3 … H1f1) /2 width=5 by coafter_isfin2_fwd/ -H3 -H1f1 #f2V #f2T #Hf2V #H
+  elim (coafter_inv_tl1 … H) -H
+  /5 width=5 by frees_bind, drops_skip, ext2_pair/
+| #f1V #f1T #f1 #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
+  elim (sor_inv_isfin3 … H1f1) //
+  elim (lifts_inv_flat1 … H2) -H2 #W #U #HVW #HTU #H destruct
+  elim (coafter_sor … H3 … H1f1)
+  /3 width=5 by coafter_isfin2_fwd, frees_flat/
 ]
-qed.
+qed-.
 
-(* Note: lemma 1250 *)
-lemma frees_bind_dx_O: ∀a,I,L,W,U,i. L.ⓑ{I}W ⊢ ⫯i ϵ 𝐅*[0]⦃U⦄ →
-                       L ⊢ i ϵ 𝐅*[0]⦃ⓑ{a,I}W.U⦄.
-#a #I #L #W #U #i #HU elim (frees_dec L W 0 i)
-/4 width=5 by frees_S, frees_bind_dx, frees_bind_sn/
+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.
 
-(* Properties on relocation *************************************************)
-
-lemma frees_lift_ge: ∀K,T,l,i. K ⊢ i ϵ𝐅*[l]⦃T⦄ →
-                     ∀L,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
-                     ∀U. ⬆[l0, m0] T ≡ U → l0 ≤ i →
-                     L ⊢ i+m0 ϵ 𝐅*[l]⦃U⦄.
-#K #T #l #i #H elim H -K -T -l -i
-[ #K #T #l #i #HnT #L #s #l0 #m0 #_ #U #HTU #Hl0i -K -s
-  @frees_eq #X #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
-| #I #K #K0 #T #V #l #i #j #Hlj #Hji #HnT #HK0 #HV #IHV #L #s #l0 #m0 #HLK #U #HTU #Hl0i
-  elim (ylt_split j l0) #H0
-  [ elim (drop_trans_lt … HLK … HK0) // -K #L0 #W #HL0 >yminus_SO2 #HLK0 #HVW
-    @(frees_be … HL0) -HL0 -HV /3 width=3 by ylt_plus_dx2_trans/
-    [ lapply (ylt_fwd_lt_O1 … H0) #H1
-      #X #HXU <(ymax_pre_sn l0 1) in HTU; /2 width=1 by ylt_fwd_le_succ1/ #HTU
-      <(ylt_inv_O1 l0) in H0; // -H1 #H0
-      elim (lift_div_le … HXU … HTU ?) -U /2 width=2 by ylt_fwd_succ2/
-    | >yplus_minus_comm_inj /2 width=1 by ylt_fwd_le/
-      <yplus_pred1 /4 width=5 by monotonic_yle_minus_dx, yle_pred, ylt_to_minus/
-    ]
-  | lapply (drop_trans_ge … HLK … HK0 ?) // -K #HLK0
-    lapply (drop_inv_gen … HLK0) >commutative_plus -HLK0 #HLK0
-    @(frees_be … HLK0) -HLK0 -IHV
-    /2 width=1 by monotonic_ylt_plus_dx, yle_plus_dx1_trans/
-    [ #X <yplus_inj #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
-    | <yplus_minus_assoc_comm_inj //
-    ]
-  ]
-]
-qed.
+(* Forward lemmas with generic slicing for local environments ***************)
 
-(* Inversion lemmas on relocation *******************************************)
-
-lemma frees_inv_lift_be: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
-                         ∀K,s,l0,m0. ⬇[s, l0, m0+1] L ≡ K →
-                         ∀T. ⬆[l0, m0+1] T ≡ U → l0 ≤ i → i ≤ l0 + m0 → ⊥.
-#L #U #l #i #H elim H -L -U -l -i
-[ #L #U #l #i #HnU #K #s #l0 #m0 #_ #T #HTU #Hl0i #Hilm0
-  elim (lift_split … HTU i m0) -HTU /2 width=2 by/
-| #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hl0i #Hilm0
-  elim (ylt_split j l0) #H1
-  [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
-    @(IHW … HKL0 … HVW)
-    [ /3 width=1 by monotonic_yle_minus_dx, yle_pred/
-    | >yplus_pred1 /2 width=1 by ylt_to_minus/
-      <yplus_minus_comm_inj /3 width=1 by monotonic_yle_minus_dx, yle_pred, ylt_fwd_le/
-    ]
-  | elim (lift_split … HTU j m0) -HTU /3 width=3 by ylt_yle_trans, ylt_fwd_le/
-  ]
-]
+lemma frees_fwd_coafter: ∀b,f2,L,U. L ⊢ 𝐅*⦃U⦄ ≘ f2 →
+                         ∀f,K. ⬇*[b, f] L ≘ K → ∀T. ⬆*[f] T ≘ U →
+                         ∀f1. K ⊢ 𝐅*⦃T⦄ ≘ f1 → f ~⊚ f1 ≘ f2.
+/4 width=11 by frees_lifts, frees_mono, coafter_eq_repl_back0/ qed-.
+
+(* 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_lift_ge: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
-                         ∀K,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
-                         ∀T. ⬆[l0, m0] T ≡ U → l0 + m0 ≤ i →
-                         K ⊢ i-m0 ϵ𝐅*[l-yinj m0]⦃T⦄.
-#L #U #l #i #H elim H -L -U -l -i
-[ #L #U #l #i #HnU #K #s #l0 #m0 #HLK #T #HTU #Hlm0i -L -s
-  elim (yle_inv_plus_inj2 … Hlm0i) -Hlm0i #Hl0im0 #Hm0i @frees_eq #X #HXT -K
-  elim (lift_trans_le … HXT … HTU) -T // >ymax_pre_sn /2 width=2 by/
-| #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hlm0i
-  elim (ylt_split j l0) #H1
-  [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
-    elim (yle_inv_plus_inj2 … Hlm0i) #H0 #Hm0i
-    @(frees_be … H) -H
-    [ /3 width=1 by yle_plus_dx1_trans, monotonic_yle_minus_dx/
-    | /2 width=3 by ylt_yle_trans/
-    | #X #HXT elim (lift_trans_ge … HXT … HTU) -T /2 width=2 by ylt_fwd_le_succ1/
-    | lapply (IHW … HKL0 … HVW ?) // -I -K -K0 -L0 -V -W -T -U -s
-      >yplus_pred1 /2 width=1 by ylt_to_minus/
-      <yplus_minus_comm_inj /3 width=1 by monotonic_yle_minus_dx, yle_pred, ylt_fwd_le/
-    ]
-  | elim (ylt_split j (l0+m0)) #H2
-    [ -L -I -W elim (yle_inv_inj2 … H1) -H1 #x #H1 #H destruct
-      lapply (ylt_plus2_to_minus_inj1 … H2) /2 width=1 by yle_inj/ #H3
-      lapply (ylt_fwd_lt_O1 … H3) -H3 #H3
-      elim (lift_split … HTU j (m0-1)) -HTU /2 width=1 by yle_inj/
-      [ >minus_minus_associative /2 width=1 by ylt_inv_inj/ <minus_n_n
-        -H2 #X #_ #H elim (HnU … H)
-      | <yminus_inj >yminus_SO2 >yplus_pred2 /2 width=1 by ylt_fwd_le_pred2/
-      ]
-    | lapply (drop_conf_ge … HLK … HLK0 ?) // -L #HK0
-      elim ( yle_inv_plus_inj2 … H2) -H2 #H2 #Hm0j
-      @(frees_be … HK0)
-      [ /2 width=1 by monotonic_yle_minus_dx/
-      | /2 width=1 by monotonic_ylt_minus_dx/
-      | #X #HXT elim (lift_trans_le … HXT … HTU) -T //
-        <yminus_inj >ymax_pre_sn /2 width=2 by/
-      | <yminus_inj >yplus_minus_assoc_comm_inj //
-        >ymax_pre_sn /3 width=5 by yle_trans, ylt_fwd_le/
-      ]
-    ]
+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 #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-.
+
+(* Note: this is used by rex_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 →
+                            ∃∃g2. L2 ⊢ 𝐅*⦃V2⦄ ≘ g2 & g2 ⊆ g1.
+#f1 #L1 #T1 #H elim H -f1 -L1 -T1
+[ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s
+  lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
+  elim (isid_inv_next … Hf1) -Hf1 //
+| #f1 #i #_ #I2 #L2 #V2 #n #H
+  elim (drops_inv_atom1 … H) -H #H destruct
+| #f1 #I1 #L1 #V1 #Hf1 #IH #I2 #L2 #V2 *
+  [ -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
+    #H destruct #g1 #Hgf1 >(injective_next … Hgf1) -g1
+    /2 width=3 by ex2_intro/
+  | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
+    #HL12 #g1 <tls_xn <tl_next_rew #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
+    /2 width=3 by ex2_intro/
+  ]
+| #f1 #I1 #L1 #Hf1 #I2 #L2 #V2 *
+  [ #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 // #H destruct
+  | #n #_ #g1 #Hgf1 elim (isid_inv_next … Hgf1) -Hgf1 <tls_xn /2 width=1 by isid_tls/
+  ]
+| #f1 #I1 #L1 #i #_ #IH #I2 #L2 #V2 *
+  [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
+  | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
+    #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
+    /2 width=3 by ex2_intro/
+  ]
+| #f1 #L1 #l #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -l
+  lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
+  elim (isid_inv_next … Hf1) -Hf1 //
+| #fV1 #fT1 #f1 #p #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
+  lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
+  elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
+  #gV1 #gT1 #Hg1
+  [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
+    /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
+  | -IHV1 #_ >tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2
+    /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/
+  ]
+| #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
+  lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
+  elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
+  #gV1 #gT1 #Hg1
+  [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
+    /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
+  | -IHV1 #_ #H2 elim (IHT1 … HL12 … H2) -IHT1 -HL12 -H2
+    /3 width=6 by sor_inv_sle_dx_trans, ex2_intro/
   ]
 ]
 qed-.