]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/syntax/lveq_voids.ma
\lambda\delta web site update for git
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / lveq_voids.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/xoa2.ma".
16 include "basic_2/syntax/voids_length.ma".
17 include "basic_2/syntax/lveq.ma".
18
19 (* EQUIVALENCE FOR LOCAL ENVIRONMENTS UP TO EXCLUSION BINDERS ***************)
20
21 (* Inversion lemmas with extension with exclusion binders *******************)
22
23 lemma lveq_inv_voids: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
24                       ∨∨ ∧∧ ⓧ*[n1]⋆ = L1 & ⓧ*[n2]⋆ = L2
25                        | ∃∃I1,I2,K1,K2,V1,n. K1 ≋ⓧ*[n, n] K2 & ⓧ*[n1](K1.ⓑ{I1}V1) = L1 & ⓧ*[n2](K2.ⓘ{I2}) = L2
26                        | ∃∃I1,I2,K1,K2,V2,n. K1 ≋ⓧ*[n, n] K2 & ⓧ*[n1](K1.ⓘ{I1}) = L1 & ⓧ*[n2](K2.ⓑ{I2}V2) = L2.
27 #L1 #L2 #n1 #n2 #H elim H -L1 -L2 -n1 -n2
28 [ /3 width=1 by conj, or3_intro0/
29 |2,3: /3 width=9 by or3_intro1, or3_intro2, ex3_6_intro/
30 |4,5: #K1 #K2 #n1 #n2 #HK12 * *
31    /3 width=9 by conj, or3_intro0, or3_intro1, or3_intro2, ex3_6_intro/
32 ]
33 qed-.
34
35 (* Eliminators with extension with exclusion binders ************************)
36
37 lemma lveq_ind_voids: ∀R:bi_relation lenv nat. (
38                          ∀n1,n2. R (ⓧ*[n1]⋆) n1 (ⓧ*[n2]⋆) n2
39                       ) → (
40                          ∀I1,I2,K1,K2,V1,n1,n2,n. K1 ≋ⓧ*[n, n] K2 → R K1 n K2 n →
41                          R (ⓧ*[n1]K1.ⓑ{I1}V1) n1 (ⓧ*[n2]K2.ⓘ{I2}) n2
42                       ) → (
43                          ∀I1,I2,K1,K2,V2,n1,n2,n. K1 ≋ⓧ*[n, n] K2 → R K1 n K2 n →
44                          R (ⓧ*[n1]K1.ⓘ{I1}) n1 (ⓧ*[n2]K2.ⓑ{I2}V2) n2
45                       ) →
46                       ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 → R L1 n1 L2 n2.
47 #R #IH1 #IH2 #IH3 #L1 #L2 @(f2_ind ?? length2 ?? L1 L2) -L1 -L2
48 #m #IH #L1 #L2 #Hm #n1 #n2 #H destruct
49 elim (lveq_inv_voids … H) -H * //
50 #I1 #I2 #K1 #K2 #V #n #HK #H1 #H2 destruct
51 /4 width=3 by lt_plus/
52 qed-.
53
54 (*
55
56 (* Properties with extension with exclusion binders *************************)
57
58 lemma lveq_voids_sn: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
59                      ∀m1. ⓧ*[m1]L1 ≋ⓧ*[m1+n1, n2] L2.
60 #L1 #L2 #n1 #n2 #HL12 #m1 elim m1 -m1 /2 width=1 by lveq_void_sn/
61 qed-.
62
63 lemma lveq_voids_dx: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
64                      ∀m2. L1 ≋ⓧ*[n1, m2+n2] ⓧ*[m2]L2.
65 #L1 #L2 #n1 #n2 #HL12 #m2 elim m2 -m2 /2 width=1 by lveq_void_dx/
66 qed-.
67
68 lemma lveq_voids: ∀L1,L2,n1,n2. L1 ≋ⓧ*[n1, n2] L2 →
69                   ∀m1,m2. ⓧ*[m1]L1 ≋ⓧ*[m1+n1, m2+n2] ⓧ*[m2]L2.
70 /3 width=1 by lveq_voids_dx, lveq_voids_sn/ qed-.
71
72 lemma lveq_voids_zero: ∀L1,L2. L1 ≋ⓧ*[0, 0] L2 →
73                        ∀n1,n2. ⓧ*[n1]L1 ≋ⓧ*[n1, n2] ⓧ*[n2]L2.
74 #L1 #L2 #HL12 #n1 #n2
75 >(plus_n_O … n1) in ⊢ (?%???); >(plus_n_O … n2) in ⊢ (???%?);
76 /2 width=1 by lveq_voids/ qed-.
77
78 (* Inversion lemmas with extension with exclusion binders *******************)
79
80 lemma lveq_inv_voids_sn: ∀L1,L2,n1,n2,m1. ⓧ*[m1]L1 ≋ⓧ*[m1+n1, n2] L2 →
81                          L1 ≋ⓧ*[n1, n2] L2.
82 #L1 #L2 #n1 #n2 #m1 elim m1 -m1 /3 width=1 by lveq_inv_void_sn/
83 qed-.
84
85 lemma lveq_inv_voids_dx: ∀L1,L2,n1,n2,m2. L1 ≋ⓧ*[n1, m2+n2] ⓧ*[m2]L2 →
86                          L1 ≋ⓧ*[n1, n2] L2.
87 #L1 #L2 #n1 #n2 #m2 elim m2 -m2 /3 width=1 by lveq_inv_void_dx/
88 qed-.
89
90 lemma lveq_inv_voids: ∀L1,L2,n1,n2,m1,m2. ⓧ*[m1]L1 ≋ⓧ*[m1+n1, m2+n2] ⓧ*[m2]L2 →
91                       L1 ≋ⓧ*[n1, n2] L2.
92 /3 width=3 by lveq_inv_voids_dx, lveq_inv_voids_sn/ qed-.
93
94 lemma lveq_inv_voids_zero: ∀L1,L2,n1,n2. ⓧ*[n1]L1 ≋ⓧ*[n1, n2] ⓧ*[n2]L2 →
95                            L1 ≋ⓧ*[0, 0] L2.
96 /2 width=3 by lveq_inv_voids/ qed-.
97
98 *)