]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_tpss.ma
We are decapitalizing the contributions' names ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / unfold / ltpss_tpss.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_2/unfold/tpss_ltps.ma".
16 include "Basic_2/unfold/ltpss.ma".
17
18 (* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************)
19
20 (* Properties concerning partial unfold on terms ****************************)
21
22 lemma ltpss_tpss_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 →
23                            ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶* U2 →
24                            d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2.
25 #L1 #L0 #d1 #e1 #H @(ltpss_ind … H) -L0 //
26 #L #L0 #_ #HL0 #IHL #T2 #U2 #d2 #e2 #HTU2 #Hde1d2
27 lapply (ltps_tpss_trans_ge … HL0 HTU2) -HL0 -HTU2 /2 width=1/
28 qed.
29
30 lemma ltpss_tps_trans_ge: ∀L1,L0,d1,e1. L1 [d1, e1] ▶* L0 →
31                           ∀T2,U2,d2,e2. L0 ⊢ T2 [d2, e2] ▶ U2 →
32                           d1 + e1 ≤ d2 → L1 ⊢ T2 [d2, e2] ▶* U2.
33 #L1 #L0 #d1 #e1 #HL10 #T2 #U2 #d2 #e2 #HTU2 #Hde1d2
34 @(ltpss_tpss_trans_ge … HL10 … Hde1d2) /2 width=1/ (**) (* /3 width=6/ is too slow *)
35 qed.
36
37 lemma ltpss_tpss_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 →
38                            ∀T2,U2. L1 ⊢ T2 [d, e] ▶* U2 → L0 ⊢ T2 [d, e] ▶* U2.
39 #L0 #L1 #d #e #H @(ltpss_ind … H) -L1
40 [ /2 width=1/
41 | #L #L1 #_ #HL1 #IHL #T2 #U2 #HTU2
42   lapply (ltps_tpss_trans_eq … HL1 HTU2) -HL1 -HTU2 /2 width=1/
43 ]
44 qed.
45
46 lemma ltpss_tps_trans_eq: ∀L0,L1,d,e. L0 [d, e] ▶* L1 →
47                           ∀T2,U2. L1 ⊢ T2 [d, e] ▶ U2 → L0 ⊢ T2 [d, e] ▶* U2.
48 /3 width=3/ qed.
49
50 lemma ltpss_tpss_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → 
51                           L0 ⊢ T2 [d2, e2] ▶* U2 → L0 [d1, e1] ▶* L1 →
52                           L1 ⊢ T2 [d2, e2] ▶* U2.
53 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #H @(ltpss_ind … H) -L1
54 [ //
55 | -HTU2 #L #L1 #_ #HL1 #IHL
56   lapply (ltps_tpss_conf_ge … HL1 IHL) -HL1 -IHL //
57 ]
58 qed.
59
60 lemma ltpss_tps_conf_ge: ∀L0,L1,T2,U2,d1,e1,d2,e2. d1 + e1 ≤ d2 → 
61                          L0 ⊢ T2 [d2, e2] ▶ U2 → L0 [d1, e1] ▶* L1 →
62                          L1 ⊢ T2 [d2, e2] ▶* U2.
63 #L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde1d2 #HTU2 #HL01
64 @(ltpss_tpss_conf_ge … Hde1d2 … HL01) /2 width=1/ (**) (* /3 width=6/ is too slow *)
65 qed.
66
67 lemma ltpss_tpss_conf_eq: ∀L0,L1,T2,U2,d,e.
68                           L0 ⊢ T2 [d, e] ▶* U2 → L0 [d, e] ▶* L1 →
69                           ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T.
70 #L0 #L1 #T2 #U2 #d #e #HTU2 #H @(ltpss_ind … H) -L1
71 [ /2 width=3/
72 | -HTU2 #L #L1 #_ #HL1 * #W2 #HTW2 #HUW2
73   elim (ltps_tpss_conf … HL1 HTW2) -HTW2 #T #HT2 #HW2T
74   elim (ltps_tpss_conf … HL1 HUW2) -HL1 -HUW2 #U #HU2 #HW2U
75   elim (tpss_conf_eq … HW2T … HW2U) -HW2T -HW2U #V #HTV #HUV
76   lapply (tpss_trans_eq … HT2 … HTV) -T
77   lapply (tpss_trans_eq … HU2 … HUV) -U /2 width=3/
78 ]
79 qed.
80
81 lemma ltpss_tps_conf_eq: ∀L0,L1,T2,U2,d,e.
82                          L0 ⊢ T2 [d, e] ▶ U2 → L0 [d, e] ▶* L1 →
83                          ∃∃T. L1 ⊢ T2 [d, e] ▶* T & L1 ⊢ U2 [d, e] ▶* T.
84 /3 width=3/ qed.
85
86 lemma ltpss_tpss_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶* T2 →
87                        ∀L2,ds,es. L1 [ds, es] ▶* L2 → 
88                        ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T.
89 #L1 #T1 #T2 #d #e #HT12 #L2 #ds #es #H @(ltpss_ind … H) -L2
90 [ /3 width=3/
91 | #L #L2 #HL1 #HL2 * #T #HT1 #HT2
92   elim (ltps_tpss_conf … HL2 HT1) -HT1 #T0 #HT10 #HT0
93   lapply (ltps_tpss_trans_eq … HL2 … HT0) -HL2 -HT0 #HT0
94   lapply (ltpss_tpss_trans_eq … HL1 … HT0) -HL1 -HT0 #HT0
95   lapply (tpss_trans_eq … HT2 … HT0) -T /2 width=3/
96 ]
97 qed.
98
99 lemma ltpss_tps_conf: ∀L1,T1,T2,d,e. L1 ⊢ T1 [d, e] ▶ T2 →
100                       ∀L2,ds,es. L1 [ds, es] ▶* L2 → 
101                       ∃∃T. L2 ⊢ T1 [d, e] ▶* T & L1 ⊢ T2 [ds, es] ▶* T.
102 /3 width=1/ qed.
103
104 (* Advanced properties ******************************************************)
105
106 lemma ltpss_tps2: ∀L1,L2,I,e.
107                   L1 [0, e] ▶* L2 → ∀V1,V2. L2 ⊢ V1 [0, e] ▶ V2 →
108                   L1. ⓑ{I} V1 [0, e + 1] ▶* L2. ⓑ{I} V2.
109 #L1 #L2 #I #e #H @(ltpss_ind … H) -L2
110 [ /3 width=1/
111 | #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12
112   elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2
113   lapply (IHL … HV1) -IHL -HV1 #HL1
114   @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
115 ]
116 qed.
117
118 lemma ltpss_tps2_lt: ∀L1,L2,I,V1,V2,e.
119                      L1 [0, e - 1] ▶* L2 → L2 ⊢ V1 [0, e - 1] ▶ V2 →
120                      0 < e → L1. ⓑ{I} V1 [0, e] ▶* L2. ⓑ{I} V2.
121 #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He
122 >(plus_minus_m_m e 1) // /2 width=1/
123 qed.
124
125 lemma ltpss_tps1: ∀L1,L2,I,d,e. L1 [d, e] ▶* L2 →
126                   ∀V1,V2. L2 ⊢ V1 [d, e] ▶ V2 →
127                   L1. ⓑ{I} V1 [d + 1, e] ▶* L2. ⓑ{I} V2.
128 #L1 #L2 #I #d #e #H @(ltpss_ind … H) -L2
129 [ /3 width=1/
130 | #L #L2 #_ #HL2 #IHL #V1 #V2 #HV12
131   elim (ltps_tps_trans … HV12 … HL2) -HV12 #V #HV1 #HV2
132   lapply (IHL … HV1) -IHL -HV1 #HL1
133   @step /2 width=5/ (**) (* /3 width=5/ is too slow *)
134 ]
135 qed.
136
137 lemma ltpss_tps1_lt: ∀L1,L2,I,V1,V2,d,e.
138                      L1 [d - 1, e] ▶* L2 → L2 ⊢ V1 [d - 1, e] ▶ V2 →
139                      0 < d → L1. ⓑ{I} V1 [d, e] ▶* L2. ⓑ{I} V2.
140 #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd
141 >(plus_minus_m_m d 1) // /2 width=1/
142 qed.
143
144 (* Advanced forward lemmas **************************************************)
145
146 lemma ltpss_fwd_tpss21: ∀e,K1,I,V1,L2. 0 < e → K1. ⓑ{I} V1 [0, e] ▶* L2 →
147                         ∃∃K2,V2. K1 [0, e - 1] ▶* K2 & K1 ⊢ V1 [0, e - 1] ▶* V2 &
148                                  L2 = K2. ⓑ{I} V2.
149 #e #K1 #I #V1 #L2 #He #H @(ltpss_ind … H) -L2
150 [ /2 width=5/
151 | #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct
152   elim (ltps_inv_tps21 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H
153   lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2
154   lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/
155 ]
156 qed-.
157
158 lemma ltpss_fwd_tpss11: ∀d,e,I,K1,V1,L2. 0 < d → K1. ⓑ{I} V1 [d, e] ▶* L2 →
159                         ∃∃K2,V2. K1 [d - 1, e] ▶* K2 &
160                                  K1 ⊢ V1 [d - 1, e] ▶* V2 &
161                                  L2 = K2. ⓑ{I} V2.
162 #d #e #K1 #I #V1 #L2 #Hd #H @(ltpss_ind … H) -L2
163 [ /2 width=5/
164 | #L #L2 #_ #HL2 * #K #V #HK1 #HV1 #H destruct
165   elim (ltps_inv_tps11 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H
166   lapply (ltps_tps_trans_eq … HV2 … HK2) -HV2 #HV2
167   lapply (ltpss_tpss_trans_eq … HK1 … HV2) -HV2 #HV2 /3 width=5/
168 ]
169 qed-.