module Agda.TypeChecking.DiscrimTree.Types where
import Control.DeepSeq
import qualified Data.Map.Strict as Map
import qualified Data.Set as Set
import Data.Map.Strict (Map)
import Data.Set (Set)
import GHC.Generics (Generic)
import Agda.Syntax.Internal
import Agda.Syntax.Position
import Agda.Utils.Impossible
import Agda.Utils.Null
data Key
= RigidK {-# UNPACK #-} !QName {-# UNPACK #-} !Int
| LocalK {-# UNPACK #-} !Int {-# UNPACK #-} !Int
| PiK
| ConstK
| SortK
| FlexK
deriving (Int -> Key -> ShowS
[Key] -> ShowS
Key -> String
(Int -> Key -> ShowS)
-> (Key -> String) -> ([Key] -> ShowS) -> Show Key
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: Int -> Key -> ShowS
showsPrec :: Int -> Key -> ShowS
$cshow :: Key -> String
show :: Key -> String
$cshowList :: [Key] -> ShowS
showList :: [Key] -> ShowS
Show, Key -> Key -> Bool
(Key -> Key -> Bool) -> (Key -> Key -> Bool) -> Eq Key
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: Key -> Key -> Bool
== :: Key -> Key -> Bool
$c/= :: Key -> Key -> Bool
/= :: Key -> Key -> Bool
Eq, Eq Key
Eq Key =>
(Key -> Key -> Ordering)
-> (Key -> Key -> Bool)
-> (Key -> Key -> Bool)
-> (Key -> Key -> Bool)
-> (Key -> Key -> Bool)
-> (Key -> Key -> Key)
-> (Key -> Key -> Key)
-> Ord Key
Key -> Key -> Bool
Key -> Key -> Ordering
Key -> Key -> Key
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
$ccompare :: Key -> Key -> Ordering
compare :: Key -> Key -> Ordering
$c< :: Key -> Key -> Bool
< :: Key -> Key -> Bool
$c<= :: Key -> Key -> Bool
<= :: Key -> Key -> Bool
$c> :: Key -> Key -> Bool
> :: Key -> Key -> Bool
$c>= :: Key -> Key -> Bool
>= :: Key -> Key -> Bool
$cmax :: Key -> Key -> Key
max :: Key -> Key -> Key
$cmin :: Key -> Key -> Key
min :: Key -> Key -> Key
Ord, (forall x. Key -> Rep Key x)
-> (forall x. Rep Key x -> Key) -> Generic Key
forall x. Rep Key x -> Key
forall x. Key -> Rep Key x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
$cfrom :: forall x. Key -> Rep Key x
from :: forall x. Key -> Rep Key x
$cto :: forall x. Rep Key x -> Key
to :: forall x. Rep Key x -> Key
Generic)
instance NFData Key
data DiscrimTree a
= EmptyDT
| DoneDT (Set a)
| CaseDT
{-# UNPACK #-} !Int
(Map Key (DiscrimTree a))
(DiscrimTree a)
deriving ((forall x. DiscrimTree a -> Rep (DiscrimTree a) x)
-> (forall x. Rep (DiscrimTree a) x -> DiscrimTree a)
-> Generic (DiscrimTree a)
forall x. Rep (DiscrimTree a) x -> DiscrimTree a
forall x. DiscrimTree a -> Rep (DiscrimTree a) x
forall a.
(forall x. a -> Rep a x) -> (forall x. Rep a x -> a) -> Generic a
forall a x. Rep (DiscrimTree a) x -> DiscrimTree a
forall a x. DiscrimTree a -> Rep (DiscrimTree a) x
$cfrom :: forall a x. DiscrimTree a -> Rep (DiscrimTree a) x
from :: forall x. DiscrimTree a -> Rep (DiscrimTree a) x
$cto :: forall a x. Rep (DiscrimTree a) x -> DiscrimTree a
to :: forall x. Rep (DiscrimTree a) x -> DiscrimTree a
Generic, DiscrimTree a -> DiscrimTree a -> Bool
(DiscrimTree a -> DiscrimTree a -> Bool)
-> (DiscrimTree a -> DiscrimTree a -> Bool) -> Eq (DiscrimTree a)
forall a. Eq a => DiscrimTree a -> DiscrimTree a -> Bool
forall a. (a -> a -> Bool) -> (a -> a -> Bool) -> Eq a
$c== :: forall a. Eq a => DiscrimTree a -> DiscrimTree a -> Bool
== :: DiscrimTree a -> DiscrimTree a -> Bool
$c/= :: forall a. Eq a => DiscrimTree a -> DiscrimTree a -> Bool
/= :: DiscrimTree a -> DiscrimTree a -> Bool
Eq, Int -> DiscrimTree a -> ShowS
[DiscrimTree a] -> ShowS
DiscrimTree a -> String
(Int -> DiscrimTree a -> ShowS)
-> (DiscrimTree a -> String)
-> ([DiscrimTree a] -> ShowS)
-> Show (DiscrimTree a)
forall a. Show a => Int -> DiscrimTree a -> ShowS
forall a. Show a => [DiscrimTree a] -> ShowS
forall a. Show a => DiscrimTree a -> String
forall a.
(Int -> a -> ShowS) -> (a -> String) -> ([a] -> ShowS) -> Show a
$cshowsPrec :: forall a. Show a => Int -> DiscrimTree a -> ShowS
showsPrec :: Int -> DiscrimTree a -> ShowS
$cshow :: forall a. Show a => DiscrimTree a -> String
show :: DiscrimTree a -> String
$cshowList :: forall a. Show a => [DiscrimTree a] -> ShowS
showList :: [DiscrimTree a] -> ShowS
Show)
instance NFData a => NFData (DiscrimTree a)
instance (KillRange a, Ord a) => KillRange (DiscrimTree a) where
killRange :: KillRangeT (DiscrimTree a)
killRange = \case
DiscrimTree a
EmptyDT -> DiscrimTree a
forall a. DiscrimTree a
EmptyDT
DoneDT Set a
s -> (Set a -> DiscrimTree a) -> Set a -> DiscrimTree a
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT Set a
s
CaseDT Int
i Map Key (DiscrimTree a)
k DiscrimTree a
o -> (Int -> Map Key (DiscrimTree a) -> KillRangeT (DiscrimTree a))
-> Int -> Map Key (DiscrimTree a) -> KillRangeT (DiscrimTree a)
forall t (b :: Bool).
(KILLRANGE t b, IsBase t ~ b, All KillRange (Domains t)) =>
t -> t
killRangeN Int -> Map Key (DiscrimTree a) -> KillRangeT (DiscrimTree a)
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
k DiscrimTree a
o
instance Null (DiscrimTree a) where
empty :: DiscrimTree a
empty = DiscrimTree a
forall a. DiscrimTree a
EmptyDT
null :: DiscrimTree a -> Bool
null = \case
DiscrimTree a
EmptyDT -> Bool
True
DiscrimTree a
_ -> Bool
False
mergeDT :: Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT :: forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT DiscrimTree a
EmptyDT DiscrimTree a
x = DiscrimTree a
x
mergeDT (DoneDT Set a
s) DiscrimTree a
x = case DiscrimTree a
x of
DiscrimTree a
EmptyDT -> Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT Set a
s
DoneDT Set a
s' -> Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT (Set a
s Set a -> Set a -> Set a
forall a. Semigroup a => a -> a -> a
<> Set a
s')
CaseDT Int
i Map Key (DiscrimTree a)
bs DiscrimTree a
x -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
bs (DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT (Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT Set a
s) DiscrimTree a
x)
mergeDT (CaseDT Int
i Map Key (DiscrimTree a)
bs DiscrimTree a
els) DiscrimTree a
x = case DiscrimTree a
x of
DiscrimTree a
EmptyDT -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
bs DiscrimTree a
els
DoneDT Set a
s -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
bs (DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT (Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT Set a
s) DiscrimTree a
els)
CaseDT Int
j Map Key (DiscrimTree a)
bs' DiscrimTree a
els' -> case Int -> Int -> Ordering
forall a. Ord a => a -> a -> Ordering
compare Int
i Int
j of
Ordering
EQ -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
j ((DiscrimTree a -> DiscrimTree a -> DiscrimTree a)
-> Map Key (DiscrimTree a)
-> Map Key (DiscrimTree a)
-> Map Key (DiscrimTree a)
forall k a. Ord k => (a -> a -> a) -> Map k a -> Map k a -> Map k a
Map.unionWith DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT Map Key (DiscrimTree a)
bs Map Key (DiscrimTree a)
bs') (DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT DiscrimTree a
els DiscrimTree a
els')
Ordering
LT -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
bs (DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT DiscrimTree a
els (Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
j Map Key (DiscrimTree a)
bs' DiscrimTree a
els'))
Ordering
GT -> Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
j Map Key (DiscrimTree a)
bs' (DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT DiscrimTree a
els' (Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
i Map Key (DiscrimTree a)
bs DiscrimTree a
els))
instance Ord a => Semigroup (DiscrimTree a) where
<> :: DiscrimTree a -> DiscrimTree a -> DiscrimTree a
(<>) = DiscrimTree a -> DiscrimTree a -> DiscrimTree a
forall a. Ord a => DiscrimTree a -> DiscrimTree a -> DiscrimTree a
mergeDT
instance Ord a => Monoid (DiscrimTree a) where
mempty :: DiscrimTree a
mempty = DiscrimTree a
forall a. DiscrimTree a
EmptyDT
singletonDT :: [Key] -> a -> DiscrimTree a
singletonDT :: forall a. [Key] -> a -> DiscrimTree a
singletonDT [Key]
key a
val = Int -> [Key] -> DiscrimTree a
go Int
0 [Key]
key where
go :: Int -> [Key] -> DiscrimTree a
go Int
focus [] = Set a -> DiscrimTree a
forall a. Set a -> DiscrimTree a
DoneDT (Set a -> DiscrimTree a) -> Set a -> DiscrimTree a
forall a b. (a -> b) -> a -> b
$ a -> Set a
forall a. a -> Set a
Set.singleton a
val
go Int
focus (Key
FlexK:[Key]
ts) = Int -> [Key] -> DiscrimTree a
go (Int
focus Int -> Int -> Int
forall a. Num a => a -> a -> a
+ Int
1) [Key]
ts
go Int
focus (Key
t:[Key]
ts) = Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
forall a.
Int -> Map Key (DiscrimTree a) -> DiscrimTree a -> DiscrimTree a
CaseDT Int
focus (Key -> DiscrimTree a -> Map Key (DiscrimTree a)
forall k a. k -> a -> Map k a
Map.singleton Key
t (Int -> [Key] -> DiscrimTree a
go Int
focus [Key]
ts)) DiscrimTree a
forall a. DiscrimTree a
EmptyDT