Skip to content

Commit

Permalink
[ fix ] Fix pattern match issue with function application in Refl (#3027
Browse files Browse the repository at this point in the history
)
  • Loading branch information
dunhamsteve authored Aug 4, 2023
1 parent f0f776c commit bde1a66
Show file tree
Hide file tree
Showing 7 changed files with 28 additions and 4 deletions.
2 changes: 2 additions & 0 deletions CHANGELOG.md
Original file line number Diff line number Diff line change
Expand Up @@ -99,6 +99,8 @@
* Fixed a bug that caused holes to appear unexpectedly during quotation of
dependent pairs.

* Fixed a bug that caused `f` to sometimes be replaced by `fx` after matching `fx = f x`.

### Library changes

#### Prelude
Expand Down
8 changes: 5 additions & 3 deletions src/Core/Case/CaseBuilder.idr
Original file line number Diff line number Diff line change
Expand Up @@ -1089,7 +1089,7 @@ mutual

export
mkPat : {auto c : Ref Ctxt Defs} -> List Pat -> ClosedTerm -> ClosedTerm -> Core Pat
mkPat args orig (Ref fc Bound n) = pure $ PLoc fc n
mkPat [] orig (Ref fc Bound n) = pure $ PLoc fc n
mkPat args orig (Ref fc (DataCon t a) n) = pure $ PCon fc n t a args
mkPat args orig (Ref fc (TyCon t a) n) = pure $ PTyCon fc n a args
mkPat args orig (Ref fc Func n)
Expand All @@ -1107,8 +1107,10 @@ mkPat args orig (Ref fc Func n)
"Unmatchable function: " ++ show n
pure $ PUnmatchable (getLoc orig) orig
mkPat args orig (Bind fc x (Pi _ _ _ s) t)
= let t' = subst (Erased fc Placeholder) t in
pure $ PArrow fc x !(mkPat [] s s) !(mkPat [] t' t')
-- For (b:Nat) -> b, the codomain looks like b [__], but we want `b` as the pattern
= case subst (Erased fc Placeholder) t of
App _ t'@(Ref fc Bound n) (Erased _ _) => pure $ PArrow fc x !(mkPat [] s s) !(mkPat [] t' t')
t' => pure $ PArrow fc x !(mkPat [] s s) !(mkPat [] t' t')
mkPat args orig (App fc fn arg)
= do parg <- mkPat [] arg arg
mkPat (parg :: args) orig fn
Expand Down
2 changes: 1 addition & 1 deletion tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -70,7 +70,7 @@ idrisTestsTermination = MkTestPool "Termination checking" [] Nothing
idrisTestsCasetree : TestPool
idrisTestsCasetree = MkTestPool "Case tree building" [] Nothing
-- Case tree building
["casetree001", "casetree002", "casetree003"]
["casetree001", "casetree002", "casetree003", "casetree004"]

idrisTestsWarning : TestPool
idrisTestsWarning = MkTestPool "Warnings" [] Nothing
Expand Down
11 changes: 11 additions & 0 deletions tests/idris2/casetree004/LocalArgs.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,11 @@
import Data.Vect

bar : (f : List a -> Nat) -> (xs : List a) -> Nat
bar f xs = f xs

-- Idris was putting fx@f for the first implicit (instead of fx@(f xs))
foo : (f : List a -> Nat) -> (xs : List a) -> (0 _ : fx = f xs) -> Nat
foo f xs Refl = bar f xs

blah : (xs : List a) -> Vect (foo List.length xs Refl) ()
blah xs = replicate (length xs) ()
3 changes: 3 additions & 0 deletions tests/idris2/casetree004/PiMatch.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
idMatch : Type -> Bool
idMatch ((x : Type) -> x) = True
idMatch _ = False
2 changes: 2 additions & 0 deletions tests/idris2/casetree004/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,2 @@
1/1: Building LocalArgs (LocalArgs.idr)
1/1: Building PiMatch (PiMatch.idr)
4 changes: 4 additions & 0 deletions tests/idris2/casetree004/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
rm -rf build

$1 --no-color --console-width 0 --no-banner LocalArgs.idr --check
$1 --no-color --console-width 0 --no-banner PiMatch.idr --check

0 comments on commit bde1a66

Please sign in to comment.