{-# LANGUAGE DataKinds #-}
{-# LANGUAGE UndecidableInstances #-}
module Agda.Syntax.Position
(
Position
, PositionWithoutFile
, Position'(..)
, SrcFile
, RangeFile(..)
, mkRangeFile
, positionInvariant
, startPos
, startPos'
, movePos
, movePosByString
, backupPos
, Interval
, IntervalWithoutFile
, Interval'(Interval, iStart', iEnd')
, intervalInvariant
, iStart
, iEnd
, posToInterval
, getIntervalFile
, iLength
, fuseIntervals
, Range
, RangeWithoutFile
, Range'(..)
, rangeInvariant
, consecutiveAndSeparated
, intervalsToRange
, intervalToRange
, rangeFromAbsolutePath
, rangeIntervals
, rangeFile
, rangeModule'
, rangeModule
, rightMargin
, noRange
, posToRange, posToRange'
, rStart, rStart'
, rEnd, rEnd'
, rangeToInterval
, rangeToIntervalWithFile
, continuous
, continuousPerLine
, PrintRange(..)
, HasRange(..)
, HasRangeWithoutFile(..)
, SetRange(..)
, KillRange(..)
, KillRangeT
, killRangeMap
, KILLRANGE(..)
, withRangeOf
, fuseRange
, fuseRanges
, beginningOf
, beginningOfFile
, interleaveRanges
) where
import Prelude hiding ( null )
import Control.DeepSeq
import Control.Monad
import Control.Monad.Writer (runWriter, tell)
import qualified Data.Foldable as Fold
import Data.Function (on)
import Data.List (sort)
import Data.Map (Map)
import qualified Data.Map as Map
import Data.Set (Set)
import qualified Data.Set as Set
import Data.Sequence (Seq)
import qualified Data.Sequence as Seq
import Data.Semigroup (Semigroup(..))
import Data.Void
import Data.Word (Word32)
import GHC.Generics (Generic)
import Agda.Syntax.TopLevelModuleName.Boot (TopLevelModuleName'(..))
import Agda.Utils.FileName
import Agda.Utils.List
import Agda.Utils.List1 (List1)
import Agda.Utils.List2 (List2)
import qualified Agda.Utils.Maybe.Strict as Strict
import Agda.Utils.Null
import Agda.Utils.Permutation
import Agda.Utils.Set1 (Set1)
import qualified Agda.Utils.Set1 as Set1
import Agda.Utils.TypeLevel (IsBase, All, Domains)
import Agda.Utils.Tuple (sortPair)
import Agda.Utils.Impossible
data Position' a = Pn
{ forall a. Position' a -> a
srcFile :: !a
, forall a. Position' a -> Word32
posPos :: !Word32
, forall a. Position' a -> Word32
posLine :: !Word32
, forall a. Position' a -> Word32
posCol :: !Word32
}
deriving (Int -> Position' a -> ShowS
[Position' a] -> ShowS
Position' a -> String
(Int -> Position' a -> ShowS)
-> (Position' a -> String)
-> ([Position' a] -> ShowS)
-> Show (Position' a)
forall a. Show a => Int -> Position' a -> ShowS
forall a. Show a => [Position' a] -> ShowS
forall a. Show a => Position' a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> Position' a -> ShowS
showsPrec :: Int -> Position' a -> ShowS
$cshow :: forall a. Show a => Position' a -> String
show :: Position' a -> String
$cshowList :: forall a. Show a => [Position' a] -> ShowS
showList :: [Position' a] -> ShowS
Show, (forall a b. (a -> b) -> Position' a -> Position' b)
-> (forall a b. a -> Position' b -> Position' a)
-> Functor Position'
forall a b. a -> Position' b -> Position' a
forall a b. (a -> b) -> Position' a -> Position' b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Position' a -> Position' b
fmap :: forall a b. (a -> b) -> Position' a -> Position' b
$c<$ :: forall a b. a -> Position' b -> Position' a
<$ :: forall a b. a -> Position' b -> Position' a
Functor, (forall m. Monoid m => Position' m -> m)
-> (forall m a. Monoid m => (a -> m) -> Position' a -> m)
-> (forall m a. Monoid m => (a -> m) -> Position' a -> m)
-> (forall a b. (a -> b -> b) -> b -> Position' a -> b)
-> (forall a b. (a -> b -> b) -> b -> Position' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Position' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Position' a -> b)
-> (forall a. (a -> a -> a) -> Position' a -> a)
-> (forall a. (a -> a -> a) -> Position' a -> a)
-> (forall a. Position' a -> [a])
-> (forall a. Position' a -> Bool)
-> (forall a. Position' a -> Int)
-> (forall a. Eq a => a -> Position' a -> Bool)
-> (forall a. Ord a => Position' a -> a)
-> (forall a. Ord a => Position' a -> a)
-> (forall a. Num a => Position' a -> a)
-> (forall a. Num a => Position' a -> a)
-> Foldable Position'
forall a. Eq a => a -> Position' a -> Bool
forall a. Num a => Position' a -> a
forall a. Ord a => Position' a -> a
forall m. Monoid m => Position' m -> m
forall a. Position' a -> Bool
forall a. Position' a -> Int
forall a. Position' a -> [a]
forall a. (a -> a -> a) -> Position' a -> a
forall m a. Monoid m => (a -> m) -> Position' a -> m
forall b a. (b -> a -> b) -> b -> Position' a -> b
forall a b. (a -> b -> b) -> b -> Position' a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Position' m -> m
fold :: forall m. Monoid m => Position' m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Position' a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Position' a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Position' a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Position' a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Position' a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Position' a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Position' a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Position' a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Position' a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Position' a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Position' a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Position' a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Position' a -> a
foldr1 :: forall a. (a -> a -> a) -> Position' a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Position' a -> a
foldl1 :: forall a. (a -> a -> a) -> Position' a -> a
$ctoList :: forall a. Position' a -> [a]
toList :: forall a. Position' a -> [a]
$cnull :: forall a. Position' a -> Bool
null :: forall a. Position' a -> Bool
$clength :: forall a. Position' a -> Int
length :: forall a. Position' a -> Int
$celem :: forall a. Eq a => a -> Position' a -> Bool
elem :: forall a. Eq a => a -> Position' a -> Bool
$cmaximum :: forall a. Ord a => Position' a -> a
maximum :: forall a. Ord a => Position' a -> a
$cminimum :: forall a. Ord a => Position' a -> a
minimum :: forall a. Ord a => Position' a -> a
$csum :: forall a. Num a => Position' a -> a
sum :: forall a. Num a => Position' a -> a
$cproduct :: forall a. Num a => Position' a -> a
product :: forall a. Num a => Position' a -> a
Foldable, Functor Position'
Foldable Position'
(Functor Position', Foldable Position') =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Position' a -> f (Position' b))
-> (forall (f :: * -> *) a.
Applicative f =>
Position' (f a) -> f (Position' a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Position' a -> m (Position' b))
-> (forall (m :: * -> *) a.
Monad m =>
Position' (m a) -> m (Position' a))
-> Traversable Position'
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a.
Monad m =>
Position' (m a) -> m (Position' a)
forall (f :: * -> *) a.
Applicative f =>
Position' (f a) -> f (Position' a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Position' a -> m (Position' b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Position' a -> f (Position' b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Position' a -> f (Position' b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Position' a -> f (Position' b)
$csequenceA :: forall (f :: * -> *) a.
Applicative f =>
Position' (f a) -> f (Position' a)
sequenceA :: forall (f :: * -> *) a.
Applicative f =>
Position' (f a) -> f (Position' a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Position' a -> m (Position' b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Position' a -> m (Position' b)
$csequence :: forall (m :: * -> *) a.
Monad m =>
Position' (m a) -> m (Position' a)
sequence :: forall (m :: * -> *) a.
Monad m =>
Position' (m a) -> m (Position' a)
Traversable, (forall x. Position' a -> Rep (Position' a) x)
-> (forall x. Rep (Position' a) x -> Position' a)
-> Generic (Position' a)
forall x. Rep (Position' a) x -> Position' a
forall x. Position' a -> Rep (Position' a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Position' a) x -> Position' a
forall a x. Position' a -> Rep (Position' a) x
$cfrom :: forall a x. Position' a -> Rep (Position' a) x
from :: forall x. Position' a -> Rep (Position' a) x
$cto :: forall a x. Rep (Position' a) x -> Position' a
to :: forall x. Rep (Position' a) x -> Position' a
Generic)
positionInvariant :: Position' a -> Bool
positionInvariant :: forall a. Position' a -> Bool
positionInvariant Position' a
p =
Position' a -> Word32
forall a. Position' a -> Word32
posPos Position' a
p Word32 -> Word32 -> Bool
forall a. Ord a => a -> a -> Bool
> Word32
0 Bool -> Bool -> Bool
&& Position' a -> Word32
forall a. Position' a -> Word32
posLine Position' a
p Word32 -> Word32 -> Bool
forall a. Ord a => a -> a -> Bool
> Word32
0 Bool -> Bool -> Bool
&& Position' a -> Word32
forall a. Position' a -> Word32
posCol Position' a
p Word32 -> Word32 -> Bool
forall a. Ord a => a -> a -> Bool
> Word32
0
importantPart :: Position' a -> (a, Word32)
importantPart :: forall a. Position' a -> (a, Word32)
importantPart Position' a
p = (Position' a -> a
forall a. Position' a -> a
srcFile Position' a
p, Position' a -> Word32
forall a. Position' a -> Word32
posPos Position' a
p)
instance Eq a => Eq (Position' a) where
== :: Position' a -> Position' a -> Bool
(==) = (a, Word32) -> (a, Word32) -> Bool
forall a. Eq a => a -> a -> Bool
(==) ((a, Word32) -> (a, Word32) -> Bool)
-> (Position' a -> (a, Word32))
-> Position' a
-> Position' a
-> Bool
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` Position' a -> (a, Word32)
forall a. Position' a -> (a, Word32)
importantPart
instance Ord a => Ord (Position' a) where
compare :: Position' a -> Position' a -> Ordering
compare = (a, Word32) -> (a, Word32) -> Ordering
forall a. Ord a => a -> a -> Ordering
compare ((a, Word32) -> (a, Word32) -> Ordering)
-> (Position' a -> (a, Word32))
-> Position' a
-> Position' a
-> Ordering
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` Position' a -> (a, Word32)
forall a. Position' a -> (a, Word32)
importantPart
type SrcFile = Strict.Maybe RangeFile
data RangeFile = RangeFile
{ RangeFile -> AbsolutePath
rangeFilePath :: !AbsolutePath
, RangeFile -> Maybe (TopLevelModuleName' Range)
rangeFileName :: !(Maybe (TopLevelModuleName' Range))
}
deriving (Int -> RangeFile -> ShowS
[RangeFile] -> ShowS
RangeFile -> String
(Int -> RangeFile -> ShowS)
-> (RangeFile -> String)
-> ([RangeFile] -> ShowS)
-> Show RangeFile
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> RangeFile -> ShowS
showsPrec :: Int -> RangeFile -> ShowS
$cshow :: RangeFile -> String
show :: RangeFile -> String
$cshowList :: [RangeFile] -> ShowS
showList :: [RangeFile] -> ShowS
Show, (forall x. RangeFile -> Rep RangeFile x)
-> (forall x. Rep RangeFile x -> RangeFile) -> Generic RangeFile
forall x. Rep RangeFile x -> RangeFile
forall x. RangeFile -> Rep RangeFile x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. RangeFile -> Rep RangeFile x
from :: forall x. RangeFile -> Rep RangeFile x
$cto :: forall x. Rep RangeFile x -> RangeFile
to :: forall x. Rep RangeFile x -> RangeFile
Generic)
mkRangeFile :: AbsolutePath -> Maybe (TopLevelModuleName' Range) -> RangeFile
mkRangeFile :: AbsolutePath -> Maybe (TopLevelModuleName' Range) -> RangeFile
mkRangeFile AbsolutePath
f Maybe (TopLevelModuleName' Range)
top = RangeFile
{ rangeFilePath :: AbsolutePath
rangeFilePath = AbsolutePath
f
, rangeFileName :: Maybe (TopLevelModuleName' Range)
rangeFileName = KillRangeT (Maybe (TopLevelModuleName' Range))
forall a. KillRange a => KillRangeT a
killRange Maybe (TopLevelModuleName' Range)
top
}
instance Eq RangeFile where
== :: RangeFile -> RangeFile -> Bool
(==) = Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Bool
forall a. Eq a => a -> a -> Bool
(==) (Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Bool)
-> (RangeFile -> Maybe (TopLevelModuleName' Range))
-> RangeFile
-> RangeFile
-> Bool
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` RangeFile -> Maybe (TopLevelModuleName' Range)
rangeFileName
instance Ord RangeFile where
compare :: RangeFile -> RangeFile -> Ordering
compare = Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Ordering
forall a. Ord a => a -> a -> Ordering
compare (Maybe (TopLevelModuleName' Range)
-> Maybe (TopLevelModuleName' Range) -> Ordering)
-> (RangeFile -> Maybe (TopLevelModuleName' Range))
-> RangeFile
-> RangeFile
-> Ordering
forall b c a. (b -> b -> c) -> (a -> b) -> a -> a -> c
`on` RangeFile -> Maybe (TopLevelModuleName' Range)
rangeFileName
instance NFData RangeFile where
rnf :: RangeFile -> ()
rnf (RangeFile AbsolutePath
_ Maybe (TopLevelModuleName' Range)
n) = Maybe (TopLevelModuleName' Range) -> ()
forall a. NFData a => a -> ()
rnf Maybe (TopLevelModuleName' Range)
n
type Position = Position' SrcFile
type PositionWithoutFile = Position' ()
instance NFData Position where
rnf :: Position -> ()
rnf = (Position -> () -> ()
forall a b. a -> b -> b
`seq` ())
instance NFData PositionWithoutFile where
rnf :: PositionWithoutFile -> ()
rnf = (PositionWithoutFile -> () -> ()
forall a b. a -> b -> b
`seq` ())
data Interval' a = Interval
{ forall a. Interval' a -> a
getIntervalFile :: a
, forall a. Interval' a -> PositionWithoutFile
iStart' :: !PositionWithoutFile
, forall a. Interval' a -> PositionWithoutFile
iEnd' :: !PositionWithoutFile
}
deriving (Int -> Interval' a -> ShowS
[Interval' a] -> ShowS
Interval' a -> String
(Int -> Interval' a -> ShowS)
-> (Interval' a -> String)
-> ([Interval' a] -> ShowS)
-> Show (Interval' a)
forall a. Show a => Int -> Interval' a -> ShowS
forall a. Show a => [Interval' a] -> ShowS
forall a. Show a => Interval' a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> Interval' a -> ShowS
showsPrec :: Int -> Interval' a -> ShowS
$cshow :: forall a. Show a => Interval' a -> String
show :: Interval' a -> String
$cshowList :: forall a. Show a => [Interval' a] -> ShowS
showList :: [Interval' a] -> ShowS
Show, Interval' a -> Interval' a -> Bool
(Interval' a -> Interval' a -> Bool)
-> (Interval' a -> Interval' a -> Bool) -> Eq (Interval' a)
forall a. Eq a => Interval' a -> Interval' a -> Bool
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: forall a. Eq a => Interval' a -> Interval' a -> Bool
== :: Interval' a -> Interval' a -> Bool
$c/= :: forall a. Eq a => Interval' a -> Interval' a -> Bool
/= :: Interval' a -> Interval' a -> Bool
Eq, Eq (Interval' a)
Eq (Interval' a) =>
(Interval' a -> Interval' a -> Ordering)
-> (Interval' a -> Interval' a -> Bool)
-> (Interval' a -> Interval' a -> Bool)
-> (Interval' a -> Interval' a -> Bool)
-> (Interval' a -> Interval' a -> Bool)
-> (Interval' a -> Interval' a -> Interval' a)
-> (Interval' a -> Interval' a -> Interval' a)
-> Ord (Interval' a)
Interval' a -> Interval' a -> Bool
Interval' a -> Interval' a -> Ordering
Interval' a -> Interval' a -> Interval' a
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
forall a. Ord a => Eq (Interval' a)
forall a. Ord a => Interval' a -> Interval' a -> Bool
forall a. Ord a => Interval' a -> Interval' a -> Ordering
forall a. Ord a => Interval' a -> Interval' a -> Interval' a
$ccompare :: forall a. Ord a => Interval' a -> Interval' a -> Ordering
compare :: Interval' a -> Interval' a -> Ordering
$c< :: forall a. Ord a => Interval' a -> Interval' a -> Bool
< :: Interval' a -> Interval' a -> Bool
$c<= :: forall a. Ord a => Interval' a -> Interval' a -> Bool
<= :: Interval' a -> Interval' a -> Bool
$c> :: forall a. Ord a => Interval' a -> Interval' a -> Bool
> :: Interval' a -> Interval' a -> Bool
$c>= :: forall a. Ord a => Interval' a -> Interval' a -> Bool
>= :: Interval' a -> Interval' a -> Bool
$cmax :: forall a. Ord a => Interval' a -> Interval' a -> Interval' a
max :: Interval' a -> Interval' a -> Interval' a
$cmin :: forall a. Ord a => Interval' a -> Interval' a -> Interval' a
min :: Interval' a -> Interval' a -> Interval' a
Ord, (forall a b. (a -> b) -> Interval' a -> Interval' b)
-> (forall a b. a -> Interval' b -> Interval' a)
-> Functor Interval'
forall a b. a -> Interval' b -> Interval' a
forall a b. (a -> b) -> Interval' a -> Interval' b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Interval' a -> Interval' b
fmap :: forall a b. (a -> b) -> Interval' a -> Interval' b
$c<$ :: forall a b. a -> Interval' b -> Interval' a
<$ :: forall a b. a -> Interval' b -> Interval' a
Functor, (forall m. Monoid m => Interval' m -> m)
-> (forall m a. Monoid m => (a -> m) -> Interval' a -> m)
-> (forall m a. Monoid m => (a -> m) -> Interval' a -> m)
-> (forall a b. (a -> b -> b) -> b -> Interval' a -> b)
-> (forall a b. (a -> b -> b) -> b -> Interval' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Interval' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Interval' a -> b)
-> (forall a. (a -> a -> a) -> Interval' a -> a)
-> (forall a. (a -> a -> a) -> Interval' a -> a)
-> (forall a. Interval' a -> [a])
-> (forall a. Interval' a -> Bool)
-> (forall a. Interval' a -> Int)
-> (forall a. Eq a => a -> Interval' a -> Bool)
-> (forall a. Ord a => Interval' a -> a)
-> (forall a. Ord a => Interval' a -> a)
-> (forall a. Num a => Interval' a -> a)
-> (forall a. Num a => Interval' a -> a)
-> Foldable Interval'
forall a. Eq a => a -> Interval' a -> Bool
forall a. Num a => Interval' a -> a
forall a. Ord a => Interval' a -> a
forall m. Monoid m => Interval' m -> m
forall a. Interval' a -> Bool
forall a. Interval' a -> Int
forall a. Interval' a -> [a]
forall a. (a -> a -> a) -> Interval' a -> a
forall m a. Monoid m => (a -> m) -> Interval' a -> m
forall b a. (b -> a -> b) -> b -> Interval' a -> b
forall a b. (a -> b -> b) -> b -> Interval' a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Interval' m -> m
fold :: forall m. Monoid m => Interval' m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Interval' a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Interval' a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Interval' a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Interval' a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Interval' a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Interval' a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Interval' a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Interval' a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Interval' a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Interval' a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Interval' a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Interval' a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Interval' a -> a
foldr1 :: forall a. (a -> a -> a) -> Interval' a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Interval' a -> a
foldl1 :: forall a. (a -> a -> a) -> Interval' a -> a
$ctoList :: forall a. Interval' a -> [a]
toList :: forall a. Interval' a -> [a]
$cnull :: forall a. Interval' a -> Bool
null :: forall a. Interval' a -> Bool
$clength :: forall a. Interval' a -> Int
length :: forall a. Interval' a -> Int
$celem :: forall a. Eq a => a -> Interval' a -> Bool
elem :: forall a. Eq a => a -> Interval' a -> Bool
$cmaximum :: forall a. Ord a => Interval' a -> a
maximum :: forall a. Ord a => Interval' a -> a
$cminimum :: forall a. Ord a => Interval' a -> a
minimum :: forall a. Ord a => Interval' a -> a
$csum :: forall a. Num a => Interval' a -> a
sum :: forall a. Num a => Interval' a -> a
$cproduct :: forall a. Num a => Interval' a -> a
product :: forall a. Num a => Interval' a -> a
Foldable, Functor Interval'
Foldable Interval'
(Functor Interval', Foldable Interval') =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Interval' a -> f (Interval' b))
-> (forall (f :: * -> *) a.
Applicative f =>
Interval' (f a) -> f (Interval' a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Interval' a -> m (Interval' b))
-> (forall (m :: * -> *) a.
Monad m =>
Interval' (m a) -> m (Interval' a))
-> Traversable Interval'
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a.
Monad m =>
Interval' (m a) -> m (Interval' a)
forall (f :: * -> *) a.
Applicative f =>
Interval' (f a) -> f (Interval' a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Interval' a -> m (Interval' b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Interval' a -> f (Interval' b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Interval' a -> f (Interval' b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Interval' a -> f (Interval' b)
$csequenceA :: forall (f :: * -> *) a.
Applicative f =>
Interval' (f a) -> f (Interval' a)
sequenceA :: forall (f :: * -> *) a.
Applicative f =>
Interval' (f a) -> f (Interval' a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Interval' a -> m (Interval' b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Interval' a -> m (Interval' b)
$csequence :: forall (m :: * -> *) a.
Monad m =>
Interval' (m a) -> m (Interval' a)
sequence :: forall (m :: * -> *) a.
Monad m =>
Interval' (m a) -> m (Interval' a)
Traversable, (forall x. Interval' a -> Rep (Interval' a) x)
-> (forall x. Rep (Interval' a) x -> Interval' a)
-> Generic (Interval' a)
forall x. Rep (Interval' a) x -> Interval' a
forall x. Interval' a -> Rep (Interval' a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Interval' a) x -> Interval' a
forall a x. Interval' a -> Rep (Interval' a) x
$cfrom :: forall a x. Interval' a -> Rep (Interval' a) x
from :: forall x. Interval' a -> Rep (Interval' a) x
$cto :: forall a x. Rep (Interval' a) x -> Interval' a
to :: forall x. Rep (Interval' a) x -> Interval' a
Generic)
type Interval = Interval' SrcFile
type IntervalWithoutFile = Interval' ()
instance NFData Interval where
rnf :: Interval -> ()
rnf = (Interval -> () -> ()
forall a b. a -> b -> b
`seq` ())
instance NFData IntervalWithoutFile where
rnf :: IntervalWithoutFile -> ()
rnf = (IntervalWithoutFile -> () -> ()
forall a b. a -> b -> b
`seq` ())
intervalInvariant :: Ord a => Interval' a -> Bool
intervalInvariant :: forall a. Ord a => Interval' a -> Bool
intervalInvariant Interval' a
i = [Bool] -> Bool
forall (t :: * -> *). Foldable t => t Bool -> Bool
and
[ Position' a -> Bool
forall a. Position' a -> Bool
positionInvariant (Position' a -> Bool) -> Position' a -> Bool
forall a b. (a -> b) -> a -> b
$ Interval' a -> Position' a
forall a. Interval' a -> Position' a
iStart Interval' a
i
, Position' a -> Bool
forall a. Position' a -> Bool
positionInvariant (Position' a -> Bool) -> Position' a -> Bool
forall a b. (a -> b) -> a -> b
$ Interval' a -> Position' a
forall a. Interval' a -> Position' a
iEnd Interval' a
i
, Interval' a -> Position' a
forall a. Interval' a -> Position' a
iStart Interval' a
i Position' a -> Position' a -> Bool
forall a. Ord a => a -> a -> Bool
<= Interval' a -> Position' a
forall a. Interval' a -> Position' a
iEnd Interval' a
i
]
iStart :: Interval' a -> Position' a
iStart :: forall a. Interval' a -> Position' a
iStart (Interval a
f PositionWithoutFile
s PositionWithoutFile
_) = a
f a -> PositionWithoutFile -> Position' a
forall a b. a -> Position' b -> Position' a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ PositionWithoutFile
s
iEnd :: Interval' a -> Position' a
iEnd :: forall a. Interval' a -> Position' a
iEnd (Interval a
f PositionWithoutFile
_ PositionWithoutFile
e) = a
f a -> PositionWithoutFile -> Position' a
forall a b. a -> Position' b -> Position' a
forall (f :: * -> *) a b. Functor f => a -> f b -> f a
<$ PositionWithoutFile
e
posToInterval ::
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
posToInterval :: forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
posToInterval a
f PositionWithoutFile
p1 PositionWithoutFile
p2 = (PositionWithoutFile -> PositionWithoutFile -> Interval' a)
-> (PositionWithoutFile, PositionWithoutFile) -> Interval' a
forall a b c. (a -> b -> c) -> (a, b) -> c
uncurry (a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval a
f) ((PositionWithoutFile, PositionWithoutFile) -> Interval' a)
-> (PositionWithoutFile, PositionWithoutFile) -> Interval' a
forall a b. (a -> b) -> a -> b
$ (PositionWithoutFile, PositionWithoutFile)
-> (PositionWithoutFile, PositionWithoutFile)
forall a. Ord a => (a, a) -> (a, a)
sortPair (PositionWithoutFile
p1, PositionWithoutFile
p2)
iLength :: Interval' a -> Word32
iLength :: forall a. Interval' a -> Word32
iLength Interval' a
i = Position' a -> Word32
forall a. Position' a -> Word32
posPos (Interval' a -> Position' a
forall a. Interval' a -> Position' a
iEnd Interval' a
i) Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
- Position' a -> Word32
forall a. Position' a -> Word32
posPos (Interval' a -> Position' a
forall a. Interval' a -> Position' a
iStart Interval' a
i)
data Range' a
= NoRange
| Range !a (Seq IntervalWithoutFile)
deriving
(Int -> Range' a -> ShowS
[Range' a] -> ShowS
Range' a -> String
(Int -> Range' a -> ShowS)
-> (Range' a -> String) -> ([Range' a] -> ShowS) -> Show (Range' a)
forall a. Show a => Int -> Range' a -> ShowS
forall a. Show a => [Range' a] -> ShowS
forall a. Show a => Range' a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> Range' a -> ShowS
showsPrec :: Int -> Range' a -> ShowS
$cshow :: forall a. Show a => Range' a -> String
show :: Range' a -> String
$cshowList :: forall a. Show a => [Range' a] -> ShowS
showList :: [Range' a] -> ShowS
Show, Range' a -> Range' a -> Bool
(Range' a -> Range' a -> Bool)
-> (Range' a -> Range' a -> Bool) -> Eq (Range' a)
forall a. Eq a => Range' a -> Range' a -> Bool
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: forall a. Eq a => Range' a -> Range' a -> Bool
== :: Range' a -> Range' a -> Bool
$c/= :: forall a. Eq a => Range' a -> Range' a -> Bool
/= :: Range' a -> Range' a -> Bool
Eq, Eq (Range' a)
Eq (Range' a) =>
(Range' a -> Range' a -> Ordering)
-> (Range' a -> Range' a -> Bool)
-> (Range' a -> Range' a -> Bool)
-> (Range' a -> Range' a -> Bool)
-> (Range' a -> Range' a -> Bool)
-> (Range' a -> Range' a -> Range' a)
-> (Range' a -> Range' a -> Range' a)
-> Ord (Range' a)
Range' a -> Range' a -> Bool
Range' a -> Range' a -> Ordering
Range' a -> Range' a -> Range' a
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
forall a. Ord a => Eq (Range' a)
forall a. Ord a => Range' a -> Range' a -> Bool
forall a. Ord a => Range' a -> Range' a -> Ordering
forall a. Ord a => Range' a -> Range' a -> Range' a
$ccompare :: forall a. Ord a => Range' a -> Range' a -> Ordering
compare :: Range' a -> Range' a -> Ordering
$c< :: forall a. Ord a => Range' a -> Range' a -> Bool
< :: Range' a -> Range' a -> Bool
$c<= :: forall a. Ord a => Range' a -> Range' a -> Bool
<= :: Range' a -> Range' a -> Bool
$c> :: forall a. Ord a => Range' a -> Range' a -> Bool
> :: Range' a -> Range' a -> Bool
$c>= :: forall a. Ord a => Range' a -> Range' a -> Bool
>= :: Range' a -> Range' a -> Bool
$cmax :: forall a. Ord a => Range' a -> Range' a -> Range' a
max :: Range' a -> Range' a -> Range' a
$cmin :: forall a. Ord a => Range' a -> Range' a -> Range' a
min :: Range' a -> Range' a -> Range' a
Ord, (forall a b. (a -> b) -> Range' a -> Range' b)
-> (forall a b. a -> Range' b -> Range' a) -> Functor Range'
forall a b. a -> Range' b -> Range' a
forall a b. (a -> b) -> Range' a -> Range' b
forall (f :: * -> *).
(forall a b. (a -> b) -> f a -> f b)
-> (forall a b. a -> f b -> f a) -> Functor f
$cfmap :: forall a b. (a -> b) -> Range' a -> Range' b
fmap :: forall a b. (a -> b) -> Range' a -> Range' b
$c<$ :: forall a b. a -> Range' b -> Range' a
<$ :: forall a b. a -> Range' b -> Range' a
Functor, (forall m. Monoid m => Range' m -> m)
-> (forall m a. Monoid m => (a -> m) -> Range' a -> m)
-> (forall m a. Monoid m => (a -> m) -> Range' a -> m)
-> (forall a b. (a -> b -> b) -> b -> Range' a -> b)
-> (forall a b. (a -> b -> b) -> b -> Range' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Range' a -> b)
-> (forall b a. (b -> a -> b) -> b -> Range' a -> b)
-> (forall a. (a -> a -> a) -> Range' a -> a)
-> (forall a. (a -> a -> a) -> Range' a -> a)
-> (forall a. Range' a -> [a])
-> (forall a. Range' a -> Bool)
-> (forall a. Range' a -> Int)
-> (forall a. Eq a => a -> Range' a -> Bool)
-> (forall a. Ord a => Range' a -> a)
-> (forall a. Ord a => Range' a -> a)
-> (forall a. Num a => Range' a -> a)
-> (forall a. Num a => Range' a -> a)
-> Foldable Range'
forall a. Eq a => a -> Range' a -> Bool
forall a. Num a => Range' a -> a
forall a. Ord a => Range' a -> a
forall m. Monoid m => Range' m -> m
forall a. Range' a -> Bool
forall a. Range' a -> Int
forall a. Range' a -> [a]
forall a. (a -> a -> a) -> Range' a -> a
forall m a. Monoid m => (a -> m) -> Range' a -> m
forall b a. (b -> a -> b) -> b -> Range' a -> b
forall a b. (a -> b -> b) -> b -> Range' a -> b
forall (t :: * -> *).
(forall m. Monoid m => t m -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall m a. Monoid m => (a -> m) -> t a -> m)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall a b. (a -> b -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall b a. (b -> a -> b) -> b -> t a -> b)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. (a -> a -> a) -> t a -> a)
-> (forall a. t a -> [a])
-> (forall a. t a -> Bool)
-> (forall a. t a -> Int)
-> (forall a. Eq a => a -> t a -> Bool)
-> (forall a. Ord a => t a -> a)
-> (forall a. Ord a => t a -> a)
-> (forall a. Num a => t a -> a)
-> (forall a. Num a => t a -> a)
-> Foldable t
$cfold :: forall m. Monoid m => Range' m -> m
fold :: forall m. Monoid m => Range' m -> m
$cfoldMap :: forall m a. Monoid m => (a -> m) -> Range' a -> m
foldMap :: forall m a. Monoid m => (a -> m) -> Range' a -> m
$cfoldMap' :: forall m a. Monoid m => (a -> m) -> Range' a -> m
foldMap' :: forall m a. Monoid m => (a -> m) -> Range' a -> m
$cfoldr :: forall a b. (a -> b -> b) -> b -> Range' a -> b
foldr :: forall a b. (a -> b -> b) -> b -> Range' a -> b
$cfoldr' :: forall a b. (a -> b -> b) -> b -> Range' a -> b
foldr' :: forall a b. (a -> b -> b) -> b -> Range' a -> b
$cfoldl :: forall b a. (b -> a -> b) -> b -> Range' a -> b
foldl :: forall b a. (b -> a -> b) -> b -> Range' a -> b
$cfoldl' :: forall b a. (b -> a -> b) -> b -> Range' a -> b
foldl' :: forall b a. (b -> a -> b) -> b -> Range' a -> b
$cfoldr1 :: forall a. (a -> a -> a) -> Range' a -> a
foldr1 :: forall a. (a -> a -> a) -> Range' a -> a
$cfoldl1 :: forall a. (a -> a -> a) -> Range' a -> a
foldl1 :: forall a. (a -> a -> a) -> Range' a -> a
$ctoList :: forall a. Range' a -> [a]
toList :: forall a. Range' a -> [a]
$cnull :: forall a. Range' a -> Bool
null :: forall a. Range' a -> Bool
$clength :: forall a. Range' a -> Int
length :: forall a. Range' a -> Int
$celem :: forall a. Eq a => a -> Range' a -> Bool
elem :: forall a. Eq a => a -> Range' a -> Bool
$cmaximum :: forall a. Ord a => Range' a -> a
maximum :: forall a. Ord a => Range' a -> a
$cminimum :: forall a. Ord a => Range' a -> a
minimum :: forall a. Ord a => Range' a -> a
$csum :: forall a. Num a => Range' a -> a
sum :: forall a. Num a => Range' a -> a
$cproduct :: forall a. Num a => Range' a -> a
product :: forall a. Num a => Range' a -> a
Foldable, Functor Range'
Foldable Range'
(Functor Range', Foldable Range') =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Range' a -> f (Range' b))
-> (forall (f :: * -> *) a.
Applicative f =>
Range' (f a) -> f (Range' a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Range' a -> m (Range' b))
-> (forall (m :: * -> *) a.
Monad m =>
Range' (m a) -> m (Range' a))
-> Traversable Range'
forall (t :: * -> *).
(Functor t, Foldable t) =>
(forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> t a -> f (t b))
-> (forall (f :: * -> *) a. Applicative f => t (f a) -> f (t a))
-> (forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> t a -> m (t b))
-> (forall (m :: * -> *) a. Monad m => t (m a) -> m (t a))
-> Traversable t
forall (m :: * -> *) a. Monad m => Range' (m a) -> m (Range' a)
forall (f :: * -> *) a.
Applicative f =>
Range' (f a) -> f (Range' a)
forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Range' a -> m (Range' b)
forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Range' a -> f (Range' b)
$ctraverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Range' a -> f (Range' b)
traverse :: forall (f :: * -> *) a b.
Applicative f =>
(a -> f b) -> Range' a -> f (Range' b)
$csequenceA :: forall (f :: * -> *) a.
Applicative f =>
Range' (f a) -> f (Range' a)
sequenceA :: forall (f :: * -> *) a.
Applicative f =>
Range' (f a) -> f (Range' a)
$cmapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Range' a -> m (Range' b)
mapM :: forall (m :: * -> *) a b.
Monad m =>
(a -> m b) -> Range' a -> m (Range' b)
$csequence :: forall (m :: * -> *) a. Monad m => Range' (m a) -> m (Range' a)
sequence :: forall (m :: * -> *) a. Monad m => Range' (m a) -> m (Range' a)
Traversable, (forall x. Range' a -> Rep (Range' a) x)
-> (forall x. Rep (Range' a) x -> Range' a) -> Generic (Range' a)
forall x. Rep (Range' a) x -> Range' a
forall x. Range' a -> Rep (Range' a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (Range' a) x -> Range' a
forall a x. Range' a -> Rep (Range' a) x
$cfrom :: forall a x. Range' a -> Rep (Range' a) x
from :: forall x. Range' a -> Rep (Range' a) x
$cto :: forall a x. Rep (Range' a) x -> Range' a
to :: forall x. Rep (Range' a) x -> Range' a
Generic)
type Range = Range' SrcFile
type RangeWithoutFile = Range' ()
instance NFData a => NFData (Range' a)
instance Null (Range' a) where
null :: Range' a -> Bool
null Range' a
NoRange = Bool
True
null Range{} = Bool
False
empty :: Range' a
empty = Range' a
forall a. Range' a
NoRange
instance Eq a => Semigroup (Range' a) where
Range' a
NoRange <> :: Range' a -> Range' a -> Range' a
<> Range' a
r = Range' a
r
Range' a
r <> Range' a
NoRange = Range' a
r
Range a
f Seq IntervalWithoutFile
is <> Range a
f' Seq IntervalWithoutFile
is'
| a
f a -> a -> Bool
forall a. Eq a => a -> a -> Bool
/= a
f' = Range' a
forall a. HasCallStack => a
__IMPOSSIBLE__
| Bool
otherwise = a -> Seq IntervalWithoutFile -> Range' a
forall a. a -> Seq IntervalWithoutFile -> Range' a
Range a
f (Seq IntervalWithoutFile
is Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. Semigroup a => a -> a -> a
<> Seq IntervalWithoutFile
is')
instance Eq a => Monoid (Range' a) where
mempty :: Range' a
mempty = Range' a
forall a. Null a => a
empty
mappend :: Range' a -> Range' a -> Range' a
mappend = Range' a -> Range' a -> Range' a
forall a. Semigroup a => a -> a -> a
(<>)
rangeIntervals :: Range' a -> [IntervalWithoutFile]
rangeIntervals :: forall a. Range' a -> [IntervalWithoutFile]
rangeIntervals Range' a
NoRange = []
rangeIntervals (Range a
_ Seq IntervalWithoutFile
is) = Seq IntervalWithoutFile -> [IntervalWithoutFile]
forall a. Seq a -> [a]
forall (t :: * -> *) a. Foldable t => t a -> [a]
Fold.toList Seq IntervalWithoutFile
is
intervalsToRange :: a -> [IntervalWithoutFile] -> Range' a
intervalsToRange :: forall a. a -> [IntervalWithoutFile] -> Range' a
intervalsToRange a
_ [] = Range' a
forall a. Range' a
NoRange
intervalsToRange a
f [IntervalWithoutFile]
is = a -> Seq IntervalWithoutFile -> Range' a
forall a. a -> Seq IntervalWithoutFile -> Range' a
Range a
f ([IntervalWithoutFile] -> Seq IntervalWithoutFile
forall a. [a] -> Seq a
Seq.fromList [IntervalWithoutFile]
is)
consecutiveAndSeparated :: Ord a => [Interval' a] -> Bool
consecutiveAndSeparated :: forall a. Ord a => [Interval' a] -> Bool
consecutiveAndSeparated [Interval' a]
is =
(Interval' a -> Bool) -> [Interval' a] -> Bool
forall (t :: * -> *) a. Foldable t => (a -> Bool) -> t a -> Bool
all Interval' a -> Bool
forall a. Ord a => Interval' a -> Bool
intervalInvariant [Interval' a]
is
Bool -> Bool -> Bool
&&
[a] -> Bool
forall a. Eq a => [a] -> Bool
allEqual ((Interval' a -> a) -> [Interval' a] -> [a]
forall a b. (a -> b) -> [a] -> [b]
map (Position' a -> a
forall a. Position' a -> a
srcFile (Position' a -> a)
-> (Interval' a -> Position' a) -> Interval' a -> a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Interval' a -> Position' a
forall a. Interval' a -> Position' a
iStart) [Interval' a]
is)
Bool -> Bool -> Bool
&&
(Interval' a -> Interval' a -> Bool) -> [Interval' a] -> Bool
forall a. (a -> a -> Bool) -> [a] -> Bool
allConsecutive (\ Interval' a
i Interval' a
j -> Interval' a -> Position' a
forall a. Interval' a -> Position' a
iEnd Interval' a
i Position' a -> Position' a -> Bool
forall a. Ord a => a -> a -> Bool
< Interval' a -> Position' a
forall a. Interval' a -> Position' a
iStart Interval' a
j) [Interval' a]
is
rangeInvariant :: Ord a => Range' a -> Bool
rangeInvariant :: forall a. Ord a => Range' a -> Bool
rangeInvariant Range' a
r =
[IntervalWithoutFile] -> Bool
forall a. Ord a => [Interval' a] -> Bool
consecutiveAndSeparated (Range' a -> [IntervalWithoutFile]
forall a. Range' a -> [IntervalWithoutFile]
rangeIntervals Range' a
r)
Bool -> Bool -> Bool
&&
case Range' a
r of
Range a
_ Seq IntervalWithoutFile
is -> Bool -> Bool
not (Seq IntervalWithoutFile -> Bool
forall a. Null a => a -> Bool
null Seq IntervalWithoutFile
is)
Range' a
NoRange -> Bool
True
rangeFile :: Range -> SrcFile
rangeFile :: Range -> SrcFile
rangeFile Range
NoRange = SrcFile
forall a. Maybe a
Strict.Nothing
rangeFile (Range SrcFile
f Seq IntervalWithoutFile
_) = SrcFile
f
rangeModule' :: Range -> Maybe (Maybe (TopLevelModuleName' Range))
rangeModule' :: Range -> Maybe (Maybe (TopLevelModuleName' Range))
rangeModule' Range
NoRange = Maybe (Maybe (TopLevelModuleName' Range))
forall a. Maybe a
Nothing
rangeModule' (Range SrcFile
f Seq IntervalWithoutFile
_) = Maybe (TopLevelModuleName' Range)
-> Maybe (Maybe (TopLevelModuleName' Range))
forall a. a -> Maybe a
Just (Maybe (TopLevelModuleName' Range)
-> Maybe (Maybe (TopLevelModuleName' Range)))
-> Maybe (TopLevelModuleName' Range)
-> Maybe (Maybe (TopLevelModuleName' Range))
forall a b. (a -> b) -> a -> b
$ case SrcFile
f of
SrcFile
Strict.Nothing -> Maybe (TopLevelModuleName' Range)
forall a. Maybe a
Nothing
Strict.Just RangeFile
f -> RangeFile -> Maybe (TopLevelModuleName' Range)
rangeFileName RangeFile
f
rangeModule :: Range -> Maybe (TopLevelModuleName' Range)
rangeModule :: Range -> Maybe (TopLevelModuleName' Range)
rangeModule = Maybe (Maybe (TopLevelModuleName' Range))
-> Maybe (TopLevelModuleName' Range)
forall (m :: * -> *) a. Monad m => m (m a) -> m a
join (Maybe (Maybe (TopLevelModuleName' Range))
-> Maybe (TopLevelModuleName' Range))
-> (Range -> Maybe (Maybe (TopLevelModuleName' Range)))
-> Range
-> Maybe (TopLevelModuleName' Range)
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range -> Maybe (Maybe (TopLevelModuleName' Range))
rangeModule'
rightMargin :: Range -> Range
rightMargin :: Range -> Range
rightMargin r :: Range
r@Range
NoRange = Range
r
rightMargin r :: Range
r@(Range SrcFile
f Seq IntervalWithoutFile
is) = case Seq IntervalWithoutFile -> ViewR IntervalWithoutFile
forall a. Seq a -> ViewR a
Seq.viewr Seq IntervalWithoutFile
is of
ViewR IntervalWithoutFile
Seq.EmptyR -> Range
forall a. HasCallStack => a
__IMPOSSIBLE__
Seq IntervalWithoutFile
_ Seq.:> Interval () PositionWithoutFile
s PositionWithoutFile
e -> SrcFile -> IntervalWithoutFile -> Range
forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange SrcFile
f (()
-> PositionWithoutFile
-> PositionWithoutFile
-> IntervalWithoutFile
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval () PositionWithoutFile
e PositionWithoutFile
e)
newtype PrintRange a = PrintRange a
deriving (PrintRange a -> PrintRange a -> Bool
(PrintRange a -> PrintRange a -> Bool)
-> (PrintRange a -> PrintRange a -> Bool) -> Eq (PrintRange a)
forall a. Eq a => PrintRange a -> PrintRange a -> Bool
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: forall a. Eq a => PrintRange a -> PrintRange a -> Bool
== :: PrintRange a -> PrintRange a -> Bool
$c/= :: forall a. Eq a => PrintRange a -> PrintRange a -> Bool
/= :: PrintRange a -> PrintRange a -> Bool
Eq, Eq (PrintRange a)
Eq (PrintRange a) =>
(PrintRange a -> PrintRange a -> Ordering)
-> (PrintRange a -> PrintRange a -> Bool)
-> (PrintRange a -> PrintRange a -> Bool)
-> (PrintRange a -> PrintRange a -> Bool)
-> (PrintRange a -> PrintRange a -> Bool)
-> (PrintRange a -> PrintRange a -> PrintRange a)
-> (PrintRange a -> PrintRange a -> PrintRange a)
-> Ord (PrintRange a)
PrintRange a -> PrintRange a -> Bool
PrintRange a -> PrintRange a -> Ordering
PrintRange a -> PrintRange a -> PrintRange a
forall a.
Eq a =>
(a -> a -> Ordering)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> Bool)
-> (a -> a -> a)
-> (a -> a -> a)
-> Ord a
forall a. Ord a => Eq (PrintRange a)
forall a. Ord a => PrintRange a -> PrintRange a -> Bool
forall a. Ord a => PrintRange a -> PrintRange a -> Ordering
forall a. Ord a => PrintRange a -> PrintRange a -> PrintRange a
$ccompare :: forall a. Ord a => PrintRange a -> PrintRange a -> Ordering
compare :: PrintRange a -> PrintRange a -> Ordering
$c< :: forall a. Ord a => PrintRange a -> PrintRange a -> Bool
< :: PrintRange a -> PrintRange a -> Bool
$c<= :: forall a. Ord a => PrintRange a -> PrintRange a -> Bool
<= :: PrintRange a -> PrintRange a -> Bool
$c> :: forall a. Ord a => PrintRange a -> PrintRange a -> Bool
> :: PrintRange a -> PrintRange a -> Bool
$c>= :: forall a. Ord a => PrintRange a -> PrintRange a -> Bool
>= :: PrintRange a -> PrintRange a -> Bool
$cmax :: forall a. Ord a => PrintRange a -> PrintRange a -> PrintRange a
max :: PrintRange a -> PrintRange a -> PrintRange a
$cmin :: forall a. Ord a => PrintRange a -> PrintRange a -> PrintRange a
min :: PrintRange a -> PrintRange a -> PrintRange a
Ord, PrintRange a -> Range
(PrintRange a -> Range) -> HasRange (PrintRange a)
forall a. HasRange a => PrintRange a -> Range
forall a. (a -> Range) -> HasRange a
$cgetRange :: forall a. HasRange a => PrintRange a -> Range
getRange :: PrintRange a -> Range
HasRange, HasRange (PrintRange a)
HasRange (PrintRange a) =>
(Range -> PrintRange a -> PrintRange a) -> SetRange (PrintRange a)
Range -> PrintRange a -> PrintRange a
forall a. SetRange a => HasRange (PrintRange a)
forall a. SetRange a => Range -> PrintRange a -> PrintRange a
forall a. HasRange a => (Range -> a -> a) -> SetRange a
$csetRange :: forall a. SetRange a => Range -> PrintRange a -> PrintRange a
setRange :: Range -> PrintRange a -> PrintRange a
SetRange, KillRangeT (PrintRange a)
KillRangeT (PrintRange a) -> KillRange (PrintRange a)
forall a. KillRange a => KillRangeT (PrintRange a)
forall a. KillRangeT a -> KillRange a
$ckillRange :: forall a. KillRange a => KillRangeT (PrintRange a)
killRange :: KillRangeT (PrintRange a)
KillRange)
class HasRange a where
getRange :: a -> Range
default getRange :: (Foldable t, HasRange b, t b ~ a) => a -> Range
getRange = (b -> Range -> Range) -> Range -> t b -> Range
forall a b. (a -> b -> b) -> b -> t a -> b
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
Fold.foldr b -> Range -> Range
forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange Range
forall a. Range' a
noRange
{-# INLINABLE getRange #-}
instance HasRange Interval where
getRange :: Interval -> Range
getRange (Interval SrcFile
f PositionWithoutFile
p1 PositionWithoutFile
p2) = SrcFile -> IntervalWithoutFile -> Range
forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange SrcFile
f (()
-> PositionWithoutFile
-> PositionWithoutFile
-> IntervalWithoutFile
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval () PositionWithoutFile
p1 PositionWithoutFile
p2)
instance HasRange Range where
getRange :: Range -> Range
getRange = Range -> Range
forall a. a -> a
id
instance HasRange () where
getRange :: () -> Range
getRange ()
_ = Range
forall a. Range' a
noRange
instance HasRange Bool where
getRange :: Bool -> Range
getRange Bool
_ = Range
forall a. Range' a
noRange
instance HasRange (TopLevelModuleName' Range) where
getRange :: TopLevelModuleName' Range -> Range
getRange = TopLevelModuleName' Range -> Range
forall range. TopLevelModuleName' range -> range
moduleNameRange
instance SetRange (TopLevelModuleName' Range) where
setRange :: Range -> TopLevelModuleName' Range -> TopLevelModuleName' Range
setRange Range
r (TopLevelModuleName Range
_ ModuleNameHash
h TopLevelModuleNameParts
x) = Range
-> ModuleNameHash
-> TopLevelModuleNameParts
-> TopLevelModuleName' Range
forall range.
range
-> ModuleNameHash
-> TopLevelModuleNameParts
-> TopLevelModuleName' range
TopLevelModuleName Range
r ModuleNameHash
h TopLevelModuleNameParts
x
instance KillRange (TopLevelModuleName' Range) where
killRange :: TopLevelModuleName' Range -> TopLevelModuleName' Range
killRange (TopLevelModuleName Range
_ ModuleNameHash
h TopLevelModuleNameParts
x) = Range
-> ModuleNameHash
-> TopLevelModuleNameParts
-> TopLevelModuleName' Range
forall range.
range
-> ModuleNameHash
-> TopLevelModuleNameParts
-> TopLevelModuleName' range
TopLevelModuleName Range
forall a. Range' a
noRange ModuleNameHash
h TopLevelModuleNameParts
x
instance HasRange a => HasRange [a]
instance HasRange a => HasRange (List1 a)
instance HasRange a => HasRange (List2 a)
instance HasRange a => HasRange (Maybe a)
instance HasRange a => HasRange (Set1 a)
instance (HasRange a, HasRange b) => HasRange (a,b) where
getRange :: (a, b) -> Range
getRange = (a -> b -> Range) -> (a, b) -> Range
forall a b c. (a -> b -> c) -> (a, b) -> c
uncurry a -> b -> Range
forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange
instance (HasRange a, HasRange b, HasRange c) => HasRange (a,b,c) where
getRange :: (a, b, c) -> Range
getRange (a
x,b
y,c
z) = (a, (b, c)) -> Range
forall a. HasRange a => a -> Range
getRange (a
x,(b
y,c
z))
instance (HasRange a, HasRange b, HasRange c, HasRange d) => HasRange (a,b,c,d) where
getRange :: (a, b, c, d) -> Range
getRange (a
x,b
y,c
z,d
w) = (a, (b, (c, d))) -> Range
forall a. HasRange a => a -> Range
getRange (a
x,(b
y,(c
z,d
w)))
instance (HasRange a, HasRange b, HasRange c, HasRange d, HasRange e) => HasRange (a,b,c,d,e) where
getRange :: (a, b, c, d, e) -> Range
getRange (a
x,b
y,c
z,d
w,e
v) = (a, (b, (c, (d, e)))) -> Range
forall a. HasRange a => a -> Range
getRange (a
x,(b
y,(c
z,(d
w,e
v))))
instance (HasRange a, HasRange b, HasRange c, HasRange d, HasRange e, HasRange f) => HasRange (a,b,c,d,e,f) where
getRange :: (a, b, c, d, e, f) -> Range
getRange (a
x,b
y,c
z,d
w,e
v,f
u) = (a, (b, (c, (d, (e, f))))) -> Range
forall a. HasRange a => a -> Range
getRange (a
x,(b
y,(c
z,(d
w,(e
v,f
u)))))
instance (HasRange a, HasRange b, HasRange c, HasRange d, HasRange e, HasRange f, HasRange g) => HasRange (a,b,c,d,e,f,g) where
getRange :: (a, b, c, d, e, f, g) -> Range
getRange (a
x,b
y,c
z,d
w,e
v,f
u,g
t) = (a, (b, (c, (d, (e, (f, g)))))) -> Range
forall a. HasRange a => a -> Range
getRange (a
x,(b
y,(c
z,(d
w,(e
v,(f
u,g
t))))))
instance (HasRange a, HasRange b) => HasRange (Either a b) where
getRange :: Either a b -> Range
getRange = (a -> Range) -> (b -> Range) -> Either a b -> Range
forall a c b. (a -> c) -> (b -> c) -> Either a b -> c
either a -> Range
forall a. HasRange a => a -> Range
getRange b -> Range
forall a. HasRange a => a -> Range
getRange
class HasRangeWithoutFile a where
getRangeWithoutFile :: a -> RangeWithoutFile
default getRangeWithoutFile :: (Foldable t, HasRangeWithoutFile b, t b ~ a) => a -> RangeWithoutFile
getRangeWithoutFile = (b -> RangeWithoutFile -> RangeWithoutFile)
-> RangeWithoutFile -> t b -> RangeWithoutFile
forall a b. (a -> b -> b) -> b -> t a -> b
forall (t :: * -> *) a b.
Foldable t =>
(a -> b -> b) -> b -> t a -> b
Fold.foldr b -> RangeWithoutFile -> RangeWithoutFile
forall u t.
(HasRangeWithoutFile u, HasRangeWithoutFile t) =>
u -> t -> RangeWithoutFile
fuseRangeWithoutFile RangeWithoutFile
forall a. Range' a
noRange
{-# INLINABLE getRangeWithoutFile #-}
instance HasRangeWithoutFile IntervalWithoutFile where
getRangeWithoutFile :: IntervalWithoutFile -> RangeWithoutFile
getRangeWithoutFile = () -> IntervalWithoutFile -> RangeWithoutFile
forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange ()
instance HasRangeWithoutFile RangeWithoutFile where
getRangeWithoutFile :: RangeWithoutFile -> RangeWithoutFile
getRangeWithoutFile = RangeWithoutFile -> RangeWithoutFile
forall a. a -> a
id
instance HasRangeWithoutFile () where
getRangeWithoutFile :: () -> RangeWithoutFile
getRangeWithoutFile ()
_ = RangeWithoutFile
forall a. Range' a
noRange
instance HasRangeWithoutFile Bool where
getRangeWithoutFile :: Bool -> RangeWithoutFile
getRangeWithoutFile Bool
_ = RangeWithoutFile
forall a. Range' a
noRange
instance HasRangeWithoutFile a => HasRangeWithoutFile [a]
instance HasRangeWithoutFile a => HasRangeWithoutFile (List1 a)
instance HasRangeWithoutFile a => HasRangeWithoutFile (List2 a)
instance HasRangeWithoutFile a => HasRangeWithoutFile (Maybe a)
instance HasRangeWithoutFile a => HasRangeWithoutFile (Set1 a)
instance (HasRangeWithoutFile a, HasRangeWithoutFile b) => HasRangeWithoutFile (a,b) where
getRangeWithoutFile :: (a, b) -> RangeWithoutFile
getRangeWithoutFile = (a -> b -> RangeWithoutFile) -> (a, b) -> RangeWithoutFile
forall a b c. (a -> b -> c) -> (a, b) -> c
uncurry a -> b -> RangeWithoutFile
forall u t.
(HasRangeWithoutFile u, HasRangeWithoutFile t) =>
u -> t -> RangeWithoutFile
fuseRangeWithoutFile
instance (HasRangeWithoutFile a, HasRangeWithoutFile b, HasRangeWithoutFile c) => HasRangeWithoutFile (a,b,c) where
getRangeWithoutFile :: (a, b, c) -> RangeWithoutFile
getRangeWithoutFile (a
x,b
y,c
z) = (a, (b, c)) -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile (a
x,(b
y,c
z))
instance (HasRangeWithoutFile a, HasRangeWithoutFile b, HasRangeWithoutFile c, HasRangeWithoutFile d) => HasRangeWithoutFile (a,b,c,d) where
getRangeWithoutFile :: (a, b, c, d) -> RangeWithoutFile
getRangeWithoutFile (a
x,b
y,c
z,d
w) = (a, (b, (c, d))) -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile (a
x,(b
y,(c
z,d
w)))
instance (HasRangeWithoutFile a, HasRangeWithoutFile b, HasRangeWithoutFile c, HasRangeWithoutFile d, HasRangeWithoutFile e) => HasRangeWithoutFile (a,b,c,d,e) where
getRangeWithoutFile :: (a, b, c, d, e) -> RangeWithoutFile
getRangeWithoutFile (a
x,b
y,c
z,d
w,e
v) = (a, (b, (c, (d, e)))) -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile (a
x,(b
y,(c
z,(d
w,e
v))))
instance (HasRangeWithoutFile a, HasRangeWithoutFile b, HasRangeWithoutFile c, HasRangeWithoutFile d, HasRangeWithoutFile e, HasRangeWithoutFile f) => HasRangeWithoutFile (a,b,c,d,e,f) where
getRangeWithoutFile :: (a, b, c, d, e, f) -> RangeWithoutFile
getRangeWithoutFile (a
x,b
y,c
z,d
w,e
v,f
u) = (a, (b, (c, (d, (e, f))))) -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile (a
x,(b
y,(c
z,(d
w,(e
v,f
u)))))
instance (HasRangeWithoutFile a, HasRangeWithoutFile b, HasRangeWithoutFile c, HasRangeWithoutFile d, HasRangeWithoutFile e, HasRangeWithoutFile f, HasRangeWithoutFile g) => HasRangeWithoutFile (a,b,c,d,e,f,g) where
getRangeWithoutFile :: (a, b, c, d, e, f, g) -> RangeWithoutFile
getRangeWithoutFile (a
x,b
y,c
z,d
w,e
v,f
u,g
t) = (a, (b, (c, (d, (e, (f, g)))))) -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile (a
x,(b
y,(c
z,(d
w,(e
v,(f
u,g
t))))))
instance (HasRangeWithoutFile a, HasRangeWithoutFile b) => HasRangeWithoutFile (Either a b) where
getRangeWithoutFile :: Either a b -> RangeWithoutFile
getRangeWithoutFile = (a -> RangeWithoutFile)
-> (b -> RangeWithoutFile) -> Either a b -> RangeWithoutFile
forall a c b. (a -> c) -> (b -> c) -> Either a b -> c
either a -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile b -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile
class HasRange a => SetRange a where
setRange :: Range -> a -> a
default setRange :: (Functor f, SetRange b, f b ~ a) => Range -> a -> a
setRange = (b -> b) -> a -> a
(b -> b) -> f b -> f b
forall a b. (a -> b) -> f a -> f b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap ((b -> b) -> a -> a) -> (Range -> b -> b) -> Range -> a -> a
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range -> b -> b
forall a. SetRange a => Range -> a -> a
setRange
instance SetRange Range where
setRange :: Range -> Range -> Range
setRange = Range -> Range -> Range
forall a b. a -> b -> a
const
instance SetRange a => SetRange [a]
instance SetRange a => SetRange (Maybe a)
class KillRange a where
killRange :: KillRangeT a
default killRange :: (Functor f, KillRange b, f b ~ a) => KillRangeT a
killRange = (b -> b) -> f b -> f b
forall a b. (a -> b) -> f a -> f b
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
fmap b -> b
forall a. KillRange a => KillRangeT a
killRange
type KillRangeT a = a -> a
class KILLRANGE t b where
killRangeN :: IsBase t ~ b => All KillRange (Domains t) =>
t -> t
instance IsBase t ~ 'True => KILLRANGE t 'True where
{-# INLINE killRangeN #-}
killRangeN :: (IsBase t ~ 'True, All KillRange (Domains t)) => t -> t
killRangeN t
v = t
v
instance KILLRANGE t (IsBase t) => KILLRANGE (a -> t) 'False where
{-# INLINE killRangeN #-}
killRangeN :: (IsBase (a -> t) ~ 'False, All KillRange (Domains (a -> t))) =>
(a -> t) -> a -> t
killRangeN a -> t
f a
a = t -> t
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN (a -> t
f (KillRangeT a
forall a. KillRange a => KillRangeT a
killRange a
a))
killRangeMap :: (KillRange k, KillRange v) => KillRangeT (Map k v)
killRangeMap :: forall k v. (KillRange k, KillRange v) => KillRangeT (Map k v)
killRangeMap = (k -> k) -> Map k v -> Map k v
forall k1 k2 a. (k1 -> k2) -> Map k1 a -> Map k2 a
Map.mapKeysMonotonic k -> k
forall a. KillRange a => KillRangeT a
killRange (Map k v -> Map k v) -> (Map k v -> Map k v) -> Map k v -> Map k v
forall b c a. (b -> c) -> (a -> b) -> a -> c
. (v -> v) -> Map k v -> Map k v
forall a b k. (a -> b) -> Map k a -> Map k b
Map.map v -> v
forall a. KillRange a => KillRangeT a
killRange
instance KillRange Range where
killRange :: Range -> Range
killRange Range
_ = Range
forall a. Range' a
noRange
instance KillRange Void where
killRange :: KillRangeT Void
killRange = KillRangeT Void
forall a. a -> a
id
instance KillRange () where
killRange :: () -> ()
killRange = () -> ()
forall a. a -> a
id
instance KillRange Bool where
killRange :: Bool -> Bool
killRange = Bool -> Bool
forall a. a -> a
id
instance KillRange Int where
killRange :: KillRangeT Int
killRange = KillRangeT Int
forall a. a -> a
id
instance KillRange Integer where
killRange :: KillRangeT Integer
killRange = KillRangeT Integer
forall a. a -> a
id
instance KillRange Permutation where
killRange :: KillRangeT Permutation
killRange = KillRangeT Permutation
forall a. a -> a
id
instance {-# OVERLAPPING #-} KillRange String where
killRange :: ShowS
killRange = ShowS
forall a. a -> a
id
instance {-# OVERLAPPABLE #-} KillRange a => KillRange [a]
instance {-# OVERLAPPABLE #-} KillRange a => KillRange (Map k a)
instance KillRange a => KillRange (Drop a)
instance KillRange a => KillRange (List1 a)
instance KillRange a => KillRange (List2 a)
instance KillRange a => KillRange (Maybe a)
instance KillRange a => KillRange (Strict.Maybe a)
instance {-# OVERLAPPABLE #-} (Ord a, KillRange a) => KillRange (Set a) where
killRange :: KillRangeT (Set a)
killRange = (a -> a) -> KillRangeT (Set a)
forall b a. Ord b => (a -> b) -> Set a -> Set b
Set.map a -> a
forall a. KillRange a => KillRangeT a
killRange
instance (Ord a, KillRange a) => KillRange (Set1 a) where
killRange :: KillRangeT (Set1 a)
killRange = (a -> a) -> KillRangeT (Set1 a)
forall b a. Ord b => (a -> b) -> NESet a -> NESet b
Set1.map a -> a
forall a. KillRange a => KillRangeT a
killRange
instance (KillRange a, KillRange b) => KillRange (a, b) where
killRange :: KillRangeT (a, b)
killRange (a
x, b
y) = (KillRangeT a
forall a. KillRange a => KillRangeT a
killRange a
x, KillRangeT b
forall a. KillRange a => KillRangeT a
killRange b
y)
instance (KillRange a, KillRange b, KillRange c) =>
KillRange (a, b, c) where
killRange :: KillRangeT (a, b, c)
killRange (a
x, b
y, c
z) = (a -> b -> c -> (a, b, c)) -> a -> b -> c -> (a, b, c)
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN (,,) a
x b
y c
z
instance (KillRange a, KillRange b, KillRange c, KillRange d) =>
KillRange (a, b, c, d) where
killRange :: KillRangeT (a, b, c, d)
killRange (a
x, b
y, c
z, d
u) = (a -> b -> c -> d -> (a, b, c, d))
-> a -> b -> c -> d -> (a, b, c, d)
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN (,,,) a
x b
y c
z d
u
instance (KillRange a, KillRange b) => KillRange (Either a b) where
killRange :: KillRangeT (Either a b)
killRange (Left a
x) = a -> Either a b
forall a b. a -> Either a b
Left (a -> Either a b) -> a -> Either a b
forall a b. (a -> b) -> a -> b
$ KillRangeT a
forall a. KillRange a => KillRangeT a
killRange a
x
killRange (Right b
x) = b -> Either a b
forall a b. b -> Either a b
Right (b -> Either a b) -> b -> Either a b
forall a b. (a -> b) -> a -> b
$ KillRangeT b
forall a. KillRange a => KillRangeT a
killRange b
x
startPos' :: a -> Position' a
startPos' :: forall a. a -> Position' a
startPos' a
f = Pn
{ srcFile :: a
srcFile = a
f
, posPos :: Word32
posPos = Word32
1
, posLine :: Word32
posLine = Word32
1
, posCol :: Word32
posCol = Word32
1
}
startPos :: Maybe RangeFile -> Position
startPos :: Maybe RangeFile -> Position
startPos = SrcFile -> Position
forall a. a -> Position' a
startPos' (SrcFile -> Position)
-> (Maybe RangeFile -> SrcFile) -> Maybe RangeFile -> Position
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Maybe RangeFile -> SrcFile
forall lazy strict. Strict lazy strict => lazy -> strict
Strict.toStrict
rangeFromAbsolutePath :: AbsolutePath -> Range
rangeFromAbsolutePath :: AbsolutePath -> Range
rangeFromAbsolutePath AbsolutePath
f = SrcFile -> PositionWithoutFile -> PositionWithoutFile -> Range
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' SrcFile
src PositionWithoutFile
p0 PositionWithoutFile
p0
where
src :: SrcFile
src = RangeFile -> SrcFile
forall a. a -> Maybe a
Strict.Just (RangeFile -> SrcFile) -> RangeFile -> SrcFile
forall a b. (a -> b) -> a -> b
$ AbsolutePath -> Maybe (TopLevelModuleName' Range) -> RangeFile
mkRangeFile AbsolutePath
f Maybe (TopLevelModuleName' Range)
forall a. Maybe a
Nothing
p0 :: PositionWithoutFile
p0 = () -> PositionWithoutFile
forall a. a -> Position' a
startPos' ()
noRange :: Range' a
noRange :: forall a. Range' a
noRange = Range' a
forall a. Range' a
NoRange
movePos :: Position' a -> Char -> Position' a
movePos :: forall a. Position' a -> Char -> Position' a
movePos (Pn a
f Word32
p Word32
l Word32
c) Char
'\n' = a -> Word32 -> Word32 -> Word32 -> Position' a
forall a. a -> Word32 -> Word32 -> Word32 -> Position' a
Pn a
f (Word32
p Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
+ Word32
1) (Word32
l Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
+ Word32
1) Word32
1
movePos (Pn a
f Word32
p Word32
l Word32
c) Char
_ = a -> Word32 -> Word32 -> Word32 -> Position' a
forall a. a -> Word32 -> Word32 -> Word32 -> Position' a
Pn a
f (Word32
p Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
+ Word32
1) Word32
l (Word32
c Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
+ Word32
1)
movePosByString :: Foldable t => Position' a -> t Char -> Position' a
movePosByString :: forall (t :: * -> *) a.
Foldable t =>
Position' a -> t Char -> Position' a
movePosByString = (Position' a -> Char -> Position' a)
-> Position' a -> t Char -> Position' a
forall b a. (b -> a -> b) -> b -> t a -> b
forall (t :: * -> *) b a.
Foldable t =>
(b -> a -> b) -> b -> t a -> b
Fold.foldl' Position' a -> Char -> Position' a
forall a. Position' a -> Char -> Position' a
movePos
backupPos :: Position' a -> Position' a
backupPos :: forall a. Position' a -> Position' a
backupPos (Pn a
f Word32
p Word32
l Word32
c) = a -> Word32 -> Word32 -> Word32 -> Position' a
forall a. a -> Word32 -> Word32 -> Word32 -> Position' a
Pn a
f (Word32
p Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
- Word32
1) Word32
l (Word32
c Word32 -> Word32 -> Word32
forall a. Num a => a -> a -> a
- Word32
1)
posToRange' ::
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' :: forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' a
f PositionWithoutFile
p1 PositionWithoutFile
p2 = a -> IntervalWithoutFile -> Range' a
forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange a
f (()
-> PositionWithoutFile
-> PositionWithoutFile
-> IntervalWithoutFile
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
posToInterval () PositionWithoutFile
p1 PositionWithoutFile
p2)
posToRange :: Position' a -> Position' a -> Range' a
posToRange :: forall a. Position' a -> Position' a -> Range' a
posToRange Position' a
p1 Position' a
p2 =
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' (Position' a -> a
forall a. Position' a -> a
srcFile Position' a
p1) (Position' a
p1 { srcFile = () }) (Position' a
p2 { srcFile = () })
intervalToRange :: a -> IntervalWithoutFile -> Range' a
intervalToRange :: forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange a
f IntervalWithoutFile
i = a -> Seq IntervalWithoutFile -> Range' a
forall a. a -> Seq IntervalWithoutFile -> Range' a
Range a
f (IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. a -> Seq a
Seq.singleton IntervalWithoutFile
i)
rangeToIntervalWithFile :: Range' a -> Maybe (Interval' a)
rangeToIntervalWithFile :: forall a. Range' a -> Maybe (Interval' a)
rangeToIntervalWithFile Range' a
NoRange = Maybe (Interval' a)
forall a. Maybe a
Nothing
rangeToIntervalWithFile (Range a
f Seq IntervalWithoutFile
is) =
case (Seq IntervalWithoutFile -> ViewL IntervalWithoutFile
forall a. Seq a -> ViewL a
Seq.viewl Seq IntervalWithoutFile
is, Seq IntervalWithoutFile -> ViewR IntervalWithoutFile
forall a. Seq a -> ViewR a
Seq.viewr Seq IntervalWithoutFile
is) of
(IntervalWithoutFile
head Seq.:< Seq IntervalWithoutFile
_, Seq IntervalWithoutFile
_ Seq.:> IntervalWithoutFile
last) -> Interval' a -> Maybe (Interval' a)
forall a. a -> Maybe a
Just (Interval' a -> Maybe (Interval' a))
-> Interval' a -> Maybe (Interval' a)
forall a b. (a -> b) -> a -> b
$ a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval a
f (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
head) (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
last)
(ViewL IntervalWithoutFile, ViewR IntervalWithoutFile)
_ -> Maybe (Interval' a)
forall a. HasCallStack => a
__IMPOSSIBLE__
rangeToInterval :: Range' a -> Maybe IntervalWithoutFile
rangeToInterval :: forall a. Range' a -> Maybe IntervalWithoutFile
rangeToInterval = RangeWithoutFile -> Maybe IntervalWithoutFile
forall a. Range' a -> Maybe (Interval' a)
rangeToIntervalWithFile (RangeWithoutFile -> Maybe IntervalWithoutFile)
-> (Range' a -> RangeWithoutFile)
-> Range' a
-> Maybe IntervalWithoutFile
forall b c a. (b -> c) -> (a -> b) -> a -> c
. Range' a -> RangeWithoutFile
forall (f :: * -> *) a. Functor f => f a -> f ()
void
continuous :: Range' a -> Range' a
continuous :: forall a. Range' a -> Range' a
continuous Range' a
NoRange = Range' a
forall a. Range' a
NoRange
continuous r :: Range' a
r@(Range a
f Seq IntervalWithoutFile
_) =
Range' a
-> (IntervalWithoutFile -> Range' a)
-> Maybe IntervalWithoutFile
-> Range' a
forall b a. b -> (a -> b) -> Maybe a -> b
maybe Range' a
forall a. HasCallStack => a
__IMPOSSIBLE__ (a -> IntervalWithoutFile -> Range' a
forall a. a -> IntervalWithoutFile -> Range' a
intervalToRange a
f) (Maybe IntervalWithoutFile -> Range' a)
-> Maybe IntervalWithoutFile -> Range' a
forall a b. (a -> b) -> a -> b
$ Range' a -> Maybe IntervalWithoutFile
forall a. Range' a -> Maybe IntervalWithoutFile
rangeToInterval Range' a
r
continuousPerLine :: Ord a => Range' a -> Range' a
continuousPerLine :: forall a. Ord a => Range' a -> Range' a
continuousPerLine r :: Range' a
r@Range' a
NoRange = Range' a
r
continuousPerLine r :: Range' a
r@(Range a
f Seq IntervalWithoutFile
_) =
a -> Seq IntervalWithoutFile -> Range' a
forall a. a -> Seq IntervalWithoutFile -> Range' a
Range a
f (([IntervalWithoutFile]
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile]))
-> [IntervalWithoutFile] -> Seq IntervalWithoutFile
forall b a. (b -> Maybe (a, b)) -> b -> Seq a
Seq.unfoldr [IntervalWithoutFile]
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile])
step (Range' a -> [IntervalWithoutFile]
forall a. Range' a -> [IntervalWithoutFile]
rangeIntervals Range' a
r))
where
step :: [IntervalWithoutFile]
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile])
step [] = Maybe (IntervalWithoutFile, [IntervalWithoutFile])
forall a. Maybe a
Nothing
step [IntervalWithoutFile
i] = (IntervalWithoutFile, [IntervalWithoutFile])
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile])
forall a. a -> Maybe a
Just (IntervalWithoutFile
i, [])
step (IntervalWithoutFile
i : is :: [IntervalWithoutFile]
is@(IntervalWithoutFile
j : [IntervalWithoutFile]
js))
| Bool
sameLine = [IntervalWithoutFile]
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile])
step (IntervalWithoutFile -> IntervalWithoutFile -> IntervalWithoutFile
fuseIntervals IntervalWithoutFile
i IntervalWithoutFile
j IntervalWithoutFile
-> [IntervalWithoutFile] -> [IntervalWithoutFile]
forall a. a -> [a] -> [a]
: [IntervalWithoutFile]
js)
| Bool
otherwise = (IntervalWithoutFile, [IntervalWithoutFile])
-> Maybe (IntervalWithoutFile, [IntervalWithoutFile])
forall a. a -> Maybe a
Just (IntervalWithoutFile
i, [IntervalWithoutFile]
is)
where
sameLine :: Bool
sameLine = PositionWithoutFile -> Word32
forall a. Position' a -> Word32
posLine (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
i) Word32 -> Word32 -> Bool
forall a. Eq a => a -> a -> Bool
== PositionWithoutFile -> Word32
forall a. Position' a -> Word32
posLine (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
j)
rStart' :: Range' a -> Maybe PositionWithoutFile
rStart' :: forall a. Range' a -> Maybe PositionWithoutFile
rStart' Range' a
r = IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart (IntervalWithoutFile -> PositionWithoutFile)
-> Maybe IntervalWithoutFile -> Maybe PositionWithoutFile
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Range' a -> Maybe IntervalWithoutFile
forall a. Range' a -> Maybe IntervalWithoutFile
rangeToInterval Range' a
r
rStart :: Range' a -> Maybe (Position' a)
rStart :: forall a. Range' a -> Maybe (Position' a)
rStart Range' a
NoRange = Maybe (Position' a)
forall a. Maybe a
Nothing
rStart r :: Range' a
r@(Range a
f Seq IntervalWithoutFile
_) = (\PositionWithoutFile
p -> PositionWithoutFile
p { srcFile = f }) (PositionWithoutFile -> Position' a)
-> Maybe PositionWithoutFile -> Maybe (Position' a)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Range' a -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe PositionWithoutFile
rStart' Range' a
r
rEnd' :: Range' a -> Maybe PositionWithoutFile
rEnd' :: forall a. Range' a -> Maybe PositionWithoutFile
rEnd' Range' a
r = IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd (IntervalWithoutFile -> PositionWithoutFile)
-> Maybe IntervalWithoutFile -> Maybe PositionWithoutFile
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Range' a -> Maybe IntervalWithoutFile
forall a. Range' a -> Maybe IntervalWithoutFile
rangeToInterval Range' a
r
rEnd :: Range' a -> Maybe (Position' a)
rEnd :: forall a. Range' a -> Maybe (Position' a)
rEnd Range' a
NoRange = Maybe (Position' a)
forall a. Maybe a
Nothing
rEnd r :: Range' a
r@(Range a
f Seq IntervalWithoutFile
_) = (\PositionWithoutFile
p -> PositionWithoutFile
p { srcFile = f }) (PositionWithoutFile -> Position' a)
-> Maybe PositionWithoutFile -> Maybe (Position' a)
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> Range' a -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe PositionWithoutFile
rEnd' Range' a
r
fuseIntervals :: IntervalWithoutFile -> IntervalWithoutFile -> IntervalWithoutFile
fuseIntervals :: IntervalWithoutFile -> IntervalWithoutFile -> IntervalWithoutFile
fuseIntervals (Interval () PositionWithoutFile
s1 PositionWithoutFile
e1) (Interval () PositionWithoutFile
s2 PositionWithoutFile
e2) = ()
-> PositionWithoutFile
-> PositionWithoutFile
-> IntervalWithoutFile
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval () (PositionWithoutFile -> PositionWithoutFile -> PositionWithoutFile
forall a. Ord a => a -> a -> a
min PositionWithoutFile
s1 PositionWithoutFile
s2) (PositionWithoutFile -> PositionWithoutFile -> PositionWithoutFile
forall a. Ord a => a -> a -> a
max PositionWithoutFile
e1 PositionWithoutFile
e2)
fuseRanges :: (Ord a) => Range' a -> Range' a -> Range' a
fuseRanges :: forall a. Ord a => Range' a -> Range' a -> Range' a
fuseRanges Range' a
NoRange Range' a
is2 = Range' a
is2
fuseRanges Range' a
is1 Range' a
NoRange = Range' a
is1
fuseRanges (Range a
f Seq IntervalWithoutFile
is1) (Range a
_ Seq IntervalWithoutFile
is2) = a -> Seq IntervalWithoutFile -> Range' a
forall a. a -> Seq IntervalWithoutFile -> Range' a
Range a
f (Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
fuse Seq IntervalWithoutFile
is1 Seq IntervalWithoutFile
is2)
where
fuse :: Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
fuse Seq IntervalWithoutFile
is1 Seq IntervalWithoutFile
is2 = case (Seq IntervalWithoutFile -> ViewL IntervalWithoutFile
forall a. Seq a -> ViewL a
Seq.viewl Seq IntervalWithoutFile
is1, Seq IntervalWithoutFile -> ViewR IntervalWithoutFile
forall a. Seq a -> ViewR a
Seq.viewr Seq IntervalWithoutFile
is1,
Seq IntervalWithoutFile -> ViewL IntervalWithoutFile
forall a. Seq a -> ViewL a
Seq.viewl Seq IntervalWithoutFile
is2, Seq IntervalWithoutFile -> ViewR IntervalWithoutFile
forall a. Seq a -> ViewR a
Seq.viewr Seq IntervalWithoutFile
is2) of
(ViewL IntervalWithoutFile
Seq.EmptyL, ViewR IntervalWithoutFile
_, ViewL IntervalWithoutFile
_, ViewR IntervalWithoutFile
_) -> Seq IntervalWithoutFile
is2
(ViewL IntervalWithoutFile
_, ViewR IntervalWithoutFile
_, ViewL IntervalWithoutFile
Seq.EmptyL, ViewR IntervalWithoutFile
_) -> Seq IntervalWithoutFile
is1
(IntervalWithoutFile
s1 Seq.:< Seq IntervalWithoutFile
r1, Seq IntervalWithoutFile
l1 Seq.:> IntervalWithoutFile
e1, IntervalWithoutFile
s2 Seq.:< Seq IntervalWithoutFile
r2, Seq IntervalWithoutFile
l2 Seq.:> IntervalWithoutFile
e2)
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
e1 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s2 -> Seq IntervalWithoutFile
is1 Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. Seq a -> Seq a -> Seq a
Seq.>< Seq IntervalWithoutFile
is2
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
e2 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s1 -> Seq IntervalWithoutFile
is2 Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. Seq a -> Seq a -> Seq a
Seq.>< Seq IntervalWithoutFile
is1
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
e1 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Eq a => a -> a -> Bool
== IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s2 -> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
mergeTouching Seq IntervalWithoutFile
l1 IntervalWithoutFile
e1 IntervalWithoutFile
s2 Seq IntervalWithoutFile
r2
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
e2 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Eq a => a -> a -> Bool
== IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s1 -> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
mergeTouching Seq IntervalWithoutFile
l2 IntervalWithoutFile
e2 IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s1 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s2 -> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
outputLeftPrefix IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1 IntervalWithoutFile
s2 Seq IntervalWithoutFile
is2
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s2 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s1 -> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
outputLeftPrefix IntervalWithoutFile
s2 Seq IntervalWithoutFile
r2 IntervalWithoutFile
s1 Seq IntervalWithoutFile
is1
| IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s1 PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s2 -> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
fuseSome IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1 IntervalWithoutFile
s2 Seq IntervalWithoutFile
r2
| Bool
otherwise -> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
fuseSome IntervalWithoutFile
s2 Seq IntervalWithoutFile
r2 IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1
(ViewL IntervalWithoutFile, ViewR IntervalWithoutFile,
ViewL IntervalWithoutFile, ViewR IntervalWithoutFile)
_ -> Seq IntervalWithoutFile
forall a. HasCallStack => a
__IMPOSSIBLE__
mergeTouching :: Seq IntervalWithoutFile
-> IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
mergeTouching Seq IntervalWithoutFile
l IntervalWithoutFile
e IntervalWithoutFile
s Seq IntervalWithoutFile
r = Seq IntervalWithoutFile
l Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. Seq a -> Seq a -> Seq a
Seq.>< IntervalWithoutFile
i IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. a -> Seq a -> Seq a
Seq.<| Seq IntervalWithoutFile
r
where
i :: IntervalWithoutFile
i = ()
-> PositionWithoutFile
-> PositionWithoutFile
-> IntervalWithoutFile
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Interval' a
Interval () (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
e) (IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s)
outputLeftPrefix :: IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
outputLeftPrefix IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1 IntervalWithoutFile
s2 Seq IntervalWithoutFile
is2 = IntervalWithoutFile
s1 IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. a -> Seq a -> Seq a
Seq.<| Seq IntervalWithoutFile
r1' Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. Seq a -> Seq a -> Seq a
Seq.>< Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
fuse Seq IntervalWithoutFile
r1'' Seq IntervalWithoutFile
is2
where
(Seq IntervalWithoutFile
r1', Seq IntervalWithoutFile
r1'') = (IntervalWithoutFile -> Bool)
-> Seq IntervalWithoutFile
-> (Seq IntervalWithoutFile, Seq IntervalWithoutFile)
forall a. (a -> Bool) -> Seq a -> (Seq a, Seq a)
Seq.spanl (\IntervalWithoutFile
s -> IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iStart IntervalWithoutFile
s2) Seq IntervalWithoutFile
r1
fuseSome :: IntervalWithoutFile
-> Seq IntervalWithoutFile
-> IntervalWithoutFile
-> Seq IntervalWithoutFile
-> Seq IntervalWithoutFile
fuseSome IntervalWithoutFile
s1 Seq IntervalWithoutFile
r1 IntervalWithoutFile
s2 Seq IntervalWithoutFile
r2 = Seq IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
fuse Seq IntervalWithoutFile
r1' (IntervalWithoutFile -> IntervalWithoutFile -> IntervalWithoutFile
fuseIntervals IntervalWithoutFile
s1 IntervalWithoutFile
s2 IntervalWithoutFile
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. a -> Seq a -> Seq a
Seq.<| Seq IntervalWithoutFile
r2)
where
r1' :: Seq IntervalWithoutFile
r1' = (IntervalWithoutFile -> Bool)
-> Seq IntervalWithoutFile -> Seq IntervalWithoutFile
forall a. (a -> Bool) -> Seq a -> Seq a
Seq.dropWhileL (\IntervalWithoutFile
s -> IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s PositionWithoutFile -> PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
<= IntervalWithoutFile -> PositionWithoutFile
forall a. Interval' a -> Position' a
iEnd IntervalWithoutFile
s2) Seq IntervalWithoutFile
r1
{-# INLINE fuseRange #-}
fuseRange :: (HasRange u, HasRange t) => u -> t -> Range
fuseRange :: forall u t. (HasRange u, HasRange t) => u -> t -> Range
fuseRange u
x t
y = Range -> Range -> Range
forall a. Ord a => Range' a -> Range' a -> Range' a
fuseRanges (u -> Range
forall a. HasRange a => a -> Range
getRange u
x) (t -> Range
forall a. HasRange a => a -> Range
getRange t
y)
{-# INLINE fuseRangeWithoutFile #-}
fuseRangeWithoutFile :: (HasRangeWithoutFile u, HasRangeWithoutFile t) => u -> t -> RangeWithoutFile
fuseRangeWithoutFile :: forall u t.
(HasRangeWithoutFile u, HasRangeWithoutFile t) =>
u -> t -> RangeWithoutFile
fuseRangeWithoutFile u
x t
y = RangeWithoutFile -> RangeWithoutFile -> RangeWithoutFile
forall a. Ord a => Range' a -> Range' a -> Range' a
fuseRanges (u -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile u
x) (t -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile t
y)
beginningOf :: Range -> Range
beginningOf :: Range -> Range
beginningOf Range
NoRange = Range
forall a. Range' a
NoRange
beginningOf r :: Range
r@(Range SrcFile
f Seq IntervalWithoutFile
_) = case Range -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe PositionWithoutFile
rStart' Range
r of
Maybe PositionWithoutFile
Nothing -> Range
forall a. HasCallStack => a
__IMPOSSIBLE__
Just PositionWithoutFile
pos -> SrcFile -> PositionWithoutFile -> PositionWithoutFile -> Range
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' SrcFile
f PositionWithoutFile
pos PositionWithoutFile
pos
beginningOfFile :: Range -> Range
beginningOfFile :: Range -> Range
beginningOfFile Range
NoRange = Range
forall a. Range' a
NoRange
beginningOfFile (Range SrcFile
f Seq IntervalWithoutFile
_) = SrcFile -> PositionWithoutFile -> PositionWithoutFile -> Range
forall a.
a -> PositionWithoutFile -> PositionWithoutFile -> Range' a
posToRange' SrcFile
f PositionWithoutFile
p PositionWithoutFile
p
where p :: PositionWithoutFile
p = () -> PositionWithoutFile
forall a. a -> Position' a
startPos' ()
withRangeOf :: (SetRange t, HasRange u) => t -> u -> t
t
x withRangeOf :: forall t u. (SetRange t, HasRange u) => t -> u -> t
`withRangeOf` u
y = Range -> t -> t
forall a. SetRange a => Range -> a -> a
setRange (u -> Range
forall a. HasRange a => a -> Range
getRange u
y) t
x
interleaveRanges :: forall a. (HasRangeWithoutFile a) => [a] -> [a] -> ([a], [(a,a)])
interleaveRanges :: forall a. HasRangeWithoutFile a => [a] -> [a] -> ([a], [(a, a)])
interleaveRanges [a]
as [a]
bs = Writer [(a, a)] [a] -> ([a], [(a, a)])
forall w a. Writer w a -> (a, w)
runWriter (Writer [(a, a)] [a] -> ([a], [(a, a)]))
-> Writer [(a, a)] [a] -> ([a], [(a, a)])
forall a b. (a -> b) -> a -> b
$ [a] -> [a] -> Writer [(a, a)] [a]
forall {m :: * -> *} {b}.
(HasRangeWithoutFile b, MonadWriter [(b, b)] m) =>
[b] -> [b] -> m [b]
go [a]
as [a]
bs
where
go :: [b] -> [b] -> m [b]
go [] [b]
as = [b] -> m [b]
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return [b]
as
go [b]
as [] = [b] -> m [b]
forall a. a -> m a
forall (m :: * -> *) a. Monad m => a -> m a
return [b]
as
go as :: [b]
as@(b
a:[b]
as') bs :: [b]
bs@(b
b:[b]
bs') =
let ra :: RangeWithoutFile
ra = b -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile b
a
rb :: RangeWithoutFile
rb = b -> RangeWithoutFile
forall a. HasRangeWithoutFile a => a -> RangeWithoutFile
getRangeWithoutFile b
b
ra0 :: Maybe PositionWithoutFile
ra0 = RangeWithoutFile -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe (Position' a)
rStart RangeWithoutFile
ra
rb0 :: Maybe PositionWithoutFile
rb0 = RangeWithoutFile -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe (Position' a)
rStart RangeWithoutFile
rb
ra1 :: Maybe PositionWithoutFile
ra1 = RangeWithoutFile -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe (Position' a)
rEnd RangeWithoutFile
ra
rb1 :: Maybe PositionWithoutFile
rb1 = RangeWithoutFile -> Maybe PositionWithoutFile
forall a. Range' a -> Maybe (Position' a)
rEnd RangeWithoutFile
rb
in
if Maybe PositionWithoutFile
ra1 Maybe PositionWithoutFile -> Maybe PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
<= Maybe PositionWithoutFile
rb0 then
(b
ab -> [b] -> [b]
forall a. a -> [a] -> [a]
:) ([b] -> [b]) -> m [b] -> m [b]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [b] -> [b] -> m [b]
go [b]
as' [b]
bs
else if Maybe PositionWithoutFile
rb1 Maybe PositionWithoutFile -> Maybe PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
<= Maybe PositionWithoutFile
ra0 then
(b
bb -> [b] -> [b]
forall a. a -> [a] -> [a]
:) ([b] -> [b]) -> m [b] -> m [b]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [b] -> [b] -> m [b]
go [b]
as [b]
bs'
else do
[(b, b)] -> m ()
forall w (m :: * -> *). MonadWriter w m => w -> m ()
tell [(b
a,b
b)]
if Maybe PositionWithoutFile
ra0 Maybe PositionWithoutFile -> Maybe PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
< Maybe PositionWithoutFile
rb0 Bool -> Bool -> Bool
|| (Maybe PositionWithoutFile
ra0 Maybe PositionWithoutFile -> Maybe PositionWithoutFile -> Bool
forall a. Eq a => a -> a -> Bool
== Maybe PositionWithoutFile
rb0 Bool -> Bool -> Bool
&& Maybe PositionWithoutFile
ra1 Maybe PositionWithoutFile -> Maybe PositionWithoutFile -> Bool
forall a. Ord a => a -> a -> Bool
<= Maybe PositionWithoutFile
rb1) then
(b
ab -> [b] -> [b]
forall a. a -> [a] -> [a]
:) ([b] -> [b]) -> m [b] -> m [b]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [b] -> [b] -> m [b]
go [b]
as' [b]
bs
else
(b
bb -> [b] -> [b]
forall a. a -> [a] -> [a]
:) ([b] -> [b]) -> m [b] -> m [b]
forall (f :: * -> *) a b. Functor f => (a -> b) -> f a -> f b
<$> [b] -> [b] -> m [b]
go [b]
as [b]
bs'