Skip to content

Commit 3f15823

Browse files
committed
[ cleanup ] Cleaning up the use of applicative syntax
1 parent ab17241 commit 3f15823

19 files changed

+64
-83
lines changed

src/Compiler/Common.idr

+2-2
Original file line numberDiff line numberDiff line change
@@ -354,7 +354,7 @@ getCompileDataWith exports doLazyAnnots phase_in tm_in
354354
ldefs ++ concat lifted_in
355355

356356
anf <- if phase >= ANF
357-
then logTime 2 "Get ANF" $ traverse (\ (n, d) => pure (n, !(toANF d))) lifted
357+
then logTime 2 "Get ANF" $ traverse (traversePair toANF) lifted
358358
else pure []
359359
vmcode <- if phase >= VMCode
360360
then logTime 2 "Get VM Code" $ pure (allDefs anf)
@@ -440,7 +440,7 @@ getIncCompileData doLazyAnnots phase
440440
else pure []
441441
let lifted = concat lifted_in
442442
anf <- if phase >= ANF
443-
then logTime 2 "Get ANF" $ traverse (\ (n, d) => pure (n, !(toANF d))) lifted
443+
then logTime 2 "Get ANF" $ traverse (traversePair toANF) lifted
444444
else pure []
445445
vmcode <- if phase >= VMCode
446446
then logTime 2 "Get VM Code" $ pure (allDefs anf)

src/Compiler/CompileExpr.idr

+3-3
Original file line numberDiff line numberDiff line change
@@ -354,9 +354,9 @@ toCExpTm n (As _ _ _ p) = toCExpTm n p
354354
-- TODO: Either make sure 'Delayed' is always Rig0, or add to typecase
355355
toCExpTm n (TDelayed fc _ _) = pure $ CErased fc
356356
toCExpTm n (TDelay fc lr _ arg)
357-
= pure (CDelay fc lr !(toCExp n arg))
357+
= CDelay fc lr <$> toCExp n arg
358358
toCExpTm n (TForce fc lr arg)
359-
= pure (CForce fc lr !(toCExp n arg))
359+
= CForce fc lr <$> toCExp n arg
360360
toCExpTm n (PrimVal fc $ PrT c) = pure $ CCon fc (UN $ Basic $ show c) TYCON Nothing [] -- Primitive type constant
361361
toCExpTm n (PrimVal fc c) = pure $ CPrimVal fc c -- Non-type constant
362362
toCExpTm n (Erased fc _) = pure $ CErased fc
@@ -604,7 +604,7 @@ getNArgs defs (NS _ (UN $ Basic "Unit")) [] = pure NUnit
604604
getNArgs defs (NS _ (UN $ Basic "Struct")) [n, args]
605605
= do NPrimVal _ (Str n') <- evalClosure defs n
606606
| nf => throw (GenericMsg (getLoc nf) "Unknown name for struct")
607-
pure (Struct n' !(getFieldArgs defs args))
607+
Struct n' <$> getFieldArgs defs args
608608
getNArgs defs n args = pure $ User n args
609609

610610
nfToCFType : {auto c : Ref Ctxt Defs} ->

src/Compiler/Scheme/Common.idr

+6-8
Original file line numberDiff line numberDiff line change
@@ -414,7 +414,7 @@ parameters (constants : SortedSet Name,
414414
!(showAlts n alts)
415415
schCaseTree i sc alts def
416416
= do tcode <- schExp (i + 1) sc
417-
defc <- maybe (pure Nothing) (\v => pure (Just !(schExp i v))) def
417+
defc <- traverseOpt (schExp i) def
418418
let n = getScrutineeTemp i
419419
if var sc
420420
then pure $ "(case (vector-ref " ++ tcode ++ " 0) "
@@ -446,8 +446,7 @@ parameters (constants : SortedSet Name,
446446
schListCase i sc alts def
447447
= do tcode <- schExp (i + 1) sc
448448
let n = getScrutineeTemp i
449-
defc <- maybe (pure Nothing)
450-
(\v => pure (Just !(schExp (i + 1) v))) def
449+
defc <- traverseOpt (schExp (i + 1)) def
451450
nil <- getNilCode alts
452451
if var sc
453452
then do cons <- getConsCode tcode alts
@@ -473,7 +472,7 @@ parameters (constants : SortedSet Name,
473472
getNilCode : List NamedConAlt -> Core (Maybe Builder)
474473
getNilCode [] = pure Nothing
475474
getNilCode (MkNConAlt _ NIL _ _ sc :: _)
476-
= pure (Just !(schExp (i + 1) sc))
475+
= Just <$> schExp (i + 1) sc
477476
getNilCode (_ :: xs) = getNilCode xs
478477

479478
getConsCode : Builder -> List NamedConAlt -> Core (Maybe Builder)
@@ -496,8 +495,7 @@ parameters (constants : SortedSet Name,
496495
schMaybeCase i sc alts def
497496
= do tcode <- schExp (i + 1) sc
498497
let n = getScrutineeTemp i
499-
defc <- maybe (pure Nothing)
500-
(\v => pure (Just !(schExp (i + 1) v))) def
498+
defc <- traverseOpt (schExp (i + 1)) def
501499
nothing <- getNothingCode alts
502500
if var sc
503501
then do just <- getJustCode tcode alts
@@ -523,7 +521,7 @@ parameters (constants : SortedSet Name,
523521
getNothingCode : List NamedConAlt -> Core (Maybe Builder)
524522
getNothingCode [] = pure Nothing
525523
getNothingCode (MkNConAlt _ NOTHING _ _ sc :: _)
526-
= pure (Just !(schExp (i + 1) sc))
524+
= Just <$> schExp (i + 1) sc
527525
getNothingCode (_ :: xs) = getNothingCode xs
528526

529527
getJustCode : Builder -> List NamedConAlt -> Core (Maybe Builder)
@@ -623,7 +621,7 @@ parameters (constants : SortedSet Name,
623621
= pure $ !(schConstAlt (i + 1) n alt) ++ " " ++
624622
!(showConstAlts n alts)
625623
schExp i (NmConstCase fc sc alts def)
626-
= do defc <- maybe (pure Nothing) (\v => pure (Just !(schExp i v))) def
624+
= do defc <- traverseOpt (schExp i) def
627625
tcode <- schExp (i + 1) sc
628626
let n = getScrutineeTemp i
629627
if var sc

src/Core/Binary.idr

+2-4
Original file line numberDiff line numberDiff line change
@@ -345,10 +345,8 @@ addGlobalDef modns filens asm (n, def)
345345
codedentry <- lookupContextEntry n (gamma defs)
346346
-- Don't update the coded entry because some names might not be
347347
-- resolved yet
348-
entry <- maybe (pure Nothing)
349-
(\ p => do x <- decode (gamma defs) (fst p) False (snd p)
350-
pure (Just x))
351-
codedentry
348+
entry <- traverseOpt (\p => decode (gamma defs) (fst p) False (snd p))
349+
codedentry
352350
unless (completeDef entry) $
353351
ignore $ addContextEntry filens n def
354352

src/Core/Case/CaseBuilder.idr

+2-2
Original file line numberDiff line numberDiff line change
@@ -1080,10 +1080,10 @@ mutual
10801080
Core (Maybe (CaseTree vars))
10811081
mixture fc fn phase (ConClauses cs rest) err
10821082
= do fallthrough <- mixture fc fn phase rest err
1083-
pure (Just !(conRule fc fn phase cs fallthrough))
1083+
Just <$> conRule fc fn phase cs fallthrough
10841084
mixture fc fn phase (VarClauses vs rest) err
10851085
= do fallthrough <- mixture fc fn phase rest err
1086-
pure (Just !(varRule fc fn phase vs fallthrough))
1086+
Just <$> varRule fc fn phase vs fallthrough
10871087
mixture fc fn {a} {todo} phase NoClauses err
10881088
= pure err
10891089

src/Core/Core.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -894,7 +894,7 @@ namespace PiInfo
894894
traverse f Explicit = pure Explicit
895895
traverse f Implicit = pure Implicit
896896
traverse f AutoImplicit = pure AutoImplicit
897-
traverse f (DefImplicit t) = pure (DefImplicit !(f t))
897+
traverse f (DefImplicit t) = DefImplicit <$> f t
898898

899899
namespace Binder
900900
export

src/Core/LinearCheck.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -472,7 +472,7 @@ mutual
472472
Core (List ArgUsage)
473473
getArgUsage topfc rig ty pats
474474
= do us <- traverse (getPUsage ty) pats
475-
pure (map snd !(combine us))
475+
map snd <$> combine us
476476
where
477477
getCaseUsage : {vs : _} ->
478478
Term ns -> Env Term vs -> List (Term vs) ->

src/Core/Normalise.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -154,7 +154,7 @@ etaContract tm = do
154154
export
155155
getValArity : Defs -> Env Term vars -> NF vars -> Core Nat
156156
getValArity defs env (NBind fc x (Pi _ _ _ _) sc)
157-
= pure (S !(getValArity defs env !(sc defs (toClosure defaultOpts env (Erased fc Placeholder)))))
157+
= S <$> getValArity defs env !(sc defs (toClosure defaultOpts env (Erased fc Placeholder)))
158158
getValArity defs env val = pure 0
159159

160160
export

src/Idris/Desugar.idr

+1-3
Original file line numberDiff line numberDiff line change
@@ -1155,9 +1155,7 @@ mutual
11551155
let isb = map (\ (info, r, n, p, tm) => (info, r, n, p, doBind bnames tm)) is'
11561156
let consb = map (\(n, tm) => (n, doBind bnames tm)) cons'
11571157

1158-
body' <- maybe (pure Nothing)
1159-
(\b => do b' <- traverse (desugarDecl ps) b
1160-
pure (Just (concat b'))) body
1158+
body' <- traverseOpt (map concat . traverse (desugarDecl ps)) body
11611159
-- calculate the name of the implementation, if it's not explicitly
11621160
-- given.
11631161
let impname = maybe (mkImplName fc tn paramsb) id impln

src/Idris/ModTree.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -137,7 +137,7 @@ getBuildMods loc done fname
137137
dm <- newRef DoneMod empty
138138
o <- newRef BuildOrder []
139139
mkBuildMods {d=dm} {o} t
140-
pure (reverse !(get BuildOrder))
140+
reverse <$> get BuildOrder
141141

142142
checkTotalReq : {auto c : Ref Ctxt Defs} ->
143143
String -> String -> TotalReq -> Core Bool

src/Idris/Package.idr

+1-1
Original file line numberDiff line numberDiff line change
@@ -196,7 +196,7 @@ field fname
196196
depends
197197
= do name <- packageName
198198
bs <- sepBy andop bound
199-
pure (MkDepends name !(mkBound (concat bs) anyBounds))
199+
MkDepends name <$> mkBound (concat bs) anyBounds
200200

201201
strField : (FC -> String -> DescField) -> String -> Rule DescField
202202
strField fieldConstructor fieldName

src/Idris/Parser.idr

+6-7
Original file line numberDiff line numberDiff line change
@@ -71,9 +71,9 @@ decoratedNamespacedSymbol fname smb =
7171

7272
parens : {b : _} -> OriginDesc -> BRule b a -> Rule a
7373
parens fname p
74-
= pure id <* decoratedSymbol fname "("
75-
<*> p
76-
<* decoratedSymbol fname ")"
74+
= id <$ decoratedSymbol fname "("
75+
<*> p
76+
<* decoratedSymbol fname ")"
7777

7878
decoratedDataTypeName : OriginDesc -> Rule Name
7979
decoratedDataTypeName fname = decorate fname Typ dataTypeName
@@ -1822,10 +1822,9 @@ typedArg fname indents
18221822
pure $ map (\(c, n, tm) => (n.val, c, Explicit, tm)) params
18231823
<|> do decoratedSymbol fname "{"
18241824
commit
1825-
info <-
1826-
(pure AutoImplicit <* decoratedKeyword fname "auto"
1827-
<|> (decoratedKeyword fname "default" *> DefImplicit <$> simpleExpr fname indents)
1828-
<|> pure Implicit)
1825+
info <- AutoImplicit <$ decoratedKeyword fname "auto"
1826+
<|> decoratedKeyword fname "default" *> DefImplicit <$> simpleExpr fname indents
1827+
<|> pure Implicit
18291828
params <- pibindListName fname indents
18301829
decoratedSymbol fname "}"
18311830
pure $ map (\(c, n, tm) => (n.val, c, info, tm)) params

src/Idris/REPL.idr

+2-2
Original file line numberDiff line numberDiff line change
@@ -1130,10 +1130,10 @@ processCatch cmd
11301130
)
11311131

11321132
parseEmptyCmd : EmptyRule (Maybe REPLCmd)
1133-
parseEmptyCmd = eoi *> (pure Nothing)
1133+
parseEmptyCmd = eoi $> Nothing
11341134

11351135
parseCmd : EmptyRule (Maybe REPLCmd)
1136-
parseCmd = do c <- command; eoi; pure $ Just c
1136+
parseCmd = Just <$> command <* eoi
11371137

11381138
export
11391139
parseRepl : String -> Either Error (Maybe REPLCmd)

src/Idris/Resugar.idr

+20-23
Original file line numberDiff line numberDiff line change
@@ -369,7 +369,7 @@ mutual
369369
toPTerm p (IAlternative fc _ _) = pure (PImplicit fc)
370370
toPTerm p (IRewrite fc rule tm)
371371
= pure (PRewrite fc !(toPTerm startPrec rule)
372-
!(toPTerm startPrec tm))
372+
!(toPTerm startPrec tm))
373373
toPTerm p (ICoerced fc tm) = toPTerm p tm
374374
toPTerm p (IPrimVal fc c) = pure (PPrimVal fc c)
375375
toPTerm p (IHole fc str) = pure (PHole fc False str)
@@ -378,19 +378,18 @@ mutual
378378
= let nm = UN (Basic v) in
379379
pure (PRef fc (MkKindedName (Just Bound) nm nm))
380380
toPTerm p (IBindHere fc _ tm) = toPTerm p tm
381-
toPTerm p (IAs fc nameFC _ n pat) = pure (PAs fc nameFC n !(toPTerm argPrec pat))
382-
toPTerm p (IMustUnify fc r pat) = pure (PDotted fc !(toPTerm argPrec pat))
381+
toPTerm p (IAs fc nameFC _ n pat) = PAs fc nameFC n <$> toPTerm argPrec pat
382+
toPTerm p (IMustUnify fc r pat) = PDotted fc <$> toPTerm argPrec pat
383383

384-
toPTerm p (IDelayed fc r ty) = pure (PDelayed fc r !(toPTerm argPrec ty))
385-
toPTerm p (IDelay fc tm) = pure (PDelay fc !(toPTerm argPrec tm))
386-
toPTerm p (IForce fc tm) = pure (PForce fc !(toPTerm argPrec tm))
387-
toPTerm p (IQuote fc tm) = pure (PQuote fc !(toPTerm argPrec tm))
384+
toPTerm p (IDelayed fc r ty) = PDelayed fc r <$> toPTerm argPrec ty
385+
toPTerm p (IDelay fc tm) = PDelay fc <$> toPTerm argPrec tm
386+
toPTerm p (IForce fc tm) = PForce fc <$> toPTerm argPrec tm
387+
toPTerm p (IQuote fc tm) = PQuote fc <$> toPTerm argPrec tm
388388
toPTerm p (IQuoteName fc n) = pure (PQuoteName fc n)
389389
toPTerm p (IQuoteDecl fc ds)
390-
= do ds' <- traverse toPDecl ds
391-
pure $ PQuoteDecl fc (catMaybes ds')
392-
toPTerm p (IUnquote fc tm) = pure (PUnquote fc !(toPTerm argPrec tm))
393-
toPTerm p (IRunElab fc _ tm) = pure (PRunElab fc !(toPTerm argPrec tm))
390+
= PQuoteDecl fc . catMaybes <$> traverse toPDecl ds
391+
toPTerm p (IUnquote fc tm) = PUnquote fc <$> toPTerm argPrec tm
392+
toPTerm p (IRunElab fc _ tm) = PRunElab fc <$> toPTerm argPrec tm
394393

395394
toPTerm p (IUnifyLog fc _ tm) = toPTerm p tm
396395
toPTerm p (Implicit fc True) = pure (PImplicit fc)
@@ -468,22 +467,22 @@ mutual
468467
flags
469468
!(traverse toPClause cs)
470469
toPClause (ImpossibleClause fc lhs)
471-
= pure (MkImpossible fc !(toPTerm startPrec lhs))
470+
= MkImpossible fc <$> toPTerm startPrec lhs
472471

473472
toPTypeDecl : {auto c : Ref Ctxt Defs} ->
474473
{auto s : Ref Syn SyntaxInfo} ->
475474
ImpTy' KindedName -> Core (PTypeDecl' KindedName)
476475
toPTypeDecl (MkImpTy fc n ty)
477-
= pure (MkFCVal fc $ MkPTy (pure ("", n)) "" !(toPTerm startPrec ty))
476+
= MkFCVal fc . MkPTy (pure ("", n)) "" <$> toPTerm startPrec ty
478477

479478
toPData : {auto c : Ref Ctxt Defs} ->
480479
{auto s : Ref Syn SyntaxInfo} ->
481480
ImpData' KindedName -> Core (PDataDecl' KindedName)
482481
toPData (MkImpData fc n ty opts cs)
483482
= pure (MkPData fc n !(traverseOpt (toPTerm startPrec) ty) opts
484-
!(traverse toPTypeDecl cs))
483+
!(traverse toPTypeDecl cs))
485484
toPData (MkImpLater fc n ty)
486-
= pure (MkPLater fc n !(toPTerm startPrec ty))
485+
= MkPLater fc n <$> toPTerm startPrec ty
487486

488487
toPField : {auto c : Ref Ctxt Defs} ->
489488
{auto s : Ref Syn SyntaxInfo} ->
@@ -528,11 +527,11 @@ mutual
528527
ImpDecl' KindedName -> Core (Maybe (PDecl' KindedName))
529528
toPDecl (IClaim (MkFCVal fc $ MkIClaimData rig vis opts ty))
530529
= do opts' <- traverse toPFnOpt opts
531-
pure (Just (PClaim (MkFCVal fc $ MkPClaim rig vis opts' !(toPTypeDecl ty))))
530+
Just . PClaim . MkFCVal fc . MkPClaim rig vis opts' <$> toPTypeDecl ty
532531
toPDecl (IData fc vis mbtot d)
533-
= pure (Just (PData fc "" vis mbtot !(toPData d)))
532+
= Just . PData fc "" vis mbtot <$> toPData d
534533
toPDecl (IDef fc n cs)
535-
= pure (Just (PDef $ MkFCVal fc !(traverse toPClause cs)))
534+
= Just . PDef . MkFCVal fc <$> traverse toPClause cs
536535
toPDecl (IParameters fc ps ds)
537536
= do ds' <- traverse toPDecl ds
538537
pure (Just (PParameters fc
@@ -545,17 +544,15 @@ mutual
545544
= do (n, ps, opts, con, fs) <- toPRecord r
546545
pure (Just (PRecord fc "" vis mbtot (MkPRecord n ps opts con fs)))
547546
toPDecl (IFail fc msg ds)
548-
= do ds' <- traverse toPDecl ds
549-
pure (Just (PFail fc msg (catMaybes ds')))
547+
= Just . PFail fc msg . catMaybes <$> traverse toPDecl ds
550548
toPDecl (INamespace fc ns ds)
551-
= do ds' <- traverse toPDecl ds
552-
pure (Just (PNamespace fc ns (catMaybes ds')))
549+
= Just . PNamespace fc ns . catMaybes <$> traverse toPDecl ds
553550
toPDecl (ITransform fc n lhs rhs)
554551
= pure (Just (PTransform fc (show n)
555552
!(toPTerm startPrec lhs)
556553
!(toPTerm startPrec rhs)))
557554
toPDecl (IRunElabDecl fc tm)
558-
= pure (Just (PRunElabDecl fc !(toPTerm startPrec tm)))
555+
= Just . PRunElabDecl fc <$> toPTerm startPrec tm
559556
toPDecl (IPragma _ _ _) = pure Nothing
560557
toPDecl (ILog _) = pure Nothing
561558
toPDecl (IBuiltin fc type name) = pure $ Just $ PBuiltin fc type name

src/TTImp/Elab/App.idr

+6-8
Original file line numberDiff line numberDiff line change
@@ -543,9 +543,8 @@ mutual
543543
logNF "elab" 10 ("Full function type") env
544544
(NBind fc x (Pi fc argRig Explicit aty) sc)
545545
logC "elab" 10
546-
(do ety <- maybe (pure Nothing)
547-
(\t => pure (Just !(toFullNames!(getTerm t))))
548-
expty
546+
(do ety <- traverseOpt (\t => toFullNames !(getTerm t))
547+
expty
549548
pure ("Overall expected type: " ++ show ety))
550549
res <- check argRig ({ topLevel := False } elabinfo)
551550
nest env arg (Just (glueClosure defs env aty))
@@ -829,11 +828,10 @@ checkApp rig elabinfo nest env fc (IVar fc' n) expargs autoargs namedargs exp
829828
logC "elab" 10
830829
(do defs <- get Ctxt
831830
fnty <- quote defs env nty
832-
exptyt <- maybe (pure Nothing)
833-
(\t => do ety <- getTerm t
834-
etynf <- normaliseHoles defs env ety
835-
pure (Just !(toFullNames etynf)))
836-
exp
831+
exptyt <- traverseOpt (\t => do ety <- getTerm t
832+
etynf <- normaliseHoles defs env ety
833+
toFullNames etynf)
834+
exp
837835
pure ("Checking application of " ++ show !(getFullName n) ++
838836
" (" ++ show n ++ ")" ++
839837
" to " ++ show expargs ++ "\n\tFunction type " ++

src/TTImp/Elab/Lazy.idr

+2-5
Original file line numberDiff line numberDiff line change
@@ -91,11 +91,8 @@ checkForce : {vars : _} ->
9191
Core (Term vars, Glued vars)
9292
checkForce rig elabinfo nest env fc tm exp
9393
= do defs <- get Ctxt
94-
expf <- maybe (pure Nothing)
95-
(\gty => do tynf <- getNF gty
96-
pure (Just (glueBack defs env
97-
(NDelayed fc LUnknown tynf))))
98-
exp
94+
expf <- traverseOpt (map (glueBack defs env . NDelayed fc LUnknown) . getNF)
95+
exp
9996
(tm', gty) <- check rig elabinfo nest env tm expf
10097
tynf <- getNF gty
10198
case tynf of

src/TTImp/Elab/Record.idr

+2-3
Original file line numberDiff line numberDiff line change
@@ -86,9 +86,8 @@ findFieldsAndTypeArgs : {auto c : Ref Ctxt Defs} ->
8686
Defs -> Name ->
8787
Core $ Maybe (List (String, Maybe Name, Maybe Name), SortedSet Name)
8888
findFieldsAndTypeArgs defs con
89-
= case !(lookupTyExact con (gamma defs)) of
90-
Just t => pure (Just !(getExpNames empty [] !(nf defs [] t)))
91-
_ => pure Nothing
89+
= traverseOpt (\t => getExpNames empty [] !(nf defs [] t))
90+
!(lookupTyExact con (gamma defs))
9291
where
9392
getExpNames : SortedSet Name ->
9493
List (String, Maybe Name, Maybe Name) ->

src/TTImp/Interactive/CaseSplit.idr

+2-2
Original file line numberDiff line numberDiff line change
@@ -250,7 +250,7 @@ getUpdates : Defs -> RawImp -> RawImp -> Core (List (Name, RawImp))
250250
getUpdates defs orig updated
251251
= do u <- newRef UPD (MkUpdates [] [])
252252
findUpdates defs orig updated
253-
pure (updates !(get UPD))
253+
updates <$> get UPD
254254

255255
mkCase : {auto c : Ref Ctxt Defs} ->
256256
{auto u : Ref UST UState} ->
@@ -285,7 +285,7 @@ mkCase {c} {u} fn orig lhs_raw
285285
log "interaction.casesplit" 3 $ "Original LHS: " ++ show orig
286286
log "interaction.casesplit" 3 $ "New LHS: " ++ show lhs'
287287

288-
pure (Valid lhs' !(getUpdates defs orig lhs')))
288+
Valid lhs' <$> getUpdates defs orig lhs')
289289
(\err =>
290290
do put Ctxt defs
291291
put UST ust

0 commit comments

Comments
 (0)