Skip to content

Commit

Permalink
[ new ] function options for case blocks (#3062)
Browse files Browse the repository at this point in the history
  • Loading branch information
gallais authored Sep 1, 2023
1 parent af7ba2f commit c52b029
Show file tree
Hide file tree
Showing 30 changed files with 363 additions and 299 deletions.
1 change: 1 addition & 0 deletions idris2api.ipkg
Original file line number Diff line number Diff line change
Expand Up @@ -257,6 +257,7 @@ modules =
TTImp.ProcessDecls,
TTImp.ProcessDecls.Totality,
TTImp.ProcessDef,
TTImp.ProcessFnOpt,
TTImp.ProcessParams,
TTImp.ProcessRecord,
TTImp.ProcessRunElab,
Expand Down
18 changes: 9 additions & 9 deletions libs/base/Language/Reflection/TTImp.idr
Original file line number Diff line number Diff line change
Expand Up @@ -39,7 +39,7 @@ mutual
ILet : FC -> (lhsFC : FC) -> Count -> Name ->
(nTy : TTImp) -> (nVal : TTImp) ->
(scope : TTImp) -> TTImp
ICase : FC -> TTImp -> (ty : TTImp) ->
ICase : FC -> List FnOpt -> TTImp -> (ty : TTImp) ->
List Clause -> TTImp
ILocal : FC -> List Decl -> TTImp -> TTImp
IUpdate : FC -> List IFieldUpdate -> TTImp -> TTImp
Expand Down Expand Up @@ -203,7 +203,7 @@ getFC (IVar fc _) = fc
getFC (IPi fc _ _ _ _ _) = fc
getFC (ILam fc _ _ _ _ _) = fc
getFC (ILet fc _ _ _ _ _ _) = fc
getFC (ICase fc _ _ _) = fc
getFC (ICase fc _ _ _ _) = fc
getFC (ILocal fc _ _) = fc
getFC (IUpdate fc _ _) = fc
getFC (IApp fc _ _) = fc
Expand Down Expand Up @@ -236,7 +236,7 @@ mapTopmostFC fcf $ IVar fc a = IVar (fcf fc) a
mapTopmostFC fcf $ IPi fc a b c d e = IPi (fcf fc) a b c d e
mapTopmostFC fcf $ ILam fc a b c d e = ILam (fcf fc) a b c d e
mapTopmostFC fcf $ ILet fc a b c d e f = ILet (fcf fc) a b c d e f
mapTopmostFC fcf $ ICase fc a b c = ICase (fcf fc) a b c
mapTopmostFC fcf $ ICase fc opts a b c = ICase (fcf fc) opts a b c
mapTopmostFC fcf $ ILocal fc a b = ILocal (fcf fc) a b
mapTopmostFC fcf $ IUpdate fc a b = IUpdate (fcf fc) a b
mapTopmostFC fcf $ IApp fc a b = IApp (fcf fc) a b
Expand Down Expand Up @@ -403,7 +403,7 @@ Eq TTImp where
c == c' && (assert_total $ i == i') && n == n' && a == a' && r == r'
ILet _ _ c n ty val s == ILet _ _ c' n' ty' val' s' =
c == c' && n == n' && ty == ty' && val == val' && s == s'
ICase _ t ty cs == ICase _ t' ty' cs'
ICase _ _ t ty cs == ICase _ _ t' ty' cs'
= t == t' && ty == ty' && (assert_total $ cs == cs')
ILocal _ ds e == ILocal _ ds' e' =
(assert_total $ ds == ds') && e == e'
Expand Down Expand Up @@ -573,7 +573,7 @@ mutual
showPrec d (ILet fc lhsFC rig nm nTy nVal scope)
= showParens (d > Open) $
"let \{showCount rig (show nm)} : \{show nTy} = \{show nVal} in \{show scope}"
showPrec d (ICase fc s ty xs)
showPrec d (ICase fc _ s ty xs)
= showParens (d > Open) $
unwords $ [ "case", show s ] ++ typeFor ty ++ [ "of", "{"
, joinBy "; " (assert_total $ map (showClause InCase) xs)
Expand Down Expand Up @@ -783,8 +783,8 @@ parameters (f : TTImp -> TTImp)
= f $ ILam fc rig (mapPiInfo pinfo) x (mapTTImp argTy) (mapTTImp lamTy)
mapTTImp (ILet fc lhsFC rig n nTy nVal scope)
= f $ ILet fc lhsFC rig n (mapTTImp nTy) (mapTTImp nVal) (mapTTImp scope)
mapTTImp (ICase fc t ty cls)
= f $ ICase fc (mapTTImp t) (mapTTImp ty) (assert_total $ map mapClause cls)
mapTTImp (ICase fc opts t ty cls)
= f $ ICase fc opts (mapTTImp t) (mapTTImp ty) (assert_total $ map mapClause cls)
mapTTImp (ILocal fc xs t)
= f $ ILocal fc (assert_total $ map mapDecl xs) (mapTTImp t)
mapTTImp (IUpdate fc upds t) = f $ IUpdate fc (assert_total map mapIFieldUpdate upds) (mapTTImp t)
Expand Down Expand Up @@ -909,8 +909,8 @@ parameters {0 m : Type -> Type} {auto mon : Monad m} (f : TTImp -> m TTImp)
= f =<< ILam fc rig <$> mapMPiInfo pinfo <*> pure x <*> mapMTTImp argTy <*> mapMTTImp lamTy
mapMTTImp (ILet fc lhsFC rig n nTy nVal scope)
= f =<< ILet fc lhsFC rig n <$> mapMTTImp nTy <*> mapMTTImp nVal <*> mapMTTImp scope
mapMTTImp (ICase fc t ty cls)
= f =<< ICase fc <$> mapMTTImp t <*> mapMTTImp ty <*> assert_total (traverse mapMClause cls)
mapMTTImp (ICase fc opts t ty cls)
= f =<< ICase fc opts <$> mapMTTImp t <*> mapMTTImp ty <*> assert_total (traverse mapMClause cls)
mapMTTImp (ILocal fc xs t)
= f =<< ILocal fc <$> assert_total (traverse mapMDecl xs) <*> mapMTTImp t
mapMTTImp (IUpdate fc upds t)
Expand Down
2 changes: 1 addition & 1 deletion src/Core/Binary.idr
Original file line number Diff line number Diff line change
Expand Up @@ -29,7 +29,7 @@ import public Libraries.Utils.Binary
||| version number if you're changing the version more than once in the same day.
export
ttcVersion : Int
ttcVersion = 20230414 * 100 + 0
ttcVersion = 20230829 * 100 + 0

export
checkTTCVersion : String -> Int -> Int -> Core ()
Expand Down
22 changes: 12 additions & 10 deletions src/Idris/Desugar.idr
Original file line number Diff line number Diff line change
Expand Up @@ -245,7 +245,7 @@ mutual
!(desugar AnyExpr (n :: ps) scope)
else pure $ ILam EmptyFC rig !(traverse (desugar AnyExpr ps) p)
(Just (MN "lamc" 0)) !(desugarB AnyExpr ps argTy) $
ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
ICase fc [] (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
[snd !(desugarClause ps True (MkPatClause fc pat scope []))]
desugarB side ps (PLam fc rig p (PRef _ n@(MN _ _)) argTy scope)
= pure $ ILam fc rig !(traverse (desugar AnyExpr ps) p)
Expand All @@ -258,21 +258,23 @@ mutual
desugarB side ps (PLam fc rig p pat argTy scope)
= pure $ ILam EmptyFC rig !(traverse (desugar AnyExpr ps) p)
(Just (MN "lamc" 0)) !(desugarB AnyExpr ps argTy) $
ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
ICase fc [] (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
[snd !(desugarClause ps True (MkPatClause fc pat scope []))]
desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal scope [])
= do whenJust (isConcreteFC prefFC) $ \nfc =>
addSemanticDecorations [(nfc, Bound, Just n)]
pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal)
!(desugar side (n :: ps) scope)
desugarB side ps (PLet fc rig pat nTy nVal scope alts)
= pure $ ICase fc !(desugarB side ps nVal) !(desugarB side ps nTy)
= pure $ ICase fc [] !(desugarB side ps nVal) !(desugarB side ps nTy)
!(traverse (map snd . desugarClause ps True)
(MkPatClause fc pat scope [] :: alts))
desugarB side ps (PCase fc x xs)
= pure $ ICase fc !(desugarB side ps x)
(Implicit (virtualiseFC fc) False)
!(traverse (map snd . desugarClause ps True) xs)
desugarB side ps (PCase fc opts scr cls)
= do opts <- traverse (desugarFnOpt ps) opts
scr <- desugarB side ps scr
let scrty = Implicit (virtualiseFC fc) False
cls <- traverse (map snd . desugarClause ps True) cls
pure $ ICase fc opts scr scrty cls
desugarB side ps (PLocal fc xs scope)
= let ps' = definedIn xs ++ ps in
pure $ ILocal fc (concat !(traverse (desugarDecl ps') xs))
Expand Down Expand Up @@ -454,7 +456,7 @@ mutual
IVar fc (UN $ Basic "MkUnit")]
desugarB side ps (PIfThenElse fc x t e)
= let fc = virtualiseFC fc in
pure $ ICase fc !(desugarB side ps x) (IVar fc (UN $ Basic "Bool"))
pure $ ICase fc [] !(desugarB side ps x) (IVar fc (UN $ Basic "Bool"))
[PatClause fc (IVar fc (UN $ Basic "True")) !(desugar side ps t),
PatClause fc (IVar fc (UN $ Basic "False")) !(desugar side ps e)]
desugarB side ps (PComprehension fc ret conds) = do
Expand Down Expand Up @@ -674,7 +676,7 @@ mutual
pure $ bindFun fc ns exp'
$ ILam EmptyFC top Explicit (Just (MN "_" 0))
(Implicit fc False)
(ICase fc (IVar patFC (MN "_" 0))
(ICase fc [] (IVar patFC (MN "_" 0))
(Implicit fc False)
(PatClause fcOriginal bpat rest'
:: alts'))
Expand All @@ -700,7 +702,7 @@ mutual
bd <- get Bang
let fc = virtualiseFC fc
pure $ bindBangs (bangNames bd) ns $
ICase fc tm' ty'
ICase fc [] tm' ty'
(PatClause fc bpat rest'
:: alts')
expandDo side ps topfc ns (DoLetLocal fc decls :: rest)
Expand Down
120 changes: 61 additions & 59 deletions src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -211,6 +211,17 @@ debugString fname = do
DebugLine => "\{show (startLine di.bounds)}"
DebugCol => "\{show (startCol di.bounds)}"

totalityOpt : OriginDesc -> Rule TotalReq
totalityOpt fname
= (decoratedKeyword fname "partial" $> PartialOK)
<|> (decoratedKeyword fname "total" $> Total)
<|> (decoratedKeyword fname "covering" $> CoveringOnly)

fnOpt : OriginDesc -> Rule PFnOpt
fnOpt fname
= do x <- totalityOpt fname
pure $ IFnOpt (Totality x)

mutual
appExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
appExpr q fname indents
Expand Down Expand Up @@ -753,7 +764,7 @@ mutual
fcCase = virtualiseFC $ boundToFC fname endCase
n = MN "lcase" 0 in
PLam fcCase top Explicit (PRef fcCase n) (PInfer fcCase) $
PCase (virtualiseFC fc) (PRef fcCase n) b.val)
PCase (virtualiseFC fc) [] (PRef fcCase n) b.val)

letBlock : OriginDesc -> IndentInfo -> Rule (WithBounds (Either LetBinder LetDecl))
letBlock fname indents = bounds (letBinder <||> letDecl) where
Expand Down Expand Up @@ -782,13 +793,14 @@ mutual

case_ : OriginDesc -> IndentInfo -> Rule PTerm
case_ fname indents
= do b <- bounds (do decoratedKeyword fname "case"
= do opts <- many (fnDirectOpt fname)
b <- bounds (do decoratedKeyword fname "case"
scr <- expr pdef fname indents
mustWork (commitKeyword fname indents "of")
alts <- block (caseAlt fname)
pure (scr, alts))
(scr, alts) <- pure b.val
pure (PCase (virtualiseFC $ boundToFC fname b) scr alts)
pure (PCase (virtualiseFC $ boundToFC fname b) opts scr alts)


caseAlt : OriginDesc -> IndentInfo -> Rule PClause
Expand Down Expand Up @@ -1055,6 +1067,52 @@ mutual
$ acc :< (line <>> [StrLiteral fc str])
<>< map (\str => [StrLiteral fc str]) (init strs)

fnDirectOpt : OriginDesc -> Rule PFnOpt
fnDirectOpt fname
= do decoratedPragma fname "hint"
pure $ IFnOpt (Hint True)
<|> do decoratedPragma fname "globalhint"
pure $ IFnOpt (GlobalHint False)
<|> do decoratedPragma fname "defaulthint"
pure $ IFnOpt (GlobalHint True)
<|> do decoratedPragma fname "inline"
commit
pure $ IFnOpt Inline
<|> do decoratedPragma fname "unsafe"
commit
pure $ IFnOpt Unsafe
<|> do decoratedPragma fname "noinline"
commit
pure $ IFnOpt NoInline
<|> do decoratedPragma fname "deprecate"
commit
pure $ IFnOpt Deprecate
<|> do decoratedPragma fname "tcinline"
commit
pure $ IFnOpt TCInline
<|> do decoratedPragma fname "extern"
pure $ IFnOpt ExternFn
<|> do decoratedPragma fname "macro"
pure $ IFnOpt Macro
<|> do decoratedPragma fname "spec"
ns <- sepBy (decoratedSymbol fname ",") name
pure $ IFnOpt (SpecArgs ns)
<|> do decoratedPragma fname "foreign"
cs <- block (expr pdef fname)
pure $ PForeign cs
<|> do (decoratedPragma fname "export"
<|> withWarning noMangleWarning
(decoratedPragma fname "nomangle"))
cs <- block (expr pdef fname)
pure $ PForeignExport cs
where
noMangleWarning : String
noMangleWarning = """
DEPRECATED: "%nomangle".
Use "%export" instead
"""


visOption : OriginDesc -> Rule Visibility
visOption fname
= (decoratedKeyword fname "public" *> decoratedKeyword fname "export" $> Public)
Expand Down Expand Up @@ -1268,12 +1326,6 @@ dataDeclBody fname indents
(col, n) <- pure b.val
simpleData fname b n indents <|> gadtData fname col b n indents

totalityOpt : OriginDesc -> Rule TotalReq
totalityOpt fname
= (decoratedKeyword fname "partial" $> PartialOK)
<|> (decoratedKeyword fname "total" $> Total)
<|> (decoratedKeyword fname "covering" $> CoveringOnly)

-- a data declaration can have a visibility and an optional totality (#1404)
dataVisOpt : OriginDesc -> EmptyRule (Visibility, Maybe TotalReq)
dataVisOpt fname
Expand Down Expand Up @@ -1507,56 +1559,6 @@ usingDecls fname indents
(us, ds) <- pure b.val
pure (PUsing (boundToFC fname b) us (collectDefs (concat ds)))

fnOpt : OriginDesc -> Rule PFnOpt
fnOpt fname
= do x <- totalityOpt fname
pure $ IFnOpt (Totality x)

fnDirectOpt : OriginDesc -> Rule PFnOpt
fnDirectOpt fname
= do decoratedPragma fname "hint"
pure $ IFnOpt (Hint True)
<|> do decoratedPragma fname "globalhint"
pure $ IFnOpt (GlobalHint False)
<|> do decoratedPragma fname "defaulthint"
pure $ IFnOpt (GlobalHint True)
<|> do decoratedPragma fname "inline"
commit
pure $ IFnOpt Inline
<|> do decoratedPragma fname "unsafe"
commit
pure $ IFnOpt Unsafe
<|> do decoratedPragma fname "noinline"
commit
pure $ IFnOpt NoInline
<|> do decoratedPragma fname "deprecate"
commit
pure $ IFnOpt Deprecate
<|> do decoratedPragma fname "tcinline"
commit
pure $ IFnOpt TCInline
<|> do decoratedPragma fname "extern"
pure $ IFnOpt ExternFn
<|> do decoratedPragma fname "macro"
pure $ IFnOpt Macro
<|> do decoratedPragma fname "spec"
ns <- sepBy (decoratedSymbol fname ",") name
pure $ IFnOpt (SpecArgs ns)
<|> do decoratedPragma fname "foreign"
cs <- block (expr pdef fname)
pure $ PForeign cs
<|> do (decoratedPragma fname "export"
<|> withWarning noMangleWarning
(decoratedPragma fname "nomangle"))
cs <- block (expr pdef fname)
pure $ PForeignExport cs
where
noMangleWarning : String
noMangleWarning = """
DEPRECATED: "%nomangle".
Use "%export" instead
"""

builtinDecl : OriginDesc -> IndentInfo -> Rule PDecl
builtinDecl fname indents
= do b <- bounds (do decoratedPragma fname "builtin"
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Pretty.idr
Original file line number Diff line number Diff line change
Expand Up @@ -281,7 +281,7 @@ mutual
; _ => hardline <+> indent 4 (vsep (prettyAlt <$> alts)) <+> hardline <+> in_
}
<+> group (align $ hang 2 $ pretty sc)
prettyPrec d (PCase _ tm cs) =
prettyPrec d (PCase _ _ tm cs) =
parenthesise (d > startPrec) $
case_ <++> pretty tm <++> of_ <+> nest 2 (
let punctuation = lcurly :: (semi <$ fromMaybe [] (tail' [1..length cs])) in
Expand Down
14 changes: 8 additions & 6 deletions src/Idris/Resugar.idr
Original file line number Diff line number Diff line change
Expand Up @@ -315,20 +315,22 @@ mutual
sc' <- toPTerm startPrec sc
let var = PRef lhsFC (MkKindedName (Just Bound) n n)
bracket p startPrec (PLet fc rig var ty' val' sc' [])
toPTerm p (ICase fc sc scty [PatClause _ lhs rhs])
toPTerm p (ICase fc _ sc scty [PatClause _ lhs rhs])
= do sc' <- toPTerm startPrec sc
lhs' <- toPTerm startPrec lhs
rhs' <- toPTerm startPrec rhs
bracket p startPrec
(PLet fc top lhs' (PImplicit fc) sc' rhs' [])
toPTerm p (ICase fc sc scty alts)
= do sc' <- toPTerm startPrec sc
toPTerm p (ICase fc opts sc scty alts)
= do opts' <- traverse toPFnOpt opts
sc' <- toPTerm startPrec sc
alts' <- traverse toPClause alts
bracket p startPrec (mkIf (PCase fc sc' alts'))
bracket p startPrec (mkIf (PCase fc opts' sc' alts'))
where
mkIf : IPTerm -> IPTerm
mkIf tm@(PCase loc sc [MkPatClause _ (PRef _ tval) t [],
MkPatClause _ (PRef _ fval) f []])
mkIf tm@(PCase loc opts sc
[ MkPatClause _ (PRef _ tval) t []
, MkPatClause _ (PRef _ fval) f []])
= if dropNS (rawName tval) == UN (Basic "True")
&& dropNS (rawName fval) == UN (Basic "False")
then PIfThenElse loc sc t f
Expand Down
6 changes: 3 additions & 3 deletions src/Idris/Syntax.idr
Original file line number Diff line number Diff line change
Expand Up @@ -107,7 +107,7 @@ mutual
PLet : FC -> RigCount -> (pat : PTerm' nm) ->
(nTy : PTerm' nm) -> (nVal : PTerm' nm) -> (scope : PTerm' nm) ->
(alts : List (PClause' nm)) -> PTerm' nm
PCase : FC -> PTerm' nm -> List (PClause' nm) -> PTerm' nm
PCase : FC -> List (PFnOpt' nm) -> PTerm' nm -> List (PClause' nm) -> PTerm' nm
PLocal : FC -> List (PDecl' nm) -> (scope : PTerm' nm) -> PTerm' nm
PUpdate : FC -> List (PFieldUpdate' nm) -> PTerm' nm
PApp : FC -> PTerm' nm -> PTerm' nm -> PTerm' nm
Expand Down Expand Up @@ -178,7 +178,7 @@ mutual
getPTermLoc (PPi fc _ _ _ _ _) = fc
getPTermLoc (PLam fc _ _ _ _ _) = fc
getPTermLoc (PLet fc _ _ _ _ _ _) = fc
getPTermLoc (PCase fc _ _) = fc
getPTermLoc (PCase fc _ _ _) = fc
getPTermLoc (PLocal fc _ _) = fc
getPTermLoc (PUpdate fc _) = fc
getPTermLoc (PApp fc _ _) = fc
Expand Down Expand Up @@ -755,7 +755,7 @@ parameters {0 nm : Type} (toName : nm -> Name)
= "let " ++ showCount rig ++ showPTermPrec d n ++ " : " ++ showPTermPrec d ty ++ " = "
++ showPTermPrec d val ++ concatMap showAlt alts ++
" in " ++ showPTermPrec d sc
showPTermPrec _ (PCase _ tm cs)
showPTermPrec _ (PCase _ _ tm cs)
= "case " ++ showPTerm tm ++ " of { " ++
showSep " ; " (map showCase cs) ++ " }"
where
Expand Down
Loading

0 comments on commit c52b029

Please sign in to comment.