Skip to content

Commit 51938ae

Browse files
committed
[ fix ] Add eta expand implicits in top-level methods
1 parent 340cf98 commit 51938ae

File tree

2 files changed

+36
-8
lines changed

2 files changed

+36
-8
lines changed

src/Idris/Elab/Interface.idr

Lines changed: 26 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -191,13 +191,22 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
191191
(MkImpTy vfc cn ty_imp))
192192
let conapp = apply (IVar vfc cname)
193193
(map (IBindVar EmptyFC) (map bindName allmeths))
194-
let fnclause = PatClause vfc
195-
(INamedApp vfc
196-
(IVar cn.fc cn.val) -- See #3409
197-
(UN $ Basic "__con")
198-
conapp
199-
)
200-
(IVar EmptyFC (methName sig.name.val))
194+
195+
let lhs = INamedApp vfc
196+
(IVar cn.fc cn.val) -- See #3409
197+
(UN $ Basic "__con")
198+
conapp
199+
let rhs = IVar EmptyFC (methName sig.name.val)
200+
201+
-- EtaExpand implicits on both sides:
202+
-- First, obtain all the implicit names in the prefix of
203+
let piNames = collectImplicitNames sig.type
204+
-- Then apply names for each argument to the lhs
205+
let lhs = namesToRawImp piNames lhs
206+
-- Do the same for the rhs
207+
let rhs = namesToRawImp piNames rhs
208+
209+
let fnclause = PatClause vfc lhs rhs
201210
let fndef = IDef vfc cn.val [fnclause]
202211
pure [tydecl, fndef]
203212
where
@@ -223,6 +232,16 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
223232
methName : Name -> Name
224233
methName n = UN (Basic $ bindName n)
225234

235+
collectImplicitNames : RawImp -> List Name
236+
collectImplicitNames (IPi _ _ Explicit _ _ ty) = []
237+
collectImplicitNames (IPi _ _ _ (Just n) _ ty) = n :: collectImplicitNames ty
238+
collectImplicitNames (IPi _ _ _ Nothing _ ty) = collectImplicitNames ty
239+
collectImplicitNames _ = []
240+
241+
namesToRawImp : List Name -> RawImp -> RawImp
242+
namesToRawImp (nm@(UN{}) :: xs) fn = namesToRawImp xs (INamedApp vfc fn nm (IVar vfc nm))
243+
namesToRawImp _ fn = fn
244+
226245
-- Get the function for chasing a constraint. This is one of the
227246
-- arguments to the record, appearing before the method arguments.
228247
getConstraintHint : {vars : _} ->
Lines changed: 10 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1,2 +1,11 @@
1+
interface Ok1 where
2+
fOk : Type -> Type -> Type
3+
14
interface Fail1 where
2-
f : Type -> {_ : Type} -> Type
5+
fFail : Type -> {_ : Type} -> Type
6+
7+
interface Ok2 where
8+
gOk : (x : Type) -> Type
9+
10+
interface Fail2 where
11+
gFail : {x : Type} -> Type

0 commit comments

Comments
 (0)