Merge pull request #36 from logos-co/goas/hash-based-balance-commitments

goas: switch balance commitments to hash based commitments
This commit is contained in:
davidrusu 2024-08-23 16:18:37 +04:00 committed by GitHub
commit da18cef987
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
19 changed files with 469 additions and 338 deletions

View File

@ -1,8 +1,11 @@
pub mod mmr;
use cl::{balance::Unit, merkle, NoteCommitment};
use ed25519_dalek::{
ed25519::{signature::SignerMut, SignatureBytes},
Signature, SigningKey, VerifyingKey, PUBLIC_KEY_LENGTH,
};
use mmr::{MMRProof, MMR};
use once_cell::sync::Lazy;
use rand_core::CryptoRngCore;
use serde::{Deserialize, Serialize};
@ -21,7 +24,7 @@ pub struct StateCommitment(pub [u8; 32]);
pub type AccountId = [u8; PUBLIC_KEY_LENGTH];
// PLACEHOLDER: this is probably going to be NMO?
pub static ZONE_CL_FUNDS_UNIT: Lazy<Unit> = Lazy::new(|| cl::note::unit_point("NMO"));
pub static ZONE_CL_FUNDS_UNIT: Lazy<Unit> = Lazy::new(|| cl::note::derive_unit("NMO"));
pub fn new_account(mut rng: impl CryptoRngCore) -> SigningKey {
SigningKey::generate(&mut rng)
@ -39,7 +42,7 @@ impl ZoneMetadata {
let mut hasher = Sha256::new();
hasher.update(self.zone_vk);
hasher.update(self.funds_vk);
hasher.update(self.unit.compress().as_bytes());
hasher.update(self.unit);
hasher.finalize().into()
}
}
@ -47,7 +50,7 @@ impl ZoneMetadata {
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct StateWitness {
pub balances: BTreeMap<AccountId, u64>,
pub included_txs: Vec<Tx>,
pub included_txs: MMR,
pub zone_metadata: ZoneMetadata,
}
@ -58,21 +61,25 @@ impl StateWitness {
pub fn state_roots(&self) -> StateRoots {
StateRoots {
tx_root: self.included_txs_root(),
included_txs: self.included_txs.clone(),
zone_id: self.zone_metadata.id(),
balance_root: self.balances_root(),
}
}
pub fn apply(self, tx: Tx) -> Self {
pub fn apply(self, tx: Tx) -> (Self, IncludedTxWitness) {
let mut state = match tx {
Tx::Withdraw(w) => self.withdraw(w),
Tx::Deposit(d) => self.deposit(d),
};
state.included_txs.push(tx);
let inclusion_proof = state.included_txs.push(&tx.to_bytes());
let tx_inclusion_proof = IncludedTxWitness {
tx,
proof: inclusion_proof,
};
state
(state, tx_inclusion_proof)
}
fn withdraw(mut self, w: Withdraw) -> Self {
@ -97,16 +104,6 @@ impl StateWitness {
self
}
pub fn included_txs_root(&self) -> [u8; 32] {
merkle::root::<MAX_TXS>(self.included_tx_merkle_leaves())
}
pub fn included_tx_witness(&self, idx: usize) -> IncludedTxWitness {
let tx = *self.included_txs.get(idx).unwrap();
let path = merkle::path(self.included_tx_merkle_leaves(), idx);
IncludedTxWitness { tx, path }
}
pub fn balances_root(&self) -> [u8; 32] {
let balance_bytes = Vec::from_iter(self.balances.iter().map(|(owner, balance)| {
let mut bytes: Vec<u8> = vec![];
@ -121,15 +118,6 @@ impl StateWitness {
pub fn total_balance(&self) -> u64 {
self.balances.values().sum()
}
fn included_tx_merkle_leaves(&self) -> [[u8; 32]; MAX_TXS] {
let tx_bytes = self
.included_txs
.iter()
.map(|t| t.to_bytes())
.collect::<Vec<_>>();
merkle::padded_leaves(&tx_bytes)
}
}
impl From<StateCommitment> for [u8; 32] {
@ -237,31 +225,28 @@ impl Tx {
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct IncludedTxWitness {
pub tx: Tx,
pub path: Vec<merkle::PathNode>,
}
impl IncludedTxWitness {
pub fn tx_root(&self) -> [u8; 32] {
let leaf = merkle::leaf(&self.tx.to_bytes());
merkle::path_root(leaf, &self.path)
}
pub proof: MMRProof,
}
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct StateRoots {
pub tx_root: [u8; 32],
pub included_txs: MMR,
pub zone_id: [u8; 32],
pub balance_root: [u8; 32],
}
impl StateRoots {
/// Merkle tree over: [txs, zoneid, balances]
pub fn verify_tx_inclusion(&self, tx_inclusion: &IncludedTxWitness) -> bool {
self.included_txs
.verify_proof(&tx_inclusion.tx.to_bytes(), &tx_inclusion.proof)
}
/// Commitment to the state roots
pub fn commit(&self) -> StateCommitment {
let leaves = cl::merkle::padded_leaves::<4>(&[
self.tx_root.to_vec(),
self.zone_id.to_vec(),
self.balance_root.to_vec(),
]);
StateCommitment(cl::merkle::root(leaves))
let mut hasher = Sha256::new();
hasher.update(self.included_txs.commit());
hasher.update(self.zone_id);
hasher.update(self.balance_root);
StateCommitment(hasher.finalize().into())
}
}

View File

@ -0,0 +1,126 @@
use cl::merkle;
use serde::{Deserialize, Serialize};
use sha2::{Digest, Sha256};
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct MMR {
pub roots: Vec<Root>,
}
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct Root {
pub root: [u8; 32],
pub height: u8,
}
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct MMRProof {
pub path: Vec<merkle::PathNode>,
}
impl MMR {
pub fn new() -> Self {
Self { roots: vec![] }
}
pub fn push(&mut self, elem: &[u8]) -> MMRProof {
let new_root = Root {
root: merkle::leaf(elem),
height: 1,
};
self.roots.push(new_root);
let mut path = vec![];
for i in (1..self.roots.len()).rev() {
if self.roots[i].height == self.roots[i - 1].height {
path.push(merkle::PathNode::Left(self.roots[i - 1].root));
self.roots[i - 1] = Root {
root: merkle::node(self.roots[i - 1].root, self.roots[i].root),
height: self.roots[i - 1].height + 1,
};
self.roots.remove(i);
} else {
break;
}
}
MMRProof { path }
}
pub fn verify_proof(&self, elem: &[u8], proof: &MMRProof) -> bool {
let path_len = proof.path.len();
let leaf = merkle::leaf(elem);
let root = merkle::path_root(leaf, &proof.path);
for mmr_root in self.roots.iter() {
if mmr_root.height == (path_len + 1) as u8 {
return mmr_root.root == root;
}
}
false
}
pub fn commit(&self) -> [u8; 32] {
let mut hasher = Sha256::new();
for mrr_root in self.roots.iter() {
hasher.update(mrr_root.root);
hasher.update(mrr_root.height.to_le_bytes());
}
hasher.finalize().into()
}
}
#[cfg(test)]
mod test {
use super::*;
#[test]
fn test_mrr_push() {
let mut mmr = MMR::new();
let proof = mmr.push(b"hello");
assert_eq!(mmr.roots.len(), 1);
assert_eq!(mmr.roots[0].height, 1);
assert_eq!(mmr.roots[0].root, merkle::leaf(b"hello"));
assert!(mmr.verify_proof(b"hello", &proof));
let proof = mmr.push(b"world");
assert_eq!(mmr.roots.len(), 1);
assert_eq!(mmr.roots[0].height, 2);
assert_eq!(
mmr.roots[0].root,
merkle::node(merkle::leaf(b"hello"), merkle::leaf(b"world"))
);
assert!(mmr.verify_proof(b"world", &proof));
let proof = mmr.push(b"!");
assert_eq!(mmr.roots.len(), 2);
assert_eq!(mmr.roots[0].height, 2);
assert_eq!(
mmr.roots[0].root,
merkle::node(merkle::leaf(b"hello"), merkle::leaf(b"world"))
);
assert_eq!(mmr.roots[1].height, 1);
assert_eq!(mmr.roots[1].root, merkle::leaf(b"!"));
assert!(mmr.verify_proof(b"!", &proof));
let proof = mmr.push(b"!");
assert_eq!(mmr.roots.len(), 1);
assert_eq!(mmr.roots[0].height, 3);
assert_eq!(
mmr.roots[0].root,
merkle::node(
merkle::node(merkle::leaf(b"hello"), merkle::leaf(b"world")),
merkle::node(merkle::leaf(b"!"), merkle::leaf(b"!"))
)
);
assert!(mmr.verify_proof(b"!", &proof));
}
}

View File

@ -1,6 +1,8 @@
use std::collections::BTreeMap;
use common::{AccountId, SignedBoundTx, StateWitness, Tx, ZoneMetadata};
use common::{
mmr::MMR, AccountId, IncludedTxWitness, SignedBoundTx, StateWitness, Tx, ZoneMetadata,
};
use goas_proof_statements::{
user_note::UserAtomicTransfer, zone_funds::SpendFundsPrivate, zone_state::ZoneStatePrivate,
};
@ -21,7 +23,7 @@ impl ZoneNotes {
) -> Self {
let state = StateWitness {
balances,
included_txs: vec![],
included_txs: MMR::new(),
zone_metadata: zone_metadata(zone_name),
};
let state_note = zone_state_utxo(&state, &mut rng);
@ -41,10 +43,9 @@ impl ZoneNotes {
cl::InputWitness::public(self.fund_note)
}
pub fn run(mut self, txs: impl IntoIterator<Item = Tx>) -> Self {
for tx in txs {
self.state = self.state.apply(tx);
}
pub fn run(mut self, tx: Tx) -> (Self, IncludedTxWitness) {
let (new_state, included_tx) = self.state.apply(tx);
self.state = new_state;
let state_in = self.state_input_witness();
self.state_note = cl::OutputWitness::public(
@ -63,7 +64,8 @@ impl ZoneNotes {
},
state_in.evolved_nonce(b"FUND_NONCE"),
);
self
(self, included_tx)
}
}
@ -115,7 +117,7 @@ pub fn zone_metadata(zone_mnemonic: &str) -> ZoneMetadata {
ZoneMetadata {
zone_vk: zone_state_death_constraint(),
funds_vk: zone_fund_death_constraint(),
unit: cl::note::unit_point(zone_mnemonic),
unit: cl::note::derive_unit(zone_mnemonic),
}
}
@ -218,7 +220,7 @@ pub fn prove_user_atomic_transfer(atomic_transfer: UserAtomicTransfer) -> ledger
#[cfg(test)]
mod tests {
use cl::{
note::unit_point, BalanceWitness, NoteWitness, NullifierNonce, OutputWitness,
note::derive_unit, BalanceWitness, NoteWitness, NullifierNonce, OutputWitness,
PartialTxWitness,
};
use common::{BoundTx, Deposit, Withdraw};
@ -231,26 +233,17 @@ mod tests {
pub fn test_prove_zone_stf() {
let mut rng = rand::thread_rng();
let zone_start = ZoneNotes::new_with_balances("ZONE", BTreeMap::from_iter([]), &mut rng);
let mut alice = common::new_account(&mut rng);
let alice_vk = alice.verifying_key().to_bytes();
let zone_start =
ZoneNotes::new_with_balances("ZONE", BTreeMap::from_iter([(alice_vk, 32)]), &mut rng);
let bind = OutputWitness::public(
NoteWitness::basic(32, *common::ZONE_CL_FUNDS_UNIT),
cl::NullifierNonce::random(&mut rng),
);
let mut alice = common::new_account(&mut rng);
let alice_vk = alice.verifying_key().to_bytes();
let signed_deposit = SignedBoundTx::sign(
BoundTx {
tx: Tx::Deposit(Deposit {
to: alice_vk,
amount: 32,
}),
bind: bind.commit_note(),
},
&mut alice,
);
let signed_withdraw = SignedBoundTx::sign(
BoundTx {
tx: Tx::Withdraw(Withdraw {
@ -262,9 +255,7 @@ mod tests {
&mut alice,
);
let zone_end = zone_start
.clone()
.run([signed_deposit.bound_tx.tx, signed_withdraw.bound_tx.tx]);
let zone_end = zone_start.clone().run(signed_withdraw.bound_tx.tx).0;
let ptx = PartialTxWitness {
inputs: vec![
@ -273,13 +264,10 @@ mod tests {
zone_start.fund_input_witness(),
],
outputs: vec![zone_end.state_note, zone_end.fund_note],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let txs = vec![
(signed_deposit, ptx.input_witness(0)),
(signed_withdraw, ptx.input_witness(0)),
];
let txs = vec![(signed_withdraw, ptx.input_witness(0))];
let proof = prove_zone_stf(
zone_start.state.clone(),
@ -303,7 +291,7 @@ mod tests {
let ptx = PartialTxWitness {
inputs: vec![zone.fund_input_witness()],
outputs: vec![zone.state_note],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let proof =
@ -322,9 +310,9 @@ mod tests {
let alice = common::new_account(&mut rng);
let alice_vk = alice.verifying_key().to_bytes();
let mut zone_a =
let zone_a =
ZoneNotes::new_with_balances("ZONE_A", BTreeMap::from_iter([(alice_vk, 40)]), &mut rng);
let mut zone_b = ZoneNotes::new_with_balances("ZONE_B", BTreeMap::new(), &mut rng);
let zone_b = ZoneNotes::new_with_balances("ZONE_B", BTreeMap::new(), &mut rng);
let user_intent = UserIntent {
zone_a_meta: zone_a.state.zone_metadata,
@ -339,17 +327,17 @@ mod tests {
},
};
let user_note = cl::InputWitness::public(cl::OutputWitness::public(
NoteWitness::new(1, unit_point("INTENT"), [0u8; 32], user_intent.commit()),
NoteWitness::new(1, derive_unit("INTENT"), [0u8; 32], user_intent.commit()),
NullifierNonce::random(&mut rng),
));
zone_a = zone_a.run([Tx::Withdraw(user_intent.withdraw)]);
zone_b = zone_b.run([Tx::Deposit(user_intent.deposit)]);
let (zone_a, withdraw_included_witnesss) = zone_a.run(Tx::Withdraw(user_intent.withdraw));
let (zone_b, deposit_included_witnesss) = zone_b.run(Tx::Deposit(user_intent.deposit));
let ptx = PartialTxWitness {
inputs: vec![user_note],
outputs: vec![zone_a.state_note, zone_b.state_note],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let user_atomic_transfer = UserAtomicTransfer {
@ -359,8 +347,8 @@ mod tests {
zone_b: ptx.output_witness(1),
zone_a_roots: zone_a.state.state_roots(),
zone_b_roots: zone_b.state.state_roots(),
withdraw_tx: zone_a.state.included_tx_witness(0),
deposit_tx: zone_b.state.included_tx_witness(0),
withdraw_tx: withdraw_included_witnesss,
deposit_tx: deposit_included_witnesss,
};
let proof = prove_user_atomic_transfer(user_atomic_transfer);

View File

@ -33,7 +33,7 @@ fn test_atomic_transfer() {
let alice_intent_out = cl::OutputWitness::public(
NoteWitness {
value: 1,
unit: cl::note::unit_point("INTENT"),
unit: cl::note::derive_unit("INTENT"),
death_constraint: executor::user_atomic_transfer_death_constraint(),
state: alice_intent.commit(),
},
@ -43,16 +43,15 @@ fn test_atomic_transfer() {
let user_ptx = cl::PartialTxWitness {
inputs: vec![],
outputs: vec![alice_intent_out],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let zone_a_end = zone_a_start
let (zone_a_end, withdraw_inclusion) = zone_a_start
.clone()
.run([Tx::Withdraw(alice_intent.withdraw)]);
.run(Tx::Withdraw(alice_intent.withdraw));
let zone_b_end = zone_b_start
.clone()
.run([Tx::Deposit(alice_intent.deposit)]);
let (zone_b_end, deposit_inclusion) =
zone_b_start.clone().run(Tx::Deposit(alice_intent.deposit));
let alice_intent_in = cl::InputWitness::public(alice_intent_out);
let atomic_transfer_ptx = cl::PartialTxWitness {
@ -69,7 +68,7 @@ fn test_atomic_transfer() {
zone_b_end.state_note,
zone_b_end.fund_note,
],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let signed_withdraw = SignedBoundTx::sign(
@ -97,8 +96,8 @@ fn test_atomic_transfer() {
zone_b: atomic_transfer_ptx.output_witness(2),
zone_a_roots: zone_a_end.state.state_roots(),
zone_b_roots: zone_b_end.state.state_roots(),
withdraw_tx: zone_a_end.state.included_tx_witness(0),
deposit_tx: zone_b_end.state.included_tx_witness(0),
withdraw_tx: withdraw_inclusion,
deposit_tx: deposit_inclusion,
}),
),
(
@ -161,18 +160,12 @@ fn test_atomic_transfer() {
assert!(atomic_transfer_proof.verify());
let bundle = cl::Bundle {
partials: vec![user_ptx.commit(), atomic_transfer_ptx.commit()],
};
let bundle_witness = BundleWitness {
balance_blinding: cl::BalanceWitness(
user_ptx.balance_blinding.0 + atomic_transfer_ptx.balance_blinding.0,
),
partials: vec![user_ptx, atomic_transfer_ptx],
};
let bundle_proof =
ledger::bundle::ProvedBundle::prove(&bundle, &bundle_witness).expect("bundle proof failed");
ledger::bundle::ProvedBundle::prove(&bundle_witness).expect("bundle proof failed");
assert!(bundle_proof.verify());
}

View File

@ -1,7 +1,7 @@
use std::collections::BTreeMap;
use cl::{BalanceWitness, NoteWitness, NullifierSecret};
use common::{new_account, BoundTx, SignedBoundTx, StateWitness, Tx, ZONE_CL_FUNDS_UNIT};
use common::{mmr::MMR, new_account, BoundTx, SignedBoundTx, StateWitness, Tx, ZONE_CL_FUNDS_UNIT};
use executor::ZoneNotes;
use ledger::death_constraint::DeathProof;
@ -20,7 +20,7 @@ fn test_deposit() {
amount: 78,
};
let zone_end = zone_start.clone().run([Tx::Deposit(deposit)]);
let zone_end = zone_start.clone().run(Tx::Deposit(deposit)).0;
let alice_deposit = cl::InputWitness::from_output(
cl::OutputWitness::random(
@ -38,7 +38,7 @@ fn test_deposit() {
let deposit_ptx = cl::PartialTxWitness {
inputs: vec![zone_start.state_input_witness(), alice_deposit],
outputs: vec![zone_end.state_note, zone_end.fund_note],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let signed_deposit = SignedBoundTx::sign(
@ -82,14 +82,15 @@ fn test_deposit() {
zone_end.state_note.note.state,
StateWitness {
balances: BTreeMap::from_iter([(alice_vk, 78)]),
included_txs: vec![Tx::Deposit(deposit)],
included_txs: {
let mut mmr = MMR::new();
mmr.push(&Tx::Deposit(deposit).to_bytes());
mmr
},
zone_metadata: zone_start.state.zone_metadata,
}
.commit()
.0
);
assert_eq!(
deposit_ptx.commit().balance,
cl::Balance::zero(deposit_ptx.balance_blinding)
);
assert!(deposit_ptx.balance().is_zero());
}

View File

@ -1,7 +1,7 @@
use std::collections::BTreeMap;
use cl::{BalanceWitness, NoteWitness, NullifierSecret};
use common::{new_account, BoundTx, SignedBoundTx, StateWitness, Tx, ZONE_CL_FUNDS_UNIT};
use common::{mmr::MMR, new_account, BoundTx, SignedBoundTx, StateWitness, Tx, ZONE_CL_FUNDS_UNIT};
use executor::ZoneNotes;
use ledger::death_constraint::DeathProof;
@ -30,7 +30,7 @@ fn test_withdrawal() {
amount: 78,
};
let zone_end = zone_start.clone().run([Tx::Withdraw(withdraw)]);
let zone_end = zone_start.clone().run(Tx::Withdraw(withdraw)).0;
let alice_withdrawal = cl::OutputWitness::random(
NoteWitness::stateless(
@ -49,7 +49,7 @@ fn test_withdrawal() {
alice_intent,
],
outputs: vec![zone_end.state_note, zone_end.fund_note, alice_withdrawal],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let signed_withdraw = SignedBoundTx::sign(
@ -102,7 +102,11 @@ fn test_withdrawal() {
zone_end.state_note.note.state,
StateWitness {
balances: BTreeMap::from_iter([(alice_vk, 22)]),
included_txs: vec![Tx::Withdraw(withdraw)],
included_txs: {
let mut mmr = MMR::new();
mmr.push(&Tx::Withdraw(withdraw).to_bytes());
mmr
},
zone_metadata: zone_start.state.zone_metadata,
}
.commit()

View File

@ -7,9 +7,9 @@
///
/// The Alice will create a partial tx that looks like this:
///
/// [fee note] -> [user note]
/// [] -> [user note]
///
/// The User Note will encode the logic that orchestrates the withdrawal from zone A
/// Thep User Note will encode the logic that orchestrates the withdrawal from zone A
/// and deposit to zone B.
///
/// The User Notes death constraint requires the following statements to be satisfied
@ -84,8 +84,8 @@ impl UserAtomicTransfer {
);
// ensure txs were included in the respective zones
assert_eq!(self.withdraw_tx.tx_root(), self.zone_a_roots.tx_root);
assert_eq!(self.deposit_tx.tx_root(), self.zone_b_roots.tx_root);
assert!(self.zone_a_roots.verify_tx_inclusion(&self.withdraw_tx));
assert!(self.zone_b_roots.verify_tx_inclusion(&self.deposit_tx));
// ensure the txs are the same ones the user requested
assert_eq!(

View File

@ -81,7 +81,7 @@ fn main() {
assert_eq!(ptx_input_witness.input_root(), input_root);
// apply the ptx
state = state.apply(bound_tx.tx)
state = state.apply(bound_tx.tx).0
}
validate_zone_transition(zone_in, zone_out, funds_out, in_state_cm, state);

View File

@ -1,164 +1,149 @@
use curve25519_dalek::{
constants::RISTRETTO_BASEPOINT_POINT, ristretto::RistrettoPoint, traits::VartimeMultiscalarMul,
Scalar,
};
use rand_core::CryptoRngCore;
use serde::{Deserialize, Serialize};
use sha2::{Digest, Sha256};
use crate::NoteWitness;
#[derive(Debug, PartialEq, Eq, Clone, Copy, Serialize, Deserialize)]
pub struct Balance(pub RistrettoPoint);
use crate::PartialTxWitness;
pub type Value = u64;
pub type Unit = RistrettoPoint;
pub type Unit = [u8; 32];
#[derive(Debug, PartialEq, Eq, Clone, Copy, Serialize, Deserialize)]
pub struct BalanceWitness(pub Scalar);
pub struct Balance([u8; 32]);
impl Balance {
/// A commitment to zero, blinded by the provided balance witness
pub fn zero(blinding: BalanceWitness) -> Self {
// Since, balance commitments are `value * UnitPoint + blinding * H`, when value=0, the commmitment is unitless.
// So we use the generator point as a stand in for the unit point.
//
// TAI: we can optimize this further from `0*G + r*H` to just `r*H` to save a point scalar mult + point addition.
let unit = curve25519_dalek::constants::RISTRETTO_BASEPOINT_POINT;
Self(balance(0, unit, blinding.0))
}
pub fn to_bytes(&self) -> [u8; 32] {
self.0.compress().to_bytes()
self.0
}
}
#[derive(Debug, PartialEq, Eq, Clone, Serialize, Deserialize)]
pub struct UnitBalance {
pub unit: Unit,
pub pos: u64,
pub neg: u64,
}
impl UnitBalance {
pub fn is_zero(&self) -> bool {
return self.pos == self.neg;
}
}
#[derive(Debug, PartialEq, Eq, Clone, Serialize, Deserialize)]
pub struct BalanceWitness {
pub balances: Vec<UnitBalance>,
pub blinding: [u8; 16],
}
impl BalanceWitness {
pub fn new(blinding: Scalar) -> Self {
Self(blinding)
pub fn random_blinding(mut rng: impl CryptoRngCore) -> [u8; 16] {
let mut blinding = [0u8; 16];
rng.fill_bytes(&mut blinding);
blinding
}
pub fn unblinded() -> Self {
Self::new(Scalar::ZERO)
pub fn zero(blinding: [u8; 16]) -> Self {
Self {
balances: Default::default(),
blinding,
}
}
pub fn random(mut rng: impl CryptoRngCore) -> Self {
Self::new(Scalar::random(&mut rng))
pub fn from_ptx(ptx: &PartialTxWitness, blinding: [u8; 16]) -> Self {
let mut balance = Self::zero(blinding);
for input in ptx.inputs.iter() {
balance.insert_negative(input.note.unit, input.note.value);
}
pub fn commit<'a>(
&self,
inputs: impl IntoIterator<Item = &'a NoteWitness>,
outputs: impl IntoIterator<Item = &'a NoteWitness>,
) -> Balance {
let (input_points, input_scalars): (Vec<_>, Vec<_>) = inputs
.into_iter()
.map(|i| (i.unit, -Scalar::from(i.value)))
.unzip();
for output in ptx.outputs.iter() {
balance.insert_positive(output.note.unit, output.note.value);
}
let (output_points, output_scalars): (Vec<_>, Vec<_>) = outputs
.into_iter()
.map(|o| (o.unit, Scalar::from(o.value)))
.unzip();
balance.clear_zeros();
let points = input_points
.into_iter()
.chain(output_points)
.chain([RISTRETTO_BASEPOINT_POINT]);
let scalars = input_scalars
.into_iter()
.chain(output_scalars)
.chain([self.0]);
balance
}
let blinded_balance = RistrettoPoint::vartime_multiscalar_mul(scalars, points);
pub fn insert_positive(&mut self, unit: Unit, value: Value) {
for unit_bal in self.balances.iter_mut() {
if unit_bal.unit == unit {
unit_bal.pos += value;
return;
}
}
Balance(blinded_balance)
}
}
pub fn balance(value: u64, unit: Unit, blinding: Scalar) -> Unit {
let value_scalar = Scalar::from(value);
// can vartime leak the number of cycles through the stark proof?
RistrettoPoint::vartime_double_scalar_mul_basepoint(&value_scalar, &unit, &blinding)
}
#[cfg(test)]
mod test {
use super::*;
use crate::note::unit_point;
#[test]
fn test_balance_zero_unitless() {
// Zero is the same across all units
let (nmo, eth) = (unit_point("NMO"), unit_point("ETH"));
let mut rng = rand::thread_rng();
let b = BalanceWitness::random(&mut rng);
assert_eq!(
b.commit([&NoteWitness::basic(0, nmo)], []),
b.commit([&NoteWitness::basic(0, eth)], []),
);
}
#[test]
fn test_balance_blinding() {
// balances are blinded
let nmo = unit_point("NMO");
let r_a = Scalar::from(12u32);
let r_b = Scalar::from(8u32);
let bal_a = BalanceWitness::new(r_a);
let bal_b = BalanceWitness::new(r_b);
let note = NoteWitness::basic(10, nmo);
let a = bal_a.commit([&note], []);
let b = bal_b.commit([&note], []);
assert_ne!(a, b);
let diff_note = NoteWitness::basic(0, nmo);
assert_eq!(
a.0 - b.0,
BalanceWitness::new(r_a - r_b).commit([&diff_note], []).0
);
}
#[test]
fn test_balance_units() {
// Unit's differentiate between values.
let (nmo, eth) = (unit_point("NMO"), unit_point("ETH"));
let b = BalanceWitness::new(Scalar::from(1337u32));
let nmo = NoteWitness::basic(10, nmo);
let eth = NoteWitness::basic(10, eth);
assert_ne!(b.commit([&nmo], []), b.commit([&eth], []));
}
#[test]
fn test_balance_homomorphism() {
let nmo = unit_point("NMO");
let mut rng = rand::thread_rng();
let b1 = BalanceWitness::random(&mut rng);
let b2 = BalanceWitness::random(&mut rng);
let b_zero = BalanceWitness::new(Scalar::ZERO);
let ten = NoteWitness::basic(10, nmo);
let eight = NoteWitness::basic(8, nmo);
let two = NoteWitness::basic(2, nmo);
let zero = NoteWitness::basic(0, nmo);
// Values of same unit are homomorphic
assert_eq!(
(b1.commit([&ten], []).0 - b1.commit([&eight], []).0),
b_zero.commit([&two], []).0
);
// Blinding factors are also homomorphic.
assert_eq!(
b1.commit([&ten], []).0 - b2.commit([&ten], []).0,
BalanceWitness::new(b1.0 - b2.0).commit([&zero], []).0
);
// Unit was not found, so we must create one.
self.balances.push(UnitBalance {
unit,
pos: value,
neg: 0,
});
}
pub fn insert_negative(&mut self, unit: Unit, value: Value) {
for unit_bal in self.balances.iter_mut() {
if unit_bal.unit == unit {
unit_bal.neg += value;
return;
}
}
self.balances.push(UnitBalance {
unit,
pos: 0,
neg: value,
});
}
pub fn clear_zeros(&mut self) {
let mut i = 0usize;
while i < self.balances.len() {
if self.balances[i].is_zero() {
self.balances.swap_remove(i);
// don't increment `i` since the last element has been swapped into the
// `i`'th place
} else {
i += 1;
}
}
}
pub fn combine(balances: impl IntoIterator<Item = Self>, blinding: [u8; 16]) -> Self {
let mut combined = BalanceWitness::zero(blinding);
for balance in balances {
for unit_bal in balance.balances.iter() {
if unit_bal.pos > unit_bal.neg {
combined.insert_positive(unit_bal.unit, unit_bal.pos - unit_bal.neg);
} else {
combined.insert_negative(unit_bal.unit, unit_bal.neg - unit_bal.pos);
}
}
}
combined.clear_zeros();
combined
}
pub fn is_zero(&self) -> bool {
self.balances.is_empty()
}
pub fn commit(&self) -> Balance {
let mut hasher = Sha256::new();
hasher.update(b"NOMOS_CL_BAL_COMMIT");
for unit_balance in self.balances.iter() {
hasher.update(unit_balance.unit);
hasher.update(unit_balance.pos.to_le_bytes());
hasher.update(unit_balance.neg.to_le_bytes());
}
hasher.update(self.blinding);
let commit_bytes: [u8; 32] = hasher.finalize().into();
Balance(commit_bytes)
}
}

View File

@ -1,6 +1,6 @@
use serde::{Deserialize, Serialize};
use crate::{partial_tx::PartialTx, Balance, BalanceWitness};
use crate::{partial_tx::PartialTx, BalanceWitness, PartialTxWitness};
/// The transaction bundle is a collection of partial transactions.
/// The goal in bundling transactions is to produce a set of partial transactions
@ -11,28 +11,29 @@ pub struct Bundle {
pub partials: Vec<PartialTx>,
}
#[derive(Debug, Clone, Copy, PartialEq, Eq, Serialize, Deserialize)]
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct BundleWitness {
pub balance_blinding: BalanceWitness,
pub partials: Vec<PartialTxWitness>,
}
impl Bundle {
pub fn balance(&self) -> Balance {
Balance(self.partials.iter().map(|ptx| ptx.balance.0).sum())
impl BundleWitness {
pub fn balance(&self) -> BalanceWitness {
BalanceWitness::combine(self.partials.iter().map(|ptx| ptx.balance()), [0u8; 16])
}
pub fn is_balanced(&self, witness: BalanceWitness) -> bool {
self.balance() == Balance::zero(witness)
pub fn commit(&self) -> Bundle {
Bundle {
partials: Vec::from_iter(self.partials.iter().map(|ptx| ptx.commit())),
}
}
}
#[cfg(test)]
mod test {
use curve25519_dalek::{constants::RISTRETTO_BASEPOINT_POINT, Scalar};
use crate::{
balance::UnitBalance,
input::InputWitness,
note::{unit_point, NoteWitness},
note::{derive_unit, NoteWitness},
nullifier::NullifierSecret,
output::OutputWitness,
partial_tx::PartialTxWitness,
@ -43,7 +44,7 @@ mod test {
#[test]
fn test_bundle_balance() {
let mut rng = rand::thread_rng();
let (nmo, eth, crv) = (unit_point("NMO"), unit_point("ETH"), unit_point("CRV"));
let (nmo, eth, crv) = (derive_unit("NMO"), derive_unit("ETH"), derive_unit("CRV"));
let nf_a = NullifierSecret::random(&mut rng);
let nf_b = NullifierSecret::random(&mut rng);
@ -63,25 +64,33 @@ mod test {
let ptx_unbalanced = PartialTxWitness {
inputs: vec![nmo_10_in, eth_23_in],
outputs: vec![crv_4840_out],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let bundle_witness = BundleWitness {
balance_blinding: ptx_unbalanced.balance_blinding,
partials: vec![ptx_unbalanced.clone()],
};
let mut bundle = Bundle {
partials: vec![ptx_unbalanced.commit()],
};
let crv_4840_out_bal = crv_4840_out.note.unit * Scalar::from(crv_4840_out.note.value);
let nmo_10_in_bal = nmo_10_in.note.unit * Scalar::from(nmo_10_in.note.value);
let eth_23_in_bal = eth_23_in.note.unit * Scalar::from(eth_23_in.note.value);
let unbalance_blinding = RISTRETTO_BASEPOINT_POINT * ptx_unbalanced.balance_blinding.0;
assert!(!bundle.is_balanced(bundle_witness.balance_blinding));
assert!(!bundle_witness.balance().is_zero());
assert_eq!(
bundle.balance().0,
crv_4840_out_bal - (nmo_10_in_bal + eth_23_in_bal) + unbalance_blinding
bundle_witness.balance().balances,
vec![
UnitBalance {
unit: nmo,
pos: 0,
neg: 10
},
UnitBalance {
unit: eth,
pos: 0,
neg: 23
},
UnitBalance {
unit: crv,
pos: 4840,
neg: 0
},
]
);
let crv_4840_in = InputWitness::from_output(crv_4840_out, nf_c);
@ -99,17 +108,14 @@ mod test {
let ptx_solved = PartialTxWitness {
inputs: vec![crv_4840_in],
outputs: vec![nmo_10_out, eth_23_out],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
bundle.partials.push(ptx_solved.commit());
let witness = BundleWitness {
balance_blinding: BalanceWitness::new(
ptx_unbalanced.balance_blinding.0 + ptx_solved.balance_blinding.0,
),
partials: vec![ptx_unbalanced, ptx_solved],
};
assert!(bundle.is_balanced(witness.balance_blinding));
assert!(witness.balance().is_zero());
assert_eq!(witness.balance().balances, vec![]);
}
}

View File

@ -18,8 +18,12 @@ pub fn death_commitment(death_constraint: &[u8]) -> DeathCommitment {
DeathCommitment(death_cm)
}
pub fn unit_point(unit: &str) -> Unit {
crate::crypto::hash_to_curve(format!("NOMOS_CL_UNIT{unit}").as_bytes())
pub fn derive_unit(unit: &str) -> Unit {
let mut hasher = Sha256::new();
hasher.update(b"NOMOS_CL_UNIT");
hasher.update(unit.as_bytes());
let unit: Unit = hasher.finalize().into();
unit
}
#[derive(Debug, Clone, Copy, PartialEq, Eq, PartialOrd, Ord, Hash, Serialize, Deserialize)]
@ -63,7 +67,7 @@ impl NoteWitness {
// COMMIT TO BALANCE
hasher.update(self.value.to_le_bytes());
hasher.update(self.unit.compress().to_bytes());
hasher.update(self.unit);
// Important! we don't commit to the balance blinding factor as that may make the notes linkable.
// COMMIT TO STATE
@ -92,7 +96,7 @@ mod test {
#[test]
fn test_note_commit_permutations() {
let (nmo, eth) = (unit_point("NMO"), unit_point("ETH"));
let (nmo, eth) = (derive_unit("NMO"), derive_unit("ETH"));
let mut rng = rand::thread_rng();

View File

@ -122,7 +122,7 @@ impl Nullifier {
#[cfg(test)]
mod test {
use crate::{note::unit_point, NoteWitness};
use crate::{note::derive_unit, NoteWitness};
use super::*;
@ -147,7 +147,7 @@ mod test {
fn test_nullifier_same_sk_different_nonce() {
let mut rng = rand::thread_rng();
let sk = NullifierSecret::random(&mut rng);
let note = NoteWitness::basic(1, unit_point("NMO"));
let note = NoteWitness::basic(1, derive_unit("NMO"));
let nonce_1 = NullifierNonce::random(&mut rng);
let nonce_2 = NullifierNonce::random(&mut rng);
@ -164,8 +164,8 @@ mod test {
fn test_same_sk_same_nonce_different_note() {
let mut rng = rand::thread_rng();
let sk = NullifierSecret::random(&mut rng);
let note_1 = NoteWitness::basic(1, unit_point("NMO"));
let note_2 = NoteWitness::basic(1, unit_point("ETH"));
let note_1 = NoteWitness::basic(1, derive_unit("NMO"));
let note_2 = NoteWitness::basic(1, derive_unit("ETH"));
let nonce = NullifierNonce::random(&mut rng);
let note_cm_1 = note_1.commit(sk.commit(), nonce);
let note_cm_2 = note_2.commit(sk.commit(), nonce);

View File

@ -43,7 +43,7 @@ pub struct PartialTx {
pub struct PartialTxWitness {
pub inputs: Vec<InputWitness>,
pub outputs: Vec<OutputWitness>,
pub balance_blinding: BalanceWitness,
pub balance_blinding: [u8; 16],
}
impl PartialTxWitness {
@ -55,18 +55,19 @@ impl PartialTxWitness {
Self {
inputs,
outputs,
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
}
}
pub fn balance(&self) -> BalanceWitness {
BalanceWitness::from_ptx(self, self.balance_blinding)
}
pub fn commit(&self) -> PartialTx {
PartialTx {
inputs: Vec::from_iter(self.inputs.iter().map(InputWitness::commit)),
outputs: Vec::from_iter(self.outputs.iter().map(OutputWitness::commit)),
balance: self.balance_blinding.commit(
self.inputs.iter().map(|i| &i.note),
self.outputs.iter().map(|o| &o.note),
),
balance: self.balance().commit(),
}
}
@ -149,10 +150,9 @@ impl PartialTxOutputWitness {
#[cfg(test)]
mod test {
use curve25519_dalek::{constants::RISTRETTO_BASEPOINT_POINT, Scalar};
use crate::{
note::{unit_point, NoteWitness},
balance::UnitBalance,
note::{derive_unit, NoteWitness},
nullifier::NullifierSecret,
};
@ -160,7 +160,7 @@ mod test {
#[test]
fn test_partial_tx_balance() {
let (nmo, eth, crv) = (unit_point("NMO"), unit_point("ETH"), unit_point("CRV"));
let (nmo, eth, crv) = (derive_unit("NMO"), derive_unit("ETH"), derive_unit("CRV"));
let mut rng = rand::thread_rng();
let nf_a = NullifierSecret::random(&mut rng);
@ -181,18 +181,34 @@ mod test {
let ptx_witness = PartialTxWitness {
inputs: vec![nmo_10, eth_23],
outputs: vec![crv_4840],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let ptx = ptx_witness.commit();
let crv_4840_bal = crv_4840.note.unit * Scalar::from(crv_4840.note.value);
let nmo_10_bal = nmo_10.note.unit * Scalar::from(nmo_10.note.value);
let eth_23_bal = eth_23.note.unit * Scalar::from(eth_23.note.value);
let blinding = RISTRETTO_BASEPOINT_POINT * ptx_witness.balance_blinding.0;
assert_eq!(
ptx.balance.0,
crv_4840_bal - (nmo_10_bal + eth_23_bal) + blinding,
ptx.balance,
BalanceWitness {
balances: vec![
UnitBalance {
unit: nmo,
pos: 0,
neg: 10
},
UnitBalance {
unit: eth,
pos: 0,
neg: 23
},
UnitBalance {
unit: crv,
pos: 4840,
neg: 0
},
],
blinding: ptx_witness.balance_blinding
}
.commit()
);
}
}

View File

@ -1,4 +1,4 @@
use cl::{note::unit_point, BalanceWitness};
use cl::{note::derive_unit, BalanceWitness};
use rand_core::CryptoRngCore;
fn receive_utxo(
@ -11,7 +11,7 @@ fn receive_utxo(
#[test]
fn test_simple_transfer() {
let nmo = unit_point("NMO");
let nmo = derive_unit("NMO");
let mut rng = rand::thread_rng();
let sender_nf_sk = cl::NullifierSecret::random(&mut rng);
@ -31,14 +31,12 @@ fn test_simple_transfer() {
let ptx_witness = cl::PartialTxWitness {
inputs: vec![cl::InputWitness::from_output(utxo, sender_nf_sk)],
outputs: vec![recipient_output, change_output],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
let ptx = ptx_witness.commit();
let bundle = cl::Bundle {
partials: vec![ptx],
let bundle = cl::BundleWitness {
partials: vec![ptx_witness],
};
assert!(bundle.is_balanced(ptx_witness.balance_blinding))
assert!(bundle.balance().is_zero())
}

View File

@ -1,3 +1,5 @@
use ledger_proof_statements::bundle::BundlePrivate;
use crate::error::{Error, Result};
pub struct ProvedBundle {
@ -6,12 +8,20 @@ pub struct ProvedBundle {
}
impl ProvedBundle {
pub fn prove(bundle: &cl::Bundle, bundle_witness: &cl::BundleWitness) -> Result<Self> {
pub fn prove(bundle_witness: &cl::BundleWitness) -> Result<Self> {
// need to show that bundle is balanced.
// i.e. the sum of ptx balances is 0
let bundle_private = BundlePrivate {
balances: bundle_witness
.partials
.iter()
.map(|ptx| ptx.balance())
.collect(),
};
let env = risc0_zkvm::ExecutorEnv::builder()
.write(&bundle_witness)
.write(&bundle_private)
.unwrap()
.build()
.unwrap();
@ -34,21 +44,21 @@ impl ProvedBundle {
let receipt = prove_info.receipt;
Ok(Self {
bundle: bundle.clone(),
bundle: bundle_witness.commit(),
risc0_receipt: receipt,
})
}
pub fn public(&self) -> Result<cl::Balance> {
pub fn public(&self) -> Result<ledger_proof_statements::bundle::BundlePublic> {
Ok(self.risc0_receipt.journal.decode()?)
}
pub fn verify(&self) -> bool {
let Ok(zero_commitment) = self.public() else {
let Ok(bundle_public) = self.public() else {
return false;
};
self.bundle.balance() == zero_commitment
Vec::from_iter(self.bundle.partials.iter().map(|ptx| ptx.balance)) == bundle_public.balances
&& self
.risc0_receipt
.verify(nomos_cl_risc0_proofs::BUNDLE_ID)

View File

@ -1,6 +1,6 @@
use std::collections::BTreeMap;
use cl::{note::unit_point, BalanceWitness};
use cl::{note::derive_unit, BalanceWitness};
use ledger::{bundle::ProvedBundle, death_constraint::DeathProof, partial_tx::ProvedPartialTx};
use rand_core::CryptoRngCore;
@ -30,7 +30,7 @@ fn receive_utxo(
#[test]
fn test_simple_transfer() {
let nmo = unit_point("NMO");
let nmo = derive_unit("NMO");
let mut rng = rand::thread_rng();
@ -58,7 +58,7 @@ fn test_simple_transfer() {
let ptx_witness = cl::PartialTxWitness {
inputs: vec![alices_input],
outputs: vec![bobs_output, change_output],
balance_blinding: BalanceWitness::random(&mut rng),
balance_blinding: BalanceWitness::random_blinding(&mut rng),
};
// Prove the death constraints for alices input (she uses the no-op death constraint)
@ -75,14 +75,10 @@ fn test_simple_transfer() {
assert!(proved_ptx.verify()); // It's a valid ptx.
let bundle = cl::Bundle {
partials: vec![ptx_witness.commit()],
};
let bundle_witness = cl::BundleWitness {
balance_blinding: ptx_witness.balance_blinding,
partials: vec![ptx_witness],
};
let proved_bundle = ProvedBundle::prove(&bundle, &bundle_witness).unwrap();
let proved_bundle = ProvedBundle::prove(&bundle_witness).unwrap();
assert!(proved_bundle.verify()); // The bundle is balanced.
}

View File

@ -0,0 +1,12 @@
use cl::{Balance, BalanceWitness};
use serde::{Deserialize, Serialize};
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct BundlePublic {
pub balances: Vec<Balance>,
}
#[derive(Debug, Clone, PartialEq, Eq, Serialize, Deserialize)]
pub struct BundlePrivate {
pub balances: Vec<BalanceWitness>,
}

View File

@ -1,2 +1,3 @@
pub mod death_constraint;
pub mod ptx;
pub mod bundle;

View File

@ -11,7 +11,13 @@
use risc0_zkvm::guest::env;
fn main() {
let bundle_witness: cl::BundleWitness = env::read();
let zero_balance = cl::Balance::zero(bundle_witness.balance_blinding);
env::commit(&zero_balance);
let bundle_private: ledger_proof_statements::bundle::BundlePrivate = env::read();
let bundle_public = ledger_proof_statements::bundle::BundlePublic {
balances: Vec::from_iter(bundle_private.balances.iter().map(|b| b.commit())),
};
assert!(cl::BalanceWitness::combine(bundle_private.balances, [0u8; 16]).is_zero());
env::commit(&bundle_public);
}