]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_ldrop.ma
- degree assignment, static type assignment, iterated static type
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / ldrop_ldrop.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/relocation/lift_lift.ma".
16 include "basic_2/relocation/ldrop.ma".
17
18 (* DROPPING *****************************************************************)
19
20 (* Main properties **********************************************************)
21
22 (* Basic_1: was: drop_mono *)
23 theorem ldrop_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 -d -e -L -L1
26 [ #d #L2 #H elim (ldrop_inv_atom1 … H) -H //
27 | #K #I #V #L2 #HL12 <(ldrop_inv_O2 … HL12) -L2 //
28 | #L #K #I #V #e #_ #IHLK #L2 #H
29   lapply (ldrop_inv_ldrop1 … H ?) -H // /2 width=1/
30 | #L #K1 #I #T #V1 #d #e #_ #HVT1 #IHLK1 #X #H
31   elim (ldrop_inv_skip1 … H) -H // <minus_plus_m_m #K2 #V2 #HLK2 #HVT2 #H destruct
32   >(lift_inj … HVT1 … HVT2) -HVT1 -HVT2
33   >(IHLK1 … HLK2) -IHLK1 -HLK2 //
34 ]
35 qed-.
36
37 (* Basic_1: was: drop_conf_ge *)
38 theorem ldrop_conf_ge: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 →
39                        ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 + e1 ≤ e2 →
40                        ⇩[0, e2 - e1] L1 ≡ L2.
41 #d1 #e1 #L #L1 #H elim H -d1 -e1 -L -L1 //
42 [ #L #K #I #V #e #_ #IHLK #e2 #L2 #H #He2
43   lapply (ldrop_inv_ldrop1 … H ?) -H /2 width=2/ #HL2
44   <minus_plus >minus_minus_comm /3 width=1/
45 | #L #K #I #V1 #V2 #d #e #_ #_ #IHLK #e2 #L2 #H #Hdee2
46   lapply (transitive_le 1 … Hdee2) // #He2
47   lapply (ldrop_inv_ldrop1 … H ?) -H // -He2 #HL2
48   lapply (transitive_le (1 + e) … Hdee2) // #Hee2
49   @ldrop_ldrop_lt >minus_minus_comm /3 width=1/ (**) (* explicit constructor *)
50 ]
51 qed.
52
53 (* Note: apparently this was missing in basic_1 *)
54 theorem ldrop_conf_be: ∀L0,L1,d1,e1. ⇩[d1, e1] L0 ≡ L1 →
55                        ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 →
56                        ∃∃L. ⇩[0, d1 + e1 - e2] L2 ≡ L & ⇩[0, d1] L1 ≡ L.
57 #L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
58 [ #d1 #L2 #e2 #H #Hd1 #_ elim (ldrop_inv_atom1 … H) -H #H1 #H2 destruct
59   <(le_n_O_to_eq … Hd1) -d1 /2 width=3/
60 | normalize #L #I #V #L2 #e2 #HL2 #_ #He2
61   lapply (le_n_O_to_eq … He2) -He2 #H destruct
62   lapply (ldrop_inv_O2 … HL2) -HL2 #H destruct /2 width=3/
63 | normalize #L0 #K0 #I #V1 #e1 #HLK0 #IHLK0 #L2 #e2 #H #_ #He21
64   lapply (ldrop_inv_O1_pair1 … H) -H * * #He2 #HL20
65   [ -IHLK0 -He21 destruct <minus_n_O /3 width=3/
66   | -HLK0 <minus_le_minus_minus_comm //
67     elim (IHLK0 … HL20) -L0 // /2 width=1/ /2 width=3/
68   ]
69 | #L0 #K0 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHLK0 #L2 #e2 #H #Hd1e2 #He2de1
70   elim (le_inv_plus_l … Hd1e2) #_ #He2
71   <minus_le_minus_minus_comm //
72   lapply (ldrop_inv_ldrop1 … H ?) -H // #HL02
73   elim (IHLK0 … HL02) -L0 /2 width=1/ /3 width=3/
74 ]
75 qed.
76
77 (* Note: apparently this was missing in basic_1 *)
78 theorem ldrop_conf_le: ∀L0,L1,d1,e1. ⇩[d1, e1] L0 ≡ L1 →
79                        ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 →
80                        ∃∃L. ⇩[0, e2] L1 ≡ L & ⇩[d1 - e2, e1] L2 ≡ L.
81 #L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1
82 [ #d1 #L2 #e2 #H
83   elim (ldrop_inv_atom1 … H) -H #H destruct /2 width=3/
84 | #K0 #I #V0 #L2 #e2 #H1 #H2
85   lapply (le_n_O_to_eq … H2) -H2 #H destruct
86   lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /2 width=3/
87 | #K0 #K1 #I #V0 #e1 #HK01 #_ #L2 #e2 #H1 #H2
88   lapply (le_n_O_to_eq … H2) -H2 #H destruct
89   lapply (ldrop_inv_pair1 … H1) -H1 #H destruct /3 width=3/
90 | #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV10 #IHK01 #L2 #e2 #H #He2d1
91   elim (ldrop_inv_O1_pair1 … H) -H *
92   [ -IHK01 -He2d1 #H1 #H2 destruct /3 width=5/
93   | -HK01 -HV10 #He2 #HK0L2
94     elim (IHK01 … HK0L2) -IHK01 -HK0L2 /2 width=1/ >minus_le_minus_minus_comm // /3 width=3/
95   ]
96 ]
97 qed.
98
99 (* Basic_1: was: drop_trans_ge *)
100 theorem ldrop_trans_ge: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L →
101                         ∀e2,L2. ⇩[0, e2] L ≡ L2 → d1 ≤ e2 → ⇩[0, e1 + e2] L1 ≡ L2.
102 #d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L //
103 [ /3 width=1/
104 | #L1 #L2 #I #V1 #V2 #d #e #H_ #_ #IHL12 #e2 #L #H #Hde2
105   lapply (lt_to_le_to_lt 0 … Hde2) // #He2
106   lapply (lt_to_le_to_lt … (e + e2) He2 ?) // #Hee2
107   lapply (ldrop_inv_ldrop1 … H ?) -H // #HL2
108   @ldrop_ldrop_lt // >le_plus_minus // @IHL12 /2 width=1/ (**) (* explicit constructor *)
109 ]
110 qed.
111
112 (* Basic_1: was: drop_trans_le *)
113 theorem ldrop_trans_le: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L →
114                         ∀e2,L2. ⇩[0, e2] L ≡ L2 → e2 ≤ d1 →
115                         ∃∃L0. ⇩[0, e2] L1 ≡ L0 & ⇩[d1 - e2, e1] L0 ≡ L2.
116 #d1 #e1 #L1 #L #H elim H -d1 -e1 -L1 -L
117 [ #d #e2 #L2 #H
118   elim (ldrop_inv_atom1 … H) -H /2 width=3/
119 | #K #I #V #e2 #L2 #HL2 #H
120   lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/
121 | #L1 #L2 #I #V #e #_ #IHL12 #e2 #L #HL2 #H
122   lapply (le_n_O_to_eq … H) -H #H destruct
123   elim (IHL12 … HL2) -IHL12 -HL2 // #L0 #H #HL0
124   lapply (ldrop_inv_O2 … H) -H #H destruct /3 width=5/
125 | #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #IHL12 #e2 #L #H #He2d
126   elim (ldrop_inv_O1_pair1 … H) -H *
127   [ -He2d -IHL12 #H1 #H2 destruct /3 width=5/
128   | -HL12 -HV12 #He2 #HL2
129     elim (IHL12 … HL2) -L2 [ >minus_le_minus_minus_comm // /3 width=3/ | /2 width=1/ ]
130   ]
131 ]
132 qed.
133
134 (* Basic_1: was: drop_conf_rev *)
135 axiom ldrop_div: ∀e1,L1,L. ⇩[0, e1] L1 ≡ L → ∀e2,L2. ⇩[0, e2] L2 ≡ L →
136                  ∃∃L0. ⇩[0, e1] L0 ≡ L2 & ⇩[e1, e2] L0 ≡ L1.
137
138 (* Advanced properties ******************************************************)
139
140 lemma l_liftable_llstar: ∀R. l_liftable R → ∀l. l_liftable (llstar … R l).
141 #R #HR #l #K #T1 #T2 #H @(lstar_ind_r … l T2 H) -l -T2
142 [ #L #d #e #_ #U1 #HTU1 #U2 #HTU2 -HR -K
143   >(lift_mono … HTU2 … HTU1) -T1 -U2 -d -e //
144 | #l #T #T2 #_ #HT2 #IHT1 #L #d #e #HLK #U1 #HTU1 #U2 #HTU2
145   elim (lift_total T d e) /3 width=11 by lstar_dx/ (**) (* auto too slow without trace *)
146 ]
147 qed.
148
149 (* Basic_1: was: drop_conf_lt *)
150 lemma ldrop_conf_lt: ∀d1,e1,L,L1. ⇩[d1, e1] L ≡ L1 →
151                      ∀e2,K2,I,V2. ⇩[0, e2] L ≡ K2. ⓑ{I} V2 →
152                      e2 < d1 → let d ≝ d1 - e2 - 1 in
153                      ∃∃K1,V1. ⇩[0, e2] L1 ≡ K1. ⓑ{I} V1 &
154                               ⇩[d, e1] K2 ≡ K1 & ⇧[d, e1] V1 ≡ V2.
155 #d1 #e1 #L #L1 #H1 #e2 #K2 #I #V2 #H2 #He2d1
156 elim (ldrop_conf_le … H1 … H2) -L [2: /2 width=2/] #K #HL1K #HK2
157 elim (ldrop_inv_skip1 … HK2) -HK2 [2: /2 width=1/] #K1 #V1 #HK21 #HV12 #H destruct /2 width=5/
158 qed.
159
160 (* Note: apparently this was missing in basic_1 *)
161 lemma ldrop_trans_lt: ∀d1,e1,L1,L. ⇩[d1, e1] L1 ≡ L →
162                       ∀e2,L2,I,V2. ⇩[0, e2] L ≡ L2.ⓑ{I}V2 →
163                       e2 < d1 → let d ≝ d1 - e2 - 1 in
164                       ∃∃L0,V0. ⇩[0, e2] L1 ≡ L0.ⓑ{I}V0 &
165                                ⇩[d, e1] L0 ≡ L2 & ⇧[d, e1] V2 ≡ V0.
166 #d1 #e1 #L1 #L #HL1 #e2 #L2 #I #V2 #HL2 #Hd21
167 elim (ldrop_trans_le … HL1 … HL2) -L [2: /2 width=1/ ] #L0 #HL10 #HL02
168 elim (ldrop_inv_skip2 … HL02) -HL02 [2: /2 width=1/ ] #L #V1 #HL2 #HV21 #H destruct /2 width=5/
169 qed-.
170
171 lemma ldrop_trans_ge_comm: ∀d1,e1,e2,L1,L2,L.
172                            ⇩[d1, e1] L1 ≡ L → ⇩[0, e2] L ≡ L2 → d1 ≤ e2 →
173                            ⇩[0, e2 + e1] L1 ≡ L2.
174 #e1 #e1 #e2 >commutative_plus /2 width=5/
175 qed.
176
177 lemma ldrop_conf_div: ∀I1,L,K,V1,e1. ⇩[0, e1] L ≡ K. ⓑ{I1} V1 →
178                       ∀I2,V2,e2. ⇩[0, e2] L ≡ K. ⓑ{I2} V2 →
179                       ∧∧ e1 = e2 & I1 = I2 & V1 = V2.
180 #I1 #L #K #V1 #e1 #HLK1 #I2 #V2 #e2 #HLK2
181 elim (le_or_ge e1 e2) #He
182 [ lapply (ldrop_conf_ge … HLK1 … HLK2 ?)
183 | lapply (ldrop_conf_ge … HLK2 … HLK1 ?)
184 ] -HLK1 -HLK2 // #HK
185 lapply (ldrop_fwd_length_minus2 … HK) #H
186 elim (discr_minus_x_xy … H) -H
187 [1,3: normalize <plus_n_Sm #H destruct ]
188 #H >H in HK; #HK
189 lapply (ldrop_inv_O2 … HK) -HK #H destruct
190 lapply (inv_eq_minus_O … H) -H /3 width=1/
191 qed-.