]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/ground/lib/star.ma
e1580c46eb241c25a832c74897ace7810c02445f
[helm.git] / matita / matita / contribs / lambdadelta / ground / lib / star.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 "basics/star1.ma".
16 include "ground/lib/relations.ma".
17
18 (* TRANSITIVE CLOSURE *******************************************************)
19
20 definition CTC: ∀A:Type[0]. ∀B. (A→relation B) → (A→relation B) ≝
21                 λA,B,R,a. TC … (R a).
22
23 definition s_r_transitive: ∀A,B. relation2 (A→relation B) (B→relation A) ≝ λA,B,R1,R2.
24                            ∀L2,T1,T2. R1 L2 T1 T2 → ∀L1. R2 T1 L1 L2 → CTC … R1 L1 T1 T2.
25
26 definition s_rs_transitive: ∀A,B. relation2 (A→relation B) (B→relation A) ≝ λA,B,R1,R2.
27                             ∀L2,T1,T2. CTC … R1 L2 T1 T2 → ∀L1. R2 T1 L1 L2 → CTC … R1 L1 T1 T2.
28
29 lemma TC_strip1: ∀A,R1,R2. confluent2 A R1 R2 →
30                  ∀a0,a1. TC … R1 a0 a1 → ∀a2. R2 a0 a2 →
31                  ∃∃a. R2 a1 a & TC … R1 a2 a.
32 #A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1
33 [ #a1 #Ha01 #a2 #Ha02
34   elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3 by inj, ex2_intro/
35 | #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02
36   elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20
37   elim (HR12 … Ha1 … Ha0) -HR12 -a /4 width=5 by step, ex2_intro/
38 ]
39 qed.
40
41 lemma TC_strip2: ∀A,R1,R2. confluent2 A R1 R2 →
42                  ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a0 a1 →
43                  ∃∃a. TC … R2 a1 a & R1 a2 a.
44 #A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2
45 [ #a2 #Ha02 #a1 #Ha01
46   elim (HR12 … Ha01 … Ha02) -HR12 -a0 /3 width=3 by inj, ex2_intro/
47 | #a #a2 #_ #Ha2 #IHa0 #a1 #Ha01
48   elim (IHa0 … Ha01) -a0 #a0 #Ha10 #Ha0
49   elim (HR12 … Ha0 … Ha2) -HR12 -a /4 width=3 by step, ex2_intro/
50 ]
51 qed.
52
53 lemma TC_confluent2: ∀A,R1,R2.
54                      confluent2 A R1 R2 → confluent2 A (TC … R1) (TC … R2).
55 #A #R1 #R2 #HR12 #a0 #a1 #H elim H -a1
56 [ #a1 #Ha01 #a2 #Ha02
57   elim (TC_strip2 … HR12 … Ha02 … Ha01) -HR12 -a0 /3 width=3 by inj, ex2_intro/
58 | #a #a1 #_ #Ha1 #IHa0 #a2 #Ha02
59   elim (IHa0 … Ha02) -a0 #a0 #Ha0 #Ha20
60   elim (TC_strip2 … HR12 … Ha0 … Ha1) -HR12 -a /4 width=5 by step, ex2_intro/
61 ]
62 qed.
63
64 lemma TC_strap1: ∀A,R1,R2. transitive2 A R1 R2 →
65                  ∀a1,a0. TC … R1 a1 a0 → ∀a2. R2 a0 a2 →
66                  ∃∃a. R2 a1 a & TC … R1 a a2.
67 #A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0
68 [ #a0 #Ha10 #a2 #Ha02
69   elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3 by inj, ex2_intro/
70 | #a #a0 #_ #Ha0 #IHa #a2 #Ha02
71   elim (HR12 … Ha0 … Ha02) -HR12 -a0 #a0 #Ha0 #Ha02
72   elim (IHa … Ha0) -a /4 width=5 by step, ex2_intro/
73 ]
74 qed.
75
76 lemma TC_strap2: ∀A,R1,R2. transitive2 A R1 R2 →
77                  ∀a0,a2. TC … R2 a0 a2 → ∀a1. R1 a1 a0 →
78                  ∃∃a. TC … R2 a1 a & R1 a a2.
79 #A #R1 #R2 #HR12 #a0 #a2 #H elim H -a2
80 [ #a2 #Ha02 #a1 #Ha10
81   elim (HR12 … Ha10 … Ha02) -HR12 -a0 /3 width=3 by inj, ex2_intro/
82 | #a #a2 #_ #Ha02 #IHa #a1 #Ha10
83   elim (IHa … Ha10) -a0 #a0 #Ha10 #Ha0
84   elim (HR12 … Ha0 … Ha02) -HR12 -a /4 width=3 by step, ex2_intro/
85 ]
86 qed.
87
88 lemma TC_transitive2: ∀A,R1,R2.
89                       transitive2 A R1 R2 → transitive2 A (TC … R1) (TC … R2).
90 #A #R1 #R2 #HR12 #a1 #a0 #H elim H -a0
91 [ #a0 #Ha10 #a2 #Ha02
92   elim (TC_strap2 … HR12 … Ha02 … Ha10) -HR12 -a0 /3 width=3 by inj, ex2_intro/
93 | #a #a0 #_ #Ha0 #IHa #a2 #Ha02
94   elim (TC_strap2 … HR12 … Ha02 … Ha0) -HR12 -a0 #a0 #Ha0 #Ha02
95   elim (IHa … Ha0) -a /4 width=5 by step, ex2_intro/
96 ]
97 qed.
98
99 lemma CTC_lsub_trans: ∀A,B,R,S. lsub_trans A B R S → lsub_trans A B (CTC … R) S.
100 #A #B #R #S #HRS #L2 #T1 #T2 #H elim H -T2 /3 width=3 by inj/
101 #T #T2 #_ #HT2 #IHT1 #L1 #HL12
102 lapply (HRS … HT2 … HL12) -HRS -HT2 /3 width=3 by step/
103 qed-.
104
105 lemma s_r_conf1_CTC1: ∀A,B,S,R. s_r_confluent1 A B S R → s_r_confluent1 A B (CTC … S) R.
106 #A #B #S #R #HSR #L1 #T1 #T2 #H @(TC_ind_dx … T1 H) -T1 /3 width=3 by/
107 qed-.
108
109 lemma s_r_trans_CTC1: ∀A,B,S,R. s_r_confluent1 A B S R →
110                       s_r_transitive A B S R → s_rs_transitive A B S R.
111 #A #B #S #R #H1SR #H2SR #L2 #T1 #T2 #H @(TC_ind_dx … T1 H) -T1 /2 width=3 by/
112 #T1 #T #HT1 #_ #IHT2 #L1 #HL12 lapply (H2SR … HT1 … HL12) -H2SR -HT1
113 /4 width=5 by s_r_conf1_CTC1, trans_TC/
114 qed-.
115
116 lemma s_r_trans_CTC2: ∀A,B,S,R. s_rs_transitive A B S R → s_r_transitive A B S (CTC … R).
117 #A #B #S #R #HSR #L2 #T1 #T2 #HT12 #L1 #H @(TC_ind_dx … L1 H) -L1 /3 width=3 by inj/
118 qed-.
119
120 lemma s_r_to_s_rs_trans: ∀A,B,S,R. s_r_transitive A B (CTC … S) R →
121                          s_rs_transitive A B S R.
122 #A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
123 elim (TC_idem … (S L1) …  T1 T2)
124 #_ #H @H @HSR //
125 qed-.
126
127 lemma s_rs_to_s_r_trans: ∀A,B,S,R. s_rs_transitive A B S R →
128                          s_r_transitive A B (CTC … S) R.
129 #A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
130 elim (TC_idem … (S L1) …  T1 T2)
131 #H #_ @H @HSR //
132 qed-.
133
134 lemma s_rs_trans_TC1: ∀A,B,S,R. s_rs_transitive A B S R →
135                       s_rs_transitive A B (CTC … S) R.
136 #A #B #S #R #HSR #L2 #T1 #T2 #HL2 #L1 #HT1
137 elim (TC_idem … (S L1) …  T1 T2)
138 elim (TC_idem … (S L2) …  T1 T2)
139 #_ #H1 #H2 #_ @H2 @HSR /3 width=3 by/
140 qed-.
141
142 (* Normal form and strong normalization *************************************)
143
144 lemma SN_to_NF: ∀A,R,S. NF_dec A R S →
145                 ∀a1. SN A R S a1 →
146                 ∃∃a2. star … R a1 a2 & NF A R S a2.
147 #A #R #S #HRS #a1 #H elim H -a1
148 #a1 #_ #IHa1 elim (HRS a1) -HRS /2 width=3 by srefl, ex2_intro/
149 * #a0 #Ha10 #Ha01 elim (IHa1 … Ha10 Ha01) -IHa1 -Ha01 /3 width=3 by star_compl, ex2_intro/
150 qed-.
151
152 (* Relations on unboxed pairs ***********************************************)
153
154 lemma bi_TC_strip: ∀A,B,R. bi_confluent A B R →
155                    ∀a0,a1,b0,b1. R a0 b0 a1 b1 → ∀a2,b2. bi_TC … R a0 b0 a2 b2 →
156                    ∃∃a,b. bi_TC … R a1 b1 a b & R a2 b2 a b.
157 #A #B #R #HR #a0 #a1 #b0 #b1 #H01 #a2 #b2 #H elim H -a2 -b2
158 [ #a2 #b2 #H02
159   elim (HR … H01 … H02) -HR -a0 -b0 /3 width=4 by ex2_2_intro, bi_inj/
160 | #a2 #b2 #a3 #b3 #_ #H23 * #a #b #H1 #H2
161   elim (HR … H23 … H2) -HR -a0 -b0 -a2 -b2 /3 width=4 by ex2_2_intro, bi_step/
162 ]
163 qed.
164
165 lemma bi_TC_confluent: ∀A,B,R. bi_confluent A B R →
166                        bi_confluent A B (bi_TC … R).
167 #A #B #R #HR #a0 #a1 #b0 #b1 #H elim H -a1 -b1
168 [ #a1 #b1 #H01 #a2 #b2 #H02
169   elim (bi_TC_strip … HR … H01 … H02) -a0 -b0 /3 width=4 by ex2_2_intro, bi_inj/
170 | #a1 #b1 #a3 #b3 #_ #H13 #IH #a2 #b2 #H02
171   elim (IH … H02) -a0 -b0 #a0 #b0 #H10 #H20
172   elim (bi_TC_strip … HR … H13 … H10) -a1 -b1 /3 width=7 by ex2_2_intro, bi_step/
173 ]
174 qed.
175
176 lemma bi_TC_decomp_r: ∀A,B. ∀R:bi_relation A B.
177                       ∀a1,a2,b1,b2. bi_TC … R a1 b1 a2 b2 →
178                       R a1 b1 a2 b2 ∨
179                       ∃∃a,b. bi_TC … R a1 b1 a b & R a b a2 b2.
180 #A #B #R #a1 #a2 #b1 #b2 * -a2 -b2 /2 width=1/ /3 width=4 by ex2_2_intro, or_intror/
181 qed-.
182
183 lemma bi_TC_decomp_l: ∀A,B. ∀R:bi_relation A B.
184                       ∀a1,a2,b1,b2. bi_TC … R a1 b1 a2 b2 →
185                       R a1 b1 a2 b2 ∨
186                       ∃∃a,b. R a1 b1 a b & bi_TC … R a b a2 b2.
187 #A #B #R #a1 #a2 #b1 #b2 #H @(bi_TC_ind_dx … a1 b1 H) -a1 -b1
188 [ /2 width=1 by or_introl/
189 | #a1 #a #b1 #b #Hab1 #Hab2 #_ /3 width=4 by ex2_2_intro, or_intror/ (**) (* auto fails without #_ *)
190 ]
191 qed-.
192
193 (* Relations on unboxed triples *********************************************)
194
195 definition tri_star: ∀A,B,C,R. tri_relation A B C ≝
196                      λA,B,C,R. tri_RC A B C (tri_TC … R).
197
198 lemma tri_star_tri_reflexive: ∀A,B,C,R. tri_reflexive A B C (tri_star … R).
199 /2 width=1 by/ qed.
200
201 lemma tri_TC_to_tri_star: ∀A,B,C,R,a1,b1,c1,a2,b2,c2.
202                           tri_TC A B C R a1 b1 c1 a2 b2 c2 →
203                           tri_star A B C R a1 b1 c1 a2 b2 c2.
204 /2 width=1 by or_introl/ qed.
205
206 lemma tri_R_to_tri_star: ∀A,B,C,R,a1,b1,c1,a2,b2,c2.
207                          R a1 b1 c1 a2 b2 c2 → tri_star A B C R a1 b1 c1 a2 b2 c2.
208 /3 width=1 by tri_TC_to_tri_star, tri_inj/ qed.
209
210 lemma tri_star_strap1: ∀A,B,C,R,a1,a,a2,b1,b,b2,c1,c,c2.
211                        tri_star A B C R a1 b1 c1 a b c →
212                        R a b c a2 b2 c2 → tri_star A B C R a1 b1 c1 a2 b2 c2.
213 #A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 *
214 [ /3 width=5 by tri_TC_to_tri_star, tri_step/
215 | * #H1 #H2 #H3 destruct /2 width=1 by tri_R_to_tri_star/
216 ]
217 qed.
218
219 lemma tri_star_strap2: ∀A,B,C,R,a1,a,a2,b1,b,b2,c1,c,c2. R a1 b1 c1 a b c →
220                        tri_star A B C R a b c a2 b2 c2 →
221                        tri_star A B C R a1 b1 c1 a2 b2 c2.
222 #A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #H *
223 [ /3 width=5 by tri_TC_to_tri_star, tri_TC_strap/
224 | * #H1 #H2 #H3 destruct /2 width=1 by tri_R_to_tri_star/
225 ]
226 qed.
227
228 lemma tri_star_to_tri_TC_to_tri_TC: ∀A,B,C,R,a1,a,a2,b1,b,b2,c1,c,c2.
229                                     tri_star A B C R a1 b1 c1 a b c →
230                                     tri_TC A B C R a b c a2 b2 c2 →
231                                     tri_TC A B C R a1 b1 c1 a2 b2 c2.
232 #A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 *
233 [ /2 width=5 by tri_TC_transitive/
234 | * #H1 #H2 #H3 destruct /2 width=1 by/
235 ]
236 qed.
237
238 lemma tri_TC_to_tri_star_to_tri_TC: ∀A,B,C,R,a1,a,a2,b1,b,b2,c1,c,c2.
239                                     tri_TC A B C R a1 b1 c1 a b c →
240                                     tri_star A B C R a b c a2 b2 c2 →
241                                     tri_TC A B C R a1 b1 c1 a2 b2 c2.
242 #A #B #C #R #a1 #a #a2 #b1 #b #b2 #c1 #c #c2 #H *
243 [ /2 width=5 by tri_TC_transitive/
244 | * #H1 #H2 #H3 destruct /2 width=1 by/
245 ]
246 qed.
247
248 lemma tri_tansitive_tri_star: ∀A,B,C,R. tri_transitive A B C (tri_star … R).
249 #A #B #C #R #a1 #a #b1 #b #c1 #c #H #a2 #b2 #c2 *
250 [ /3 width=5 by tri_star_to_tri_TC_to_tri_TC, tri_TC_to_tri_star/
251 | * #H1 #H2 #H3 destruct /2 width=1 by/
252 ]
253 qed.
254
255 lemma tri_star_ind: ∀A,B,C,R,a1,b1,c1. ∀P:relation3 A B C. P a1 b1 c1 →
256                     (∀a,a2,b,b2,c,c2. tri_star … R a1 b1 c1 a b c → R a b c a2 b2 c2 → P a b c → P a2 b2 c2) →
257                     ∀a2,b2,c2. tri_star … R a1 b1 c1 a2 b2 c2 → P a2 b2 c2.
258 #A #B #C #R #a1 #b1 #c1 #P #H #IH #a2 #b2 #c2 *
259 [ #H12 elim H12 -a2 -b2 -c2 /3 width=6 by tri_TC_to_tri_star/
260 | * #H1 #H2 #H3 destruct //
261 ]
262 qed-.
263
264 lemma tri_star_ind_dx: ∀A,B,C,R,a2,b2,c2. ∀P:relation3 A B C. P a2 b2 c2 →
265                        (∀a1,a,b1,b,c1,c. R a1 b1 c1 a b c → tri_star … R a b c a2 b2 c2 → P a b c → P a1 b1 c1) →
266                        ∀a1,b1,c1. tri_star … R a1 b1 c1 a2 b2 c2 → P a1 b1 c1.
267 #A #B #C #R #a2 #b2 #c2 #P #H #IH #a1 #b1 #c1 *
268 [ #H12 @(tri_TC_ind_dx … a1 b1 c1 H12) -a1 -b1 -c1 /3 width=6 by tri_TC_to_tri_star/
269 | * #H1 #H2 #H3 destruct //
270 ]
271 qed-.