]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/static/reqx.ma
update in ground_2, static_2, basic_2, apps_2, alpha_1
[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/stareqsn_3.ma".
16 include "static_2/syntax/teqx_ext.ma".
17 include "static_2/static/rex.ma".
18
19 (* SORT-IRRELEVANT EQUIVALENCE FOR LOCAL ENVIRONMENTS ON REFERRED ENTRIES ***)
20
21 definition reqx: relation3 term lenv lenv โ‰
22                  rex cdeq.
23
24 interpretation
25    "sort-irrelevant equivalence on referred entries (local environment)"
26    'StarEqSn T L1 L2 = (reqx T L1 L2).
27
28 interpretation
29    "sort-irrelevant ranged equivalence (local environment)"
30    'StarEqSn f L1 L2 = (sex cdeq_ext cfull f L1 L2).
31
32 (* Basic properties ***********************************************************)
33
34 lemma frees_teqx_conf_reqx: โˆ€f,L1,T1. L1 โŠข ๐…+โชT1โซ โ‰˜ f โ†’ โˆ€T2. T1 โ‰› T2 โ†’
35                             โˆ€L2. L1 โ‰›[f] L2 โ†’ L2 โŠข ๐…+โชT2โซ โ‰˜ f.
36 #f #L1 #T1 #H elim H -f -L1 -T1
37 [ #f #L1 #s1 #Hf #X #H1 #L2 #_
38   elim (teqx_inv_sort1 โ€ฆ H1) -H1 #s2 #H destruct
39   /2 width=3 by frees_sort/
40 | #f #i #Hf #X #H1
41   >(teqx_inv_lref1 โ€ฆ H1) -X #Y #H2
42   >(sex_inv_atom1 โ€ฆ H2) -Y
43   /2 width=1 by frees_atom/
44 | #f #I #L1 #V1 #_ #IH #X #H1
45   >(teqx_inv_lref1 โ€ฆ H1) -X #Y #H2
46   elim (sex_inv_next1 โ€ฆ H2) -H2 #Z #L2 #HL12 #HZ #H destruct
47   elim (ext2_inv_pair_sn โ€ฆ HZ) -HZ #V2 #HV12 #H destruct
48   /3 width=1 by frees_pair/
49 | #f #I #L1 #Hf #X #H1
50   >(teqx_inv_lref1 โ€ฆ H1) -X #Y #H2
51   elim (sex_inv_next1 โ€ฆ H2) -H2 #Z #L2 #_ #HZ #H destruct
52   >(ext2_inv_unit_sn โ€ฆ HZ) -Z /2 width=1 by frees_unit/
53 | #f #I #L1 #i #_ #IH #X #H1
54   >(teqx_inv_lref1 โ€ฆ H1) -X #Y #H2
55   elim (sex_inv_push1 โ€ฆ H2) -H2 #J #L2 #HL12 #_ #H destruct
56   /3 width=1 by frees_lref/
57 | #f #L1 #l #Hf #X #H1 #L2 #_
58   >(teqx_inv_gref1 โ€ฆ H1) -X /2 width=1 by frees_gref/
59 | #f1V #f1T #f1 #p #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
60   elim (teqx_inv_pair1 โ€ฆ H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
61   /6 width=5 by frees_bind, sex_inv_tl, ext2_pair, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/
62 | #f1V #f1T #f1 #I #L1 #V1 #T1 #_ #_ #Hf1 #IHV #IHT #X #H1
63   elim (teqx_inv_pair1 โ€ฆ H1) -H1 #V2 #T2 #HV12 #HT12 #H1 #L2 #HL12 destruct
64   /5 width=5 by frees_flat, sle_sex_trans, sor_inv_sle_dx, sor_inv_sle_sn/
65 ]
66 qed-.
67
68 lemma frees_teqx_conf: โˆ€f,L,T1. L โŠข ๐…+โชT1โซ โ‰˜ f โ†’
69                        โˆ€T2. T1 โ‰› T2 โ†’ L โŠข ๐…+โชT2โซ โ‰˜ f.
70 /4 width=7 by frees_teqx_conf_reqx, sex_refl, ext2_refl/ qed-.
71
72 lemma frees_reqx_conf: โˆ€f,L1,T. L1 โŠข ๐…+โชTโซ โ‰˜ f โ†’
73                        โˆ€L2. L1 โ‰›[f] L2 โ†’ L2 โŠข ๐…+โชTโซ โ‰˜ f.
74 /2 width=7 by frees_teqx_conf_reqx, teqx_refl/ qed-.
75
76 lemma teqx_rex_conf (R): s_r_confluent1 โ€ฆ cdeq (rex R).
77 #R #L1 #T1 #T2 #HT12 #L2 *
78 /3 width=5 by frees_teqx_conf, ex2_intro/
79 qed-.
80
81 lemma teqx_rex_div (R): โˆ€T1,T2. T1 โ‰› T2 โ†’
82                         โˆ€L1,L2. L1 โชค[R,T2] L2 โ†’ L1 โชค[R,T1] L2.
83 /3 width=5 by teqx_rex_conf, teqx_sym/ qed-.
84
85 lemma teqx_reqx_conf: s_r_confluent1 โ€ฆ cdeq reqx.
86 /2 width=5 by teqx_rex_conf/ qed-.
87
88 lemma teqx_reqx_div: โˆ€T1,T2. T1 โ‰› T2 โ†’
89                      โˆ€L1,L2. L1 โ‰›[T2] L2 โ†’ L1 โ‰›[T1] L2.
90 /2 width=5 by teqx_rex_div/ qed-.
91
92 lemma reqx_atom: โˆ€I. โ‹† โ‰›[โ“ช[I]] โ‹†.
93 /2 width=1 by rex_atom/ qed.
94
95 lemma reqx_sort: โˆ€I1,I2,L1,L2,s.
96                  L1 โ‰›[โ‹†s] L2 โ†’ L1.โ“˜[I1] โ‰›[โ‹†s] L2.โ“˜[I2].
97 /2 width=1 by rex_sort/ qed.
98
99 lemma reqx_pair: โˆ€I,L1,L2,V1,V2.
100                  L1 โ‰›[V1] L2 โ†’ V1 โ‰› V2 โ†’ L1.โ“‘[I]V1 โ‰›[#0] L2.โ“‘[I]V2.
101 /2 width=1 by rex_pair/ qed.
102
103 lemma reqx_unit: โˆ€f,I,L1,L2. ๐ˆโชfโซ โ†’ L1 โ‰›[f] L2 โ†’
104                  L1.โ“ค[I] โ‰›[#0] L2.โ“ค[I].
105 /2 width=3 by rex_unit/ qed.
106
107 lemma reqx_lref: โˆ€I1,I2,L1,L2,i.
108                  L1 โ‰›[#i] L2 โ†’ L1.โ“˜[I1] โ‰›[#โ†‘i] L2.โ“˜[I2].
109 /2 width=1 by rex_lref/ qed.
110
111 lemma reqx_gref: โˆ€I1,I2,L1,L2,l.
112                  L1 โ‰›[ยงl] L2 โ†’ L1.โ“˜[I1] โ‰›[ยงl] L2.โ“˜[I2].
113 /2 width=1 by rex_gref/ qed.
114
115 lemma reqx_bind_repl_dx: โˆ€I,I1,L1,L2.โˆ€T:term.
116                          L1.โ“˜[I] โ‰›[T] L2.โ“˜[I1] โ†’
117                          โˆ€I2. I โ‰› I2 โ†’
118                          L1.โ“˜[I] โ‰›[T] L2.โ“˜[I2].
119 /2 width=2 by rex_bind_repl_dx/ qed-.
120
121 (* Basic inversion lemmas ***************************************************)
122
123 lemma reqx_inv_atom_sn: โˆ€Y2. โˆ€T:term. โ‹† โ‰›[T] Y2 โ†’ Y2 = โ‹†.
124 /2 width=3 by rex_inv_atom_sn/ qed-.
125
126 lemma reqx_inv_atom_dx: โˆ€Y1. โˆ€T:term. Y1 โ‰›[T] โ‹† โ†’ Y1 = โ‹†.
127 /2 width=3 by rex_inv_atom_dx/ qed-.
128
129 lemma reqx_inv_zero:
130       โˆ€Y1,Y2. Y1 โ‰›[#0] Y2 โ†’
131       โˆจโˆจ โˆงโˆง Y1 = โ‹† & Y2 = โ‹†
132        | โˆƒโˆƒI,L1,L2,V1,V2. L1 โ‰›[V1] L2 & V1 โ‰› V2 & Y1 = L1.โ“‘[I]V1 & Y2 = L2.โ“‘[I]V2
133        | โˆƒโˆƒf,I,L1,L2. ๐ˆโชfโซ & L1 โ‰›[f] L2 & Y1 = L1.โ“ค[I] & Y2 = L2.โ“ค[I].
134 #Y1 #Y2 #H elim (rex_inv_zero โ€ฆ H) -H *
135 /3 width=9 by or3_intro0, or3_intro1, or3_intro2, ex4_5_intro, ex4_4_intro, conj/
136 qed-.
137
138 lemma reqx_inv_lref: โˆ€Y1,Y2,i. Y1 โ‰›[#โ†‘i] Y2 โ†’
139                      โˆจโˆจ โˆงโˆง Y1 = โ‹† & Y2 = โ‹†
140                       | โˆƒโˆƒI1,I2,L1,L2. L1 โ‰›[#i] L2 &
141                                        Y1 = L1.โ“˜[I1] & Y2 = L2.โ“˜[I2].
142 /2 width=1 by rex_inv_lref/ qed-.
143
144 (* Basic_2A1: uses: lleq_inv_bind lleq_inv_bind_O *)
145 lemma reqx_inv_bind: โˆ€p,I,L1,L2,V,T. L1 โ‰›[โ“‘[p,I]V.T] L2 โ†’
146                      โˆงโˆง L1 โ‰›[V] L2 & L1.โ“‘[I]V โ‰›[T] L2.โ“‘[I]V.
147 /2 width=2 by rex_inv_bind/ qed-.
148
149 (* Basic_2A1: uses: lleq_inv_flat *)
150 lemma reqx_inv_flat: โˆ€I,L1,L2,V,T. L1 โ‰›[โ“•[I]V.T] L2 โ†’
151                      โˆงโˆง L1 โ‰›[V] L2 & L1 โ‰›[T] L2.
152 /2 width=2 by rex_inv_flat/ qed-.
153
154 (* Advanced inversion lemmas ************************************************)
155
156 lemma reqx_inv_zero_pair_sn: โˆ€I,Y2,L1,V1. L1.โ“‘[I]V1 โ‰›[#0] Y2 โ†’
157                              โˆƒโˆƒL2,V2. L1 โ‰›[V1] L2 & V1 โ‰› V2 & Y2 = L2.โ“‘[I]V2.
158 /2 width=1 by rex_inv_zero_pair_sn/ qed-.
159
160 lemma reqx_inv_zero_pair_dx: โˆ€I,Y1,L2,V2. Y1 โ‰›[#0] L2.โ“‘[I]V2 โ†’
161                              โˆƒโˆƒL1,V1. L1 โ‰›[V1] L2 & V1 โ‰› V2 & Y1 = L1.โ“‘[I]V1.
162 /2 width=1 by rex_inv_zero_pair_dx/ qed-.
163
164 lemma reqx_inv_lref_bind_sn: โˆ€I1,Y2,L1,i. L1.โ“˜[I1] โ‰›[#โ†‘i] Y2 โ†’
165                              โˆƒโˆƒI2,L2. L1 โ‰›[#i] L2 & Y2 = L2.โ“˜[I2].
166 /2 width=2 by rex_inv_lref_bind_sn/ qed-.
167
168 lemma reqx_inv_lref_bind_dx: โˆ€I2,Y1,L2,i. Y1 โ‰›[#โ†‘i] L2.โ“˜[I2] โ†’
169                              โˆƒโˆƒI1,L1. L1 โ‰›[#i] L2 & Y1 = L1.โ“˜[I1].
170 /2 width=2 by rex_inv_lref_bind_dx/ qed-.
171
172 (* Basic forward lemmas *****************************************************)
173
174 lemma reqx_fwd_zero_pair: โˆ€I,K1,K2,V1,V2.
175                           K1.โ“‘[I]V1 โ‰›[#0] K2.โ“‘[I]V2 โ†’ K1 โ‰›[V1] K2.
176 /2 width=3 by rex_fwd_zero_pair/ qed-.
177
178 (* Basic_2A1: uses: lleq_fwd_bind_sn lleq_fwd_flat_sn *)
179 lemma reqx_fwd_pair_sn: โˆ€I,L1,L2,V,T. L1 โ‰›[โ‘ก[I]V.T] L2 โ†’ L1 โ‰›[V] L2.
180 /2 width=3 by rex_fwd_pair_sn/ qed-.
181
182 (* Basic_2A1: uses: lleq_fwd_bind_dx lleq_fwd_bind_O_dx *)
183 lemma reqx_fwd_bind_dx: โˆ€p,I,L1,L2,V,T.
184                         L1 โ‰›[โ“‘[p,I]V.T] L2 โ†’ L1.โ“‘[I]V โ‰›[T] L2.โ“‘[I]V.
185 /2 width=2 by rex_fwd_bind_dx/ qed-.
186
187 (* Basic_2A1: uses: lleq_fwd_flat_dx *)
188 lemma reqx_fwd_flat_dx: โˆ€I,L1,L2,V,T. L1 โ‰›[โ“•[I]V.T] L2 โ†’ L1 โ‰›[T] L2.
189 /2 width=3 by rex_fwd_flat_dx/ qed-.
190
191 lemma reqx_fwd_dx: โˆ€I2,L1,K2. โˆ€T:term. L1 โ‰›[T] K2.โ“˜[I2] โ†’
192                    โˆƒโˆƒI1,K1. L1 = K1.โ“˜[I1].
193 /2 width=5 by rex_fwd_dx/ qed-.