1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "ground/xoa/ex_1_2.ma".
16 include "ground/xoa/ex_3_2.ma".
17 include "ground/arith/wf1_ind_plt.ma".
18 include "static_2/notation/relations/tildeminus_2.ma".
19 include "static_2/syntax/term_weight.ma".
21 (* SORT-IRRELEVANT WHD EQUIVALENCE ON TERMS *********************************)
23 inductive teqw: relation term ≝
24 | teqw_sort: ∀s1,s2. teqw (⋆s1) (⋆s2)
25 | teqw_lref: ∀i. teqw (#i) (#i)
26 | teqw_gref: ∀l. teqw (§l) (§l)
27 | teqw_abbr: ∀p,V1,V2,T1,T2. (p=Ⓣ→teqw T1 T2) → teqw (ⓓ[p]V1.T1) (ⓓ[p]V2.T2)
28 | teqw_abst: ∀p,V1,V2,T1,T2. teqw (ⓛ[p]V1.T1) (ⓛ[p]V2.T2)
29 | teqw_appl: ∀V1,V2,T1,T2. teqw T1 T2 → teqw (ⓐV1.T1) (ⓐV2.T2)
30 | teqw_cast: ∀V1,V2,T1,T2. teqw V1 V2 → teqw T1 T2 → teqw (ⓝV1.T1) (ⓝV2.T2)
34 "context-free tail sort-irrelevant equivalence (term)"
35 'TildeMinus T1 T2 = (teqw T1 T2).
37 (* Basic properties *********************************************************)
39 lemma teqw_abbr_pos: ∀V1,V2,T1,T2. T1 ≃ T2 → +ⓓV1.T1 ≃ +ⓓV2.T2.
40 /3 width=1 by teqw_abbr/ qed.
42 lemma teqw_abbr_neg: ∀V1,V2,T1,T2. -ⓓV1.T1 ≃ -ⓓV2.T2.
44 @teqw_abbr #H destruct
47 lemma teqw_refl: reflexive … teqw.
48 #T elim T -T * [||| #p * | * ]
49 /2 width=1 by teqw_sort, teqw_lref, teqw_gref, teqw_abbr, teqw_abst, teqw_appl, teqw_cast/
52 lemma teqw_sym: symmetric … teqw.
53 #T1 #T2 #H elim H -T1 -T2
54 /3 width=3 by teqw_sort, teqw_lref, teqw_gref, teqw_abbr, teqw_abst, teqw_appl, teqw_cast/
57 (* Left basic inversion lemmas **********************************************)
59 fact teqw_inv_sort_sn_aux:
60 ∀X,Y. X ≃ Y → ∀s1. X = ⋆s1 → ∃s2. Y = ⋆s2.
62 [1 : #s1 #s2 #s #H destruct /2 width=2 by ex_intro/
63 |2,3: #i #s #H destruct
64 |4 : #p #V1 #V2 #T1 #T2 #_ #s #H destruct
65 |5 : #p #V1 #V2 #T1 #T2 #s #H destruct
66 |6 : #V1 #V2 #T1 #T2 #_ #s #H destruct
67 |7 : #V1 #V2 #T1 #T2 #_ #_ #s #H destruct
71 lemma teqw_inv_sort_sn:
72 ∀Y,s1. ⋆s1 ≃ Y → ∃s2. Y = ⋆s2.
73 /2 width=4 by teqw_inv_sort_sn_aux/ qed-.
75 fact teqw_inv_lref_sn_aux:
76 ∀X,Y. X ≃ Y → ∀i. X = #i → Y = #i.
78 [1 : #s1 #s2 #j #H destruct
80 |4 : #p #V1 #V2 #T1 #T2 #_ #j #H destruct
81 |5 : #p #V1 #V2 #T1 #T2 #j #H destruct
82 |6 : #V1 #V2 #T1 #T2 #_ #j #H destruct
83 |7 : #V1 #V2 #T1 #T2 #_ #_ #j #H destruct
87 lemma teqw_inv_lref_sn: ∀Y,i. #i ≃ Y → Y = #i.
88 /2 width=5 by teqw_inv_lref_sn_aux/ qed-.
90 fact teqw_inv_gref_sn_aux:
91 ∀X,Y. X ≃ Y → ∀l. X = §l → Y = §l.
93 [1 : #s1 #s2 #k #H destruct
95 |4 : #p #V1 #V2 #T1 #T2 #_ #k #H destruct
96 |5 : #p #V1 #V2 #T1 #T2 #k #H destruct
97 |6 : #V1 #V2 #T1 #T2 #_ #k #H destruct
98 |7 : #V1 #V2 #T1 #T2 #_ #_ #j #H destruct
102 lemma teqw_inv_gref_sn:
103 ∀Y,l. §l ≃ Y → Y = §l.
104 /2 width=5 by teqw_inv_gref_sn_aux/ qed-.
106 fact teqw_inv_abbr_sn_aux:
107 ∀X,Y. X ≃ Y → ∀p,V1,T1. X = ⓓ[p]V1.T1 →
108 ∃∃V2,T2. p = Ⓣ → T1 ≃ T2 & Y = ⓓ[p]V2.T2.
110 [1 : #s1 #s2 #q #W1 #U1 #H destruct
111 |2,3: #i #q #W1 #U1 #H destruct
112 |4 : #p #V1 #V2 #T1 #T2 #HT #q #W1 #U1 #H destruct /3 width=4 by ex2_2_intro/
113 |5 : #p #V1 #V2 #T1 #T2 #q #W1 #U1 #H destruct
114 |6 : #V1 #V2 #T1 #T2 #_ #q #W1 #U1 #H destruct
115 |7 : #V1 #V2 #T1 #T2 #_ #_ #q #W1 #U1 #H destruct
119 lemma teqw_inv_abbr_sn:
120 ∀p,V1,T1,Y. ⓓ[p]V1.T1 ≃ Y →
121 ∃∃V2,T2. p = Ⓣ → T1 ≃ T2 & Y = ⓓ[p]V2.T2.
122 /2 width=4 by teqw_inv_abbr_sn_aux/ qed-.
124 fact teqw_inv_abst_sn_aux:
125 ∀X,Y. X ≃ Y → ∀p,V1,T1. X = ⓛ[p]V1.T1 →
126 ∃∃V2,T2. Y = ⓛ[p]V2.T2.
128 [1 : #s1 #s2 #q #W1 #U1 #H destruct
129 |2,3: #i #q #W1 #U1 #H destruct
130 |4 : #p #V1 #V2 #T1 #T2 #_ #q #W1 #U1 #H destruct
131 |5 : #p #V1 #V2 #T1 #T2 #q #W1 #U1 #H destruct /2 width=3 by ex1_2_intro/
132 |6 : #V1 #V2 #T1 #T2 #_ #q #W1 #U1 #H destruct
133 |7 : #V1 #V2 #T1 #T2 #_ #_ #q #W1 #U1 #H destruct
137 lemma teqw_inv_abst_sn:
138 ∀p,V1,T1,Y. ⓛ[p]V1.T1 ≃ Y →
139 ∃∃V2,T2. Y = ⓛ[p]V2.T2.
140 /2 width=5 by teqw_inv_abst_sn_aux/ qed-.
142 fact teqw_inv_appl_sn_aux:
143 ∀X,Y. X ≃ Y → ∀V1,T1. X = ⓐV1.T1 →
144 ∃∃V2,T2. T1 ≃ T2 & Y = ⓐV2.T2.
146 [1 : #s1 #s2 #W1 #U1 #H destruct
147 |2,3: #i #W1 #U1 #H destruct
148 |4 : #p #V1 #V2 #T1 #T2 #HT #W1 #U1 #H destruct
149 |5 : #p #V1 #V2 #T1 #T2 #W1 #U1 #H destruct
150 |6 : #V1 #V2 #T1 #T2 #HT #W1 #U1 #H destruct /2 width=4 by ex2_2_intro/
151 |7 : #V1 #V2 #T1 #T2 #_ #_ #W1 #U1 #H destruct
155 lemma teqw_inv_appl_sn:
156 ∀V1,T1,Y. ⓐV1.T1 ≃ Y →
157 ∃∃V2,T2. T1 ≃ T2 & Y = ⓐV2.T2.
158 /2 width=4 by teqw_inv_appl_sn_aux/ qed-.
160 fact teqw_inv_cast_sn_aux:
161 ∀X,Y. X ≃ Y → ∀V1,T1. X = ⓝV1.T1 →
162 ∃∃V2,T2. V1 ≃ V2 & T1 ≃ T2 & Y = ⓝV2.T2.
164 [1 : #s1 #s2 #W1 #U1 #H destruct
165 |2,3: #i #W1 #U1 #H destruct
166 |4 : #p #V1 #V2 #T1 #T2 #_ #W1 #U1 #H destruct
167 |5 : #p #V1 #V2 #T1 #T2 #W1 #U1 #H destruct
168 |6 : #V1 #V2 #T1 #T2 #_ #W1 #U1 #H destruct
169 |7 : #V1 #V2 #T1 #T2 #HV #HT #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
173 lemma teqw_inv_cast_sn:
174 ∀V1,T1,Y. ⓝV1.T1 ≃ Y →
175 ∃∃V2,T2. V1 ≃ V2 & T1 ≃ T2 & Y = ⓝV2.T2.
176 /2 width=3 by teqw_inv_cast_sn_aux/ qed-.
178 (* Advanced inversion lemmas ************************************************)
180 lemma teqw_inv_abbr_pos_sn:
181 ∀V1,T1,Y. +ⓓV1.T1 ≃ Y → ∃∃V2,T2. T1 ≃ T2 & Y = +ⓓV2.T2.
183 elim (teqw_inv_abbr_sn … H) -H #V2 #T2
184 /3 width=4 by ex2_2_intro/
187 lemma teqw_inv_abbr_neg_sn:
188 ∀V1,T1,Y. -ⓓV1.T1 ≃ Y → ∃∃V2,T2. Y = -ⓓV2.T2.
190 elim (teqw_inv_abbr_sn … H) -H #V2 #T2 #_
191 /2 width=3 by ex1_2_intro/
194 lemma teqw_inv_abbr_pos_bi:
195 ∀V1,V2,T1,T2. +ⓓV1.T1 ≃ +ⓓV2.T2 → T1 ≃ T2.
197 elim (teqw_inv_abbr_pos_sn … H) -H #W2 #U2 #HTU #H destruct //
200 lemma teqw_inv_appl_bi:
201 ∀V1,V2,T1,T2. ⓐV1.T1 ≃ ⓐV2.T2 → T1 ≃ T2.
203 elim (teqw_inv_appl_sn … H) -H #W2 #U2 #HTU #H destruct //
206 lemma teqw_inv_cast_bi:
207 ∀V1,V2,T1,T2. ⓝV1.T1 ≃ ⓝV2.T2 → ∧∧ V1 ≃ V2 & T1 ≃ T2.
209 elim (teqw_inv_cast_sn … H) -H #W2 #U2 #HVW #HTU #H destruct
213 lemma teqw_inv_cast_xy_y: ∀T,V. ⓝV.T ≃ T → ⊥.
214 @(wf1_ind_plt … tw) #p #IH #T #Hp #V #H destruct
215 elim (teqw_inv_cast_sn … H) -H #X1 #X2 #_ #HX2 #H destruct -V
219 (* Advanced forward lemmas **************************************************)
221 lemma teqw_fwd_pair_sn (I):
222 ∀V1,T1,X2. ②[I]V1.T1 ≃ X2 → ∃∃V2,T2. X2 = ②[I]V2.T2.
223 * [ #p ] * [ cases p -p ] #V1 #T1 #X2 #H
224 [ elim (teqw_inv_abbr_pos_sn … H) -H #V2 #T2 #_ #H
225 | elim (teqw_inv_abbr_neg_sn … H) -H #V2 #T2 #H
226 | elim (teqw_inv_abst_sn … H) -H #V2 #T2 #H
227 | elim (teqw_inv_appl_sn … H) -H #V2 #T2 #_ #H
228 | elim (teqw_inv_cast_sn … H) -H #V2 #T2 #_ #_ #H
229 ] /2 width=3 by ex1_2_intro/
232 lemma teqw_fwd_pair_bi (I1) (I2):
233 ∀V1,V2,T1,T2. ②[I1]V1.T1 ≃ ②[I2]V2.T2 → I1 = I2.
234 #I1 #I2 #V1 #V2 #T1 #T2 #H
235 elim (teqw_fwd_pair_sn … H) -H #W2 #U2 #H destruct //
238 (* Advanced properties ******************************************************)
240 lemma teqw_dec: ∀T1,T2. Decidable (T1 ≃ T2).
241 #T1 elim T1 -T1 [ * #s1 | #I1 #V1 #T1 #IHV #IHT ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ]
242 [ /3 width=1 by teqw_sort, or_introl/
245 elim (teqw_inv_sort_sn … H) -H #x #H destruct
248 lapply (teqw_inv_lref_sn … H) -H #H destruct
250 elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
252 lapply (teqw_inv_lref_sn … H) -H #H destruct /2 width=1 by/
255 lapply (teqw_inv_gref_sn … H) -H #H destruct
257 elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
259 lapply (teqw_inv_gref_sn … H) -H #H destruct /2 width=1 by/
262 elim (teqw_fwd_pair_sn … H) -H #X1 #X2 #H destruct
264 elim (eq_item2_dec I1 I2) #HI12 destruct
265 [ cases I2 -I2 [ #p ] * [ cases p -p ]
266 [ elim (IHT T2) -IHT #HT12
267 [ /3 width=1 by teqw_abbr_pos, or_introl/
268 | /4 width=3 by teqw_inv_abbr_pos_bi, or_intror/
270 | /3 width=1 by teqw_abbr_neg, or_introl/
271 | /3 width=1 by teqw_abst, or_introl/
272 | elim (IHT T2) -IHT #HT12
273 [ /3 width=1 by teqw_appl, or_introl/
274 | /4 width=3 by teqw_inv_appl_bi, or_intror/
276 | elim (IHV V2) -IHV #HV12
277 elim (IHT T2) -IHT #HT12
278 [1: /3 width=1 by teqw_cast, or_introl/
280 elim (teqw_inv_cast_bi … H) -H #HV12 #HT12
284 | /4 width=5 by teqw_fwd_pair_bi, or_intror/