]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/multiple/lifts_lift.ma
2c693baaa79f566c131d6991c96b02ce42b007e3
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / multiple / lifts_lift.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "ground_2/ynat/ynat_max.ma".
16 include "basic_2/substitution/lift_lift.ma".
17 include "basic_2/multiple/mr2_minus.ma".
18 include "basic_2/multiple/lifts.ma".
19
20 (* GENERIC TERM RELOCATION **************************************************)
21
22 (* Properties concerning basic term relocation ******************************)
23
24 (* Basic_1: was: lift1_xhg (right to left) *)
25 lemma lifts_lift_trans_le: ∀T1,T,cs. ⬆*[cs] T1 ≡ T → ∀T2. ⬆[0, 1] T ≡ T2 →
26                            ∃∃T0. ⬆[0, 1] T1 ≡ T0 & ⬆*[cs + 1] T0 ≡ T2.
27 #T1 #T #cs #H elim H -T1 -T -cs
28 [ /2 width=3 by lifts_nil, ex2_intro/
29 | #T1 #T3 #T #cs #l #m #HT13 #_ #IHT13 #T2 #HT2
30   elim (IHT13 … HT2) -T #T #HT3 #HT2
31   elim (lift_trans_le … HT13 … HT3) -T3 /3 width=5 by lifts_cons, ex2_intro/
32 ]
33 qed-.
34
35 (* Basic_1: was: lift1_free (right to left) *)
36 lemma lifts_lift_trans: ∀cs,i,i0. @⦃i, cs⦄ ≡ i0 →
37                         ∀cs0. cs + 1 ▭ i + 1 ≡ cs0 + 1 →
38                         ∀T1,T0. ⬆*[cs0] T1 ≡ T0 →
39                         ∀T2. ⬆[O, i0 + 1] T0 ≡ T2 →
40                         ∃∃T. ⬆[0, i + 1] T1 ≡ T & ⬆*[cs] T ≡ T2.
41 #cs elim cs -cs
42 [ #i #x #H1 #cs0 #H2 #T1 #T0 #HT10 #T2
43   <(at_inv_nil … H1) -x #HT02
44   lapply (minuss_inv_nil1 … H2) -H2 #H
45   >(pluss_inv_nil2 … H) in HT10; -cs0 #H
46   >(lifts_inv_nil … H) -T1 /2 width=3 by lifts_nil, ex2_intro/
47 | #l #m #cs #IHcs #i #i0 #H1 #cs0 #H2 #T1 #T0 #HT10 #T2 #HT02
48   elim (at_inv_cons … H1) -H1 * #Hil #Hi0
49   [ elim (minuss_inv_cons1_lt … H2) -H2 /2 width=1 by ylt_succ/ #cs1 #Hcs1
50     <yplus_inj >yplus_SO2 >yplus_SO2 >yminus_succ #H
51     elim (pluss_inv_cons2 … H) -H #cs2 #H1 #H2 destruct
52     elim (lifts_inv_cons … HT10) -HT10 #T #HT1 #HT0
53     elim (IHcs … Hi0 … Hcs1 … HT0 … HT02) -IHcs -Hi0 -Hcs1 -T0 #T0 #HT0 #HT02
54     elim (lift_trans_le … HT1 … HT0) -T /2 width=1 by/ #T #HT1
55     <yminus_plus2 <yplus_inj >yplus_SO2 >ymax_pre_sn
56     /3 width=5 by lifts_cons, ex2_intro, ylt_fwd_le_succ1/
57   | >commutative_plus in Hi0; #Hi0
58     lapply (minuss_inv_cons1_ge … H2 ?) -H2 /2 width=1 by yle_succ/ <associative_plus #Hcs0
59     elim (IHcs … Hi0 … Hcs0 … HT10 … HT02) -IHcs -Hi0 -Hcs0 -T0 #T0 #HT0 #HT02
60     elim (lift_split … HT0 l (i+1)) -HT0 /3 width=5 by lifts_cons, yle_succ, yle_pred_sn, ex2_intro/
61   ]
62 ]
63 qed-.