]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/lambda/subterms/boolean.ma
decentralized notation in lambda
[helm.git] / matita / matita / lib / lambda / subterms / boolean.ma
index 66447a3bb29b49febcbf790ef5cd1ada2d8cf3da..566bfb0a3f976d2fd6680eaa2f2524e6a9c40528 100644 (file)
 
 include "lambda/subterms/carrier.ma".
 
+include "lambda/notation/functions/projectup_2.ma".
+
+include "lambda/xoa/ex_3_1.ma".
+include "lambda/xoa/ex_4_2.ma".
+
 (* BOOLEAN (EMPTY OR FULL) SUBSET *******************************************)
 
 let rec boolean b M on M ≝ match M with
@@ -25,10 +30,6 @@ let rec boolean b M on M ≝ match M with
 interpretation "boolean subset (subterms)"
    'ProjectUp b M = (boolean b M).
 
-notation "hvbox( { term 46 b } ⇑ break term 46 M)"
-   non associative with precedence 46
-   for @{ 'ProjectUp $b $M }.
-
 lemma boolean_inv_vref: ∀j,c,b,M. {b}⇑ M = {c}#j → b = c ∧ M = #j.
 #j #c #b * normalize
 [ #i #H destruct /2 width=1/