]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/equivalence/cpcs_cpcs.ma
- lambda_delta: programmed renaming to lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / equivalence / cpcs_cpcs.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/computation/cprs_lift.ma".
16 include "basic_2/computation/cprs_cprs.ma".
17 include "basic_2/conversion/cpc_cpc.ma".
18 include "basic_2/equivalence/cpcs_cprs.ma".
19
20 (* CONTEXT-SENSITIVE PARALLEL EQUIVALENCE ON TERMS **************************)
21
22 (* Advanced inversion lemmas ************************************************)
23
24 lemma cpcs_inv_cprs: ∀L,T1,T2. L ⊢ T1 ⬌* T2 →
25                      ∃∃T. L ⊢ T1 ➡* T & L ⊢ T2 ➡* T.
26 #L #T1 #T2 #H @(cpcs_ind … H) -T2
27 [ /3 width=3/
28 | #T #T2 #_ #HT2 * #T0 #HT10 elim HT2 -HT2 #HT2 #HT0
29   [ elim (cprs_strip … HT0 … HT2) -T #T #HT0 #HT2
30     lapply (cprs_strap1 … HT10 … HT0) -T0 /2 width=3/
31   | lapply (cprs_strap2 … HT2 … HT0) -T /2 width=3/
32   ]
33 ]
34 qed-.
35
36 (* Basic_1: was: pc3_gen_sort *)
37 lemma cpcs_inv_sort: ∀L,k1,k2. L ⊢ ⋆k1 ⬌* ⋆k2 → k1 = k2.
38 #L #k1 #k2 #H
39 elim (cpcs_inv_cprs … H) -H #T #H1
40 >(cprs_inv_sort1 … H1) -T #H2
41 lapply (cprs_inv_sort1 … H2) -L #H destruct //
42 qed-.
43
44 (* Basic_1: was: pc3_gen_sort_abst *)
45 lemma cpcs_inv_sort_abst: ∀a,L,W,T,k. L ⊢ ⋆k ⬌* ⓛ{a}W.T → ⊥.
46 #a #L #W #T #k #H
47 elim (cpcs_inv_cprs … H) -H #X #H1
48 >(cprs_inv_sort1 … H1) -X #H2
49 elim (cprs_inv_abst1 Abst W … H2) -H2 #W0 #T0 #_ #_ #H destruct
50 qed-.
51
52 (* Basic_1: was: pc3_gen_abst *)
53 lemma cpcs_inv_abst: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → ∀I,V.
54                      ∧∧ L ⊢ W1 ⬌* W2 & L. ②{I}V ⊢ T1 ⬌* T2 & a1 = a2.
55 #a1 #a2 #L #W1 #W2 #T1 #T2 #H #I #V
56 elim (cpcs_inv_cprs … H) -H #T #H1 #H2
57 elim (cprs_inv_abst1 I V … H1) -H1 #W0 #T0 #HW10 #HT10 #H destruct
58 elim (cprs_inv_abst1 I V … H2) -H2 #W #T #HW2 #HT2 #H destruct /3 width=3/
59 qed-.
60
61 (* Basic_1: was: pc3_gen_abst_shift *)
62 lemma cpcs_inv_abst_shift: ∀a1,a2,L,W1,W2,T1,T2. L ⊢ ⓛ{a1}W1.T1 ⬌* ⓛ{a2}W2.T2 → ∀W.
63                            ∧∧ L ⊢ W1 ⬌* W2 & L. ⓛW ⊢ T1 ⬌* T2 & a1 = a2.
64 #a1 #a2 #L #W1 #W2 #T1 #T2 #H #W
65 lapply (cpcs_inv_abst … H Abst W) -H //
66 qed.
67
68 lemma cpcs_inv_abst1: ∀a,L,W1,T1,T. L ⊢ ⓛ{a}W1.T1 ⬌* T →
69                       ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2.
70 #a #L #W1 #T1 #T #H
71 elim (cpcs_inv_cprs … H) -H #X #H1 #H2
72 elim (cprs_inv_abst1 Abst W1 … H1) -H1 #W2 #T2 #HW12 #HT12 #H destruct
73 @(ex2_2_intro … H2) -H2 /2 width=2/ (**) (* explicit constructor, /3 width=6/ is slow *)
74 qed-.
75
76 lemma cpcs_inv_abst2: ∀a,L,W1,T1,T. L ⊢ T ⬌* ⓛ{a}W1.T1 →
77                       ∃∃W2,T2. L ⊢ T ➡* ⓛ{a}W2.T2 & L ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2.
78 /3 width=1 by cpcs_inv_abst1, cpcs_sym/ qed-.
79
80 (* Basic_1: was: pc3_gen_lift *)
81 lemma cpcs_inv_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K →
82                      ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 →
83                      L ⊢ U1 ⬌* U2 → K ⊢ T1 ⬌* T2.
84 #L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HU12
85 elim (cpcs_inv_cprs … HU12) -HU12 #U #HU1 #HU2
86 elim (cprs_inv_lift1 … HLK … HTU1 … HU1) -U1 #T #HTU #HT1
87 elim (cprs_inv_lift1 … HLK … HTU2 … HU2) -L -U2 #X #HXU
88 >(lift_inj … HXU … HTU) -X -U -d -e /2 width=3/
89 qed-.
90
91 (* Advanced properties ******************************************************)
92
93 lemma cpr_cprs_conf: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡ T2 → L ⊢ T1 ⬌* T2.
94 #L #T #T1 #T2 #HT1 #HT2
95 elim (cprs_strip … HT1 … HT2) /2 width=3 by cpr_cprs_div/
96 qed-.
97
98 lemma cprs_cpr_conf: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡ T2 → L ⊢ T2 ⬌* T1.
99 #L #T #T1 #T2 #HT1 #HT2
100 elim (cprs_strip … HT1 … HT2) /2 width=3 by cprs_cpr_div/
101 qed-.
102
103 lemma cprs_conf: ∀L,T,T1,T2. L ⊢ T ➡* T1 → L ⊢ T ➡* T2 → L ⊢ T1 ⬌* T2.
104 #L #T #T1 #T2 #HT1 #HT2
105 elim (cprs_conf … HT1 … HT2) /2 width=3/
106 qed-.
107
108 (* Basic_1: was only: pc3_thin_dx *)
109 lemma cpcs_flat: ∀L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L ⊢ T1 ⬌* T2 →
110                  ∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2.
111 #L #V1 #V2 #HV12 #T1 #T2 #HT12 #I
112 elim (cpcs_inv_cprs … HV12) -HV12 #V #HV1 #HV2
113 elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_flat, cprs_div/ (**) (* /3 width=5/ is too slow *)
114 qed.
115
116 lemma cpcs_flat_dx_tpr_rev: ∀L,V1,V2. V2 ➡ V1 → ∀T1,T2. L ⊢ T1 ⬌* T2 →
117                             ∀I. L ⊢ ⓕ{I}V1. T1 ⬌* ⓕ{I}V2. T2.
118 /3 width=1/ qed.
119
120 lemma cpcs_abst: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 →
121                  ∀V,T1,T2. L.ⓛV ⊢ T1 ⬌* T2 → L ⊢ ⓛ{a}V1. T1 ⬌* ⓛ{a}V2. T2.
122 #a #L #V1 #V2 #HV12 #V #T1 #T2 #HT12
123 elim (cpcs_inv_cprs … HV12) -HV12
124 elim (cpcs_inv_cprs … HT12) -HT12
125 /3 width=6 by cprs_div, cprs_abst/ (**) (* /3 width=6/ is a bit slow *)
126 qed.
127
128 lemma cpcs_abbr_dx: ∀a,L,V,T1,T2. L.ⓓV ⊢ T1 ⬌* T2 → L ⊢ ⓓ{a}V. T1 ⬌* ⓓ{a}V. T2.
129 #a #L #V #T1 #T2 #HT12
130 elim (cpcs_inv_cprs … HT12) -HT12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
131 qed.
132
133 lemma cpcs_bind_dx: ∀a,I,L,V,T1,T2. L.ⓑ{I}V ⊢ T1 ⬌* T2 →
134                     L ⊢ ⓑ{a,I}V. T1 ⬌* ⓑ{a,I}V. T2.
135 #a * /2 width=1/ /2 width=2/ qed.
136
137 lemma cpcs_abbr_sn: ∀a,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓓ{a}V1. T ⬌* ⓓ{a}V2. T.
138 #a #L #V1 #V2 #T #HV12
139 elim (cpcs_inv_cprs … HV12) -HV12 /3 width=5 by cprs_div, cprs_abbr1/ (**) (* /3 width=5/ is a bit slow *)
140 qed.
141
142 lemma cpcs_bind_sn: ∀a,I,L,V1,V2,T. L ⊢ V1 ⬌* V2 → L ⊢ ⓑ{a,I}V1. T ⬌* ⓑ{a,I}V2. T.
143 #a * /2 width=1/ /2 width=2/ qed.
144
145 lemma cpcs_beta_dx: ∀a,L,V1,V2,W,T1,T2.
146                     L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ⬌* T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ⬌* ⓓ{a}V2.T2.
147 #a #L #V1 #V2 #W #T1 #T2 #HV12 #HT12
148 elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
149 lapply (cprs_beta_dx … HV12 HT1 a) -HV12 -HT1 #HT1
150 lapply (cprs_lsubs_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2
151 @(cprs_div … HT1) /2 width=1/
152 qed.
153
154 lemma cpcs_beta_dx_tpr_rev: ∀a,L,V1,V2,W,T1,T2.
155                             V1 ➡ V2 → L.ⓛW ⊢ T2 ⬌* T1 →
156                             L ⊢ ⓓ{a}V2.T2 ⬌* ⓐV1.ⓛ{a}W.T1.
157 /4 width=1/ qed.
158
159 (* Note: it does not hold replacing |L1| with |L2| *)
160 lemma cpcs_lsubs_trans: ∀L1,T1,T2. L1 ⊢ T1 ⬌* T2 →
161                         ∀L2. L2 ≼ [0, |L1|] L1 → L2 ⊢ T1 ⬌* T2.
162 #L1 #T1 #T2 #HT12
163 elim (cpcs_inv_cprs … HT12) -HT12
164 /3 width=5 by cprs_div, cprs_lsubs_trans/ (**) (* /3 width=5/ is a bit slow *)
165 qed.
166
167 (* Basic_1: was: pc3_lift *)
168 lemma cpcs_lift: ∀L,K,d,e. ⇩[d, e] L ≡ K →
169                  ∀T1,U1. ⇧[d, e] T1 ≡ U1 → ∀T2,U2. ⇧[d, e] T2 ≡ U2 →
170                  K ⊢ T1 ⬌* T2 → L ⊢ U1 ⬌* U2.
171 #L #K #d #e #HLK #T1 #U1 #HTU1 #T2 #U2 #HTU2 #HT12
172 elim (cpcs_inv_cprs … HT12) -HT12 #T #HT1 #HT2
173 elim (lift_total T d e) #U #HTU
174 lapply (cprs_lift … HLK … HTU1 … HT1 … HTU) -T1 #HU1
175 lapply (cprs_lift … HLK … HTU2 … HT2 … HTU) -K -T2 -T -d -e /2 width=3/
176 qed.
177
178 lemma cpcs_strip: ∀L,T1,T. L ⊢ T ⬌* T1 → ∀T2. L ⊢ T ⬌ T2 →
179                   ∃∃T0. L ⊢ T1 ⬌ T0 & L ⊢ T2 ⬌* T0.
180 /3 width=3/ qed.
181
182 (* Main properties **********************************************************)
183
184 (* Basic_1: was pc3_t *)
185 theorem cpcs_trans: ∀L,T1,T. L ⊢ T1 ⬌* T → ∀T2. L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
186 /2 width=3/ qed.
187
188 theorem cpcs_canc_sn: ∀L,T,T1,T2. L ⊢ T ⬌* T1 → L ⊢ T ⬌* T2 → L ⊢ T1 ⬌* T2.
189 /3 width=3 by cpcs_trans, cpcs_sym/ qed. (**) (* /3 width=3/ is too slow *)
190
191 theorem cpcs_canc_dx: ∀L,T,T1,T2. L ⊢ T1 ⬌* T → L ⊢ T2 ⬌* T → L ⊢ T1 ⬌* T2.
192 /3 width=3 by cpcs_trans, cpcs_sym/ qed. (**) (* /3 width=3/ is too slow *)
193
194 lemma cpcs_abbr1: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV1 ⊢ T1 ⬌* T2 →
195                   L ⊢ ⓓ{a}V1. T1 ⬌* ⓓ{a}V2. T2.
196 #a #L #V1 #V2 #HV12 #T1 #T2 #HT12
197 @(cpcs_trans … (ⓓ{a}V1.T2)) /2 width=1/
198 qed.
199
200 lemma cpcs_abbr2: ∀a,L,V1,V2. L ⊢ V1 ⬌* V2 → ∀T1,T2. L.ⓓV2 ⊢ T1 ⬌* T2 →
201                   L ⊢ ⓓ{a}V1. T1 ⬌* ⓓ{a}V2. T2.
202 #a #L #V1 #V2 #HV12 #T1 #T2 #HT12
203 @(cpcs_trans … (ⓓ{a}V2.T1)) /2 width=1/
204 qed.