]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2A/multiple/lifts_lift.ma
update in binararies for λδ
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / multiple / lifts_lift.ma
index 52f1a63954a81e16eef4dc018f20da8756ba016b..c831d85fe4af20b3e04ae2ac8fffba78a865697a 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
+include "ground/relocation/mr2_minus.ma".
 include "basic_2A/substitution/lift_lift.ma".
-include "basic_2A/multiple/mr2_minus.ma".
 include "basic_2A/multiple/lifts.ma".
 
 (* GENERIC TERM RELOCATION **************************************************)
 
 (* Properties concerning basic term relocation ******************************)
 
-(* Basic_1: was: lift1_xhg (right to left) *)
 lemma lifts_lift_trans_le: ∀T1,T,cs. ⬆*[cs] T1 ≡ T → ∀T2. ⬆[0, 1] T ≡ T2 →
                            ∃∃T0. ⬆[0, 1] T1 ≡ T0 & ⬆*[cs + 1] T0 ≡ T2.
 #T1 #T #cs #H elim H -T1 -T -cs
@@ -31,9 +30,8 @@ lemma lifts_lift_trans_le: ∀T1,T,cs. ⬆*[cs] T1 ≡ T → ∀T2. ⬆[0, 1] T
 ]
 qed-.
 
-(* Basic_1: was: lift1_free (right to left) *)
-lemma lifts_lift_trans: ∀cs,i,i0. @⦃i, cs⦄ ≡ i0 →
-                        ∀cs0. cs + 1 ▭ i + 1 ≡ cs0 + 1 →
+lemma lifts_lift_trans: ∀cs,i,i0. @❪i, cs❫ ≘ i0 →
+                        ∀cs0. cs + 1 ▭ i + 1 ≘ cs0 + 1 →
                         ∀T1,T0. ⬆*[cs0] T1 ≡ T0 →
                         ∀T2. ⬆[O, i0 + 1] T0 ≡ T2 →
                         ∃∃T. ⬆[0, i + 1] T1 ≡ T & ⬆*[cs] T ≡ T2.