]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/static/reqx.ma
λδ site update
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / reqx.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 "static_2/notation/relations/approxeqsn_3.ma".
16 include "static_2/syntax/teqx_ext.ma".
17 include "static_2/static/reqg.ma".
18
19 (* SORT-IRRELEVANT EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ***)
20 (*
21 definition reqx: relation3 … ≝
22            reqg sfull.
23 *)
24 interpretation
25   "sort-irrelevant equivalence on referred entries (local environment)"
26   'ApproxEqSn T L1 L2 = (reqg sfull T L1 L2).
27
28 interpretation
29   "sort-irrelevant ranged equivalence (local environment)"
30   'StarEqSn f L1 L2 = (sex ceqx_ext cfull f L1 L2).
31
32 (* Basic properties ***********************************************************)
33 (*
34 lemma frees_teqx_conf_reqx:
35       ∀f,L1,T1. L1 ⊢ 𝐅+❪T1❫ ≘ f → ∀T2. T1 ≛ T2 →
36       ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅+❪T2❫ ≘ f.
37 #f #L1 #T1 #H elim H -f -L1 -T1
38 [ #f #L1 #s1 #Hf #X #H1 #L2 #_
39   elim (teqx_inv_sort1 … H1) -H1 #s2 #H destruct
40   /2 width=3 by frees_sort/
41 | #f #i #Hf #X #H1
42   >(teqx_inv_lref1 … H1) -X #Y #H2
43   >(sex_inv_atom1 … H2) -Y
44   /2 width=1 by frees_atom/
45 | #f #I #L1 #V1 #_ #IH #X #H1
46   >(teqx_inv_lref1 … H1) -X #Y #H2
47   elim (sex_inv_next1 … H2) -H2 #Z #L2 #HL12 #HZ #H destruct
48   elim (ext2_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
49   /3 width=1 by frees_pair/
50 | #f #I #L1 #Hf #X #H1
51   >(teqx_inv_lref1 … H1) -X #Y #H2
52   elim (sex_inv_next1 … H2) -H2 #Z #L2 #_ #HZ #H destruct
53   >(ext2_inv_unit_sn … HZ) -Z /2 width=1 by frees_unit/
54 | #f #I #L1 #i #_ #IH #X #H1
55   >(teqx_inv_lref1 … H1) -X #Y #H2
56   elim (sex_inv_push1 … H2) -H2 #J #L2 #HL12 #_ #H destruct
57   /3 width=1 by frees_lref/
58 | #f #L1 #l #Hf #X #H1 #L2 #_
59   >(teqx_inv_gref1 … H1) -X /2 width=1 by frees_gref/
60 | #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
61   elim (teqx_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
62   /6 width=5 by frees_bind, sex_inv_tl, ext2_pair, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/
63 | #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
64   elim (teqx_inv_pair1 … H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
65   /5 width=5 by frees_flat, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/
66 ]
67 qed-.
68
69 lemma frees_teqx_conf:
70       ∀f,L,T1. L ⊢ 𝐅+❪T1❫ ≘ f →
71       ∀T2. T1 ≛ T2 → L ⊢ 𝐅+❪T2❫ ≘ f.
72 /4 width=7 by frees_teqx_conf_reqx, sex_refl, ext2_refl/ qed-.
73
74 lemma frees_reqx_conf:
75       ∀f,L1,T. L1 ⊢ 𝐅+❪T❫ ≘ f →
76       ∀L2. L1 ≛[f] L2 → L2 ⊢ 𝐅+❪T❫ ≘ f.
77 /2 width=7 by frees_teqx_conf_reqx, teqx_refl/ qed-.
78
79 lemma teqx_rex_conf_sn (R):
80       s_r_confluent1 … cdeq (rex R).
81 #R #L1 #T1 #T2 #HT12 #L2 *
82 /3 width=5 by frees_teqx_conf, ex2_intro/
83 qed-.
84
85 lemma teqx_rex_div (R):
86       ∀T1,T2. T1 ≛ T2 →
87       ∀L1,L2. L1 ⪤[R,T2] L2 → L1 ⪤[R,T1] L2.
88 /3 width=5 by teqx_rex_conf_sn, teqx_sym/ qed-.
89
90 lemma teqx_reqx_conf_sn:
91       s_r_confluent1 … cdeq reqx.
92 /2 width=5 by teqx_rex_conf_sn/ qed-.
93
94 lemma teqx_reqx_div:
95       ∀T1,T2. T1 ≛ T2 →
96       ∀L1,L2. L1 ≛[T2] L2 → L1 ≛[T1] L2.
97 /2 width=5 by teqx_rex_div/ qed-.
98
99 lemma reqx_atom: ∀I. ⋆ ≛[⓪[I]] ⋆.
100 /2 width=1 by rex_atom/ qed.
101
102 lemma reqx_sort:
103       ∀I1,I2,L1,L2,s.
104       L1 ≛[⋆s] L2 → L1.ⓘ[I1] ≛[⋆s] L2.ⓘ[I2].
105 /2 width=1 by rex_sort/ qed.
106
107 lemma reqx_pair:
108       ∀I,L1,L2,V1,V2.
109       L1 ≛[V1] L2 → V1 ≛ V2 → L1.ⓑ[I]V1 ≛[#0] L2.ⓑ[I]V2.
110 /2 width=1 by rex_pair/ qed.
111
112 lemma reqx_unit:
113       ∀f,I,L1,L2. 𝐈❪f❫ → L1 ≛[f] L2 →
114       L1.ⓤ[I] ≛[#0] L2.ⓤ[I].
115 /2 width=3 by rex_unit/ qed.
116
117 lemma reqx_lref:
118       ∀I1,I2,L1,L2,i.
119       L1 ≛[#i] L2 → L1.ⓘ[I1] ≛[#↑i] L2.ⓘ[I2].
120 /2 width=1 by rex_lref/ qed.
121
122 lemma reqx_gref:
123       ∀I1,I2,L1,L2,l.
124       L1 ≛[§l] L2 → L1.ⓘ[I1] ≛[§l] L2.ⓘ[I2].
125 /2 width=1 by rex_gref/ qed.
126
127 lemma reqx_bind_repl_dx:
128       ∀I,I1,L1,L2.∀T:term. L1.ⓘ[I] ≛[T] L2.ⓘ[I1] →
129       ∀I2. I ≛ I2 → L1.ⓘ[I] ≛[T] L2.ⓘ[I2].
130 /2 width=2 by rex_bind_repl_dx/ qed-.
131 *)
132 lemma reqg_reqx (S) (T):
133       ∀L1,L2. L1 ≛[S,T] L2 → L1 ≅[T] L2.
134 /2 width=3 by reqg_co/ qed.
135 (*
136 (* Basic inversion lemmas ***************************************************)
137
138 lemma reqx_inv_atom_sn:
139       ∀Y2. ∀T:term. ⋆ ≛[T] Y2 → Y2 = ⋆.
140 /2 width=3 by rex_inv_atom_sn/ qed-.
141
142 lemma reqx_inv_atom_dx:
143       ∀Y1. ∀T:term. Y1 ≛[T] ⋆ → Y1 = ⋆.
144 /2 width=3 by rex_inv_atom_dx/ qed-.
145
146 lemma reqx_inv_zero:
147       ∀Y1,Y2. Y1 ≛[#0] Y2 →
148       ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆
149        | ∃∃I,L1,L2,V1,V2. L1 ≛[V1] L2 & V1 ≛ V2 & Y1 = L1.ⓑ[I]V1 & Y2 = L2.ⓑ[I]V2
150        | ∃∃f,I,L1,L2. 𝐈❪f❫ & L1 ≛[f] L2 & Y1 = L1.ⓤ[I] & Y2 = L2.ⓤ[I].
151 #Y1 #Y2 #H elim (rex_inv_zero … H) -H *
152 /3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/
153 qed-.
154
155 lemma reqx_inv_lref:
156       ∀Y1,Y2,i. Y1 ≛[#↑i] Y2 →
157       ∨∨ ∧∧ Y1 = ⋆ & Y2 = ⋆
158        | ∃∃I1,I2,L1,L2. L1 ≛[#i] L2 & Y1 = L1.ⓘ[I1] & Y2 = L2.ⓘ[I2].
159 /2 width=1 by rex_inv_lref/ qed-.
160
161 (* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *)
162 lemma reqx_inv_bind:
163       ∀p,I,L1,L2,V,T. L1 ≛[ⓑ[p,I]V.T] L2 →
164       ∧∧ L1 ≛[V] L2 & L1.ⓑ[I]V ≛[T] L2.ⓑ[I]V.
165 /2 width=2 by rex_inv_bind/ qed-.
166
167 (* Basic_2A1: uses: lleq_inv_flat *)
168 lemma reqx_inv_flat:
169       ∀I,L1,L2,V,T. L1 ≛[ⓕ[I]V.T] L2 →
170       ∧∧ L1 ≛[V] L2 & L1 ≛[T] L2.
171 /2 width=2 by rex_inv_flat/ qed-.
172
173 (* Advanced inversion lemmas ************************************************)
174
175 lemma reqx_inv_zero_pair_sn:
176       ∀I,Y2,L1,V1. L1.ⓑ[I]V1 ≛[#0] Y2 →
177       ∃∃L2,V2. L1 ≛[V1] L2 & V1 ≛ V2 & Y2 = L2.ⓑ[I]V2.
178 /2 width=1 by rex_inv_zero_pair_sn/ qed-.
179
180 lemma reqx_inv_zero_pair_dx:
181       ∀I,Y1,L2,V2. Y1 ≛[#0] L2.ⓑ[I]V2 →
182       ∃∃L1,V1. L1 ≛[V1] L2 & V1 ≛ V2 & Y1 = L1.ⓑ[I]V1.
183 /2 width=1 by rex_inv_zero_pair_dx/ qed-.
184
185 lemma reqx_inv_lref_bind_sn:
186       ∀I1,Y2,L1,i. L1.ⓘ[I1] ≛[#↑i] Y2 →
187       ∃∃I2,L2. L1 ≛[#i] L2 & Y2 = L2.ⓘ[I2].
188 /2 width=2 by rex_inv_lref_bind_sn/ qed-.
189
190 lemma reqx_inv_lref_bind_dx:
191       ∀I2,Y1,L2,i. Y1 ≛[#↑i] L2.ⓘ[I2] →
192       ∃∃I1,L1. L1 ≛[#i] L2 & Y1 = L1.ⓘ[I1].
193 /2 width=2 by rex_inv_lref_bind_dx/ qed-.
194
195 (* Basic forward lemmas *****************************************************)
196
197 lemma reqx_fwd_zero_pair:
198       ∀I,K1,K2,V1,V2.
199       K1.ⓑ[I]V1 ≛[#0] K2.ⓑ[I]V2 → K1 ≛[V1] K2.
200 /2 width=3 by rex_fwd_zero_pair/ qed-.
201
202 (* Basic_2A1: uses: lleq_fwd_bind_sn lleq_fwd_flat_sn *)
203 lemma reqx_fwd_pair_sn:
204       ∀I,L1,L2,V,T. L1 ≛[②[I]V.T] L2 → L1 ≛[V] L2.
205 /2 width=3 by rex_fwd_pair_sn/ qed-.
206
207 (* Basic_2A1: uses: lleq_fwd_bind_dx lleq_fwd_bind_O_dx *)
208 lemma reqx_fwd_bind_dx:
209       ∀p,I,L1,L2,V,T.
210       L1 ≛[ⓑ[p,I]V.T] L2 → L1.ⓑ[I]V ≛[T] L2.ⓑ[I]V.
211 /2 width=2 by rex_fwd_bind_dx/ qed-.
212
213 (* Basic_2A1: uses: lleq_fwd_flat_dx *)
214 lemma reqx_fwd_flat_dx:
215       ∀I,L1,L2,V,T. L1 ≛[ⓕ[I]V.T] L2 → L1 ≛[T] L2.
216 /2 width=3 by rex_fwd_flat_dx/ qed-.
217
218 lemma reqx_fwd_dx:
219       ∀I2,L1,K2. ∀T:term. L1 ≛[T] K2.ⓘ[I2] →
220       ∃∃I1,K1. L1 = K1.ⓘ[I1].
221 /2 width=5 by rex_fwd_dx/ qed-.
222 *)