]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cpt.ma
milestone update in basic_2, update in ground and static_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cpt.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/xoa/ex_4_3.ma".
16 include "ground/steps/rtc_ist_shift.ma".
17 include "ground/steps/rtc_ist_plus.ma".
18 include "ground/steps/rtc_ist_max.ma".
19 include "static_2/syntax/sh.ma".
20 include "basic_2/notation/relations/pty_6.ma".
21 include "basic_2/rt_transition/cpg.ma".
22
23 (* T-BOUND CONTEXT-SENSITIVE PARALLEL T-TRANSITION FOR TERMS ****************)
24
25 definition cpt (h) (G) (L) (n): relation2 term term ≝
26            λT1,T2. ∃∃c. 𝐓❪n,c❫ & ❪G,L❫ ⊢ T1 ⬈[sh_is_next h,eq …,c] T2.
27
28 interpretation
29   "t-bound context-sensitive parallel t-transition (term)"
30   'PTy h n G L T1 T2 = (cpt h G L n T1 T2).
31
32 (* Basic properties *********************************************************)
33
34 lemma cpt_ess (h) (G) (L):
35       ∀s. ❪G,L❫ ⊢ ⋆s ⬆[h,1] ⋆(⫯[h]s).
36 /3 width=3 by cpg_ess, ex2_intro/ qed.
37
38 lemma cpt_delta (h) (n) (G) (K):
39       ∀V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,n] V2 →
40       ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓓV1❫ ⊢ #0 ⬆[h,n] W2.
41 #h #n #G #K #V1 #V2 *
42 /3 width=5 by cpg_delta, ex2_intro/
43 qed.
44
45 lemma cpt_ell (h) (n) (G) (K):
46       ∀V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,n] V2 →
47       ∀W2. ⇧[1] V2 ≘ W2 → ❪G,K.ⓛV1❫ ⊢ #0 ⬆[h,↑n] W2.
48 #h #n #G #K #V1 #V2 *
49 /3 width=5 by cpg_ell, ex2_intro, ist_succ/
50 qed.
51
52 lemma cpt_lref (h) (n) (G) (K):
53       ∀T,i. ❪G,K❫ ⊢ #i ⬆[h,n] T → ∀U. ⇧[1] T ≘ U →
54       ∀I. ❪G,K.ⓘ[I]❫ ⊢ #↑i ⬆[h,n] U.
55 #h #n #G #K #T #i *
56 /3 width=5 by cpg_lref, ex2_intro/
57 qed.
58
59 lemma cpt_bind (h) (n) (G) (L):
60       ∀V1,V2. ❪G,L❫ ⊢ V1 ⬆[h,0] V2 → ∀I,T1,T2. ❪G,L.ⓑ[I]V1❫ ⊢ T1 ⬆[h,n] T2 →
61       ∀p. ❪G,L❫ ⊢ ⓑ[p,I]V1.T1 ⬆[h,n] ⓑ[p,I]V2.T2.
62 #h #n #G #L #V1 #V2 * #cV #HcV #HV12 #I #T1 #T2 *
63 /3 width=5 by cpg_bind, ist_max_O1, ex2_intro/
64 qed.
65
66 lemma cpt_appl (h) (n) (G) (L):
67       ∀V1,V2. ❪G,L❫ ⊢ V1 ⬆[h,0] V2 →
68       ∀T1,T2. ❪G,L❫ ⊢ T1 ⬆[h,n] T2 → ❪G,L❫ ⊢ ⓐV1.T1 ⬆[h,n] ⓐV2.T2.
69 #h #n #G #L #V1 #V2 * #cV #HcV #HV12 #T1 #T2 *
70 /3 width=5 by ist_max_O1, cpg_appl, ex2_intro/
71 qed.
72
73 lemma cpt_cast (h) (n) (G) (L):
74       ∀U1,U2. ❪G,L❫ ⊢ U1 ⬆[h,n] U2 →
75       ∀T1,T2. ❪G,L❫ ⊢ T1 ⬆[h,n] T2 → ❪G,L❫ ⊢ ⓝU1.T1 ⬆[h,n] ⓝU2.T2.
76 #h #n #G #L #U1 #U2 * #cU #HcU #HU12 #T1 #T2 *
77 /3 width=6 by cpg_cast, ex2_intro/
78 qed.
79
80 lemma cpt_ee (h) (n) (G) (L):
81       ∀U1,U2. ❪G,L❫ ⊢ U1 ⬆[h,n] U2 → ∀T. ❪G,L❫ ⊢ ⓝU1.T ⬆[h,↑n] U2.
82 #h #n #G #L #V1 #V2 *
83 /3 width=3 by cpg_ee, ist_succ, ex2_intro/
84 qed.
85
86 lemma cpt_refl (h) (G) (L): reflexive … (cpt h G L 0).
87 /3 width=3 by cpg_refl, ex2_intro/ qed.
88
89 (* Advanced properties ******************************************************)
90
91 lemma cpt_sort (h) (G) (L):
92       ∀n. n ≤ 1 → ∀s. ❪G,L❫ ⊢ ⋆s ⬆[h,n] ⋆((next h)^n s).
93 #h #G #L * //
94 #n #H #s <(le_n_O_to_eq n) /2 width=1 by le_S_S_to_le/
95 qed.
96
97 (* Basic inversion lemmas ***************************************************)
98
99 lemma cpt_inv_atom_sn (h) (n) (J) (G) (L):
100       ∀X2. ❪G,L❫ ⊢ ⓪[J] ⬆[h,n] X2 →
101       ∨∨ ∧∧ X2 = ⓪[J] & n = 0
102        | ∃∃s. X2 = ⋆(⫯[h]s) & J = Sort s & n =1
103        | ∃∃K,V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,n] V2 & ⇧[1] V2 ≘ X2 & L = K.ⓓV1 & J = LRef 0
104        | ∃∃m,K,V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,m] V2 & ⇧[1] V2 ≘ X2 & L = K.ⓛV1 & J = LRef 0 & n = ↑m
105        | ∃∃I,K,T,i. ❪G,K❫ ⊢ #i ⬆[h,n] T & ⇧[1] T ≘ X2 & L = K.ⓘ[I] & J = LRef (↑i).
106 #h #n #J #G #L #X2 * #c #Hc #H
107 elim (cpg_inv_atom1 … H) -H *
108 [ #H1 #H2 destruct /3 width=1 by or5_intro0, conj/
109 | #s1 #s2 #H1 #H2 #H3 #H4 destruct /3 width=3 by or5_intro1, ex3_intro/
110 | #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 #H3 destruct
111   /4 width=6 by or5_intro2, ex4_3_intro, ex2_intro/
112 | #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 #H3 destruct
113   elim (ist_inv_plus_SO_dx … H3) -H3 [| // ] #m #Hc #H destruct
114   /4 width=9 by or5_intro3, ex5_4_intro, ex2_intro/
115 | #I #K #V2 #i #HV2 #HVT2 #H1 #H2 destruct
116   /4 width=8 by or5_intro4, ex4_4_intro, ex2_intro/
117 ]
118 qed-.
119
120 lemma cpt_inv_sort_sn (h) (n) (G) (L) (s):
121       ∀X2. ❪G,L❫ ⊢ ⋆s ⬆[h,n] X2 →
122       ∧∧ X2 = ⋆(((next h)^n) s) & n ≤ 1.
123 #h #n #G #L #s #X2 * #c #Hc #H
124 elim (cpg_inv_sort1 … H) -H * #H1 #H2 destruct
125 [ /2 width=1 by conj/
126 | #H1 #H2 destruct /2 width=1 by conj/
127
128 qed-.
129
130 lemma cpt_inv_zero_sn (h) (n) (G) (L):
131       ∀X2. ❪G,L❫ ⊢ #0 ⬆[h,n] X2 →
132       ∨∨ ∧∧ X2 = #0 & n = 0
133        | ∃∃K,V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,n] V2 & ⇧[1] V2 ≘ X2 & L = K.ⓓV1
134        | ∃∃m,K,V1,V2. ❪G,K❫ ⊢ V1 ⬆[h,m] V2 & ⇧[1] V2 ≘ X2 & L = K.ⓛV1 & n = ↑m.
135 #h #n #G #L #X2 * #c #Hc #H elim (cpg_inv_zero1 … H) -H *
136 [ #H1 #H2 destruct /4 width=1 by ist_inv_00, or3_intro0, conj/
137 | #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 destruct
138   /4 width=8 by or3_intro1, ex3_3_intro, ex2_intro/
139 | #cV #K #V1 #V2 #HV12 #HVT2 #H1 #H2 destruct
140   elim (ist_inv_plus_SO_dx … H2) -H2 // #m #Hc #H destruct
141   /4 width=8 by or3_intro2, ex4_4_intro, ex2_intro/
142 ]
143 qed-.
144
145 lemma cpt_inv_zero_sn_unit (h) (n) (I) (K) (G):
146       ∀X2. ❪G,K.ⓤ[I]❫ ⊢ #0 ⬆[h,n] X2 → ∧∧ X2 = #0 & n = 0.
147 #h #n #I #G #K #X2 #H
148 elim (cpt_inv_zero_sn … H) -H *
149 [ #H1 #H2 destruct /2 width=1 by conj/
150 | #Y #X1 #X2 #_ #_ #H destruct
151 | #m #Y #X1 #X2 #_ #_ #H destruct
152 ]
153 qed.
154
155 lemma cpt_inv_lref_sn (h) (n) (G) (L) (i):
156       ∀X2. ❪G,L❫ ⊢ #↑i ⬆[h,n] X2 →
157       ∨∨ ∧∧ X2 = #(↑i) & n = 0
158        | ∃∃I,K,T. ❪G,K❫ ⊢ #i ⬆[h,n] T & ⇧[1] T ≘ X2 & L = K.ⓘ[I].
159 #h #n #G #L #i #X2 * #c #Hc #H elim (cpg_inv_lref1 … H) -H *
160 [ #H1 #H2 destruct /4 width=1 by ist_inv_00, or_introl, conj/
161 | #I #K #V2 #HV2 #HVT2 #H destruct
162  /4 width=6 by ex3_3_intro, ex2_intro, or_intror/
163 ]
164 qed-.
165
166 lemma cpt_inv_lref_sn_ctop (h) (n) (G) (i):
167       ∀X2. ❪G,⋆❫ ⊢ #i ⬆[h,n] X2 → ∧∧ X2 = #i & n = 0.
168 #h #n #G * [| #i ] #X2 #H
169 [ elim (cpt_inv_zero_sn … H) -H *
170   [ #H1 #H2 destruct /2 width=1 by conj/
171   | #Y #X1 #X2 #_ #_ #H destruct
172   | #m #Y #X1 #X2 #_ #_ #H destruct
173   ]
174 | elim (cpt_inv_lref_sn … H) -H *
175   [ #H1 #H2 destruct /2 width=1 by conj/
176   | #Z #Y #X0 #_ #_ #H destruct
177   ]
178 ]
179 qed.
180
181 lemma cpt_inv_gref_sn (h) (n) (G) (L) (l):
182       ∀X2. ❪G,L❫ ⊢ §l ⬆[h,n] X2 → ∧∧ X2 = §l & n = 0.
183 #h #n #G #L #l #X2 * #c #Hc #H elim (cpg_inv_gref1 … H) -H
184 #H1 #H2 destruct /2 width=1 by conj/
185 qed-.
186
187 lemma cpt_inv_bind_sn (h) (n) (p) (I) (G) (L) (V1) (T1):
188       ∀X2. ❪G,L❫ ⊢ ⓑ[p,I]V1.T1 ⬆[h,n] X2 →
189       ∃∃V2,T2. ❪G,L❫ ⊢ V1 ⬆[h,0] V2 & ❪G,L.ⓑ[I]V1❫ ⊢ T1 ⬆[h,n] T2
190              & X2 = ⓑ[p,I]V2.T2.
191 #h #n #p #I #G #L #V1 #T1 #X2 * #c #Hc #H
192 elim (cpg_inv_bind1 … H) -H *
193 [ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
194   elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
195   elim (ist_inv_shift … HcV) -HcV #HcV #H destruct
196   /3 width=5 by ex3_2_intro, ex2_intro/
197 | #cT #T2 #_ #_ #_ #_ #H destruct
198   elim (ist_inv_plus_10_dx … H)
199 ]
200 qed-.
201
202 lemma cpt_inv_appl_sn (h) (n) (G) (L) (V1) (T1):
203       ∀X2. ❪G,L❫ ⊢ ⓐV1.T1 ⬆[h,n] X2 →
204       ∃∃V2,T2. ❪G,L❫ ⊢ V1 ⬆[h,0] V2 & ❪G,L❫ ⊢ T1 ⬆[h,n] T2 & X2 = ⓐV2.T2.
205 #h #n #G #L #V1 #T1 #X2 * #c #Hc #H elim (cpg_inv_appl1 … H) -H *
206 [ #cV #cT #V2 #T2 #HV12 #HT12 #H1 #H2 destruct
207   elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
208   elim (ist_inv_shift … HcV) -HcV #HcV #H destruct
209   /3 width=5 by ex3_2_intro, ex2_intro/
210 | #cV #cW #cU #p #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #_ #H destruct
211   elim (ist_inv_plus_10_dx … H)
212 | #cV #cW #cU #p #V #V2 #W1 #W2 #U1 #U2 #_ #_ #_ #_ #_ #_ #H destruct
213   elim (ist_inv_plus_10_dx … H)
214 ]
215 qed-.
216
217 lemma cpt_inv_cast_sn (h) (n) (G) (L) (V1) (T1):
218       ∀X2. ❪G,L❫ ⊢ ⓝV1.T1 ⬆[h,n] X2 →
219       ∨∨ ∃∃V2,T2. ❪G,L❫ ⊢ V1 ⬆[h,n] V2 & ❪G,L❫ ⊢ T1 ⬆[h,n] T2 & X2 = ⓝV2.T2
220        | ∃∃m. ❪G,L❫ ⊢ V1 ⬆[h,m] X2 & n = ↑m.
221 #h #n #G #L #V1 #T1 #X2 * #c #Hc #H elim (cpg_inv_cast1 … H) -H *
222 [ #cV #cT #V2 #T2 #HV12 #HT12 #HcVT #H1 #H2 destruct
223   elim (ist_inv_max … H2) -H2 #nV #nT #HcV #HcT #H destruct
224   <idempotent_max
225   /4 width=5 by or_introl, ex3_2_intro, ex2_intro/
226 | #cT #_ #H destruct
227   elim (ist_inv_plus_10_dx … H)
228 | #cV #H12 #H destruct
229   elim (ist_inv_plus_SO_dx … H) -H [| // ] #m #Hm #H destruct
230   /4 width=3 by ex2_intro, or_intror/
231 ]
232 qed-.