]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_2A1/nta/nta_thin.etc
milestone uupdate in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_2A1 / nta / nta_thin.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/unfold/thin_ldrop.ma".
16 include "basic_2/equivalence/cpcs_delift.ma".
17 include "basic_2/dynamic/nta_lift.ma".
18
19 (* NATIVE TYPE ASSIGNMENT ON TERMS ******************************************)
20
21 (* Properties on basic local environment thinning ***************************)
22
23 (* Note: this is known as the substitution lemma *)
24 (* Basic_1: was only: ty3_gen_cabbr *)
25 lemma nta_thin_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
26                      ∀L2,d,e. ≽ [d, e] L1 → L1 ▼*[d, e] ≡ L2 →
27                      ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
28                               L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
29 #h #L1 #T1 #U1 #H elim H -L1 -T1 -U1
30 [ /2 width=5/
31 | #L1 #K1 #V1 #W1 #U1 #i #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL1 #HL12
32   elim (lt_or_ge i d) #Hdi [ -HVW1 ]
33   [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
34     lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
35     elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
36     elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct
37     elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #X2 #W2 #HVW2 #H #HW12
38     lapply (delift_mono … H … HV12) -H -HV12 #H destruct
39     elim (lift_total W2 0 (i+1)) #U2 #HWU2
40     lapply (ldrop_fwd_ldrop2 … HLK1) -V1 #HLK1
41     lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
42     >minus_plus <plus_minus_m_m // /3 width=6/
43   | elim (lt_or_ge i (d+e)) #Hide [ -HVW1 | -Hdi -IHVW1 -HL1 ]
44     [ lapply (sfr_ldrop_trans_be_up … HLK1 … HL1 ? ?) -HL1 // /2 width=2/ <minus_n_O #H
45       elim (sfr_inv_bind … H ?) -H /2 width=1/ #HK1 #_
46       elim (thin_ldrop_conf_be … HL12 … HLK1 ? ?) -HL12 /2 width=2/ #K2 #H #HLK2
47       lapply (thin_inv_thin1 … H ?) -H /2 width=1/ #HK12
48       elim (IHVW1 … HK1 HK12) -IHVW1 -HK1 -HK12 #V2 #W2 #HVW2 #HV12 #HW12
49       elim (lift_total V2 0 d) #T2 #HVT2
50       elim (lift_total W2 0 d) #U2 #HWU2
51       elim (lift_total W2 0 (i+1)) #U #HW2U
52       lapply (nta_lift … HVW2 … HLK2 … HVT2 … HWU2) -HVW2 -HLK2 #HTU2
53       lapply (ldrop_fwd_ldrop2 … HLK1) #HLK0
54       lapply (delift_lift_ge … HW12 … HLK0 HWU1 … HW2U) -HW12 -HLK0 -HWU1 // >minus_plus #HU1
55       lapply (lift_conf_be … HWU2 … HW2U ?) -W2 /2 width=1/ #HU2
56       lapply (delift_lift_div_be … HU1 … HU2 ? ?) -U // /2 width=1/ /3 width=8/
57     | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
58       lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
59       elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
60       <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
61       <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
62     ]
63   ]
64 | #L1 #K1 #W1 #V1 #U1 #i #HLK1 #HWV1 #HWU1 #IHWV1 #L2 #d #e #HL1 #HL12
65   elim (lt_or_ge i d) #Hdi [ -HWV1 | -IHWV1 ]
66   [ lapply (sfr_ldrop_trans_ge … HLK1 … HL1 ?) -HL1 /2 width=2/ #H
67     lapply (sfr_inv_skip … H ?) -H /2 width=1/ #HK1
68     elim (thin_ldrop_conf_le … HL12 … HLK1 ?) -HL12 /2 width=2/ #X #H #HLK2
69     elim (thin_inv_delift1 … H ?) -H /2 width=1/ #K2 #W2 #HK12 #HW12 #H destruct
70     elim (IHWV1 … HK1 HK12) -IHWV1 -HK1 -HK12 #X2 #V2 #HWV2 #H #_
71     lapply (delift_mono … H … HW12) -H #H destruct
72     elim (lift_total W2 0 (i+1)) #U2 #HWU2
73     lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1
74     lapply (delift_lift_ge … HW12 … HLK1 HWU1 … HWU2) -HW12 -HLK1 -HWU1 //
75     >minus_plus <plus_minus_m_m // /3 width=6/
76   | elim (lt_or_ge i (d+e)) #Hide [ -HWV1 -HWU1 -HL12 | -Hdi -HL1 ]
77     [ lapply (sfr_inv_ldrop … HLK1 … HL1 ? ?) -HLK1 -HL1 // -Hdi -Hide #H destruct 
78     | lapply (transitive_le … (i+1) Hide ?) /2 width=1/ #Hdei
79       lapply (thin_ldrop_conf_ge … HL12 … HLK1 ?) -HL12 -HLK1 // #HL2K1
80       elim (lift_split … HWU1 d (i+1-e) ? ? ?) -HWU1 // /2 width=1/ #W
81       <plus_minus in ⊢ (??%??→?); /2 width=2/ #HW1
82       <minus_minus // /2 width=2/ -Hdei >commutative_plus <minus_n_n /3 width=6/
83     ]
84   ]
85 | #I #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
86   elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #_
87   elim (IHTU1 (L2.ⓑ{I}V2) (d+1) e ? ?) -IHTU1 /2 width=1/ -HL1 -HL12 #T2 #U2 #HTU2 #HT12 #HU12
88   lapply (delift_lsubs_trans … HT12 (L1.ⓑ{I}V2) ?) -HT12 /2 width=1/
89   lapply (delift_lsubs_trans … HU12 (L1.ⓑ{I}V2) ?) -HU12 /2 width=1/ /3 width=7/
90 | #L1 #V1 #W1 #T1 #U1 #_ #_ #IHVW1 #IHTU1 #L2 #d #e #HL1 #HL12
91   elim (IHVW1 … HL1 HL12) -IHVW1 #V2 #W2 #HVW2 #HV12 #HW12
92   elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 #X2 #Y2 #HXY2 #HX2 #HY2
93   elim (delift_inv_bind1 … HX2) -HX2 #Z21 #T2 #HZ21 #HT12 #H destruct
94   elim (delift_inv_bind1 … HY2) -HY2 #Z22 #U2 #HZ22 #HU12 #H destruct
95   lapply (delift_mono … HZ21 … HW12) -HZ21 #H destruct
96   lapply (delift_mono … HZ22 … HW12) -HZ22 #H destruct
97   @(ex3_2_intro … (ⓐV2.ⓛW2.T2) (ⓐV2.ⓛW2.U2)) /3 width=1/ (**) (* explict constructor, /4 depth=?/ is too slow *)
98 | #L1 #V1 #W1 #T1 #U1 #_ #_ #IHTU1 #IHUW1 #L2 #d #e #HL1 #HL12
99   elim (IHTU1 … HL1 HL12) -IHTU1 #T2 #U2 #HTU2 #HT12 #HU12
100   elim (IHUW1 … HL1 HL12) -IHUW1 -HL1 -HL12 #X2 #W2 #HXW2 #H #HW12
101   elim (delift_inv_flat1 … H) -H #V2 #Y2 #HV12 #HY2 #H destruct
102   lapply (delift_mono … HY2 … HU12) -HY2 #H destruct /3 width=7/
103 | #L1 #T1 #U1 #_ #IHTU1 #L2 #d #e #HL1 #HL12
104   elim (IHTU1 … HL1 HL12) -IHTU1 -HL1 -HL12 /3 width=5/
105 | #L1 #T1 #U11 #U12 #V1 #_ #HU112 #_ #IHT1 #IHU12 #L2 #d #e #HL1 #HL12
106   elim (IHT1 … HL1 HL12) -IHT1 #T2 #U21 #HT2 #HT12 #HU121
107   elim (IHU12 … HL1 HL12) -IHU12 -HL1 #U22 #V2 #HU22 #HU122 #_
108   lapply (thin_cpcs_delift_mono … HU112 … HL12 … HU121 … HU122) -HU112 -HL12 -HU121 /3 width=5/
109 ]
110 qed.
111
112 lemma nta_ldrop_conf: ∀h,L1,T1,U1. ⦃h, L1⦄ ⊢ T1 : U1 →
113                       ∀L2,d,e. ≽ [d, e] L1 → ⇩[d, e] L1 ≡ L2 →
114                       ∃∃T2,U2. ⦃h, L2⦄ ⊢ T2 : U2 &
115                                L1 ⊢ T1 ▼*[d, e] ≡ T2 & L1 ⊢ U1 ▼*[d, e] ≡ U2.
116 /3 width=1/ qed.