netstack3_tcp/
sack_scoreboard.rs

1// Copyright 2025 The Fuchsia Authors. All rights reserved.
2// Use of this source code is governed by a BSD-style license that can be
3// found in the LICENSE file.
4
5//! Implements the selective acknowledgement scoreboard data structure as
6//! described in [RFC 6675].
7//!
8//! [RFC 6675]: https://datatracker.ietf.org/doc/html/rfc6675
9
10use netstack3_base::{Mss, SackBlocks, SeqNum};
11
12use crate::internal::congestion::DUP_ACK_THRESHOLD;
13use crate::internal::seq_ranges::{FirstHoleResult, SeqRange, SeqRanges};
14
15#[derive(Debug, Default)]
16#[cfg_attr(test, derive(PartialEq, Eq))]
17pub(crate) struct SackScoreboard {
18    /// The ranges for which we've received selective acknowledgements.
19    acked_ranges: SeqRanges<()>,
20    /// Stores the number of bytes assumed to be in transit according to the definition of Pipe
21    /// defined in [RFC 6675 section 4].
22    ///
23    /// [RFC 6675 section 4]: https://datatracker.ietf.org/doc/html/rfc6675#section-4
24    pipe: u32,
25    /// A sequence number before which all holes should be considered lost as
26    /// per the IsLost definition in [RFC 66752 section 4].
27    ///
28    /// `None` if none of the tracked holes are lost, or if there are no holes
29    /// in `acked_ranges`.
30    ///
31    /// [RFC 6675 section 4]:
32    ///     https://datatracker.ietf.org/doc/html/rfc6675#section-4
33    is_lost_seqnum_end: Option<SeqNum>,
34}
35
36impl SackScoreboard {
37    /// Processes an incoming ACK and updates the scoreboard.
38    ///
39    /// - `ack` is the cumulative acknowledgement in the received segment.
40    /// - `snd_nxt` is the value of SND.NXT or the highest sequence number that
41    ///   has been sent out.
42    /// - `high_rxt` is the equivalent to the HighRxt value defined in [RFC 6675
43    ///   section 2], but we define it as the next available sequence number for
44    ///   retransmission (off by one from the RFC definition). It is expected to
45    ///   only be available if loss recovery is initiated.
46    /// - `sack_blocks` are the selective ack blocks informed by the peer in the
47    ///   segment.
48    /// - `smss` is the send size mss.
49    ///
50    /// Any blocks referencing data before `ack` or after `snd_nxt` are
51    /// *ignored* as bad data. We chose to ignore any blocks after `snd_nxt`
52    /// here so the SACK recovery algorithm works as described in [RFC 6675].
53    /// Note that [RFC 6675 section 5.1] calls out that the algorithm described
54    /// in the RFC is not suited to deal with a retransmit timeout, so to avoid
55    /// an improper pipe calculation we ignore any blocks past SND.NXT.
56    ///
57    /// Returns `true` if this segment should be considered a duplicate as per
58    /// the definition in [RFC 6675 section 2].
59    ///
60    /// [RFC 6675]: https://datatracker.ietf.org/doc/html/rfc6675
61    /// [RFC 6675 section 5.1]:
62    ///     https://datatracker.ietf.org/doc/html/rfc6675#section-5.1
63    /// [RFC 6675 Section 2]:
64    ///     https://datatracker.ietf.org/doc/html/rfc6675#section-2
65    pub(crate) fn process_ack(
66        &mut self,
67        ack: SeqNum,
68        snd_nxt: SeqNum,
69        high_rxt: Option<SeqNum>,
70        sack_blocks: &SackBlocks,
71        smss: Mss,
72    ) -> bool {
73        let Self { acked_ranges, pipe, is_lost_seqnum_end } = self;
74
75        // If we receive an ACK that is after SND.NXT, this must be a very
76        // delayed acknowledgement post a retransmission event. The state
77        // machine will eventually move SND.NXT to account for this, but this
78        // violates the scoreboard's expectations.
79        //
80        // Because we need to ensure the pipe is updated accordingly and any
81        // previous SACK ranges are cleared, process this as if it was a full
82        // cumulative ack.
83        let snd_nxt = snd_nxt.latest(ack);
84
85        // Fast exit if there's nothing interesting to do.
86        if acked_ranges.is_empty() && sack_blocks.is_empty() {
87            // Ack must not be after snd_nxt.
88            *pipe = u32::try_from(snd_nxt - ack).unwrap();
89            return false;
90        }
91
92        // Update the scoreboard with the cumulative acknowledgement.
93        //
94        // A note here: we discard all the sacked ranges that start at or after
95        // the acknowledged number. If there is intersection we must assume that
96        // the peer reneged the block.
97        acked_ranges.discard_starting_at_or_before(ack);
98
99        // Insert each valid block in acked ranges.
100        let new = sack_blocks.iter_skip_invalid().fold(false, |new, sack_block| {
101            // NB: SackBlock type here ensures this is a valid non-empty range.
102            let (start, end) = sack_block.into_parts();
103            if start.before_or_eq(ack) || end.after(snd_nxt) {
104                // Ignore block that is not in the expected range [ack, snd_nxt].
105                return new;
106            }
107            let changed = acked_ranges.insert(start..end, ());
108
109            new || changed
110        });
111
112        let sacked_byte_threshold = sacked_bytes_threshold(smss);
113        let high_rxt = high_rxt.unwrap_or(ack);
114        let get_pipe_increment = |hole: SeqRange<bool>| {
115            // From RFC 6675, where S1 is a single
116            // sequence number:
117            //
118            // (a) If IsLost (S1) returns false: Pipe is incremented by 1
119            // octet.
120            // (b) If S1 <= HighRxt: Pipe is incremented by 1 octet.
121            let mut pipe = 0u32;
122            let is_lost = *(hole.meta());
123            if !is_lost {
124                pipe = pipe.saturating_add(hole.len());
125            }
126
127            if let Some(hole) = hole.cap_right(high_rxt) {
128                pipe = pipe.saturating_add(hole.len());
129            }
130            pipe
131        };
132
133        enum IsLostPivotInfo {
134            Looking { sacked_count: usize, sacked_bytes: u32 },
135            Found(SeqNum),
136        }
137
138        // Recalculate pipe and update IsLost in the collection.
139        //
140        // We iterate acked_ranges in reverse order so we can fold over the
141        // total number of ranges and SACKed bytes that come *after* the range
142        // we operate on at each point.
143        let (new_pipe, pivot_info, later_start) = acked_ranges.iter_mut().rev().fold(
144            (0u32, IsLostPivotInfo::Looking { sacked_count: 0, sacked_bytes: 0 }, snd_nxt),
145            |(pipe, pivot_info, later_start), acked_range| {
146                let (pivot_info, later_is_lost) = match pivot_info {
147                    IsLostPivotInfo::Looking { sacked_count, sacked_bytes } => {
148                        let sacked_count = sacked_count + 1;
149                        let sacked_bytes = sacked_bytes + acked_range.len();
150                        // From RFC 6675, IsLost is defined as:
151                        //
152                        //  The routine returns true when either DupThresh
153                        //  discontiguous SACKed sequences have arrived
154                        //  above 'SeqNum' or more than (DupThresh - 1) *
155                        //  SMSS bytes with sequence numbers greater than
156                        //  'SeqNum' have been SACKed.
157                        //
158                        // is_lost here tells us whether all sequence
159                        // numbers before the start of the current
160                        // acked_range are lost.
161                        let is_lost = sacked_count >= usize::from(DUP_ACK_THRESHOLD)
162                            || sacked_bytes > sacked_byte_threshold;
163                        let pivot_info = if is_lost {
164                            IsLostPivotInfo::Found(acked_range.start())
165                        } else {
166                            IsLostPivotInfo::Looking { sacked_count, sacked_bytes }
167                        };
168                        (pivot_info, false)
169                    }
170                    // We already found the sequence number before which all
171                    // holes are lost.
172                    IsLostPivotInfo::Found(seq_num) => (IsLostPivotInfo::Found(seq_num), true),
173                };
174
175                // Increment pipe. From RFC 6675:
176                //
177                //  After initializing pipe to zero, the following steps are
178                //  taken for each octet 'S1' in the sequence space between
179                //  HighACK and HighData that has not been SACKed[...]
180                //
181                // So pipe is only calculated for the gaps between the acked
182                // ranges, i.e., from the current end to the start of the
183                // later block.
184                let pipe = if let Some(hole) =
185                    SeqRange::new(acked_range.end()..later_start, later_is_lost)
186                {
187                    pipe.saturating_add(get_pipe_increment(hole))
188                } else {
189                    // An empty hole can only happen in the first iteration,
190                    // when the right edge is SND.NXT.
191                    assert_eq!(later_start, snd_nxt);
192                    pipe
193                };
194
195                (pipe, pivot_info, acked_range.start())
196            },
197        );
198        *is_lost_seqnum_end = match pivot_info {
199            IsLostPivotInfo::Looking { sacked_count: _, sacked_bytes: _ } => None,
200            IsLostPivotInfo::Found(seq_num) => Some(seq_num),
201        };
202
203        // Add the final hole between cumulative ack and first sack block
204        // and finalize the pipe value.
205        *pipe = match SeqRange::new(ack..later_start, is_lost_seqnum_end.is_some()) {
206            Some(first_hole) => new_pipe.saturating_add(get_pipe_increment(first_hole)),
207            None => {
208                // An empty first hole can only happen if we don't have any
209                // sack blocks, and ACK is equal to SND.NXT.
210                assert_eq!(ack, snd_nxt);
211                new_pipe
212            }
213        };
214
215        new
216    }
217
218    pub(crate) fn has_sack_info(&self) -> bool {
219        !self.acked_ranges.is_empty()
220    }
221
222    /// Helper to check rule (2) from [RFC 6675 section 5]:
223    ///
224    /// > (2) If DupAcks < DupThresh but IsLost (HighACK + 1) returns true --
225    /// > indicating at least three segments have arrived above the current
226    /// > cumulative acknowledgment point, which is taken to indicate loss -- go
227    /// > to step (4).
228    ///
229    /// [RFC 6675 section 5]: https://datatracker.ietf.org/doc/html/rfc6675#section-4
230    pub(crate) fn is_first_hole_lost(&self) -> bool {
231        // If we have defined any value for the lost sequence number it means
232        // the first hole must be lost.
233        self.is_lost_seqnum_end.is_some()
234    }
235
236    pub(crate) fn pipe(&self) -> u32 {
237        self.pipe
238    }
239
240    /// Increments the pipe value kept by the scoreboard by `value`.
241    ///
242    /// Note that [`SackScoreboard::process_ack`] always updates the pipe value
243    /// based on the scoreboard. Whenever a segment is sent, we must increment
244    /// the pipe value so the estimate of total bytes in transit is always up to
245    /// date until the next ACK arrives.
246    pub(crate) fn increment_pipe(&mut self, value: u32) {
247        self.pipe = self.pipe.saturating_add(value);
248    }
249
250    /// Returns the right-side-bounded unsacked range starting at or later than
251    /// `marker`.
252    ///
253    /// Returns `None` if `mark` is not a hole bounded to the right side by a
254    /// received SACK.
255    ///
256    /// Returns a [`SeqRange`] whose metadata is a boolean indicating if this
257    /// range is considered lost.
258    pub(crate) fn first_unsacked_range_from(&self, mark: SeqNum) -> Option<SeqRange<bool>> {
259        let Self { acked_ranges, pipe: _, is_lost_seqnum_end } = self;
260        match acked_ranges.first_hole_on_or_after(mark) {
261            FirstHoleResult::None => None,
262            FirstHoleResult::Right(right) => {
263                SeqRange::new(mark..right.start(), is_lost_seqnum_end.is_some())
264            }
265            FirstHoleResult::Both(left, right) => {
266                let left = left.end().latest(mark);
267                SeqRange::new(
268                    left..right.start(),
269                    is_lost_seqnum_end.is_some_and(|s| left.before(s)),
270                )
271            }
272        }
273    }
274
275    /// Returns the end of the sequence number range in the scoreboard, if there
276    /// are any ranges tracked.
277    pub(crate) fn right_edge(&self) -> Option<SeqNum> {
278        let Self { acked_ranges, pipe: _, is_lost_seqnum_end: _ } = self;
279        acked_ranges.last().map(|seq_range| seq_range.end())
280    }
281
282    pub(crate) fn on_retransmission_timeout(&mut self) {
283        let Self { acked_ranges, pipe, is_lost_seqnum_end } = self;
284        // RFC 2018 says that we MUST clear all SACK information on a
285        // retransmission timeout.
286        //
287        // RFC 6675 changes that to a SHOULD keep SACK information on a
288        // retransmission timeout, but doesn't quite specify how to deal with
289        // the SACKed ranges post the timeout. Notably, the pipe estimate is
290        // very clearly off post an RTO.
291        //
292        // Given that, the conservative thing to do here is to clear the
293        // scoreboard and reset the pipe so estimates can be based again on the
294        // rewound value of SND.NXT and the eventually retransmitted SACK blocks
295        // that we may get post the RTO event. Note that `process_ack` ignores
296        // any SACK blocks post SND.NXT in order to maintain the pipe variable
297        // sensible as well.
298        //
299        // See:
300        // - https://datatracker.ietf.org/doc/html/rfc2018
301        // - https://datatracker.ietf.org/doc/html/rfc6675
302        *pipe = 0;
303        acked_ranges.clear();
304        *is_lost_seqnum_end = None;
305    }
306
307    pub(crate) fn on_mss_update(
308        &mut self,
309        snd_una: SeqNum,
310        snd_nxt: SeqNum,
311        high_rxt: Option<SeqNum>,
312        mss: Mss,
313    ) {
314        // When MSS updates, we must recalculate so we know what frames are
315        // considered lost or not.
316        //
317        // Notably, this will also update the pipe variable so we have a new
318        // estimate of bytes in flight with a new value for snd_nxt.
319        //
320        // Given we don't detect renegging, this is equivalent to processing an
321        // ACK at the given parameters and without any SACK blocks.
322        let _: bool = self.process_ack(snd_una, snd_nxt, high_rxt, &SackBlocks::EMPTY, mss);
323    }
324}
325
326/// Returns the threshold over which a sequence number is considered lost per
327/// the definition of `IsLost` in [RFC 6675 section 4].
328///
329/// [RFC 6675 section 4]:
330///     https://datatracker.ietf.org/doc/html/rfc6675#section-4
331fn sacked_bytes_threshold(mss: Mss) -> u32 {
332    u32::from(DUP_ACK_THRESHOLD - 1) * u32::from(mss)
333}
334
335#[cfg(test)]
336mod test {
337    use core::num::NonZeroU16;
338    use core::ops::Range;
339    use test_case::test_case;
340
341    use super::*;
342    use crate::internal::seq_ranges::SeqRange;
343    use crate::internal::testutil;
344
345    const TEST_MSS: Mss = Mss(NonZeroU16::new(50).unwrap());
346
347    fn seq_ranges(iter: impl IntoIterator<Item = Range<u32>>) -> SeqRanges<()> {
348        iter.into_iter()
349            .map(|Range { start, end }| {
350                SeqRange::new(SeqNum::new(start)..SeqNum::new(end), ()).unwrap()
351            })
352            .collect()
353    }
354
355    impl SackScoreboard {
356        fn sacked_bytes(&self) -> u32 {
357            self.acked_ranges.iter().map(|seq_range| seq_range.len()).sum()
358        }
359    }
360
361    #[test]
362    fn process_ack_noop_if_empty() {
363        let mut sb = SackScoreboard::default();
364        let ack = SeqNum::new(1);
365        let snd_nxt = SeqNum::new(100);
366        let high_rxt = None;
367        assert!(!sb.process_ack(ack, snd_nxt, high_rxt, &SackBlocks::default(), TEST_MSS));
368        assert!(sb.acked_ranges.is_empty());
369        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap());
370    }
371
372    #[test]
373    fn process_ack_ignores_bad_blocks() {
374        let mut sb = SackScoreboard::default();
375        let ack = SeqNum::new(5);
376        let snd_nxt = SeqNum::new(100);
377        let high_rxt = None;
378        // Ignores everything that doesn't match the cumulative ack.
379        assert!(!sb.process_ack(
380            ack,
381            snd_nxt,
382            high_rxt,
383            &testutil::sack_blocks([0..1, 4..6, 5..10]),
384            TEST_MSS
385        ));
386        assert!(sb.acked_ranges.is_empty());
387
388        // Ignores everything past snd_nxt.
389        assert!(!sb.process_ack(
390            ack,
391            snd_nxt,
392            high_rxt,
393            &testutil::sack_blocks([100..200, 50..150]),
394            TEST_MSS
395        ));
396        assert!(sb.acked_ranges.is_empty());
397    }
398
399    #[test]
400    fn process_ack_cumulative_ack() {
401        let mut sb = SackScoreboard::default();
402        let ack = SeqNum::new(5);
403        let snd_nxt = SeqNum::new(100);
404        let high_rxt = None;
405        let blocks = testutil::sack_blocks([20..30]);
406        assert!(sb.process_ack(ack, snd_nxt, high_rxt, &blocks, TEST_MSS));
407        let expect_ranges = seq_ranges([20..30]);
408        assert_eq!(sb.acked_ranges, expect_ranges);
409        assert_eq!(sb.is_lost_seqnum_end, None);
410        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap() - sb.sacked_bytes());
411
412        let ack = SeqNum::new(10);
413        assert!(!sb.process_ack(ack, snd_nxt, high_rxt, &blocks, TEST_MSS));
414        assert_eq!(sb.acked_ranges, expect_ranges);
415        assert_eq!(sb.is_lost_seqnum_end, None);
416        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap() - sb.sacked_bytes());
417    }
418
419    #[test]
420    fn process_ack_is_lost_dup_thresh() {
421        let mut sb = SackScoreboard::default();
422        let ack = SeqNum::new(5);
423        let snd_nxt = SeqNum::new(100);
424        let high_rxt = None;
425
426        let block1 = 20..30;
427        let block2 = 35..40;
428        let block3 = 45..50;
429
430        assert!(sb.process_ack(
431            ack,
432            snd_nxt,
433            high_rxt,
434            &testutil::sack_blocks([block1.clone(), block2.clone(), block3.clone()]),
435            TEST_MSS
436        ));
437        assert_eq!(sb.acked_ranges, seq_ranges([block1.clone(), block2, block3]));
438        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(block1.start)));
439        assert_eq!(
440            sb.pipe,
441            u32::try_from(snd_nxt - ack).unwrap()
442                - sb.sacked_bytes()
443                - (block1.start - u32::from(ack))
444        );
445    }
446
447    #[test]
448    fn process_ack_pipe_rule_a() {
449        let mut sb = SackScoreboard::default();
450        let ack = SeqNum::new(5);
451        let snd_nxt = SeqNum::new(500);
452        let high_rxt = None;
453        let small_block = 20..30;
454        let large_block_start = 35;
455        let large_block = large_block_start..(large_block_start + sacked_bytes_threshold(TEST_MSS));
456
457        assert!(sb.process_ack(
458            ack,
459            snd_nxt,
460            high_rxt,
461            &testutil::sack_blocks([small_block.clone(), large_block.clone()]),
462            TEST_MSS
463        ));
464        // Large block is exactly at the limit of the hole to its left being
465        // considered lost as well.
466        assert_eq!(sb.acked_ranges, seq_ranges([small_block.clone(), large_block.clone()]));
467        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(small_block.start)));
468        assert_eq!(
469            sb.pipe,
470            u32::try_from(snd_nxt - ack).unwrap()
471                - sb.sacked_bytes()
472                - (small_block.start - u32::from(ack))
473        );
474
475        // Now increase the large block by one.
476        let large_block = large_block.start..(large_block.end + 1);
477        assert!(sb.process_ack(
478            ack,
479            snd_nxt,
480            high_rxt,
481            &testutil::sack_blocks([small_block.clone(), large_block.clone()]),
482            TEST_MSS
483        ));
484        // Now the hole to the left of large block is also considered lost.
485        assert_eq!(sb.acked_ranges, seq_ranges([small_block.clone(), large_block.clone()]));
486        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(large_block.start)));
487        assert_eq!(
488            sb.pipe,
489            u32::try_from(snd_nxt - ack).unwrap()
490                - sb.sacked_bytes()
491                - (small_block.start - u32::from(ack))
492                - (large_block.start - small_block.end)
493        );
494    }
495
496    #[test]
497    fn process_ack_pipe_rule_b() {
498        let ack = SeqNum::new(5);
499        let snd_nxt = SeqNum::new(500);
500        let first_block = 20..30;
501        let second_block = 40..50;
502
503        let blocks = testutil::sack_blocks([first_block.clone(), second_block.clone()]);
504
505        // Extract the baseline pipe that if we receive an ACK with the
506        // parameters above but without a HighRxt value.
507        let baseline = {
508            let mut sb = SackScoreboard::default();
509            assert!(sb.process_ack(ack, snd_nxt, None, &blocks, TEST_MSS));
510            sb.pipe
511        };
512
513        // Drive HighRxt across the entire possible sequence number range that
514        // we expect to see it and check the pipe value is changing accordingly.
515        let hole1 = (u32::from(ack)..first_block.start).map(|seq| (seq, true));
516        let block1 = first_block.clone().map(|seq| (seq, false));
517        let hole2 = (first_block.end..second_block.start).map(|seq| (seq, true));
518        let block2 = second_block.map(|seq| (seq, false));
519        // Shift expecting an increment one over because HighRxt starting at
520        // HighAck is expected to be a zero contribution. This aligns the
521        // off-by-one in the expectations.
522        let iter =
523            hole1.chain(block1).chain(hole2).chain(block2).scan(false, |prev, (seq, sacked)| {
524                let expect_increment = core::mem::replace(prev, sacked);
525                Some((seq, expect_increment))
526            });
527
528        let _: u32 = iter.fold(0u32, |total, (seq, expect_increment)| {
529            let total = total + u32::from(expect_increment);
530            let mut sb = SackScoreboard::default();
531            assert!(sb.process_ack(ack, snd_nxt, Some(SeqNum::new(seq)), &blocks, TEST_MSS));
532            assert_eq!(sb.pipe - baseline, total, "at {seq}");
533            total
534        });
535    }
536
537    #[test]
538    fn process_ack_simple() {
539        let mut sb = SackScoreboard::default();
540        let ack = SeqNum::new(5);
541        let snd_nxt = SeqNum::new(500);
542        let high_rxt = None;
543
544        // Receive a single cumulative ack up to ack.
545        assert!(!sb.process_ack(ack, snd_nxt, high_rxt, &SackBlocks::default(), TEST_MSS));
546        assert_eq!(sb.acked_ranges, SeqRanges::default());
547        assert_eq!(sb.is_lost_seqnum_end, None);
548        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap());
549
550        // Cumulative ack doesn't move, 1 SACK range signaling loss is received.
551        let sack1 = 10..(10 + sacked_bytes_threshold(TEST_MSS) + 1);
552        assert!(sb.process_ack(
553            ack,
554            snd_nxt,
555            high_rxt,
556            &testutil::sack_blocks([sack1.clone()]),
557            TEST_MSS
558        ));
559        assert_eq!(sb.acked_ranges, seq_ranges([sack1.clone()]));
560        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(sack1.start)));
561        assert_eq!(
562            sb.pipe,
563            u32::try_from(snd_nxt - ack).unwrap()
564                - sb.sacked_bytes()
565                - (sack1.start - u32::from(ack))
566        );
567
568        // Another SACK range comes in, at the end of this transmission block.
569        let sack2 = (u32::from(snd_nxt) - u32::from(TEST_MSS))..u32::from(snd_nxt);
570        assert!(sb.process_ack(
571            ack,
572            snd_nxt,
573            high_rxt,
574            &testutil::sack_blocks([sack1.clone(), sack2.clone()]),
575            TEST_MSS
576        ));
577        assert_eq!(sb.acked_ranges, seq_ranges([sack1.clone(), sack2.clone()]));
578        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(sack1.start)));
579        assert_eq!(
580            sb.pipe,
581            u32::try_from(snd_nxt - ack).unwrap()
582                - sb.sacked_bytes()
583                - (sack1.start - u32::from(ack))
584        );
585
586        // Cumulative acknowledge the first SACK range.
587        let ack = SeqNum::new(sack1.end);
588        assert!(!sb.process_ack(
589            ack,
590            snd_nxt,
591            high_rxt,
592            &testutil::sack_blocks([sack2.clone()]),
593            TEST_MSS
594        ));
595        assert_eq!(sb.acked_ranges, seq_ranges([sack2]));
596        assert_eq!(sb.is_lost_seqnum_end, None);
597        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap() - sb.sacked_bytes());
598
599        // Cumulative acknowledge all the transmission.
600        assert!(!sb.process_ack(snd_nxt, snd_nxt, high_rxt, &SackBlocks::default(), TEST_MSS));
601        assert_eq!(sb.acked_ranges, SeqRanges::default());
602        assert_eq!(sb.is_lost_seqnum_end, None);
603        assert_eq!(sb.pipe, 0);
604    }
605
606    #[test]
607    fn ack_after_snd_nxt() {
608        let mut sb = SackScoreboard::default();
609        let ack = SeqNum::new(5);
610        let snd_nxt = SeqNum::new(500);
611        let high_rxt = None;
612        let block = 10..20;
613        assert!(sb.process_ack(
614            ack,
615            snd_nxt,
616            high_rxt,
617            &testutil::sack_blocks([block.clone()]),
618            TEST_MSS
619        ));
620        assert_eq!(sb.acked_ranges, seq_ranges([block.clone()]));
621        assert_eq!(sb.is_lost_seqnum_end, None);
622        assert_eq!(sb.pipe, u32::try_from(snd_nxt - ack).unwrap() - sb.sacked_bytes());
623
624        // SND.NXT rewinds after RTO.
625        let snd_nxt = ack;
626        // But we receive an ACK post the kept block.
627        let ack = SeqNum::new(block.end);
628        assert!(ack.after(snd_nxt));
629        assert!(!sb.process_ack(ack, snd_nxt, high_rxt, &SackBlocks::default(), TEST_MSS));
630        assert_eq!(sb.acked_ranges, SeqRanges::default());
631        assert_eq!(sb.is_lost_seqnum_end, None);
632        assert_eq!(sb.pipe, 0);
633    }
634
635    #[test]
636    fn first_unsacked_range_from() {
637        let mut sb = SackScoreboard::default();
638        let ack = SeqNum::new(5);
639        let snd_nxt = SeqNum::new(60);
640        let high_rxt = None;
641        let block1 = 10..20;
642        let block2 = 30..40;
643        let block3 = 50..60;
644        assert!(sb.process_ack(
645            ack,
646            snd_nxt,
647            high_rxt,
648            &testutil::sack_blocks([block1.clone(), block2.clone(), block3.clone()]),
649            TEST_MSS
650        ));
651        assert_eq!(sb.acked_ranges, seq_ranges([block1.clone(), block2.clone(), block3.clone()]));
652        assert_eq!(sb.is_lost_seqnum_end, Some(SeqNum::new(block1.start)));
653        for high_rxt in u32::from(ack)..u32::from(snd_nxt) {
654            let expect = if high_rxt < block3.start {
655                let lost = high_rxt < block1.start;
656                let (start, end) = if high_rxt < block1.start {
657                    (high_rxt, block1.start)
658                } else if high_rxt < block2.start {
659                    (block1.end.max(high_rxt), block2.start)
660                } else {
661                    (block2.end.max(high_rxt), block3.start)
662                };
663                Some(SeqRange::new(SeqNum::new(start)..SeqNum::new(end), lost).unwrap())
664            } else {
665                None
666            };
667            assert_eq!(
668                sb.first_unsacked_range_from(SeqNum::new(high_rxt)),
669                expect,
670                "high_rxt={high_rxt}"
671            );
672        }
673    }
674
675    #[test_case(0)]
676    #[test_case(1)]
677    #[test_case(2)]
678    #[test_case(3)]
679    fn lost_holes(count: usize) {
680        let mss = u32::from(TEST_MSS);
681        let mut sb = SackScoreboard::default();
682        let ack = SeqNum::new(0);
683
684        // Process enough evenly spaced SACK blocks (each 1 MSS and spaced by
685        // 1 MSS) until we have enough gaps to match count.
686        let mut start = u32::from(ack) + mss;
687        let sack_blocks_count = usize::from(DUP_ACK_THRESHOLD) + count - 1;
688        for _ in 0..sack_blocks_count {
689            let block = start..(start + mss);
690            assert!(sb.process_ack(
691                ack,
692                SeqNum::new(block.end),
693                None,
694                &testutil::sack_blocks([block.clone()]),
695                TEST_MSS
696            ));
697            start = block.end + mss;
698        }
699        assert_eq!(sb.acked_ranges.iter().count(), sack_blocks_count);
700
701        // The IsLost marker should be the start of the (count-1)-th sacked
702        // block.
703        let expect =
704            count.checked_sub(1).map(|i| sb.acked_ranges.iter().skip(i).next().unwrap().start());
705        assert_eq!(sb.is_lost_seqnum_end, expect);
706
707        // Verify the IsLost check when getting each unsacked range back.
708        let mut start = ack;
709        for i in 0..sack_blocks_count {
710            let expect_lost = i < count;
711            let expect_range = SeqRange::new(start..(start + mss), expect_lost).unwrap();
712            assert_eq!(sb.first_unsacked_range_from(start), Some(expect_range.clone()));
713            start = expect_range.end() + mss;
714        }
715    }
716}