]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/substitution/drop_defs.ma
first main property of drop closed
[helm.git] / matita / matita / lib / lambda-delta / substitution / drop_defs.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic
3     ||A||  Library of Mathematics, developed at the Computer Science
4     ||T||  Department of the University of Bologna, Italy.
5     ||I||
6     ||T||
7     ||A||  This file is distributed under the terms of the
8     \   /  GNU General Public License Version 2
9      \ /
10       V_______________________________________________________________ *)
11
12 include "lambda-delta/syntax/lenv.ma".
13 include "lambda-delta/substitution/lift_defs.ma".
14
15 (* DROPPING *****************************************************************)
16
17 inductive drop: lenv → nat → nat → lenv → Prop ≝
18 | drop_refl: ∀L. drop L 0 0 L
19 | drop_drop: ∀L1,L2,I,V,e. drop L1 0 e L2 → drop (L1. 𝕓{I} V) 0 (e + 1) L2
20 | drop_skip: ∀L1,L2,I,V1,V2,d,e.
21              drop L1 d e L2 → ↑[d,e] V2 ≡ V1 →
22              drop (L1. 𝕓{I} V1) (d + 1) e (L2. 𝕓{I} V2)
23 .
24
25 interpretation "dropping" 'RLift L2 d e L1 = (drop L1 d e L2).
26
27 (* the basic inversion lemmas ***********************************************)
28
29 lemma drop_inv_drop1_aux: ∀d,e,L2,L1. ↑[d, e] L2 ≡ L1 → 0 < e → d = 0 →
30                           ∀K,I,V. L1 = K. 𝕓{I} V → ↑[d, e - 1] L2 ≡ K.
31 #d #e #L2 #L1 #H elim H -H d e L2 L1
32 [ #L #H elim (lt_refl_false … H)
33 | #L1 #L2 #I #V #e #HL12 #_ #_ #_ #K #J #W #H destruct -L1 I V //
34 | #L1 #L2 #I #V1 #V2 #d #e #_ #_ #_ #_ #H elim (plus_S_eq_O_false … H)
35 ]
36 qed.
37
38 lemma drop_inv_drop1: ∀e,L2,K,I,V. ↑[0, e] L2 ≡ K. 𝕓{I} V → 0 < e →
39                                    ↑[0, e - 1] L2 ≡ K.
40 /2 width=5/ qed.
41
42 lemma drop_inv_skip2_aux: ∀d,e,L1,L2. ↑[d, e] L2 ≡ L1 → 0 < d →
43                           ∀I,K2,V2. L2 = K2. 𝕓{I} V2 →
44                           ∃∃K1,V1. ↑[d - 1, e] K2 ≡ K1 &
45                                    ↑[d - 1, e] V2 ≡ V1 & 
46                                    L1 = K1. 𝕓{I} V1.
47 #d #e #L1 #L2 #H elim H -H d e L1 L2
48 [ #L #H elim (lt_refl_false … H)
49 | #L1 #L2 #I #V #e #_ #_ #H elim (lt_refl_false … H)
50 | #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #_ #I #L2 #V2 #H destruct -X Y Z;
51   /2 width=5/
52 ]
53 qed.
54
55 lemma drop_inv_skip2: ∀d,e,I,L1,K2,V2. ↑[d, e] K2. 𝕓{I} V2 ≡ L1 → 0 < d →
56                       ∃∃K1,V1. ↑[d - 1, e] K2 ≡ K1 & ↑[d - 1, e] V2 ≡ V1 &
57                                L1 = K1. 𝕓{I} V1.
58 /2/ qed.