@@ -191,13 +191,22 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
191
191
(MkImpTy vfc cn ty_imp))
192
192
let conapp = apply (IVar vfc cname)
193
193
(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
201
210
let fndef = IDef vfc cn. val [fnclause]
202
211
pure [tydecl, fndef]
203
212
where
@@ -223,6 +232,16 @@ getMethToplevel {vars} env vis iname cname constraints allmeths params sig
223
232
methName : Name -> Name
224
233
methName n = UN (Basic $ bindName n)
225
234
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
+
226
245
-- Get the function for chasing a constraint. This is one of the
227
246
-- arguments to the record, appearing before the method arguments.
228
247
getConstraintHint : {vars : _} ->
0 commit comments