Skip to content

Commit

Permalink
[ fix ] case spliting under implicit/auto parameter
Browse files Browse the repository at this point in the history
  • Loading branch information
dunhamsteve authored and gallais committed Jun 13, 2024
1 parent f8adee7 commit 055568b
Show file tree
Hide file tree
Showing 6 changed files with 16 additions and 2 deletions.
2 changes: 1 addition & 1 deletion src/TTImp/Interactive/CaseSplit.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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')
Expand Down
2 changes: 1 addition & 1 deletion tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -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",
Expand Down
5 changes: 5 additions & 0 deletions tests/idris2/misc/params004/Issue2331.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,5 @@
parameters {auto _ : ()}

f : Nat -> Nat
f m = ?a

4 changes: 4 additions & 0 deletions tests/idris2/misc/params004/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
1/1: Building Issue2331 (Issue2331.idr)
Main> f 0 = ?a_0
f (S k) = ?a_1
Main> Bye for now!
2 changes: 2 additions & 0 deletions tests/idris2/misc/params004/input
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
:cs 4 5 m
:q
3 changes: 3 additions & 0 deletions tests/idris2/misc/params004/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
. ../../../testutils.sh

idris2 Issue2331.idr < input

0 comments on commit 055568b

Please sign in to comment.