]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/syntax/teqx.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / teqx.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_2/xoa/ex_3_2.ma".
16 include "static_2/notation/relations/stareq_2.ma".
17 include "static_2/syntax/term.ma".
18
19 (* SORT-IRRELEVANT EQUIVALENCE ON TERMS *************************************)
20
21 inductive teqx: relation term ≝
22 | teqx_sort: ∀s1,s2. teqx (⋆s1) (⋆s2)
23 | teqx_lref: ∀i. teqx (#i) (#i)
24 | teqx_gref: ∀l. teqx (§l) (§l)
25 | teqx_pair: ∀I,V1,V2,T1,T2. teqx V1 V2 → teqx T1 T2 → teqx (②[I]V1.T1) (②[I]V2.T2)
26 .
27
28 interpretation
29    "context-free sort-irrelevant equivalence (term)"
30    'StarEq T1 T2 = (teqx T1 T2).
31
32 (* Basic properties *********************************************************)
33
34 lemma teqx_refl: reflexive … teqx.
35 #T elim T -T /2 width=1 by teqx_pair/
36 * /2 width=1 by teqx_lref, teqx_gref/
37 qed.
38
39 lemma teqx_sym: symmetric … teqx.
40 #T1 #T2 #H elim H -T1 -T2
41 /2 width=3 by teqx_sort, teqx_lref, teqx_gref, teqx_pair/
42 qed-.
43
44 (* Basic inversion lemmas ***************************************************)
45
46 fact teqx_inv_sort1_aux: ∀X,Y. X ≛ Y → ∀s1. X = ⋆s1 →
47                          ∃s2. Y = ⋆s2.
48 #X #Y * -X -Y
49 [ #s1 #s2 #s #H destruct /2 width=2 by ex_intro/
50 | #i #s #H destruct
51 | #l #s #H destruct
52 | #I #V1 #V2 #T1 #T2 #_ #_ #s #H destruct
53 ]
54 qed-.
55
56 lemma teqx_inv_sort1: ∀Y,s1. ⋆s1 ≛ Y →
57                       ∃s2. Y = ⋆s2.
58 /2 width=4 by teqx_inv_sort1_aux/ qed-.
59
60 fact teqx_inv_lref1_aux: ∀X,Y. X ≛ Y → ∀i. X = #i → Y = #i.
61 #X #Y * -X -Y //
62 [ #s1 #s2 #j #H destruct
63 | #I #V1 #V2 #T1 #T2 #_ #_ #j #H destruct
64 ]
65 qed-.
66
67 lemma teqx_inv_lref1: ∀Y,i. #i ≛ Y → Y = #i.
68 /2 width=5 by teqx_inv_lref1_aux/ qed-.
69
70 fact teqx_inv_gref1_aux: ∀X,Y. X ≛ Y → ∀l. X = §l → Y = §l.
71 #X #Y * -X -Y //
72 [ #s1 #s2 #k #H destruct
73 | #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct
74 ]
75 qed-.
76
77 lemma teqx_inv_gref1: ∀Y,l. §l ≛ Y → Y = §l.
78 /2 width=5 by teqx_inv_gref1_aux/ qed-.
79
80 fact teqx_inv_pair1_aux: ∀X,Y. X ≛ Y → ∀I,V1,T1. X = ②[I]V1.T1 →
81                          ∃∃V2,T2. V1 ≛ V2 & T1 ≛ T2 & Y = ②[I]V2.T2.
82 #X #Y * -X -Y
83 [ #s1 #s2 #J #W1 #U1 #H destruct
84 | #i #J #W1 #U1 #H destruct
85 | #l #J #W1 #U1 #H destruct
86 | #I #V1 #V2 #T1 #T2 #HV #HT #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/
87 ]
88 qed-.
89
90 lemma teqx_inv_pair1: ∀I,V1,T1,Y. ②[I]V1.T1 ≛ Y →
91                       ∃∃V2,T2. V1 ≛ V2 & T1 ≛ T2 & Y = ②[I]V2.T2.
92 /2 width=3 by teqx_inv_pair1_aux/ qed-.
93
94 lemma teqx_inv_sort2: ∀X1,s2. X1 ≛ ⋆s2 →
95                       ∃s1. X1 = ⋆s1.
96 #X1 #s2 #H
97 elim (teqx_inv_sort1 X1 s2)
98 /2 width=2 by teqx_sym, ex_intro/
99 qed-.
100
101 lemma teqx_inv_pair2: ∀I,X1,V2,T2. X1 ≛ ②[I]V2.T2 →
102                       ∃∃V1,T1. V1 ≛ V2 & T1 ≛ T2 & X1 = ②[I]V1.T1.
103 #I #X1 #V2 #T2 #H
104 elim (teqx_inv_pair1 I V2 T2 X1)
105 [ #V1 #T1 #HV #HT #H destruct ]
106 /3 width=5 by teqx_sym, ex3_2_intro/
107 qed-.
108
109 (* Advanced inversion lemmas ************************************************)
110
111 lemma teqx_inv_pair: ∀I1,I2,V1,V2,T1,T2. ②[I1]V1.T1 ≛ ②[I2]V2.T2 →
112                      ∧∧ I1 = I2 & V1 ≛ V2 & T1 ≛ T2.
113 #I1 #I2 #V1 #V2 #T1 #T2 #H elim (teqx_inv_pair1 … H) -H
114 #V0 #T0 #HV #HT #H destruct /2 width=1 by and3_intro/
115 qed-.
116
117 lemma teqx_inv_pair_xy_x: ∀I,V,T. ②[I]V.T ≛ V → ⊥.
118 #I #V elim V -V
119 [ #J #T #H elim (teqx_inv_pair1 … H) -H #X #Y #_ #_ #H destruct
120 | #J #X #Y #IHX #_ #T #H elim (teqx_inv_pair … H) -H #H #HY #_ destruct /2 width=2 by/
121 ]
122 qed-.
123
124 lemma teqx_inv_pair_xy_y: ∀I,T,V. ②[I]V.T ≛ T → ⊥.
125 #I #T elim T -T
126 [ #J #V #H elim (teqx_inv_pair1 … H) -H #X #Y #_ #_ #H destruct
127 | #J #X #Y #_ #IHY #V #H elim (teqx_inv_pair … H) -H #H #_ #HY destruct /2 width=2 by/
128 ]
129 qed-.
130
131 (* Basic forward lemmas *****************************************************)
132
133 lemma teqx_fwd_atom1: ∀I,Y. ⓪[I] ≛ Y → ∃J. Y = ⓪[J].
134 * #x #Y #H [ elim (teqx_inv_sort1 … H) -H ]
135 /3 width=4 by teqx_inv_gref1, teqx_inv_lref1, ex_intro/
136 qed-.
137
138 (* Advanced properties ******************************************************)
139
140 lemma teqx_dec: ∀T1,T2. Decidable (T1 ≛ T2).
141 #T1 elim T1 -T1 [ * #s1 | #I1 #V1 #T1 #IHV #IHT ] * [1,3,5,7: * #s2 |*: #I2 #V2 #T2 ]
142 [ /3 width=1 by teqx_sort, or_introl/
143 |2,3,13:
144   @or_intror #H
145   elim (teqx_inv_sort1 … H) -H #x #H destruct
146 |4,6,14:
147   @or_intror #H
148   lapply (teqx_inv_lref1 … H) -H #H destruct
149 |5:
150   elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
151   @or_intror #H
152   lapply (teqx_inv_lref1 … H) -H #H destruct /2 width=1 by/
153 |7,8,15:
154   @or_intror #H
155   lapply (teqx_inv_gref1 … H) -H #H destruct
156 |9:
157   elim (eq_nat_dec s1 s2) #Hs12 destruct /2 width=1 by or_introl/
158   @or_intror #H
159   lapply (teqx_inv_gref1 … H) -H #H destruct /2 width=1 by/
160 |10,11,12:
161   @or_intror #H
162   elim (teqx_inv_pair1 … H) -H #X1 #X2 #_ #_ #H destruct
163 |16:
164   elim (eq_item2_dec I1 I2) #HI12 destruct
165   [ elim (IHV V2) -IHV #HV12
166     elim (IHT T2) -IHT #HT12
167     [ /3 width=1 by teqx_pair, or_introl/ ]
168   ]
169   @or_intror #H
170   elim (teqx_inv_pair … H) -H /2 width=1 by/
171 ]
172 qed-.
173
174 (* Negated inversion lemmas *************************************************)
175
176 lemma tneqx_inv_pair: ∀I1,I2,V1,V2,T1,T2.
177                       (②[I1]V1.T1 ≛ ②[I2]V2.T2 → ⊥) →
178                       ∨∨ I1 = I2 → ⊥
179                       |  (V1 ≛ V2 → ⊥)
180                       |  (T1 ≛ T2 → ⊥).
181 #I1 #I2 #V1 #V2 #T1 #T2 #H12
182 elim (eq_item2_dec I1 I2) /3 width=1 by or3_intro0/ #H destruct
183 elim (teqx_dec V1 V2) /3 width=1 by or3_intro1/
184 elim (teqx_dec T1 T2) /4 width=1 by teqx_pair, or3_intro2/
185 qed-.