Idris2/tests/idris2/linear012/linholes.idr
Edwin Brady ff7d3a0246 Use precise inference for hole types
That is, don't generalise multiplicities, because we need the hole type
to be precise wrt multiplicities. Resolves #189
2020-06-28 22:16:15 +01:00

16 lines
411 B
Idris

infixr 6 -*
(-*) : Type -> Type -> Type
(-*) a b = (1 _ : a) -> b
-- Test that the types of holes are not overly generalised wrt the
-- multiplicities in the function type
foo1 : (((1 _ : a) -> (1 _ : b) -> (1 _ : c) -> ()) -> ()) -> ()
foo1 f = f ?foo1h
foo2 : (((1 _ : a) -> (1 _ : b) -> (1 _ : c) -> ()) -> ()) -> ()
foo2 f = f (id ?foo2h)
foo3 : ((a -* b -* c -* ()) -> ()) -> ()
foo3 f = f (id ?foo3h)