]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/substitution/drop_drop.ma
b8a790fb597402712c04b7ba40647aa1922b9601
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / substitution / drop_drop.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 "Basic_2/substitution/lift_lift.ma".
16 include "Basic_2/substitution/drop.ma".
17
18 (* DROPPING *****************************************************************)
19
20 (* Main properties **********************************************************)
21
22 (* Basic_1: was: drop_mono *)
23 theorem drop_mono: ∀d,e,L,L1. ↓[d, e] L ≡ L1 →
24                    ∀L2. ↓[d, e] L ≡ L2 → L1 = L2.
25 #d #e #L #L1 #H elim H -H d e L L1
26 [ #d #e #L2 #H
27   >(drop_inv_atom1 … H) -H L2 //
28 | #K #I #V #L2 #HL12
29    <(drop_inv_refl … HL12) -HL12 L2 //
30 | #L #K #I #V #e #_ #IHLK #L2 #H
31   lapply (drop_inv_drop1 … H ?) -H /2/
32 | #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H
33   elim (drop_inv_skip1 … H ?) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct -X
34   >(lift_inj … HVT1 … HVT2) -HVT1 HVT2
35   >(IHLK1 … HLK2) -IHLK1 HLK2 // 
36 ]
37 qed.
38
39 (* Basic_1: was: drop_conf_ge *)
40 theorem drop_conf_ge: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 →
41                       ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
42                       ↓[0, e2 - e1] L1 ≡ L2.
43 #d1 #e1 #L #L1 #H elim H -H d1 e1 L L1
44 [ #d #e #e2 #L2 #H
45   >(drop_inv_atom1 … H) -H L2 //
46 | //
47 | #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2
48   lapply (drop_inv_drop1 … H ?) -H /2/ #HL2
49   <minus_plus_comm /3/
50 | #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2
51   lapply (transitive_le 1 … Hdee2) // #He2
52   lapply (drop_inv_drop1 … H ?) -H // -He2 #HL2
53   lapply (transitive_le (1 + e) … Hdee2) // #Hee2
54   @drop_drop_lt >minus_minus_comm /3/ (**) (* explicit constructor *)
55 ]
56 qed.
57
58 (* Basic_1: was: drop_conf_lt *)
59 theorem drop_conf_lt: ∀d1,e1,L,L1. ↓[d1, e1] L ≡ L1 →
60                       ∀e2,K2,I,V2. ↓[0, e2] L ≡ K2. 𝕓{I} V2 →
61                       e2 < d1 → let d ≝ d1 - e2 - 1 in
62                       ∃∃K1,V1. ↓[0, e2] L1 ≡ K1. 𝕓{I} V1 &
63                                ↓[d, e1] K2 ≡ K1 & ↑[d, e1] V1 ≡ V2.
64 #d1 #e1 #L #L1 #H elim H -H d1 e1 L L1
65 [ #d #e #e2 #K2 #I #V2 #H
66   lapply (drop_inv_atom1 … H) -H #H destruct
67 | #L #I #V #e2 #K2 #J #V2 #_ #H
68   elim (lt_zero_false … H)
69 | #L1 #L2 #I #V #e #_ #_ #e2 #K2 #J #V2 #_ #H
70   elim (lt_zero_false … H)
71 | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #K2 #J #V #H #He2d
72   elim (drop_inv_O1 … H) -H *
73   [ -IHL12 He2d #H1 #H2 destruct -e2 K2 J V /2 width=5/
74   | -HL12 -HV12 #He #HLK
75     elim (IHL12 … HLK ?) -IHL12 HLK [ <minus_minus /3 width=5/ | /2/ ] (**) (* a bit slow *)
76   ]
77 ]
78 qed.
79
80 (* Basic_1: was: drop_trans_le *)
81 theorem drop_trans_le: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L →
82                        ∀e2,L2. ↓[0, e2] L ≡ L2 → e2 ≤ d1 →
83                        ∃∃L0. ↓[0, e2] L1 ≡ L0 & ↓[d1 - e2, e1] L0 ≡ L2.
84 #d1 #e1 #L1 #L #H elim H -H d1 e1 L1 L
85 [ #d #e #e2 #L2 #H
86   >(drop_inv_atom1 … H) -H L2 /2/
87 | #K #I #V #e2 #L2 #HL2 #H
88   lapply (le_O_to_eq_O … H) -H #H destruct -e2 /2/
89 | #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H
90   lapply (le_O_to_eq_O … H) -H #H destruct -e2;
91   elim (IHL12 … HL2 ?) -IHL12 HL2 // #L0 #H #HL0
92   lapply (drop_inv_refl … H) -H #H destruct -L1 /3 width=5/
93 | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d
94   elim (drop_inv_O1 … H) -H *
95   [ -He2d IHL12 #H1 #H2 destruct -e2 L /3 width=5/
96   | -HL12 HV12 #He2 #HL2
97     elim (IHL12 … HL2 ?) -IHL12 HL2 L2
98     [ >minus_le_minus_minus_comm // /3/ | /2/ ]
99   ]
100 ]
101 qed.
102
103 (* Basic_1: was: drop_trans_ge *)
104 theorem drop_trans_ge: ∀d1,e1,L1,L. ↓[d1, e1] L1 ≡ L →
105                        ∀e2,L2. ↓[0, e2] L ≡ L2 → d1 ≤ e2 → ↓[0, e1 + e2] L1 ≡ L2.
106 #d1 #e1 #L1 #L #H elim H -H d1 e1 L1 L
107 [ #d #e #e2 #L2 #H
108   >(drop_inv_atom1 … H) -H L2 //
109 | //
110 | /3/
111 | #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2
112   lapply (lt_to_le_to_lt 0 … Hde2) // #He2
113   lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2
114   lapply (drop_inv_drop1 … H ?) -H // #HL2
115   @drop_drop_lt // >le_plus_minus // @IHL12 /2/ (**) (* explicit constructor *)
116 ]
117 qed.
118
119 theorem drop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L.
120                             ↓[d1, e1] L1 ≡ L → ↓[0, e2] L ≡ L2 → d1 ≤ e2 →
121                             ↓[0, e2 + e1] L1 ≡ L2.
122 #e1 #e1 #e2 >commutative_plus /2 width=5/
123 qed.
124
125 (* Basic_1: was: drop_conf_rev *)
126 axiom drop_div: ∀e1,L1,L. ↓[0, e1] L1 ≡ L → ∀e2,L2. ↓[0, e2] L2 ≡ L →
127                 ∃∃L0. ↓[0, e1] L0 ≡ L2 & ↓[e1, e2] L0 ≡ L1.