]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/substitution/thin_defs.ma
first main property of drop closed
[helm.git] / matita / matita / lib / lambda-delta / substitution / thin_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/substitution/subst_defs.ma".
13
14 (* THINNING *****************************************************************)
15
16 inductive thin: lenv → nat → nat → lenv → Prop ≝
17 | thin_refl: ∀L. thin L 0 0 L
18 | thin_thin: ∀L1,L2,I,V,e. thin L1 0 e L2 → thin (L1. 𝕓{I} V) 0 (e + 1) L2
19 | thin_skip: ∀L1,L2,I,V1,V2,d,e.
20              thin L1 d e L2 → L1 ⊢ ↓[d,e] V1 ≡ V2 →
21              thin (L1. 𝕓{I} V1) (d + 1) e (L2. 𝕓{I} V2)
22 .
23
24 interpretation "thinning" 'RSubst L1 d e L2 = (thin L1 d e L2).
25
26 (* the basic inversion lemmas ***********************************************)
27
28 lemma thin_inv_skip2_aux: ∀d,e,L1,L2. ↓[d, e] L1 ≡ L2 → 0 < d →
29                           ∀I,K2,V2. L2 = K2. 𝕓{I} V2 →
30                           ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 &
31                                    K1 ⊢ ↓[d - 1, e] V1 ≡ V2 & 
32                                    L1 = K1. 𝕓{I} V1.
33 #d #e #L1 #L2 #H elim H -H d e L1 L2
34 [ #L #H elim (lt_refl_false … H)
35 | #L1 #L2 #I #V #e #_ #_ #H elim (lt_refl_false … H)
36 | #L1 #X #Y #V1 #Z #d #e #HL12 #HV12 #_ #_ #I #L2 #V2 #H destruct -X Y Z;
37   /2 width=5/
38 ]
39 qed.
40
41 lemma thin_inv_skip2: ∀d,e,I,L1,K2,V2. ↓[d, e] L1 ≡ K2. 𝕓{I} V2 → 0 < d →
42                       ∃∃K1,V1. ↓[d - 1, e] K1 ≡ K2 & K1 ⊢ ↓[d - 1, e] V1 ≡ V2 &
43                                L1 = K1. 𝕓{I} V1.
44 /2/ qed.