]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/etc_2A1/cofrees/cofrees_lift.etc
update in binaries for λδ
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / etc_2A1 / cofrees / cofrees_lift.etc
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/cpys_lift.ma".
16 include "basic_2/substitution/cofrees.ma".
17
18 (* CONTEXT-SENSITIVE EXCLUSION FROM FREE VARIABLES **************************)
19
20 (* Advanced inversion lemmas ************************************************)
21
22 lemma cofrees_inv_lref_be: ∀L,d,i,j. L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄ → d ≤ yinj j → j < i →
23                            ∀I,K,W. ⇩[j]L ≡ K.ⓑ{I}W → K ⊢ i-j-1 ~ϵ 𝐅*[yinj 0]⦃W⦄.
24 #L #d #i #j #Hj #Hdj #Hji #I #K #W1 #HLK #W2 #HW12 elim (lift_total W2 0 (j+1))
25 #X2 #HWX2 elim (Hj X2) /2 width=7 by cpys_subst_Y2/ -I -L -K -W1 -d
26 #Z2 #HZX2 elim (lift_div_le … HWX2 (i-j-1) 1 Z2) -HWX2 /2 width=2 by ex_intro/
27 >minus_plus <plus_minus_m_m //
28 qed-.
29
30 lemma cofrees_inv_be: ∀L,U,d,i. L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ∀j. (∀T. ⇧[j, 1] T ≡ U → ⊥) →
31                       ∀I,K,W. ⇩[j]L ≡ K.ⓑ{I}W → d ≤ yinj j → j < i → K ⊢ i-j-1 ~ϵ 𝐅*[yinj 0]⦃W⦄.
32 #L #U @(f2_ind … rfw … L U) -L -U
33 #n #IH #L * *
34 [ -IH #k #_ #d #i #_ #j #H elim (H (⋆k)) -H //
35 | -IH #j #_ #d #i #Hi0 #j0 #H <(nlift_inv_lref_be_SO … H) -j0
36   /2 width=9 by cofrees_inv_lref_be/
37 | -IH #p #_ #d #i #_ #j #H elim (H (§p)) -H //
38 | #a #J #W #U #Hn #d #i #H1 #j #H2 #I #K #V #HLK #Hdj #Hji destruct
39   elim (cofrees_inv_bind … H1) -H1 #HW #HU
40   elim (nlift_inv_bind … H2) -H2 [ -HU /3 width=9 by/ ]
41   -HW #HnU lapply (IH … HU … HnU I K V ? ? ?)
42   /2 width=1 by ldrop_drop, yle_succ, lt_minus_to_plus/ -a -I -J -L -W -U -d
43   >minus_plus_plus_l //
44 | #J #W #U #Hn #d #i #H1 #j #H2 #I #K #V #HLK #Hdj #Hji destruct
45   elim (cofrees_inv_flat … H1) -H1 #HW #HU
46   elim (nlift_inv_flat … H2) -H2 [ /3 width=9 by/ ]
47   #HnU @(IH … HU … HnU … HLK) // (**) (* full auto fails *)
48 ]
49 qed-.
50
51 (* Advanced properties ******************************************************)
52
53 lemma cofrees_lref_skip: ∀L,d,i,j. j < i → yinj j < d → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄.
54 #L #d #i #j #Hji #Hjd #X #H elim (cpys_inv_lref1_Y2 … H) -H
55 [ #H destruct /3 width=2 by lift_lref_lt, ex_intro/
56 | * #I #K #W1 #W2 #Hdj elim (ylt_yle_false … Hdj) -i -I -L -K -W1 -W2 -X //
57 ]
58 qed.
59
60 lemma cofrees_lref_lt: ∀L,d,i,j. i < j → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄.
61 #L #d #i #j #Hij #X #H elim (cpys_inv_lref1_Y2 … H) -H
62 [ #H destruct /3 width=2 by lift_lref_ge_minus, ex_intro/
63 | * #I #K #V1 #V2 #_ #_ #_ #H -I -L -K -V1 -d
64   elim (lift_split … H i j) /2 width=2 by lt_to_le, ex_intro/
65 ]
66 qed.
67
68 lemma cofrees_lref_gt: ∀I,L,K,W,d,i,j. j < i → ⇩[j] L ≡ K.ⓑ{I}W → 
69                        K ⊢ (i-j-1) ~ϵ 𝐅*[O]⦃W⦄ → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄.
70 #I #L #K #W1 #d #i #j #Hji #HLK #HW1 #X #H elim (cpys_inv_lref1_Y2 … H) -H
71 [ #H destruct /3 width=2 by lift_lref_lt, ex_intro/
72 | * #I0 #K0 #W0 #W2 #Hdj #HLK0 #HW12 #HW2 lapply (ldrop_mono … HLK0 … HLK) -L
73   #H destruct elim (HW1 … HW12) -I -K -W1 -d
74   #V2 #HVW2 elim (lift_trans_le … HVW2 … HW2) -W2 //
75   >minus_plus <plus_minus_m_m /2 width=2 by ex_intro/
76 ]
77 qed.
78
79 lemma cofrees_lref_free: ∀L,d,i,j. |L| ≤ j → j < i → L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄.
80 #L #d #i #j #Hj #Hji #X #H elim (cpys_inv_lref1_Y2 … H) -H
81 [ #H destruct /3 width=2 by lift_lref_lt, ex_intro/
82 | * #I #K #W1 #W2 #_ #HLK lapply (ldrop_fwd_length_lt2 … HLK) -I
83   #H elim (lt_refl_false j) -d -i -K -W1 -W2 -X /2 width=3 by lt_to_le_to_lt/
84 ]
85 qed.
86
87 (* Advanced negated inversion lemmas ****************************************)
88
89 lemma frees_inv_lref_gt: ∀L,d,i,j. j < i → (L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄ → ⊥) →
90                          ∃∃I,K,W. ⇩[j] L ≡ K.ⓑ{I}W & (K ⊢ (i-j-1) ~ϵ 𝐅*[0]⦃W⦄ → ⊥) & d ≤ yinj j.
91 #L #d #i #j #Hji #H elim (ylt_split j d) #Hjd
92 [ elim H -H /2 width=6 by cofrees_lref_skip/ 
93 | elim (lt_or_ge j (|L|)) #Hj
94   [ elim (ldrop_O1_lt … Hj) -Hj /4 width=10 by cofrees_lref_gt, ex3_3_intro/
95   | elim H -H /2 width=6 by cofrees_lref_free/
96   ]
97 ]
98 qed-.
99
100 lemma frees_inv_lref_free: ∀L,d,i,j. (L ⊢ i ~ϵ 𝐅*[d]⦃#j⦄  → ⊥) → |L| ≤ j → j = i.
101 #L #d #i #j #H #Hj elim (lt_or_eq_or_gt i j) //
102 #Hij elim H -H /2 width=6 by cofrees_lref_lt, cofrees_lref_free/
103 qed-.
104
105 lemma frees_inv_gen: ∀L,U,d,i. (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) →
106                      ∃∃U0.  ⦃⋆, L⦄ ⊢ U ▶*[d, ∞] U0 & (∀T. ⇧[i, 1] T ≡ U0 → ⊥).
107 #L #U @(f2_ind … rfw … L U) -L -U
108 #n #IH #L * *
109 [ -IH #k #_ #d #i #H elim H -H //
110 | #j #Hn #d #i #H elim (lt_or_eq_or_gt i j)
111   [ -n #Hij elim H -H /2 width=5 by cofrees_lref_lt/
112   | -H -n #H destruct /3 width=7 by lift_inv_lref2_be, ex2_intro/
113   | #Hji elim (frees_inv_lref_gt … H) // -H
114     #I #K #W1 #HLK #H #Hdj elim (IH … H) /2 width=3 by ldrop_fwd_rfw/ -H -n
115     #W2 #HW12 #HnW2 elim (lift_total W2 0 (j+1))
116     #U2 #HWU2 @(ex2_intro … U2) /2 width=7 by cpys_subst_Y2/ -I -L -K -W1 -d
117     #T2 #HTU2 elim (lift_div_le … HWU2 (i-j-1) 1 T2) /2 width=2 by/ -W2
118     >minus_plus <plus_minus_m_m //
119   ]
120 | -IH #p #_ #d #i #H elim H -H //
121 | #a #I #W #U #Hn #d #i #H elim (frees_inv_bind … H) -H
122   #H elim (IH … H) // -H -n
123   /4 width=9 by cpys_bind, nlift_bind_dx, nlift_bind_sn, ex2_intro/
124 | #I #W #U #Hn #d #i #H elim (frees_inv_flat … H) -H
125   #H elim (IH … H) // -H -n
126   /4 width=9 by cpys_flat, nlift_flat_dx, nlift_flat_sn, ex2_intro/
127 ]
128 qed-.
129
130 lemma frees_ind: ∀L,d,i. ∀R:predicate term.
131                  (∀U1. (∀T1. ⇧[i, 1] T1 ≡ U1 → ⊥) → R U1) → 
132                  (∀U1,U2. ⦃⋆, L⦄ ⊢ U1 ▶[d, ∞] U2 → (L ⊢ i ~ϵ 𝐅*[d]⦃U2⦄ → ⊥) → R U2 → R U1) →
133                  ∀U. (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → R U.
134 #L #d #i #R #IH1 #IH2 #U1 #H elim (frees_inv_gen … H) -H
135 #U2 #H #HnU2 @(cpys_ind_dx … H) -U1 /4 width=8 by cofrees_inv_gen/
136 qed-.
137
138 (* Advanced negated properties **********************************************)
139
140 lemma frees_be: ∀I,L,K,W,j. ⇩[j]L ≡ K.ⓑ{I}W →
141                 ∀i. j < i → (K ⊢ i-j-1 ~ϵ 𝐅*[yinj 0]⦃W⦄ → ⊥) →
142                 ∀U. (∀T. ⇧[j, 1] T ≡ U → ⊥) →
143                 ∀d. d ≤ yinj j → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥).
144 /4 width=11 by cofrees_inv_be/ qed-.
145
146 (* Relocation properties ****************************************************)
147
148 lemma cofrees_lift_be: ∀d0,e0,i. d0 ≤ i → i ≤ d0 + e0 →
149                        ∀L,K,s. ⇩[s, d0, e0+1] L ≡ K → ∀T,U. ⇧[d0, e0+1] T ≡ U →
150                        ∀d. L ⊢ i ~ϵ 𝐅*[d]⦃U⦄.
151 #d0 #e0 #i #Hd0i #Hide0 #L #K #s #HLK #T1 #U1 #HTU1 #d #U2 #HU12
152 elim (yle_split d0 d) #H1
153 [ elim (yle_split d (d0+e0+1)) #H2
154   [ letin cpys_inv ≝ cpys_inv_lift1_ge_up
155   | letin cpys_inv ≝ cpys_inv_lift1_ge
156   ]
157 | letin cpys_inv ≝ cpys_inv_lift1_be
158 ]
159 elim (cpys_inv … HU12 … HLK … HTU1) // #T2 #_ #HTU2 -s -L -K -U1 -T1 -d
160 elim (lift_split … HTU2 i e0) /2 width=2 by ex_intro/
161 qed.
162
163 lemma cofrees_lift_ge: ∀d0,e0,i. d0 + e0 ≤ i →
164                        ∀L,K,s. ⇩[s, d0, e0] L ≡ K → ∀T,U. ⇧[d0, e0] T ≡ U →
165                        ∀d. K ⊢ i-e0 ~ϵ 𝐅*[d-yinj e0]⦃T⦄ → L ⊢ i ~ϵ 𝐅*[d]⦃U⦄.
166 #d0 #e0 #i #Hde0i #L #K #s #HLK #T1 #U1 #HTU1 #d #HT1 #U2 #HU12
167 elim (le_inv_plus_l … Hde0i) -Hde0i #Hd0ie0 #He0i
168 elim (yle_split d0 d) #H1
169 [ elim (yle_split d (d0+e0)) #H2
170   [ elim (cpys_inv_lift1_ge_up … HU12 … HLK … HTU1) // >yplus_inj >yminus_Y_inj #T2 #HT12
171     lapply (cpys_weak … HT12 (d-yinj e0) (∞) ? ?) /2 width=1 by yle_plus2_to_minus_inj2/ -HT12
172   | elim (cpys_inv_lift1_ge … HU12 … HLK … HTU1) // #T2
173   ]
174 | elim (cpys_inv_lift1_be … HU12 … HLK … HTU1) // >yminus_Y_inj #T2 #HT12
175   lapply (cpys_weak … HT12 (d-yinj e0) (∞) ? ?) // -HT12
176 ]
177 -s -L #HT12 #HTU2
178 elim (HT1 … HT12) -T1 #V2 #HVT2
179 elim (lift_trans_le … HVT2 … HTU2 ?) // <plus_minus_m_m /2 width=2 by ex_intro/
180 qed.