]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/substitution/lift_neg.ma
6c797befb7d0d3a3ecdecc75b11c231a66c7b5b3
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / substitution / lift_neg.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/substitution/lift.ma".
16
17 (* BASIC TERM RELOCATION ****************************************************)
18
19 (* Properties on negated basic relocation ***********************************)
20
21 lemma nlift_lref_be_SO: ∀X,i. ⬆[yinj i, 1] X ≡ #i → ⊥.
22 /3 width=7 by lift_inv_lref2_be, ylt_inj/ qed-.
23
24 lemma nlift_bind_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) →
25                      ∀a,I,U. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥).
26 #W #l #m #HW #a #I #U #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
27 qed-.
28
29 lemma nlift_bind_dx: ∀U,l,m. (∀T. ⬆[⫯l, m] T ≡ U → ⊥) →
30                      ∀a,I,W. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥).
31 #U #l #m #HU #a #I #W #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/
32 qed-.
33
34 lemma nlift_flat_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) →
35                      ∀I,U. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥).
36 #W #l #m #HW #I #U #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
37 qed-.
38
39 lemma nlift_flat_dx: ∀U,l,m. (∀T. ⬆[l, m] T ≡ U → ⊥) →
40                      ∀I,W. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥).
41 #U #l #m #HU #I #W #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/
42 qed-.
43
44 (* Inversion lemmas on negated basic relocation *****************************)
45
46 lemma nlift_inv_lref_be_SO: ∀i,j. (∀X. ⬆[i, 1] X ≡ #j → ⊥) → yinj j = i.
47 * [2: #j #H elim H -H // ]
48 #i #j elim (lt_or_eq_or_gt i j) // #Hij #H
49 [ elim (H (#(j-1))) -H /3 width=1 by lift_lref_ge_minus, yle_inj/
50 | elim (H (#j)) -H /3 width=1 by lift_lref_lt, ylt_inj/
51 ]
52 qed-.
53
54 lemma nlift_inv_bind: ∀a,I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥) →
55                       (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[⫯l, m] T ≡ U → ⊥).
56 #a #I #W #U #l #m #H elim (is_lift_dec W l m)
57 [ * /4 width=2 by lift_bind, or_intror/
58 | /4 width=2 by ex_intro, or_introl/
59 ]
60 qed-.
61
62 lemma nlift_inv_flat: ∀I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥) →
63                       (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[l, m] T ≡ U → ⊥).
64 #I #W #U #l #m #H elim (is_lift_dec W l m)
65 [ * /4 width=2 by lift_flat, or_intror/
66 | /4 width=2 by ex_intro, or_introl/
67 ]
68 qed-.