]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/unfold/lstas_lift.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / unfold / lstas_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 "basic_2A/substitution/drop_drop.ma".
16 include "basic_2A/unfold/lstas.ma".
17
18 (* NAT-ITERATED STATIC TYPE ASSIGNMENT FOR TERMS ****************************)
19
20 (* Properties on relocation *************************************************)
21
22 (* Basic_1: was just: sty0_lift *)
23 lemma lstas_lift: ∀h,G,d. d_liftable (lstas h G d).
24 #h #G #d #L1 #T1 #U1 #H elim H -G -L1 -T1 -U1 -d
25 [ #G #L1 #d #k #L2 #s #l #m #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 #d #HLK1 #_ #HW1 #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
29   elim (lift_inv_lref1 … H) * #Hil #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/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
33     /3 width=9 by lstas_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 lstas_ldef, drop_inv_gen/
36   ]
37 | #G #L1 #K1 #V1 #W1 #i #HLK1 #_ #IHVW1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
38   >(lift_mono … HWU2 … H) -U2
39   elim (lift_inv_lref1 … H) * #Hil #H destruct
40   [ elim (lift_total W1 (l-i-1) m) #W2 #HW12
41     elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
42     elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #V2 #HK21 #HV12 #H destruct
43     /3 width=10 by lstas_zero/
44   | lapply (drop_trans_ge … HL21 … HLK1 ?) -L1
45     /3 width=10 by lstas_zero, drop_inv_gen/
46   ]
47 | #G #L1 #K1 #W1 #V1 #W #i #d #HLK1 #_ #HW1 #IHWV1 #L2 #s #l #m #HL21 #X #H #U2 #HWU2
48   elim (lift_inv_lref1 … H) * #Hil #H destruct
49   [ elim (lift_trans_ge … HW1 … HWU2) -W // <minus_plus #W #HW1 #HWU2
50     elim (drop_trans_le … HL21 … HLK1) -L1 /2 width=2 by lt_to_le/ #X #HLK2 #H
51     elim (drop_inv_skip2 … H) -H /2 width=1 by lt_plus_to_minus_r/ -Hil #K2 #W2 #HK21 #HW12 #H destruct
52     /3 width=9 by lstas_succ/
53   | lapply (lift_trans_be … HW1 … HWU2 ? ?) -W /2 width=1 by le_S/ #HW1U2
54     lapply (drop_trans_ge … HL21 … HLK1 ?) -L1 /3 width=9 by lstas_succ, drop_inv_gen/
55   ]
56 | #a #I #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
57   elim (lift_inv_bind1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
58   elim (lift_inv_bind1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
59   lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by lstas_bind, drop_skip/
60 | #G #L1 #V1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X1 #H1 #X2 #H2
61   elim (lift_inv_flat1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H destruct
62   elim (lift_inv_flat1 … H2) -H2 #X #U2 #H1 #HU12 #H2 destruct
63   lapply (lift_mono … H1 … HV12) -H1 #H destruct /4 width=6 by lstas_appl/
64 | #G #L1 #W1 #T1 #U1 #d #_ #IHTU1 #L2 #s #l #m #HL21 #X #H #U2 #HU12
65   elim (lift_inv_flat1 … H) -H #W2 #T2 #_ #HT12 #H destruct /3 width=6 by lstas_cast/
66 ]
67 qed.
68
69 (* Inversion lemmas on relocation *******************************************)
70
71 (* Note: apparently this was missing in basic_1 *)
72 lemma lstas_inv_lift1: ∀h,G,d. d_deliftable_sn (lstas h G d).
73 #h #G #d #L2 #T2 #U2 #H elim H -G -L2 -T2 -U2 -d
74 [ #G #L2 #d #k #L1 #s #l #m #_ #X #H
75   >(lift_inv_sort2 … H) -X /2 width=3 by lstas_sort, lift_sort, ex2_intro/
76 | #G #L2 #K2 #V2 #W2 #W #i #d #HLK2 #HVW2 #HW2 #IHVW2 #L1 #s #l #m #HL21 #X #H
77   elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HVW2 | -IHVW2 ]
78   [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #V1 #HLK1 #HK21 #HV12
79     elim (IHVW2 … HK21 … HV12) -K2 -V2 #W1 #HW12 #HVW1
80     elim (lift_trans_le … HW12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by lstas_ldef, ex2_intro/
81   | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
82     elim (le_inv_plus_l … Hil) -Hil #Hlim #mi
83     elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by le_S_S, le_S/
84     #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by lstas_ldef, le_S, ex2_intro/
85   ]
86 | #G #L2 #K2 #W2 #V2 #i #HLK2 #HWV2 #IHWV2 #L1 #s #l #m #HL21 #X #H
87   elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
88   [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
89     elim (IHWV2 … HK21 … HW12) -K2
90     /3 width=5 by lstas_zero, lift_lref_lt, ex2_intro/
91   | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2
92     /3 width=5 by lstas_zero, lift_lref_ge_minus, ex2_intro/
93   ]
94 | #G #L2 #K2 #W2 #V2 #W #i #d #HLK2 #HWV2 #HW2 #IHWV2 #L1 #s #l #m #HL21 #X #H
95   elim (lift_inv_lref2 … H) * #Hil #H destruct [ -HWV2 | -IHWV2 ]
96   [ elim (drop_conf_lt … HL21 … HLK2) -L2 // #K1 #W1 #HLK1 #HK21 #HW12
97     elim (IHWV2 … HK21 … HW12) -K2 #V1 #HV12 #HWV1
98     elim (lift_trans_le … HV12 … HW2) -W2 // >minus_plus <plus_minus_m_m /3 width=8 by lstas_succ, ex2_intro/
99   | lapply (drop_conf_ge … HL21 … HLK2 ?) -L2 // #HL1K2
100     elim (le_inv_plus_l … Hil) -Hil #Hlim #mi
101     elim (lift_split … HW2 l (i-m+1)) -HW2 /2 width=1 by le_S_S, le_S/
102     #W0 #HW20 <le_plus_minus_comm // >minus_minus_m_m /3 width=8 by lstas_succ, le_S, ex2_intro/
103   ]
104 | #a #I #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
105   elim (lift_inv_bind2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
106   elim (IHTU2 (L1.ⓑ{I}V1) … HT12) -IHTU2 -HT12 /3 width=5 by lstas_bind, drop_skip, lift_bind, ex2_intro/
107 | #G #L2 #V2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
108   elim (lift_inv_flat2 … H) -H #V1 #T1 #HV12 #HT12 #H destruct
109   elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=5 by lstas_appl, lift_flat, ex2_intro/
110 | #G #L2 #W2 #T2 #U2 #d #_ #IHTU2 #L1 #s #l #m #HL21 #X #H
111   elim (lift_inv_flat2 … H) -H #W1 #T1 #_ #HT12 #H destruct
112   elim (IHTU2 … HL21 … HT12) -L2 -HT12 /3 width=3 by lstas_cast, ex2_intro/
113 ]
114 qed-.
115
116 (* Advanced inversion lemmas ************************************************)
117
118 lemma lstas_split_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → ∀d1,d2. d = d1 + d2 →
119                        ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T & ⦃G, L⦄ ⊢ T •*[h, d2] T2.
120 #h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d
121 [ #G #L #d #k #d1 #d2 #H destruct
122   >commutative_plus >iter_plus /2 width=3 by lstas_sort, ex2_intro/
123 | #G #L #K #V1 #V2 #U2 #i #d #HLK #_ #VU2 #IHV12 #d1 #d2 #H destruct
124   elim (IHV12 d1 d2) -IHV12 // #V
125   elim (lift_total V 0 (i+1))
126   lapply (drop_fwd_drop2 … HLK)
127   /3 width=12 by lstas_lift, lstas_ldef, ex2_intro/
128 | #G #L #K #W1 #W2 #i #HLK #HW12 #_ #d1 #d2 #H
129   elim (zero_eq_plus … H) -H #H1 #H2 destruct
130   /3 width=5 by lstas_zero, ex2_intro/
131 | #G #L #K #W1 #W2 #U2 #i #d #HLK #HW12 #HWU2 #IHW12 #d1 @(nat_ind_plus … d1) -d1
132   [ #d2 normalize #H destruct
133     elim (IHW12 0 d) -IHW12 //
134     lapply (drop_fwd_drop2 … HLK)
135     /3 width=8 by lstas_succ, lstas_zero, ex2_intro/
136   | #d1 #_ #d2 <plus_plus_comm_23 #H lapply (injective_plus_l … H) -H #H
137     elim (IHW12 … H) -d #W
138     elim (lift_total W 0 (i+1))
139     lapply (drop_fwd_drop2 … HLK)
140     /3 width=12 by lstas_lift, lstas_succ, ex2_intro/
141   ]
142 | #a #I #G #L #V #T #U #d #_ #IHTU #d1 #d2 #H
143   elim (IHTU … H) -d /3 width=3 by lstas_bind, ex2_intro/
144 | #G #L #V #T #U #d #_ #IHTU #d1 #d2 #H
145   elim (IHTU … H) -d /3 width=3 by lstas_appl, ex2_intro/
146 | #G #L #W #T #U #d #_ #IHTU #d1 #d2 #H
147   elim (IHTU … H) -d /3 width=3 by lstas_cast, ex2_intro/
148 ]
149 qed-.
150
151 lemma lstas_split: ∀h,G,L,T1,T2,d1,d2. ⦃G, L⦄ ⊢ T1 •*[h, d1 + d2] T2 →
152                    ∃∃T. ⦃G, L⦄ ⊢ T1 •*[h, d1] T & ⦃G, L⦄ ⊢ T •*[h, d2] T2.
153 /2 width=3 by lstas_split_aux/ qed-.
154
155 (* Advanced properties ******************************************************)
156
157 lemma lstas_lstas: ∀h,G,L,T,T1,d1. ⦃G, L⦄ ⊢ T •*[h, d1] T1 →
158                    ∀d2. ∃T2. ⦃G, L⦄ ⊢ T •*[h, d2] T2.
159 #h #G #L #T #T1 #d1 #H elim H -G -L -T -T1 -d1
160 [ /2 width=2 by lstas_sort, ex_intro/
161 | #G #L #K #V #V1 #U1 #i #d1 #HLK #_ #HVU1 #IHV1 #d2
162   elim (IHV1 d2) -IHV1 #V2
163   elim (lift_total V2 0 (i+1))
164   /3 width=7 by ex_intro, lstas_ldef/
165 | #G #L #K #W #W1 #i #HLK #HW1 #IHW1 #d2
166   @(nat_ind_plus … d2) -d2 /3 width=5 by lstas_zero, ex_intro/
167   #d2 #_ elim (IHW1 d2) -IHW1 #W2
168   elim (lift_total W2 0 (i+1))
169   /3 width=7 by lstas_succ, ex_intro/
170 | #G #L #K #W #W1 #U1 #i #d #HLK #_ #_ #IHW1 #d2
171   @(nat_ind_plus … d2) -d2
172   [ elim (IHW1 0) -IHW1 /3 width=5 by lstas_zero, ex_intro/
173   | #d2 #_ elim (IHW1 d2) -IHW1
174     #W2 elim (lift_total W2 0 (i+1)) /3 width=7 by ex_intro, lstas_succ/
175   ]
176 | #a #I #G #L #V #T #U #d #_ #IHTU #d2
177   elim (IHTU d2) -IHTU /3 width=2 by lstas_bind, ex_intro/
178 | #G #L #V #T #U #d #_ #IHTU #d2
179   elim (IHTU d2) -IHTU /3 width=2 by lstas_appl, ex_intro/
180 | #G #L #W #T #U #d #_ #IHTU #d2
181   elim (IHTU d2) -IHTU /3 width=2 by lstas_cast, ex_intro/
182 ]
183 qed-.