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