]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/subterms/booleanized.ma
decentralized notation in lambda
[helm.git] / matita / matita / lib / lambda / subterms / booleanized.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 "lambda/subterms/boolean.ma".
16
17 include "lambda/notation/functions/projectsame_2.ma".
18
19 include "lambda/xoa/ex_4_3.ma".
20
21 (* BOOLEANIZED SUBSET (EMPTY OR FULL) ***************************************)
22
23 definition booleanized: bool → subterms → subterms ≝
24    λb,F. {b}⇑⇓F.
25
26 interpretation "booleanized (subterms)"
27    'ProjectSame b F = (booleanized b F).
28
29 lemma booleanized_inv_vref: ∀j,c,b,F. {b}⇕ F = {c}#j →
30                             ∃∃b1. b = c & F = {b1}#j.
31 #j #c #b #F #H
32 elim (boolean_inv_vref … H) -H #H0 #H
33 elim (carrier_inv_vref … H) -H /2 width=2/
34 qed-.
35
36 lemma booleanized_inv_abst: ∀U,c,b,F. {b}⇕ F = {c}𝛌.U →
37                             ∃∃b1,T. b = c & {b}⇕T = U & F = {b1}𝛌.T.
38 #U #c #b #F #H
39 elim (boolean_inv_abst … H) -H #C #H0 #H1 #H
40 elim (carrier_inv_abst … H) -H #b1 #U1 #H3 destruct /2 width=4/
41 qed-.
42
43 lemma booleanized_inv_appl: ∀W,U,c,b,F. {b}⇕ F = {c}@W.U →
44                             ∃∃b1,V,T. b = c & {b}⇕V = W & {b}⇕T = U & F = {b1}@V.T.
45 #W #U #c #b #F #H
46 elim (boolean_inv_appl … H) -H #D #C #H0 #H1 #H2 #H
47 elim (carrier_inv_appl … H) -H #b1 #W1 #U1 #H3 #H4 destruct /2 width=6/
48 qed-.
49
50 lemma booleanized_booleanized: ∀c,b,F. {b}⇕ {c}⇕ F = {b}⇕ F.
51 normalize //
52 qed.
53
54 lemma booleanized_lift: ∀b,h,F,d. {b}⇕ ↑[d, h] F = ↑[d, h] {b}⇕ F.
55 normalize //
56 qed.
57
58 lemma booleanized_dsubst: ∀b,G,F,d. {b}⇕ [d ↙ G] F = [d ↙ {b}⇕ G] {b}⇕ F.
59 normalize //
60 qed.