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