Skip to content

Commit

Permalink
[ fix idris-lang#1887 ] Parse _ <- as DoBind instead of DoBinPat
Browse files Browse the repository at this point in the history
  • Loading branch information
gallais committed Aug 31, 2021
1 parent 2428b35 commit cf85798
Show file tree
Hide file tree
Showing 7 changed files with 40 additions and 3 deletions.
6 changes: 3 additions & 3 deletions src/Idris/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -843,13 +843,13 @@ mutual

validPatternVar : Name -> EmptyRule ()
validPatternVar (UN n)
= if lowerFirst n then pure ()
else fail "Not a pattern variable"
= unless (lowerFirst n || n == "_") $
fail "Not a pattern variable"
validPatternVar _ = fail "Not a pattern variable"

doAct : OriginDesc -> IndentInfo -> Rule (List PDo)
doAct fname indents
= do b <- bounds (do n <- bounds name
= do b <- bounds (do n <- bounds (name <|> UN "_" <$ symbol "_")
-- If the name doesn't begin with a lower case letter, we should
-- treat this as a pattern, so fail
validPatternVar n.val
Expand Down
5 changes: 5 additions & 0 deletions src/Libraries/Text/Parser.idr
Original file line number Diff line number Diff line change
Expand Up @@ -226,3 +226,8 @@ public export
when : Bool -> Lazy (Grammar state token False ()) -> Grammar state token False ()
when True y = y
when False y = pure ()

public export
%inline
unless : Bool -> Lazy (Grammar state token False ()) -> Grammar state token False ()
unless = when . not
1 change: 1 addition & 0 deletions tests/Main.idr
Original file line number Diff line number Diff line change
Expand Up @@ -41,6 +41,7 @@ idrisTestsBasic = MkTestPool "Fundamental language features" [] Nothing
"basic046", "basic047", "basic049", "basic050",
"basic051", "basic052", "basic053", "basic054", "basic055",
"basic056", "basic057", "basic058", "basic059", "basic060",
"basic061",
"interpolation001", "interpolation002"]

idrisTestsCoverage : TestPool
Expand Down
16 changes: 16 additions & 0 deletions tests/idris2/basic061/IgnoreDo.idr
Original file line number Diff line number Diff line change
@@ -0,0 +1,16 @@
module IgnoreDo

bound : Maybe () -> Maybe b -> Maybe b
bound m n = do
x <- m
n

ignored : Maybe () -> Maybe b -> Maybe b
ignored m n = do
_ <- m
n

seqd : Maybe () -> Maybe b -> Maybe b
seqd m n = do
m
n
8 changes: 8 additions & 0 deletions tests/idris2/basic061/expected
Original file line number Diff line number Diff line change
@@ -0,0 +1,8 @@
1/1: Building IgnoreDo (IgnoreDo.idr)
IgnoreDo> IgnoreDo.bound : Maybe () -> Maybe b -> Maybe b
bound m n = m >>= (\x => n)
IgnoreDo> IgnoreDo.ignored : Maybe () -> Maybe b -> Maybe b
ignored m n = m >>= (\_ => n)
IgnoreDo> IgnoreDo.seqd : Maybe () -> Maybe b -> Maybe b
seqd m n = m >> Delay n
IgnoreDo> Bye for now!
4 changes: 4 additions & 0 deletions tests/idris2/basic061/input
Original file line number Diff line number Diff line change
@@ -0,0 +1,4 @@
:printdef bound
:printdef ignored
:printdef seqd
:q
3 changes: 3 additions & 0 deletions tests/idris2/basic061/run
Original file line number Diff line number Diff line change
@@ -0,0 +1,3 @@
rm -rf build

$1 --no-color --console-width 0 --no-banner IgnoreDo.idr < input

0 comments on commit cf85798

Please sign in to comment.