miden_crypto/merkle/store/
mod.rs

1//! Merkle store for efficiently storing multiple Merkle trees with common subtrees.
2
3use alloc::vec::Vec;
4use core::borrow::Borrow;
5
6use super::{
7    EmptySubtreeRoots, InnerNodeInfo, MerkleError, MerklePath, MerkleProof, MerkleTree, NodeIndex,
8    PartialMerkleTree, RootPath, Rpo256, Word,
9    mmr::Mmr,
10    smt::{SimpleSmt, Smt},
11};
12use crate::{
13    Map,
14    utils::{ByteReader, ByteWriter, Deserializable, DeserializationError, Serializable},
15};
16
17#[cfg(test)]
18mod tests;
19
20// MERKLE STORE
21// ================================================================================================
22
23#[derive(Debug, Default, Copy, Clone, Eq, PartialEq)]
24#[cfg_attr(feature = "serde", derive(serde::Deserialize, serde::Serialize))]
25pub struct StoreNode {
26    left: Word,
27    right: Word,
28}
29
30/// An in-memory data store for Merkelized data.
31///
32/// This is a in memory data store for Merkle trees, this store allows all the nodes of multiple
33/// trees to live as long as necessary and without duplication, this allows the implementation of
34/// space efficient persistent data structures.
35///
36/// Example usage:
37///
38/// ```rust
39/// # use miden_crypto::{ZERO, Felt, Word};
40/// # use miden_crypto::merkle::{NodeIndex, MerkleTree, store::MerkleStore};
41/// # use miden_crypto::hash::rpo::Rpo256;
42/// # use miden_crypto::field::PrimeCharacteristicRing;
43/// # const fn int_to_node(value: u64) -> Word {
44/// #     Word::new([Felt::new(value), ZERO, ZERO, ZERO])
45/// # }
46/// # let A = int_to_node(1);
47/// # let B = int_to_node(2);
48/// # let C = int_to_node(3);
49/// # let D = int_to_node(4);
50/// # let E = int_to_node(5);
51/// # let F = int_to_node(6);
52/// # let G = int_to_node(7);
53/// # let H0 = int_to_node(8);
54/// # let H1 = int_to_node(9);
55/// # let T0 = MerkleTree::new([A, B, C, D, E, F, G, H0].to_vec()).expect("even number of leaves provided");
56/// # let T1 = MerkleTree::new([A, B, C, D, E, F, G, H1].to_vec()).expect("even number of leaves provided");
57/// # let ROOT0 = T0.root();
58/// # let ROOT1 = T1.root();
59/// let mut store: MerkleStore = MerkleStore::new();
60///
61/// // the store is initialized with the SMT empty nodes
62/// assert_eq!(store.num_internal_nodes(), 255);
63///
64/// let tree1 = MerkleTree::new(vec![A, B, C, D, E, F, G, H0]).unwrap();
65/// let tree2 = MerkleTree::new(vec![A, B, C, D, E, F, G, H1]).unwrap();
66///
67/// // populates the store with two merkle trees, common nodes are shared
68/// store.extend(tree1.inner_nodes());
69/// store.extend(tree2.inner_nodes());
70///
71/// // every leaf except the last are the same
72/// for i in 0..7 {
73///     let idx0 = NodeIndex::new(3, i).unwrap();
74///     let d0 = store.get_node(ROOT0, idx0).unwrap();
75///     let idx1 = NodeIndex::new(3, i).unwrap();
76///     let d1 = store.get_node(ROOT1, idx1).unwrap();
77///     assert_eq!(d0, d1, "Both trees have the same leaf at pos {i}");
78/// }
79///
80/// // The leaves A-B-C-D are the same for both trees, so are their 2 immediate parents
81/// for i in 0..4 {
82///     let idx0 = NodeIndex::new(3, i).unwrap();
83///     let d0 = store.get_path(ROOT0, idx0).unwrap();
84///     let idx1 = NodeIndex::new(3, i).unwrap();
85///     let d1 = store.get_path(ROOT1, idx1).unwrap();
86///     assert_eq!(d0.path[0..2], d1.path[0..2], "Both sub-trees are equal up to two levels");
87/// }
88///
89/// // Common internal nodes are shared, the two added trees have a total of 30, but the store has
90/// // only 10 new entries, corresponding to the 10 unique internal nodes of these trees.
91/// assert_eq!(store.num_internal_nodes() - 255, 10);
92/// ```
93#[derive(Debug, Clone, Eq, PartialEq)]
94#[cfg_attr(feature = "serde", derive(serde::Deserialize, serde::Serialize))]
95pub struct MerkleStore {
96    nodes: Map<Word, StoreNode>,
97}
98
99impl Default for MerkleStore {
100    fn default() -> Self {
101        Self::new()
102    }
103}
104
105impl MerkleStore {
106    // CONSTRUCTORS
107    // --------------------------------------------------------------------------------------------
108
109    /// Creates an empty `MerkleStore` instance.
110    pub fn new() -> MerkleStore {
111        // pre-populate the store with the empty hashes
112        let nodes = empty_hashes().collect();
113        MerkleStore { nodes }
114    }
115
116    // PUBLIC ACCESSORS
117    // --------------------------------------------------------------------------------------------
118
119    /// Return a count of the non-leaf nodes in the store.
120    pub fn num_internal_nodes(&self) -> usize {
121        self.nodes.len()
122    }
123
124    /// Returns the node at `index` rooted on the tree `root`.
125    ///
126    /// # Errors
127    /// This method can return the following errors:
128    /// - `RootNotInStore` if the `root` is not present in the store.
129    /// - `NodeNotInStore` if a node needed to traverse from `root` to `index` is not present in the
130    ///   store.
131    pub fn get_node(&self, root: Word, index: NodeIndex) -> Result<Word, MerkleError> {
132        let mut hash = root;
133
134        // corner case: check the root is in the store when called with index `NodeIndex::root()`
135        self.nodes.get(&hash).ok_or(MerkleError::RootNotInStore(hash))?;
136
137        for i in (0..index.depth()).rev() {
138            let node = self
139                .nodes
140                .get(&hash)
141                .ok_or(MerkleError::NodeIndexNotFoundInStore(hash, index))?;
142
143            let is_right = index.is_nth_bit_odd(i);
144            hash = if is_right { node.right } else { node.left };
145        }
146
147        Ok(hash)
148    }
149
150    /// Returns the node at the specified `index` and its opening to the `root`.
151    ///
152    /// The path starts at the sibling of the target leaf.
153    ///
154    /// # Errors
155    /// This method can return the following errors:
156    /// - `RootNotInStore` if the `root` is not present in the store.
157    /// - `NodeNotInStore` if a node needed to traverse from `root` to `index` is not present in the
158    ///   store.
159    pub fn get_path(&self, root: Word, index: NodeIndex) -> Result<MerkleProof, MerkleError> {
160        let mut hash = root;
161        let mut path = Vec::with_capacity(index.depth().into());
162
163        // corner case: check the root is in the store when called with index `NodeIndex::root()`
164        self.nodes.get(&hash).ok_or(MerkleError::RootNotInStore(hash))?;
165
166        for i in (0..index.depth()).rev() {
167            let node = self
168                .nodes
169                .get(&hash)
170                .ok_or(MerkleError::NodeIndexNotFoundInStore(hash, index))?;
171
172            let is_right = index.is_nth_bit_odd(i);
173            hash = if is_right {
174                path.push(node.left);
175                node.right
176            } else {
177                path.push(node.right);
178                node.left
179            }
180        }
181
182        // the path is computed from root to leaf, so it must be reversed
183        path.reverse();
184
185        Ok(MerkleProof::new(hash, MerklePath::new(path)))
186    }
187
188    /// Returns `true` if a valid path exists from `root` to the specified `index`, `false`
189    /// otherwise.
190    ///
191    /// This method checks if all nodes needed to traverse from `root` to `index` are present in the
192    /// store, without building the actual path. It is more efficient than `get_path` when only
193    /// existence verification is needed.
194    pub fn has_path(&self, root: Word, index: NodeIndex) -> bool {
195        // check if the root exists
196        if !self.nodes.contains_key(&root) {
197            return false;
198        }
199
200        // traverse from root to index
201        let mut hash = root;
202        for i in (0..index.depth()).rev() {
203            let node = match self.nodes.get(&hash) {
204                Some(node) => node,
205                None => return false,
206            };
207
208            let is_right = index.is_nth_bit_odd(i);
209            hash = if is_right { node.right } else { node.left };
210        }
211
212        true
213    }
214
215    // LEAF TRAVERSAL
216    // --------------------------------------------------------------------------------------------
217
218    /// Returns the depth of the first leaf or an empty node encountered while traversing the tree
219    /// from the specified root down according to the provided index.
220    ///
221    /// The `tree_depth` parameter specifies the depth of the tree rooted at `root`. The
222    /// maximum value the argument accepts is [u64::BITS].
223    ///
224    /// # Errors
225    /// Will return an error if:
226    /// - The provided root is not found.
227    /// - The provided `tree_depth` is greater than 64.
228    /// - The provided `index` is not valid for a depth equivalent to `tree_depth`.
229    /// - No leaf or an empty node was found while traversing the tree down to `tree_depth`.
230    pub fn get_leaf_depth(
231        &self,
232        root: Word,
233        tree_depth: u8,
234        index: u64,
235    ) -> Result<u8, MerkleError> {
236        // validate depth and index
237        if tree_depth > 64 {
238            return Err(MerkleError::DepthTooBig(tree_depth as u64));
239        }
240        NodeIndex::new(tree_depth, index)?;
241
242        // check if the root exists, providing the proper error report if it doesn't
243        let empty = EmptySubtreeRoots::empty_hashes(tree_depth);
244        let mut hash = root;
245        if !self.nodes.contains_key(&hash) {
246            return Err(MerkleError::RootNotInStore(hash));
247        }
248
249        // we traverse from root to leaf, so the path is reversed
250        let mut path = (index << (64 - tree_depth)).reverse_bits();
251
252        // iterate every depth and reconstruct the path from root to leaf
253        for depth in 0..=tree_depth {
254            // we short-circuit if an empty node has been found
255            if hash == empty[depth as usize] {
256                return Ok(depth);
257            }
258
259            // fetch the children pair, mapped by its parent hash
260            let children = match self.nodes.get(&hash) {
261                Some(node) => node,
262                None => return Ok(depth),
263            };
264
265            // traverse down
266            hash = if path & 1 == 0 { children.left } else { children.right };
267            path >>= 1;
268        }
269
270        // return an error because we exhausted the index but didn't find either a leaf or an
271        // empty node
272        Err(MerkleError::DepthTooBig(tree_depth as u64 + 1))
273    }
274
275    /// Returns index and value of a leaf node which is the only leaf node in a subtree defined by
276    /// the provided root. If the subtree contains zero or more than one leaf nodes None is
277    /// returned.
278    ///
279    /// The `tree_depth` parameter specifies the depth of the parent tree such that `root` is
280    /// located in this tree at `root_index`. The maximum value the argument accepts is
281    /// [u64::BITS].
282    ///
283    /// # Errors
284    /// Will return an error if:
285    /// - The provided root is not found.
286    /// - The provided `tree_depth` is greater than 64.
287    /// - The provided `root_index` has depth greater than `tree_depth`.
288    /// - A lone node at depth `tree_depth` is not a leaf node.
289    pub fn find_lone_leaf(
290        &self,
291        root: Word,
292        root_index: NodeIndex,
293        tree_depth: u8,
294    ) -> Result<Option<(NodeIndex, Word)>, MerkleError> {
295        // we set max depth at u64::BITS as this is the largest meaningful value for a 64-bit index
296        const MAX_DEPTH: u8 = u64::BITS as u8;
297        if tree_depth > MAX_DEPTH {
298            return Err(MerkleError::DepthTooBig(tree_depth as u64));
299        }
300        let empty = EmptySubtreeRoots::empty_hashes(MAX_DEPTH);
301
302        let mut node = root;
303        if !self.nodes.contains_key(&node) {
304            return Err(MerkleError::RootNotInStore(node));
305        }
306
307        let mut index = root_index;
308        if index.depth() > tree_depth {
309            return Err(MerkleError::DepthTooBig(index.depth() as u64));
310        }
311
312        // traverse down following the path of single non-empty nodes; this works because if a
313        // node has two empty children it cannot contain a lone leaf. similarly if a node has
314        // two non-empty children it must contain at least two leaves.
315        for depth in index.depth()..tree_depth {
316            // if the node is a leaf, return; otherwise, examine the node's children
317            let children = match self.nodes.get(&node) {
318                Some(node) => node,
319                None => return Ok(Some((index, node))),
320            };
321
322            let empty_node = empty[depth as usize + 1];
323            node = if children.left != empty_node && children.right == empty_node {
324                index = index.left_child();
325                children.left
326            } else if children.left == empty_node && children.right != empty_node {
327                index = index.right_child();
328                children.right
329            } else {
330                return Ok(None);
331            };
332        }
333
334        // if we are here, we got to `tree_depth`; thus, either the current node is a leaf node,
335        // and so we return it, or it is an internal node, and then we return an error
336        if self.nodes.contains_key(&node) {
337            Err(MerkleError::DepthTooBig(tree_depth as u64 + 1))
338        } else {
339            Ok(Some((index, node)))
340        }
341    }
342
343    // DATA EXTRACTORS
344    // --------------------------------------------------------------------------------------------
345
346    /// Returns a subset of this Merkle store such that the returned Merkle store contains all
347    /// nodes which are descendants of the specified roots.
348    ///
349    /// The roots for which no descendants exist in this Merkle store are ignored.
350    pub fn subset<I, R>(&self, roots: I) -> MerkleStore
351    where
352        I: Iterator<Item = R>,
353        R: Borrow<Word>,
354    {
355        let mut store = MerkleStore::new();
356        for root in roots {
357            let root = *root.borrow();
358            store.clone_tree_from(root, self);
359        }
360        store
361    }
362
363    /// Iterator over the inner nodes of the [MerkleStore].
364    pub fn inner_nodes(&self) -> impl Iterator<Item = InnerNodeInfo> + '_ {
365        self.nodes
366            .iter()
367            .map(|(r, n)| InnerNodeInfo { value: *r, left: n.left, right: n.right })
368    }
369
370    /// Iterator over the non-empty leaves of the Merkle tree associated with the specified `root`
371    /// and `max_depth`.
372    pub fn non_empty_leaves(
373        &self,
374        root: Word,
375        max_depth: u8,
376    ) -> impl Iterator<Item = (NodeIndex, Word)> + '_ {
377        let empty_roots = EmptySubtreeRoots::empty_hashes(max_depth);
378        let mut stack = Vec::new();
379        stack.push((NodeIndex::new_unchecked(0, 0), root));
380
381        core::iter::from_fn(move || {
382            while let Some((index, node_hash)) = stack.pop() {
383                // if we are at the max depth then we have reached a leaf
384                if index.depth() == max_depth {
385                    return Some((index, node_hash));
386                }
387
388                // fetch the nodes children and push them onto the stack if they are not the roots
389                // of empty subtrees
390                if let Some(node) = self.nodes.get(&node_hash) {
391                    if !empty_roots.contains(&node.left) {
392                        stack.push((index.left_child(), node.left));
393                    }
394                    if !empty_roots.contains(&node.right) {
395                        stack.push((index.right_child(), node.right));
396                    }
397
398                // if the node is not in the store assume it is a leaf
399                } else {
400                    return Some((index, node_hash));
401                }
402            }
403
404            None
405        })
406    }
407
408    // STATE MUTATORS
409    // --------------------------------------------------------------------------------------------
410
411    /// Adds all the nodes of a Merkle path represented by `path`, opening to `node`. Returns the
412    /// new root.
413    ///
414    /// This will compute the sibling elements determined by the Merkle `path` and `node`, and
415    /// include all the nodes into the store.
416    pub fn add_merkle_path(
417        &mut self,
418        index: u64,
419        node: Word,
420        path: MerklePath,
421    ) -> Result<Word, MerkleError> {
422        let root = path.authenticated_nodes(index, node)?.fold(Word::default(), |_, node| {
423            let value: Word = node.value;
424            let left: Word = node.left;
425            let right: Word = node.right;
426
427            debug_assert_eq!(Rpo256::merge(&[left, right]), value);
428            self.nodes.insert(value, StoreNode { left, right });
429
430            node.value
431        });
432        Ok(root)
433    }
434
435    /// Adds all the nodes of multiple Merkle paths into the store.
436    ///
437    /// This will compute the sibling elements for each Merkle `path` and include all the nodes
438    /// into the store.
439    ///
440    /// For further reference, check [MerkleStore::add_merkle_path].
441    pub fn add_merkle_paths<I>(&mut self, paths: I) -> Result<(), MerkleError>
442    where
443        I: IntoIterator<Item = (u64, Word, MerklePath)>,
444    {
445        for (index_value, node, path) in paths.into_iter() {
446            self.add_merkle_path(index_value, node, path)?;
447        }
448        Ok(())
449    }
450
451    /// Sets a node to `value`.
452    ///
453    /// # Errors
454    /// This method can return the following errors:
455    /// - `RootNotInStore` if the `root` is not present in the store.
456    /// - `NodeNotInStore` if a node needed to traverse from `root` to `index` is not present in the
457    ///   store.
458    pub fn set_node(
459        &mut self,
460        mut root: Word,
461        index: NodeIndex,
462        value: Word,
463    ) -> Result<RootPath, MerkleError> {
464        let node = value;
465        let MerkleProof { value, path } = self.get_path(root, index)?;
466
467        // performs the update only if the node value differs from the opening
468        if node != value {
469            root = self.add_merkle_path(index.value(), node, path.clone())?;
470        }
471
472        Ok(RootPath { root, path })
473    }
474
475    /// Merges two elements and adds the resulting node into the store.
476    ///
477    /// Merges arbitrary values. They may be leaves, nodes, or a mixture of both.
478    pub fn merge_roots(&mut self, left_root: Word, right_root: Word) -> Result<Word, MerkleError> {
479        let parent = Rpo256::merge(&[left_root, right_root]);
480        self.nodes.insert(parent, StoreNode { left: left_root, right: right_root });
481
482        Ok(parent)
483    }
484
485    // HELPER METHODS
486    // --------------------------------------------------------------------------------------------
487
488    /// Returns the inner storage of this MerkleStore while consuming `self`.
489    pub fn into_inner(self) -> Map<Word, StoreNode> {
490        self.nodes
491    }
492
493    /// Recursively clones a tree with the specified root from the specified source into self.
494    ///
495    /// If the source store does not contain a tree with the specified root, this is a noop.
496    fn clone_tree_from(&mut self, root: Word, source: &Self) {
497        // process the node only if it is in the source
498        if let Some(node) = source.nodes.get(&root) {
499            // if the node has already been inserted, no need to process it further as all of its
500            // descendants should be already cloned from the source store
501            if self.nodes.insert(root, *node).is_none() {
502                self.clone_tree_from(node.left, source);
503                self.clone_tree_from(node.right, source);
504            }
505        }
506    }
507}
508
509// CONVERSIONS
510// ================================================================================================
511
512impl From<&MerkleTree> for MerkleStore {
513    fn from(value: &MerkleTree) -> Self {
514        let nodes = combine_nodes_with_empty_hashes(value.inner_nodes()).collect();
515        Self { nodes }
516    }
517}
518
519impl<const DEPTH: u8> From<&SimpleSmt<DEPTH>> for MerkleStore {
520    fn from(value: &SimpleSmt<DEPTH>) -> Self {
521        let nodes = combine_nodes_with_empty_hashes(value.inner_nodes()).collect();
522        Self { nodes }
523    }
524}
525
526impl From<&Smt> for MerkleStore {
527    fn from(value: &Smt) -> Self {
528        let nodes = combine_nodes_with_empty_hashes(value.inner_nodes()).collect();
529        Self { nodes }
530    }
531}
532
533impl From<&Mmr> for MerkleStore {
534    fn from(value: &Mmr) -> Self {
535        let nodes = combine_nodes_with_empty_hashes(value.inner_nodes()).collect();
536        Self { nodes }
537    }
538}
539
540impl From<&PartialMerkleTree> for MerkleStore {
541    fn from(value: &PartialMerkleTree) -> Self {
542        let nodes = combine_nodes_with_empty_hashes(value.inner_nodes()).collect();
543        Self { nodes }
544    }
545}
546
547impl FromIterator<InnerNodeInfo> for MerkleStore {
548    fn from_iter<I: IntoIterator<Item = InnerNodeInfo>>(iter: I) -> Self {
549        let nodes = combine_nodes_with_empty_hashes(iter).collect();
550        Self { nodes }
551    }
552}
553
554impl FromIterator<(Word, StoreNode)> for MerkleStore {
555    fn from_iter<I: IntoIterator<Item = (Word, StoreNode)>>(iter: I) -> Self {
556        let nodes = iter.into_iter().chain(empty_hashes()).collect();
557        Self { nodes }
558    }
559}
560
561// ITERATORS
562// ================================================================================================
563impl Extend<InnerNodeInfo> for MerkleStore {
564    fn extend<I: IntoIterator<Item = InnerNodeInfo>>(&mut self, iter: I) {
565        self.nodes.extend(
566            iter.into_iter()
567                .map(|info| (info.value, StoreNode { left: info.left, right: info.right })),
568        );
569    }
570}
571
572// SERIALIZATION
573// ================================================================================================
574
575impl Serializable for StoreNode {
576    fn write_into<W: ByteWriter>(&self, target: &mut W) {
577        self.left.write_into(target);
578        self.right.write_into(target);
579    }
580}
581
582impl Deserializable for StoreNode {
583    fn read_from<R: ByteReader>(source: &mut R) -> Result<Self, DeserializationError> {
584        let left = Word::read_from(source)?;
585        let right = Word::read_from(source)?;
586        Ok(StoreNode { left, right })
587    }
588}
589
590impl Serializable for MerkleStore {
591    fn write_into<W: ByteWriter>(&self, target: &mut W) {
592        target.write_u64(self.nodes.len() as u64);
593
594        for (k, v) in self.nodes.iter() {
595            k.write_into(target);
596            v.write_into(target);
597        }
598    }
599}
600
601impl Deserializable for MerkleStore {
602    fn read_from<R: ByteReader>(source: &mut R) -> Result<Self, DeserializationError> {
603        let len = source.read_u64()?;
604        let mut nodes: Vec<(Word, StoreNode)> = Vec::with_capacity(len as usize);
605
606        for _ in 0..len {
607            let key = Word::read_from(source)?;
608            let value = StoreNode::read_from(source)?;
609            nodes.push((key, value));
610        }
611
612        Ok(nodes.into_iter().collect())
613    }
614}
615
616// HELPER FUNCTIONS
617// ================================================================================================
618
619/// Creates empty hashes for all the subtrees of a tree with a max depth of 255.
620fn empty_hashes() -> impl Iterator<Item = (Word, StoreNode)> {
621    let subtrees = EmptySubtreeRoots::empty_hashes(255);
622    subtrees
623        .iter()
624        .rev()
625        .copied()
626        .zip(subtrees.iter().rev().skip(1).copied())
627        .map(|(child, parent)| (parent, StoreNode { left: child, right: child }))
628}
629
630/// Consumes an iterator of [InnerNodeInfo] and returns an iterator of `(value, node)` tuples
631/// which includes the nodes associate with roots of empty subtrees up to a depth of 255.
632fn combine_nodes_with_empty_hashes(
633    nodes: impl IntoIterator<Item = InnerNodeInfo>,
634) -> impl Iterator<Item = (Word, StoreNode)> {
635    nodes
636        .into_iter()
637        .map(|info| (info.value, StoreNode { left: info.left, right: info.right }))
638        .chain(empty_hashes())
639}