]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/etc/sta/sta_lift.etc
milestone update in ground_2 and basic_2A
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / etc / sta / sta_lift.etc
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 "basic_2/substitution/drop_drop.ma".
16 include "basic_2/static/sta.ma".
17
18 (* STATIC TYPE ASSIGNMENT ON TERMS ******************************************)
19
20 (* Properties on relocation *************************************************)
21
22 (* Basic_1: was: sty0_lift *)
23 lemma sta_lift: ∀h,G. l_liftable (sta h G).
24 #h #G #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1
25 [ #G #L1 #k #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
26   >(lift_inv_sort1 … H1) -X1
27   >(lift_inv_sort1 … H2) -X2 //
28 | #G #L1 #K1 #V1 #W1 #W #i #HLK1 #_ #HW1 #IHVW1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
29   elim (lift_inv_lref1 … H) * #Hid #H destruct
30   [ elim (lift_trans_ge … HW1 … HWU2) -W // #W2 #HW12 #HWU2
31     elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
32     elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #V2 #HK21 #HV12 #H destruct
33     /3 width=9 by sta_ldef/
34   | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
35     lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldef, drop_inv_gen/
36   ]
37 | #G #L1 #K1 #W1 #V1 #W #i #HLK1 #_ #HW1 #IHWV1 #L2 #s #d #e #HL21 #X #H #U2 #HWU2
38   elim (lift_inv_lref1 … H) * #Hid #H destruct
39   [ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
40     elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
41     elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hid #K2 #W2 #HK21 #HW12 #H destruct
42     lapply (lift_mono … HW1 … HW12) -HW1 #H destruct
43     elim (lift_total V1 (d-i-1) e) /3 width=9 by sta_ldec/
44   | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
45     lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by sta_ldec, drop_inv_gen/
46   ]
47 | #a #I #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
48   elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
49   elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
50   lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_bind, drop_skip/
51 | #G #L1 #V1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X1 #H1 #X2 #H2
52   elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
53   elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
54   lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by sta_appl/
55 | #G #L1 #W1 #T1 #U1 #_ #IHTU1 #L2 #s #d #e #HL21 #X #H #U2 #HU12
56   elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by sta_cast/
57 ]
58 qed.
59
60 (* Note: apparently this was missing in basic_1 *)
61 lemma sta_inv_lift1: ∀h,G. l_deliftable_sn (sta h G).
62 #h #G #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2
63 [ #G #L2 #k #L1 #s #d #e #_ #X #H
64   >(lift_inv_sort2 … H) -X /2 width=3 by sta_sort, lift_sort, ex2_intro/
65 | #G #L2 #K2 #V2 #W2 #W #i #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #d #e #HL21 #X #H
66   elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HVW2 | -IHVW2 ]
67   [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
68     elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
69     elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldef, ex2_intro/
70   | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
71     elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
72     elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
73     #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldef, le_S, ex2_intro/
74   ]
75 | #G #L2 #K2 #W2 #V2 #W #i #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #d #e #HL21 #X #H
76   elim (lift_inv_lref2 … H) * #Hid #H destruct [ -HWV2 | -IHWV2 ]
77   [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
78     elim (IHWV2 … HK21 … HW12) -K2 #V1 #_ #HWV1
79     elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by sta_ldec, ex2_intro/
80   | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
81     elim (le_inv_plus_l … Hid) -Hid #Hdie #ei
82     elim (lift_split … HW2 d (i-e+1)) -HW2 /2 width=1 by le_S_S, le_S/
83     #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by sta_ldec, le_S, ex2_intro/
84   ]
85 | #a #I #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
86   elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
87   elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by sta_bind, drop_skip, lift_bind, ex2_intro/
88 | #G #L2 #V2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
89   elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
90   elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by sta_appl, lift_flat, ex2_intro/
91 | #G #L2 #W2 #T2 #U2 #_ #IHTU2 #L1 #s #d #e #HL21 #X #H
92   elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
93   elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by sta_cast, ex2_intro/
94 ]
95 qed-.
96
97 (* Advanced forward lemmas **************************************************)
98
99 (* Basic_1: was: sty0_correct *)
100 lemma sta_fwd_correct: ∀h,G,L,T,U. ⦃G, L⦄ ⊢ T •[h] U → ∃T0. ⦃G, L⦄ ⊢ U •[h] T0.
101 #h #G #L #T #U #H elim H -G -L -T -U
102 [ /2 width=2/
103 | #G #L #K #V #W #W0 #i #HLK #_ #HW0 * #V0 #HWV0
104   lapply (drop_fwd_drop2 … HLK) -HLK #HLK
105   elim (lift_total V0 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
106 | #G #L #K #W #V #V0 #i #HLK #HWV #HWV0 #_
107   lapply (drop_fwd_drop2 … HLK) -HLK #HLK
108   elim (lift_total V 0 (i+1)) /3 width=11 by ex_intro, sta_lift/
109 | #a #I #G #L #V #T #U #_ * /3 width=2 by sta_bind, ex_intro/
110 | #G #L #V #T #U #_ * #T0 #HUT0 /3 width=2 by sta_appl, ex_intro/
111 | #G #L #W #T #U #_ * /2 width=2 by ex_intro/
112 ]
113 qed-.