Skip to content

Commit

Permalink
Implement Comparer trait for the Database (#83)
Browse files Browse the repository at this point in the history
  • Loading branch information
tirix authored Apr 5, 2022
1 parent 25f1f55 commit 69ec97b
Show file tree
Hide file tree
Showing 4 changed files with 25 additions and 9 deletions.
1 change: 1 addition & 0 deletions src/lib.rs
Original file line number Diff line number Diff line change
Expand Up @@ -89,6 +89,7 @@ pub use formula::Formula;
pub use formula::FormulaRef;
pub use formula::Label;
pub use formula::Symbol;
pub use segment::Comparer;
pub use segment_set::SourceInfo;
pub use statement::as_str;
pub use statement::Span;
Expand Down
7 changes: 2 additions & 5 deletions src/outline.rs
Original file line number Diff line number Diff line change
Expand Up @@ -347,11 +347,8 @@ impl<'a> OutlineNodeRef<'a> {
let end = if let Some(next_stmt) = self.next_after_children() {
next_stmt.get_statement().span().start - 1
} else {
database
.parse_result()
.source_info(stmt.segment.id)
.span
.end
let segment_span = database.parse_result().source_info(stmt.segment.id).span;
segment_span.end - segment_span.start
};
Span { start, end }
}
Expand Down
24 changes: 22 additions & 2 deletions src/segment.rs
Original file line number Diff line number Diff line change
Expand Up @@ -13,7 +13,7 @@ use crate::{
StatementAddress, StatementIndex, StatementIter, SymbolDef, TokenAddress, DUMMY_STATEMENT,
NO_STATEMENT,
},
StatementRef,
Database, StatementRef,
};

/// Data structure which tracks the logical order of segment IDs, since they are
Expand Down Expand Up @@ -127,16 +127,18 @@ impl<'a> DoubleEndedIterator for SegmentIter<'a> {
}

/// A trait for objects which can be used to order other datatypes.
pub(crate) trait Comparer<T> {
pub trait Comparer<T> {
/// Compares two objects, like `Ord::cmp`, but with additional state data
/// from the comparer that can be used for the ordering.
fn cmp(&self, left: &T, right: &T) -> Ordering;

/// Returns true if the `left` argument is (strictly) less than the `right` argument
#[inline]
fn lt(&self, left: &T, right: &T) -> bool {
self.cmp(left, right) == Ordering::Less
}

/// Returns true if the `left` argument is less or equal to the `right` argument
#[inline]
fn le(&self, left: &T, right: &T) -> bool {
self.cmp(left, right) != Ordering::Greater
Expand All @@ -163,6 +165,24 @@ impl Comparer<TokenAddress> for SegmentOrder {
}
}

impl Comparer<SegmentId> for Database {
fn cmp(&self, left: &SegmentId, right: &SegmentId) -> Ordering {
self.parse_result().order.cmp(left, right)
}
}

impl Comparer<StatementAddress> for Database {
fn cmp(&self, left: &StatementAddress, right: &StatementAddress) -> Ordering {
self.parse_result().order.cmp(left, right)
}
}

impl Comparer<TokenAddress> for Database {
fn cmp(&self, left: &TokenAddress, right: &TokenAddress) -> Ordering {
self.parse_result().order.cmp(left, right)
}
}

impl<'a, T, C: Comparer<T>> Comparer<T> for &'a C {
fn cmp(&self, left: &T, right: &T) -> Ordering {
(*self).cmp(left, right)
Expand Down
2 changes: 0 additions & 2 deletions src/statement.rs
Original file line number Diff line number Diff line change
Expand Up @@ -153,9 +153,7 @@ impl StatementAddress {
pub const fn new(segment_id: SegmentId, index: StatementIndex) -> Self {
StatementAddress { segment_id, index }
}
}

impl StatementAddress {
/// Convert a statement address into a statement range from here to the
/// logical end of the database.
#[inline]
Expand Down

0 comments on commit 69ec97b

Please sign in to comment.