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/leq.ma".
16 include "Basic-2/substitution/lift.ma".
18 (* DROPPING *****************************************************************)
20 (* Basic-1: includes: drop_skip_bind *)
21 inductive drop: nat → nat → relation lenv ≝
22 | drop_atom: ∀d,e. drop d e (⋆) (⋆)
23 | drop_pair: ∀L,I,V. drop 0 0 (L. 𝕓{I} V) (L. 𝕓{I} V)
24 | drop_drop: ∀L1,L2,I,V,e. drop 0 e L1 L2 → drop 0 (e + 1) (L1. 𝕓{I} V) L2
25 | drop_skip: ∀L1,L2,I,V1,V2,d,e.
26 drop d e L1 L2 → ↑[d,e] V2 ≡ V1 →
27 drop (d + 1) e (L1. 𝕓{I} V1) (L2. 𝕓{I} V2)
30 interpretation "dropping" 'RDrop d e L1 L2 = (drop d e L1 L2).
32 (* Basic inversion lemmas ***************************************************)
34 fact drop_inv_refl_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 → e = 0 → L1 = L2.
35 #d #e #L1 #L2 * -d e L1 L2
38 | #L1 #L2 #I #V #e #_ #_ #H
39 elim (plus_S_eq_O_false … H)
40 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H
41 elim (plus_S_eq_O_false … H)
45 (* Basic-1: was: drop_gen_refl *)
46 lemma drop_inv_refl: ∀L1,L2. ↓[0, 0] L1 ≡ L2 → L1 = L2.
49 fact drop_inv_atom1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → L1 = ⋆ →
51 #d #e #L1 #L2 * -d e L1 L2
53 | #L #I #V #H destruct
54 | #L1 #L2 #I #V #e #_ #H destruct
55 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct
59 (* Basic-1: was: drop_gen_sort *)
60 lemma drop_inv_atom1: ∀d,e,L2. ↓[d, e] ⋆ ≡ L2 → L2 = ⋆.
63 fact drop_inv_O1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → d = 0 →
64 ∀K,I,V. L1 = K. 𝕓{I} V →
65 (e = 0 ∧ L2 = K. 𝕓{I} V) ∨
66 (0 < e ∧ ↓[d, e - 1] K ≡ L2).
67 #d #e #L1 #L2 * -d e L1 L2
68 [ #d #e #_ #K #I #V #H destruct
69 | #L #I #V #_ #K #J #W #HX destruct -L I V /3/
70 | #L1 #L2 #I #V #e #HL12 #_ #K #J #W #H destruct -L1 I V /3/
71 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H elim (plus_S_eq_O_false … H)
75 lemma drop_inv_O1: ∀e,K,I,V,L2. ↓[0, e] K. 𝕓{I} V ≡ L2 →
76 (e = 0 ∧ L2 = K. 𝕓{I} V) ∨
77 (0 < e ∧ ↓[0, e - 1] K ≡ L2).
80 (* Basic-1: was: drop_gen_drop *)
81 lemma drop_inv_drop1: ∀e,K,I,V,L2.
82 ↓[0, e] K. 𝕓{I} V ≡ L2 → 0 < e → ↓[0, e - 1] K ≡ L2.
83 #e #K #I #V #L2 #H #He
84 elim (drop_inv_O1 … H) -H * // #H destruct -e;
85 elim (lt_refl_false … He)
88 fact drop_inv_skip1_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d →
89 ∀I,K1,V1. L1 = K1. 𝕓{I} V1 →
90 ∃∃K2,V2. ↓[d - 1, e] K1 ≡ K2 &
93 #d #e #L1 #L2 * -d e L1 L2
94 [ #d #e #_ #I #K #V #H destruct
95 | #L #I #V #H elim (lt_refl_false … H)
96 | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H)
97 | #X #L2 #Y #Z #V2 #d #e #HL12 #HV12 #_ #I #L1 #V1 #H destruct -X Y Z
102 (* Basic-1: was: drop_gen_skip_l *)
103 lemma drop_inv_skip1: ∀d,e,I,K1,V1,L2. ↓[d, e] K1. 𝕓{I} V1 ≡ L2 → 0 < d →
104 ∃∃K2,V2. ↓[d - 1, e] K1 ≡ K2 &
105 ↑[d - 1, e] V2 ≡ V1 &
109 fact drop_inv_skip2_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d →
110 ∀I,K2,V2. L2 = K2. 𝕓{I} V2 →
111 ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 &
112 ↑[d - 1, e] V2 ≡ V1 &
114 #d #e #L1 #L2 * -d e L1 L2
115 [ #d #e #_ #I #K #V #H destruct
116 | #L #I #V #H elim (lt_refl_false … H)
117 | #L1 #L2 #I #V #e #_ #H elim (lt_refl_false … H)
118 | #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #I #L2 #V2 #H destruct -X Y Z
123 (* Basic-1: was: drop_gen_skip_r *)
124 lemma drop_inv_skip2: ∀d,e,I,L1,K2,V2. ↓[d, e] L1 ≡ K2. 𝕓{I} V2 → 0 < d →
125 ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & ↑[d - 1, e] V2 ≡ V1 &
129 (* Basic properties *********************************************************)
131 (* Basic-1: was by definition: drop_refl *)
132 lemma drop_refl: ∀L. ↓[0, 0] L ≡ L.
136 lemma drop_drop_lt: ∀L1,L2,I,V,e.
137 ↓[0, e - 1] L1 ≡ L2 → 0 < e → ↓[0, e] L1. 𝕓{I} V ≡ L2.
138 #L1 #L2 #I #V #e #HL12 #He >(plus_minus_m_m e 1) /2/
141 lemma drop_leq_drop1: ∀L1,L2,d,e. L1 [d, e] ≈ L2 →
142 ∀I,K1,V,i. ↓[0, i] L1 ≡ K1. 𝕓{I} V →
144 ∃∃K2. K1 [0, d + e - i - 1] ≈ K2 &
145 ↓[0, i] L2 ≡ K2. 𝕓{I} V.
146 #L1 #L2 #d #e #H elim H -H L1 L2 d e
147 [ #d #e #I #K1 #V #i #H
148 lapply (drop_inv_atom1 … H) -H #H destruct
149 | #L1 #L2 #I #K1 #V #i #_ #_ #H
150 elim (lt_zero_false … H)
151 | #L1 #L2 #I #V #e #HL12 #IHL12 #J #K1 #W #i #H #_ #Hie
152 elim (drop_inv_O1 … H) -H * #Hi #HLK1
153 [ -IHL12 Hie; destruct -i K1 J W;
154 <minus_n_O <minus_plus_m_m /2/
156 elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 // [2: /2/ ] -Hie >arith_g1 // /3/
158 | #L1 #L2 #I1 #I2 #V1 #V2 #d #e #_ #IHL12 #I #K1 #V #i #H #Hdi >plus_plus_comm_23 #Hide
159 lapply (plus_S_le_to_pos … Hdi) #Hi
160 lapply (drop_inv_drop1 … H ?) -H // #HLK1
161 elim (IHL12 … HLK1 ? ?) -IHL12 HLK1 [2: /2/ |3: /2/ ] -Hdi Hide >arith_g1 // /3/
165 (* Basic forvard lemmas *****************************************************)
167 (* Basic-1: was: drop_S *)
168 lemma drop_fwd_drop2: ∀L1,I2,K2,V2,e. ↓[O, e] L1 ≡ K2. 𝕓{I2} V2 →
171 [ #I2 #K2 #V2 #e #H lapply (drop_inv_atom1 … H) -H #H destruct
172 | #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H
173 elim (drop_inv_O1 … H) -H * #He #H
174 [ -IHL1; destruct -e K2 I2 V2 /2/
175 | @drop_drop >(plus_minus_m_m e 1) /2/
180 lemma drop_fwd_drop2_length: ∀L1,I2,K2,V2,e.
181 ↓[0, e] L1 ≡ K2. 𝕓{I2} V2 → e < |L1|.
183 [ #I2 #K2 #V2 #e #H lapply (drop_inv_atom1 … H) -H #H destruct
184 | #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #e #H
185 elim (drop_inv_O1 … H) -H * #He #H
186 [ -IHL1; destruct -e K2 I2 V2 //
187 | lapply (IHL1 … H) -IHL1 H #HeK1 whd in ⊢ (? ? %) /2/
192 lemma drop_fwd_O1_length: ∀L1,L2,e. ↓[0, e] L1 ≡ L2 → |L2| = |L1| - e.
194 [ #L2 #e #H >(drop_inv_atom1 … H) -H //
195 | #K1 #I1 #V1 #IHL1 #L2 #e #H
196 elim (drop_inv_O1 … H) -H * #He #H
197 [ -IHL1; destruct -e L2 //
198 | lapply (IHL1 … H) -IHL1 H #H >H -H; normalize
199 >minus_le_minus_minus_comm //
204 (* Basic-1: removed theorems 49:
206 cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf
207 drop_clear drop_clear_O drop_clear_S
208 clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r
209 clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle
210 getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans
211 getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt
212 getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev
213 drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge
214 getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O
215 getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le
216 getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono