]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_conversion/cpce_drops.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_conversion / cpce_drops.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 "ground_2/xoa/ex_7_8.ma".
16 include "basic_2/rt_computation/cpms_drops.ma".
17 include "basic_2/rt_conversion/cpce.ma".
18
19 (* CONTEXT-SENSITIVE PARALLEL ETA-CONVERSION FOR TERMS **********************)
20
21 (* Advanced properties ******************************************************)
22
23 lemma cpce_ldef_drops (h) (G) (K) (V):
24       ∀i,L. ⇩*[i] L ≘ K.ⓓV → ⦃G,L⦄ ⊢ #i ⬌η[h] #i.
25 #h #G #K #V #i elim i -i
26 [ #L #HLK
27   lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
28   /2 width=1 by cpce_ldef/
29 | #i #IH #L #HLK
30   elim (drops_inv_succ … HLK) -HLK #Z #Y #HYK #H destruct
31   /3 width=3 by cpce_lref/
32 ]
33 qed.
34
35 lemma cpce_ldec_drops (h) (G) (K) (W):
36       (∀n,p,V,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥) →
37       ∀i,L. ⇩*[i] L ≘ K.ⓛW → ⦃G,L⦄ ⊢ #i ⬌η[h] #i.
38 #h #G #K #W #HW #i elim i -i
39 [ #L #HLK
40   lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
41   /3 width=5 by cpce_ldec/
42 | #i #IH #L #HLK
43   elim (drops_inv_succ … HLK) -HLK #Z #Y #HYK #H destruct
44   /3 width=3 by cpce_lref/
45 ]
46 qed.
47
48 lemma cpce_eta_drops (h) (G) (K) (W):
49       ∀n,p,V,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U →
50       ∀W1. ⦃G,K⦄ ⊢ W ⬌η[h] W1 → ∀V1. ⦃G,K⦄ ⊢ V ⬌η[h] V1 →
51       ∀i,L. ⇩*[i] L ≘ K.ⓛW → ∀W2. ⇧*[↑i] W1 ≘ W2 →
52       ∀V2. ⇧*[↑i] V1 ≘ V2 → ⦃G,L⦄ ⊢ #i ⬌η[h] ⓝW2.+ⓛV2.ⓐ#0.#↑i.
53 #h #G #K #W #n #p #V #U #HWU #W1 #HW1 #V1 #HV1 #i elim i -i
54 [ #L #HLK #W2 #HW12 #V2 #HV12
55   lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
56   /2 width=8 by cpce_eta/
57 | #i #IH #L #HLK #W2 #HW12 #V2 #HV12
58   elim (drops_inv_succ … HLK) -HLK #I #Y #HYK #H destruct
59   elim (lifts_split_trans … HW12 (𝐔❴↑i❵) (𝐔❴1❵)) [| // ] #XW #HXW1 #HXW2
60   elim (lifts_split_trans … HV12 (𝐔❴↑i❵) (𝐔❴1❵)) [| // ] #XV #HXV1 #HXV2
61   /6 width=9 by cpce_lref, lifts_push_lref, lifts_bind, lifts_flat/
62 ]
63 qed.
64
65 lemma cpce_lref_drops (h) (G) (K) (i):
66       ∀T. ⦃G,K⦄ ⊢ #i ⬌η[h] T → ∀j,L. ⇩*[j] L ≘ K →
67       ∀U. ⇧*[j] T ≘ U → ⦃G,L⦄ ⊢ #(j+i) ⬌η[h] U.
68 #h #G #K #i #T #Hi #j elim j -j
69 [ #L #HLK #U #HTU
70   lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
71   lapply (lifts_fwd_isid … HTU ?) -HTU [ // ] #H destruct //
72 | #j #IH #Y #HYK #X #HTX -Hi
73   elim (drops_inv_succ … HYK) -HYK #I #L #HLK #H destruct
74   elim (lifts_split_trans … HTX (𝐔❴j❵) (𝐔❴1❵)) [| // ] #U #HTU #HUX
75   /3 width=3 by cpce_lref/
76 ]
77 qed-.
78
79 (* Advanced inversion lemmas ************************************************)
80
81 axiom cpce_inv_lref_sn_drops_pair (h) (G) (i) (L):
82       ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
83       ∀I,K,W. ⇩*[i] L ≘ K.ⓑ{I}W →
84       ∨∨ ∧∧ Abbr = I & #i = X2
85        | ∧∧ Abst = I & ∀n,p,V,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥ & #i = X2
86        | ∃∃n,p,W1,W2,V,V1,V2,U. Abst = I & ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U
87                               & ⦃G,K⦄ ⊢ W ⬌η[h] W1 & ⇧*[↑i] W1 ≘ W2
88                               & ⦃G,K⦄ ⊢ V ⬌η[h] V1 & ⇧*[↑i] V1 ≘ V2
89                               & ⓝW2.+ⓛV2.ⓐ#0.#(↑i) = X2.
90
91 axiom cpce_inv_lref_sn_drops_ldef (h) (G) (i) (L):
92       ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
93       ∀K,V. ⇩*[i] L ≘ K.ⓓV → #i = X2.
94
95 axiom cpce_inv_lref_sn_drops_ldec (h) (G) (i) (L):
96       ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
97       ∀K,W. ⇩*[i] L ≘ K.ⓛW →
98       ∨∨ ∧∧ ∀n,p,V,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥ & #i = X2
99        | ∃∃n,p,W1,W2,V,V1,V2,U. ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U
100                               & ⦃G,K⦄ ⊢ W ⬌η[h] W1 & ⇧*[↑i] W1 ≘ W2
101                               & ⦃G,K⦄ ⊢ V ⬌η[h] V1 & ⇧*[↑i] V1 ≘ V2
102                               & ⓝW2.+ⓛV2.ⓐ#0.#(↑i) = X2.
103 (*
104 #h #G #i elim i -i
105 [ #L #X2 #HX2 #I #K #HLK
106   lapply (drops_fwd_isid … HLK ?) -HLK [ // ] #H destruct
107   /2 width=1 by cpce_inv_zero_sn/
108 | #i #IH #L0 #X0 #HX0 #J #K #H0
109   elim (drops_inv_succ … H0) -H0 #I #L #HLK #H destruct
110   elim (cpce_inv_lref_sn … HX0) -HX0 #X2 #HX2 #HX20
111   elim (IH … HX2 … HLK) -IH -I -L *
112   [ #HJ #H destruct
113     lapply (lifts_inv_lref1_uni … HX20) -HX20 #H destruct
114     /4 width=7 by or_introl, conj/
115   | #n #p #W #V1 #V2 #W2 #U #HWU #HV12 #HVW2 #H1 #H2 destruct
116     elim (lifts_inv_bind1 … HX20) -HX20 #X2 #X #HWX2 #HX #H destruct
117     elim (lifts_inv_flat1 … HX) -HX #X0 #X1 #H0 #H1 #H destruct
118     lapply (lifts_inv_push_zero_sn … H0) -H0 #H destruct
119     elim (lifts_inv_push_succ_sn … H1) -H1 #j #Hj #H destruct
120     lapply (lifts_inv_lref1_uni … Hj) -Hj #H destruct
121     /4 width=12 by lifts_trans_uni, ex5_7_intro, or_intror/
122   ]
123 ]
124 qed-.
125
126 lemma cpce_inv_zero_sn_drops (h) (G) (i) (L):
127       ∀X2. ⦃G,L⦄ ⊢ #i ⬌η[h] X2 →
128       ∀I,K. ⇩*[i] L ≘ K.ⓘ{I} →
129       (∀n,p,W,V,U. I = BPair Abst W → ⦃G,K⦄ ⊢ W ➡*[n,h] ⓛ{p}V.U → ⊥) →
130       #i = X2.
131 #h #G #i #L #X2 #HX2 #I #K #HLK #HI
132 elim (cpce_inv_lref_sn_drops_bind … HX2 … HLK) -L *
133 [ #_ #H //
134 | #n #p #W #V1 #V2 #W2 #U #HWU #_ #_ #H destruct
135   elim (HI … HWU) -n -p -K -X2 -V1 -V2 -W2 -U -i //
136 ]
137 qed-.
138 *)
139 (* Properties with uniform slicing for local environments *******************)
140
141 axiom cpce_lifts_sn (h) (G):
142       d_liftable2_sn … lifts (cpce h G).
143 (*
144 #h #G #K #T1 #T2 #H elim H -G -K -T1 -T2
145 [ #G #K #s #b #f #L #HLK #X #HX
146   lapply (lifts_inv_sort1 … HX) -HX #H destruct
147   /2 width=3 by cpce_sort, lifts_sort, ex2_intro/
148 | #G #i #b #f #L #HLK #X #HX
149   elim (lifts_inv_lref1 … HX) -HX #j #Hf #H destruct
150   @(ex2_intro … (#j))
151   [ /2 width=1 by lifts_lref/
152   | @cpce_zero_drops #n #p #Y #W #V #U #HY #_
153     elim (drops_inv_atom2 … HLK) -HLK #j1 #g #HLK #Hg
154     elim (after_at_fwd … Hf … Hg) -f #j2 #_ #Hj -g -i
155     lapply (at_inv_uni … Hj) -Hj #H destruct
156     lapply (drops_conf … HLK … HY ??) -L [3:|*: // ] #H
157     elim (drops_inv_atom1 … H) -H #H #_ destruct
158   ]
159 | #G #K #I #HI #b #f #L #HLK #X #HX
160   elim (lifts_inv_lref1 … HX) -HX #j #Hf #H destruct
161   @(ex2_intro … (#j))
162   [ /2 width=1 by lifts_lref/
163   | elim (drops_split_trans_bind2 … HLK … Hf) -HLK -Hf #J #Y1 #HY1 #HK #HIJ
164     @cpce_zero_drops #n #p #Y2 #W #V #U #HY2 #HWU
165     lapply (drops_mono … HY2 … HY1) -L #H destruct
166     elim (liftsb_inv_pair_dx … HIJ) -HIJ #X #HXW #H destruct
167     elim (cpms_inv_lifts_sn … HWU … HK … HXW) -b -Y1 -W #X0 #H #HXU
168     elim (lifts_inv_bind2 … H) -H #V0 #U0 #_ #_ #H destruct -f -j -V -U
169     /2 width=7 by/
170   ]
171 | #n #p #G #K #W #V1 #V2 #W2 #U #HWU #_ #HVW2 #IH #b #f #L #HLK #X #HX
172   elim (lifts_inv_lref1 … HX) -HX #j #Hf #H destruct
173   elim (drops_split_trans_bind2 … HLK … Hf) -HLK #J #Y #HY #HK #HIJ
174   elim (liftsb_inv_pair_sn … HIJ) -HIJ #W0 #HW0 #H destruct
175   elim (cpms_lifts_sn … HWU … HK … HW0) -HWU -HW0 #X #H #HWU0
176   elim (lifts_inv_bind1 … H) -H #V0 #U0 #HV10 #HU0 #H destruct
177   elim (IH … HK … HV10) -IH -HK -HV10 #VX #HV2X #HV0X
178   elim (lifts_total W2 f) #WX2 #HWX2
179   lapply (lifts_trans … HVW2 … HWX2 ??) [3:|*: // ] -HVW2 #HVX2
180   @(ex2_intro … (+ⓛWX2.ⓐ#O.#(↑j)))
181   [ /5 width=1 by lifts_lref, lifts_bind, lifts_flat, at_S1/
182   | /4 width=18 by cpce_eta_drops, lifts_conf, after_uni_succ_dx/
183   ]
184 | #I #G #K #T #U #i #_ #HTU #IH #b #f #L #HLK #X #HX
185   elim (lifts_inv_lref1 … HX) -HX #x #Hf #H destruct
186   elim (at_inv_nxx … Hf) -Hf [|*: // ] #j #Hf #H destruct
187   elim (drops_split_trans_bind2 … HLK) -HLK [|*: // ] #Z #Y #HLY #HYK #_ -I
188   lapply (drops_isuni_fwd_drop2 … HLY) -HLY [ // ] #HLY  
189   elim (IH … HYK) -IH -HYK [|*: /2 width=2 by lifts_lref/ ] -i #T0 #HT0 #Hj
190   elim (lifts_total U f) #U0 #HU0
191   lapply (lifts_trans … HTU … HU0 ??) [3:|*: // ] -HTU #HTU0
192   lapply (lifts_conf … HT0 … HTU0 ??) -T
193   [3:|*: /2 width=3 by after_uni_succ_dx/ ] #HTU0 >plus_S1
194   /3 width=7 by cpce_lref_drops, ex2_intro/
195 | #G #K #l #b #f #L #HLK #X #HX
196   lapply (lifts_inv_gref1 … HX) -HX #H destruct
197   /2 width=3 by cpce_gref, lifts_gref, ex2_intro/
198 | #p #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #b #f #L #HLK #X #HX
199   elim (lifts_inv_bind1 … HX) -HX #W1 #U1 #HVW1 #HTU1 #H destruct
200   elim (IHV … HLK … HVW1) -IHV #W2 #HVW2 #HW12
201   elim (IHT … HTU1) -IHT -HTU1 [|*: /3 width=3 by drops_skip, ext2_pair/ ] -HVW1 #U2 #HTU2 #HU12
202   /3 width=5 by cpce_bind, lifts_bind, ex2_intro/
203 | #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #b #f #L #HLK #X #HX
204   elim (lifts_inv_flat1 … HX) -HX #W1 #U1 #HVW1 #HTU1 #H destruct
205   elim (IHV … HLK … HVW1) -IHV -HVW1 #W2 #HVW2 #HW12
206   elim (IHT … HLK … HTU1) -IHT -HTU1 -HLK #U2 #HTU2 #HU12
207   /3 width=5 by cpce_flat, lifts_flat, ex2_intro/
208 ]
209 qed-.
210 *)
211 lemma cpce_lifts_bi (h) (G):
212       d_liftable2_bi … lifts (cpce h G).
213 /3 width=12 by cpce_lifts_sn, d_liftable2_sn_bi, lifts_mono/ qed-.
214
215 (* Inversion lemmas with uniform slicing for local environments *************)
216
217 axiom cpce_inv_lifts_sn (h) (G):
218       d_deliftable2_sn … lifts (cpce h G).
219 (*
220 #h #G #K #T1 #T2 #H elim H -G -K -T1 -T2
221 [ #G #K #s #b #f #L #HLK #X #HX
222   lapply (lifts_inv_sort2 … HX) -HX #H destruct
223   /2 width=3 by cpce_sort, lifts_sort, ex2_intro/
224 | #G #i #b #f #L #HLK #X #HX
225   elim (lifts_inv_lref2 … HX) -HX #j #Hf #H destruct
226   @(ex2_intro … (#j))
227   [ /2 width=1 by lifts_lref/
228   | @cpce_zero_drops #n #p #Y #W #V #U #HY #_ -n -p -G -V -U -i
229     elim (drops_inv_atom1 … HLK) -HLK #H #_ destruct -b -f
230     elim (drops_inv_atom1 … HY) -HY #H #_ destruct
231   ]
232 | #G #K #I #HI #b #f #L #HLK #X #HX
233   elim (lifts_inv_lref2 … HX) -HX #j #Hf #H destruct
234   @(ex2_intro … (#j))
235   [ /2 width=1 by lifts_lref/
236   | elim (at_inv_xxp … Hf) -Hf [| // ] #g #H1 #H2 destruct
237     elim (drops_inv_skip1 … HLK) -HLK #J #Y #HKY #HIJ #H destruct
238     @cpce_zero #n #p #W #V #U #H #HWU destruct
239     elim (liftsb_inv_pair_sn … HIJ) -HIJ #X #HXW #H destruct
240     elim (cpms_lifts_sn … HWU … HKY … HXW) -b -Y -W #X0 #H #HXU
241     elim (lifts_inv_bind1 … H) -H #V0 #U0 #_ #_ #H destruct -V -U
242     /2 width=7 by/
243   ]
244 | #n #p #G #K #W #V1 #V2 #W2 #U #HWU #_ #HVW2 #IH #b #f #L #HLK #X #HX
245   elim (lifts_inv_lref2 … HX) -HX #j #Hf #H destruct
246   elim (at_inv_xxp … Hf) -Hf [| // ] #g #H1 #H2 destruct
247   elim (drops_inv_skip1 … HLK) -HLK #J #Y #HKY #HIJ #H destruct
248   elim (liftsb_inv_pair_dx … HIJ) -HIJ #W0 #HW0 #H destruct
249   elim (cpms_inv_lifts_sn … HWU … HKY … HW0) -HWU -HW0 #X #H #HWU0
250   elim (lifts_inv_bind2 … H) -H #V0 #U0 #HV10 #HU0 #H destruct
251   elim (IH … HKY … HV10) -IH -HKY -HV10 #VX #HV2X #HV0X
252   lapply (lifts_trans … HV2X … HVW2 (↑g) ?)
253   [ /3 width=5 by after_isid_sn, after_next/ ] -V2 #H
254   elim (lifts_split_trans … H 𝐔❴1❵ (⫯g) ?)
255   [| /3 width=7 by after_isid_dx, after_push/ ] #VX2 #HVX2 #HVW2
256   /5 width=10 by cpce_eta, lifts_flat, lifts_bind, lifts_lref, ex2_intro/
257 | #I #G #K #T #U #i #_ #HTU #IH #b #f #L #HLK #X #HX
258   elim (lifts_inv_lref2 … HX) -HX #x #Hf #H destruct
259 (**) (* this part should be a lemma *)
260   elim (at_inv_xxn … Hf) -Hf [2,4: // ] *
261   [ #g #j #Hij #H1 #H2 destruct
262     elim (drops_inv_skip1 … HLK) -HLK #J #Y #HLK #_ #H destruct -I
263   | #g #Hij #H destruct
264     lapply (drops_inv_drop1 … HLK) -HLK #HLK
265   ]
266 (**)
267   elim (IH … HLK) -IH -HLK [1,4:|*: /2 width=2 by lifts_lref/ ] -i #T0 #HT0 #Hj
268   lapply (lifts_trans … HT0 … HTU (↑g) ?)
269   [1,3: /3 width=5 by after_isid_sn, after_next/ ] -T #H
270   elim (lifts_split_trans … H 𝐔❴1❵ (⫯g) ?)
271   [2,4: /3 width=7 by after_isid_dx, after_push/ ] #U0 #HTU0 #HU0
272   /3 width=5 by cpce_lref, ex2_intro/
273 | #G #K #l #b #f #L #HLK #X #HX
274   lapply (lifts_inv_gref2 … HX) -HX #H destruct
275   /2 width=3 by cpce_gref, lifts_gref, ex2_intro/
276 | #p #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #b #f #L #HLK #X #HX
277   elim (lifts_inv_bind2 … HX) -HX #W1 #U1 #HVW1 #HTU1 #H destruct
278   elim (IHV … HLK … HVW1) -IHV #W2 #HVW2 #HW12
279   elim (IHT … HTU1) -IHT -HTU1 [|*: /3 width=3 by drops_skip, ext2_pair/ ] -HVW1 #U2 #HTU2 #HU12
280   /3 width=5 by cpce_bind, lifts_bind, ex2_intro/
281 | #I #G #K #V1 #V2 #T1 #T2 #_ #_ #IHV #IHT #b #f #L #HLK #X #HX
282   elim (lifts_inv_flat2 … HX) -HX #W1 #U1 #HVW1 #HTU1 #H destruct
283   elim (IHV … HLK … HVW1) -IHV -HVW1 #W2 #HVW2 #HW12
284   elim (IHT … HLK … HTU1) -IHT -HTU1 -HLK #U2 #HTU2 #HU12
285   /3 width=5 by cpce_flat, lifts_flat, ex2_intro/
286 ]
287 qed-.
288 *)
289 lemma cpce_inv_lifts_bi (h) (G):
290       d_deliftable2_bi … lifts (cpce h G).
291 /3 width=12 by cpce_inv_lifts_sn, d_deliftable2_sn_bi, lifts_inj/ qed-.