]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/substitution/drop_drop.ma
parked material ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / drop_drop.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_2/substitution/lift_lift.ma".
16 include "basic_2/substitution/drop.ma".
17
18 (* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
19
20 (* Main properties **********************************************************)
21
22 (* Basic_1: was: drop_mono *)
23 theorem drop_mono: ∀L,L1,s1,l,m. ⬇[s1, l, m] L ≡ L1 →
24                    ∀L2,s2. ⬇[s2, l, m] L ≡ L2 → L1 = L2.
25 #L #L1 #s1 #l #m #H elim H -L -L1 -l -m
26 [ #l #m #Hm #L2 #s2 #H elim (drop_inv_atom1 … H) -H //
27 | #I #K #V #L2 #s2 #HL12 <(drop_inv_O2 … HL12) -L2 //
28 | #I #L #K #V #m #_ #IHLK #L2 #s2 #H
29   lapply (drop_inv_drop1 … H) -H /2 width=2 by/
30 | #I #L #K1 #T #V1 #l #m #_ #HVT1 #IHLK1 #X #s2 #H
31   elim (drop_inv_skip1 … H) -H // >ypred_succ #K2 #V2 #HLK2 #HVT2 #H destruct
32   >(lift_inj … HVT1 … HVT2) -HVT1 -HVT2
33   >(IHLK1 … HLK2) -IHLK1 -HLK2 //
34 ]
35 qed-.
36
37 (* Basic_1: was: drop_conf_ge *)
38 theorem drop_conf_ge: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 →
39                       ∀L2,s2,m2. ⬇[s2, 0, m2] L ≡ L2 → l1 + m1 ≤ m2 →
40                       ⬇[s2, 0, m2 - m1] L1 ≡ L2.
41 #L #L1 #s1 #l1 #m1 #H elim H -L -L1 -l1 -m1 //
42 [ #l #m #_ #L2 #s2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
43   #H #Hm destruct
44   @drop_atom #H >Hm // (**) (* explicit constructor *)
45 | #I #L #K #V #m #_ #IHLK #L2 #s2 #m2 #H >yplus_O1 <yplus_inj #Hm2
46   lapply (drop_inv_drop1_lt … H ?) -H /3 width=2 by yle_inv_inj, ltn_to_ltO/ #HL2
47   lapply (yle_plus1_to_minus_inj2 … Hm2) -Hm2 #Hm2
48   <minus_plus >minus_minus_comm @IHLK //
49 | #I #L #K #V1 #V2 #l #m #_ #_ #IHLK #L2 #s2 #m2 #H #Hlmm2
50   lapply (yle_plus1_to_minus_inj2 … Hlmm2) #Hlm2m
51   lapply (ylt_yle_trans 0 … Hlm2m ?) // -Hlm2m #Hm2m
52   >yplus_succ1 in Hlmm2; #Hlmm2
53   elim (yle_inv_succ1 … Hlmm2) -Hlmm2 #Hlmm2 #Hm2
54   lapply (drop_inv_drop1_lt … H ?) -H /2 width=1 by ylt_inv_inj/ -Hm2 #HL2
55   @drop_drop_lt /2 width=1 by ylt_inv_inj/ >minus_minus_comm
56   <yminus_SO2 in Hlmm2; /2 width=1 by/
57 ]
58 qed.
59 (*
60 (* Note: apparently this was missing in basic_1 *)
61 theorem drop_conf_be: ∀L0,L1,s1,l1,m1. ⬇[s1, yinj l1, m1] L0 ≡ L1 →
62                       ∀L2,m2. ⬇[m2] L0 ≡ L2 → l1 ≤ m2 → m2 ≤ l1 + m1 →
63                       ∃∃L. ⬇[s1, yinj 0, l1 + m1 - m2] L2 ≡ L & ⬇[l1] L1 ≡ L.
64 #L0 #L1 #s1 #l1 #m1 #H elim H -L0 -L1 -l1 -m1
65 [ #l1 #m1 #Hm1 #L2 #m2 #H #Hl1 #_ elim (drop_inv_atom1 … H) -H #H #Hm2 destruct
66   >(Hm2 ?) in Hl1; // -Hm2 #Hl1 <(le_n_O_to_eq … Hl1) -l1
67   /4 width=3 by drop_atom, ex2_intro/
68 | normalize #I #L #V #L2 #m2 #HL2 #_ #Hm2
69   lapply (le_n_O_to_eq … Hm2) -Hm2 #H destruct
70   lapply (drop_inv_O2 … HL2) -HL2 #H destruct /2 width=3 by drop_pair, ex2_intro/
71 | normalize #I #L0 #K0 #V1 #m1 #HLK0 #IHLK0 #L2 #m2 #H #_ #Hm21
72   lapply (drop_inv_O1_pair1 … H) -H * * #Hm2 #HL20
73   [ -IHLK0 -Hm21 destruct <minus_n_O /3 width=3 by drop_drop, ex2_intro/
74   | -HLK0 <minus_le_minus_minus_comm //
75     elim (IHLK0 … HL20) -L0 /2 width=3 by monotonic_pred, ex2_intro/
76   ]
77 | #I #L0 #K0 #V0 #V1 #l1 #m1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #m2 #H #Hl1m2 #Hm2lm1
78   elim (le_inv_plus_l … Hl1m2) #_ #Hm2
79   <minus_le_minus_minus_comm //
80   lapply (drop_inv_drop1_lt … H ?) -H // #HL02
81   elim (IHLK0 … HL02) -L0 /3 width=3 by drop_drop, monotonic_pred, ex2_intro/
82 ]
83 qed-.
84 *)
85 (* Note: apparently this was missing in basic_1 *)
86 theorem drop_conf_le: ∀L0,L1,s1,l1,m1. ⬇[s1, l1, m1] L0 ≡ L1 →
87                       ∀L2,s2,m2. ⬇[s2, 0, m2] L0 ≡ L2 → m2 ≤ l1 →
88                       ∃∃L. ⬇[s2, 0, m2] L1 ≡ L & ⬇[s1, l1 - m2, m1] L2 ≡ L.
89 #L0 #L1 #s1 #l1 #m1 #H elim H -L0 -L1 -l1 -m1
90 [ #l1 #m1 #Hm1 #L2 #s2 #m2 #H elim (drop_inv_atom1 … H) -H
91   #H #Hm2 #_ destruct /4 width=3 by drop_atom, ex2_intro/
92 | #I #K0 #V0 #L2 #s2 #m2 #H1 #H2
93   lapply (yle_inv_O2  … H2) -H2 #H destruct
94   lapply (drop_inv_pair1 … H1) -H1 #H destruct /2 width=3 by drop_pair, ex2_intro/
95 | #I #K0 #K1 #V0 #m1 #HK01 #_ #L2 #s2 #m2 #H1 #H2
96   lapply (yle_inv_O2 … H2) -H2 #H destruct
97   lapply (drop_inv_pair1 … H1) -H1 #H destruct /3 width=3 by drop_drop, ex2_intro/
98 | #I #K0 #K1 #V0 #V1 #l1 #m1 #HK01 #HV10 #IHK01 #L2 #s2 #m2 #H #Hm2l1
99   elim (drop_inv_O1_pair1 … H) -H *
100   [ -IHK01 -Hm2l1 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
101   | -HK01 -HV10 #Hm2 #HK0L2
102     lapply (yle_inv_succ2 … Hm2l1) -Hm2l1 <yminus_SO2 #Hm2l1
103     elim (IHK01 … HK0L2) -IHK01 -HK0L2 //
104     <yminus_inj >yplus_minus_assoc_comm_inj /2 width=1 by yle_inj/
105     >yplus_SO2 /3 width=3 by drop_drop_lt, ex2_intro/  
106   ]
107 ]
108 qed-.
109
110 (* Note: with "s2", the conclusion parameter is "s1 ∨ s2" *)
111 (* Basic_1: was: drop_trans_ge *)
112 theorem drop_trans_ge: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
113                        ∀L2,m2. ⬇[m2] L ≡ L2 → l1 ≤ m2 → ⬇[s1, 0, m1 + m2] L1 ≡ L2.
114 #L1 #L #s1 #l1 #m1 #H elim H -L1 -L -l1 -m1
115 [ #l1 #m1 #Hm1 #L2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
116   #H #Hm2 destruct /4 width=1 by drop_atom, eq_f2/
117 | /2 width=1 by drop_gen/
118 | /3 width=1 by drop_drop/
119 | #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 #L #m2 #H #Hlm2
120   elim (yle_inv_succ1 … Hlm2) -Hlm2 #Hlm2 #Hm2
121   lapply (ylt_O … Hm2) -Hm2 #Hm2
122   lapply (lt_to_le_to_lt … (m + m2) Hm2 ?) // #Hmm2
123   lapply (drop_inv_drop1_lt … H ?) -H // #HL2
124   @drop_drop_lt // >le_plus_minus <yminus_SO2 in Hlm2; /2 width=1 by/
125 ]
126 qed.
127
128 (* Basic_1: was: drop_trans_le *)
129 theorem drop_trans_le: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
130                        ∀L2,s2,m2. ⬇[s2, 0, m2] L ≡ L2 → m2 ≤ l1 →
131                        ∃∃L0. ⬇[s2, 0, m2] L1 ≡ L0 & ⬇[s1, l1 - m2, m1] L0 ≡ L2.
132 #L1 #L #s1 #l1 #m1 #H elim H -L1 -L -l1 -m1
133 [ #l1 #m1 #Hm1 #L2 #s2 #m2 #H #_ elim (drop_inv_atom1 … H) -H
134   #H #Hm2 destruct /4 width=3 by drop_atom, ex2_intro/
135 | #I #K #V #L2 #s2 #m2 #HL2 #H lapply (yle_inv_O2 … H) -H
136   #H destruct /2 width=3 by drop_pair, ex2_intro/
137 | #I #L1 #L2 #V #m #_ #IHL12 #L #s2 #m2 #HL2 #H lapply (yle_inv_O2 … H) -H
138   #H destruct elim (IHL12 … HL2) -IHL12 -HL2 //
139   #L0 #H #HL0 lapply (drop_inv_O2 … H) -H #H destruct
140   /3 width=5 by drop_pair, drop_drop, ex2_intro/
141 | #I #L1 #L2 #V1 #V2 #l #m #HL12 #HV12 #IHL12 #L #s2 #m2 #H #Hm2l
142   elim (drop_inv_O1_pair1 … H) -H *
143   [ -Hm2l -IHL12 #H1 #H2 destruct /3 width=5 by drop_pair, drop_skip, ex2_intro/
144   | -HL12 -HV12 #Hm2 #HL2
145     lapply (yle_inv_succ2 … Hm2l) -Hm2l <yminus_SO2 #Hm2l
146     elim (IHL12 … HL2) -L2 //
147     <yminus_inj >yplus_minus_assoc_comm_inj /3 width=3 by drop_drop_lt, yle_inj, ex2_intro/
148   ]
149 ]
150 qed-.
151
152 (* Advanced properties ******************************************************)
153
154 lemma d_liftable_llstar: ∀R. d_liftable R → ∀d. d_liftable (llstar … R d).
155 #R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2
156 [ #L #s #l #m #_ #U1 #HTU1 #U2 #HTU2 -HR -K
157   >(lift_mono … HTU2 … HTU1) -T1 -U2 -l -m //
158 | #d #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #HTU1 #U2 #HTU2
159   elim (lift_total T l m) /3 width=12 by lstar_dx/
160 ]
161 qed-.
162
163 (* Basic_1: was: drop_conf_lt *)
164 lemma drop_conf_lt: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 →
165                     ∀I,K2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ K2.ⓑ{I}V2 →
166                     m2 < l1 → let l ≝ ⫰(l1 - m2) in
167                     ∃∃K1,V1. ⬇[s2, 0, m2] L1 ≡ K1.ⓑ{I}V1 &
168                              ⬇[s1, l, m1] K2 ≡ K1 & ⬆[l, m1] V1 ≡ V2.
169 #L #L1 #s1 #l1 #m1 #H1 #I #K2 #V2 #s2 #m2 #H2 #Hm2l1
170 elim (drop_conf_le … H1 … H2) -L /2 width=2 by ylt_fwd_le/ #K #HL1K #HK2
171 elim (drop_inv_skip1 … HK2) -HK2 /2 width=1 by ylt_to_minus/
172 #K1 #V1 #HK21 #HV12 #H destruct /2 width=5 by ex3_2_intro/
173 qed-.
174
175 (* Note: apparently this was missing in basic_1 *)
176 lemma drop_trans_lt: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L →
177                      ∀I,L2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ L2.ⓑ{I}V2 →
178                      m2 < l1 → let l ≝ l1 - m2 - 1 in
179                      ∃∃L0,V0. ⬇[s2, 0, m2] L1 ≡ L0.ⓑ{I}V0 &
180                               ⬇[s1, l, m1] L0 ≡ L2 & ⬆[l, m1] V2 ≡ V0.
181 #L1 #L #s1 #l1 #m1 #HL1 #I #L2 #V2 #s2 #m2 #HL2 #Hl21
182 elim (drop_trans_le … HL1 … HL2) -L /2 width=1 by ylt_fwd_le/ #L0 #HL10 #HL02
183 elim (drop_inv_skip2 … HL02) -HL02 /2 width=1 by ylt_to_minus/
184 #L #V1 #HL2 #HV21 #H destruct /2 width=5 by ex3_2_intro/
185 qed-.
186
187 lemma drop_trans_ge_comm: ∀L1,L,L2,s1,l1,m1,m2.
188                           ⬇[s1, l1, m1] L1 ≡ L → ⬇[m2] L ≡ L2 → l1 ≤ m2 →
189                           ⬇[s1, 0, m2 + m1] L1 ≡ L2.
190 #L1 #L #L2 #s1 #l1 #m1 #m2
191 >commutative_plus /2 width=5 by drop_trans_ge/
192 qed.
193
194 lemma drop_conf_div: ∀I1,L,K,V1,m1. ⬇[m1] L ≡ K.ⓑ{I1}V1 →
195                      ∀I2,V2,m2. ⬇[m2] L ≡ K.ⓑ{I2}V2 →
196                      ∧∧ m1 = m2 & I1 = I2 & V1 = V2.
197 #I1 #L #K #V1 #m1 #HLK1 #I2 #V2 #m2 #HLK2
198 elim (le_or_ge m1 m2) #Hm
199 [ lapply (drop_conf_ge … HLK1 … HLK2 ?)
200 | lapply (drop_conf_ge … HLK2 … HLK1 ?)
201 ] -HLK1 -HLK2 /2 width=1 by yle_inj/ #HK
202 lapply (drop_fwd_length_minus2 … HK) #H
203 elim (discr_minus_x_xy … H) -H
204 [1,3: normalize <plus_n_Sm #H destruct ]
205 #H >H in HK; #HK
206 lapply (drop_inv_O2 … HK) -HK #H destruct
207 lapply (inv_eq_minus_O … H) -H /3 width=1 by le_to_le_to_eq, and3_intro/
208 qed-.
209
210 (* Advanced forward lemmas **************************************************)
211
212 lemma drop_fwd_be: ∀L,K,s,l,m,i. ⬇[s, l, m] L ≡ K → |K| ≤ i → yinj i < l → |L| ≤ i.
213 #L #K #s #l #m #i #HLK #HK #Hl elim (lt_or_ge i (|L|)) //
214 #HL elim (drop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL
215 elim (drop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hl
216 #K1 #V1 #HK1 #_ #_ lapply (drop_fwd_length_lt2 … HK1) -I -K1 -V1
217 #H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
218 qed-.