]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/lambda/subterms/carrier.ma
decentralized notation in lambda
[helm.git] / matita / matita / lib / lambda / subterms / carrier.ma
index 8936792e2eec1c0ab2ace0ba056d9a406b77156e..3f4c0499181fa17b33cc40284fb184772c443677 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "terms/relocating_substitution.ma".
-include "subterms/relocating_substitution.ma".
+include "lambda/terms/relocating_substitution.ma".
+include "lambda/subterms/relocating_substitution.ma".
+
+include "lambda/notation/functions/projectdown_1.ma".
+
+include "lambda/xoa/ex_2_2.ma".
 
 (* CARRIER (UNDERLYING TERM) ************************************************)
 
@@ -26,10 +30,6 @@ let rec carrier F on F ≝ match F with
 interpretation "carrier (subterms)"
    'ProjectDown F = (carrier F).
 
-notation "hvbox(⇓ term 46 F)"
-   non associative with precedence 46
-   for @{ 'ProjectDown $F }.
-
 lemma carrier_inv_vref: ∀j,F. ⇓F = #j → ∃b. F = {b}#j.
 #j * normalize
 [ #b #i #H destruct /2 width=2/