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