]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/multiple/frees_lift.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / multiple / frees_lift.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_2A/substitution/drop_drop.ma".
16 include "basic_2A/multiple/frees.ma".
17
18 (* CONTEXT-SENSITIVE FREE VARIABLES *****************************************)
19
20 (* Advanced properties ******************************************************)
21
22 lemma frees_dec: ∀L,U,l,i. Decidable (frees l L U i).
23 #L #U @(f2_ind … rfw … L U) -L -U
24 #x #IH #L * *
25 [ -IH /3 width=5 by frees_inv_sort, or_intror/
26 | #j #Hx #l #i elim (lt_or_eq_or_gt i j) #Hji
27   [ -x @or_intror #H elim (lt_refl_false i)
28     lapply (frees_inv_lref_ge … H ?) -L -l /2 width=1 by lt_to_le/
29   | -x /2 width=1 by or_introl/
30   | elim (ylt_split j l) #Hli
31     [ -x @or_intror #H elim (lt_refl_false i)
32       lapply (frees_inv_lref_skip … H ?) -L //
33     | elim (lt_or_ge j (|L|)) #Hj
34       [ elim (drop_O1_lt (Ⓕ) L j) // -Hj #I #K #W #HLK destruct
35         elim (IH K W … 0 (i-j-1)) -IH [1,3: /3 width=5 by frees_lref_be, drop_fwd_rfw, or_introl/ ] #HnW
36         @or_intror #H elim (frees_inv_lref_lt … H) // #Z #Y #X #_ #HLY -l
37         lapply (drop_mono … HLY … HLK) -L #H destruct /2 width=1 by/  
38       | -x @or_intror #H elim (lt_refl_false i)
39         lapply (frees_inv_lref_free … H ?) -l //
40       ]
41     ]
42   ]
43 | -IH /3 width=5 by frees_inv_gref, or_intror/
44 | #a #I #W #U #Hx #l #i destruct
45   elim (IH L W … l i) [1,3: /3 width=1 by frees_bind_sn, or_introl/ ] #HnW
46   elim (IH (L.ⓑ{I}W) U … (⫯l) (i+1)) -IH [1,3: /3 width=1 by frees_bind_dx, or_introl/ ] #HnU
47   @or_intror #H elim (frees_inv_bind … H) -H /2 width=1 by/
48 | #I #W #U #Hx #l #i destruct
49   elim (IH L W … l i) [1,3: /3 width=1 by frees_flat_sn, or_introl/ ] #HnW
50   elim (IH L U … l i) -IH [1,3: /3 width=1 by frees_flat_dx, or_introl/ ] #HnU
51   @or_intror #H elim (frees_inv_flat … H) -H /2 width=1 by/
52 ]
53 qed-.
54
55 lemma frees_S: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[yinj l]⦃U⦄ → ∀I,K,W. ⬇[l] L ≡ K.ⓑ{I}W →
56                (K ⊢ i-l-1 ϵ 𝐅*[0]⦃W⦄ → ⊥) → L ⊢ i ϵ 𝐅*[⫯l]⦃U⦄.
57 #L #U #l #i #H elim (frees_inv … H) -H /3 width=2 by frees_eq/
58 * #I #K #W #j #Hlj #Hji #HnU #HLK #HW #I0 #K0 #W0 #HLK0 #HnW0
59 lapply (yle_inv_inj … Hlj) -Hlj #Hlj
60 elim (le_to_or_lt_eq … Hlj) -Hlj
61 [ -I0 -K0 -W0 /3 width=9 by frees_be, yle_inj/
62 | -Hji -HnU #H destruct
63   lapply (drop_mono … HLK0 … HLK) #H destruct -I
64   elim HnW0 -L -U -HnW0 //
65 ]
66 qed.
67
68 (* Note: lemma 1250 *)
69 lemma frees_bind_dx_O: ∀a,I,L,W,U,i. L.ⓑ{I}W ⊢ i+1 ϵ 𝐅*[0]⦃U⦄ →
70                        L ⊢ i ϵ 𝐅*[0]⦃ⓑ{a,I}W.U⦄.
71 #a #I #L #W #U #i #HU elim (frees_dec L W 0 i)
72 /4 width=5 by frees_S, frees_bind_dx, frees_bind_sn/
73 qed.
74
75 (* Properties on relocation *************************************************)
76
77 lemma frees_lift_ge: ∀K,T,l,i. K ⊢ i ϵ𝐅*[l]⦃T⦄ →
78                      ∀L,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
79                      ∀U. ⬆[l0, m0] T ≡ U → l0 ≤ i →
80                      L ⊢ i+m0 ϵ 𝐅*[l]⦃U⦄.
81 #K #T #l #i #H elim H -K -T -l -i
82 [ #K #T #l #i #HnT #L #s #l0 #m0 #_ #U #HTU #Hl0i -K -s
83   @frees_eq #X #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
84 | #I #K #K0 #T #V #l #i #j #Hlj #Hji #HnT #HK0 #HV #IHV #L #s #l0 #m0 #HLK #U #HTU #Hl0i
85   elim (lt_or_ge j l0) #H1
86   [ elim (drop_trans_lt … HLK … HK0) // -K #L0 #W #HL0 #HLK0 #HVW
87     @(frees_be … HL0) -HL0 -HV
88     /3 width=3 by lt_plus_to_minus_r, lt_to_le_to_lt/
89     [ #X #HXU >(plus_minus_m_m l0 1) in HTU; /2 width=2 by ltn_to_ltO/ #HTU
90       elim (lift_div_le … HXU … HTU ?) -U /2 width=2 by monotonic_pred/
91     | >minus_plus <plus_minus // <minus_plus
92       /3 width=5 by monotonic_le_minus_l2/
93     ]
94   | lapply (drop_trans_ge … HLK … HK0 ?) // -K #HLK0
95     lapply (drop_inv_gen … HLK0) >commutative_plus -HLK0 #HLK0
96     @(frees_be … HLK0) -HLK0 -IHV
97     /2 width=1 by yle_plus_dx1_trans, lt_minus_to_plus/
98     #X #HXU elim (lift_div_le … HTU … HXU) -U /2 width=2 by/
99   ]
100 ]
101 qed.
102
103 (* Inversion lemmas on relocation *******************************************)
104
105 lemma frees_inv_lift_be: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
106                          ∀K,s,l0,m0. ⬇[s, l0, m0+1] L ≡ K →
107                          ∀T. ⬆[l0, m0+1] T ≡ U → l0 ≤ i → i ≤ l0 + m0 → ⊥.
108 #L #U #l #i #H elim H -L -U -l -i
109 [ #L #U #l #i #HnU #K #s #l0 #m0 #_ #T #HTU #Hl0i #Hilm0
110   elim (lift_split … HTU i m0) -HTU /2 width=2 by/
111 | #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hl0i #Hilm0
112   elim (lt_or_ge j l0) #H1
113   [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
114     @(IHW … HKL0 … HVW)
115     [ /2 width=1 by monotonic_le_minus_l2/
116     | >minus_plus >minus_plus >plus_minus /2 width=1 by monotonic_le_minus_l/
117     ]
118   | elim (lift_split … HTU j m0) -HTU /3 width=3 by lt_to_le_to_lt, lt_to_le/
119   ]
120 ]
121 qed-.
122
123 lemma frees_inv_lift_ge: ∀L,U,l,i. L ⊢ i ϵ 𝐅*[l]⦃U⦄ →
124                          ∀K,s,l0,m0. ⬇[s, l0, m0] L ≡ K →
125                          ∀T. ⬆[l0, m0] T ≡ U → l0 + m0 ≤ i →
126                          K ⊢ i-m0 ϵ𝐅*[l-yinj m0]⦃T⦄.
127 #L #U #l #i #H elim H -L -U -l -i
128 [ #L #U #l #i #HnU #K #s #l0 #m0 #HLK #T #HTU #Hlm0i -L -s
129   elim (le_inv_plus_l … Hlm0i) -Hlm0i #Hl0im0 #Hm0i @frees_eq #X #HXT -K
130   elim (lift_trans_le … HXT … HTU) -T // <plus_minus_m_m /2 width=2 by/
131 | #I #L #K0 #U #W #l #i #j #Hli #Hij #HnU #HLK0 #_ #IHW #K #s #l0 #m0 #HLK #T #HTU #Hlm0i
132   elim (lt_or_ge j l0) #H1
133   [ elim (drop_conf_lt … HLK … HLK0) -L // #L0 #V #H #HKL0 #HVW
134     elim (le_inv_plus_l … Hlm0i) #H0 #Hm0i
135     @(frees_be … H) -H
136     [ /3 width=1 by yle_plus_dx1_trans, monotonic_yle_minus_dx/
137     | /2 width=3 by lt_to_le_to_lt/
138     | #X #HXT elim (lift_trans_ge … HXT … HTU) -T /2 width=2 by/
139     | lapply (IHW … HKL0 … HVW ?) // -I -K -K0 -L0 -V -W -T -U -s
140       >minus_plus >minus_plus >plus_minus /2 width=1 by monotonic_le_minus_l/
141     ]
142   | elim (lt_or_ge j (l0+m0)) [ >commutative_plus |] #H2
143     [ -L -I -W lapply (lt_plus_to_minus ???? H2) // -H2 #H2
144       elim (lift_split … HTU j (m0-1)) -HTU //
145       [ >minus_minus_associative /2 width=2 by ltn_to_ltO/ <minus_n_n
146         #X #_ #H elim (HnU … H)
147       | >commutative_plus /3 width=1 by le_minus_to_plus, monotonic_pred/
148       ]
149     | lapply (drop_conf_ge … HLK … HLK0 ?) // -L #HK0
150       elim (le_inv_plus_l … H2) -H2 #H2 #Hm0j
151       @(frees_be … HK0)
152       [ /2 width=1 by monotonic_yle_minus_dx/
153       | /2 width=1 by monotonic_lt_minus_l/
154       | #X #HXT elim (lift_trans_le … HXT … HTU) -T // <plus_minus_m_m /2 width=2 by/
155       | >arith_b1 /2 width=5 by/
156       ]
157     ]
158   ]
159 ]
160 qed-.