Skip to content

Commit

Permalink
Merge pull request #2027 from david-christiansen/issue/2019
Browse files Browse the repository at this point in the history
Better documentation display for instances
  • Loading branch information
david-christiansen committed Mar 20, 2015
2 parents be1569e + beda221 commit abe0f32
Show file tree
Hide file tree
Showing 7 changed files with 112 additions and 21 deletions.
5 changes: 4 additions & 1 deletion idris.cabal
Original file line number Diff line number Diff line change
Expand Up @@ -631,7 +631,10 @@ Extra-source-files:
test/docs002/input
test/docs002/*.idr
test/docs002/expected

test/docs003/run
test/docs003/input
test/docs003/*.idr
test/docs003/expected


source-repository head
Expand Down
2 changes: 1 addition & 1 deletion src/Idris/Delaborate.hs
Original file line number Diff line number Diff line change
@@ -1,6 +1,6 @@
{-# LANGUAGE PatternGuards #-}
{-# OPTIONS_GHC -fwarn-incomplete-patterns #-}
module Idris.Delaborate (bugaddr, delab, delab', delabMV, delabTy, delabTy', fancifyAnnots, pprintDelab, pprintDelabTy, pprintErr) where
module Idris.Delaborate (annName, bugaddr, delab, delab', delabMV, delabTy, delabTy', fancifyAnnots, pprintDelab, pprintDelabTy, pprintErr) where

-- Convert core TT back into high level syntax, primarily for display
-- purposes.
Expand Down
74 changes: 55 additions & 19 deletions src/Idris/Docs.hs
Original file line number Diff line number Diff line change
@@ -1,4 +1,5 @@
{-# LANGUAGE DeriveFunctor, PatternGuards #-}
{-# LANGUAGE DeriveFunctor, PatternGuards, MultiWayIf #-}
{-# OPTIONS_GHC -fwarn-incomplete-patterns #-}
module Idris.Docs (pprintDocs, getDocs, pprintConstDocs, FunDoc, FunDoc'(..), Docs, Docs'(..)) where

import Idris.AbsSyntax
Expand All @@ -10,6 +11,8 @@ import Idris.Docstrings (Docstring, emptyDocstring, noDocs, nullDocstring, rende

import Util.Pretty

import Control.Arrow (first)

import Data.Maybe
import Data.List
import qualified Data.Text as T
Expand All @@ -32,9 +35,10 @@ data Docs' d = FunDoc (FunDoc' d)
| ClassDoc Name d -- class docs
[FunDoc' d] -- method docs
[(Name, Maybe d)] -- parameters and their docstrings
[(PTerm, (d, [(Name, d)]))] -- instances
[(Maybe Name, PTerm, (d, [(Name, d)]))] -- instances: name for named instances, the constraint term, the docs
[PTerm] -- subclasses
[PTerm] -- superclasses
| NamedInstanceDoc Name (FunDoc' d) -- name is class
| ModDoc [String] -- Module name
d
deriving Functor
Expand Down Expand Up @@ -102,7 +106,11 @@ pprintDocs ist (ClassDoc n doc meths params instances subclasses superclasses)
<$>
nest 4 (text "Instances:" <$>
vsep (if null instances then [text "<no instances>"]
else map pprintInstance instances))
else map pprintInstance normalInstances))
<>
(if null namedInstances then empty
else line <$> nest 4 (text "Named instances:" <$>
vsep (map pprintInstance namedInstances)))
<>
(if null subclasses then empty
else line <$> nest 4 (text "Subclasses:" <$>
Expand All @@ -113,14 +121,17 @@ pprintDocs ist (ClassDoc n doc meths params instances subclasses superclasses)
vsep (map dumpInstance superclasses)))
where
params' = zip pNames (repeat False)

(normalInstances, namedInstances) = partition (\(n, _, _) -> not $ isJust n)
instances

pNames = map fst params

ppo = ppOptionIst ist
infixes = idris_infixes ist

pprintInstance (term, (doc, argDocs)) =
nest 4 (dumpInstance term <>
pprintInstance (mname, term, (doc, argDocs)) =
nest 4 (iname mname <> dumpInstance term <>
(if nullDocstring doc
then empty
else line <>
Expand All @@ -133,6 +144,10 @@ pprintDocs ist (ClassDoc n doc meths params instances subclasses superclasses)
then empty
else line <> vsep (map (prettyInstanceParam (map fst argDocs)) argDocs))


iname Nothing = empty
iname (Just n) = annName n <+> colon <> space

prettyInstanceParam params (name, doc) =
if nullDocstring doc
then empty
Expand Down Expand Up @@ -170,6 +185,9 @@ pprintDocs ist (ClassDoc n doc meths params instances subclasses superclasses)
then vsep (map (\(nm,md) -> prettyName True False params' nm <+> maybe empty (showDoc ist) md) params)
else hsep (punctuate comma (map (prettyName True False params' . fst) params))

pprintDocs ist (NamedInstanceDoc _cls doc)
= nest 4 (text "Named instance:" <$> pprintFD ist doc)

pprintDocs ist (ModDoc mod docs)
= nest 4 $ text "Module" <+> text (concat (intersperse "." mod)) <> colon <$>
renderDocstring (renderDocTerm (pprintDelab ist) (normaliseAll (tt_ctxt ist) [])) docs
Expand All @@ -190,13 +208,23 @@ getDocs n@(NS n' ns) w | n' == modDocName
" do not exist! This shouldn't have happened and is a bug."
getDocs n w
= do i <- getIState
docs <- case lookupCtxt n (idris_classes i) of
[ci] -> docClass n ci
_ -> case lookupCtxt n (idris_datatypes i) of
[ti] -> docData n ti
_ -> do fd <- docFun n
return (FunDoc fd)
docs <- if | Just ci <- lookupCtxtExact n (idris_classes i)
-> docClass n ci
| Just ti <- lookupCtxtExact n (idris_datatypes i)
-> docData n ti
| Just class_ <- classNameForInst i n
-> do fd <- docFun n
return $ NamedInstanceDoc class_ fd
| otherwise
-> do fd <- docFun n
return (FunDoc fd)
return $ fmap (howMuch w) docs
where classNameForInst :: IState -> Name -> Maybe Name
classNameForInst ist n =
listToMaybe [ cn
| (cn, ci) <- toAlist (idris_classes ist)
, n `elem` class_instances ci
]

docData :: Name -> TypeInfo -> Idris Docs
docData n ti
Expand All @@ -209,19 +237,26 @@ docClass n ci
= do i <- getIState
let docStrings = listToMaybe $ lookupCtxt n $ idris_docstrings i
docstr = maybe emptyDocstring fst docStrings
params = map (\pn -> (pn, docStrings >>= (lookup pn . snd))) (class_params ci)
instanceDocs = map (fromMaybe (emptyDocstring, []) .
listToMaybe .
flip lookupCtxt (idris_docstrings i))
(class_instances ci)
instances = zip (map (delabTy i) (class_instances ci)) instanceDocs
(subclasses, instances') = partition (isSubclass . fst) instances
params = map (\pn -> (pn, docStrings >>= (lookup pn . snd)))
(class_params ci)
docsForInstance inst = fromMaybe (emptyDocstring, []) .
flip lookupCtxtExact (idris_docstrings i) $
inst
instances = map (\inst -> (namedInst inst,
delabTy i inst,
docsForInstance inst))
(nub (class_instances ci))
(subclasses, instances') = partition (isSubclass . (\(_,tm,_) -> tm)) instances
superclasses = catMaybes $ map getDInst (class_default_superclasses ci)
mdocs <- mapM (docFun . fst) (class_methods ci)
return $ ClassDoc
n docstr mdocs params
instances' (map fst subclasses) superclasses
instances' (map (\(_,tm,_) -> tm) subclasses) superclasses
where
namedInst (NS n ns) = fmap (flip NS ns) (namedInst n)
namedInst n@(UN _) = Just n
namedInst _ = Nothing

getDInst (PInstance _ _ _ _ _ _ _ t _ _) = Just t
getDInst _ = Nothing

Expand Down Expand Up @@ -250,6 +285,7 @@ docFun n
where funName :: Name -> String
funName (UN n) = str n
funName (NS n _) = funName n
funName n = show n

getPArgNames :: PTerm -> [(Name, Docstring DocTerm)] -> [(Name, PTerm, Plicity, Maybe (Docstring DocTerm))]
getPArgNames (PPi plicity name ty body) ds =
Expand Down
9 changes: 9 additions & 0 deletions test/docs003/docs003.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,9 @@
module docs003

instance [mine] Functor List where
map m [] = []
map m (x :: xs) = m x :: map m xs

||| More functors!
instance [another] Functor List where
map f xs = map @{mine} f xs
37 changes: 37 additions & 0 deletions test/docs003/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,37 @@
Type checking ./docs003.idr
Type class Functor
Functors

Parameters:
f -- the action of the functor on objects

Methods:
map : Functor f => (m : a -> b) -> f a -> f b
The action of the functor on morphisms

Instances:
Functor List
Functor Stream
Functor Provider
Functor Binder
Functor PrimIO
Functor (IO' ffi)
Functor Maybe
Functor (Either e)

Named instances:
docs003.mine : Functor List
docs003.another : Functor List
More functors!

Subclasses:
Traversable f
Applicative f
Named instance:
mine : Functor List


Named instance:
another : Functor List
More functors!

3 changes: 3 additions & 0 deletions test/docs003/input
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
:doc Functor
:doc mine
:doc another
3 changes: 3 additions & 0 deletions test/docs003/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
#!/usr/bin/env bash
idris --consolewidth 80 --quiet --nocolor docs003.idr < input
rm *.ibc

0 comments on commit abe0f32

Please sign in to comment.