]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/static/frees_drops.ma
more additions and corrections for the article
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / frees_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/relocation/nstream_coafter.ma".
16 include "static_2/relocation/drops_drops.ma".
17 include "static_2/static/frees_fqup.ma".
18
19 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
20
21 (* Advanced properties ******************************************************)
22
23 lemma frees_atom_drops: 
24       ∀b,L,i. ⬇*[b,𝐔❴i❵] L ≘ ⋆ →
25       ∀f. 𝐈⦃f⦄ → L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i]↑f.
26 #b #L elim L -L /2 width=1 by frees_atom/
27 #L #I #IH *
28 [ #H lapply (drops_fwd_isid … H ?) -H // #H destruct
29 | /4 width=3 by frees_lref, drops_inv_drop1/
30 ]
31 qed.
32
33 lemma frees_pair_drops:
34       ∀f,K,V. K ⊢ 𝐅+⦃V⦄ ≘ f → 
35       ∀i,I,L. ⬇*[i] L ≘ K.ⓑ{I}V → L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i] ↑f.
36 #f #K #V #Hf #i elim i -i
37 [ #I #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_pair/
38 | #i #IH #I #L #H elim (drops_inv_succ … H) -H /3 width=2 by frees_lref/
39 ]
40 qed.
41
42 lemma frees_unit_drops:
43       ∀f.  𝐈⦃f⦄ → ∀I,K,i,L. ⬇*[i] L ≘ K.ⓤ{I} →
44       L ⊢ 𝐅+⦃#i⦄ ≘ ⫯*[i] ↑f.
45 #f #Hf #I #K #i elim i -i
46 [ #L #H lapply (drops_fwd_isid … H ?) -H /2 width=1 by frees_unit/
47 | #i #IH #Y #H elim (drops_inv_succ … H) -H
48   #J #L #HLK #H destruct /3 width=1 by frees_lref/
49 ]
50 qed.
51 (*
52 lemma frees_sort_pushs:
53       ∀f,K,s. K ⊢ 𝐅+⦃⋆s⦄ ≘ f →
54       ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅+⦃⋆s⦄ ≘ ⫯*[i] f.
55 #f #K #s #Hf #i elim i -i
56 [ #L #H lapply (drops_fwd_isid … H ?) -H //
57 | #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_sort/
58 ]
59 qed.
60 *)
61 lemma frees_lref_pushs:
62       ∀f,K,j. K ⊢ 𝐅+⦃#j⦄ ≘ f →
63       ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅+⦃#(i+j)⦄ ≘ ⫯*[i] f.
64 #f #K #j #Hf #i elim i -i
65 [ #L #H lapply (drops_fwd_isid … H ?) -H //
66 | #i #IH #L #H elim (drops_inv_succ … H) -H
67   #I #Y #HYK #H destruct /3 width=1 by frees_lref/
68 ]
69 qed.
70 (*
71 lemma frees_gref_pushs:
72       ∀f,K,l. K ⊢ 𝐅+⦃§l⦄ ≘ f →
73       ∀i,L. ⬇*[i] L ≘ K → L ⊢ 𝐅+⦃§l⦄ ≘ ⫯*[i] f.
74 #f #K #l #Hf #i elim i -i
75 [ #L #H lapply (drops_fwd_isid … H ?) -H //
76 | #i #IH #L #H elim (drops_inv_succ … H) -H /3 width=1 by frees_gref/
77 ]
78 qed.
79 *)
80 (* Advanced inversion lemmas ************************************************)
81
82 lemma frees_inv_lref_drops:
83       ∀L,i,f. L ⊢ 𝐅+⦃#i⦄ ≘ f →
84       ∨∨ ∃∃g. ⬇*[Ⓕ,𝐔❴i❵] L ≘ ⋆ & 𝐈⦃g⦄ & f = ⫯*[i] ↑g
85        | ∃∃g,I,K,V. K ⊢ 𝐅+⦃V⦄ ≘ g & ⬇*[i] L ≘ K.ⓑ{I}V & f = ⫯*[i] ↑g
86        | ∃∃g,I,K. ⬇*[i] L ≘ K.ⓤ{I} & 𝐈⦃g⦄ & f = ⫯*[i] ↑g.
87 #L elim L -L
88 [ #i #g | #L #I #IH * [ #g cases I -I [ #I | #I #V ] -IH | #i #g ] ] #H
89 [ elim (frees_inv_atom … H) -H #f #Hf #H destruct
90   /3 width=3 by or3_intro0, ex3_intro/
91 | elim (frees_inv_unit … H) -H #f #Hf #H destruct
92   /4 width=3 by drops_refl, or3_intro2, ex3_3_intro/
93 | elim (frees_inv_pair … H) -H #f #Hf #H destruct
94   /4 width=7 by drops_refl, or3_intro1, ex3_4_intro/
95 | elim (frees_inv_lref … H) -H #f #Hf #H destruct
96   elim (IH … Hf) -IH -Hf *
97   [ /4 width=3 by drops_drop, or3_intro0, ex3_intro/
98   | /4 width=7 by drops_drop, or3_intro1, ex3_4_intro/
99   | /4 width=3 by drops_drop, or3_intro2, ex3_3_intro/
100   ]
101 ]
102 qed-.
103
104 (* Properties with generic slicing for local environments *******************)
105
106 lemma frees_lifts:
107       ∀b,f1,K,T. K ⊢ 𝐅+⦃T⦄ ≘ f1 →
108       ∀f,L. ⬇*[b,f] L ≘ K → ∀U. ⬆*[f] T ≘ U →
109       ∀f2. f ~⊚ f1 ≘ f2 → L ⊢ 𝐅+⦃U⦄ ≘ f2.
110 #b #f1 #K #T #H lapply (frees_fwd_isfin … H) elim H -f1 -K -T
111 [ #f1 #K #s #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
112   lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
113   >(lifts_inv_sort1 … H2) -U /2 width=1 by frees_sort/
114 | #f1 #i #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
115   elim (lifts_inv_lref1 … H2) -H2 #j #Hij #H destruct
116   elim (coafter_fwd_xnx_pushs … Hij H3) -H3 #g2 #Hg2 #H2 destruct
117   lapply (coafter_isid_inv_dx … Hg2 … Hf1) -f1 #Hf2
118   elim (drops_inv_atom2 … H1) -H1 #n #g #H1 #Hf
119   elim (after_at_fwd … Hij … Hf) -f #x #_ #Hj -g -i
120   lapply (at_inv_uni … Hj) -Hj #H destruct
121   /3 width=8 by frees_atom_drops, drops_trans/
122 | #f1 #I #K #V #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
123   lapply (isfin_inv_next … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
124   lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
125   elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #H
126   elim (liftsb_inv_pair_sn … H) -H #W #HVW #H destruct
127   elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
128   lapply (IH … HYK … HVW … H3) -IH -H3 -HYK -HVW //
129   /2 width=5 by frees_pair_drops/
130 | #f1 #I #K #Hf1 #_ #f #L #H1 #U #H2 #f2 #H3
131   lapply (lifts_inv_lref1 … H2) -H2 * #j #Hf #H destruct
132   elim (coafter_fwd_xnx_pushs … Hf H3) -H3 #g2 #H3 #H2 destruct
133   lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hg2
134   elim (drops_split_trans_bind2 … H1 … Hf) -H1 -Hf #Z #Y #HLY #_ #H
135   lapply (liftsb_inv_unit_sn … H) -H #H destruct
136   /2 width=3 by frees_unit_drops/
137 | #f1 #I #K #i #_ #IH #Hf1 #f #L #H1 #U #H2 #f2 #H3
138   lapply (isfin_inv_push … Hf1 ??) -Hf1 [3: |*: // ] #Hf1
139   lapply (lifts_inv_lref1 … H2) -H2 * #x #Hf #H destruct
140   elim (at_inv_nxx … Hf) -Hf [ |*: // ] #j #Hf #H destruct
141   elim (drops_split_trans_bind2 … H1) -H1 [ |*: // ] #Z #Y #HLY #HYK #_
142   elim (coafter_fwd_xpx_pushs … 0 … H3) [ |*: // ] #g2 #H3 #H2 destruct
143   lapply (drops_isuni_fwd_drop2 … HLY) -HLY // #HLY
144   lapply (IH … HYK … H3) -IH -H3 -HYK [4: |*: /2 width=2 by lifts_lref/ ]
145   >plus_S1 /2 width=3 by frees_lref_pushs/ (**) (* full auto fails *)
146 | #f1 #K #l #Hf1 #_ #f #L #HLK #U #H2 #f2 #H3
147   lapply (coafter_isid_inv_dx … H3 … Hf1) -f1 #Hf2
148   >(lifts_inv_gref1 … H2) -U /2 width=1 by frees_gref/
149 | #f1V #f1T #f1 #p #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
150   elim (sor_inv_isfin3 … H1f1) // #Hf1V #H
151   lapply (isfin_inv_tl … H) -H
152   elim (lifts_inv_bind1 … H2) -H2 #W #U #HVW #HTU #H destruct
153   elim (coafter_sor … H3 … H1f1) /2 width=5 by coafter_isfin2_fwd/ -H3 -H1f1 #f2V #f2T #Hf2V #H
154   elim (coafter_inv_tl1 … H) -H
155   /5 width=5 by frees_bind, drops_skip, ext2_pair/
156 | #f1V #f1T #f1 #I #K #V #T #_ #_ #H1f1 #IHV #IHT #H2f1 #f #L #H1 #Y #H2 #f2 #H3
157   elim (sor_inv_isfin3 … H1f1) //
158   elim (lifts_inv_flat1 … H2) -H2 #W #U #HVW #HTU #H destruct
159   elim (coafter_sor … H3 … H1f1)
160   /3 width=5 by coafter_isfin2_fwd, frees_flat/
161 ]
162 qed-.
163
164 lemma frees_lifts_SO:
165       ∀b,L,K. ⬇*[b,𝐔❴1❵] L ≘ K → ∀T,U. ⬆*[1] T ≘ U →
166       ∀f. K ⊢ 𝐅+⦃T⦄ ≘ f → L ⊢ 𝐅+⦃U⦄ ≘ ⫯f.
167 #b #L #K #HLK #T #U #HTU #f #Hf
168 @(frees_lifts b … Hf … HTU) //  (**) (* auto fails *)
169 qed.
170
171 (* Forward lemmas with generic slicing for local environments ***************)
172
173 lemma frees_fwd_coafter:
174       ∀b,f2,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f2 →
175       ∀f,K. ⬇*[b,f] L ≘ K → ∀T. ⬆*[f] T ≘ U →
176       ∀f1. K ⊢ 𝐅+⦃T⦄ ≘ f1 → f ~⊚ f1 ≘ f2.
177 /4 width=11 by frees_lifts, frees_mono, coafter_eq_repl_back0/ qed-.
178
179 (* Inversion lemmas with generic slicing for local environments *************)
180
181 lemma frees_inv_lifts_ex:
182       ∀b,f2,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f2 →
183       ∀f,K. ⬇*[b,f] L ≘ K → ∀T. ⬆*[f] T ≘ U →
184       ∃∃f1. f ~⊚ f1 ≘ f2 & K ⊢ 𝐅+⦃T⦄ ≘ f1.
185 #b #f2 #L #U #Hf2 #f #K #HLK #T elim (frees_total K T)
186 /3 width=9 by frees_fwd_coafter, ex2_intro/
187 qed-.
188
189 lemma frees_inv_lifts_SO:
190       ∀b,f,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f →
191       ∀K. ⬇*[b,𝐔❴1❵] L ≘ K → ∀T. ⬆*[1] T ≘ U →
192       K ⊢ 𝐅+⦃T⦄ ≘ ⫱f.
193 #b #f #L #U #H #K #HLK #T #HTU elim(frees_inv_lifts_ex … H … HLK … HTU) -b -L -U
194 #f1 #Hf #Hf1 elim (coafter_inv_nxx … Hf) -Hf
195 /3 width=5 by frees_eq_repl_back, coafter_isid_inv_sn/
196 qed-.
197
198 lemma frees_inv_lifts:
199       ∀b,f2,L,U. L ⊢ 𝐅+⦃U⦄ ≘ f2 →
200       ∀f,K. ⬇*[b,f] L ≘ K → ∀T. ⬆*[f] T ≘ U →
201       ∀f1. f ~⊚ f1 ≘ f2 → K ⊢ 𝐅+⦃T⦄ ≘ f1.
202 #b #f2 #L #U #H #f #K #HLK #T #HTU #f1 #Hf2 elim (frees_inv_lifts_ex … H … HLK … HTU) -b -L -U
203 /3 width=7 by frees_eq_repl_back, coafter_inj/
204 qed-.
205
206 (* Note: this is used by rex_conf and might be modified *)
207 lemma frees_inv_drops_next:
208       ∀f1,L1,T1. L1 ⊢ 𝐅+⦃T1⦄ ≘ f1 →
209       ∀I2,L2,V2,n. ⬇*[n] L1 ≘ L2.ⓑ{I2}V2 →
210       ∀g1. ↑g1 = ⫱*[n] f1 →
211       ∃∃g2. L2 ⊢ 𝐅+⦃V2⦄ ≘ g2 & g2 ⊆ g1.
212 #f1 #L1 #T1 #H elim H -f1 -L1 -T1
213 [ #f1 #L1 #s #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -s
214   lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
215   elim (isid_inv_next … Hf1) -Hf1 //
216 | #f1 #i #_ #I2 #L2 #V2 #n #H
217   elim (drops_inv_atom1 … H) -H #H destruct
218 | #f1 #I1 #L1 #V1 #Hf1 #IH #I2 #L2 #V2 *
219   [ -IH #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 //
220     #H destruct #g1 #Hgf1 >(injective_next … Hgf1) -g1
221     /2 width=3 by ex2_intro/
222   | -Hf1 #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
223     #HL12 #g1 <tls_xn <tl_next_rew #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
224     /2 width=3 by ex2_intro/
225   ]
226 | #f1 #I1 #L1 #Hf1 #I2 #L2 #V2 *
227   [ #HL12 lapply (drops_fwd_isid … HL12 ?) -HL12 // #H destruct
228   | #n #_ #g1 #Hgf1 elim (isid_inv_next … Hgf1) -Hgf1 <tls_xn /2 width=1 by isid_tls/
229   ]
230 | #f1 #I1 #L1 #i #_ #IH #I2 #L2 #V2 *
231   [ -IH #_ #g1 #Hgf1 elim (discr_next_push … Hgf1)
232   | #n #HL12 lapply (drops_inv_drop1 … HL12) -HL12
233     #HL12 #g1 <tls_xn #Hgf1 elim (IH … HL12 … Hgf1) -IH -HL12 -Hgf1
234     /2 width=3 by ex2_intro/
235   ]
236 | #f1 #L1 #l #Hf1 #I2 #L2 #V2 #n #_ #g1 #H1 -I2 -L1 -l
237   lapply (isid_tls n … Hf1) -Hf1 <H1 -f1 #Hf1
238   elim (isid_inv_next … Hf1) -Hf1 //
239 | #fV1 #fT1 #f1 #p #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
240   lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
241   elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
242   #gV1 #gT1 #Hg1
243   [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
244     /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
245   | -IHV1 #_ >tls_xn #H2 elim (IHT1 … H2) -IHT1 -H2
246     /3 width=6 by drops_drop, sor_inv_sle_dx_trans, ex2_intro/
247   ]
248 | #fV1 #fT1 #f1 #I1 #L1 #V1 #T1 #_ #_ #Hf1 #IHV1 #IHT1 #I2 #L2 #V2 #n #HL12 #g1 #Hgf1
249   lapply (sor_tls … Hf1 n) -Hf1 <Hgf1 -Hgf1 #Hf1
250   elim (sor_xxn_tl … Hf1) [1,2: * |*: // ] -Hf1
251   #gV1 #gT1 #Hg1
252   [ -IHT1 #H1 #_ elim (IHV1 … HL12 … H1) -IHV1 -HL12 -H1
253     /3 width=6 by sor_inv_sle_sn_trans, ex2_intro/
254   | -IHV1 #_ #H2 elim (IHT1 … HL12 … H2) -IHT1 -HL12 -H2
255     /3 width=6 by sor_inv_sle_dx_trans, ex2_intro/
256   ]
257 ]
258 qed-.