]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda/subterms/boolean.ma
66447a3bb29b49febcbf790ef5cd1ada2d8cf3da
[helm.git] / matita / matita / lib / lambda / subterms / boolean.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/carrier.ma".
16
17 (* BOOLEAN (EMPTY OR FULL) SUBSET *******************************************)
18
19 let rec boolean b M on M ≝ match M with
20 [ VRef i   ⇒ {b}#i
21 | Abst A   ⇒ {b}𝛌.(boolean b A)
22 | Appl B A ⇒ {b}@(boolean b B).(boolean b A)
23 ].
24
25 interpretation "boolean subset (subterms)"
26    'ProjectUp b M = (boolean b M).
27
28 notation "hvbox( { term 46 b } ⇑ break term 46 M)"
29    non associative with precedence 46
30    for @{ 'ProjectUp $b $M }.
31
32 lemma boolean_inv_vref: ∀j,c,b,M. {b}⇑ M = {c}#j → b = c ∧ M = #j.
33 #j #c #b * normalize
34 [ #i #H destruct /2 width=1/
35 | #A #H destruct
36 | #B #A #H destruct
37 ]
38 qed-.
39
40 lemma boolean_inv_abst: ∀U,c,b,M. {b}⇑ M = {c}𝛌.U →
41                         ∃∃C. b = c & {b}⇑C = U & M = 𝛌.C.
42 #U #c #b * normalize
43 [ #i #H destruct
44 | #A #H destruct /2 width=3/
45 | #B #A #H destruct
46 ]
47 qed-.
48
49 lemma boolean_inv_appl: ∀W,U,c,b,M. {b}⇑ M = {c}@W.U →
50                         ∃∃D,C. b = c & {b}⇑D = W & {b}⇑C = U & M = @D.C.
51 #W #U #c #b * normalize
52 [ #i #H destruct
53 | #A #H destruct
54 | #B #A #H destruct /2 width=5/
55 ]
56 qed-.
57
58 lemma boolean_lift: ∀b,h,M,d. {b}⇑ ↑[d, h] M = ↑[d, h] {b}⇑ M.
59 #b #h #M elim M -M normalize //
60 qed.
61
62 lemma boolean_dsubst: ∀b,N,M,d. {b}⇑ [d ↙ N] M = [d ↙ {b}⇑ N] {b}⇑ M.
63 #b #N #M elim M -M [2,3: normalize // ]
64 #i #d elim (lt_or_eq_or_gt i d) #Hid
65 [ >(sdsubst_vref_lt … Hid) >(dsubst_vref_lt … Hid) //
66 | destruct normalize //
67 | >(sdsubst_vref_gt … Hid) >(dsubst_vref_gt … Hid) //
68 ]
69 qed.
70
71 lemma carrier_boolean: ∀b,M. ⇓ {b}⇑ M = M.
72 #b #M elim M -M normalize //
73 qed.