]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma
syntactic components detached from basic_2 become static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / syntax / append_length.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/syntax/lenv_length.ma".
16 include "static_2/syntax/append.ma".
17
18 (* APPEND FOR LOCAL ENVIRONMENTS ********************************************)
19
20 (* Properties with length for local environments ****************************)
21
22 lemma append_length: ∀L1,L2. |L1 + L2| = |L1| + |L2|.
23 #L1 #L2 elim L2 -L2 //
24 #L2 #I >append_bind >length_bind >length_bind //
25 qed.
26
27 lemma ltail_length: ∀I,L. |ⓘ{I}.L| = ↑|L|.
28 #I #L >append_length //
29 qed.
30
31 (* Advanced inversion lemmas on length for local environments ***************)
32
33 (* Basic_2A1: was: length_inv_pos_dx_ltail *)
34 lemma length_inv_succ_dx_ltail: ∀L,n. |L| = ↑n →
35                                 ∃∃I,K. |K| = n & L = ⓘ{I}.K.
36 #Y #n #H elim (length_inv_succ_dx … H) -H #I #L #Hn #HLK destruct
37 elim (lenv_case_tail … L) [2: * #K #J ]
38 #H destruct /2 width=4 by ex2_2_intro/
39 qed-.
40
41 (* Basic_2A1: was: length_inv_pos_sn_ltail *)
42 lemma length_inv_succ_sn_ltail: ∀L,n. ↑n = |L| →
43                                 ∃∃I,K. n = |K| & L = ⓘ{I}.K.
44 #Y #n #H elim (length_inv_succ_sn … H) -H #I #L #Hn #HLK destruct
45 elim (lenv_case_tail … L) [2: * #K #J ]
46 #H destruct /2 width=4 by ex2_2_intro/
47 qed-.
48
49 (* Inversion lemmas with length for local environments **********************)
50
51 (* Basic_2A1: was: append_inj_sn *)
52 lemma append_inj_length_sn: ∀K1,K2,L1,L2. L1 + K1 = L2 + K2 → |K1| = |K2| →
53                             L1 = L2 ∧ K1 = K2.
54 #K1 elim K1 -K1
55 [ * /2 width=1 by conj/
56   #K2 #I2 #L1 #L2 #_ >length_atom >length_bind
57   #H destruct
58 | #K1 #I1 #IH *
59   [ #L1 #L2 #_ >length_atom >length_bind
60     #H destruct
61   | #K2 #I2 #L1 #L2 #H1 >length_bind >length_bind #H2
62     elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *)
63     elim (IH … H1) -IH -H1 /3 width=4 by conj/
64   ]
65 ]
66 qed-.
67
68 (* Note: lemma 750 *)
69 (* Basic_2A1: was: append_inj_dx *)
70 lemma append_inj_length_dx: ∀K1,K2,L1,L2. L1 + K1 = L2 + K2 → |L1| = |L2| →
71                             L1 = L2 ∧ K1 = K2.
72 #K1 elim K1 -K1
73 [ * /2 width=1 by conj/
74   #K2 #I2 #L1 #L2 >append_atom >append_bind #H destruct
75   >length_bind >append_length >plus_n_Sm
76   #H elim (plus_xSy_x_false … H)
77 | #K1 #I1 #IH *
78   [ #L1 #L2 >append_bind >append_atom #H destruct
79     >length_bind >append_length >plus_n_Sm #H
80     lapply (discr_plus_x_xy … H) -H #H destruct
81   | #K2 #I2 #L1 #L2 >append_bind >append_bind #H1 #H2
82     elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *)
83     elim (IH … H1) -IH -H1 /2 width=1 by conj/
84   ]
85 ]
86 qed-.
87
88 (* Advanced inversion lemmas ************************************************)
89
90 lemma append_inj_dx: ∀L,K1,K2. L+K1 = L+K2 → K1 = K2.
91 #L #K1 #K2 #H elim (append_inj_length_dx … H) -H //
92 qed-.
93
94 lemma append_inv_refl_dx: ∀L,K. L+K = L → K = ⋆.
95 #L #K #H elim (append_inj_dx … (⋆) … H) //
96 qed-.
97
98 lemma append_inv_pair_dx: ∀I,L,K,V. L+K = L.ⓑ{I}V → K = ⋆.ⓑ{I}V.
99 #I #L #K #V #H elim (append_inj_dx … (⋆.ⓑ{I}V) … H) //
100 qed-.
101
102 (* Basic eliminators ********************************************************)
103
104 (* Basic_1: was: c_tail_ind *)
105 (* Basic_2A1: was: lenv_ind_alt *) 
106 lemma lenv_ind_tail: ∀Q:predicate lenv.
107                      Q (⋆) → (∀I,L. Q L → Q (ⓘ{I}.L)) → ∀L. Q L.
108 #Q #IH1 #IH2 #L @(f_ind … length … L) -L #x #IHx * //
109 #L #I -IH1 #H destruct
110 elim (lenv_case_tail … L) [2: * #K #J ]
111 #H destruct /3 width=1 by/
112 qed-.