Safe Haskell | None |
---|---|
Language | Haskell2010 |
Agda.TypeChecking.SizedTypes.Syntax
Description
Syntax of size expressions and constraints.
Synopsis
- newtype Offset = O Int
- newtype Rigid = RigidId {}
- newtype Flex = FlexId {}
- data SizeExpr' rigid flex
- type SizeExpr = SizeExpr' Rigid Flex
- data Cmp
- data Constraint' rigid flex = Constraint {}
- type Constraint = Constraint' Rigid Flex
- data Polarity
- data PolarityAssignment flex = PolarityAssignment Polarity flex
- type Polarities flex = Map flex Polarity
- emptyPolarities :: Polarities flex
- polaritiesFromAssignments :: Ord flex => [PolarityAssignment flex] -> Polarities flex
- getPolarity :: Ord flex => Polarities flex -> flex -> Polarity
- newtype Solution rigid flex = Solution {
- theSolution :: Map flex (SizeExpr' rigid flex)
- emptySolution :: Solution r f
- class Substitute r f a where
- type CTrans r f = Constraint' r f -> Maybe [Constraint' r f]
- simplify1 :: (Pretty f, Pretty r, Eq r) => CTrans r f -> CTrans r f
- ifLe :: Cmp -> a -> a -> a
- compareOffset :: Offset -> Cmp -> Offset -> Bool
- class ValidOffset a where
- validOffset :: a -> Bool
- class TruncateOffset a where
- truncateOffset :: a -> a
- class Ord (RigidOf a) => Rigids a where
- class Ord (FlexOf a) => Flexs a where
- data NamedRigid = NamedRigid {
- rigidName :: String
- rigidIndex :: Int
- data SizeMeta = SizeMeta {
- sizeMetaId :: MetaId
- sizeMetaArgs :: [Int]
- type DBSizeExpr = SizeExpr' NamedRigid SizeMeta
- type SizeConstraint = Constraint' NamedRigid SizeMeta
- data HypSizeConstraint = HypSizeConstraint {}
Syntax
Constant finite sizes n >= 0
.
Instances
Pretty Offset Source # | |||||
TruncateOffset Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods truncateOffset :: Offset -> Offset Source # | |||||
ValidOffset Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods validOffset :: Offset -> Bool Source # | |||||
MeetSemiLattice Offset Source # | |||||
Negative Offset Source # | |||||
NFData Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
Enum Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
Generic Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
| |||||
Num Offset Source # | |||||
Show Offset Source # | |||||
Eq Offset Source # | |||||
Ord Offset Source # | |||||
Plus Offset Offset Offset Source # | |||||
Plus Offset Weight Weight Source # | |||||
Plus Weight Offset Weight Source # | |||||
Plus (SizeExpr' r f) Offset (SizeExpr' r f) Source # | Add offset to size expression. | ||||
type Rep Offset Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax |
Fixed size variables i
.
Size meta variables X
to solve for.
data SizeExpr' rigid flex Source #
Size expressions appearing in constraints.
Constructors
Const | Constant number |
Rigid | Variable plus offset |
Infty | Infinity |
Flex | Meta variable |
Instances
Comparison operator, e.g. for size expression.
Instances
Pretty Cmp Source # | |
Dioid Cmp Source # | |
MeetSemiLattice Cmp Source # | |
Top Cmp Source # | |
NFData Cmp Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax | |
Bounded Cmp Source # | |
Enum Cmp Source # | |
Generic Cmp Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax | |
Show Cmp Source # | |
Eq Cmp Source # | |
Ord Cmp Source # | |
type Rep Cmp Source # | |
data Constraint' rigid flex Source #
Constraint: an inequation between size expressions,
e.g. X < ∞
or i + 3 ≤ j
.
Constructors
Constraint | |
Instances
PrettyTCM SizeConstraint Source # | Assumes we are in the right context. | ||||
Defined in Agda.TypeChecking.SizedTypes.Pretty Methods prettyTCM :: MonadPretty m => SizeConstraint -> m Doc Source # | |||||
Subst SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg SizeConstraint) -> SizeConstraint -> SizeConstraint Source # | |||||
Ord f => Substitute r f (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods subst :: Solution r f -> Constraint' r f -> Constraint' r f Source # | |||||
Functor (Constraint' rigid) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods fmap :: (a -> b) -> Constraint' rigid a -> Constraint' rigid b # (<$) :: a -> Constraint' rigid b -> Constraint' rigid a # | |||||
Foldable (Constraint' rigid) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods fold :: Monoid m => Constraint' rigid m -> m # foldMap :: Monoid m => (a -> m) -> Constraint' rigid a -> m # foldMap' :: Monoid m => (a -> m) -> Constraint' rigid a -> m # foldr :: (a -> b -> b) -> b -> Constraint' rigid a -> b # foldr' :: (a -> b -> b) -> b -> Constraint' rigid a -> b # foldl :: (b -> a -> b) -> b -> Constraint' rigid a -> b # foldl' :: (b -> a -> b) -> b -> Constraint' rigid a -> b # foldr1 :: (a -> a -> a) -> Constraint' rigid a -> a # foldl1 :: (a -> a -> a) -> Constraint' rigid a -> a # toList :: Constraint' rigid a -> [a] # null :: Constraint' rigid a -> Bool # length :: Constraint' rigid a -> Int # elem :: Eq a => a -> Constraint' rigid a -> Bool # maximum :: Ord a => Constraint' rigid a -> a # minimum :: Ord a => Constraint' rigid a -> a # sum :: Num a => Constraint' rigid a -> a # product :: Num a => Constraint' rigid a -> a # | |||||
Traversable (Constraint' rigid) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods traverse :: Applicative f => (a -> f b) -> Constraint' rigid a -> f (Constraint' rigid b) # sequenceA :: Applicative f => Constraint' rigid (f a) -> f (Constraint' rigid a) # mapM :: Monad m => (a -> m b) -> Constraint' rigid a -> m (Constraint' rigid b) # sequence :: Monad m => Constraint' rigid (m a) -> m (Constraint' rigid a) # | |||||
(Pretty r, Pretty f) => Pretty (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods pretty :: Constraint' r f -> Doc Source # prettyPrec :: Int -> Constraint' r f -> Doc Source # prettyList :: [Constraint' r f] -> Doc Source # | |||||
Ord flex => Flexs (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods flexs :: Constraint' rigid flex -> Set (FlexOf (Constraint' rigid flex)) Source # | |||||
Ord r => Rigids (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods rigids :: Constraint' r f -> Set (RigidOf (Constraint' r f)) Source # | |||||
(NFData r, NFData f) => NFData (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods rnf :: Constraint' r f -> () # | |||||
Generic (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods from :: Constraint' rigid flex -> Rep (Constraint' rigid flex) x # to :: Rep (Constraint' rigid flex) x -> Constraint' rigid flex # | |||||
(Show rigid, Show flex) => Show (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods showsPrec :: Int -> Constraint' rigid flex -> ShowS # show :: Constraint' rigid flex -> String # showList :: [Constraint' rigid flex] -> ShowS # | |||||
type SubstArg SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve | |||||
type FlexOf (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
type RigidOf (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
type Rep (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax type Rep (Constraint' rigid flex) = D1 ('MetaData "Constraint'" "Agda.TypeChecking.SizedTypes.Syntax" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "Constraint" 'PrefixI 'True) (S1 ('MetaSel ('Just "leftExpr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SizeExpr' rigid flex)) :*: (S1 ('MetaSel ('Just "cmp") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Cmp) :*: S1 ('MetaSel ('Just "rightExpr") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 (SizeExpr' rigid flex))))) |
type Constraint = Constraint' Rigid Flex Source #
Polarities to specify solutions.
What type of solution are we looking for?
data PolarityAssignment flex Source #
Assigning a polarity to a flexible variable.
Constructors
PolarityAssignment Polarity flex |
Instances
Pretty flex => Pretty (PolarityAssignment flex) Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods pretty :: PolarityAssignment flex -> Doc Source # prettyPrec :: Int -> PolarityAssignment flex -> Doc Source # prettyList :: [PolarityAssignment flex] -> Doc Source # |
type Polarities flex = Map flex Polarity Source #
Type of solution wanted for each flexible.
emptyPolarities :: Polarities flex Source #
polaritiesFromAssignments :: Ord flex => [PolarityAssignment flex] -> Polarities flex Source #
getPolarity :: Ord flex => Polarities flex -> flex -> Polarity Source #
Default polarity is Least
.
Solutions.
newtype Solution rigid flex Source #
Partial substitution from flexible variables to size expression.
Constructors
Solution | |
Fields
|
emptySolution :: Solution r f Source #
class Substitute r f a where Source #
Executing a substitution.
Instances
Substitute r f a => Substitute r f [a] Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax | |
Ord f => Substitute r f (Constraint' r f) Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods subst :: Solution r f -> Constraint' r f -> Constraint' r f Source # | |
Ord f => Substitute r f (SizeExpr' r f) Source # | |
Ord f => Substitute r f (Solution r f) Source # | |
Substitute r f a => Substitute r f (Map k a) Source # | |
Constraint simplification
type CTrans r f = Constraint' r f -> Maybe [Constraint' r f] Source #
simplify1 :: (Pretty f, Pretty r, Eq r) => CTrans r f -> CTrans r f Source #
Returns Nothing
if the constraint is contradictory.
Printing
Wellformedness
class ValidOffset a where Source #
Offsets + n
must be non-negative
Methods
validOffset :: a -> Bool Source #
Instances
ValidOffset Offset Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods validOffset :: Offset -> Bool Source # | |
ValidOffset (SizeExpr' r f) Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods validOffset :: SizeExpr' r f -> Bool Source # |
class TruncateOffset a where Source #
Make offsets non-negative by rounding up.
Methods
truncateOffset :: a -> a Source #
Instances
TruncateOffset Offset Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods truncateOffset :: Offset -> Offset Source # | |
TruncateOffset (SizeExpr' r f) Source # | |
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods truncateOffset :: SizeExpr' r f -> SizeExpr' r f Source # |
Computing variable sets
class Ord (RigidOf a) => Rigids a where Source #
The rigid variables contained in a pice of syntax.
Instances
Rigids a => Rigids [a] Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
| |||||
Ord r => Rigids (Constraint' r f) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods rigids :: Constraint' r f -> Set (RigidOf (Constraint' r f)) Source # | |||||
Ord r => Rigids (SizeExpr' r f) Source # | |||||
class Ord (FlexOf a) => Flexs a where Source #
The flexibe variables contained in a pice of syntax.
Instances
Flexs HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods flexs :: HypSizeConstraint -> Set (FlexOf HypSizeConstraint) Source # | |||||
Flexs a => Flexs [a] Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
| |||||
Ord flex => Flexs (Constraint' rigid flex) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods flexs :: Constraint' rigid flex -> Set (FlexOf (Constraint' rigid flex)) Source # | |||||
Ord flex => Flexs (SizeExpr' rigid flex) Source # | |||||
data NamedRigid Source #
Identifiers for rigid variables.
Constructors
NamedRigid | |
Fields
|
Instances
Pretty NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods pretty :: NamedRigid -> Doc Source # prettyPrec :: Int -> NamedRigid -> Doc Source # prettyList :: [NamedRigid] -> Doc Source # | |||||
PrettyTCM SizeConstraint Source # | Assumes we are in the right context. | ||||
Defined in Agda.TypeChecking.SizedTypes.Pretty Methods prettyTCM :: MonadPretty m => SizeConstraint -> m Doc Source # | |||||
Subst SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg SizeConstraint) -> SizeConstraint -> SizeConstraint Source # | |||||
NFData NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods rnf :: NamedRigid -> () # | |||||
Generic NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
| |||||
Show NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods showsPrec :: Int -> NamedRigid -> ShowS # show :: NamedRigid -> String # showList :: [NamedRigid] -> ShowS # | |||||
Eq NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
Ord NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods compare :: NamedRigid -> NamedRigid -> Ordering # (<) :: NamedRigid -> NamedRigid -> Bool # (<=) :: NamedRigid -> NamedRigid -> Bool # (>) :: NamedRigid -> NamedRigid -> Bool # (>=) :: NamedRigid -> NamedRigid -> Bool # max :: NamedRigid -> NamedRigid -> NamedRigid # min :: NamedRigid -> NamedRigid -> NamedRigid # | |||||
Plus NamedRigid Int NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods plus :: NamedRigid -> Int -> NamedRigid Source # | |||||
Subst (SizeExpr' NamedRigid SizeMeta) Source # | Only for | ||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg (SizeExpr' NamedRigid SizeMeta)) -> SizeExpr' NamedRigid SizeMeta -> SizeExpr' NamedRigid SizeMeta Source # | |||||
type SubstArg SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve | |||||
type Rep NamedRigid Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax type Rep NamedRigid = D1 ('MetaData "NamedRigid" "Agda.TypeChecking.SizedTypes.Syntax" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "NamedRigid" 'PrefixI 'True) (S1 ('MetaSel ('Just "rigidName") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 String) :*: S1 ('MetaSel ('Just "rigidIndex") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Int))) | |||||
type SubstArg (SizeExpr' NamedRigid SizeMeta) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve |
Size metas in size expressions.
Constructors
SizeMeta | |
Fields
|
Instances
Pretty SizeMeta Source # | |||||
PrettyTCM SizeConstraint Source # | Assumes we are in the right context. | ||||
Defined in Agda.TypeChecking.SizedTypes.Pretty Methods prettyTCM :: MonadPretty m => SizeConstraint -> m Doc Source # | |||||
PrettyTCM SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Pretty | |||||
Subst SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg SizeConstraint) -> SizeConstraint -> SizeConstraint Source # | |||||
Subst SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg SizeMeta) -> SizeMeta -> SizeMeta Source # | |||||
NFData SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
Generic SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
| |||||
Show SizeMeta Source # | |||||
Eq SizeMeta Source # | An equality which ignores the meta arguments. | ||||
Ord SizeMeta Source # | An order which ignores the meta arguments. | ||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
Subst (SizeExpr' NamedRigid SizeMeta) Source # | Only for | ||||
Defined in Agda.TypeChecking.SizedTypes.Solve Associated Types
Methods applySubst :: Substitution' (SubstArg (SizeExpr' NamedRigid SizeMeta)) -> SizeExpr' NamedRigid SizeMeta -> SizeExpr' NamedRigid SizeMeta Source # | |||||
type SubstArg SizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve | |||||
type SubstArg SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve | |||||
type Rep SizeMeta Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax type Rep SizeMeta = D1 ('MetaData "SizeMeta" "Agda.TypeChecking.SizedTypes.Syntax" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "SizeMeta" 'PrefixI 'True) (S1 ('MetaSel ('Just "sizeMetaId") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 MetaId) :*: S1 ('MetaSel ('Just "sizeMetaArgs") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Int]))) | |||||
type SubstArg (SizeExpr' NamedRigid SizeMeta) Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Solve |
type DBSizeExpr = SizeExpr' NamedRigid SizeMeta Source #
Size expression with de Bruijn indices.
data HypSizeConstraint Source #
Size constraint with de Bruijn indices.
Constructors
HypSizeConstraint | |
Fields
|
Instances
PrettyTCM HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Pretty Methods prettyTCM :: MonadPretty m => HypSizeConstraint -> m Doc Source # | |||||
Flexs HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods flexs :: HypSizeConstraint -> Set (FlexOf HypSizeConstraint) Source # | |||||
NFData HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods rnf :: HypSizeConstraint -> () # | |||||
Generic HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Associated Types
Methods from :: HypSizeConstraint -> Rep HypSizeConstraint x # to :: Rep HypSizeConstraint x -> HypSizeConstraint # | |||||
Show HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax Methods showsPrec :: Int -> HypSizeConstraint -> ShowS # show :: HypSizeConstraint -> String # showList :: [HypSizeConstraint] -> ShowS # | |||||
type FlexOf HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax | |||||
type Rep HypSizeConstraint Source # | |||||
Defined in Agda.TypeChecking.SizedTypes.Syntax type Rep HypSizeConstraint = D1 ('MetaData "HypSizeConstraint" "Agda.TypeChecking.SizedTypes.Syntax" "Agda-2.8.0-inplace" 'False) (C1 ('MetaCons "HypSizeConstraint" 'PrefixI 'True) ((S1 ('MetaSel ('Just "sizeContext") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 Context) :*: S1 ('MetaSel ('Just "sizeHypIds") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [Nat])) :*: (S1 ('MetaSel ('Just "sizeHypotheses") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 [SizeConstraint]) :*: S1 ('MetaSel ('Just "sizeConstraint") 'NoSourceUnpackedness 'NoSourceStrictness 'DecidedLazy) (Rec0 SizeConstraint)))) |