]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/substitution/ldrop_ldrop.ma
- some refactoring and minor additions
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / ldrop_ldrop.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/ldrop.ma".
17
18 (* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************)
19
20 (* Main properties **********************************************************)
21
22 (* Basic_1: was: drop_mono *)
23 theorem ldrop_mono: ∀L,L1,s1,d,e. ⇩[s1, d, e] L ≡ L1 →
24                     ∀L2,s2. ⇩[s2, d, e] L ≡ L2 → L1 = L2.
25 #L #L1 #s1 #d #e #H elim H -L -L1 -d -e
26 [ #d #e #He #L2 #s2 #H elim (ldrop_inv_atom1 … H) -H //
27 | #I #K #V #L2 #s2 #HL12 <(ldrop_inv_O2 … HL12) -L2 //
28 | #I #L #K #V #e #_ #IHLK #L2 #s2 #H
29   lapply (ldrop_inv_drop1 … H) -H /2 width=2 by/
30 | #I #L #K1 #T #V1 #d #e #_ #HVT1 #IHLK1 #X #s2 #H
31   elim (ldrop_inv_skip1 … H) -H // <minus_plus_m_m #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 ldrop_conf_ge: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
39                        ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
40                        ⇩[s2, 0, e2 - e1] L1 ≡ L2.
41 #L #L1 #s1 #d1 #e1 #H elim H -L -L1 -d1 -e1 //
42 [ #d #e #_ #L2 #s2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
43   #H #He destruct
44   @ldrop_atom #H >He // (**) (* explicit constructor *)
45 | #I #L #K #V #e #_ #IHLK #L2 #s2 #e2 #H #He2
46   lapply (ldrop_inv_drop1_lt … H ?) -H /2 width=2 by ltn_to_ltO/ #HL2
47   <minus_plus >minus_minus_comm /3 width=1 by monotonic_pred/
48 | #I #L #K #V1 #V2 #d #e #_ #_ #IHLK #L2 #s2 #e2 #H #Hdee2
49   lapply (transitive_le 1 … Hdee2) // #He2
50   lapply (ldrop_inv_drop1_lt … H ?) -H // -He2 #HL2
51   lapply (transitive_le (1+e) … Hdee2) // #Hee2
52   @ldrop_drop_lt >minus_minus_comm /3 width=1 by lt_minus_to_plus_r, monotonic_le_minus_r, monotonic_pred/ (**) (* explicit constructor *)
53 ]
54 qed.
55
56 (* Note: apparently this was missing in basic_1 *)
57 theorem ldrop_conf_be: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
58                        ∀L2,e2. ⇩[e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 →
59                        ∃∃L. ⇩[s1, 0, d1 + e1 - e2] L2 ≡ L & ⇩[d1] L1 ≡ L.
60 #L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
61 [ #d1 #e1 #He1 #L2 #e2 #H #Hd1 #_ elim (ldrop_inv_atom1 … H) -H #H #He2 destruct
62   >(He2 ?) in Hd1; // -He2 #Hd1 <(le_n_O_to_eq … Hd1) -d1
63   /4 width=3 by ldrop_atom, ex2_intro/
64 | normalize #I #L #V #L2 #e2 #HL2 #_ #He2
65   lapply (le_n_O_to_eq … He2) -He2 #H destruct
66   lapply (ldrop_inv_O2 … HL2) -HL2 #H destruct /2 width=3 by ldrop_pair, ex2_intro/
67 | normalize #I #L0 #K0 #V1 #e1 #HLK0 #IHLK0 #L2 #e2 #H #_ #He21
68   lapply (ldrop_inv_O1_pair1 … H) -H * * #He2 #HL20
69   [ -IHLK0 -He21 destruct <minus_n_O /3 width=3 by ldrop_drop, ex2_intro/
70   | -HLK0 <minus_le_minus_minus_comm //
71     elim (IHLK0 … HL20) -L0 /2 width=3 by monotonic_pred, ex2_intro/
72   ]
73 | #I #L0 #K0 #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #e2 #H #Hd1e2 #He2de1
74   elim (le_inv_plus_l … Hd1e2) #_ #He2
75   <minus_le_minus_minus_comm //
76   lapply (ldrop_inv_drop1_lt … H ?) -H // #HL02
77   elim (IHLK0 … HL02) -L0 /3 width=3 by ldrop_drop, monotonic_pred, ex2_intro/
78 ]
79 qed-.
80
81 (* Note: apparently this was missing in basic_1 *)
82 theorem ldrop_conf_le: ∀L0,L1,s1,d1,e1. ⇩[s1, d1, e1] L0 ≡ L1 →
83                        ∀L2,s2,e2. ⇩[s2, 0, e2] L0 ≡ L2 → e2 ≤ d1 →
84                        ∃∃L. ⇩[s2, 0, e2] L1 ≡ L & ⇩[s1, d1 - e2, e1] L2 ≡ L.
85 #L0 #L1 #s1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
86 [ #d1 #e1 #He1 #L2 #s2 #e2 #H elim (ldrop_inv_atom1 … H) -H
87   #H #He2 #_ destruct /4 width=3 by ldrop_atom, ex2_intro/
88 | #I #K0 #V0 #L2 #s2 #e2 #H1 #H2
89   lapply (le_n_O_to_eq … H2) -H2 #H destruct
90   lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /2 width=3 by ldrop_pair, ex2_intro/
91 | #I #K0 #K1 #V0 #e1 #HK01 #_ #L2 #s2 #e2 #H1 #H2
92   lapply (le_n_O_to_eq … H2) -H2 #H destruct
93   lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /3 width=3 by ldrop_drop, ex2_intro/
94 | #I #K0 #K1 #V0 #V1 #d1 #e1 #HK01 #HV10 #IHK01 #L2 #s2 #e2 #H #He2d1
95   elim (ldrop_inv_O1_pair1 … H) -H *
96   [ -IHK01 -He2d1 #H1 #H2 destruct /3 width=5 by ldrop_pair, ldrop_skip, ex2_intro/
97   | -HK01 -HV10 #He2 #HK0L2
98     elim (IHK01 … HK0L2) -IHK01 -HK0L2 /2 width=1 by monotonic_pred/
99     >minus_le_minus_minus_comm /3 width=3 by ldrop_drop_lt, ex2_intro/
100   ]
101 ]
102 qed-.
103
104 (* Note: with "s2", the conclusion parameter is "s1 ∨ s2" *)
105 (* Basic_1: was: drop_trans_ge *)
106 theorem ldrop_trans_ge: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
107                         ∀L2,e2. ⇩[e2] L ≡ L2 → d1 ≤ e2 → ⇩[s1, 0, e1 + e2] L1 ≡ L2.
108 #L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
109 [ #d1 #e1 #He1 #L2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
110   #H #He2 destruct /4 width=1 by ldrop_atom, eq_f2/
111 | /2 width=1 by ldrop_gen/
112 | /3 width=1 by ldrop_drop/
113 | #I #L1 #L2 #V1 #V2 #d #e #_ #_ #IHL12 #L #e2 #H #Hde2
114   lapply (lt_to_le_to_lt 0 … Hde2) // #He2
115   lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2
116   lapply (ldrop_inv_drop1_lt … H ?) -H // #HL2
117   @ldrop_drop_lt // >le_plus_minus /3 width=1 by monotonic_pred/
118 ]
119 qed.
120
121 (* Basic_1: was: drop_trans_le *)
122 theorem ldrop_trans_le: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
123                         ∀L2,s2,e2. ⇩[s2, 0, e2] L ≡ L2 → e2 ≤ d1 →
124                         ∃∃L0. ⇩[s2, 0, e2] L1 ≡ L0 & ⇩[s1, d1 - e2, e1] L0 ≡ L2.
125 #L1 #L #s1 #d1 #e1 #H elim H -L1 -L -d1 -e1
126 [ #d1 #e1 #He1 #L2 #s2 #e2 #H #_ elim (ldrop_inv_atom1 … H) -H
127   #H #He2 destruct /4 width=3 by ldrop_atom, ex2_intro/
128 | #I #K #V #L2 #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
129   #H destruct /2 width=3 by ldrop_pair, ex2_intro/
130 | #I #L1 #L2 #V #e #_ #IHL12 #L #s2 #e2 #HL2 #H lapply (le_n_O_to_eq … H) -H
131   #H destruct elim (IHL12 … HL2) -IHL12 -HL2 //
132   #L0 #H #HL0 lapply (ldrop_inv_O2 … H) -H #H destruct
133   /3 width=5 by ldrop_pair, ldrop_drop, ex2_intro/
134 | #I #L1 #L2 #V1 #V2 #d #e #HL12 #HV12 #IHL12 #L #s2 #e2 #H #He2d
135   elim (ldrop_inv_O1_pair1 … H) -H *
136   [ -He2d -IHL12 #H1 #H2 destruct /3 width=5 by ldrop_pair, ldrop_skip, ex2_intro/
137   | -HL12 -HV12 #He2 #HL2
138     elim (IHL12 … HL2) -L2 [ >minus_le_minus_minus_comm // /3 width=3 by ldrop_drop_lt, ex2_intro/ | /2 width=1 by monotonic_pred/ ]
139   ]
140 ]
141 qed-.
142
143 (* Advanced properties ******************************************************)
144
145 lemma l_liftable_llstar: ∀R. l_liftable R → ∀l. l_liftable (llstar … R l).
146 #R #HR #l #K #T1 #T2 #H @(lstar_ind_r … l T2 H) -l -T2
147 [ #L #s #d #e #_ #U1 #HTU1 #U2 #HTU2 -HR -K
148   >(lift_mono … HTU2 … HTU1) -T1 -U2 -d -e //
149 | #l #T #T2 #_ #HT2 #IHT1 #L #s #d #e #HLK #U1 #HTU1 #U2 #HTU2
150   elim (lift_total T d e) /3 width=12 by lstar_dx/
151 ]
152 qed-.
153
154 (* Basic_1: was: drop_conf_lt *)
155 lemma ldrop_conf_lt: ∀L,L1,s1,d1,e1. ⇩[s1, d1, e1] L ≡ L1 →
156                      ∀I,K2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ K2.ⓑ{I}V2 →
157                      e2 < d1 → let d ≝ d1 - e2 - 1 in
158                      ∃∃K1,V1. ⇩[s2, 0, e2] L1 ≡ K1.ⓑ{I}V1 &
159                               ⇩[s1, d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2.
160 #L #L1 #s1 #d1 #e1 #H1 #I #K2 #V2 #s2 #e2 #H2 #He2d1
161 elim (ldrop_conf_le … H1 … H2) -L /2 width=2 by lt_to_le/ #K #HL1K #HK2
162 elim (ldrop_inv_skip1 … HK2) -HK2 /2 width=1 by lt_plus_to_minus_r/
163 #K1 #V1 #HK21 #HV12 #H destruct /2 width=5 by ex3_2_intro/
164 qed-.
165
166 (* Note: apparently this was missing in basic_1 *)
167 lemma ldrop_trans_lt: ∀L1,L,s1,d1,e1. ⇩[s1, d1, e1] L1 ≡ L →
168                       ∀I,L2,V2,s2,e2. ⇩[s2, 0, e2] L ≡ L2.ⓑ{I}V2 →
169                       e2 < d1 → let d ≝ d1 - e2 - 1 in
170                       ∃∃L0,V0. ⇩[s2, 0, e2] L1 ≡ L0.ⓑ{I}V0 &
171                                ⇩[s1, d, e1] L0 ≡ L2 & ⇧[d, e1] V2 ≡ V0.
172 #L1 #L #s1 #d1 #e1 #HL1 #I #L2 #V2 #s2 #e2 #HL2 #Hd21
173 elim (ldrop_trans_le … HL1 … HL2) -L /2 width=1 by lt_to_le/ #L0 #HL10 #HL02
174 elim (ldrop_inv_skip2 … HL02) -HL02 /2 width=1 by lt_plus_to_minus_r/ #L #V1 #HL2 #HV21 #H destruct /2 width=5 by ex3_2_intro/
175 qed-.
176
177 lemma ldrop_trans_ge_comm: ∀L1,L,L2,s1,d1,e1,e2.
178                            ⇩[s1, d1, e1] L1 ≡ L → ⇩[e2] L ≡ L2 → d1 ≤ e2 →
179                            ⇩[s1, 0, e2 + e1] L1 ≡ L2.
180 #L1 #L #L2 #s1 #d1 #e1 #e2
181 >commutative_plus /2 width=5 by ldrop_trans_ge/
182 qed.
183
184 lemma ldrop_conf_div: ∀I1,L,K,V1,e1. ⇩[e1] L ≡ K.ⓑ{I1}V1 →
185                       ∀I2,V2,e2. ⇩[e2] L ≡ K.ⓑ{I2}V2 →
186                       ∧∧ e1 = e2 & I1 = I2 & V1 = V2.
187 #I1 #L #K #V1 #e1 #HLK1 #I2 #V2 #e2 #HLK2
188 elim (le_or_ge e1 e2) #He
189 [ lapply (ldrop_conf_ge … HLK1 … HLK2 ?)
190 | lapply (ldrop_conf_ge … HLK2 … HLK1 ?)
191 ] -HLK1 -HLK2 // #HK
192 lapply (ldrop_fwd_length_minus2 … HK) #H
193 elim (discr_minus_x_xy … H) -H
194 [1,3: normalize <plus_n_Sm #H destruct ]
195 #H >H in HK; #HK
196 lapply (ldrop_inv_O2 … HK) -HK #H destruct
197 lapply (inv_eq_minus_O … H) -H /3 width=1 by le_to_le_to_eq, and3_intro/
198 qed-.
199
200 (* Advanced forward lemmas **************************************************)
201
202 lemma ldrop_fwd_be: ∀L,K,s,d,e,i. ⇩[s, d, e] L ≡ K → |K| ≤ i → i < d → |L| ≤ i.
203 #L #K #s #d #e #i #HLK #HK #Hd elim (lt_or_ge i (|L|)) //
204 #HL elim (ldrop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL
205 elim (ldrop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hd
206 #K1 #V1 #HK1 #_ #_ lapply (ldrop_fwd_length_lt2 … HK1) -I -K1 -V1
207 #H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/
208 qed-.