]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/syntax/voids.ma
766eefd5c2ccfa8c1027b62cfb4b7ded9bffc94e
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / syntax / 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 "basic_2/notation/functions/voidstar_2.ma".
16 include "basic_2/syntax/lenv.ma".
17
18 (* EXTENSION OF A LOCAL ENVIRONMENT WITH EXCLUSION BINDERS ******************)
19
20 rec definition voids (L:lenv) (n:nat) on n: lenv ≝ match n with
21 [ O ⇒ L | S m ⇒ (voids L m).ⓧ ].
22
23 interpretation "extension with exclusion binders (local environment)"
24    'VoidStar n L = (voids L n).
25
26 (* Basic properties *********************************************************)
27
28 lemma voids_zero: ∀L. L = ⓧ*[0]L.
29 // qed.
30
31 lemma voids_succ: ∀L,n. (ⓧ*[n]L).ⓧ = ⓧ*[⫯n]L.
32 // qed.
33
34 (* Advanced properties ******************************************************)
35
36 lemma voids_next: ∀n,L. ⓧ*[n](L.ⓧ) = ⓧ*[⫯n]L.
37 #n elim n -n //
38 qed.
39
40 (* Main inversion properties ************************************************)
41
42 theorem voids_inj: ∀n. injective … (λL. ⓧ*[n]L).
43 #n elim n -n //
44 #n #IH #L1 #L2
45 <voids_succ <voids_succ #H
46 elim (destruct_lbind_lbind_aux … H) -H (**) (* destruct lemma needed *)
47 /2 width=1 by/
48 qed-.