[ fix ] case spliting under implicit/auto parameter

This commit is contained in:
Steve Dunham 2024-06-12 20:16:48 -07:00 committed by G. Allais
parent f8adee7059
commit 055568be28
6 changed files with 16 additions and 2 deletions

View File

@ -197,7 +197,7 @@ newLHS : {auto c : Ref Ctxt Defs} ->
newLHS fc envlen allvars var con tm
= do let (f, args) = getFnArgs tm []
let keep = map (const (Explicit fc (Implicit fc True)))
(take envlen args)
(mapMaybe isExplicit $ take envlen args)
let ups = drop envlen args
ups' <- traverse (update allvars var con) ups
pure $ apply f (keep ++ ups')

View File

@ -132,7 +132,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
-- Namespace blocks
"namespace001", "namespace002", "namespace003", "namespace004", "namespace005",
-- Parameters blocks
"params001", "params002", "params003",
"params001", "params002", "params003", "params004",
-- Larger programs arising from real usage. Typically things with
-- interesting interactions between features
"real001", "real002",

View File

@ -0,0 +1,5 @@
parameters {auto _ : ()}
f : Nat -> Nat
f m = ?a

View File

@ -0,0 +1,4 @@
1/1: Building Issue2331 (Issue2331.idr)
Main> f 0 = ?a_0
f (S k) = ?a_1
Main> Bye for now!

View File

@ -0,0 +1,2 @@
:cs 4 5 m
:q

View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
idris2 Issue2331.idr < input