Skip to content

Commit 5135ae4

Browse files
committed
[ cleanup ] Small cleanup Core code
1 parent ef9938f commit 5135ae4

File tree

5 files changed

+38
-68
lines changed

5 files changed

+38
-68
lines changed

src/Compiler/RefC/RefC.idr

+6-10
Original file line numberDiff line numberDiff line change
@@ -593,11 +593,9 @@ mutual
593593
decreaseIndentation
594594
pure "} else ") "" alts
595595

596-
case mDef of
597-
Nothing => pure ()
598-
Just body => do
599-
emit emptyFC "} else {"
600-
concaseBody env switchReturnVar "" [] body tailPosition
596+
whenJust mDef $ \body => do
597+
emit emptyFC "} else {"
598+
concaseBody env switchReturnVar "" [] body tailPosition
601599
emit emptyFC "}"
602600
pure switchReturnVar
603601

@@ -626,11 +624,9 @@ mutual
626624
pure "} else ") "" alts
627625
pure ()
628626

629-
case def of
630-
Nothing => pure ()
631-
Just body => do
632-
emit emptyFC "} else {"
633-
concaseBody env switchReturnVar "" [] body tailPosition
627+
whenJust def $ \body => do
628+
emit emptyFC "} else {"
629+
concaseBody env switchReturnVar "" [] body tailPosition
634630
emit emptyFC "}"
635631
pure switchReturnVar
636632

src/Core/Context.idr

+3-4
Original file line numberDiff line numberDiff line change
@@ -1385,10 +1385,9 @@ updateDef n fdef
13851385
= do defs <- get Ctxt
13861386
Just gdef <- lookupCtxtExact n (gamma defs)
13871387
| Nothing => pure ()
1388-
case fdef (definition gdef) of
1389-
Nothing => pure ()
1390-
Just def' => ignore $ addDef n ({ definition := def',
1391-
schemeExpr := Nothing } gdef)
1388+
whenJust (fdef $ definition gdef) $ \def' =>
1389+
ignore $ addDef n $ { definition := def',
1390+
schemeExpr := Nothing } gdef
13921391

13931392
export
13941393
updateTy : {auto c : Ref Ctxt Defs} ->

src/Core/Core.idr

+24-45
Original file line numberDiff line numberDiff line change
@@ -670,14 +670,12 @@ record Core t where
670670
runCore : IO (Either Error t)
671671

672672
export
673-
coreRun : Core a ->
674-
(Error -> IO b) -> (a -> IO b) -> IO b
675-
coreRun (MkCore act) err ok
676-
= either err ok !act
673+
coreRun : Core a -> (Error -> IO b) -> (a -> IO b) -> IO b
674+
coreRun (MkCore act) err ok = either err ok !act
677675

678676
export
679677
coreFail : Error -> Core a
680-
coreFail e = MkCore (pure (Left e))
678+
coreFail = MkCore . pure . Left
681679

682680
export
683681
wrapError : (Error -> Error) -> Core a -> Core a
@@ -687,8 +685,7 @@ wrapError fe (MkCore prog) = MkCore $ mapFst fe <$> prog
687685
export
688686
%inline
689687
coreLift : IO a -> Core a
690-
coreLift op = MkCore (do op' <- op
691-
pure (Right op'))
688+
coreLift = MkCore . map Right
692689

693690
{- Monad, Applicative, Traversable are specialised by hand for Core.
694691
In theory, this shouldn't be necessary, but it turns out that Idris 1 doesn't
@@ -702,11 +699,11 @@ in the next version (i.e., in this project...)! -}
702699
-- Functor (specialised)
703700
export %inline
704701
map : (a -> b) -> Core a -> Core b
705-
map f (MkCore a) = MkCore (map (map f) a)
702+
map f (MkCore a) = MkCore $ map (map f) a
706703

707704
export %inline
708705
(<$>) : (a -> b) -> Core a -> Core b
709-
(<$>) f (MkCore a) = MkCore (map (map f) a)
706+
(<$>) = map
710707

711708
export %inline
712709
(<$) : b -> Core a -> Core b
@@ -718,7 +715,7 @@ export %inline
718715

719716
export %inline
720717
ignore : Core a -> Core ()
721-
ignore = map (const ())
718+
ignore = map $ const ()
722719

723720
-- This would be better if we restrict it to a limited set of IO operations
724721
export
@@ -729,11 +726,9 @@ coreLift_ op = ignore (coreLift op)
729726
-- Monad (specialised)
730727
export %inline
731728
(>>=) : Core a -> (a -> Core b) -> Core b
732-
(>>=) (MkCore act) f
733-
= MkCore (act >>=
734-
\case
735-
Left err => pure $ Left err
736-
Right val => runCore $ f val)
729+
MkCore act >>= f = MkCore $ act >>= \case
730+
Left err => pure $ Left err
731+
Right val => runCore $ f val
737732

738733
export %inline
739734
(>>) : Core () -> Core a -> Core a
@@ -757,26 +752,25 @@ export %inline
757752
-- Applicative (specialised)
758753
export %inline
759754
pure : a -> Core a
760-
pure x = MkCore (pure (pure x))
755+
pure = MkCore . pure . Right
761756

762757
export
763758
(<*>) : Core (a -> b) -> Core a -> Core b
764-
(<*>) (MkCore f) (MkCore a) = MkCore [| f <*> a |]
759+
MkCore f <*> MkCore a = MkCore [| f <*> a |]
765760

766761
export
767762
(*>) : Core a -> Core b -> Core b
768-
(*>) (MkCore a) (MkCore b) = MkCore [| a *> b |]
763+
MkCore a *> MkCore b = MkCore [| a *> b |]
769764

770765
export
771766
(<*) : Core a -> Core b -> Core a
772-
(<*) (MkCore a) (MkCore b) = MkCore [| a <* b |]
767+
MkCore a <* MkCore b = MkCore [| a <* b |]
773768

774769
export %inline
775770
when : Bool -> Lazy (Core ()) -> Core ()
776771
when True f = f
777772
when False f = pure ()
778773

779-
780774
export %inline
781775
unless : Bool -> Lazy (Core ()) -> Core ()
782776
unless = when . not
@@ -811,11 +805,9 @@ interface Catchable m t | m where
811805

812806
export
813807
Catchable Core Error where
814-
catch (MkCore prog) h
815-
= MkCore ( do p' <- prog
816-
case p' of
817-
Left e => let MkCore he = h e in he
818-
Right val => pure (Right val))
808+
catch (MkCore prog) h = MkCore $ prog >>= \case
809+
Left e => runCore (h e)
810+
Right val => pure (Right val)
819811
breakpoint (MkCore prog) = MkCore (pure <$> prog)
820812
throw = coreFail
821813

@@ -827,8 +819,7 @@ foldlC fm a0 = foldl (\ma,b => ma >>= flip fm b) (pure a0)
827819
-- Traversable (specialised)
828820
traverse' : (a -> Core b) -> List a -> List b -> Core (List b)
829821
traverse' f [] acc = pure (reverse acc)
830-
traverse' f (x :: xs) acc
831-
= traverse' f xs (!(f x) :: acc)
822+
traverse' f (x :: xs) acc = traverse' f xs (!(f x) :: acc)
832823

833824
%inline
834825
export
@@ -851,15 +842,12 @@ for = flip traverse
851842

852843
export
853844
traverseList1 : (a -> Core b) -> List1 a -> Core (List1 b)
854-
traverseList1 f xxs
855-
= let x = head xxs
856-
xs = tail xxs in
857-
[| f x ::: traverse f xs |]
845+
traverseList1 f (x ::: xs) = [| f x ::: traverse f xs |]
858846

859847
export
860848
traverseSnocList : (a -> Core b) -> SnocList a -> Core (SnocList b)
861849
traverseSnocList f [<] = pure [<]
862-
traverseSnocList f (as :< a) = (:<) <$> traverseSnocList f as <*> f a
850+
traverseSnocList f (as :< a) = [| traverseSnocList f as :< f a |]
863851

864852
export
865853
traverseVect : (a -> Core b) -> Vect n a -> Core (Vect n b)
@@ -879,9 +867,8 @@ traversePair f (w, a) = (w,) <$> f a
879867
export
880868
traverse_ : (a -> Core b) -> List a -> Core ()
881869
traverse_ f [] = pure ()
882-
traverse_ f (x :: xs)
883-
= Core.do ignore (f x)
884-
traverse_ f xs
870+
traverse_ f (x :: xs) = ignore (f x) >> traverse_ f xs
871+
885872
%inline
886873
export
887874
for_ : List a -> (a -> Core ()) -> Core ()
@@ -890,20 +877,12 @@ for_ = flip traverse_
890877
%inline
891878
export
892879
sequence : List (Core a) -> Core (List a)
893-
sequence (x :: xs)
894-
= do
895-
x' <- x
896-
xs' <- sequence xs
897-
pure (x' :: xs')
880+
sequence (x :: xs) = [| x :: sequence xs |]
898881
sequence [] = pure []
899882

900883
export
901884
traverseList1_ : (a -> Core b) -> List1 a -> Core ()
902-
traverseList1_ f xxs
903-
= do let x = head xxs
904-
let xs = tail xxs
905-
ignore (f x)
906-
traverse_ f xs
885+
traverseList1_ f (x ::: xs) = ignore (f x) >> traverse_ f xs
907886

908887
%inline export
909888
traverseFC : (a -> Core b) -> WithFC a -> Core (WithFC b)

src/Idris/Driver.idr

+2-4
Original file line numberDiff line numberDiff line change
@@ -99,8 +99,7 @@ updateREPLOpts
9999

100100
tryYaffle : List CLOpt -> Core Bool
101101
tryYaffle [] = pure False
102-
tryYaffle (Yaffle f :: _) = do yaffleMain f []
103-
pure True
102+
tryYaffle (Yaffle f :: _) = yaffleMain f [] $> True
104103
tryYaffle (c :: cs) = tryYaffle cs
105104

106105
ignoreMissingIpkg : List CLOpt -> Bool
@@ -110,8 +109,7 @@ ignoreMissingIpkg (c :: cs) = ignoreMissingIpkg cs
110109

111110
tryTTM : List CLOpt -> Core Bool
112111
tryTTM [] = pure False
113-
tryTTM (Metadata f :: _) = do dumpTTM f
114-
pure True
112+
tryTTM (Metadata f :: _) = dumpTTM f $> True
115113
tryTTM (c :: cs) = tryTTM cs
116114

117115

src/Idris/Package.idr

+3-5
Original file line numberDiff line numberDiff line change
@@ -889,11 +889,9 @@ runRepl fname = do
889889
pure (PhysicalIdrSrc modIdent)
890890
) fname
891891
m <- newRef MD (initMetadata origin)
892-
case fname of
893-
Nothing => pure ()
894-
Just fn => do
895-
errs <- loadMainFile fn
896-
displayStartupErrors errs
892+
whenJust fname $ \fn => do
893+
errs <- loadMainFile fn
894+
displayStartupErrors errs
897895
repl {u} {s}
898896

899897
||| If the user did not provide a package file we can look in the working

0 commit comments

Comments
 (0)