Skip to content

Commit

Permalink
Fix type class name propagation issue
Browse files Browse the repository at this point in the history
We weren't shadowing names correctly when a class was parameterised on a
single variable and one of the methods bound the same name. This almost
never happens, but leads to very odd error messages when it does.

Fixes #2026
  • Loading branch information
edwinb committed Mar 20, 2015
1 parent 63e3ef9 commit be1569e
Show file tree
Hide file tree
Showing 5 changed files with 22 additions and 1 deletion.
3 changes: 3 additions & 0 deletions idris.cabal
Original file line number Diff line number Diff line change
Expand Up @@ -298,6 +298,9 @@ Extra-source-files:
test/reg059/run
test/reg059/*.idr
test/reg059/expected
test/reg060/run
test/reg060/*.idr
test/reg060/expected

test/basic001/run
test/basic001/*.idr
Expand Down
5 changes: 4 additions & 1 deletion src/Idris/Elab/Instance.hs
Original file line number Diff line number Diff line change
Expand Up @@ -88,6 +88,7 @@ elabInstance info syn doc argDocs what fc cs n ps t expn ds = do
-- where block
wparams <- mapM (\p -> case p of
PApp _ _ args -> getWParams (map getTm args)
a@(PRef fc f) -> getWParams [a]
_ -> return []) ps
let pnames = map pname (concat (nub wparams))
let superclassInstances = map (substInstance ips pnames) (class_default_superclasses ci)
Expand Down Expand Up @@ -242,7 +243,9 @@ elabInstance info syn doc argDocs what fc cs n ps t expn ds = do
decorate ns iname (UN n) = NS (SN (MethodN (UN n))) ns
decorate ns iname (NS (UN n) s) = NS (SN (MethodN (UN n))) ns

mkTyDecl (n, op, t, _) = PTy emptyDocstring [] syn fc op n t
mkTyDecl (n, op, t, _)
= PTy emptyDocstring [] syn fc op n
(mkUniqueNames [] t)

conbind :: [(Name, PTerm)] -> PTerm -> PTerm
conbind ((c,ty) : ns) x = PPi constraint c ty (conbind ns x)
Expand Down
Empty file added test/reg060/expected
Empty file.
12 changes: 12 additions & 0 deletions test/reg060/reg060.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@

class MyFunctor (f : Type -> Type) where
mymap : (m : a -> b) -> f a -> f b

data Foo x y = Bar y

instance MyFunctor (Foo m) where
mymap m x = ?wibble

instance [foo] Functor m => MyFunctor m where
mymap m x = map m x

3 changes: 3 additions & 0 deletions test/reg060/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
#!/usr/bin/env bash
idris $@ reg060.idr --check
rm -f *.ibc

0 comments on commit be1569e

Please sign in to comment.