commonware_cryptography/bls12381/
dkg.rs

1//! Distributed Key Generation (DKG) and Resharing protocol for the BLS12-381 curve.
2//!
3//! This module implements an interactive Distributed Key Generation (DKG) and Resharing protocol
4//! for the BLS12-381 curve. Unlike other constructions, this construction does not require encrypted
5//! shares to be publicly broadcast to complete a DKG/Reshare. Shares, instead, are sent directly
6//! between dealers and players over an encrypted channel (which can be instantiated
7//! with [commonware-p2p](https://docs.rs/commonware-p2p)).
8//!
9//! The DKG is based on the "Joint-Feldman" construction from "Secure Distributed Key
10//! Generation for Discrete-Log Based Cryptosystems" (GJKR99) and Resharing is based
11//! on the construction described in "Redistributing secret shares to new access structures
12//! and its applications" (Desmedt97).
13//!
14//! # Overview
15//!
16//! The protocol involves _dealers_ and _players_. The dealers are trying to jointly create a shared
17//! key, and then distribute it among the players. The dealers may have pre-existing shares of a key
18//! from a previous round, in which case the goal is to re-distribute that key among the players,
19//! with fresh randomness.
20//!
21//! The protocol is also designed such that an external observer can figure out whether the protocol
22//! succeeded or failed, and learn of the public outputs of the protocol. This includes
23//! the participants in the protocol, and the public polynomial committing to the key
24//! and its sharing.
25//!
26//! # Usage
27//!
28//! ## Core Types
29//!
30//! * [`Info`]: Configuration for a DKG/Reshare round, containing the dealers, players, and optional previous output
31//! * [`Output`]: The public result of a successful DKG round, containing the public polynomial and player list
32//! * [`Share`]: A player's private share of the distributed key (from `primitives::group`)
33//! * [`Dealer`]: State machine for a dealer participating in the protocol
34//! * [`Player`]: State machine for a player receiving shares
35//! * [`SignedDealerLog`]: A dealer's signed transcript of their interactions with players
36//!
37//! ## Message Types
38//!
39//! * [`DealerPubMsg`]: Public commitment polynomial sent from dealer to all players
40//! * [`DealerPrivMsg`]: Private share sent from dealer to a specific player
41//! * [`PlayerAck`]: Acknowledgement sent from player back to dealer
42//! * [`DealerLog`]: Complete log of a dealer's interactions (commitments and acks/reveals)
43//!
44//! ## Protocol Flow
45//!
46//! ### Step 1: Initialize Round
47//!
48//! Create a [`Info`] using [`Info::new`] with:
49//! - Round number (should increment sequentially, including for failed rounds)
50//! - Optional previous [`Output`] (for resharing)
51//! - List of dealers (must be >= quorum of previous round if resharing)
52//! - List of players who will receive shares
53//!
54//! ### Step 2: Dealer Phase
55//!
56//! Each dealer calls [`Dealer::start`] which returns:
57//! - A [`Dealer`] instance for tracking state
58//! - A [`DealerPubMsg`] containing the polynomial commitment to broadcast
59//! - A vector of `(player_id, DealerPrivMsg)` pairs to send privately
60//!
61//! The [`DealerPubMsg`] contains a public polynomial commitment of degree `2f` where `f = max_faults(n)`.
62//! Each [`DealerPrivMsg`] contains a scalar evaluation of the dealer's private polynomial at the player's index.
63//!
64//! ### Step 3: Player Verification
65//!
66//! Each player creates a [`Player`] instance via [`Player::new`], then for each dealer message:
67//! - Call [`Player::dealer_message`] with the [`DealerPubMsg`] and [`DealerPrivMsg`]
68//! - If valid, this returns a [`PlayerAck`] containing a signature over `(dealer, commitment)`
69//! - The player verifies that the private share matches the public commitment evaluation
70//!
71//! ### Step 4: Dealer Collection
72//!
73//! Each dealer:
74//! - Calls [`Dealer::receive_player_ack`] for each acknowledgement received
75//! - After timeout, calls [`Dealer::finalize`] to produce a [`SignedDealerLog`]
76//! - The log contains the commitment and either acks or reveals for each player
77//!
78//! ### Step 5: Finalization
79//!
80//! With collected [`SignedDealerLog`]s:
81//! - Call [`SignedDealerLog::check`] to verify and extract [`DealerLog`]s
82//! - Players call [`Player::finalize`] with all logs to compute their [`Share`] and [`Output`]
83//! - Observers call [`observe`] with all logs to compute just the [`Output`]
84//!
85//! The [`Output`] contains:
86//! - The final public polynomial (sum of dealer polynomials for DKG, interpolation for reshare),
87//! - The list of dealers who distributed shares,
88//! - The list of players who received shares,
89//! - The set of players whose shares may have been revealed,
90//! - A digest of the round's [`Info`] (including the counter, and the list of dealers and players).
91//!
92//! ## Trusted Dealing Functions
93//!
94//! As a convenience (for tests, etc.), this module also provides functions for
95//! generating shares using a trusted dealer.
96//!
97//! - [`deal`]: given a list of players, generates an [`Output`] like the DKG would,
98//! - [`deal_anonymous`]: a lower-level version that produces a polynomial directly,
99//!   and doesn't require public keys for the players.
100//!
101//! # Caveats
102//!
103//! ## Synchrony Assumption
104//!
105//! Under synchrony (where `t` is the maximum amount of time it takes for a message to be sent between any two participants),
106//! this construction can be used to maintain a shared secret where at least `f + 1` honest players must participate to
107//! recover the shared secret (`2f + 1` threshold where at most `f` players are Byzantine). To see how this is true,
108//! first consider that in any successful round there must exist `2f + 1` commitments with at most `f` reveals. This implies
109//! that all players must have acknowledged or have access to a reveal for each of the `2f + 1` selected commitments (allowing
110//! them to derive their share). Next, consider that when the network is synchronous that all `2f + 1` honest players send
111//! acknowledgements to honest dealers before `2t`. Because `2f + 1` commitments must be chosen, at least `f + 1` commitments
112//! must be from honest dealers (where no honest player dealing is revealed). Even if the remaining `f` commitments are from
113//! Byzantine dealers, there will not be enough dealings to recover the derived share of any honest player (at most `f` of
114//! `2f + 1` dealings publicly revealed). Given all `2f + 1` honest players have access to their shares and it is not possible
115//! for a Byzantine player to derive any honest player's share, this claim holds.
116//!
117//! If the network is not synchronous, however, Byzantine players can collude to recover a shared secret with the
118//! participation of a single honest player (rather than `f + 1`) and `f + 1` honest players will each be able to derive
119//! the shared secret (if the Byzantine players reveal their shares). To see how this could be, consider a network where
120//! `f` honest participants are in one partition and (`f + 1` honest and `f` Byzantine participants) are in another. All
121//! `f` Byzantine players acknowledge dealings from the `f + 1` honest dealers. Participants in the second partition will
122//! complete a round and all the reveals will belong to the same set of `f` honest players (that are in the first partition).
123//! A colluding Byzantine adversary will then have access to their acknowledged `f` shares and the revealed `f` shares
124//! (requiring only the participation of a single honest player that was in their partition to recover the shared secret).
125//! If the Byzantine adversary reveals all of their (still private) shares at this time, each of the `f + 1` honest players
126//! that were in the second partition will be able to derive the shared secret without collusion (using their private share
127//! and the `2f` public shares). It will not be possible for any external observer, however, to recover the shared secret.
128//!
129//! ### Future Work: Dropping the Synchrony Assumption?
130//!
131//! It is possible to design a DKG/Resharing scheme that maintains a shared secret where at least `f + 1` honest players
132//! must participate to recover the shared secret that doesn't require a synchrony assumption (`2f + 1` threshold
133//! where at most `f` players are Byzantine). However, known constructions that satisfy this requirement require both
134//! broadcasting encrypted dealings publicly and employing Zero-Knowledge Proofs (ZKPs) to attest that encrypted dealings
135//! were generated correctly ([Groth21](https://eprint.iacr.org/2021/339), [Kate23](https://eprint.iacr.org/2023/451)).
136//!
137//! As of January 2025, these constructions are still considered novel (2-3 years in production), require stronger
138//! cryptographic assumptions, don't scale to hundreds of participants (unless dealers have powerful hardware), and provide
139//! observers the opportunity to brute force decrypt shares (even if honest players are online).
140//!
141//! ## Handling Complaints
142//!
143//! This crate does not provide an integrated mechanism for tracking complaints from players (of malicious dealers). However, it is
144//! possible to implement your own mechanism and to manually disqualify dealers from a given round in the arbiter. This decision was made
145//! because the mechanism for communicating commitments/shares/acknowledgements is highly dependent on the context in which this
146//! construction is used.
147//!
148//! In practice:
149//! - [`Player::dealer_message`] returns `None` for invalid messages (implicit complaint)
150//! - [`Dealer::receive_player_ack`] validates acknowledgements
151//! - Other custom mechanisms can exclude dealers before calling [`observe`] or [`Player::finalize`],
152//!   to enforce other rules for "misbehavior" beyond what the DKG does already.
153//!
154//! ## Non-Uniform Distribution
155//!
156//! The Joint-Feldman DKG protocol does not guarantee a uniformly random secret key is generated. An adversary
157//! can introduce `O(lg N)` bits of bias into the key with `O(poly(N))` amount of computation. For uses
158//! like signing, threshold encryption, where the security of the scheme reduces to that of
159//! the underlying assumption that cryptographic constructions using the curve are secure (i.e.
160//! that the Discrete Logarithm Problem, or stronger variants, are hard), then this caveat does
161//! not affect the security of the scheme. This must be taken into account when integrating this
162//! component into more esoteric schemes.
163//!
164//! This choice was explicitly made, because the best known protocols guaranteeing a uniform output
165//! require an extra round of broadcast ([GJKR02](https://www.researchgate.net/publication/2558744_Revisiting_the_Distributed_Key_Generation_for_Discrete-Log_Based_Cryptosystems),
166//! [BK25](https://eprint.iacr.org/2025/819)).
167//!
168//! ## Share Reveals
169//!
170//! In order to prevent malicious dealers from withholding shares from players, we
171//! require the dealers reveal the shares for which they did not receive acks.
172//! Because of the synchrony assumption above, this will only happen if either:
173//! - the dealer is malicious, not sending a share, but honestly revealing,
174//! - or, the player is malicious, not sending an ack when they should.
175//!
176//! Thus, for honest players, in the worst case, `f` reveals get created, because
177//! they correctly did not ack the `f` malicious dealers who failed to send them
178//! a share. In that case, their final share remains secret, because it is the linear
179//! combination of at least `f + 1` shares received from dealers.
180//!
181//! # Example
182//!
183//! ```
184//! use commonware_cryptography::bls12381::{
185//!     dkg::{Dealer, Player, Info, SignedDealerLog, observe},
186//!     primitives::{variant::MinSig, sharing::Mode},
187//! };
188//! use commonware_cryptography::{ed25519, Signer};
189//! use commonware_math::algebra::Random;
190//! use commonware_utils::{ordered::Set, TryCollect};
191//! use std::collections::BTreeMap;
192//! use rand::SeedableRng;
193//! use rand_chacha::ChaCha8Rng;
194//!
195//! # fn main() -> Result<(), Box<dyn std::error::Error>> {
196//! let mut rng = ChaCha8Rng::seed_from_u64(42);
197//!
198//! // Generate 4 Ed25519 private keys for participants
199//! let mut private_keys = Vec::new();
200//! for _ in 0..4 {
201//!     let private_key = ed25519::PrivateKey::random(&mut rng);
202//!     private_keys.push(private_key);
203//! }
204//!
205//! // All 4 participants are both dealers and players in initial DKG
206//! let dealer_set: Set<ed25519::PublicKey> = private_keys.iter()
207//!     .map(|k| k.public_key())
208//!     .try_collect()?;
209//! let player_set = dealer_set.clone();
210//!
211//! // Step 1: Create round info for initial DKG
212//! let info = Info::<MinSig, ed25519::PublicKey>::new(
213//!     b"application-namespace",
214//!     0,                        // round number
215//!     None,                     // no previous output (initial DKG)
216//!     Mode::default(),   // sharing mode
217//!     dealer_set.clone(),       // dealers
218//!     player_set.clone(),       // players
219//! )?;
220//!
221//! // Step 2: Initialize players
222//! let mut players = BTreeMap::new();
223//! for private_key in &private_keys {
224//!     let player = Player::<MinSig, ed25519::PrivateKey>::new(
225//!         info.clone(),
226//!         private_key.clone(),
227//!     )?;
228//!     players.insert(private_key.public_key(), player);
229//! }
230//!
231//! // Step 3: Run dealer protocol for each participant
232//! let mut dealer_logs = BTreeMap::new();
233//! for dealer_priv in &private_keys {
234//!     // Each dealer generates messages for all players
235//!     let (mut dealer, pub_msg, priv_msgs) = Dealer::start(
236//!         &mut rng,
237//!         info.clone(),
238//!         dealer_priv.clone(),
239//!         None,  // no previous share for initial DKG
240//!     )?;
241//!
242//!     // Distribute messages to players and collect acknowledgements
243//!     for (player_pk, priv_msg) in priv_msgs {
244//!         if let Some(player) = players.get_mut(&player_pk) {
245//!             if let Some(ack) = player.dealer_message(
246//!                 dealer_priv.public_key(),
247//!                 pub_msg.clone(),
248//!                 priv_msg,
249//!             ) {
250//!                 dealer.receive_player_ack(player_pk, ack)?;
251//!             }
252//!         }
253//!     }
254//!
255//!     // Finalize dealer and verify log
256//!     let signed_log = dealer.finalize();
257//!     if let Some((dealer_pk, log)) = signed_log.check(&info) {
258//!         dealer_logs.insert(dealer_pk, log);
259//!     }
260//! }
261//!
262//! // Step 4: Players finalize to get their shares
263//! let mut player_shares = BTreeMap::new();
264//! for (player_pk, player) in players {
265//!     let (output, share) = player.finalize(
266//!       dealer_logs.clone(),
267//!       1 // Increase this for parallelism.
268//!     )?;
269//!     println!("Player {:?} got share at index {}", player_pk, share.index);
270//!     player_shares.insert(player_pk, share);
271//! }
272//!
273//! // Step 5: Observer can also compute the public output
274//! let observer_output = observe::<MinSig, ed25519::PublicKey>(
275//!     info,
276//!     dealer_logs,
277//!     1 // Increase this for parallelism.
278//! )?;
279//! println!("DKG completed with threshold {}", observer_output.quorum());
280//! # Ok(())
281//! # }
282//! ```
283//!
284//! For a complete example with resharing, see [commonware-reshare](https://docs.rs/commonware-reshare).
285use super::primitives::group::Share;
286use crate::{
287    bls12381::primitives::{
288        group::Scalar,
289        sharing::{Mode, Sharing},
290        variant::Variant,
291    },
292    transcript::{Summary, Transcript},
293    PublicKey, Signer,
294};
295use commonware_codec::{Encode, EncodeSize, RangeCfg, Read, ReadExt, Write};
296use commonware_math::{
297    algebra::{Additive, CryptoGroup, Random},
298    poly::{Interpolator, Poly},
299};
300use commonware_utils::{
301    ordered::{Map, Quorum, Set},
302    TryCollect, NZU32,
303};
304use core::num::NonZeroU32;
305use rand_core::CryptoRngCore;
306use std::collections::BTreeMap;
307use thiserror::Error;
308
309const NAMESPACE: &[u8] = b"_COMMONWARE_CRYPTOGRAPHY_BLS12381_DKG";
310const SIG_ACK: &[u8] = b"ack";
311const SIG_LOG: &[u8] = b"log";
312
313/// The error type for the DKG protocol.
314///
315/// The only error which can happen through no fault of your own is
316/// [`Error::DkgFailed`]. Everything else only happens if you use a configuration
317/// for [`Info`] or [`Dealer`] which is invalid in some way.
318#[derive(Debug, Error)]
319pub enum Error {
320    #[error("missing dealer's share from the previous round")]
321    MissingDealerShare,
322    #[error("player is not present in the list of players")]
323    UnknownPlayer,
324    #[error("dealer is not present in the previous list of players")]
325    UnknownDealer(String),
326    #[error("invalid number of dealers: {0}")]
327    NumDealers(usize),
328    #[error("invalid number of players: {0}")]
329    NumPlayers(usize),
330    #[error("dkg failed for some reason")]
331    DkgFailed,
332}
333
334/// The output of a successful DKG.
335#[derive(Debug, Clone, PartialEq, Eq)]
336pub struct Output<V: Variant, P> {
337    summary: Summary,
338    public: Sharing<V>,
339    dealers: Set<P>,
340    players: Set<P>,
341    revealed: Set<P>,
342}
343
344impl<V: Variant, P: Ord> Output<V, P> {
345    fn share_commitment(&self, player: &P) -> Option<V::Public> {
346        self.public.partial_public(self.players.index(player)?).ok()
347    }
348
349    /// Return the quorum, i.e. the number of players needed to reconstruct the key.
350    pub fn quorum(&self) -> u32 {
351        self.players.quorum()
352    }
353
354    /// Get the public polynomial associated with this output.
355    ///
356    /// This is useful for verifying partial signatures, with [crate::bls12381::primitives::ops::partial_verify_message].
357    pub const fn public(&self) -> &Sharing<V> {
358        &self.public
359    }
360
361    /// Return the dealers who were selected in this round of the DKG.
362    pub const fn dealers(&self) -> &Set<P> {
363        &self.dealers
364    }
365
366    /// Return the players who participated in this round of the DKG, and should have shares.
367    pub const fn players(&self) -> &Set<P> {
368        &self.players
369    }
370
371    /// Return the set of players whose shares may have been revealed.
372    ///
373    /// These are players who had more than `max_faults` reveals.
374    pub const fn revealed(&self) -> &Set<P> {
375        &self.revealed
376    }
377}
378
379impl<V: Variant, P: PublicKey> EncodeSize for Output<V, P> {
380    fn encode_size(&self) -> usize {
381        self.summary.encode_size()
382            + self.public.encode_size()
383            + self.dealers.encode_size()
384            + self.players.encode_size()
385            + self.revealed.encode_size()
386    }
387}
388
389impl<V: Variant, P: PublicKey> Write for Output<V, P> {
390    fn write(&self, buf: &mut impl bytes::BufMut) {
391        self.summary.write(buf);
392        self.public.write(buf);
393        self.dealers.write(buf);
394        self.players.write(buf);
395        self.revealed.write(buf);
396    }
397}
398
399impl<V: Variant, P: PublicKey> Read for Output<V, P> {
400    type Cfg = NonZeroU32;
401
402    fn read_cfg(
403        buf: &mut impl bytes::Buf,
404        &max_participants: &Self::Cfg,
405    ) -> Result<Self, commonware_codec::Error> {
406        let max_participants_usize = max_participants.get() as usize;
407        Ok(Self {
408            summary: ReadExt::read(buf)?,
409            public: Read::read_cfg(buf, &max_participants)?,
410            dealers: Read::read_cfg(buf, &(RangeCfg::new(1..=max_participants_usize), ()))?, // at least one dealer must be part of a dealing
411            players: Read::read_cfg(buf, &(RangeCfg::new(1..=max_participants_usize), ()))?, // at least one player must be part of a dealing
412            revealed: Read::read_cfg(buf, &(RangeCfg::new(0..=max_participants_usize), ()))?, // there may not be any reveals
413        })
414    }
415}
416
417#[cfg(feature = "arbitrary")]
418impl<P: PublicKey, V: Variant> arbitrary::Arbitrary<'_> for Output<V, P>
419where
420    P: for<'a> arbitrary::Arbitrary<'a> + Ord,
421    V::Public: for<'a> arbitrary::Arbitrary<'a>,
422{
423    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
424        let summary = u.arbitrary()?;
425        let public: Sharing<V> = u.arbitrary()?;
426        let total = public.total().get() as usize;
427
428        let num_dealers = u.int_in_range(1..=total * 2)?;
429        let dealers = Set::try_from(
430            u.arbitrary_iter::<P>()?
431                .take(num_dealers)
432                .collect::<Result<Vec<_>, _>>()?,
433        )
434        .map_err(|_| arbitrary::Error::IncorrectFormat)?;
435
436        let players = Set::try_from(
437            u.arbitrary_iter::<P>()?
438                .take(total)
439                .collect::<Result<Vec<_>, _>>()?,
440        )
441        .map_err(|_| arbitrary::Error::IncorrectFormat)?;
442
443        let max_revealed = commonware_utils::max_faults(total as u32) as usize;
444        let revealed = Set::from_iter_dedup(
445            players
446                .iter()
447                .filter(|_| u.arbitrary::<bool>().unwrap_or(false))
448                .take(max_revealed)
449                .cloned(),
450        );
451
452        Ok(Self {
453            summary,
454            public,
455            dealers,
456            players,
457            revealed,
458        })
459    }
460}
461
462/// Information about the current round of the DKG.
463///
464/// This is used to bind signatures to the current round, and to provide the
465/// information that dealers, players, and observers need to perform their actions.
466#[derive(Debug, Clone)]
467pub struct Info<V: Variant, P: PublicKey> {
468    summary: Summary,
469    round: u64,
470    previous: Option<Output<V, P>>,
471    mode: Mode,
472    dealers: Set<P>,
473    players: Set<P>,
474}
475
476impl<V: Variant, P: PublicKey> PartialEq for Info<V, P> {
477    fn eq(&self, other: &Self) -> bool {
478        self.summary == other.summary
479    }
480}
481
482impl<V: Variant, P: PublicKey> Info<V, P> {
483    /// Figure out what the dealer share should be.
484    ///
485    /// If there's no previous round, we need a random value, hence `rng`.
486    ///
487    /// However, if there is a previous round, we expect a share, hence `Result`.
488    fn unwrap_or_random_share(
489        &self,
490        mut rng: impl CryptoRngCore,
491        share: Option<Scalar>,
492    ) -> Result<Scalar, Error> {
493        let out = match (self.previous.as_ref(), share) {
494            (None, None) => Scalar::random(&mut rng),
495            (_, Some(x)) => x,
496            (Some(_), None) => return Err(Error::MissingDealerShare),
497        };
498        Ok(out)
499    }
500
501    const fn num_players(&self) -> NonZeroU32 {
502        // Will not panic because we check that the number of players is non-empty in `new`
503        NZU32!(self.players.len() as u32)
504    }
505
506    fn degree(&self) -> u32 {
507        self.players.quorum().saturating_sub(1)
508    }
509
510    fn required_commitments(&self) -> u32 {
511        let dealer_quorum = self.dealers.quorum();
512        let prev_quorum = self
513            .previous
514            .as_ref()
515            .map(Output::quorum)
516            .unwrap_or(u32::MIN);
517        dealer_quorum.max(prev_quorum)
518    }
519
520    fn max_reveals(&self) -> u32 {
521        self.players.max_faults()
522    }
523
524    fn player_index(&self, player: &P) -> Result<u32, Error> {
525        self.players.index(player).ok_or(Error::UnknownPlayer)
526    }
527
528    fn dealer_index(&self, dealer: &P) -> Result<u32, Error> {
529        self.dealers
530            .index(dealer)
531            .ok_or(Error::UnknownDealer(format!("{dealer:?}")))
532    }
533
534    fn player_scalar(&self, player: &P) -> Result<Scalar, Error> {
535        Ok(self
536            .mode
537            .scalar(self.num_players(), self.player_index(player)?)
538            .expect("player index should be < num_players"))
539    }
540
541    #[must_use]
542    fn check_dealer_pub_msg(&self, dealer: &P, pub_msg: &DealerPubMsg<V>) -> bool {
543        if self.degree() != pub_msg.commitment.degree_exact() {
544            return false;
545        }
546        if let Some(previous) = self.previous.as_ref() {
547            let Some(share_commitment) = previous.share_commitment(dealer) else {
548                return false;
549            };
550            if *pub_msg.commitment.constant() != share_commitment {
551                return false;
552            }
553        }
554        true
555    }
556
557    #[must_use]
558    fn check_dealer_priv_msg(
559        &self,
560        player: &P,
561        pub_msg: &DealerPubMsg<V>,
562        priv_msg: &DealerPrivMsg,
563    ) -> bool {
564        let Ok(scalar) = self.player_scalar(player) else {
565            return false;
566        };
567        let expected = pub_msg.commitment.eval_msm(&scalar);
568        expected == V::Public::generator() * &priv_msg.share
569    }
570}
571
572impl<V: Variant, P: PublicKey> Info<V, P> {
573    /// Create a new [`Info`].
574    ///
575    /// `namespace` must be provided to isolate different applications
576    /// performing DKGs from each other.
577    /// `round` should be a counter, always incrementing, even for failed DKGs.
578    /// `previous` should be the result of the previous successful DKG.
579    /// `dealers` should be the list of public keys for the dealers. This MUST
580    /// be a subset of the previous round's players.
581    /// `players` should be the list of public keys for the players.
582    pub fn new(
583        namespace: &[u8],
584        round: u64,
585        previous: Option<Output<V, P>>,
586        mode: Mode,
587        dealers: Set<P>,
588        players: Set<P>,
589    ) -> Result<Self, Error> {
590        let participant_range = 1..u32::MAX as usize;
591        if !participant_range.contains(&dealers.len()) {
592            return Err(Error::NumDealers(dealers.len()));
593        }
594        if !participant_range.contains(&players.len()) {
595            return Err(Error::NumPlayers(players.len()));
596        }
597        if let Some(previous) = previous.as_ref() {
598            if let Some(unknown) = dealers
599                .iter()
600                .find(|d| previous.players.position(d).is_none())
601            {
602                return Err(Error::UnknownDealer(format!("{unknown:?}")));
603            }
604            if dealers.len() < previous.quorum() as usize {
605                return Err(Error::NumDealers(dealers.len()));
606            }
607        }
608        let summary = Transcript::new(NAMESPACE)
609            .commit(namespace)
610            .commit(round.encode())
611            .commit(previous.encode())
612            .commit(dealers.encode())
613            .commit(players.encode())
614            .summarize();
615        Ok(Self {
616            summary,
617            round,
618            previous,
619            mode,
620            dealers,
621            players,
622        })
623    }
624
625    /// Return the round number for this round.
626    ///
627    /// Round numbers should increase sequentially.
628    pub const fn round(&self) -> u64 {
629        self.round
630    }
631}
632
633#[derive(Clone, Debug)]
634pub struct DealerPubMsg<V: Variant> {
635    commitment: Poly<V::Public>,
636}
637
638impl<V: Variant> PartialEq for DealerPubMsg<V> {
639    fn eq(&self, other: &Self) -> bool {
640        self.commitment == other.commitment
641    }
642}
643
644impl<V: Variant> Eq for DealerPubMsg<V> {}
645
646impl<V: Variant> EncodeSize for DealerPubMsg<V> {
647    fn encode_size(&self) -> usize {
648        self.commitment.encode_size()
649    }
650}
651
652impl<V: Variant> Write for DealerPubMsg<V> {
653    fn write(&self, buf: &mut impl bytes::BufMut) {
654        self.commitment.write(buf);
655    }
656}
657
658impl<V: Variant> Read for DealerPubMsg<V> {
659    type Cfg = NonZeroU32;
660
661    fn read_cfg(
662        buf: &mut impl bytes::Buf,
663        &max_size: &Self::Cfg,
664    ) -> Result<Self, commonware_codec::Error> {
665        Ok(Self {
666            commitment: Read::read_cfg(buf, &(RangeCfg::from(NZU32!(1)..=max_size), ()))?,
667        })
668    }
669}
670
671#[cfg(feature = "arbitrary")]
672impl<V: Variant> arbitrary::Arbitrary<'_> for DealerPubMsg<V>
673where
674    V::Public: for<'a> arbitrary::Arbitrary<'a>,
675{
676    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
677        let commitment = u.arbitrary()?;
678        Ok(Self { commitment })
679    }
680}
681
682#[derive(Clone, PartialEq, Eq)]
683pub struct DealerPrivMsg {
684    share: Scalar,
685}
686
687impl std::fmt::Debug for DealerPrivMsg {
688    fn fmt(&self, f: &mut core::fmt::Formatter<'_>) -> core::fmt::Result {
689        write!(f, "DealerPrivMsg(REDACTED)")
690    }
691}
692
693impl EncodeSize for DealerPrivMsg {
694    fn encode_size(&self) -> usize {
695        self.share.encode_size()
696    }
697}
698
699impl Write for DealerPrivMsg {
700    fn write(&self, buf: &mut impl bytes::BufMut) {
701        self.share.write(buf);
702    }
703}
704
705impl Read for DealerPrivMsg {
706    type Cfg = ();
707
708    fn read_cfg(
709        buf: &mut impl bytes::Buf,
710        _cfg: &Self::Cfg,
711    ) -> Result<Self, commonware_codec::Error> {
712        Ok(Self {
713            share: ReadExt::read(buf)?,
714        })
715    }
716}
717
718#[cfg(feature = "arbitrary")]
719impl arbitrary::Arbitrary<'_> for DealerPrivMsg {
720    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
721        let share = u.arbitrary()?;
722        Ok(Self { share })
723    }
724}
725
726#[derive(Clone, Debug)]
727pub struct PlayerAck<P: PublicKey> {
728    sig: P::Signature,
729}
730
731impl<P: PublicKey> PartialEq for PlayerAck<P> {
732    fn eq(&self, other: &Self) -> bool {
733        self.sig == other.sig
734    }
735}
736
737impl<P: PublicKey> EncodeSize for PlayerAck<P> {
738    fn encode_size(&self) -> usize {
739        self.sig.encode_size()
740    }
741}
742
743impl<P: PublicKey> Write for PlayerAck<P> {
744    fn write(&self, buf: &mut impl bytes::BufMut) {
745        self.sig.write(buf);
746    }
747}
748
749impl<P: PublicKey> Read for PlayerAck<P> {
750    type Cfg = ();
751
752    fn read_cfg(
753        buf: &mut impl bytes::Buf,
754        _cfg: &Self::Cfg,
755    ) -> Result<Self, commonware_codec::Error> {
756        Ok(Self {
757            sig: ReadExt::read(buf)?,
758        })
759    }
760}
761
762#[cfg(feature = "arbitrary")]
763impl<P: PublicKey> arbitrary::Arbitrary<'_> for PlayerAck<P>
764where
765    P::Signature: for<'a> arbitrary::Arbitrary<'a>,
766{
767    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
768        let sig = u.arbitrary()?;
769        Ok(Self { sig })
770    }
771}
772
773#[derive(Clone, PartialEq)]
774enum AckOrReveal<P: PublicKey> {
775    Ack(PlayerAck<P>),
776    Reveal(DealerPrivMsg),
777}
778
779impl<P: PublicKey> AckOrReveal<P> {
780    const fn is_reveal(&self) -> bool {
781        matches!(*self, Self::Reveal(_))
782    }
783}
784
785impl<P: PublicKey> std::fmt::Debug for AckOrReveal<P> {
786    fn fmt(&self, f: &mut core::fmt::Formatter<'_>) -> core::fmt::Result {
787        match self {
788            Self::Ack(x) => write!(f, "Ack({x:?})"),
789            Self::Reveal(_) => write!(f, "Reveal(REDACTED)"),
790        }
791    }
792}
793
794impl<P: PublicKey> EncodeSize for AckOrReveal<P> {
795    fn encode_size(&self) -> usize {
796        1 + match self {
797            Self::Ack(x) => x.encode_size(),
798            Self::Reveal(x) => x.encode_size(),
799        }
800    }
801}
802
803impl<P: PublicKey> Write for AckOrReveal<P> {
804    fn write(&self, buf: &mut impl bytes::BufMut) {
805        match self {
806            Self::Ack(x) => {
807                0u8.write(buf);
808                x.write(buf);
809            }
810            Self::Reveal(x) => {
811                1u8.write(buf);
812                x.write(buf);
813            }
814        }
815    }
816}
817
818impl<P: PublicKey> Read for AckOrReveal<P> {
819    type Cfg = ();
820
821    fn read_cfg(
822        buf: &mut impl bytes::Buf,
823        _cfg: &Self::Cfg,
824    ) -> Result<Self, commonware_codec::Error> {
825        let tag = u8::read(buf)?;
826        match tag {
827            0 => Ok(Self::Ack(ReadExt::read(buf)?)),
828            1 => Ok(Self::Reveal(ReadExt::read(buf)?)),
829            x => Err(commonware_codec::Error::InvalidEnum(x)),
830        }
831    }
832}
833
834#[cfg(feature = "arbitrary")]
835impl<P: PublicKey> arbitrary::Arbitrary<'_> for AckOrReveal<P>
836where
837    P: for<'a> arbitrary::Arbitrary<'a>,
838    P::Signature: for<'a> arbitrary::Arbitrary<'a>,
839{
840    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
841        let choice = u.int_in_range(0..=1)?;
842        match choice {
843            0 => {
844                let ack = u.arbitrary()?;
845                Ok(Self::Ack(ack))
846            }
847            1 => {
848                let reveal = u.arbitrary()?;
849                Ok(Self::Reveal(reveal))
850            }
851            _ => unreachable!(),
852        }
853    }
854}
855
856#[derive(Clone, Debug)]
857enum DealerResult<P: PublicKey> {
858    Ok(Map<P, AckOrReveal<P>>),
859    TooManyReveals,
860}
861
862impl<P: PublicKey> PartialEq for DealerResult<P> {
863    fn eq(&self, other: &Self) -> bool {
864        match (self, other) {
865            (Self::Ok(x), Self::Ok(y)) => x == y,
866            (Self::TooManyReveals, Self::TooManyReveals) => true,
867            _ => false,
868        }
869    }
870}
871
872impl<P: PublicKey> EncodeSize for DealerResult<P> {
873    fn encode_size(&self) -> usize {
874        1 + match self {
875            Self::Ok(r) => r.encode_size(),
876            Self::TooManyReveals => 0,
877        }
878    }
879}
880
881impl<P: PublicKey> Write for DealerResult<P> {
882    fn write(&self, buf: &mut impl bytes::BufMut) {
883        match self {
884            Self::Ok(r) => {
885                0u8.write(buf);
886                r.write(buf);
887            }
888            Self::TooManyReveals => {
889                1u8.write(buf);
890            }
891        }
892    }
893}
894
895impl<P: PublicKey> Read for DealerResult<P> {
896    type Cfg = NonZeroU32;
897
898    fn read_cfg(
899        buf: &mut impl bytes::Buf,
900        &max_players: &Self::Cfg,
901    ) -> Result<Self, commonware_codec::Error> {
902        let tag = u8::read(buf)?;
903        match tag {
904            0 => Ok(Self::Ok(Read::read_cfg(
905                buf,
906                &(RangeCfg::from(0..=max_players.get() as usize), (), ()),
907            )?)),
908            1 => Ok(Self::TooManyReveals),
909            x => Err(commonware_codec::Error::InvalidEnum(x)),
910        }
911    }
912}
913
914#[cfg(feature = "arbitrary")]
915impl<P: PublicKey> arbitrary::Arbitrary<'_> for DealerResult<P>
916where
917    P: for<'a> arbitrary::Arbitrary<'a>,
918    P::Signature: for<'a> arbitrary::Arbitrary<'a>,
919{
920    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
921        let choice = u.int_in_range(0..=1)?;
922        match choice {
923            0 => {
924                use commonware_utils::TryFromIterator;
925                use std::collections::HashMap;
926
927                let base: HashMap<P, AckOrReveal<P>> = u.arbitrary()?;
928                let map = Map::try_from_iter(base.into_iter())
929                    .map_err(|_| arbitrary::Error::IncorrectFormat)?;
930
931                Ok(Self::Ok(map))
932            }
933            1 => Ok(Self::TooManyReveals),
934            _ => unreachable!(),
935        }
936    }
937}
938
939#[derive(Clone, Debug)]
940pub struct DealerLog<V: Variant, P: PublicKey> {
941    pub_msg: DealerPubMsg<V>,
942    results: DealerResult<P>,
943}
944
945impl<V: Variant, P: PublicKey> PartialEq for DealerLog<V, P> {
946    fn eq(&self, other: &Self) -> bool {
947        self.pub_msg == other.pub_msg && self.results == other.results
948    }
949}
950
951impl<V: Variant, P: PublicKey> EncodeSize for DealerLog<V, P> {
952    fn encode_size(&self) -> usize {
953        self.pub_msg.encode_size() + self.results.encode_size()
954    }
955}
956
957impl<V: Variant, P: PublicKey> Write for DealerLog<V, P> {
958    fn write(&self, buf: &mut impl bytes::BufMut) {
959        self.pub_msg.write(buf);
960        self.results.write(buf);
961    }
962}
963
964impl<V: Variant, P: PublicKey> Read for DealerLog<V, P> {
965    type Cfg = NonZeroU32;
966
967    fn read_cfg(
968        buf: &mut impl bytes::Buf,
969        cfg: &Self::Cfg,
970    ) -> Result<Self, commonware_codec::Error> {
971        Ok(Self {
972            pub_msg: Read::read_cfg(buf, cfg)?,
973            results: Read::read_cfg(buf, cfg)?,
974        })
975    }
976}
977
978impl<V: Variant, P: PublicKey> DealerLog<V, P> {
979    fn get_reveal(&self, player: &P) -> Option<&DealerPrivMsg> {
980        let DealerResult::Ok(results) = &self.results else {
981            return None;
982        };
983        match results.get_value(player) {
984            Some(AckOrReveal::Reveal(priv_msg)) => Some(priv_msg),
985            _ => None,
986        }
987    }
988
989    fn zip_players<'a, 'b>(
990        &'a self,
991        players: &'b Set<P>,
992    ) -> Option<impl Iterator<Item = (&'b P, &'a AckOrReveal<P>)>> {
993        match &self.results {
994            DealerResult::TooManyReveals => None,
995            DealerResult::Ok(results) => {
996                // We don't check this on deserialization.
997                if results.keys() != players {
998                    return None;
999                }
1000                Some(players.iter().zip(results.values().iter()))
1001            }
1002        }
1003    }
1004}
1005
1006#[cfg(feature = "arbitrary")]
1007impl<V: Variant, P: PublicKey> arbitrary::Arbitrary<'_> for DealerLog<V, P>
1008where
1009    P: for<'a> arbitrary::Arbitrary<'a>,
1010    V::Public: for<'a> arbitrary::Arbitrary<'a>,
1011    P::Signature: for<'a> arbitrary::Arbitrary<'a>,
1012{
1013    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
1014        let pub_msg = u.arbitrary()?;
1015        let results = u.arbitrary()?;
1016        Ok(Self { pub_msg, results })
1017    }
1018}
1019
1020/// A [`DealerLog`], but identified to and signed by a dealer.
1021///
1022/// The [`SignedDealerLog::check`] method allows extracting a public key (the dealer)
1023/// and a [`DealerLog`] from this struct.
1024///
1025/// This avoids having to trust some other party or process for knowing that a
1026/// dealer actually produced a log.
1027#[derive(Clone, Debug)]
1028pub struct SignedDealerLog<V: Variant, S: Signer> {
1029    dealer: S::PublicKey,
1030    log: DealerLog<V, S::PublicKey>,
1031    sig: S::Signature,
1032}
1033
1034impl<V: Variant, S: Signer> PartialEq for SignedDealerLog<V, S> {
1035    fn eq(&self, other: &Self) -> bool {
1036        self.dealer == other.dealer && self.log == other.log && self.sig == other.sig
1037    }
1038}
1039
1040impl<V: Variant, S: Signer> SignedDealerLog<V, S> {
1041    fn sign(sk: &S, info: &Info<V, S::PublicKey>, log: DealerLog<V, S::PublicKey>) -> Self {
1042        let sig = transcript_for_log(info, &log).sign(sk);
1043        Self {
1044            dealer: sk.public_key(),
1045            log,
1046            sig,
1047        }
1048    }
1049
1050    /// Check this log for a particular round.
1051    ///
1052    /// This will produce the public key of the dealer that signed this log,
1053    /// and the underlying log that they signed.
1054    ///
1055    /// This will return [`Option::None`] if the check fails.
1056    #[allow(clippy::type_complexity)]
1057    pub fn check(
1058        self,
1059        info: &Info<V, S::PublicKey>,
1060    ) -> Option<(S::PublicKey, DealerLog<V, S::PublicKey>)> {
1061        if !transcript_for_log(info, &self.log).verify(&self.dealer, &self.sig) {
1062            return None;
1063        }
1064        Some((self.dealer, self.log))
1065    }
1066}
1067
1068impl<V: Variant, S: Signer> EncodeSize for SignedDealerLog<V, S> {
1069    fn encode_size(&self) -> usize {
1070        self.dealer.encode_size() + self.log.encode_size() + self.sig.encode_size()
1071    }
1072}
1073
1074impl<V: Variant, S: Signer> Write for SignedDealerLog<V, S> {
1075    fn write(&self, buf: &mut impl bytes::BufMut) {
1076        self.dealer.write(buf);
1077        self.log.write(buf);
1078        self.sig.write(buf);
1079    }
1080}
1081
1082impl<V: Variant, S: Signer> Read for SignedDealerLog<V, S> {
1083    type Cfg = NonZeroU32;
1084
1085    fn read_cfg(
1086        buf: &mut impl bytes::Buf,
1087        cfg: &Self::Cfg,
1088    ) -> Result<Self, commonware_codec::Error> {
1089        Ok(Self {
1090            dealer: ReadExt::read(buf)?,
1091            log: Read::read_cfg(buf, cfg)?,
1092            sig: ReadExt::read(buf)?,
1093        })
1094    }
1095}
1096
1097#[cfg(feature = "arbitrary")]
1098impl<V: Variant, S: Signer> arbitrary::Arbitrary<'_> for SignedDealerLog<V, S>
1099where
1100    S::PublicKey: for<'a> arbitrary::Arbitrary<'a>,
1101    V::Public: for<'a> arbitrary::Arbitrary<'a>,
1102    S::Signature: for<'a> arbitrary::Arbitrary<'a>,
1103{
1104    fn arbitrary(u: &mut arbitrary::Unstructured<'_>) -> arbitrary::Result<Self> {
1105        let dealer = u.arbitrary()?;
1106        let log = u.arbitrary()?;
1107        let sig = u.arbitrary()?;
1108        Ok(Self { dealer, log, sig })
1109    }
1110}
1111
1112fn transcript_for_round<V: Variant, P: PublicKey>(info: &Info<V, P>) -> Transcript {
1113    Transcript::resume(info.summary)
1114}
1115
1116fn transcript_for_ack<V: Variant, P: PublicKey>(
1117    transcript: &Transcript,
1118    dealer: &P,
1119    pub_msg: &DealerPubMsg<V>,
1120) -> Transcript {
1121    let mut out = transcript.fork(SIG_ACK);
1122    out.commit(dealer.encode());
1123    out.commit(pub_msg.encode());
1124    out
1125}
1126
1127fn transcript_for_log<V: Variant, P: PublicKey>(
1128    info: &Info<V, P>,
1129    log: &DealerLog<V, P>,
1130) -> Transcript {
1131    let mut out = transcript_for_round(info).fork(SIG_LOG);
1132    out.commit(log.encode());
1133    out
1134}
1135
1136pub struct Dealer<V: Variant, S: Signer> {
1137    me: S,
1138    info: Info<V, S::PublicKey>,
1139    pub_msg: DealerPubMsg<V>,
1140    results: Map<S::PublicKey, AckOrReveal<S::PublicKey>>,
1141    transcript: Transcript,
1142}
1143
1144impl<V: Variant, S: Signer> Dealer<V, S> {
1145    /// Create a [`Dealer`].
1146    ///
1147    /// This needs randomness, to generate a dealing.
1148    ///
1149    /// We also need the dealer's private key, in order to produce the [`SignedDealerLog`].
1150    ///
1151    /// If we're doing a reshare, the dealer should have a share from the previous round.
1152    ///
1153    /// This will produce the [`Dealer`], a [`DealerPubMsg`] to send to every player,
1154    /// and a list of [`DealerPrivMsg`]s, along with which players those need to
1155    /// be sent to.
1156    ///
1157    /// The public message can be sent in the clear, but it's important that players
1158    /// know which dealer sent what public message. You MUST ensure that dealers
1159    /// cannot impersonate each-other when sending this message.
1160    ///
1161    /// The private message MUST be sent encrypted (or, in some other way, privately)
1162    /// to the target player. Similarly, that player MUST be convinced that this dealer
1163    /// sent it that message, without any possibility of impersonation. A simple way
1164    /// to provide both guarantees is through an authenticated channel, e.g. via
1165    /// [crate::handshake], or [commonware-p2p](https://docs.rs/commonware-p2p/latest/commonware_p2p/).
1166    #[allow(clippy::type_complexity)]
1167    pub fn start(
1168        mut rng: impl CryptoRngCore,
1169        info: Info<V, S::PublicKey>,
1170        me: S,
1171        share: Option<Share>,
1172    ) -> Result<(Self, DealerPubMsg<V>, Vec<(S::PublicKey, DealerPrivMsg)>), Error> {
1173        // Check that this dealer is defined in the round.
1174        info.dealer_index(&me.public_key())?;
1175        let share = info.unwrap_or_random_share(&mut rng, share.map(|x| x.private))?;
1176        let my_poly = Poly::new_with_constant(&mut rng, info.degree(), share);
1177        let priv_msgs = info
1178            .players
1179            .iter()
1180            .map(|pk| {
1181                (
1182                    pk.clone(),
1183                    DealerPrivMsg {
1184                        share: my_poly
1185                            .eval_msm(&info.player_scalar(pk).expect("player should exist")),
1186                    },
1187                )
1188            })
1189            .collect::<Vec<_>>();
1190        let results: Map<_, _> = priv_msgs
1191            .clone()
1192            .into_iter()
1193            .map(|(pk, priv_msg)| (pk, AckOrReveal::Reveal(priv_msg)))
1194            .try_collect()
1195            .expect("players are unique");
1196        let commitment = Poly::commit(my_poly);
1197        let pub_msg = DealerPubMsg { commitment };
1198        let transcript = {
1199            let t = transcript_for_round(&info);
1200            transcript_for_ack(&t, &me.public_key(), &pub_msg)
1201        };
1202        let this = Self {
1203            me,
1204            info,
1205            pub_msg: pub_msg.clone(),
1206            results,
1207            transcript,
1208        };
1209        Ok((this, pub_msg, priv_msgs))
1210    }
1211
1212    /// Process an acknowledgement from a player.
1213    ///
1214    /// Acknowledgements should really only be processed once per player,
1215    /// but this method is idempotent nonetheless.
1216    pub fn receive_player_ack(
1217        &mut self,
1218        player: S::PublicKey,
1219        ack: PlayerAck<S::PublicKey>,
1220    ) -> Result<(), Error> {
1221        let res_mut = self
1222            .results
1223            .get_value_mut(&player)
1224            .ok_or(Error::UnknownPlayer)?;
1225        if self.transcript.verify(&player, &ack.sig) {
1226            *res_mut = AckOrReveal::Ack(ack);
1227        }
1228        Ok(())
1229    }
1230
1231    /// Finalize the dealer, producing a signed log.
1232    ///
1233    /// This should be called at the point where no more acks will be processed.
1234    pub fn finalize(self) -> SignedDealerLog<V, S> {
1235        let reveals = self
1236            .results
1237            .values()
1238            .iter()
1239            .filter(|x| x.is_reveal())
1240            .count() as u32;
1241        // Omit results if there are too many reveals.
1242        let results = if reveals > self.info.max_reveals() {
1243            DealerResult::TooManyReveals
1244        } else {
1245            DealerResult::Ok(self.results)
1246        };
1247        let log = DealerLog {
1248            pub_msg: self.pub_msg,
1249            results,
1250        };
1251        SignedDealerLog::sign(&self.me, &self.info, log)
1252    }
1253}
1254
1255#[allow(clippy::type_complexity)]
1256fn select<V: Variant, P: PublicKey>(
1257    info: &Info<V, P>,
1258    logs: BTreeMap<P, DealerLog<V, P>>,
1259) -> Result<Map<P, DealerLog<V, P>>, Error> {
1260    let required_commitments = info.required_commitments() as usize;
1261    let transcript = transcript_for_round(info);
1262    let out = logs
1263        .into_iter()
1264        .filter_map(|(dealer, log)| {
1265            info.dealer_index(&dealer).ok()?;
1266            if !info.check_dealer_pub_msg(&dealer, &log.pub_msg) {
1267                return None;
1268            }
1269            let results_iter = log.zip_players(&info.players)?;
1270            let transcript = transcript_for_ack(&transcript, &dealer, &log.pub_msg);
1271            let mut reveal_count = 0;
1272            let max_reveals = info.max_reveals();
1273            for (player, result) in results_iter {
1274                match result {
1275                    AckOrReveal::Ack(ack) => {
1276                        if !transcript.verify(player, &ack.sig) {
1277                            return None;
1278                        }
1279                    }
1280                    AckOrReveal::Reveal(priv_msg) => {
1281                        reveal_count += 1;
1282                        if reveal_count > max_reveals {
1283                            return None;
1284                        }
1285                        if !info.check_dealer_priv_msg(player, &log.pub_msg, priv_msg) {
1286                            return None;
1287                        }
1288                    }
1289                }
1290            }
1291            Some((dealer, log))
1292        })
1293        .take(required_commitments)
1294        .try_collect::<Map<_, _>>()
1295        .expect("logs has at most one entry per dealer");
1296    if out.len() < required_commitments {
1297        return Err(Error::DkgFailed);
1298    }
1299    Ok(out)
1300}
1301
1302struct ObserveInner<V: Variant, P: PublicKey> {
1303    output: Output<V, P>,
1304    weights: Option<Interpolator<P, Scalar>>,
1305}
1306
1307impl<V: Variant, P: PublicKey> ObserveInner<V, P> {
1308    fn reckon(
1309        info: Info<V, P>,
1310        selected: Map<P, DealerLog<V, P>>,
1311        concurrency: usize,
1312    ) -> Result<Self, Error> {
1313        // Track players with too many reveals
1314        let max_faults = info.players.max_faults();
1315        let mut reveal_counts: BTreeMap<P, u32> = BTreeMap::new();
1316        let mut revealed = Vec::new();
1317        for log in selected.values() {
1318            let Some(iter) = log.zip_players(&info.players) else {
1319                continue;
1320            };
1321            for (player, result) in iter {
1322                if !result.is_reveal() {
1323                    continue;
1324                }
1325                let count = reveal_counts.entry(player.clone()).or_insert(0);
1326                *count += 1;
1327                if *count == max_faults + 1 {
1328                    revealed.push(player.clone());
1329                }
1330            }
1331        }
1332        let revealed: Set<P> = revealed
1333            .into_iter()
1334            .try_collect()
1335            .expect("players are unique");
1336
1337        // Extract dealers before consuming selected
1338        let dealers: Set<P> = selected
1339            .keys()
1340            .iter()
1341            .cloned()
1342            .try_collect()
1343            .expect("selected dealers are unique");
1344
1345        // Recover the public polynomial
1346        let (public, weights) = if let Some(previous) = info.previous.as_ref() {
1347            let weights = previous
1348                .public()
1349                .mode()
1350                .subset_interpolator(previous.players(), selected.keys())
1351                .expect("the result of select should produce a valid subset");
1352            let commitments = selected
1353                .into_iter()
1354                .map(|(dealer, log)| (dealer, log.pub_msg.commitment))
1355                .try_collect::<Map<_, _>>()
1356                .expect("Map should have unique keys");
1357            let public = weights
1358                .interpolate(&commitments, concurrency)
1359                .expect("select checks that enough points have been provided");
1360            if previous.public().public() != public.constant() {
1361                return Err(Error::DkgFailed);
1362            }
1363            (public, Some(weights))
1364        } else {
1365            let mut public = Poly::zero();
1366            for log in selected.values() {
1367                public += &log.pub_msg.commitment;
1368            }
1369            (public, None)
1370        };
1371        let n = info.players.len() as u32;
1372        let output = Output {
1373            summary: info.summary,
1374            public: Sharing::new(info.mode, NZU32!(n), public),
1375            dealers,
1376            players: info.players,
1377            revealed,
1378        };
1379        Ok(Self { output, weights })
1380    }
1381}
1382
1383/// Observe the result of a DKG, using the public results.
1384///
1385/// The log mapping dealers to their log is the shared piece of information
1386/// that the participants (players, observers) of the DKG must all agree on.
1387///
1388/// From this log, we can (potentially, as the DKG can fail) compute the public output.
1389///
1390/// This will only ever return [`Error::DkgFailed`].
1391pub fn observe<V: Variant, P: PublicKey>(
1392    info: Info<V, P>,
1393    logs: BTreeMap<P, DealerLog<V, P>>,
1394    concurrency: usize,
1395) -> Result<Output<V, P>, Error> {
1396    let selected = select(&info, logs)?;
1397    ObserveInner::<V, P>::reckon(info, selected, concurrency).map(|x| x.output)
1398}
1399
1400/// Represents a player in the DKG / reshare process.
1401///
1402/// The player is attempting to get a share of the key.
1403///
1404/// They need not have participated in prior rounds.
1405pub struct Player<V: Variant, S: Signer> {
1406    me: S,
1407    me_pub: S::PublicKey,
1408    info: Info<V, S::PublicKey>,
1409    index: u32,
1410    transcript: Transcript,
1411    view: BTreeMap<S::PublicKey, (DealerPubMsg<V>, DealerPrivMsg)>,
1412}
1413
1414impl<V: Variant, S: Signer> Player<V, S> {
1415    /// Create a new [`Player`].
1416    ///
1417    /// We need the player's private key in order to sign messages.
1418    pub fn new(info: Info<V, S::PublicKey>, me: S) -> Result<Self, Error> {
1419        let me_pub = me.public_key();
1420        Ok(Self {
1421            index: info.player_index(&me_pub)?,
1422            me,
1423            me_pub,
1424            transcript: transcript_for_round(&info),
1425            info,
1426            view: BTreeMap::new(),
1427        })
1428    }
1429
1430    /// Process a message from a dealer.
1431    ///
1432    /// It's important that nobody can impersonate the dealer, and that the
1433    /// private message was not exposed to anyone else. A convenient way to
1434    /// provide this is by using an authenticated channel, e.g. via
1435    /// [crate::handshake], or [commonware-p2p](https://docs.rs/commonware-p2p/latest/commonware_p2p/).
1436    pub fn dealer_message(
1437        &mut self,
1438        dealer: S::PublicKey,
1439        pub_msg: DealerPubMsg<V>,
1440        priv_msg: DealerPrivMsg,
1441    ) -> Option<PlayerAck<S::PublicKey>> {
1442        if self.view.contains_key(&dealer) {
1443            return None;
1444        }
1445        self.info.dealer_index(&dealer).ok()?;
1446        if !self.info.check_dealer_pub_msg(&dealer, &pub_msg) {
1447            return None;
1448        }
1449        if !self
1450            .info
1451            .check_dealer_priv_msg(&self.me_pub, &pub_msg, &priv_msg)
1452        {
1453            return None;
1454        }
1455        let sig = transcript_for_ack(&self.transcript, &dealer, &pub_msg).sign(&self.me);
1456        self.view.insert(dealer, (pub_msg, priv_msg));
1457        Some(PlayerAck { sig })
1458    }
1459
1460    /// Finalize the player, producing an output, and a share.
1461    ///
1462    /// This should agree with [`observe`], in terms of `Ok` vs `Err` and the
1463    /// public output, so long as the logs agree. It's crucial that the players
1464    /// come to agreement, in some way, on exactly which logs they need to use
1465    /// for finalize.
1466    ///
1467    /// This will only ever return [`Error::DkgFailed`].
1468    pub fn finalize(
1469        self,
1470        logs: BTreeMap<S::PublicKey, DealerLog<V, S::PublicKey>>,
1471        concurrency: usize,
1472    ) -> Result<(Output<V, S::PublicKey>, Share), Error> {
1473        let selected = select(&self.info, logs)?;
1474        let dealings = selected
1475            .iter_pairs()
1476            .map(|(dealer, log)| {
1477                let share = self
1478                    .view
1479                    .get(dealer)
1480                    .map(|(_, priv_msg)| priv_msg.share.clone())
1481                    .unwrap_or_else(|| {
1482                        log.get_reveal(&self.me_pub).map_or_else(
1483                            || {
1484                                unreachable!(
1485                                    "select didn't check dealer reveal, or we're not a player?"
1486                                )
1487                            },
1488                            |priv_msg| priv_msg.share.clone(),
1489                        )
1490                    });
1491                (dealer.clone(), share)
1492            })
1493            .try_collect::<Map<_, _>>()
1494            .expect("select produces at most one entry per dealer");
1495        let ObserveInner { output, weights } =
1496            ObserveInner::<V, S::PublicKey>::reckon(self.info, selected, concurrency)?;
1497        let private = weights.map_or_else(
1498            || {
1499                let mut out = <Scalar as Additive>::zero();
1500                for s in dealings.values() {
1501                    out += s;
1502                }
1503                out
1504            },
1505            |weights| {
1506                weights
1507                    .interpolate(&dealings, concurrency)
1508                    .expect("select ensures that we can recover")
1509            },
1510        );
1511        let share = Share {
1512            index: self.index,
1513            private,
1514        };
1515        Ok((output, share))
1516    }
1517}
1518
1519/// The result of dealing shares to players.
1520pub type DealResult<V, P> = Result<(Output<V, P>, Map<P, Share>), Error>;
1521
1522/// Simply distribute shares at random, instead of performing a distributed protocol.
1523pub fn deal<V: Variant, P: Clone + Ord>(
1524    mut rng: impl CryptoRngCore,
1525    mode: Mode,
1526    players: Set<P>,
1527) -> DealResult<V, P> {
1528    if players.is_empty() {
1529        return Err(Error::NumPlayers(0));
1530    }
1531    let n = NZU32!(players.len() as u32);
1532    let t = players.quorum();
1533    let private = Poly::new(&mut rng, t - 1);
1534    let shares: Map<_, _> = players
1535        .iter()
1536        .enumerate()
1537        .map(|(i, p)| {
1538            let i = i as u32;
1539            let eval = private.eval_msm(&mode.scalar(n, i).expect("player index should be valid"));
1540            let share = Share {
1541                index: i,
1542                private: eval,
1543            };
1544            (p.clone(), share)
1545        })
1546        .try_collect()
1547        .expect("players are unique");
1548    let output = Output {
1549        summary: Summary::random(&mut rng),
1550        public: Sharing::new(mode, n, Poly::commit(private)),
1551        dealers: players.clone(),
1552        players,
1553        revealed: Set::default(),
1554    };
1555    Ok((output, shares))
1556}
1557
1558/// Like [`deal`], but without linking the result to specific public keys.
1559///
1560/// This can be more convenient for testing, where you don't want to go through
1561/// the trouble of generating signing keys. The downside is that the result isn't
1562/// compatible with subsequent DKGs, which need an [`Output`].
1563pub fn deal_anonymous<V: Variant>(
1564    rng: impl CryptoRngCore,
1565    mode: Mode,
1566    n: NonZeroU32,
1567) -> (Sharing<V>, Vec<Share>) {
1568    let players = (0..n.get()).try_collect().unwrap();
1569    let (output, shares) = deal::<V, _>(rng, mode, players).unwrap();
1570    (output.public().clone(), shares.values().to_vec())
1571}
1572
1573#[cfg(any(feature = "arbitrary", test))]
1574mod test_plan {
1575    use super::*;
1576    use crate::{
1577        bls12381::primitives::{
1578            ops::{
1579                partial_sign_message, partial_verify_message, threshold_signature_recover,
1580                verify_message,
1581            },
1582            variant::Variant,
1583        },
1584        ed25519, PublicKey,
1585    };
1586    use anyhow::anyhow;
1587    use bytes::BytesMut;
1588    use commonware_utils::{max_faults, quorum, TryCollect};
1589    use core::num::NonZeroI32;
1590    use rand::{rngs::StdRng, SeedableRng as _};
1591    use std::collections::BTreeSet;
1592
1593    /// Apply a mask to some bytes, returning whether or not a modification happened
1594    fn apply_mask(bytes: &mut BytesMut, mask: &[u8]) -> bool {
1595        let mut modified = false;
1596        for (l, &r) in bytes.iter_mut().zip(mask.iter()) {
1597            modified |= r != 0;
1598            *l ^= r;
1599        }
1600        modified
1601    }
1602
1603    #[derive(Clone, Default, Debug)]
1604    pub struct Masks {
1605        pub info_summary: Vec<u8>,
1606        pub dealer: Vec<u8>,
1607        pub pub_msg: Vec<u8>,
1608        pub log: Vec<u8>,
1609    }
1610
1611    impl Masks {
1612        fn transcript_for_round<V: Variant, P: PublicKey>(
1613            &self,
1614            info: &Info<V, P>,
1615        ) -> anyhow::Result<(bool, Transcript)> {
1616            let mut summary_bs = info.summary.encode();
1617            let modified = apply_mask(&mut summary_bs, &self.info_summary);
1618            let summary = Summary::read(&mut summary_bs)?;
1619            Ok((modified, Transcript::resume(summary)))
1620        }
1621
1622        fn transcript_for_player_ack<V: Variant, P: PublicKey>(
1623            &self,
1624            info: &Info<V, P>,
1625            dealer: &P,
1626            pub_msg: &DealerPubMsg<V>,
1627        ) -> anyhow::Result<(bool, Transcript)> {
1628            let (mut modified, transcript) = self.transcript_for_round(info)?;
1629            let mut transcript = transcript.fork(SIG_ACK);
1630
1631            let mut dealer_bs = dealer.encode();
1632            modified |= apply_mask(&mut dealer_bs, &self.dealer);
1633            transcript.commit(&mut dealer_bs);
1634
1635            let mut pub_msg_bs = pub_msg.encode();
1636            modified |= apply_mask(&mut pub_msg_bs, &self.pub_msg);
1637            transcript.commit(&mut pub_msg_bs);
1638
1639            Ok((modified, transcript))
1640        }
1641
1642        fn transcript_for_signed_dealer_log<V: Variant, P: PublicKey>(
1643            &self,
1644            info: &Info<V, P>,
1645            log: &DealerLog<V, P>,
1646        ) -> anyhow::Result<(bool, Transcript)> {
1647            let (mut modified, transcript) = self.transcript_for_round(info)?;
1648            let mut transcript = transcript.fork(SIG_LOG);
1649
1650            let mut log_bs = log.encode();
1651            modified |= apply_mask(&mut log_bs, &self.log);
1652            transcript.commit(&mut log_bs);
1653
1654            Ok((modified, transcript))
1655        }
1656    }
1657
1658    /// A round in the DKG test plan.
1659    #[derive(Debug, Default)]
1660    pub struct Round {
1661        dealers: Vec<u32>,
1662        players: Vec<u32>,
1663        no_acks: BTreeSet<(u32, u32)>,
1664        bad_shares: BTreeSet<(u32, u32)>,
1665        bad_player_sigs: BTreeMap<(u32, u32), Masks>,
1666        bad_reveals: BTreeSet<(u32, u32)>,
1667        bad_dealer_sigs: BTreeMap<u32, Masks>,
1668        replace_shares: BTreeSet<u32>,
1669        shift_degrees: BTreeMap<u32, NonZeroI32>,
1670    }
1671
1672    impl Round {
1673        pub fn new(dealers: Vec<u32>, players: Vec<u32>) -> Self {
1674            Self {
1675                dealers,
1676                players,
1677                ..Default::default()
1678            }
1679        }
1680
1681        pub fn no_ack(mut self, dealer: u32, player: u32) -> Self {
1682            self.no_acks.insert((dealer, player));
1683            self
1684        }
1685
1686        pub fn bad_share(mut self, dealer: u32, player: u32) -> Self {
1687            self.bad_shares.insert((dealer, player));
1688            self
1689        }
1690
1691        pub fn bad_player_sig(mut self, dealer: u32, player: u32, masks: Masks) -> Self {
1692            self.bad_player_sigs.insert((dealer, player), masks);
1693            self
1694        }
1695
1696        pub fn bad_reveal(mut self, dealer: u32, player: u32) -> Self {
1697            self.bad_reveals.insert((dealer, player));
1698            self
1699        }
1700
1701        pub fn bad_dealer_sig(mut self, dealer: u32, masks: Masks) -> Self {
1702            self.bad_dealer_sigs.insert(dealer, masks);
1703            self
1704        }
1705
1706        pub fn replace_share(mut self, dealer: u32) -> Self {
1707            self.replace_shares.insert(dealer);
1708            self
1709        }
1710
1711        pub fn shift_degree(mut self, dealer: u32, shift: NonZeroI32) -> Self {
1712            self.shift_degrees.insert(dealer, shift);
1713            self
1714        }
1715
1716        /// Validate that this round is well-formed given the number of participants
1717        /// and the previous successful round's players.
1718        pub fn validate(
1719            &self,
1720            num_participants: u32,
1721            previous_players: Option<&[u32]>,
1722        ) -> anyhow::Result<()> {
1723            if self.dealers.is_empty() {
1724                return Err(anyhow!("dealers is empty"));
1725            }
1726            if self.players.is_empty() {
1727                return Err(anyhow!("players is empty"));
1728            }
1729            // Check dealer/player ranges
1730            for &d in &self.dealers {
1731                if d >= num_participants {
1732                    return Err(anyhow!("dealer {d} out of range [1, {num_participants}]"));
1733                }
1734            }
1735            for &p in &self.players {
1736                if p >= num_participants {
1737                    return Err(anyhow!("player {p} out of range [1, {num_participants}]"));
1738                }
1739            }
1740
1741            // If there's a previous round, check dealer constraints
1742            if let Some(prev_players) = previous_players {
1743                // Every dealer must have been a player in the previous round
1744                for &d in &self.dealers {
1745                    if !prev_players.contains(&d) {
1746                        return Err(anyhow!("dealer {d} was not a player in previous round"));
1747                    }
1748                }
1749                // Must have >= quorum(prev_players) dealers
1750                let required = quorum(prev_players.len() as u32);
1751                if (self.dealers.len() as u32) < required {
1752                    return Err(anyhow!(
1753                        "not enough dealers: have {}, need {} (quorum of {} previous players)",
1754                        self.dealers.len(),
1755                        required,
1756                        prev_players.len()
1757                    ));
1758                }
1759            }
1760
1761            Ok(())
1762        }
1763
1764        fn bad(&self, previous_successful_round: bool, dealer: u32) -> bool {
1765            if self.replace_shares.contains(&dealer) && previous_successful_round {
1766                return true;
1767            }
1768            if let Some(shift) = self.shift_degrees.get(&dealer) {
1769                let degree = quorum(self.players.len() as u32) as i32 - 1;
1770                // We shift the degree, but saturate at 0, so it's possible
1771                // that the shift isn't actually doing anything.
1772                //
1773                // This is effectively the same as checking degree == 0 && shift < 0,
1774                // but matches what ends up happening a bit better.
1775                if (degree + shift.get()).max(0) != degree {
1776                    return true;
1777                }
1778            }
1779            if self.bad_reveals.iter().any(|&(d, _)| d == dealer) {
1780                return true;
1781            }
1782            let revealed_players = self
1783                .bad_shares
1784                .iter()
1785                .copied()
1786                .chain(self.no_acks.iter().copied())
1787                .filter_map(|(d, p)| if d == dealer { Some(p) } else { None })
1788                .collect::<BTreeSet<_>>();
1789            revealed_players.len() as u32 > max_faults(self.players.len() as u32)
1790        }
1791
1792        /// Determine if this round is expected to fail.
1793        fn expect_failure(&self, previous_successful_round: Option<u32>) -> bool {
1794            let good_dealer_count = self
1795                .dealers
1796                .iter()
1797                .filter(|&&d| !self.bad(previous_successful_round.is_some(), d))
1798                .count();
1799            let required = previous_successful_round
1800                .map(quorum)
1801                .unwrap_or_default()
1802                .max(quorum(self.dealers.len() as u32)) as usize;
1803            good_dealer_count < required
1804        }
1805    }
1806
1807    /// A DKG test plan consisting of multiple rounds.
1808    #[derive(Debug)]
1809    pub struct Plan {
1810        num_participants: NonZeroU32,
1811        rounds: Vec<Round>,
1812    }
1813
1814    impl Plan {
1815        pub const fn new(num_participants: NonZeroU32) -> Self {
1816            Self {
1817                num_participants,
1818                rounds: Vec::new(),
1819            }
1820        }
1821
1822        pub fn with(mut self, round: Round) -> Self {
1823            self.rounds.push(round);
1824            self
1825        }
1826
1827        /// Validate the entire plan.
1828        fn validate(&self) -> anyhow::Result<()> {
1829            let mut last_successful_players: Option<Vec<u32>> = None;
1830
1831            for round in &self.rounds {
1832                round.validate(
1833                    self.num_participants.get(),
1834                    last_successful_players.as_deref(),
1835                )?;
1836
1837                // If this round is expected to succeed, update last_successful_players
1838                if !round.expect_failure(last_successful_players.as_ref().map(|x| x.len() as u32)) {
1839                    last_successful_players = Some(round.players.clone());
1840                }
1841            }
1842            Ok(())
1843        }
1844
1845        /// Run the test plan with a given seed.
1846        pub fn run<V: Variant>(self, seed: u64) -> anyhow::Result<()> {
1847            self.validate()?;
1848
1849            let mut rng = StdRng::seed_from_u64(seed);
1850
1851            // Generate keys for all participants (1-indexed to num_participants)
1852            let keys = (0..self.num_participants.get())
1853                .map(|_| ed25519::PrivateKey::random(&mut rng))
1854                .collect::<Vec<_>>();
1855
1856            // Precompute mapping from public key to key index to avoid confusion
1857            // between key indices and positions in sorted Sets.
1858            let pk_to_key_idx: BTreeMap<ed25519::PublicKey, u32> = keys
1859                .iter()
1860                .enumerate()
1861                .map(|(i, k)| (k.public_key(), i as u32))
1862                .collect();
1863
1864            // The max_read_size needs to account for shifted polynomial degrees.
1865            // Find the maximum positive shift across all rounds.
1866            let max_shift = self
1867                .rounds
1868                .iter()
1869                .flat_map(|r| r.shift_degrees.values())
1870                .map(|s| s.get())
1871                .max()
1872                .unwrap_or(0)
1873                .max(0) as u32;
1874            let max_read_size =
1875                NonZeroU32::new(self.num_participants.get() + max_shift).expect("non-zero");
1876
1877            let mut previous_output: Option<Output<V, ed25519::PublicKey>> = None;
1878            let mut shares: BTreeMap<ed25519::PublicKey, Share> = BTreeMap::new();
1879            let mut threshold_public_key: Option<V::Public> = None;
1880
1881            for (i_round, round) in self.rounds.into_iter().enumerate() {
1882                let previous_successful_round =
1883                    previous_output.as_ref().map(|o| o.players.len() as u32);
1884
1885                let dealer_set = round
1886                    .dealers
1887                    .iter()
1888                    .map(|&i| keys[i as usize].public_key())
1889                    .try_collect::<Set<_>>()
1890                    .unwrap();
1891                let player_set: Set<ed25519::PublicKey> = round
1892                    .players
1893                    .iter()
1894                    .map(|&i| keys[i as usize].public_key())
1895                    .try_collect()
1896                    .unwrap();
1897
1898                // Create round info
1899                let info = Info::new(
1900                    &[],
1901                    i_round as u64,
1902                    previous_output.clone(),
1903                    Default::default(),
1904                    dealer_set.clone(),
1905                    player_set.clone(),
1906                )?;
1907
1908                let mut players: Map<_, _> = round
1909                    .players
1910                    .iter()
1911                    .map(|&i| {
1912                        let sk = keys[i as usize].clone();
1913                        let pk = sk.public_key();
1914                        let player = Player::new(info.clone(), sk)?;
1915                        Ok((pk, player))
1916                    })
1917                    .collect::<anyhow::Result<Vec<_>>>()?
1918                    .try_into()
1919                    .unwrap();
1920
1921                // Run dealer protocol
1922                let mut dealer_logs = BTreeMap::new();
1923                for &i_dealer in &round.dealers {
1924                    let sk = keys[i_dealer as usize].clone();
1925                    let pk = sk.public_key();
1926                    let share = match (shares.get(&pk), round.replace_shares.contains(&i_dealer)) {
1927                        (None, _) => None,
1928                        (Some(s), false) => Some(s.clone()),
1929                        (Some(_), true) => Some(Share {
1930                            index: i_dealer,
1931                            private: Scalar::random(&mut rng),
1932                        }),
1933                    };
1934
1935                    // Start dealer (with potential modifications)
1936                    let (mut dealer, pub_msg, mut priv_msgs) = if let Some(shift) =
1937                        round.shift_degrees.get(&i_dealer)
1938                    {
1939                        // Create dealer with shifted degree
1940                        let degree =
1941                            u32::try_from(info.degree() as i32 + shift.get()).unwrap_or_default();
1942
1943                        // Manually create the dealer with adjusted polynomial
1944                        let share = info
1945                            .unwrap_or_random_share(&mut rng, share.map(|s| s.private))
1946                            .expect("Failed to generate dealer share");
1947
1948                        let my_poly = Poly::new_with_constant(&mut rng, degree, share);
1949                        let priv_msgs = info
1950                            .players
1951                            .iter()
1952                            .map(|pk| {
1953                                (
1954                                    pk.clone(),
1955                                    DealerPrivMsg {
1956                                        share: my_poly.eval_msm(
1957                                            &info.player_scalar(pk).expect("player should exist"),
1958                                        ),
1959                                    },
1960                                )
1961                            })
1962                            .collect::<Vec<_>>();
1963                        let results: Map<_, _> = priv_msgs
1964                            .iter()
1965                            .map(|(pk, pm)| (pk.clone(), AckOrReveal::Reveal(pm.clone())))
1966                            .try_collect()
1967                            .unwrap();
1968                        let commitment = Poly::commit(my_poly);
1969                        let pub_msg = DealerPubMsg { commitment };
1970                        let transcript = {
1971                            let t = transcript_for_round(&info);
1972                            transcript_for_ack(&t, &pk, &pub_msg)
1973                        };
1974                        let dealer = Dealer {
1975                            me: sk.clone(),
1976                            info: info.clone(),
1977                            pub_msg: pub_msg.clone(),
1978                            results,
1979                            transcript,
1980                        };
1981                        (dealer, pub_msg, priv_msgs)
1982                    } else {
1983                        Dealer::start(&mut rng, info.clone(), sk.clone(), share)?
1984                    };
1985
1986                    // Apply BadShare perturbations
1987                    for (player, priv_msg) in &mut priv_msgs {
1988                        let player_key_idx = pk_to_key_idx[player];
1989                        if round.bad_shares.contains(&(i_dealer, player_key_idx)) {
1990                            priv_msg.share = Scalar::random(&mut rng);
1991                        }
1992                    }
1993                    assert_eq!(priv_msgs.len(), players.len());
1994
1995                    // Process player acks
1996                    let mut num_reveals = players.len() as u32;
1997                    for (player_pk, priv_msg) in priv_msgs {
1998                        // Check priv msg encoding.
1999                        assert_eq!(priv_msg, ReadExt::read(&mut priv_msg.encode())?);
2000
2001                        let i_player = players
2002                            .index(&player_pk)
2003                            .ok_or_else(|| anyhow!("unknown player: {:?}", &player_pk))?;
2004                        let player_key_idx = pk_to_key_idx[&player_pk];
2005                        let player = &mut players.values_mut()[i_player as usize];
2006
2007                        let ack = player.dealer_message(pk.clone(), pub_msg.clone(), priv_msg);
2008                        assert_eq!(ack, ReadExt::read(&mut ack.encode())?);
2009                        if let Some(ack) = ack {
2010                            let masks = round
2011                                .bad_player_sigs
2012                                .get(&(i_dealer, player_key_idx))
2013                                .cloned()
2014                                .unwrap_or_default();
2015                            let (modified, transcript) =
2016                                masks.transcript_for_player_ack(&info, &pk, &pub_msg)?;
2017                            assert_eq!(transcript.verify(&player_pk, &ack.sig), !modified);
2018
2019                            // Skip receiving ack if NoAck perturbation
2020                            if !round.no_acks.contains(&(i_dealer, player_key_idx)) {
2021                                dealer.receive_player_ack(player_pk, ack)?;
2022                                num_reveals -= 1;
2023                            }
2024                        } else {
2025                            assert!(
2026                                round.bad_shares.contains(&(i_dealer, player_key_idx))
2027                                    || round.bad(previous_successful_round.is_some(), i_dealer)
2028                            );
2029                        }
2030                    }
2031
2032                    // Finalize dealer
2033                    let signed_log = dealer.finalize();
2034                    assert_eq!(
2035                        signed_log,
2036                        Read::read_cfg(&mut signed_log.encode(), &max_read_size)?
2037                    );
2038
2039                    // Check for BadDealerSig
2040                    let masks = round
2041                        .bad_dealer_sigs
2042                        .get(&i_dealer)
2043                        .cloned()
2044                        .unwrap_or_default();
2045                    let (modified, transcript) =
2046                        masks.transcript_for_signed_dealer_log(&info, &signed_log.log)?;
2047                    assert_eq!(transcript.verify(&pk, &signed_log.sig), !modified);
2048                    let (found_pk, mut log) = signed_log
2049                        .check(&info)
2050                        .ok_or_else(|| anyhow!("signed log should verify"))?;
2051                    assert_eq!(pk, found_pk);
2052                    // Apply BadReveal perturbations
2053                    match &mut log.results {
2054                        DealerResult::TooManyReveals => {
2055                            assert!(num_reveals > info.max_reveals());
2056                        }
2057                        DealerResult::Ok(results) => {
2058                            assert_eq!(results.len(), players.len());
2059                            for &i_player in &round.players {
2060                                if !round.bad_reveals.contains(&(i_dealer, i_player)) {
2061                                    continue;
2062                                }
2063                                let player_pk = keys[i_player as usize].public_key();
2064                                *results
2065                                    .get_value_mut(&player_pk)
2066                                    .ok_or_else(|| anyhow!("unknown player: {:?}", &player_pk))? =
2067                                    AckOrReveal::Reveal(DealerPrivMsg {
2068                                        share: Scalar::random(&mut rng),
2069                                    });
2070                            }
2071                        }
2072                    }
2073                    dealer_logs.insert(pk, log);
2074                }
2075
2076                // Make sure that bad dealers are not selected.
2077                let selection = select(&info, dealer_logs.clone());
2078                if let Ok(ref selection) = selection {
2079                    let good_pks = selection
2080                        .iter_pairs()
2081                        .map(|(pk, _)| pk.clone())
2082                        .collect::<BTreeSet<_>>();
2083                    for &i_dealer in &round.dealers {
2084                        if round.bad(previous_successful_round.is_some(), i_dealer) {
2085                            assert!(!good_pks.contains(&keys[i_dealer as usize].public_key()));
2086                        }
2087                    }
2088                }
2089                // Run observer
2090                let observe_result = observe(info.clone(), dealer_logs.clone(), 1);
2091                if round.expect_failure(previous_successful_round) {
2092                    assert!(
2093                        observe_result.is_err(),
2094                        "Round {i_round} should have failed but succeeded",
2095                    );
2096                    continue;
2097                }
2098                let observer_output = observe_result?;
2099                let selection = selection.expect("select should succeed if observe succeeded");
2100
2101                // Compute expected dealers: good dealers up to required_commitments
2102                // The select function iterates dealer_logs (BTreeMap) in public key order
2103                let required_commitments = info.required_commitments() as usize;
2104                let expected_dealers: Set<ed25519::PublicKey> = dealer_set
2105                    .iter()
2106                    .filter(|pk| {
2107                        let i = keys.iter().position(|k| &k.public_key() == *pk).unwrap() as u32;
2108                        !round.bad(previous_successful_round.is_some(), i)
2109                    })
2110                    .take(required_commitments)
2111                    .cloned()
2112                    .try_collect()
2113                    .expect("dealers are unique");
2114                let expected_dealer_indices: BTreeSet<u32> = expected_dealers
2115                    .iter()
2116                    .filter_map(|pk| {
2117                        keys.iter()
2118                            .position(|k| &k.public_key() == pk)
2119                            .map(|i| i as u32)
2120                    })
2121                    .collect();
2122                assert_eq!(
2123                    observer_output.dealers(),
2124                    &expected_dealers,
2125                    "Output dealers should match expected good dealers"
2126                );
2127
2128                // Map selected dealers to their key indices (for later use)
2129                let selected_dealers: BTreeSet<u32> = selection
2130                    .keys()
2131                    .iter()
2132                    .filter_map(|pk| {
2133                        keys.iter()
2134                            .position(|k| &k.public_key() == pk)
2135                            .map(|i| i as u32)
2136                    })
2137                    .collect();
2138                assert_eq!(
2139                    selected_dealers, expected_dealer_indices,
2140                    "Selection should match expected dealers"
2141                );
2142                let selected_players: Set<ed25519::PublicKey> = round
2143                    .players
2144                    .iter()
2145                    .map(|&i| keys[i as usize].public_key())
2146                    .try_collect()
2147                    .expect("players are unique");
2148
2149                // Compute expected reveals
2150                let mut expected_reveals: BTreeMap<ed25519::PublicKey, u32> = BTreeMap::new();
2151                for &(dealer_idx, player_key_idx) in
2152                    round.no_acks.iter().chain(round.bad_shares.iter())
2153                {
2154                    if !selected_dealers.contains(&dealer_idx) {
2155                        continue;
2156                    }
2157                    let pk = keys[player_key_idx as usize].public_key();
2158                    if selected_players.position(&pk).is_none() {
2159                        continue;
2160                    }
2161                    *expected_reveals.entry(pk).or_insert(0) += 1;
2162                }
2163
2164                // Verify each player's revealed status
2165                let max_faults = selected_players.max_faults();
2166                for player in player_set.iter() {
2167                    let expected = expected_reveals.get(player).copied().unwrap_or(0) > max_faults;
2168                    let actual = observer_output.revealed().position(player).is_some();
2169                    assert_eq!(expected, actual, "Unexpected outcome for player {player:?} (expected={expected}, actual={actual})");
2170                }
2171
2172                // Finalize each player
2173                for (player_pk, player) in players.into_iter() {
2174                    let (player_output, share) = player
2175                        .finalize(dealer_logs.clone(), 1)
2176                        .expect("Player finalize should succeed");
2177
2178                    assert_eq!(
2179                        player_output, observer_output,
2180                        "Player output should match observer output"
2181                    );
2182
2183                    // Verify share matches public polynomial
2184                    let expected_public = observer_output
2185                        .public
2186                        .partial_public(share.index)
2187                        .expect("share index should be valid");
2188                    let actual_public = share.public::<V>();
2189                    assert_eq!(
2190                        expected_public, actual_public,
2191                        "Share should match public polynomial"
2192                    );
2193
2194                    shares.insert(player_pk.clone(), share);
2195                }
2196
2197                // Initialize or verify threshold public key
2198                let current_public = *observer_output.public().public();
2199                match threshold_public_key {
2200                    None => threshold_public_key = Some(current_public),
2201                    Some(tpk) => {
2202                        assert_eq!(
2203                            tpk, current_public,
2204                            "Public key should remain constant across reshares"
2205                        );
2206                    }
2207                }
2208
2209                // Generate and verify threshold signature
2210                let test_message = format!("test message round {i_round}").into_bytes();
2211                let namespace = Some(&b"test"[..]);
2212
2213                let mut partial_sigs = Vec::new();
2214                for &i_player in &round.players {
2215                    let share = &shares[&keys[i_player as usize].public_key()];
2216                    let partial_sig = partial_sign_message::<V>(share, namespace, &test_message);
2217
2218                    partial_verify_message::<V>(
2219                        &observer_output.public,
2220                        namespace,
2221                        &test_message,
2222                        &partial_sig,
2223                    )
2224                    .expect("Partial signature verification should succeed");
2225
2226                    partial_sigs.push(partial_sig);
2227                }
2228
2229                let threshold = observer_output.quorum();
2230                let threshold_sig = threshold_signature_recover::<V, _>(
2231                    &observer_output.public,
2232                    &partial_sigs[0..threshold as usize],
2233                )
2234                .expect("Should recover threshold signature");
2235
2236                // Verify against the saved public key
2237                verify_message::<V>(
2238                    threshold_public_key.as_ref().unwrap(),
2239                    namespace,
2240                    &test_message,
2241                    &threshold_sig,
2242                )
2243                .expect("Threshold signature verification should succeed");
2244
2245                // Update state for next round
2246                previous_output = Some(observer_output);
2247            }
2248            Ok(())
2249        }
2250    }
2251
2252    #[cfg(feature = "arbitrary")]
2253    mod impl_arbitrary {
2254        use super::*;
2255        use arbitrary::{Arbitrary, Unstructured};
2256        use core::ops::ControlFlow;
2257
2258        const MAX_NUM_PARTICIPANTS: u32 = 20;
2259        const MAX_ROUNDS: u32 = 10;
2260
2261        fn arbitrary_masks<'a>(u: &mut Unstructured<'a>) -> arbitrary::Result<Masks> {
2262            Ok(Masks {
2263                info_summary: Arbitrary::arbitrary(u)?,
2264                dealer: Arbitrary::arbitrary(u)?,
2265                pub_msg: Arbitrary::arbitrary(u)?,
2266                log: Arbitrary::arbitrary(u)?,
2267            })
2268        }
2269
2270        /// Pick at most `num` elements at random from `data`, returning them.
2271        ///
2272        /// This needs mutable access to perform a shuffle.
2273        ///
2274        fn pick<'a, T>(
2275            u: &mut Unstructured<'a>,
2276            num: usize,
2277            mut data: Vec<T>,
2278        ) -> arbitrary::Result<Vec<T>> {
2279            let len = data.len();
2280            let num = num.min(len);
2281            // Invariant: 0..start is a random subset of data.
2282            for start in 0..num {
2283                data.swap(start, u.int_in_range(start..=len - 1)?);
2284            }
2285            data.truncate(num);
2286            Ok(data)
2287        }
2288
2289        fn arbitrary_round<'a>(
2290            u: &mut Unstructured<'a>,
2291            num_participants: u32,
2292            last_successful_players: Option<&Set<u32>>,
2293        ) -> arbitrary::Result<Round> {
2294            let dealers = if let Some(players) = last_successful_players {
2295                let to_pick = u.int_in_range(players.quorum() as usize..=players.len())?;
2296                pick(u, to_pick, players.into_iter().copied().collect())?
2297            } else {
2298                let to_pick = u.int_in_range(1..=num_participants as usize)?;
2299                pick(u, to_pick, (0..num_participants).collect())?
2300            };
2301            let players = {
2302                let to_pick = u.int_in_range(1..=num_participants as usize)?;
2303                pick(u, to_pick, (0..num_participants).collect())?
2304            };
2305            let pairs = dealers
2306                .iter()
2307                .flat_map(|d| players.iter().map(|p| (*d, *p)))
2308                .collect::<Vec<_>>();
2309            let pick_pair_set = |u: &mut Unstructured<'a>| {
2310                let num = u.int_in_range(0..=pairs.len())?;
2311                if num == 0 {
2312                    return Ok(BTreeSet::new());
2313                }
2314                Ok(pick(u, num, pairs.clone())?.into_iter().collect())
2315            };
2316            let pick_dealer_set = |u: &mut Unstructured<'a>| {
2317                let num = u.int_in_range(0..=dealers.len())?;
2318                if num == 0 {
2319                    return Ok(BTreeSet::new());
2320                }
2321                Ok(pick(u, num, dealers.clone())?.into_iter().collect())
2322            };
2323            let round = Round {
2324                no_acks: pick_pair_set(u)?,
2325                bad_shares: pick_pair_set(u)?,
2326                bad_player_sigs: {
2327                    let indices = pick_pair_set(u)?;
2328                    indices
2329                        .into_iter()
2330                        .map(|k| Ok((k, arbitrary_masks(u)?)))
2331                        .collect::<arbitrary::Result<_>>()?
2332                },
2333                bad_reveals: pick_pair_set(u)?,
2334                bad_dealer_sigs: {
2335                    let indices = pick_dealer_set(u)?;
2336                    indices
2337                        .into_iter()
2338                        .map(|k| Ok((k, arbitrary_masks(u)?)))
2339                        .collect::<arbitrary::Result<_>>()?
2340                },
2341                replace_shares: pick_dealer_set(u)?,
2342                shift_degrees: {
2343                    let indices = pick_dealer_set(u)?;
2344                    indices
2345                        .into_iter()
2346                        .map(|k| {
2347                            let expected = quorum(players.len() as u32) as i32 - 1;
2348                            let shift = u.int_in_range(1..=expected.max(1))?;
2349                            let shift = if bool::arbitrary(u)? { -shift } else { shift };
2350                            Ok((k, NonZeroI32::new(shift).expect("checked to not be zero")))
2351                        })
2352                        .collect::<arbitrary::Result<_>>()?
2353                },
2354                dealers,
2355                players,
2356            };
2357            Ok(round)
2358        }
2359
2360        impl<'a> Arbitrary<'a> for Plan {
2361            fn arbitrary(u: &mut Unstructured<'a>) -> arbitrary::Result<Self> {
2362                let num_participants = u.int_in_range(1..=MAX_NUM_PARTICIPANTS)?;
2363                let mut rounds = Vec::new();
2364                let mut last_successful_players: Option<Set<u32>> = None;
2365                u.arbitrary_loop(None, Some(MAX_ROUNDS), |u| {
2366                    let round =
2367                        arbitrary_round(u, num_participants, last_successful_players.as_ref())?;
2368                    if !round
2369                        .expect_failure(last_successful_players.as_ref().map(|x| x.len() as u32))
2370                    {
2371                        last_successful_players =
2372                            Some(round.players.iter().copied().try_collect().unwrap());
2373                    }
2374                    rounds.push(round);
2375                    Ok(ControlFlow::Continue(()))
2376                })?;
2377                let plan = Self {
2378                    num_participants: NZU32!(num_participants),
2379                    rounds,
2380                };
2381                plan.validate()
2382                    .map_err(|_| arbitrary::Error::IncorrectFormat)?;
2383                Ok(plan)
2384            }
2385        }
2386    }
2387}
2388
2389#[cfg(feature = "arbitrary")]
2390pub use test_plan::Plan as FuzzPlan;
2391
2392#[cfg(test)]
2393mod test {
2394    use super::{test_plan::*, *};
2395    use crate::{bls12381::primitives::variant::MinPk, ed25519};
2396    use anyhow::anyhow;
2397    use core::num::NonZeroI32;
2398    use rand::SeedableRng;
2399    use rand_chacha::ChaCha8Rng;
2400
2401    #[test]
2402    fn single_round() -> anyhow::Result<()> {
2403        Plan::new(NZU32!(4))
2404            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2405            .run::<MinPk>(0)
2406    }
2407
2408    #[test]
2409    fn multiple_rounds() -> anyhow::Result<()> {
2410        Plan::new(NZU32!(4))
2411            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2412            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2413            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2414            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2415            .run::<MinPk>(0)
2416    }
2417
2418    #[test]
2419    fn changing_committee() -> anyhow::Result<()> {
2420        Plan::new(NonZeroU32::new(5).unwrap())
2421            .with(Round::new(vec![0, 1, 2], vec![1, 2, 3]))
2422            .with(Round::new(vec![1, 2, 3], vec![2, 3, 4]))
2423            .with(Round::new(vec![2, 3, 4], vec![3, 4, 0]))
2424            .with(Round::new(vec![3, 4, 0], vec![4, 0, 1]))
2425            .run::<MinPk>(0)
2426    }
2427
2428    #[test]
2429    fn missing_ack() -> anyhow::Result<()> {
2430        // With 4 players, max_faults = 1, so 1 missing ack per dealer is OK
2431        Plan::new(NonZeroU32::new(4).unwrap())
2432            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).no_ack(0, 0))
2433            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).no_ack(0, 1))
2434            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).no_ack(0, 2))
2435            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).no_ack(0, 3))
2436            .run::<MinPk>(0)
2437    }
2438
2439    #[test]
2440    fn increasing_decreasing_committee() -> anyhow::Result<()> {
2441        Plan::new(NonZeroU32::new(5).unwrap())
2442            .with(Round::new(vec![0, 1], vec![0, 1, 2]))
2443            .with(Round::new(vec![0, 1, 2], vec![0, 1, 2, 3]))
2444            .with(Round::new(vec![0, 1, 2], vec![0, 1]))
2445            .with(Round::new(vec![0, 1], vec![0, 1, 2, 3, 4]))
2446            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1]))
2447            .run::<MinPk>(0)
2448    }
2449
2450    #[test]
2451    fn bad_reveal_fails() -> anyhow::Result<()> {
2452        Plan::new(NonZeroU32::new(4).unwrap())
2453            .with(Round::new(vec![0], vec![0, 1, 2, 3]).bad_reveal(0, 1))
2454            .run::<MinPk>(0)
2455    }
2456
2457    #[test]
2458    fn bad_share() -> anyhow::Result<()> {
2459        Plan::new(NonZeroU32::new(4).unwrap())
2460            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).bad_share(0, 1))
2461            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).bad_share(0, 2))
2462            .run::<MinPk>(0)
2463    }
2464
2465    #[test]
2466    fn shift_degree_fails() -> anyhow::Result<()> {
2467        Plan::new(NonZeroU32::new(4).unwrap())
2468            .with(Round::new(vec![0], vec![0, 1, 2, 3]).shift_degree(
2469                0,
2470                NonZeroI32::new(1).ok_or_else(|| anyhow!("invalid NZI32"))?,
2471            ))
2472            .run::<MinPk>(0)
2473    }
2474
2475    #[test]
2476    fn replace_share_fails() -> anyhow::Result<()> {
2477        Plan::new(NonZeroU32::new(4).unwrap())
2478            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]))
2479            .with(Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3]).replace_share(0))
2480            .run::<MinPk>(0)
2481    }
2482
2483    #[test]
2484    fn too_many_reveals_dealer() -> anyhow::Result<()> {
2485        Plan::new(NonZeroU32::new(4).unwrap())
2486            .with(
2487                Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3])
2488                    .no_ack(0, 0)
2489                    .no_ack(0, 1),
2490            )
2491            .run::<MinPk>(0)
2492    }
2493
2494    #[test]
2495    fn too_many_reveals_player() -> anyhow::Result<()> {
2496        Plan::new(NonZeroU32::new(4).unwrap())
2497            .with(
2498                Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3])
2499                    .no_ack(0, 0)
2500                    .no_ack(1, 0)
2501                    .no_ack(3, 0),
2502            )
2503            .run::<MinPk>(0)
2504    }
2505
2506    #[test]
2507    fn bad_sigs() -> anyhow::Result<()> {
2508        Plan::new(NonZeroU32::new(4).unwrap())
2509            .with(
2510                Round::new(vec![0, 1, 2, 3], vec![0, 1, 2, 3])
2511                    .bad_dealer_sig(
2512                        0,
2513                        Masks {
2514                            log: vec![0xFF; 8],
2515                            ..Default::default()
2516                        },
2517                    )
2518                    .bad_player_sig(
2519                        0,
2520                        1,
2521                        Masks {
2522                            pub_msg: vec![0xFF; 8],
2523                            ..Default::default()
2524                        },
2525                    ),
2526            )
2527            .run::<MinPk>(0)
2528    }
2529
2530    #[test]
2531    fn signed_dealer_log_commitment() -> Result<(), Error> {
2532        let sk = ed25519::PrivateKey::from_seed(0);
2533        let pk = sk.public_key();
2534        let info = Info::<MinPk, _>::new(
2535            &[],
2536            0,
2537            None,
2538            Default::default(),
2539            vec![sk.public_key()].try_into().unwrap(),
2540            vec![sk.public_key()].try_into().unwrap(),
2541        )?;
2542        let mut log0 = {
2543            let (dealer, _, _) = Dealer::start(
2544                &mut ChaCha8Rng::seed_from_u64(0),
2545                info.clone(),
2546                sk.clone(),
2547                None,
2548            )?;
2549            dealer.finalize()
2550        };
2551        let mut log1 = {
2552            let (mut dealer, pub_msg, priv_msgs) = Dealer::start(
2553                &mut ChaCha8Rng::seed_from_u64(0),
2554                info.clone(),
2555                sk.clone(),
2556                None,
2557            )?;
2558            let mut player = Player::new(info.clone(), sk)?;
2559            let ack = player
2560                .dealer_message(pk.clone(), pub_msg, priv_msgs[0].1.clone())
2561                .unwrap();
2562            dealer.receive_player_ack(pk, ack)?;
2563            dealer.finalize()
2564        };
2565        std::mem::swap(&mut log0.log, &mut log1.log);
2566        assert!(log0.check(&info).is_none());
2567        assert!(log1.check(&info).is_none());
2568
2569        Ok(())
2570    }
2571
2572    #[cfg(feature = "arbitrary")]
2573    mod conformance {
2574        use super::*;
2575        use commonware_codec::conformance::CodecConformance;
2576
2577        commonware_conformance::conformance_tests! {
2578            CodecConformance<Output<MinPk, ed25519::PublicKey>>,
2579            CodecConformance<DealerPubMsg<MinPk>>,
2580            CodecConformance<DealerPrivMsg>,
2581            CodecConformance<PlayerAck<ed25519::PublicKey>>,
2582            CodecConformance<AckOrReveal<ed25519::PublicKey>>,
2583            CodecConformance<DealerResult<ed25519::PublicKey>>,
2584            CodecConformance<DealerLog<MinPk, ed25519::PublicKey>>,
2585            CodecConformance<SignedDealerLog<MinPk, ed25519::PrivateKey>>,
2586        }
2587    }
2588}