]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/lib/lambda/paths/dst_computation.ma
decentralized notation in lambda
[helm.git] / matita / matita / lib / lambda / paths / dst_computation.ma
index 85874e21cdf6a8d430adbc6422d5d63867591446..a9cc98f0978323866af765ffbcfea83d3524a800 100644 (file)
@@ -15,6 +15,8 @@
 include "lambda/paths/standard_trace.ma".
 include "lambda/paths/labeled_sequential_computation.ma".
 
+include "lambda/notation/relations/decomposedstd_2.ma".
+
 (* DECOMPOSED STANDARD COMPUTATION ***********************************************)
 
 (* Note: this is the "standard" computation of:
@@ -29,10 +31,6 @@ inductive dst: relation term ≝
 interpretation "decomposed standard computation"
     'DecomposedStd M N = (dst M N).
 
-notation "hvbox( M break ⓢ↦* term 46 N )"
-   non associative with precedence 45
-   for @{ 'DecomposedStd $M $N }.
-
 lemma dst_inv_lref: ∀M,N. M ⓢ↦* N → ∀j. #j = N →
                     ∃∃s. is_whd s & M ↦*[s] #j.
 #M #N * -M -N