{ {-# LANGUAGE TupleSections #-} {-| The parser is generated by Happy (). - - Ideally, ranges should be as precise as possible, to get messages that - emphasize precisely the faulting term(s) upon error. - - However, interactive highlighting is only applied at the end of each - mutual block, keywords are only highlighted once (see - `TypeChecking.Rules.Decl'). So if the ranges of two declarations - interleave, one must ensure that keyword ranges are not included in - the intersection. (Otherwise they are uncolored by the interactive - highlighting.) - -} module Agda.Syntax.Parser.Parser ( moduleParser , exprParser , tokensParser , tests ) where import Control.Monad import Data.Char import Data.Functor import Data.List import Data.Maybe import qualified Data.Traversable as T import Agda.Syntax.Position hiding (tests) import Agda.Syntax.Parser.Monad import Agda.Syntax.Parser.Lexer import Agda.Syntax.Parser.Tokens import Agda.Syntax.Concrete as C import Agda.Syntax.Concrete.Pretty () import Agda.Syntax.Common hiding (Arg, Dom, NamedArg) import qualified Agda.Syntax.Common as Common import Agda.Syntax.Fixity import Agda.Syntax.Notation import Agda.Syntax.Literal import Agda.Utils.Hash import Agda.Utils.List (spanJust) import Agda.Utils.Monad import Agda.Utils.QuickCheck import Agda.Utils.Singleton import Agda.Utils.TestHelpers import Agda.Utils.Tuple } %name tokensParser Tokens %name exprParser Expr %name moduleParser File %tokentype { Token } %monad { Parser } %lexer { lexer } { TokEOF } %expect 1 -- shift/reduce for \ x y z -> foo = bar -- shifting means it'll parse as \ x y z -> (foo = bar) rather than -- (\ x y z -> foo) = bar -- This is a trick to get rid of shift/reduce conflicts arising because we want -- to parse things like "m >>= \x -> k x". See the Expr rule for more -- information. %nonassoc LOWEST %nonassoc '->' %token 'let' { TokKeyword KwLet $$ } 'in' { TokKeyword KwIn $$ } 'where' { TokKeyword KwWhere $$ } 'with' { TokKeyword KwWith $$ } 'rewrite' { TokKeyword KwRewrite $$ } 'postulate' { TokKeyword KwPostulate $$ } 'primitive' { TokKeyword KwPrimitive $$ } 'open' { TokKeyword KwOpen $$ } 'import' { TokKeyword KwImport $$ } 'using' { TokKeyword KwUsing $$ } 'hiding' { TokKeyword KwHiding $$ } 'renaming' { TokKeyword KwRenaming $$ } 'to' { TokKeyword KwTo $$ } 'public' { TokKeyword KwPublic $$ } 'module' { TokKeyword KwModule $$ } 'data' { TokKeyword KwData $$ } 'codata' { TokKeyword KwCoData $$ } 'record' { TokKeyword KwRecord $$ } 'constructor' { TokKeyword KwConstructor $$ } 'inductive' { TokKeyword KwInductive $$ } 'coinductive' { TokKeyword KwCoInductive $$ } 'field' { TokKeyword KwField $$ } 'infix' { TokKeyword KwInfix $$ } 'infixl' { TokKeyword KwInfixL $$ } 'infixr' { TokKeyword KwInfixR $$ } 'mutual' { TokKeyword KwMutual $$ } 'abstract' { TokKeyword KwAbstract $$ } 'private' { TokKeyword KwPrivate $$ } 'instance' { TokKeyword KwInstance $$ } 'Prop' { TokKeyword KwProp $$ } 'Set' { TokKeyword KwSet $$ } 'forall' { TokKeyword KwForall $$ } 'syntax' { TokKeyword KwSyntax $$ } 'pattern' { TokKeyword KwPatternSyn $$ } 'OPTIONS' { TokKeyword KwOPTIONS $$ } 'BUILTIN' { TokKeyword KwBUILTIN $$ } 'REWRITE' { TokKeyword KwREWRITE $$ } 'IMPORT' { TokKeyword KwIMPORT $$ } 'IMPOSSIBLE' { TokKeyword KwIMPOSSIBLE $$ } 'ETA' { TokKeyword KwETA $$ } 'NO_TERMINATION_CHECK' { TokKeyword KwNO_TERMINATION_CHECK $$ } 'NON_TERMINATING' { TokKeyword KwNON_TERMINATING $$ } 'TERMINATING' { TokKeyword KwTERMINATING $$ } 'MEASURE' { TokKeyword KwMEASURE $$ } 'COMPILED' { TokKeyword KwCOMPILED $$ } 'COMPILED_EXPORT' { TokKeyword KwCOMPILED_EXPORT $$ } 'COMPILED_DATA' { TokKeyword KwCOMPILED_DATA $$ } 'COMPILED_TYPE' { TokKeyword KwCOMPILED_TYPE $$ } 'COMPILED_EPIC' { TokKeyword KwCOMPILED_EPIC $$ } 'COMPILED_JS' { TokKeyword KwCOMPILED_JS $$ } 'STATIC' { TokKeyword KwSTATIC $$ } 'quoteGoal' { TokKeyword KwQuoteGoal $$ } 'quoteContext' { TokKeyword KwQuoteContext $$ } 'quote' { TokKeyword KwQuote $$ } 'quoteTerm' { TokKeyword KwQuoteTerm $$ } 'tactic' { TokKeyword KwTactic $$ } 'unquote' { TokKeyword KwUnquote $$ } 'unquoteDecl' { TokKeyword KwUnquoteDecl $$ } setN { TokSetN $$ } tex { TokTeX $$ } comment { TokComment $$ } '...' { TokSymbol SymEllipsis $$ } '..' { TokSymbol SymDotDot $$ } '.' { TokSymbol SymDot $$ } ';' { TokSymbol SymSemi $$ } ':' { TokSymbol SymColon $$ } '=' { TokSymbol SymEqual $$ } '_' { TokSymbol SymUnderscore $$ } '?' { TokSymbol SymQuestionMark $$ } '->' { TokSymbol SymArrow $$ } '\\' { TokSymbol SymLambda $$ } '@' { TokSymbol SymAs $$ } '|' { TokSymbol SymBar $$ } '(' { TokSymbol SymOpenParen $$ } ')' { TokSymbol SymCloseParen $$ } '{{' { TokSymbol SymDoubleOpenBrace $$ } '}}' { TokSymbol SymDoubleCloseBrace $$ } '{' { TokSymbol SymOpenBrace $$ } '}' { TokSymbol SymCloseBrace $$ } -- ':{' { TokSymbol SymColonBrace $$ } vopen { TokSymbol SymOpenVirtualBrace $$ } vclose { TokSymbol SymCloseVirtualBrace $$ } vsemi { TokSymbol SymVirtualSemi $$ } '{-#' { TokSymbol SymOpenPragma $$ } '#-}' { TokSymbol SymClosePragma $$ } id { TokId $$ } q_id { TokQId $$ } string { TokString $$ } literal { TokLiteral $$ } %% {-------------------------------------------------------------------------- Parsing the token stream. Used by the TeX compiler. --------------------------------------------------------------------------} -- Parse a list of tokens. Tokens :: { [Token] } Tokens : TokensR { reverse $1 } -- Happy is much better at parsing left recursive grammars (constant -- stack size vs. linear stack size for right recursive). TokensR :: { [Token] } TokensR : TokensR Token { $2 : $1 } | { [] } -- Parse single token. Token :: { Token } Token : 'let' { TokKeyword KwLet $1 } | 'in' { TokKeyword KwIn $1 } | 'where' { TokKeyword KwWhere $1 } | 'with' { TokKeyword KwWith $1 } | 'rewrite' { TokKeyword KwRewrite $1 } | 'postulate' { TokKeyword KwPostulate $1 } | 'primitive' { TokKeyword KwPrimitive $1 } | 'open' { TokKeyword KwOpen $1 } | 'import' { TokKeyword KwImport $1 } | 'using' { TokKeyword KwUsing $1 } | 'hiding' { TokKeyword KwHiding $1 } | 'renaming' { TokKeyword KwRenaming $1 } | 'to' { TokKeyword KwTo $1 } | 'public' { TokKeyword KwPublic $1 } | 'module' { TokKeyword KwModule $1 } | 'data' { TokKeyword KwData $1 } | 'codata' { TokKeyword KwCoData $1 } | 'record' { TokKeyword KwRecord $1 } | 'constructor' { TokKeyword KwConstructor $1 } | 'inductive' { TokKeyword KwInductive $1 } | 'coinductive' { TokKeyword KwCoInductive $1 } | 'field' { TokKeyword KwField $1 } | 'infix' { TokKeyword KwInfix $1 } | 'infixl' { TokKeyword KwInfixL $1 } | 'infixr' { TokKeyword KwInfixR $1 } | 'mutual' { TokKeyword KwMutual $1 } | 'abstract' { TokKeyword KwAbstract $1 } | 'private' { TokKeyword KwPrivate $1 } | 'instance' { TokKeyword KwInstance $1 } | 'Prop' { TokKeyword KwProp $1 } | 'Set' { TokKeyword KwSet $1 } | 'forall' { TokKeyword KwForall $1 } | 'syntax' { TokKeyword KwSyntax $1 } | 'pattern' { TokKeyword KwPatternSyn $1 } | 'OPTIONS' { TokKeyword KwOPTIONS $1 } | 'BUILTIN' { TokKeyword KwBUILTIN $1 } | 'REWRITE' { TokKeyword KwREWRITE $1 } | 'IMPORT' { TokKeyword KwIMPORT $1 } | 'COMPILED' { TokKeyword KwCOMPILED $1 } | 'COMPILED_EXPORT' { TokKeyword KwCOMPILED_EXPORT $1 } | 'COMPILED_DATA'{ TokKeyword KwCOMPILED_DATA $1 } | 'COMPILED_TYPE'{ TokKeyword KwCOMPILED_TYPE $1 } | 'COMPILED_EPIC'{ TokKeyword KwCOMPILED_EPIC $1 } | 'COMPILED_JS' { TokKeyword KwCOMPILED_JS $1 } | 'STATIC' { TokKeyword KwSTATIC $1 } | 'IMPOSSIBLE' { TokKeyword KwIMPOSSIBLE $1 } | 'ETA' { TokKeyword KwETA $1 } | 'NO_TERMINATION_CHECK' { TokKeyword KwNO_TERMINATION_CHECK $1 } | 'NON_TERMINATING' { TokKeyword KwNON_TERMINATING $1 } | 'TERMINATING' { TokKeyword KwTERMINATING $1 } | 'MEASURE' { TokKeyword KwMEASURE $1 } | 'quoteGoal' { TokKeyword KwQuoteGoal $1 } | 'quoteContext' { TokKeyword KwQuoteContext $1 } | 'quote' { TokKeyword KwQuote $1 } | 'quoteTerm' { TokKeyword KwQuoteTerm $1 } | 'tactic' { TokKeyword KwTactic $1 } | 'unquote' { TokKeyword KwUnquote $1 } | 'unquoteDecl' { TokKeyword KwUnquoteDecl $1 } | setN { TokSetN $1 } | tex { TokTeX $1 } | comment { TokComment $1 } | '...' { TokSymbol SymEllipsis $1 } | '..' { TokSymbol SymDotDot $1 } | '.' { TokSymbol SymDot $1 } | ';' { TokSymbol SymSemi $1 } | ':' { TokSymbol SymColon $1 } | '=' { TokSymbol SymEqual $1 } | '_' { TokSymbol SymUnderscore $1 } | '?' { TokSymbol SymQuestionMark $1 } | '->' { TokSymbol SymArrow $1 } | '\\' { TokSymbol SymLambda $1 } | '@' { TokSymbol SymAs $1 } | '|' { TokSymbol SymBar $1 } | '(' { TokSymbol SymOpenParen $1 } | ')' { TokSymbol SymCloseParen $1 } | '{{' { TokSymbol SymDoubleOpenBrace $1 } | '}}' { TokSymbol SymDoubleCloseBrace $1 } | '{' { TokSymbol SymOpenBrace $1 } | '}' { TokSymbol SymCloseBrace $1 } | vopen { TokSymbol SymOpenVirtualBrace $1 } | vclose { TokSymbol SymCloseVirtualBrace $1 } | vsemi { TokSymbol SymVirtualSemi $1 } | '{-#' { TokSymbol SymOpenPragma $1 } | '#-}' { TokSymbol SymClosePragma $1 } | id { TokId $1 } | q_id { TokQId $1 } | string { TokString $1 } | literal { TokLiteral $1 } {-------------------------------------------------------------------------- Top level --------------------------------------------------------------------------} File :: { ([Pragma], [Declaration]) } File : vopen TopLevel maybe_vclose { takeOptionsPragmas $2 } maybe_vclose : {- empty -} { () } | vclose { () } {-------------------------------------------------------------------------- Meta rules --------------------------------------------------------------------------} -- The first token in a file decides the indentation of the top-level layout -- block. Or not. It will if we allow the top-level module to be omitted. -- topen : {- empty -} {% pushCurrentContext } {- A layout block might have to be closed by a parse error. Example: let x = e in e' Here the 'let' starts a layout block which should end before the 'in'. The problem is that the lexer doesn't know this, so there is no virtual close brace. However when the parser sees the 'in' there will be a parse error. This is our cue to close the layout block. -} close : vclose { () } | error {% popContext } -- You can use concrete semi colons in a layout block started with a virtual -- brace, so we don't have to distinguish between the two semi colons. You can't -- use a virtual semi colon in a block started by a concrete brace, but this is -- simply because the lexer will not generate virtual semis in this case. semi : ';' { $1 } | vsemi { $1 } -- Enter the 'imp_dir' lex state, where we can parse the keyword 'to'. beginImpDir :: { () } beginImpDir : {- empty -} {% pushLexState imp_dir } {-------------------------------------------------------------------------- Helper rules --------------------------------------------------------------------------} -- An integer. Used in fixity declarations. Int :: { Integer } Int : literal {% case $1 of { LitInt _ i -> return i; _ -> fail $ "Expected integer" } } | id {% case $1 of { (_, s) -> case readM s of { Right i -> return i; Left (err :: String) -> fail $ "Expected integer" } } } {-------------------------------------------------------------------------- Names --------------------------------------------------------------------------} -- A name is really a sequence of parts, but the lexer just sees it as a -- string, so we have to do the translation here. Id :: { Name } Id : id {% mkName $1 } -- Space separated list of one or more identifiers. SpaceIds :: { [Name] } SpaceIds : Id SpaceIds { $1 : $2 } | Id { [$1] } -- When looking for a double closed brace, we accept either a single token '}}' -- (which is what the unicode character "RIGHT WHITE CURLY BRACKET" is -- postprocessed into in LexActions.hs), but also two consecutive tokens '}' -- (which a string '}}' is lexed to). This small hack allows us to keep -- "record { a = record { }}" working. In the second case, we check that the two -- tokens '}' are immediately consecutive. DoubleCloseBrace :: { Range } DoubleCloseBrace : '}}' { getRange $1 } | '}' '}' {% if posPos (fromJust (rEnd (getRange $2))) - posPos (fromJust (rStart (getRange $1))) > 2 then parseErrorAt (fromJust (rStart (getRange $2))) "Expecting '}}', found separated '}'s." else return $ getRange ($1, $2) } -- A possibly dotted identifier. MaybeDottedId :: { Arg Name } MaybeDottedId : '.' Id { setRelevance Irrelevant $ defaultArg $2 } | Id { defaultArg $1 } -- Space separated list of one or more possibly dotted identifiers. MaybeDottedIds :: { [Arg Name] } MaybeDottedIds : MaybeDottedId MaybeDottedIds { $1 : $2 } | MaybeDottedId { [$1] } -- Space separated list of one or more identifiers, some of which may -- be surrounded by braces or dotted. ArgIds :: { [Arg Name] } ArgIds : MaybeDottedId ArgIds { $1 : $2 } | MaybeDottedId { [$1] } | '{{' MaybeDottedIds DoubleCloseBrace ArgIds { map makeInstance $2 ++ $4 } | '{{' MaybeDottedIds DoubleCloseBrace { map makeInstance $2 } | '{' MaybeDottedIds '}' ArgIds { map hide $2 ++ $4 } | '{' MaybeDottedIds '}' { map hide $2 } | '.' '{' SpaceIds '}' ArgIds { map (hide . setRelevance Irrelevant . defaultArg) $3 ++ $5 } | '.' '{' SpaceIds '}' { map (hide . setRelevance Irrelevant . defaultArg) $3 } | '.' '{{' SpaceIds DoubleCloseBrace ArgIds { map (makeInstance . setRelevance Irrelevant . defaultArg) $3 ++ $5 } | '.' '{{' SpaceIds DoubleCloseBrace { map (makeInstance . setRelevance Irrelevant . defaultArg) $3 } | '..' '{' SpaceIds '}' ArgIds { map (hide . setRelevance NonStrict . defaultArg) $3 ++ $5 } | '..' '{' SpaceIds '}' { map (hide . setRelevance NonStrict . defaultArg) $3 } | '..' '{{' SpaceIds DoubleCloseBrace ArgIds { map (makeInstance . setRelevance NonStrict . defaultArg) $3 ++ $5 } | '..' '{{' SpaceIds DoubleCloseBrace { map (makeInstance . setRelevance NonStrict . defaultArg) $3 } QId :: { QName } QId : q_id {% mkQName $1 } | Id { QName $1 } -- A module name is just a qualified name ModuleName :: { QName } ModuleName : QId { $1 } -- A binding variable. Can be '_' BId :: { Name } BId : Id { $1 } | '_' { Name (getRange $1) [Hole] } {- UNUSED -- A binding variable. Can be '_' MaybeDottedBId :: { (Relevance, Name) } MaybeDottedBId : BId { (Relevant , $1) } | '.' BId { (Irrelevant, $2) } | '..' BId { (NonStrict, $2) } -} -- Space separated list of binding identifiers. Used in fixity -- declarations infixl 100 + - SpaceBIds :: { [Name] } SpaceBIds : BId SpaceBIds { $1 : $2 } | BId { [$1] } {- DOES PRODUCE REDUCE/REDUCE CONFLICTS! -- Space-separated list of binding identifiers. Used in dependent -- function spaces: (x y z : Nat) -> ... -- (Used to be comma-separated; hence the name) -- QUESTION: Should this be replaced by SpaceBIds above? --CommaBIds :: { [(Relevance,Name)] } CommaBIds :: { [Name] } CommaBIds : CommaBIds BId { $1 ++ [$2] } -- SWITCHING DOES NOT HELP | BId { [$1] } -} -- Space-separated list of binding identifiers. Used in dependent -- function spaces: (x y z : Nat) -> ... -- (Used to be comma-separated; hence the name) -- QUESTION: Should this be replaced by SpaceBIds above? -- Andreas, 2011-04-07 the trick avoids reduce/reduce conflicts -- when parsing (x y z : A) -> B -- at point (x y it is not clear whether x y is an application or -- a variable list. We could be parsing (x y z) -> B -- with ((x y) z) being a type. CommaBIds :: { [Name] } CommaBIds : CommaBIdAndAbsurds { case $1 of Left ns -> ns Right _ -> fail $ "expected sequence of bound identifiers, not absurd pattern" } CommaBIdAndAbsurds :: { Either [Name] [Expr] } CommaBIdAndAbsurds : Application {% let getName :: Expr -> Maybe Name getName (Ident (QName x)) = Just x getName (Underscore r _) = Just (Name r [Hole]) getName _ = Nothing isAbsurd :: Expr -> Bool isAbsurd (Absurd _) = True isAbsurd (HiddenArg _ (Named _ e)) = isAbsurd e isAbsurd (InstanceArg _ (Named _ e)) = isAbsurd e isAbsurd (Paren _ expr) = isAbsurd expr isAbsurd (RawApp _ exprs) = any isAbsurd exprs isAbsurd _ = False in if any isAbsurd $1 then return $ Right $1 else case mapM getName $1 of Just good -> return $ Left good Nothing -> fail $ "expected sequence of bound identifiers" } -- Parse a sequence of identifiers, including hiding info. -- Does not include instance arguments. -- E.g. x {y z} _ {v} -- To be used in typed bindings, like (x {y z} _ {v} : Nat). BIdsWithHiding :: { [WithHiding Name] } BIdsWithHiding : Application {% let -- interpret an expression as name getName :: Expr -> Maybe Name getName (Ident (QName x)) = Just x getName (Underscore r _) = Just (Name r [Hole]) getName _ = Nothing getNames :: Expr -> Maybe [Name] getNames (RawApp _ es) = mapM getName es getNames e = singleton `fmap` getName e -- interpret an expression as name or list of hidden names getName1 :: Expr -> Maybe [WithHiding Name] getName1 (Ident (QName x)) = Just [WithHiding NotHidden x] getName1 (Underscore r _) = Just [WithHiding NotHidden $ Name r [Hole]] getName1 (HiddenArg _ (Named Nothing e)) = map (WithHiding Hidden) `fmap` getNames e getName1 _ = Nothing in case mapM getName1 $1 of Just good -> return $ concat good Nothing -> fail $ "expected sequence of possibly hidden bound identifiers" } -- Space separated list of strings in a pragma. PragmaStrings :: { [String] } PragmaStrings : {- empty -} { [] } | string PragmaStrings { snd $1 : $2 } PragmaString :: { String } PragmaString : string { snd $1 } PragmaName :: { Name } PragmaName : string {% mkName $1 } PragmaQName :: { QName } PragmaQName : string {% fmap QName (mkName $1) } {-------------------------------------------------------------------------- Expressions (terms and types) --------------------------------------------------------------------------} {- Expressions. You might expect lambdas and lets to appear in the first expression category (lowest precedence). The reason they don't is that we want to parse things like m >>= \x -> k x This will leads to a conflict in the following case m >>= \x -> k x >>= \y -> k' y At the second '>>=' we can either shift or reduce. We solve this problem using Happy's precedence directives. The rule 'Expr -> Expr1' (which is the rule you shouldn't use to reduce when seeing '>>=') is given LOWEST precedence. The terminals '->' and op (which is what you should shift) is given higher precedence. -} -- Top level: Function types. Expr :: { Expr } Expr : TeleArrow Expr { Pi $1 $2 } | 'forall' ForallBindings Expr { forallPi $2 $3 } | Application3 '->' Expr { Fun (getRange ($1,$2,$3)) (RawApp (getRange $1) $1) $3 } | Expr1 '=' Expr { Equal (getRange ($1, $2, $3)) $1 $3 } | Expr1 %prec LOWEST { $1 } -- Level 1: Application Expr1 : WithExprs {% case $1 of { [e] -> return e ; e : es -> return $ WithApp (fuseRange e es) e es ; [] -> fail "impossible: empty with expressions" } } WithExprs :: { [Expr] } WithExprs : Application3 '|' WithExprs { RawApp (getRange $1) $1 : $3 } | Application { [RawApp (getRange $1) $1] } Application :: { [Expr] } Application : Expr2 { [$1] } | Expr3 Application { $1 : $2 } -- Level 2: Lambdas and lets Expr2 : '\\' LamBindings Expr { Lam (getRange ($1,$2,$3)) $2 $3 } | ExtendedOrAbsurdLam { $1 } | 'let' Declarations 'in' Expr { Let (getRange ($1,$2,$3,$4)) $2 $4 } | Expr3 { $1 } | 'quoteGoal' Id 'in' Expr { QuoteGoal (getRange ($1,$2,$3,$4)) $2 $4 } | 'quoteContext' Id 'in' Expr { QuoteContext (getRange ($1,$2,$3,$4)) $2 $4 } | 'tactic' Application3 { Tactic (getRange ($1, $2)) (RawApp (getRange $2) $2) [] } | 'tactic' Application3 '|' WithExprs { Tactic (getRange ($1, $2, $3, $4)) (RawApp (getRange $2) $2) $4 } ExtendedOrAbsurdLam :: { Expr } ExtendedOrAbsurdLam : '\\' '{' LamClauses '}' { ExtendedLam (getRange ($1,$2,$3,$4)) (reverse $3) } | '\\' AbsurdLamBindings {% case $2 of Left (bs, h) -> if null bs then return $ AbsurdLam r h else return $ Lam r bs (AbsurdLam r h) where r = fuseRange $1 bs Right es -> do -- it is of the form @\ { p1 ... () }@ p <- exprToLHS (RawApp (getRange es) es); return $ ExtendedLam (fuseRange $1 es) [(p [] [], AbsurdRHS, NoWhere)] } Application3 :: { [Expr] } Application3 : Expr3 { [$1] } | Expr3 Application3 { $1 : $2 } -- Level 3: Atoms Expr3Curly : '{' Expr '}' { HiddenArg (getRange ($1,$2,$3)) (maybeNamed $2) } | '{' '}' { let r = fuseRange $1 $2 in HiddenArg r $ unnamed $ Absurd r } Expr3NoCurly : QId { Ident $1 } | literal { Lit $1 } | '?' { QuestionMark (getRange $1) Nothing } | '_' { Underscore (getRange $1) Nothing } | 'Prop' { Prop (getRange $1) } | 'Set' { Set (getRange $1) } | 'quote' { Quote (getRange $1) } | 'quoteTerm' { QuoteTerm (getRange $1) } | 'unquote' { Unquote (getRange $1) } | setN { SetN (getRange (fst $1)) (snd $1) } | '{{' Expr DoubleCloseBrace { InstanceArg (getRange ($1,$2,$3)) (maybeNamed $2) } | '(' Expr ')' { Paren (getRange ($1,$2,$3)) $2 } | '(' ')' { Absurd (fuseRange $1 $2) } | '{{' DoubleCloseBrace { let r = fuseRange $1 $2 in InstanceArg r $ unnamed $ Absurd r } | Id '@' Expr3 { As (getRange ($1,$2,$3)) $1 $3 } | '.' Expr3 { Dot (fuseRange $1 $2) $2 } | 'record' '{' FieldAssignments '}' { Rec (getRange ($1,$2,$3,$4)) $3 } | 'record' Expr3NoCurly '{' FieldAssignments '}' { RecUpdate (getRange ($1,$2,$3,$4,$5)) $2 $4 } Expr3 : Expr3Curly { $1 } | Expr3NoCurly { $1 } FieldAssignments :: { [(Name, Expr)] } FieldAssignments : {- empty -} { [] } | FieldAssignments1 { $1 } FieldAssignments1 :: { [(Name, Expr)] } FieldAssignments1 : FieldAssignment { [$1] } | FieldAssignment ';' FieldAssignments1 { $1 : $3 } FieldAssignment :: { (Name, Expr) } FieldAssignment : Id '=' Expr { ($1, $3) } {-------------------------------------------------------------------------- Bindings --------------------------------------------------------------------------} -- "Delta ->" to avoid conflict between Delta -> Gamma and Delta -> A. TeleArrow : Telescope1 '->' { $1 } Telescope1 : TypedBindingss { {-TeleBind-} $1 } TypedBindingss :: { [TypedBindings] } TypedBindingss : TypedBindings TypedBindingss { $1 : $2 } | TypedBindings { [$1] } -- A typed binding is either (x1 .. xn : A) or {y1 .. ym : B} -- Andreas, 2011-04-07: or .(x1 .. xn : A) or .{y1 .. ym : B} -- Andreas, 2011-04-27: or ..(x1 .. xn : A) or ..{y1 .. ym : B} TypedBindings :: { TypedBindings } TypedBindings : '.' '(' TBindWithHiding ')' { setRange (getRange ($2,$3,$4)) $ setRelevance Irrelevant $3 } | '.' '{' TBind '}' { setRange (getRange ($2,$3,$4)) $ setHiding Hidden $ setRelevance Irrelevant $3 } | '.' '{{' TBind DoubleCloseBrace { setRange (getRange ($2,$3,$4)) $ setHiding Instance $ setRelevance Irrelevant $3 } | '..' '(' TBindWithHiding ')' { setRange (getRange ($2,$3,$4)) $ setRelevance NonStrict $3 } | '..' '{' TBind '}' { setRange (getRange ($2,$3,$4)) $ setHiding Hidden $ setRelevance NonStrict $3 } | '..' '{{' TBind DoubleCloseBrace { setRange (getRange ($2,$3,$4)) $ setHiding Instance $ setRelevance NonStrict $3 } | '(' TBindWithHiding ')' { setRange (getRange ($1,$2,$3)) $2 } | '{{' TBind DoubleCloseBrace { setRange (getRange ($1,$2,$3)) $ setHiding Instance $2 } | '{' TBind '}' { setRange (getRange ($1,$2,$3)) $ setHiding Hidden $2 } | '(' Open ')' { tLet (getRange ($1,$3)) $2 } | '(' 'let' Declarations ')' { tLet (getRange ($1,$4)) $3 } -- x1 .. xn : A -- x1 .. xn :{i1 i2 ..} A TBind :: { TypedBindings } TBind : CommaBIds ':' Expr { let r = getRange ($1,$2,$3) -- the range is approximate only for TypedBindings in TypedBindings r $ defaultArg $ TBind r (map (pure . mkBoundName_) $1) $3 } -- | Colors are not yet allowed in the syntax. -- | CommaBIds ':{' Colors '}' Expr { ( $3, TBind (getRange ($1,$2,$3,$4,$5)) (map mkBoundName_ $1) $5 ) } {- Colors :: { [Color] } Colors : QId Colors { Ident $1 : $2 } | QId { [Ident $1] } -} -- x {y z} _ {v} : A TBindWithHiding :: { TypedBindings } TBindWithHiding : BIdsWithHiding ':' Expr { let r = getRange ($1,$2,$3) -- the range is approximate only for TypedBindings in TypedBindings r $ defaultArg $ TBind r (map (fmap mkBoundName_) $1) $3 } -- A non-empty sequence of lambda bindings. LamBindings :: { [LamBinding] } LamBindings : LamBinds '->' {% case reverse $1 of Left _ : _ -> parseError "Absurd lambda cannot have a body." _ : _ -> return [ b | Right b <- $1 ] [] -> parsePanic "Empty LamBinds" } AbsurdLamBindings :: { Either ([LamBinding], Hiding) [Expr] } AbsurdLamBindings : LamBindsAbsurd {% case $1 of Left lb -> case reverse lb of Right _ : _ -> parseError "Missing body for lambda" Left h : _ -> return $ Left ([ b | Right b <- init lb], h) _ -> parseError "Unsupported variant of lambda" Right es -> return $ Right es } -- absurd lambda is represented by @Left hiding@ LamBinds :: { [Either Hiding LamBinding] } LamBinds : DomainFreeBinding LamBinds { map Right $1 ++ $2 } | TypedBindings LamBinds { Right (DomainFull $1) : $2 } | DomainFreeBinding { map Right $1 } | TypedBindings { [Right $ DomainFull $1] } | '(' ')' { [Left NotHidden] } | '{' '}' { [Left Hidden] } | '{{' DoubleCloseBrace { [Left Instance] } -- Like LamBinds, but could also parse an absurd LHS of an extended lambda @{ p1 ... () }@ LamBindsAbsurd :: { Either [Either Hiding LamBinding] [Expr] } LamBindsAbsurd : DomainFreeBinding LamBinds { Left $ map Right $1 ++ $2 } | TypedBindings LamBinds { Left $ Right (DomainFull $1) : $2 } | DomainFreeBindingAbsurd { case $1 of Left lb -> Left $ map Right lb Right es -> Right es } | TypedBindings { Left [Right $ DomainFull $1] } | '(' ')' { Left [Left NotHidden] } | '{' '}' { Left [Left Hidden] } | '{{' DoubleCloseBrace { Left [Left Instance] } -- FNF, 2011-05-05: No where clauses in extended lambdas for now NonAbsurdLamClause :: { (LHS,RHS,WhereClause) } NonAbsurdLamClause : Application3 '->' Expr {% do p <- exprToLHS (RawApp (getRange $1) $1) ; return (p [] [], RHS $3, NoWhere) } AbsurdLamClause :: { (LHS,RHS,WhereClause) } AbsurdLamClause -- FNF, 2011-05-09: By being more liberal here, we avoid shift/reduce and reduce/reduce errors. -- Later stages such as scope checking will complain if we let something through which we should not : Application {% do p <- exprToLHS (RawApp (getRange $1) $1); return (p [] [], AbsurdRHS, NoWhere) } LamClause :: { (LHS,RHS,WhereClause) } LamClause : NonAbsurdLamClause { $1 } | AbsurdLamClause { $1 } -- Parses all extended lambda clauses except for a single absurd clause, which is taken care of -- in AbsurdLambda LamClauses :: { [(LHS,RHS,WhereClause)] } LamClauses : LamClauses semi LamClause { $3 : $1 } | AbsurdLamClause semi LamClause { [$3, $1] } | NonAbsurdLamClause { [$1] } -- | {- empty -} { [] } ForallBindings :: { [LamBinding] } ForallBindings : TypedUntypedBindings1 '->' { $1 } -- A non-empty sequence of possibly untyped bindings. TypedUntypedBindings1 :: { [LamBinding] } TypedUntypedBindings1 : DomainFreeBinding TypedUntypedBindings1 { $1 ++ $2 } | TypedBindings TypedUntypedBindings1 { DomainFull $1 : $2 } | DomainFreeBinding { $1 } | TypedBindings { [DomainFull $1] } -- A possibly empty sequence of possibly untyped bindings. -- This is used as telescope in data and record decls. TypedUntypedBindings :: { [LamBinding] } TypedUntypedBindings : DomainFreeBinding TypedUntypedBindings { $1 ++ $2 } | TypedBindings TypedUntypedBindings { DomainFull $1 : $2 } | { [] } -- A domain free binding is either x or {x1 .. xn} DomainFreeBinding :: { [LamBinding] } DomainFreeBinding : DomainFreeBindingAbsurd { case $1 of Left lbs -> lbs Right _ -> fail "expected sequence of bound identifiers, not absurd pattern" } {- : BId { [DomainFree NotHidden Relevant $ mkBoundName_ $1] } | '.' BId { [DomainFree NotHidden Irrelevant $ mkBoundName_ $2] } | '..' BId { [DomainFree NotHidden NonStrict $ mkBoundName_ $2] } | '{' CommaBIds '}' { map (DomainFree Hidden Relevant . mkBoundName_) $2 } | '{{' CommaBIds DoubleCloseBrace { map (DomainFree (setHiding Instance defaultArgInfo) . mkBoundName_) $2 } | '.' '{' CommaBIds '}' { map (DomainFree Hidden Irrelevant . mkBoundName_) $3 } | '.' '{{' CommaBIds DoubleCloseBrace { map (DomainFree Instance Irrelevant . mkBoundName_) $3 } | '..' '{' CommaBIds '}' { map (DomainFree Hidden NonStrict . mkBoundName_) $3 } | '..' '{{' CommaBIds DoubleCloseBrace { map (DomainFree Instance NonStrict . mkBoundName_) $3 } | '..' '{{' CommaBIds DoubleCloseBrace { map (DomainFree Instance NonStrict . mkBoundName_) $3 } -} -- A domain free binding is either x or {x1 .. xn} DomainFreeBindingAbsurd :: { Either [LamBinding] [Expr]} DomainFreeBindingAbsurd : BId { Left [DomainFree defaultArgInfo $ mkBoundName_ $1] } | '.' BId { Left [DomainFree (setRelevance Irrelevant $ defaultArgInfo) $ mkBoundName_ $2] } | '..' BId { Left [DomainFree (setRelevance NonStrict $ defaultArgInfo) $ mkBoundName_ $2] } | '{' CommaBIdAndAbsurds '}' { either (Left . map (DomainFree (setHiding Hidden $ defaultArgInfo) . mkBoundName_)) Right $2 } | '{{' CommaBIds DoubleCloseBrace { Left $ map (DomainFree (setHiding Instance $ defaultArgInfo) . mkBoundName_) $2 } | '.' '{' CommaBIds '}' { Left $ map (DomainFree (setHiding Hidden $ setRelevance Irrelevant $ defaultArgInfo) . mkBoundName_) $3 } | '.' '{{' CommaBIds DoubleCloseBrace { Left $ map (DomainFree (setHiding Instance $ setRelevance Irrelevant $ defaultArgInfo) . mkBoundName_) $3 } | '..' '{' CommaBIds '}' { Left $ map (DomainFree (setHiding Hidden $ setRelevance NonStrict $ defaultArgInfo) . mkBoundName_) $3 } | '..' '{{' CommaBIds DoubleCloseBrace { Left $ map (DomainFree (setHiding Instance $ setRelevance NonStrict $ defaultArgInfo) . mkBoundName_) $3 } {-------------------------------------------------------------------------- Modules and imports --------------------------------------------------------------------------} -- Import directives ImportDirective :: { ImportDirective } ImportDirective : ImportDirectives {% mergeImportDirectives $1 } ImportDirectives :: { [ImportDirective] } ImportDirectives : ImportDirective1 ImportDirectives { $1 : $2 } | {- empty -} { [] } ImportDirective1 :: { ImportDirective } : 'public' { defaultImportDir { importDirRange = getRange $1, publicOpen = True } } | UsingOrHiding { defaultImportDir { importDirRange = snd $1, usingOrHiding = fst $1 } } | RenamingDir { defaultImportDir { importDirRange = snd $1, renaming = fst $1 } } UsingOrHiding :: { (UsingOrHiding , Range) } UsingOrHiding : 'using' '(' CommaImportNames ')' { (Using $3 , getRange ($1,$2,$3,$4)) } -- using can have an empty list | 'hiding' '(' CommaImportNames ')' { (Hiding $3 , getRange ($1,$2,$3,$4)) } -- if you want to hide nothing that's fine, isn't it? RenamingDir :: { ([Renaming] , Range) } RenamingDir : 'renaming' '(' Renamings ')' { ($3 , getRange ($1,$2,$3,$4)) } | 'renaming' '(' ')' { ([] , getRange ($1,$2,$3)) } -- Renamings of the form 'x to y' Renamings :: { [Renaming] } Renamings : Renaming ';' Renamings { $1 : $3 } | Renaming { [$1] } Renaming :: { Renaming } Renaming : ImportName_ 'to' Id { Renaming $1 $3 (getRange $2) } -- We need a special imported name here, since we have to trigger -- the imp_dir state exactly one token before the 'to' ImportName_ :: { ImportedName } ImportName_ : beginImpDir Id { ImportedName $2 } | 'module' beginImpDir Id { ImportedModule $3 } ImportName :: { ImportedName } ImportName : Id { ImportedName $1 } | 'module' Id { ImportedModule $2 } -- Actually semi-colon separated CommaImportNames :: { [ImportedName] } CommaImportNames : {- empty -} { [] } | CommaImportNames1 { $1 } CommaImportNames1 : ImportName { [$1] } | ImportName ';' CommaImportNames1 { $1 : $3 } {-------------------------------------------------------------------------- Function clauses --------------------------------------------------------------------------} -- A left hand side of a function clause. We parse it as an expression, and -- then check that it is a valid left hand side. LHS :: { LHS } LHS : Expr1 RewriteEquations WithExpressions {% exprToLHS $1 >>= \p -> return (p $2 $3) } | '...' WithPats RewriteEquations WithExpressions { Ellipsis (getRange ($1,$2,$3,$4)) $2 $3 $4 } WithPats :: { [Pattern] } WithPats : {- empty -} { [] } | '|' Application3 WithPats {% exprToPattern (RawApp (getRange $2) $2) >>= \p -> return (p : $3) } WithExpressions :: { [Expr] } WithExpressions : {- empty -} { [] } | 'with' Expr { case $2 of { WithApp _ e es -> e : es; e -> [e] } } RewriteEquations :: { [Expr] } RewriteEquations : {- empty -} { [] } | 'rewrite' Expr1 { case $2 of { WithApp _ e es -> e : es; e -> [e] } } -- Where clauses are optional. WhereClause :: { WhereClause } WhereClause : {- empty -} { NoWhere } | 'where' Declarations0 { AnyWhere $2 } | 'module' Id 'where' Declarations0 { SomeWhere $2 $4 } | 'module' Underscore 'where' Declarations0 { SomeWhere $2 $4 } {-------------------------------------------------------------------------- Different kinds of declarations --------------------------------------------------------------------------} -- Top-level definitions. Declaration :: { [Declaration] } Declaration : Fields { $1 } | FunClause { $1 } | Data { [$1] } | DataSig { [$1] } -- lone data type signature in mutual block | Record { [$1] } | RecordSig { [$1] } -- lone record signature in mutual block | Infix { [$1] } | Mutual { [$1] } | Abstract { [$1] } | Private { [$1] } | Instance { [$1] } | Postulate { [$1] } | Primitive { [$1] } | Open { $1 } -- | Import { [$1] } | ModuleMacro { [$1] } | Module { [$1] } | Pragma { [$1] } | Syntax { [$1] } | PatternSyn { [$1] } | UnquoteDecl { [$1] } {-------------------------------------------------------------------------- Individual declarations --------------------------------------------------------------------------} -- Type signatures of the form "n1 n2 n3 ... : Type", with at least -- one bound name. TypeSigs :: { [Declaration] } TypeSigs : SpaceIds ':' Expr { map (\ x -> TypeSig defaultArgInfo x $3) $1 } -- A variant of TypeSigs where any sub-sequence of names can be marked -- as hidden or irrelevant using braces and dots: -- {n1 .n2} n3 .n4 {n5} .{n6 n7} ... : Type. ArgTypeSigs :: { [Arg Declaration] } ArgTypeSigs : ArgIds ':' Expr { map (fmap (\ x -> TypeSig defaultArgInfo x $3)) $1 } -- Function declarations. The left hand side is parsed as an expression to allow -- declarations like 'x::xs ++ ys = e', when '::' has higher precedence than '++'. -- FunClause also handle possibly dotted type signatures. FunClause :: { [Declaration] } FunClause : LHS RHS WhereClause {% funClauseOrTypeSigs $1 $2 $3 } RHS :: { RHSOrTypeSigs } RHS : '=' Expr { JustRHS (RHS $2) } | ':' Expr { TypeSigsRHS $2 } | {- empty -} { JustRHS AbsurdRHS } -- Data declaration. Can be local. Data :: { Declaration } Data : 'data' Id TypedUntypedBindings ':' Expr 'where' Constructors { Data (getRange ($1,$2,$3,$4,$5,$6,$7)) Inductive $2 $3 (Just $5) $7 } | 'codata' Id TypedUntypedBindings ':' Expr 'where' Constructors { Data (getRange ($1,$2,$3,$4,$5,$6,$7)) CoInductive $2 $3 (Just $5) $7 } -- New cases when we already had a DataSig. Then one can omit the sort. | 'data' Id TypedUntypedBindings 'where' Constructors { Data (getRange ($1,$2,$3,$4,$5)) Inductive $2 $3 Nothing $5 } | 'codata' Id TypedUntypedBindings 'where' Constructors { Data (getRange ($1,$2,$3,$4,$5)) CoInductive $2 $3 Nothing $5 } -- Data type signature. Found in mutual blocks. DataSig :: { Declaration } DataSig : 'data' Id TypedUntypedBindings ':' Expr { DataSig (getRange ($1,$2,$3,$4,$5)) Inductive $2 $3 $5 } -- Andreas, 2012-03-16: The Expr3NoCurly instead of Id in everything -- following 'record' is to remove the (harmless) shift/reduce conflict -- introduced by record update expressions. -- Record declarations. Record :: { Declaration } Record : 'record' Expr3NoCurly TypedUntypedBindings ':' Expr 'where' RecordDeclarations {% exprToName $2 >>= \ n -> return $ Record (getRange ($1,$2,$3,$4,$5,$6,$7)) n (fst3 $7) (snd3 $7) $3 (Just $5) (thd3 $7) } | 'record' Expr3NoCurly TypedUntypedBindings 'where' RecordDeclarations {% exprToName $2 >>= \ n -> return $ Record (getRange ($1,$2,$3,$4,$5)) n (fst3 $5) (snd3 $5) $3 Nothing (thd3 $5) } -- Record type signature. In mutual blocks. RecordSig :: { Declaration } RecordSig : 'record' Expr3NoCurly TypedUntypedBindings ':' Expr {% exprToName $2 >>= \ n -> return $ RecordSig (getRange ($1,$2,$3,$4,$5)) n $3 $5 } -- Declaration of record constructor name. RecordConstructorName :: { Name } RecordConstructorName : 'constructor' Id { $2 } -- Fixity declarations. Infix :: { Declaration } Infix : 'infix' Int SpaceBIds { Infix (Fixity (getRange ($1,$3)) $2 NonAssoc) $3 } | 'infixl' Int SpaceBIds { Infix (Fixity (getRange ($1,$3)) $2 LeftAssoc) $3 } | 'infixr' Int SpaceBIds { Infix (Fixity (getRange ($1,$3)) $2 RightAssoc) $3 } -- Field declarations. Fields :: { [Declaration] } Fields : 'field' ArgTypeSignatures { let toField (Common.Arg info (TypeSig _ x t)) = Field x (Common.Arg info t) in map toField $2 } --REM { let toField (h, TypeSig x t) = Field h x t in map toField $2 } -- Mutually recursive declarations. Mutual :: { Declaration } Mutual : 'mutual' Declarations { Mutual (fuseRange $1 $2) $2 } -- Abstract declarations. Abstract :: { Declaration } Abstract : 'abstract' Declarations { Abstract (fuseRange $1 $2) $2 } -- Private can only appear on the top-level (or rather the module level). Private :: { Declaration } Private : 'private' Declarations { Private (fuseRange $1 $2) $2 } -- Instance declarations. Instance :: { Declaration } Instance : 'instance' Declarations { InstanceB (fuseRange $1 $2) $2 } -- Postulates. Postulate :: { Declaration } Postulate : 'postulate' Declarations { Postulate (fuseRange $1 $2) $2 } -- Primitives. Can only contain type signatures. Primitive :: { Declaration } Primitive : 'primitive' TypeSignatures { Primitive (fuseRange $1 $2) $2 } -- Unquoting declarations. UnquoteDecl :: { Declaration } UnquoteDecl : 'unquoteDecl' Id '=' Expr { UnquoteDecl (fuseRange $1 $4) $2 $4 } -- Syntax declaration (To declare eg. mixfix binders) Syntax :: { Declaration } Syntax : 'syntax' Id HoleNames '=' SimpleIds {% case $2 of Name _ [_] -> case mkNotation $3 (map rangedThing $5) of Left err -> parseError $ "Malformed syntax declaration: " ++ err Right n -> return $ Syntax $2 n _ -> parseError "Syntax declarations are allowed only for simple names (without holes)" } -- Pattern synonyms. PatternSyn :: { Declaration } PatternSyn : 'pattern' Id PatternSynArgs '=' Expr {% do p <- exprToPattern $5 return (PatternSyn (getRange ($1,$2,$3,$4,$5)) $2 $3 p) } PatternSynArgs :: { [Arg Name] } PatternSynArgs : {- empty -} { [] } | LamBinds {% patternSynArgs $1 } SimpleIds :: { [RString] } SimpleIds : SimpleId { [$1] } | SimpleIds SimpleId {$1 ++ [$2]} HoleNames :: { [NamedArg HoleName] } HoleNames : HoleName { [$1] } | HoleNames HoleName {$1 ++ [$2]} HoleName :: { NamedArg HoleName } HoleName : SimpleTopHole { defaultNamedArg $1 } | '{' SimpleHole '}' { setHiding Hidden $ defaultNamedArg $2 } | '{{' SimpleHole '}}' { setHiding Instance $ defaultNamedArg $2 } | '{' SimpleId '=' SimpleHole '}' { setHiding Hidden $ defaultArg $ named $2 $4 } | '{{' SimpleId '=' SimpleHole '}}' { setHiding Instance $ defaultArg $ named $2 $4 } SimpleTopHole :: { HoleName } SimpleTopHole : SimpleId { ExprHole (rangedThing $1) } | '(' '\\' SimpleId '->' SimpleId ')' { LambdaHole (rangedThing $3) (rangedThing $5) } SimpleHole :: { HoleName } SimpleHole : SimpleId { ExprHole (rangedThing $1) } | '\\' SimpleId '->' SimpleId { LambdaHole (rangedThing $2) (rangedThing $4) } -- Variable name hole to be implemented later. -- Discard the interval. SimpleId :: { RString } SimpleId : id { Ranged (getRange $ fst $1) (stringToRawName $ snd $1) } MaybeOpen :: { Maybe Range } MaybeOpen : 'open' { Just (getRange $1) } | {- empty -} { Nothing } -- Open Open :: { [Declaration] } Open : MaybeOpen 'import' ModuleName OpenArgs ImportDirective {% let { doOpen = maybe DontOpen (const DoOpen) $1 ; m = $3 ; es = $4 ; dir = $5 ; r = getRange (m, es, dir) ; mr = getRange m ; unique = hashString $ show $ (Nothing :: Maybe ()) <$ r -- turn range into unique id, but delete file path -- which is absolute and messes up suite of failing tests -- (different hashs on different installations) -- TODO: Don't use (insecure) hashes in this way. ; fresh = Name mr [ Id $ stringToRawName $ ".#" ++ show m ++ "-" ++ show unique ] ; impStm asR = Import mr m (Just (AsName fresh asR)) DontOpen defaultImportDir ; appStm m' es = let r = getRange (m, es) in Private r [ ModuleMacro r m' (SectionApp (getRange es) [] (RawApp (getRange es) (Ident (QName fresh) : es))) doOpen dir ] ; (initArgs, last2Args) = splitAt (length es - 2) es ; parseAsClause = case last2Args of { [ Ident (QName (Name asR [Id x])) , Ident (QName m') ] | rawNameToString x == "as" -> Just (asR, m') ; _ -> Nothing } } in case es of { [] -> return [Import mr m Nothing doOpen dir] ; _ | Just (asR, m') <- parseAsClause -> if null initArgs then return [ Import (getRange (m, asR, m', dir)) m (Just (AsName m' asR)) doOpen dir ] else return [ impStm asR, appStm m' initArgs ] | DontOpen <- doOpen -> parseErrorAt (fromJust $ rStart $ getRange $2) "An import statement with module instantiation does not actually import the module. This statement achieves nothing. Either add the `open' keyword or bind the instantiated module with an `as' clause." | otherwise -> return [ impStm noRange , appStm (noName $ beginningOf $ getRange m) es ] } } |'open' ModuleName OpenArgs ImportDirective { let { m = $2 ; es = $3 ; dir = $4 ; r = getRange (m, es, dir) } in [ case es of { [] -> Open r m dir ; _ -> Private r [ ModuleMacro r (noName $ beginningOf $ getRange m) (SectionApp (getRange (m , es)) [] (RawApp (fuseRange m es) (Ident m : es))) DoOpen dir ] } ] } | 'open' ModuleName '{{' '...' DoubleCloseBrace ImportDirective { let r = getRange $2 in [ Private r [ ModuleMacro r (noName $ beginningOf $ getRange $2) (RecordModuleIFS r $2) DoOpen $6 ] ] } OpenArgs :: { [Expr] } OpenArgs : {- empty -} { [] } | Expr3 OpenArgs { $1 : $2 } ModuleApplication :: { [TypedBindings] -> Parser ModuleApplication } ModuleApplication : ModuleName '{{' '...' DoubleCloseBrace { (\ts -> if null ts then return $ RecordModuleIFS (getRange ($1,$2,$3,$4)) $1 else parseError "No bindings allowed for record module with non-canonical implicits" ) } | ModuleName OpenArgs { (\ts -> return $ SectionApp (getRange ($1, $2)) ts (RawApp (fuseRange $1 $2) (Ident $1 : $2)) ) } -- Module instantiation ModuleMacro :: { Declaration } ModuleMacro : 'module' ModuleName TypedUntypedBindings '=' ModuleApplication ImportDirective {% do { ma <- $5 (map addType $3) ; name <- ensureUnqual $2 ; return $ ModuleMacro (getRange ($1, $2, ma, $6)) name ma DontOpen $6 } } | 'open' 'module' Id TypedUntypedBindings '=' ModuleApplication ImportDirective {% do {ma <- $6 (map addType $4); return $ ModuleMacro (getRange ($1, $2, $3, ma, $7)) $3 ma DoOpen $7 } } -- Module Module :: { Declaration } Module : 'module' ModuleName TypedUntypedBindings 'where' Declarations0 { Module (getRange ($1,$2,$3,$4,$5)) $2 (map addType $3) $5 } | 'module' Underscore TypedUntypedBindings 'where' Declarations0 { Module (getRange ($1,$2,$3,$4,$5)) (QName $2) (map addType $3) $5 } Underscore :: { Name } Underscore : '_' { noName (getRange $1) } TopLevel :: { [Declaration] } TopLevel : TopDeclarations { figureOutTopLevelModule $1 } Pragma :: { Declaration } Pragma : DeclarationPragma { Pragma $1 } DeclarationPragma :: { Pragma } DeclarationPragma : BuiltinPragma { $1 } | RewritePragma { $1 } | CompiledPragma { $1 } | CompiledExportPragma { $1 } | CompiledDataPragma { $1 } | CompiledTypePragma { $1 } | CompiledEpicPragma { $1 } | CompiledJSPragma { $1 } | StaticPragma { $1 } | ImportPragma { $1 } | ImpossiblePragma { $1 } | RecordEtaPragma { $1 } | TerminatingPragma { $1 } | NonTerminatingPragma { $1 } | NoTerminationCheckPragma { $1 } | MeasurePragma { $1 } | OptionsPragma { $1 } -- Andreas, 2014-03-06 -- OPTIONS pragma not allowed everywhere, but don't give parse error. -- Give better error during type checking instead. OptionsPragma :: { Pragma } OptionsPragma : '{-#' 'OPTIONS' PragmaStrings '#-}' { OptionsPragma (getRange ($1,$2,$4)) $3 } BuiltinPragma :: { Pragma } BuiltinPragma : '{-#' 'BUILTIN' string PragmaQName '#-}' { BuiltinPragma (getRange ($1,$2,fst $3,$4,$5)) (snd $3) (Ident $4) } -- Extra rule to accept keword REWRITE also as built-in: | '{-#' 'BUILTIN' 'REWRITE' PragmaQName '#-}' { BuiltinPragma (getRange ($1,$2,$3,$4,$5)) "REWRITE" (Ident $4) } RewritePragma :: { Pragma } RewritePragma : '{-#' 'REWRITE' PragmaQName '#-}' { RewritePragma (getRange ($1,$2,$3,$4)) $3 } CompiledPragma :: { Pragma } CompiledPragma : '{-#' 'COMPILED' PragmaQName PragmaStrings '#-}' { CompiledPragma (getRange ($1,$2,$3,$5)) $3 (unwords $4) } CompiledExportPragma :: { Pragma } CompiledExportPragma : '{-#' 'COMPILED_EXPORT' PragmaQName PragmaString '#-}' { CompiledExportPragma (getRange ($1,$2,$3,$5)) $3 $4 } CompiledTypePragma :: { Pragma } CompiledTypePragma : '{-#' 'COMPILED_TYPE' PragmaQName PragmaStrings '#-}' { CompiledTypePragma (getRange ($1,$2,$3,$5)) $3 (unwords $4) } CompiledDataPragma :: { Pragma } CompiledDataPragma : '{-#' 'COMPILED_DATA' PragmaQName string PragmaStrings '#-}' { CompiledDataPragma (getRange ($1,$2,$3,fst $4,$6)) $3 (snd $4) $5 } CompiledEpicPragma :: { Pragma } CompiledEpicPragma : '{-#' 'COMPILED_EPIC' PragmaQName PragmaStrings '#-}' { CompiledEpicPragma (getRange ($1,$2,$3,$5)) $3 (unwords $4) } CompiledJSPragma :: { Pragma } CompiledJSPragma : '{-#' 'COMPILED_JS' PragmaQName PragmaStrings '#-}' { CompiledJSPragma (getRange ($1,$2,$3,$5)) $3 (unwords $4) } StaticPragma :: { Pragma } StaticPragma : '{-#' 'STATIC' PragmaQName '#-}' { StaticPragma (getRange ($1,$2,$3,$4)) $3 } RecordEtaPragma :: { Pragma } RecordEtaPragma : '{-#' 'ETA' PragmaQName '#-}' { EtaPragma (getRange ($1,$2,$3,$4)) $3 } NoTerminationCheckPragma :: { Pragma } NoTerminationCheckPragma : '{-#' 'NO_TERMINATION_CHECK' '#-}' { TerminationCheckPragma (getRange ($1,$2,$3)) NoTerminationCheck } NonTerminatingPragma :: { Pragma } NonTerminatingPragma : '{-#' 'NON_TERMINATING' '#-}' { TerminationCheckPragma (getRange ($1,$2,$3)) NonTerminating } TerminatingPragma :: { Pragma } TerminatingPragma : '{-#' 'TERMINATING' '#-}' { TerminationCheckPragma (getRange ($1,$2,$3)) Terminating } MeasurePragma :: { Pragma } MeasurePragma : '{-#' 'MEASURE' PragmaName '#-}' { let r = getRange ($1, $2, $3, $4) in TerminationCheckPragma r (TerminationMeasure r $3) } ImportPragma :: { Pragma } ImportPragma : '{-#' 'IMPORT' string '#-}' {% let s = snd $3 in if validHaskellModuleName s then return $ ImportPragma (getRange ($1,$2,fst $3,$4)) s else parseError $ "Malformed module name: " ++ s ++ "." } ImpossiblePragma :: { Pragma } : '{-#' 'IMPOSSIBLE' '#-}' { ImpossiblePragma (getRange ($1,$2,$3)) } {-------------------------------------------------------------------------- Sequences of declarations --------------------------------------------------------------------------} -- Non-empty list of type signatures, with several identifiers allowed -- for every signature. TypeSignatures :: { [TypeSignature] } TypeSignatures : vopen TypeSignatures1 close { reverse $2 } -- Inside the layout block. TypeSignatures1 :: { [TypeSignature] } TypeSignatures1 : TypeSignatures1 semi TypeSigs { reverse $3 ++ $1 } | TypeSigs { reverse $1 } -- A variant of TypeSignatures which uses ArgTypeSigs instead of -- TypeSigs. ArgTypeSignatures :: { [Arg TypeSignature] } ArgTypeSignatures : vopen ArgTypeSignatures1 close { reverse $2 } -- Inside the layout block. ArgTypeSignatures1 :: { [Arg TypeSignature] } ArgTypeSignatures1 : ArgTypeSignatures1 semi ArgTypeSigs { reverse $3 ++ $1 } | ArgTypeSigs { reverse $1 } -- Constructors are type signatures. But constructor lists can be empty. Constructors :: { [Constructor] } Constructors : vopen close { [] } | TypeSignatures { $1 } -- Record declarations, including an optional record constructor name. RecordDeclarations :: { (Maybe (Ranged Induction), Maybe Name, [Declaration]) } RecordDeclarations : vopen close { (Nothing, Nothing, []) } | vopen RecordConstructorName close { (Nothing, Just $2, []) } | vopen RecordConstructorName semi Declarations1 close { (Nothing, Just $2, $4) } | vopen Declarations1 close { (Nothing, Nothing, $2) } | vopen RecordInduction close { (Just $2, Nothing, []) } | vopen RecordInduction semi RecordConstructorName close { (Just $2, Just $4, []) } | vopen RecordInduction semi RecordConstructorName semi Declarations1 close { (Just $2, Just $4, $6) } | vopen RecordInduction semi Declarations1 close { (Just $2, Nothing, $4) } -- Declaration of record as 'inductive' or 'coinductive'. RecordInduction :: { Ranged Induction } RecordInduction : 'inductive' { Ranged (getRange $1) Inductive } | 'coinductive' { Ranged (getRange $1) CoInductive } -- Arbitrary declarations Declarations :: { [Declaration] } Declarations : vopen Declarations1 close { $2 } -- Arbitrary declarations Declarations0 :: { [Declaration] } Declarations0 : vopen close { [] } | Declarations { $1 } Declarations1 :: { [Declaration] } Declarations1 : Declaration semi Declarations1 { $1 ++ $3 } | Declaration { $1 } TopDeclarations :: { [Declaration] } TopDeclarations : {- empty -} { [] } | Declarations1 { $1 } { {-------------------------------------------------------------------------- Parsers --------------------------------------------------------------------------} -- | Parse the token stream. Used by the TeX compiler. tokensParser :: Parser [Token] -- | Parse an expression. Could be used in interactions. exprParser :: Parser Expr -- | Parse a module. moduleParser :: Parser Module {-------------------------------------------------------------------------- Happy stuff --------------------------------------------------------------------------} -- | Required by Happy. happyError :: Parser a happyError = parseError "Parse error" {-------------------------------------------------------------------------- Utility functions --------------------------------------------------------------------------} -- | Grab leading OPTIONS pragmas. takeOptionsPragmas :: [Declaration] -> ([Pragma], [Declaration]) takeOptionsPragmas = spanJust $ \ d -> case d of Pragma p@OptionsPragma{} -> Just p _ -> Nothing -- | Insert a top-level module if there is none. figureOutTopLevelModule :: [Declaration] -> [Declaration] figureOutTopLevelModule ds = case span isAllowedBeforeModule ds of (ds0, Module r m tel ds1 : ds2) -> ds0 ++ [Module r m tel $ ds1 ++ ds2] (ds0, ds1) -> ds0 ++ [Module (getRange ds1) (QName noName_) [] ds1] where isAllowedBeforeModule (Pragma OptionsPragma{}) = True isAllowedBeforeModule (Private _ ds) = all isAllowedBeforeModule ds isAllowedBeforeModule Import{} = True isAllowedBeforeModule ModuleMacro{} = True isAllowedBeforeModule Open{} = True isAllowedBeforeModule _ = False -- | Create a name from a string. mkName :: (Interval, String) -> Parser Name mkName (i, s) = do let xs = C.stringNameParts s mapM_ isValidId xs unless (alternating xs) $ fail $ "a name cannot contain two consecutive underscores" return $ Name (getRange i) xs where isValidId Hole = return () isValidId (Id y) = do let x = rawNameToString y case parse defaultParseFlags [0] (lexer return) x of ParseOk _ (TokId _) -> return () _ -> fail $ "in the name " ++ s ++ ", the part " ++ x ++ " is not valid" -- we know that there are no two Ids in a row alternating (Hole : Hole : _) = False alternating (_ : xs) = alternating xs alternating [] = True -- | Create a qualified name from a list of strings mkQName :: [(Interval, String)] -> Parser QName mkQName ss = do xs <- mapM mkName ss return $ foldr Qual (QName $ last xs) (init xs) ensureUnqual :: QName -> Parser Name ensureUnqual (QName x) = return x ensureUnqual q@Qual{} = parseError' (rStart $ getRange q) "Qualified name not allowed here" -- | Match a particular name. isName :: String -> (Interval, String) -> Parser () isName s (_,s') | s == s' = return () | otherwise = fail $ "expected " ++ s ++ ", found " ++ s' -- | Build a forall pi (forall x y z -> ...) forallPi :: [LamBinding] -> Expr -> Expr forallPi bs e = Pi (map addType bs) e -- | Build a telescoping let (let Ds) tLet :: Range -> [Declaration] -> TypedBindings tLet r = TypedBindings r . Common.Arg defaultArgInfo . TLet r -- | Converts lambda bindings to typed bindings. addType :: LamBinding -> TypedBindings addType (DomainFull b) = b addType (DomainFree info x) = TypedBindings r $ Common.Arg info $ TBind r [pure x] $ Underscore r Nothing where r = getRange x mergeImportDirectives :: [ImportDirective] -> Parser ImportDirective mergeImportDirectives is = do i <- foldl merge (return defaultImportDir) is verifyImportDirective i where merge mi i2 = do i1 <- mi let err = parseError' (rStart $ getRange i2) "Cannot mix using and hiding module directives" uh <- case (usingOrHiding i1, usingOrHiding i2) of (Hiding [], u) -> return u (u, Hiding []) -> return u (Using{}, Hiding{}) -> err (Hiding{}, Using{}) -> err (Using xs, Using ys) -> return $ Using (xs ++ ys) (Hiding xs, Hiding ys) -> return $ Hiding (xs ++ ys) return $ ImportDirective { importDirRange = fuseRange i1 i2 , usingOrHiding = uh , renaming = renaming i1 ++ renaming i2 , publicOpen = publicOpen i1 || publicOpen i2 } -- | Check that an import directive doesn't contain repeated names verifyImportDirective :: ImportDirective -> Parser ImportDirective verifyImportDirective i = case filter ((>1) . length) $ group $ sort xs of [] -> return i yss -> let Just pos = rStart $ getRange $ head $ concat yss in parseErrorAt pos $ "Repeated name" ++ s ++ " in import directive: " ++ concat (intersperse ", " $ map (show . head) yss) where s = case yss of [_] -> "" _ -> "s" where xs = names (usingOrHiding i) ++ map renFrom (renaming i) names (Using xs) = xs names (Hiding xs) = xs -- | Breaks up a string into substrings. Returns every maximal -- subsequence of zero or more characters distinct from @'.'@. -- -- > splitOnDots "" == [""] -- > splitOnDots "foo.bar" == ["foo", "bar"] -- > splitOnDots ".foo.bar" == ["", "foo", "bar"] -- > splitOnDots "foo.bar." == ["foo", "bar", ""] -- > splitOnDots "foo..bar" == ["foo", "", "bar"] splitOnDots :: String -> [String] splitOnDots "" = [""] splitOnDots ('.' : s) = [] : splitOnDots s splitOnDots (c : s) = case splitOnDots s of p : ps -> (c : p) : ps prop_splitOnDots = and [ splitOnDots "" == [""] , splitOnDots "foo.bar" == ["foo", "bar"] , splitOnDots ".foo.bar" == ["", "foo", "bar"] , splitOnDots "foo.bar." == ["foo", "bar", ""] , splitOnDots "foo..bar" == ["foo", "", "bar"] ] -- | Returns 'True' iff the name is a valid Haskell (hierarchical) -- module name. validHaskellModuleName :: String -> Bool validHaskellModuleName = all ok . splitOnDots where -- Checks if a dot-less module name is well-formed. ok :: String -> Bool ok [] = False ok (c : s) = isUpper c && all (\c -> isLower c || c == '_' || isUpper c || generalCategory c == DecimalNumber || c == '\'') s {-------------------------------------------------------------------------- Patterns --------------------------------------------------------------------------} -- | Turn an expression into a left hand side. exprToLHS :: Expr -> Parser ([Expr] -> [Expr] -> LHS) exprToLHS e = case e of WithApp r e es -> LHS <$> exprToPattern e <*> mapM exprToPattern es _ -> LHS <$> exprToPattern e <*> return [] -- | Turn an expression into a pattern. Fails if the expression is not a -- valid pattern. exprToPattern :: Expr -> Parser Pattern exprToPattern e = case e of Ident x -> return $ IdentP x App _ e1 e2 -> AppP <$> exprToPattern e1 <*> T.mapM (T.mapM exprToPattern) e2 Paren r e -> ParenP r <$> exprToPattern e Underscore r _ -> return $ WildP r Absurd r -> return $ AbsurdP r As r x e -> AsP r x <$> exprToPattern e Dot r (HiddenArg _ e) -> return $ HiddenP r $ fmap (DotP r) e Dot r e -> return $ DotP r e Lit l -> return $ LitP l HiddenArg r e -> HiddenP r <$> T.mapM exprToPattern e InstanceArg r e -> InstanceP r <$> T.mapM exprToPattern e RawApp r es -> RawAppP r <$> mapM exprToPattern es Quote r -> return $ QuoteP r _ -> let Just pos = rStart $ getRange e in parseErrorAt pos $ "Not a valid pattern: " ++ show e opAppExprToPattern :: OpApp Expr -> Parser Pattern opAppExprToPattern (SyntaxBindingLambda _ _ _) = parseError "Syntax binding lambda cannot appear in a pattern" opAppExprToPattern (Ordinary e) = exprToPattern e -- | Turn an expression into a name. Fails if the expression is not a -- valid identifier. exprToName :: Expr -> Parser Name exprToName (Ident (QName x)) = return x exprToName e = let Just pos = rStart $ getRange e in parseErrorAt pos $ "Not a valid identifier: " ++ show e stripSingletonRawApp :: Expr -> Expr stripSingletonRawApp (RawApp _ [e]) = stripSingletonRawApp e stripSingletonRawApp e = e isEqual :: Expr -> Maybe (Expr, Expr) isEqual e = case stripSingletonRawApp e of Equal _ a b -> Just (stripSingletonRawApp a, stripSingletonRawApp b) _ -> Nothing maybeNamed :: Expr -> Named_ Expr maybeNamed e = case isEqual e of Just (Ident (QName x), b) -> named (Ranged (getRange x) (nameToRawName x)) b _ -> unnamed e patternSynArgs :: [Either Hiding LamBinding] -> Parser [Arg Name] patternSynArgs = mapM pSynArg where pSynArg Left{} = parseError "Absurd patterns are not allowed in pattern synonyms" pSynArg (Right DomainFull{}) = parseError "Unexpected type signature in pattern synonym argument" pSynArg (Right (DomainFree a x)) | getHiding a `notElem` [Hidden, NotHidden] = parseError $ show (getHiding a) ++ " arguments not allowed to pattern synonyms" | getRelevance a /= Relevant = parseError "Arguments to pattern synonyms must be relevant" | otherwise = return $ Common.Arg a (boundName x) parsePanic s = parseError $ "Internal parser error: " ++ s ++ ". Please report this as a bug." {- RHS or type signature -} data RHSOrTypeSigs = JustRHS RHS | TypeSigsRHS Expr namesOfPattern :: Pattern -> Maybe [(C.ArgInfo, Name)] namesOfPattern (IdentP (QName i)) = Just [(defaultArgInfo, i)] namesOfPattern (WildP r) = Just [(defaultArgInfo, C.noName r)] namesOfPattern (DotP _ (Ident (QName i))) = Just [(setRelevance Irrelevant defaultArgInfo, i)] namesOfPattern (RawAppP _ ps) = fmap concat $ mapM namesOfPattern ps namesOfPattern _ = Nothing funClauseOrTypeSigs :: LHS -> RHSOrTypeSigs -> WhereClause -> Parser [Declaration] funClauseOrTypeSigs lhs (JustRHS rhs) wh = return [FunClause lhs rhs wh] funClauseOrTypeSigs lhs (TypeSigsRHS e) wh | NoWhere <- wh, LHS p [] [] [] <- lhs, Just names <- namesOfPattern p = return $ map (\(x,y) -> TypeSig x y e) names | otherwise = parseError "Illegal function clause or type signature" {-------------------------------------------------------------------------- Tests --------------------------------------------------------------------------} -- | Test suite. tests :: IO Bool tests = runTests "Agda.Syntax.Parser.Parser" [ quickCheck' prop_splitOnDots ] }