]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/substitution/drop_main.ma
72d980e9bf183784f494a1d69a8ef30253bf7a66
[helm.git] / matita / matita / lib / lambda-delta / substitution / drop_main.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 "lambda-delta/substitution/drop_defs.ma".
16
17 (* DROPPING *****************************************************************)
18
19 (* the main properties ******************************************************)
20
21 lemma drop_conf_ge: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L →
22                     ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 + e1 ≤ e2 →
23                     ↑[0, e2 - e1] L2 ≡ L1.
24 #d1 #e1 #L #L1 #H elim H -H d1 e1 L L1
25 [ //
26 | #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2
27   lapply (drop_inv_drop1 … H ?) -H /2/ #HL2
28   <minus_plus_comm /3/
29 | #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2
30   lapply (transitive_le 1 … Hdee2) // #He2
31   lapply (drop_inv_drop1 … H ?) -H // -He2 #HL2
32   lapply (transitive_le (1+e) … Hdee2) // #Hee2
33   @drop_drop_lt >minus_minus_comm /3/ (**) (* explicit constructor *)
34 ]
35 qed.
36
37 lemma drop_conf_lt: ∀d1,e1,L,L1. ↑[d1, e1] L1 ≡ L →
38                     ∀e2,K2,I,V2. ↑[0, e2] K2. 𝕓{I} V2 ≡ L →
39                     e2 < d1 → let d ≝ d1 - e2 - 1 in
40                     ∃∃K1,V1. ↑[0, e2] K1. 𝕓{I} V1 ≡ L1 & 
41                              ↑[d, e1] K1 ≡ K2 & ↑[d,e1] V1 ≡ V2.
42 #d1 #e1 #L #L1 #H elim H -H d1 e1 L L1
43 [ #L0 #e2 #K2 #I #V2 #_ #H
44   elim (lt_zero_false … H)
45 | #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H
46   elim (lt_zero_false … H)
47 | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d
48   elim (drop_inv_O1 … H) -H *
49   [ -IHL12 He2d #H1 #H2 destruct -e2 K2 J V /2 width=5/
50   | -HL12 -HV12 #He #HLK
51     elim (IHL12 … HLK ?) -IHL12 HLK [ <minus_minus /3 width=5/ | /2/ ] (**) (* a bit slow *)
52   ]
53 ]
54 qed.
55
56 lemma drop_trans_le: ∀d1,e1,L1. ∀L:lenv. ↑[d1, e1] L ≡ L1 →
57                      ∀e2,L2. ↑[0, e2] L2 ≡ L → e2 ≤ d1 →
58                      ∃∃L0. ↑[0, e2] L0 ≡ L1 & ↑[d1 - e2, e1] L2 ≡ L0.
59 #d1 #e1 #L1 #L #H elim H -H d1 e1 L1 L
60 [ #L #e2 #L2 #HL2 #H
61   lapply (le_O_to_eq_O … H) -H #H destruct -e2 /2/
62 | #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H
63   lapply (le_O_to_eq_O … H) -H #H destruct -e2;
64   elim (IHL12 … HL2 ?) -IHL12 HL2 // #L0 #H #HL0
65   lapply (drop_inv_refl … H) -H #H destruct -L1 /3 width=5/
66 | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d
67   elim (drop_inv_O1 … H) -H *
68   [ -He2d IHL12 #H1 #H2 destruct -e2 L /3 width=5/
69   | -HL12 HV12 #He2 #HL2
70     elim (IHL12 … HL2 ?) -IHL12 HL2 L2
71     [ <minus_le_minus_minus_comm /3/ | /2/ ]
72   ]
73 ]
74 qed.
75
76 axiom drop_trans_ge: ∀d1,e1,L1,L. ↑[d1, e1] L ≡ L1 →
77                      ∀e2,L2. ↑[0, e2] L2 ≡ L → d1 ≤ e2 → ↑[0, e1 + e2] L2 ≡ L1.