1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2/grammar/cl_weight.ma".
16 include "basic_2/substitution/lift.ma".
17 include "basic_2/substitution/lsubs.ma".
19 (* LOCAL ENVIRONMENT SLICING ************************************************)
21 (* Basic_1: includes: drop_skip_bind *)
22 inductive ldrop: nat → nat → relation lenv ≝
23 | ldrop_atom : ∀d,e. ldrop d e (⋆) (⋆)
24 | ldrop_pair : ∀L,I,V. ldrop 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V)
25 | ldrop_ldrop: ∀L1,L2,I,V,e. ldrop 0 e L1 L2 → ldrop 0 (e + 1) (L1. ⓑ{I} V) L2
26 | ldrop_skip : ∀L1,L2,I,V1,V2,d,e.
27 ldrop d e L1 L2 → ⇧[d,e] V2 ≡ V1 →
28 ldrop (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2)
31 interpretation "local slicing" 'RDrop d e L1 L2 = (ldrop d e L1 L2).
33 (* Basic inversion lemmas ***************************************************)
35 fact ldrop_inv_refl_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2.
36 #d #e #L1 #L2 * -d -e -L1 -L2
39 | #L1 #L2 #I #V #e #_ #_ >commutative_plus normalize #H destruct
40 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct
44 (* Basic_1: was: drop_gen_refl *)
45 lemma ldrop_inv_refl: ∀L1,L2. ⇩[0, 0] L1 ≡ L2 → L1 = L2.
48 fact ldrop_inv_atom1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → L1 = ⋆ →
50 #d #e #L1 #L2 * -d -e -L1 -L2
52 | #L #I #V #H destruct
53 | #L1 #L2 #I #V #e #_ #H destruct
54 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct
58 (* Basic_1: was: drop_gen_sort *)
59 lemma ldrop_inv_atom1: ∀d,e,L2. ⇩[d, e] ⋆ ≡ L2 → L2 = ⋆.
62 fact ldrop_inv_O1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → d = 0 →
63 ∀K,I,V. L1 = K. ⓑ{I} V →
64 (e = 0 ∧ L2 = K. ⓑ{I} V) ∨
65 (0 < e ∧ ⇩[d, e - 1] K ≡ L2).
66 #d #e #L1 #L2 * -d -e -L1 -L2
67 [ #d #e #_ #K #I #V #H destruct
68 | #L #I #V #_ #K #J #W #HX destruct /3 width=1/
69 | #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct /3 width=1/
70 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct
74 lemma ldrop_inv_O1: ∀e,K,I,V,L2. ⇩[0, e] K. ⓑ{I} V ≡ L2 →
75 (e = 0 ∧ L2 = K. ⓑ{I} V) ∨
76 (0 < e ∧ ⇩[0, e - 1] K ≡ L2).
79 lemma ldrop_inv_pair1: ∀K,I,V,L2. ⇩[0, 0] K. ⓑ{I} V ≡ L2 → L2 = K. ⓑ{I} V.
81 elim (ldrop_inv_O1 … H) -H * // #H destruct
82 elim (lt_refl_false … H)
85 (* Basic_1: was: drop_gen_drop *)
86 lemma ldrop_inv_ldrop1: ∀e,K,I,V,L2.
87 ⇩[0, e] K. ⓑ{I} V ≡ L2 → 0 < e → ⇩[0, e - 1] K ≡ L2.
88 #e #K #I #V #L2 #H #He
89 elim (ldrop_inv_O1 … H) -H * // #H destruct
90 elim (lt_refl_false … He)
93 fact ldrop_inv_skip1_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d →
94 ∀I,K1,V1. L1 = K1. ⓑ{I} V1 →
95 ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 &
98 #d #e #L1 #L2 * -d -e -L1 -L2
99 [ #d #e #_ #I #K #V #H destruct
100 | #L #I #V #H elim (lt_refl_false … H)
101 | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H)
102 | #X #L2 #Y #Z #V2 #d #e #HL12 #HV12 #_ #I #L1 #V1 #H destruct /2 width=5/
106 (* Basic_1: was: drop_gen_skip_l *)
107 lemma ldrop_inv_skip1: ∀d,e,I,K1,V1,L2. ⇩[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d →
108 ∃∃K2,V2. ⇩[d - 1, e] K1 ≡ K2 &
109 ⇧[d - 1, e] V2 ≡ V1 &
113 fact ldrop_inv_skip2_aux: ∀d,e,L1,L2. ⇩[d, e] L1 ≡ L2 → 0 < d →
114 ∀I,K2,V2. L2 = K2. ⓑ{I} V2 →
115 ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 &
116 ⇧[d - 1, e] V2 ≡ V1 &
118 #d #e #L1 #L2 * -d -e -L1 -L2
119 [ #d #e #_ #I #K #V #H destruct
120 | #L #I #V #H elim (lt_refl_false … H)
121 | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H)
122 | #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct /2 width=5/
126 (* Basic_1: was: drop_gen_skip_r *)
127 lemma ldrop_inv_skip2: ∀d,e,I,L1,K2,V2. ⇩[d, e] L1 ≡ K2. ⓑ{I} V2 → 0 < d →
128 ∃∃K1,V1. ⇩[d - 1, e] K1 ≡ K2 & ⇧[d - 1, e] V2 ≡ V1 &
132 (* Basic properties *********************************************************)
134 (* Basic_1: was by definition: drop_refl *)
135 lemma ldrop_refl: ∀L. ⇩[0, 0] L ≡ L.
139 lemma ldrop_ldrop_lt: ∀L1,L2,I,V,e.
140 ⇩[0, e - 1] L1 ≡ L2 → 0 < e → ⇩[0, e] L1. ⓑ{I} V ≡ L2.
141 #L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) // /2 width=1/
144 lemma ldrop_skip_lt: ∀L1,L2,I,V1,V2,d,e.
145 ⇩[d - 1, e] L1 ≡ L2 → ⇧[d - 1, e] V2 ≡ V1 → 0 < d →
146 ⇩[d, e] L1. ⓑ{I} V1 ≡ L2. ⓑ{I} V2.
147 #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV21 #Hd >(plus_minus_m_m d 1) // /2 width=1/
150 lemma ldrop_O1: ∀L,i. i < |L| → ∃∃I,K,V. ⇩[0, i] L ≡ K.ⓑ{I}V.
152 [ #i #H elim (lt_zero_false … H)
153 | #L #I #V #IHL #i @(nat_ind_plus … i) -i /2 width=4/ #i #_ #H
154 lapply (lt_plus_to_lt_l … H) -H #Hi
155 elim (IHL i ?) // /3 width=4/
159 lemma ldrop_lsubs_ldrop2_abbr: ∀L1,L2,d,e. L1 ≼ [d, e] L2 →
160 ∀K2,V,i. ⇩[0, i] L2 ≡ K2. ⓓV →
162 ∃∃K1. K1 ≼ [0, d + e - i - 1] K2 &
164 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e
166 lapply (ldrop_inv_atom1 … H) -H #H destruct
167 | #L1 #L2 #K1 #V #i #_ #_ #H
168 elim (lt_zero_false … H)
169 | #L1 #L2 #V #e #HL12 #IHL12 #K1 #W #i #H #_ #Hie
170 elim (ldrop_inv_O1 … H) -H * #Hi #HLK1
171 [ -IHL12 -Hie destruct
172 <minus_n_O <minus_plus_m_m // /2 width=3/
174 elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hie >minus_minus_comm >arith_b1 // /4 width=3/
176 | #L1 #L2 #I #V1 #V2 #e #_ #IHL12 #K1 #W #i #H #_ #Hie
177 elim (ldrop_inv_O1 … H) -H * #Hi #HLK1
178 [ -IHL12 -Hie -Hi destruct
179 | elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hie >minus_minus_comm >arith_b1 // /3 width=3/
181 | #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide
182 elim (le_inv_plus_l … Hdi) #Hdim #Hi
183 lapply (ldrop_inv_ldrop1 … H ?) -H // #HLK1
184 elim (IHL12 … HLK1 ? ?) -IHL12 -HLK1 // /2 width=1/ -Hdi -Hide >minus_minus_comm >arith_b1 // /3 width=3/
188 (* Basic forvard lemmas *****************************************************)
190 (* Basic_1: was: drop_S *)
191 lemma ldrop_fwd_ldrop2: ∀L1,I2,K2,V2,e. ⇩[O, e] L1 ≡ K2. ⓑ{I2} V2 →
194 [ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct
195 | #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H
196 elim (ldrop_inv_O1 … H) -H * #He #H
197 [ -IHL1 destruct /2 width=1/
198 | @ldrop_ldrop >(plus_minus_m_m e 1) // /2 width=3/
203 lemma ldrop_fwd_lw: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → #{L2} ≤ #{L1}.
204 #L1 #L2 #d #e #H elim H -L1 -L2 -d -e // normalize
206 | #L1 #L2 #I #V1 #V2 #d #e #_ #HV21 #IHL12
207 >(tw_lift … HV21) -HV21 /2 width=1/
211 lemma ldrop_pair2_fwd_cw: ∀I,L,K,V,d,e. ⇩[d, e] L ≡ K. ⓑ{I} V →
212 ∀T. #{K, V} < #{L, T}.
213 #I #L #K #V #d #e #H #T
214 lapply (ldrop_fwd_lw … H) -H #H
215 @(le_to_lt_to_lt … H) -H /3 width=1/
218 lemma ldrop_fwd_ldrop2_length: ∀L1,I2,K2,V2,e.
219 ⇩[0, e] L1 ≡ K2. ⓑ{I2} V2 → e < |L1|.
221 [ #I2 #K2 #V2 #e #H lapply (ldrop_inv_atom1 … H) -H #H destruct
222 | #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H
223 elim (ldrop_inv_O1 … H) -H * #He #H
225 | lapply (IHL1 … H) -IHL1 -H #HeK1 whd in ⊢ (? ? %); /2 width=1/
230 lemma ldrop_fwd_O1_length: ∀L1,L2,e. ⇩[0, e] L1 ≡ L2 → |L2| = |L1| - e.
232 [ #L2 #e #H >(ldrop_inv_atom1 … H) -H //
233 | #K1 #I1 #V1 #IHL1 #L2 #e #H
234 elim (ldrop_inv_O1 … H) -H * #He #H
236 | lapply (IHL1 … H) -IHL1 -H #H >H -H normalize
237 >minus_le_minus_minus_comm //
242 (* Basic_1: removed theorems 50:
243 drop_ctail drop_skip_flat
244 cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf
245 drop_clear drop_clear_O drop_clear_S
246 clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r
247 clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle
248 getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans
249 getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt
250 getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev
251 drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge
252 getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O
253 getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le
254 getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono