Fix case split in parameter blocks

We need to make sure variables are bound as PVar, in the end, otherwise
the case split machinery doesn't know how to handle them.
This commit is contained in:
Edwin Brady 2021-07-10 19:13:27 +01:00
parent 26cdfc7830
commit 4ca8caeb13
7 changed files with 20 additions and 3 deletions

View File

@ -48,11 +48,16 @@ defined {vars = x :: xs} n (b :: env)
pure (MkIsDefined rig (Later prf))
Just Refl => Just (MkIsDefined (multiplicity b) First)
-- Bind additional pattern variables in an LHS, when checking an LHS in an
-- outer environment
export
bindEnv : {vars : _} -> FC -> Env Term vars -> (tm : Term vars) -> ClosedTerm
bindEnv loc [] tm = tm
bindEnv loc (b :: env) tm
= bindEnv loc env (Bind loc _ b tm)
= bindEnv loc env (Bind loc _ (PVar (binderLoc b)
(multiplicity b)
Explicit
(binderType b)) tm)
revOnto : (xs, vs : _) -> reverseOnto xs vs = reverse vs ++ xs
revOnto xs [] = Refl

View File

@ -22,7 +22,7 @@ extend : {extvs : _} ->
NestedNames extvs ->
Term extvs ->
(vars' ** (SubVars vs vars', Env Term vars', NestedNames vars'))
extend env p nest (Bind fc n b@(Pi _ _ _ _) sc)
extend env p nest (Bind _ n b@(Pi fc c pi ty) sc)
= extend (b :: env) (DropCons p) (weaken nest) sc
extend env p nest tm = (_ ** (p, env, nest))

View File

@ -174,7 +174,7 @@ idrisTests = MkTestPool "Misc" [] Nothing
-- Namespace blocks
"namespace001",
-- Parameters blocks
"params001","params002",
"params001", "params002", "params003",
-- Packages and ipkg files
"pkg001", "pkg002", "pkg003", "pkg004", "pkg005", "pkg006", "pkg007",
"pkg008", "pkg009", "pkg010",

View File

@ -0,0 +1,3 @@
parameters (X : Nat)
foo : Bool -> ()
foo n = ?help

View File

@ -0,0 +1,4 @@
1/1: Building casesplit (casesplit.idr)
Main> foo False = ?help_1
foo True = ?help_2
Main> Bye for now!

View File

@ -0,0 +1,2 @@
:cs 3 7 n
:q

3
tests/idris2/params003/run Executable file
View File

@ -0,0 +1,3 @@
$1 --no-color --console-width 0 --no-banner casesplit.idr < input
rm -rf build