[ breaking ] remove parsing of dangling binders (#1711)

* [ breaking ] remove parsing of dangling binders

It used to be the case that

```
ID : Type -> Type
ID a = a

test : ID (a : Type) -> a -> a
test = \ a, x => x
```

and

```
head : List $ a -> Maybe a
head [] = Nothing
head (x :: _) = Just x
```

were accepted but these are now rejected because:

* `ID (a : Type) -> a -> a` is parsed as `(ID (a : Type)) -> a -> a`
* `List $ a -> Maybe a` is parsed as `List (a -> Maybe a)`

Similarly if you want to use a lambda / rewrite / let expression as
part of the last argument of an application, the use of `$` or parens
is now mandatory.

This should hopefully allow us to make progress on #1703
This commit is contained in:
G. Allais 2021-08-10 19:24:32 +01:00 committed by GitHub
parent c9f9e1e667
commit 21c6f4fb79
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
65 changed files with 514 additions and 341 deletions

View File

@ -94,22 +94,22 @@ public export
MonadError e m => MonadError e (ReaderT r m) where
throwError = lift . throwError
catchError (MkReaderT m) f =
MkReaderT \e => catchError (m e) (runReaderT e . f)
MkReaderT $ \e => catchError (m e) (runReaderT e . f)
public export
MonadError e m => MonadError e (StateT r m) where
throwError = lift . throwError
catchError (ST m) f =
ST \s => catchError (m s) (runStateT s . f)
ST $ \s => catchError (m s) (runStateT s . f)
public export
MonadError e m => MonadError e (RWST r w s m) where
throwError = lift . throwError
catchError (MkRWST m) f =
MkRWST \r,w,s => catchError (m r w s) (\e => unRWST (f e) r w s)
MkRWST $ \r,w,s => catchError (m r w s) (\e => unRWST (f e) r w s)
public export
MonadError e m => MonadError e (WriterT w m) where
throwError = lift . throwError
catchError (MkWriterT m) f =
MkWriterT \w => catchError (m w) (\e => unWriterT (f e) w)
MkWriterT $ \w => catchError (m w) (\e => unWriterT (f e) w)

View File

@ -43,14 +43,14 @@ execRWST m r s = (\(_,s',w) => (s',w)) <$> runRWST m r s
public export %inline
mapRWST : (Functor n, Monoid w, Semigroup w')
=> (m (a, s, w) -> n (b, s, w')) -> RWST r w s m a -> RWST r w' s n b
mapRWST f m = MkRWST \r,s,w =>
mapRWST f m = MkRWST $ \r,s,w =>
(\(a,s',w') => (a,s',w <+> w')) <$> f (runRWST m r s)
||| `withRWST f m` executes action `m` with an initial environment
||| and state modified by applying `f`.
public export %inline
withRWST : (r' -> s -> (r, s)) -> RWST r w s m a -> RWST r' w s m a
withRWST f m = MkRWST \r,s => uncurry (unRWST m) (f r s)
withRWST f m = MkRWST $ \r,s => uncurry (unRWST m) (f r s)
--------------------------------------------------------------------------------
-- RWS Functions
@ -70,8 +70,8 @@ runRWS m r s = runIdentity (runRWST m r s)
||| Construct an RWS computation from a function. (The inverse of `runRWS`.)
public export %inline
rws : Semigroup w => (r -> s -> (a, s, w)) -> RWS r w s a
rws f = MkRWST \r,s,w => let (a, s', w') = f r s
in Id (a, s', w <+> w')
rws f = MkRWST $ \r,s,w => let (a, s', w') = f r s
in Id (a, s', w <+> w')
||| Evaluate a computation with the given initial state and environment,
||| returning the final value and output, discarding the final state.
@ -92,7 +92,7 @@ execRWS m r s = let (_,s1,w) = runRWS m r s
public export %inline
mapRWS : (Monoid w, Semigroup w')
=> ((a, s, w) -> (b, s, w')) -> RWS r w s a -> RWS r w' s b
mapRWS f = mapRWST \(Id p) => Id (f p)
mapRWS f = mapRWST $ \(Id p) => Id (f p)
||| `withRWS f m` executes action `m` with an initial environment
||| and state modified by applying `f`.
@ -106,29 +106,29 @@ withRWS = withRWST
public export %inline
Functor m => Functor (RWST r w s m) where
map f m = MkRWST \r,s,w => (\(a,s',w') => (f a,s',w')) <$> unRWST m r s w
map f m = MkRWST $ \r,s,w => (\(a,s',w') => (f a,s',w')) <$> unRWST m r s w
public export %inline
Monad m => Applicative (RWST r w s m) where
pure a = MkRWST \_,s,w => pure (a,s,w)
pure a = MkRWST $ \_,s,w => pure (a,s,w)
MkRWST mf <*> MkRWST mx =
MkRWST \r,s,w => do (f,s1,w1) <- mf r s w
(a,s2,w2) <- mx r s1 w1
pure (f a,s2,w2)
MkRWST $ \r,s,w => do (f,s1,w1) <- mf r s w
(a,s2,w2) <- mx r s1 w1
pure (f a,s2,w2)
public export %inline
(Monad m, Alternative m) => Alternative (RWST r w s m) where
empty = MkRWST \_,_,_ => empty
MkRWST m <|> MkRWST n = MkRWST \r,s,w => m r s w <|> n r s w
empty = MkRWST $ \_,_,_ => empty
MkRWST m <|> MkRWST n = MkRWST $ \r,s,w => m r s w <|> n r s w
public export %inline
Monad m => Monad (RWST r w s m) where
m >>= k = MkRWST \r,s,w => do (a,s1,w1) <- unRWST m r s w
unRWST (k a) r s1 w1
m >>= k = MkRWST $ \r,s,w => do (a,s1,w1) <- unRWST m r s w
unRWST (k a) r s1 w1
public export %inline
MonadTrans (RWST r w s) where
lift m = MkRWST \_,s,w => map (\a => (a,s,w)) m
lift m = MkRWST $ \_,s,w => map (\a => (a,s,w)) m
public export %inline
HasIO m => HasIO (RWST r w s m) where

View File

@ -37,8 +37,8 @@ Monad m => MonadReader stateType (ReaderT stateType m) where
public export %inline
Monad m => MonadReader r (RWST r w s m) where
ask = MkRWST \r,s,w => pure (r,s,w)
local f m = MkRWST \r,s,w => unRWST m (f r) s w
ask = MkRWST $ \r,s,w => pure (r,s,w)
local f m = MkRWST $ \r,s,w => unRWST m (f r) s w
public export %inline
MonadReader r m => MonadReader r (EitherT e m) where
@ -64,4 +64,4 @@ MonadReader r m => MonadReader r (WriterT w m) where
-- this should require a Monoid instance to further
-- accumulate values, while the implementation of
-- MonadReader for RWST does no such thing.
local f (MkWriterT m) = MkWriterT \w => local f (m w)
local f (MkWriterT m) = MkWriterT $ \w => local f (m w)

View File

@ -14,7 +14,7 @@ record ReaderT (stateType : Type) (m : Type -> Type) (a : Type) where
||| Transform the computation inside a @ReaderT@.
public export %inline
mapReaderT : (m a -> n b) -> ReaderT r m a -> ReaderT r n b
mapReaderT f m = MkReaderT \st => f (runReaderT' m st)
mapReaderT f m = MkReaderT $ \st => f (runReaderT' m st)
||| Unwrap and apply a ReaderT monad computation
public export

View File

@ -63,9 +63,9 @@ MonadState s m => MonadState s (MaybeT m) where
public export %inline
Monad m => MonadState s (RWST r w s m) where
get = MkRWST \_,s,w => pure (s,s,w)
put s = MkRWST \_,_,w => pure ((),s,w)
state f = MkRWST \_,s,w => let (s',a) = f s in pure (a,s',w)
get = MkRWST $ \_,s,w => pure (s,s,w)
put s = MkRWST $ \_,_,w => pure ((),s,w)
state f = MkRWST $ \_,s,w => let (s',a) = f s in pure (a,s',w)
public export %inline
MonadState s m => MonadState s (ReaderT r m) where

View File

@ -61,7 +61,7 @@ execState s = fst . runState s
||| the given function.
public export %inline
mapState : ((s, a) -> (s, b)) -> State s a -> State s b
mapState f = mapStateT \(Id p) => Id (f p)
mapState f = mapStateT $ \(Id p) => Id (f p)
--------------------------------------------------------------------------------
-- Implementations

View File

@ -44,7 +44,7 @@ execWriterT = map snd . runWriterT
public export %inline
mapWriterT : (Functor n, Monoid w, Semigroup w')
=> (m (a, w) -> n (b, w')) -> WriterT w m a -> WriterT w' n b
mapWriterT f m = MkWriterT \w =>
mapWriterT f m = MkWriterT $ \w =>
(\(a,w') => (a,w <+> w')) <$> f (runWriterT m)
--------------------------------------------------------------------------------
@ -72,7 +72,7 @@ execWriter = runIdentity . execWriterT
public export %inline
mapWriter : (Monoid w, Semigroup w')
=> ((a, w) -> (b, w')) -> Writer w a -> Writer w' b
mapWriter f = mapWriterT \(Id p) => Id (f p)
mapWriter f = mapWriterT $ \(Id p) => Id (f p)
--------------------------------------------------------------------------------
-- Implementations
@ -80,29 +80,29 @@ mapWriter f = mapWriterT \(Id p) => Id (f p)
public export %inline
Functor m => Functor (WriterT w m) where
map f m = MkWriterT \w => (\(a,w') => (f a,w')) <$> unWriterT m w
map f m = MkWriterT $ \w => (\(a,w') => (f a,w')) <$> unWriterT m w
public export %inline
Monad m => Applicative (WriterT w m) where
pure a = MkWriterT \w => pure (a,w)
pure a = MkWriterT $ \w => pure (a,w)
MkWriterT mf <*> MkWriterT mx =
MkWriterT \w => do (f,w1) <- mf w
(a,w2) <- mx w1
pure (f a,w2)
MkWriterT $ \w => do (f,w1) <- mf w
(a,w2) <- mx w1
pure (f a,w2)
public export %inline
(Monad m, Alternative m) => Alternative (WriterT w m) where
empty = MkWriterT \_ => empty
MkWriterT m <|> MkWriterT n = MkWriterT \w => m w <|> n w
empty = MkWriterT $ \_ => empty
MkWriterT m <|> MkWriterT n = MkWriterT $ \w => m w <|> n w
public export %inline
Monad m => Monad (WriterT w m) where
m >>= k = MkWriterT \w => do (a,w1) <- unWriterT m w
unWriterT (k a) w1
m >>= k = MkWriterT $ \w => do (a,w1) <- unWriterT m w
unWriterT (k a) w1
public export %inline
MonadTrans (WriterT w) where
lift m = MkWriterT \w => map (\a => (a,w)) m
lift m = MkWriterT $ \w => map (\a => (a,w)) m
public export %inline
HasIO m => HasIO (WriterT w m) where

View File

@ -60,47 +60,47 @@ public export %inline
(Monoid w, Monad m) => MonadWriter w (WriterT w m) where
writer = writerT . pure
listen m = MkWriterT \w =>
listen m = MkWriterT $ \w =>
(\(a,w') => ((a,w'),w <+> w')) <$> runWriterT m
tell w' = writer ((), w')
pass m = MkWriterT \w =>
pass m = MkWriterT $ \w =>
(\((a,f),w') => (a,w <+> f w')) <$> runWriterT m
public export %inline
(Monoid w, Monad m) => MonadWriter w (RWST r w s m) where
writer (a,w') = MkRWST \_,s,w => pure (a,s,w <+> w')
writer (a,w') = MkRWST $ \_,s,w => pure (a,s,w <+> w')
tell w' = writer ((), w')
listen m = MkRWST \r,s,w =>
listen m = MkRWST $ \r,s,w =>
(\(a,s',w') => ((a,w'),s',w <+> w')) <$> runRWST m r s
pass m = MkRWST \r,s,w =>
pass m = MkRWST $ \r,s,w =>
(\((a,f),s',w') => (a,s',w <+> f w')) <$> runRWST m r s
public export %inline
MonadWriter w m => MonadWriter w (EitherT e m) where
writer = lift . writer
tell = lift . tell
listen = mapEitherT \m => do (e,w) <- listen m
pure $ map (\a => (a,w)) e
listen = mapEitherT $ \m => do (e,w) <- listen m
pure $ map (\a => (a,w)) e
pass = mapEitherT \m => pass $ do Right (r,f) <- m
| Left l => pure $ (Left l, id)
pure (Right r, f)
pass = mapEitherT $ \m => pass $ do Right (r,f) <- m
| Left l => pure $ (Left l, id)
pure (Right r, f)
public export %inline
MonadWriter w m => MonadWriter w (MaybeT m) where
writer = lift . writer
tell = lift . tell
listen = mapMaybeT \m => do (e,w) <- listen m
pure $ map (\a => (a,w)) e
listen = mapMaybeT $ \m => do (e,w) <- listen m
pure $ map (\a => (a,w)) e
pass = mapMaybeT \m => pass $ do Just (r,f) <- m
| Nothing => pure $ (Nothing, id)
pure (Just r, f)
pass = mapMaybeT $ \m => pass $ do Just (r,f) <- m
| Nothing => pure $ (Nothing, id)
pure (Just r, f)
public export %inline
MonadWriter w m => MonadWriter w (ReaderT r m) where
writer = lift . writer
@ -112,8 +112,8 @@ public export %inline
MonadWriter w m => MonadWriter w (StateT s m) where
writer = lift . writer
tell = lift . tell
listen (ST m) = ST \s => do ((s',a),w) <- listen (m s)
pure (s',(a,w))
listen (ST m) = ST $ \s => do ((s',a),w) <- listen (m s)
pure (s',(a,w))
pass (ST m) = ST \s => pass $ do (s',(a,f)) <- m s
pure ((s',a),f)
pass (ST m) = ST $ \s => pass $ do (s',(a,f)) <- m s
pure ((s',a),f)

View File

@ -215,7 +215,7 @@ inBounds k [] = No uninhabited
inBounds Z (x::xs) = Yes InFirst
inBounds (S k) (x::xs) = case inBounds k xs of
Yes p => Yes $ InLater p
No up => No \(InLater p) => up p
No up => No $ \(InLater p) => up p
||| Find a particular element of a Colist using InBounds
|||

View File

@ -45,8 +45,8 @@ using (k : Nat)
connex {y = FZ} _ = Right $ FromNatPrf LTEZero
connex {x = FS k} {y = FS j} prf =
case connex {rel = FinLTE} $ prf . (cong FS) of
Left $ FromNatPrf p => Left $ FromNatPrf $ LTESucc p
Right $ FromNatPrf p => Right $ FromNatPrf $ LTESucc p
Left (FromNatPrf p) => Left $ FromNatPrf $ LTESucc p
Right (FromNatPrf p) => Right $ FromNatPrf $ LTESucc p
public export
Decidable 2 [Fin k, Fin k] FinLTE where

View File

@ -49,7 +49,7 @@ Semigroup a => Semigroup (Morphism r a) where
export
Monoid a => Monoid (Morphism r a) where
neutral = Mor \_ => neutral
neutral = Mor $ \_ => neutral
export
Semigroup (Endomorphism a) where
@ -77,16 +77,16 @@ Monad f => Monad (Kleislimorphism f a) where
public export
Contravariant (Op b) where
contramap f (MkOp g) = MkOp (g . f)
v >$ (MkOp f) = MkOp \_ => f v
v >$ (MkOp f) = MkOp $ \_ => f v
-- Applicative is a bit too strong, but there is no suitable superclass
export
(Semigroup a, Applicative f) => Semigroup (Kleislimorphism f r a) where
f <+> g = Kleisli \r => (<+>) <$> (applyKleisli f) r <*> (applyKleisli g) r
f <+> g = Kleisli $ \r => (<+>) <$> (applyKleisli f) r <*> (applyKleisli g) r
export
(Monoid a, Applicative f) => Monoid (Kleislimorphism f r a) where
neutral = Kleisli \_ => pure neutral
neutral = Kleisli $ \_ => pure neutral
export
Cast (Endomorphism a) (Morphism a a) where

View File

@ -91,7 +91,7 @@ run prog = runK prog pure
export
Functor io => Functor (L io) where
map fn act = Bind act \a' => PureW (fn a')
map fn act = Bind act $ \a' => PureW (fn a')
export
Applicative io => Applicative (L io) where

View File

@ -104,14 +104,14 @@ export
left >>> right = MkValidator (validateT left >=> validateT right)
Monad m => Alternative (ValidatorT m a) where
left <|> right = MkValidator \x => MkEitherT $ do
left <|> right = MkValidator $ \x => MkEitherT $ do
case !(runEitherT $ validateT left x) of
(Right r) => pure $ Right r
(Left e) => case !(runEitherT $ validateT right x) of
(Right r) => pure $ Right r
(Left e') => pure $ Left (e <+> " / " <+> e')
empty = MkValidator \x => MkEitherT $ pure (Left "invalid")
empty = MkValidator $ \x => MkEitherT $ pure (Left "invalid")
||| Alter the input before validation using given function.
export
@ -124,7 +124,7 @@ contramap f v = MkValidator (validateT v . f)
||| raw input in case it was helpful.
export
decide : Monad m => (t -> String) -> ((x : t) -> Dec (p x)) -> PropValidator m t p
decide msg dec = MkPropValidator \x => case dec x of
decide msg dec = MkPropValidator $ \x => case dec x of
Yes prf => pure prf
No _ => left (msg x)
@ -132,7 +132,7 @@ decide msg dec = MkPropValidator \x => case dec x of
||| converting it into b.
export
fromMaybe : Monad m => (a -> String) -> (a -> Maybe b) -> ValidatorT m a b
fromMaybe e f = MkValidator \a => case f a of
fromMaybe e f = MkValidator $ \a => case f a of
Nothing => left $ e a
Just b => pure b
@ -177,7 +177,7 @@ length l = MkValidator (validateVector l)
||| Verify that certain values are equal.
export
equal : (DecEq t, Monad m) => (a : t) -> PropValidator m t (\b => a = b)
equal a = MkPropValidator \b => case decEq a b of
equal a = MkPropValidator $ \b => case decEq a b of
Yes prf => pure prf
No _ => left "Values are not equal."

View File

@ -47,7 +47,7 @@ find x (y :: xs) with (decEq x y)
find x (x :: xs) | Yes Refl = Yes (Element Z Z)
find x (y :: xs) | No neqxy = case find x xs of
Yes (Element n prf) => Yes (Element (S n) (S prf))
No notInxs => No \case
No notInxs => No $ \case
(Element Z p) => void (neqxy (inverseZ p))
(Element (S n) prf) => absurd (notInxs (Element n (inverseS prf)))

View File

@ -26,7 +26,7 @@ elemAppLorR : (xs, ys : List a)
elemAppLorR [] [] prf = absurd prf
elemAppLorR [] _ prf = Right prf
elemAppLorR (x :: xs) [] prf =
Left rewrite sym $ appendNilRightNeutral xs in prf
Left $ rewrite sym $ appendNilRightNeutral xs in prf
elemAppLorR (x :: xs) _ Here = Left Here
elemAppLorR (x :: xs) ys (There prf) = mapFst There $ elemAppLorR xs ys prf

View File

@ -96,8 +96,9 @@ appendSameRightInjective xs ys (z::zs) = rewrite appendAssociative xs [z] zs in
export
appendNonEmptyLeftNotEq : (zs, xs : List a) -> NonEmpty xs => Not (zs = xs ++ zs)
appendNonEmptyLeftNotEq [] (_::_) Refl impossible
appendNonEmptyLeftNotEq (z::zs) (_::xs) prf = appendNonEmptyLeftNotEq zs (xs ++ [z]) @{SnocNonEmpty xs z}
rewrite sym $ appendAssociative xs [z] zs in snd $ consInjective prf
appendNonEmptyLeftNotEq (z::zs) (_::xs) prf
= appendNonEmptyLeftNotEq zs (xs ++ [z]) @{SnocNonEmpty xs z}
$ rewrite sym $ appendAssociative xs [z] zs in snd $ consInjective prf
||| List cannot be equal to itself appended with some non-empty list.
export

View File

@ -65,7 +65,7 @@ splitBalancedHelper revLs rs [] prf = MkSplitBal balancedLeftsAndRights
mkBalancedEq Refl
splitBalancedHelper revLs [] (x :: xs) prf =
absurd $
the (0 = S (plus (length revLs) (length xs)))
the (0 = S (plus (length revLs) (length xs))) $
rewrite plusSuccRightSucc (length revLs) (length xs) in
prf
splitBalancedHelper revLs (x :: rs) [lastItem] prf =

View File

@ -13,7 +13,7 @@ PEM p = Either p (Not p)
||| Double negation elimination
DNE : Type -> Type
DNE p = Not $ Not p -> p
DNE p = Not (Not p) -> p
||| The consensus theorem (at least, the interesting part)
Consensus : Type -> Type -> Type -> Type
@ -96,7 +96,7 @@ switchPEM (Right r) f = f r
EDN : DN p -> DNE p -> p
EDN f g = g f
pemDNE : DNE $ PEM p -> PEM p
pemDNE : DNE (PEM p) -> PEM p
pemDNE = EDN pemDN
-- It's possible to prove the theorems assuming Peirce's law, but some

View File

@ -212,7 +212,7 @@ multiplesModuloZero : (fuel, predn, k : Nat)
-> (enough : fuel `GTE` k * (S predn) )
-> mod' fuel (k * (S predn)) predn = 0
multiplesModuloZero 0 predn k enough =
let (k_eq_z, _) = multiplicationLemma k predn 0
let (k_eq_z, _) = multiplicationLemma k predn 0 $
rewrite plusZeroRightNeutral (k * (S predn)) in
enough
in rewrite k_eq_z in

View File

@ -210,7 +210,7 @@ factorLteNumber (CofactorExists (S k) prf) =
export
plusDivisorAlsoFactor : Factor p n -> Factor p (n + p)
plusDivisorAlsoFactor (CofactorExists q prf) =
CofactorExists (S q)
CofactorExists (S q) $
rewrite plusCommutative n p in
rewrite multRightSuccPlus p q in
cong (plus p) prf
@ -222,7 +222,7 @@ plusDivisorNeitherFactor (ZeroNotFactorS k) =
rewrite plusZeroRightNeutral k in
ZeroNotFactorS k
plusDivisorNeitherFactor (ProperRemExists q r remPrf) =
ProperRemExists (S q) r
ProperRemExists (S q) r $
rewrite multRightSuccPlus p q in
rewrite sym $ plusAssociative p (p * q) (S $ finToNat r) in
rewrite plusCommutative p ((p * q) + S (finToNat r)) in
@ -234,7 +234,7 @@ export
multNAlsoFactor : Factor p n -> (a : Nat) -> {auto aok : LTE 1 a} -> Factor p (n * a)
multNAlsoFactor _ Z = absurd $ succNotLTEzero aok
multNAlsoFactor (CofactorExists q prf) (S a) =
CofactorExists (q * S a)
CofactorExists (q * S a) $
rewrite prf in
sym $ multAssociative p q (S a)
@ -254,15 +254,14 @@ plusFactor (CofactorExists qn prfN) (CofactorExists qm prfM) =
export
minusFactor : {b : Nat} -> Factor p (a + b) -> Factor p a -> Factor p b
minusFactor (CofactorExists qab prfAB) (CofactorExists qa prfA) =
CofactorExists (minus qab qa) (
CofactorExists (minus qab qa) $
rewrite multDistributesOverMinusRight p qab qa in
rewrite sym prfA in
rewrite sym prfAB in
replace {p = \x => b = minus (a + b) x} (plusZeroRightNeutral a)
replace {p = \x => b = minus (a + b) x} (plusZeroRightNeutral a) $
rewrite plusMinusLeftCancel a b 0 in
rewrite minusZeroRight b in
Refl
)
||| A decision procedure for whether of not p is a factor of n.
export
@ -454,25 +453,25 @@ divByGcdHelper : (a, b, c : Nat) -> GCD (S a) (S a * S b) (S a * c) -> GCD 1 (S
divByGcdHelper a b c (MkGCD _ greatest) =
MkGCD (CommonFactorExists 1 (oneIsFactor (S b)) (oneIsFactor c)) $
\q, (CommonFactorExists q (CofactorExists qb prfQB) (CofactorExists qc prfQC)) =>
let qFab = CofactorExists qb
let qFab = CofactorExists qb $
rewrite multCommutative q (S a) in
rewrite sym $ multAssociative (S a) q qb in
rewrite sym $ prfQB in
Refl
qFac = CofactorExists qc
qFac = CofactorExists qc $
rewrite multCommutative q (S a) in
rewrite sym $ multAssociative (S a) q qc in
rewrite sym $ prfQC in
Refl
CofactorExists f prfQAfA =
greatest (q * S a) (CommonFactorExists (q * S a) qFab qFac)
qf1 = multOneSoleNeutral a (f * q)
qf1 = multOneSoleNeutral a (f * q) $
rewrite multCommutative f q in
rewrite multAssociative (S a) q f in
rewrite sym $ multCommutative q (S a) in
prfQAfA
in
CofactorExists f
CofactorExists f $
rewrite multCommutative q f in
sym qf1

View File

@ -21,7 +21,7 @@ LTESuccInjectiveMonotone m n (RFalse not_m_lte_n) = RFalse $ \case
export
lteReflection : (a, b : Nat) -> Reflects (a `LTE` b) (a `lte` b)
lteReflection 0 b = RTrue LTEZero
lteReflection (S k) 0 = RFalse \sk_lte_z => absurd sk_lte_z
lteReflection (S k) 0 = RFalse $ \sk_lte_z => absurd sk_lte_z
lteReflection (S a) (S b) = LTESuccInjectiveMonotone a b (lteReflection a b)
export
@ -48,7 +48,7 @@ notltIsNotLT a = notlteIsNotLTE (S a)
export
notlteIsLT : (a, b : Nat) -> a `lte` b = False -> b `LT` a
notlteIsLT a b prf = notLTImpliesGTE
notlteIsLT a b prf = notLTImpliesGTE $
\prf' =>
(invert $ replace {p = Reflects (S a `LTE` S b)} prf
$ lteReflection (S a) (S b)) prf'

View File

@ -136,9 +136,10 @@ projection : {0 gamma : Left.Telescope k}
-> Environment (gamma |++ delta)
-> Environment gamma
projection {n = 0 } {delta = [] } env = env
projection {n = S n} {delta = ty :: delta} env = let (env' ** _) = projection {n} {delta}
rewrite succLemma n k in env
in env'
projection {n = S n} {delta = ty :: delta} env
= let (env' ** _) = projection {n} {delta}
$ rewrite succLemma n k in env
in env'
infixl 4 .=

View File

@ -133,10 +133,10 @@ public export
decEq (Invalid x) (Valid y) = No uninhabited
decEq (Valid x) (Valid y) with (decEq x y)
decEq (Valid _) (Valid _) | Yes p = rewrite p in Yes Refl
decEq (Valid _) (Valid _) | No up = No \case Refl => up Refl
decEq (Valid _) (Valid _) | No up = No $ \case Refl => up Refl
decEq (Invalid x) (Invalid y) with (decEq x y)
decEq (Invalid _) (Invalid _) | Yes p = rewrite p in Yes Refl
decEq (Invalid _) (Invalid _) | No up = No \case Refl => up Refl
decEq (Invalid _) (Invalid _) | No up = No $ \case Refl => up Refl
--- Convenience representations ---

View File

@ -15,7 +15,7 @@ import Syntax.PreorderReasoning
||| `map` functoriality: identity preservation
export
mapId : (xs : Vect n a) -> map Prelude.id xs = xs
mapId xs = vectorExtensionality _ _ \i => indexNaturality _ _ _
mapId xs = vectorExtensionality _ _ $ \i => indexNaturality _ _ _
||| `mapWtihPos f` represents post-composition the tabulated function `f`
export
@ -29,7 +29,7 @@ export
mapTabulate : (f : a -> b) -> (g : Fin n -> a)
-> tabulate (f . g) = map f (tabulate g)
mapTabulate f g = irrelevantEq $
vectorExtensionality _ _ \i => Calc $
vectorExtensionality _ _ $ \i => Calc $
|~ index i (tabulate (f . g))
~~ f (g i) ...(indexTabulate _ _)
~~ f (index i $ tabulate g) ...(cong f (sym $ indexTabulate _ _))
@ -39,7 +39,7 @@ mapTabulate f g = irrelevantEq $
export
tabulateConstantly : (x : a) -> Fin.tabulate {len} (const x) === replicate len x
tabulateConstantly x = irrelevantEq $
vectorExtensionality _ _ \i => Calc $
vectorExtensionality _ _ $ \i => Calc $
|~ index i (Fin.tabulate (const x))
~~ x ...(indexTabulate _ _)
~~ index i (replicate _ x) ...(sym $ indexReplicate _ _)
@ -49,7 +49,7 @@ export
mapRestrictedExtensional : (f, g : a -> b) -> (xs : Vect n a)
-> (prf : (i : Fin n) -> f (index i xs) = g (index i xs))
-> map f xs = map g xs
mapRestrictedExtensional f g xs prf = vectorExtensionality _ _ \i => Calc $
mapRestrictedExtensional f g xs prf = vectorExtensionality _ _ $ \i => Calc $
|~ index i (map f xs)
~~ f (index i xs) ...( indexNaturality _ _ _)
~~ g (index i xs) ...(prf _)
@ -76,4 +76,6 @@ mapWithElemExtensional : (xs : Vect n a) -> (f, g : (x : a) -> (0 _ : x `Elem`
-> ((x : a) -> (0 pos : x `Elem` xs) -> f x pos = g x pos)
-> mapWithElem xs f = mapWithElem xs g
mapWithElemExtensional [] f g prf = Refl
mapWithElemExtensional (x :: xs) f g prf = cong2 (::) (prf x Here) (mapWithElemExtensional xs _ _ (\x,pos => prf x (There pos)))
mapWithElemExtensional (x :: xs) f g prf
= cong2 (::) (prf x Here)
(mapWithElemExtensional xs _ _ (\x,pos => prf x (There pos)))

View File

@ -50,7 +50,7 @@ smallerPlusL m k = rewrite sym (plusSuccRightSucc m k) in LTESucc $ LTESucc $ lt
smallerPlusR : (m : Nat) -> (k : Nat) -> LTE (S (S k)) (S (m + S k))
smallerPlusR m k = rewrite sym (plusSuccRightSucc m k) in
LTESucc $ LTESucc rewrite plusCommutative m k in lteAddRight _
LTESucc $ LTESucc $ rewrite plusCommutative m k in lteAddRight _
||| Covering function for the `SplitRec` view
||| Constructs the view in O(n lg n)

View File

@ -63,7 +63,7 @@ decideTransform :
-> (posDec : IsDecidable n ts r)
-> IsDecidable n ts (chain {ts} t r)
decideTransform tDec posDec =
curryAll \xs =>
curryAll $ \xs =>
replace {p = id} (chainUncurry (chain t r) Dec xs) $
replace {p = Dec} (chainUncurry r t xs) $
tDec $ replace {p = id} (sym $ chainUncurry r Dec xs) $
@ -85,8 +85,8 @@ notExistsNotForall :
-> Dec ((x : a) -> p x)
notExistsNotForall dec decEx =
case decEx of
Yes (x ** nx) => No \f => nx $ f x
No notNot => Yes \x => case (dec x) of
Yes (x ** nx) => No $ \f => nx $ f x
No notNot => Yes $ \x => case (dec x) of
Yes px => px
No nx => void $ notNot $ (x ** nx)

View File

@ -108,7 +108,7 @@ fmtOpt (MkOpt sos los ad descr) =
||| the header (first argument) and the options described by the
||| second argument.
public export
usageInfo : (header : String) -> List $ OptDescr a -> String
usageInfo : (header : String) -> List (OptDescr a) -> String
usageInfo header optDescr =
let (ss,ls,ds) = (unzip3 . concatMap fmtOpt) optDescr
paste = \x,y,z => " " ++ x ++ " " ++ y ++ " " ++ z
@ -126,7 +126,7 @@ usageInfo header optDescr =
-- Error Formatting
--------------------------------------------------------------------------------
errAmbig : List $ OptDescr a -> (optStr : String) -> OptKind a
errAmbig : List (OptDescr a) -> (optStr : String) -> OptKind a
errAmbig ods s = let h = "option `" ++ s ++ "' is ambiguous; could be one of:"
in OptErr (usageInfo h ods)
@ -167,7 +167,7 @@ Functor Result where
map f = record { options $= map f }
OptFun : Type -> Type
OptFun a = List String -> List $ OptDescr a -> (OptKind a,List String)
OptFun a = List String -> List (OptDescr a) -> (OptKind a,List String)
longOpt : String -> OptFun a
longOpt ls rs descs =
@ -231,7 +231,7 @@ getNext a r _ = (NonOpt $ pack a,r)
||| `System.getArgs`).
export
getOpt : ArgOrder a -- non-option handling
-> List $ OptDescr a -- option descriptors
-> List (OptDescr a) -- option descriptors
-> (args : List String) -- the command-line arguments
-> Result a
getOpt _ _ [] = emptyRes

View File

@ -187,7 +187,7 @@ mutual
anf vs (LUnderApp fc n m args)
= anfArgs fc vs args (AUnderApp fc n m)
anf vs (LApp fc lazy f a)
= anfArgs fc vs [f, a]
= anfArgs fc vs [f, a] $
\case
[fvar, avar] => AApp fc lazy fvar avar
_ => ACrash fc "Can't happen (AApp)"

View File

@ -137,7 +137,7 @@ anyOne : {vars : _} ->
Core (Term vars)
anyOne fc env top [] = throw (CantSolveGoal fc [] top)
anyOne fc env top [elab]
= catch elab
= catch elab $
\case
err@(CantSolveGoal _ _ _) => throw err
_ => throw $ CantSolveGoal fc [] top
@ -151,7 +151,7 @@ exactlyOne : {vars : _} ->
List (Core (Term vars)) ->
Core (Term vars)
exactlyOne fc env top target [elab]
= catch elab
= catch elab $
\case
err@(CantSolveGoal _ _ _) => throw err
_ => throw $ CantSolveGoal fc [] top

View File

@ -894,7 +894,7 @@ getScore : {ns : _} ->
getScore fc phase name npss
= do catch (do sameType fc phase name (mkEnv fc ns) npss
pure (Right ()))
\case
$ \case
CaseCompile _ _ err => pure $ Left err
err => throw err

View File

@ -61,6 +61,7 @@ Pretty DotReason where
public export
data Warning : Type where
ParserWarning : FC -> String -> Warning
UnreachableClause : {vars : _} ->
FC -> Env Term vars -> Term vars -> Warning
ShadowingGlobalDefs : FC -> List1 (String, List1 Name) -> Warning
@ -181,6 +182,7 @@ Show TTCErrorMsg where
export
Show Warning where
show (ParserWarning _ msg) = msg
show (UnreachableClause _ _ _) = ":Unreachable clause"
show (ShadowingGlobalDefs _ _) = ":Shadowing names"
show (Deprecated name) = ":Deprecated " ++ name
@ -358,6 +360,7 @@ Show Error where
export
getWarningLoc : Warning -> Maybe FC
getWarningLoc (ParserWarning fc _) = Just fc
getWarningLoc (UnreachableClause fc _ _) = Just fc
getWarningLoc (ShadowingGlobalDefs fc _) = Just fc
getWarningLoc (Deprecated _) = Nothing

View File

@ -595,7 +595,7 @@ tryInstantiate {newvars} loc mode env mname mref num mdef locs otm tm
pure $ (Bind bfc x (Lam fc c Explicit (Erased bfc False)) <$> sc')
mkDef vs vars soln (Bind bfc x b@(Let _ c val ty) sc)
= do mbsc' <- mkDef vs (ICons Nothing vars) soln sc
flip traverseOpt mbsc' \sc' =>
flip traverseOpt mbsc' $ \sc' =>
case shrinkTerm sc' (DropCons SubRefl) of
Just scs => pure scs
Nothing => pure $ Bind bfc x b sc'
@ -1490,7 +1490,7 @@ retryGuess mode smode (hid, (loc, hname))
ignore $ addDef (Resolved hid) gdef
removeGuess hid
pure True)
\case
$ \case
DeterminingArg _ n i _ _ =>
do logTerm "unify.retry" 5
("Failed (det " ++ show hname ++ " " ++ show n ++ ")")

View File

@ -195,7 +195,7 @@ mutual
!(desugarB side ps' retTy)
desugarB side ps (PLam fc rig p pat@(PRef prefFC n@(UN nm)) argTy scope)
= if lowerFirst nm || nm == "_"
then do whenJust (isConcreteFC prefFC) \nfc
then do whenJust (isConcreteFC prefFC) $ \nfc
=> addSemanticDecorations [(nfc, Bound, Just n)]
pure $ ILam fc rig !(traverse (desugar AnyExpr ps) p)
(Just n) !(desugarB AnyExpr ps argTy)
@ -218,7 +218,7 @@ mutual
ICase fc (IVar EmptyFC (MN "lamc" 0)) (Implicit fc False)
[snd !(desugarClause ps True (MkPatClause fc pat scope []))]
desugarB side ps (PLet fc rig (PRef prefFC n) nTy nVal scope [])
= do whenJust (isConcreteFC prefFC) \nfc =>
= do whenJust (isConcreteFC prefFC) $ \nfc =>
addSemanticDecorations [(nfc, Bound, Just n)]
pure $ ILet fc prefFC rig n !(desugarB side ps nTy) !(desugarB side ps nVal)
!(desugar side (n :: ps) scope)
@ -376,7 +376,7 @@ mutual
= do r' <- desugarB side ps r
let pval = apply (IVar opFC mkdpairname) [IVar nameFC n, r']
let vfc = virtualiseFC nameFC
whenJust (isConcreteFC nameFC) \nfc =>
whenJust (isConcreteFC nameFC) $ \nfc =>
addSemanticDefault (nfc, Bound, Just n)
pure $ IAlternative fc (UniqueDefault pval)
[apply (IVar opFC dpairname)
@ -578,7 +578,7 @@ mutual
expandDo side ps topfc ns (DoBind fc nameFC n tm :: rest)
= do tm' <- desugarDo side ps ns tm
rest' <- expandDo side ps topfc ns rest
whenJust (isConcreteFC nameFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
whenJust (isConcreteFC nameFC) $ \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
pure $ bindFun fc ns tm'
$ ILam nameFC top Explicit (Just n)
(Implicit (virtualiseFC fc) False) rest'
@ -603,7 +603,7 @@ mutual
tm' <- desugarB side ps tm
ty' <- desugarDo side ps ns ty
rest' <- expandDo side ps topfc ns rest
whenJust (isConcreteFC lhsFC) \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
whenJust (isConcreteFC lhsFC) $ \nfc => addSemanticDecorations [(nfc, Bound, Just n)]
let bind = ILet fc (virtualiseFC lhsFC) rig n ty' tm' rest'
bd <- get Bang
pure $ bindBangs (bangNames bd) ns bind

View File

@ -485,7 +485,7 @@ elabInterface {vars} ifc vis env nest constraints iname params dets mcon body
= do if n /= n' then pure (IVar fc n') else do
log "ide-mode.highlight" 7 $
"elabDefault is trying to add Function: " ++ show n ++ " (" ++ show fc ++")"
whenJust (isConcreteFC fc) \nfc => do
whenJust (isConcreteFC fc) $ \nfc => do
log "ide-mode.highlight" 7 $ "elabDefault is adding Function: " ++ show n
addSemanticDecorations [(nfc, Function, Just n)]
pure (IVar fc dn)

View File

@ -159,6 +159,8 @@ pwarning : {auto c : Ref Ctxt Defs} ->
{auto s : Ref Syn SyntaxInfo} ->
{auto o : Ref ROpts REPLOpts} ->
Warning -> Core (Doc IdrisAnn)
pwarning (ParserWarning fc msg)
= pure $ pretty msg <+> line <+> !(ploc fc)
pwarning (UnreachableClause fc env tm)
= pure $ errorDesc (reflow "Unreachable clause:"
<++> code !(pshow env tm))

View File

@ -81,7 +81,7 @@ ideParser : {e : _} ->
String -> Grammar SemanticDecorations Token e ty -> Either Error ty
ideParser str p
= do toks <- mapFst (fromLexError (Virtual Interactive)) $ idelex str
(decor, (parsed, _)) <- mapFst (fromParsingErrors (Virtual Interactive)) $ parseWith p toks
(_, _, (parsed, _)) <- mapFst (fromParsingErrors (Virtual Interactive)) $ parseWith p toks
Right parsed

View File

@ -26,7 +26,6 @@ decorate fname decor rule = do
act [((fname, (start res, end res)), decor, Nothing)]
pure res.val
decorationFromBounded : OriginDesc -> Decoration -> WithBounds a -> ASemanticDecoration
decorationFromBounded fname decor bnds
= ((fname, start bnds, end bnds), decor, Nothing)
@ -109,7 +108,7 @@ atom fname
pure (PType (boundToFC fname x))
<|> do x <- bounds $ name
pure (PRef (boundToFC fname x) x.val)
<|> do x <- bounds $ dependentDecorate fname constant \c =>
<|> do x <- bounds $ dependentDecorate fname constant $ \c =>
if isPrimType c
then Typ
else Data
@ -227,7 +226,7 @@ mutual
$ do x <- bounds (decoratedSimpleBinderName fname)
let fc = boundToFC fname x
option (NamedArg (UN x.val) $ PRef fc (UN x.val))
$ do tm <- decoratedSymbol fname "=" *> expr pdef fname indents
$ do tm <- decoratedSymbol fname "=" *> typeExpr pdef fname indents
pure (NamedArg (UN x.val) tm)
matchAny <- option [] (if isCons list then
do decoratedSymbol fname ","
@ -243,7 +242,7 @@ mutual
<|> do decoratedSymbol fname "@{"
commit
tm <- expr pdef fname indents
tm <- typeExpr pdef fname indents
decoratedSymbol fname "}"
pure [UnnamedAutoArg tm]
@ -282,7 +281,13 @@ mutual
in POp fc opFC (UN "=") l.val r.val
else fail "= not allowed")
<|>
(do b <- bounds [| MkPair (continue indents *> bounds iOperator) (opExpr q fname indents) |]
(do b <- bounds $ do
continue indents
op <- bounds iOperator
e <- case op.val of
UN "$" => typeExpr q fname indents
_ => expr q fname indents
pure (op, e)
(op, r) <- pure b.val
let fc = boundToFC fname (mergeBounds l b)
let opFC = boundToFC fname op
@ -293,11 +298,11 @@ mutual
dpairType fname start indents
= do loc <- bounds (do x <- decoratedSimpleBinderName fname
decoratedSymbol fname ":"
ty <- expr pdef fname indents
ty <- typeExpr pdef fname indents
pure (x, ty))
(x, ty) <- pure loc.val
op <- bounds (symbol "**")
rest <- bounds (nestedDpair fname loc indents <|> expr pdef fname indents)
rest <- bounds (nestedDpair fname loc indents <|> typeExpr pdef fname indents)
pure (PDPair (boundToFC fname (mergeBounds start rest))
(boundToFC fname op)
(PRef (boundToFC fname loc) (UN x))
@ -342,7 +347,7 @@ mutual
-- dependent pairs with type annotation (so, the type form)
<|> do dpairType fname s indents <* (decorate fname Typ $ symbol ")")
<* act [(toNonEmptyFC $ boundToFC fname s, Typ, Nothing)]
<|> do e <- bounds (expr pdef fname indents)
<|> do e <- bounds (typeExpr pdef fname indents)
-- dependent pairs with no type annotation
(do loc <- bounds (symbol "**")
rest <- bounds ((nestedDpair fname loc indents <|> expr pdef fname indents) <* symbol ")")
@ -430,7 +435,7 @@ mutual
nonEmptyTuple : OriginDesc -> WithBounds t -> IndentInfo -> PTerm -> Rule PTerm
nonEmptyTuple fname s indents e
= do vals <- some $ do b <- bounds (symbol ",")
exp <- optional (expr pdef fname indents)
exp <- optional (typeExpr pdef fname indents)
pure (boundToFC fname b, exp)
end <- continueWithDecorated fname indents ")"
act [(toNonEmptyFC (boundToFC fname s), Keyword, Nothing)]
@ -495,18 +500,33 @@ mutual
(x, expr) <- pure b.val
pure (PAs (boundToFC fname b) (boundToFC fname x) (UN x.val) expr)
<|> atom fname
<|> binder fname indents
<|> rewrite_ fname indents
<|> record_ fname indents
<|> singlelineStr pdef fname indents
<|> multilineStr pdef fname indents
<|> do b <- bounds (decoratedSymbol fname ".(" *> commit *> expr pdef fname indents <* decoratedSymbol fname ")")
<|> do b <- bounds $ do
decoratedSymbol fname ".("
commit
t <- typeExpr pdef fname indents
decoratedSymbol fname ")"
pure t
pure (PDotted (boundToFC fname b) b.val)
<|> do b <- bounds (decoratedSymbol fname "`(" *> expr pdef fname indents <* decoratedSymbol fname ")")
<|> do b <- bounds $ do
decoratedSymbol fname "`("
t <- typeExpr pdef fname indents
decoratedSymbol fname ")"
pure t
pure (PQuote (boundToFC fname b) b.val)
<|> do b <- bounds (decoratedSymbol fname "`{" *> name <* decoratedSymbol fname "}")
<|> do b <- bounds $ do
decoratedSymbol fname "`{"
t <- name
decoratedSymbol fname "}"
pure t
pure (PQuoteName (boundToFC fname b) b.val)
<|> do b <- bounds (decoratedSymbol fname "`[" *> nonEmptyBlock (topDecl fname) <* decoratedSymbol fname "]")
<|> do b <- bounds $ do
decoratedSymbol fname "`["
ts <- nonEmptyBlock (topDecl fname)
decoratedSymbol fname "]"
pure ts
pure (PQuoteDecl (boundToFC fname b) (collectDefs (concat b.val)))
<|> do b <- bounds (decoratedSymbol fname "~" *> simpleExpr fname indents)
pure (PUnquote (boundToFC fname b) b.val)
@ -518,10 +538,12 @@ mutual
listExpr fname start indents
<|> do b <- bounds (decoratedSymbol fname "!" *> simpleExpr fname indents)
pure (PBang (virtualiseFC $ boundToFC fname b) b.val)
<|> do b <- bounds (decoratedSymbol fname "[|" *> expr pdef fname indents <* decoratedSymbol fname "|]")
<|> do b <- bounds $ do
decoratedSymbol fname "[|"
t <- expr pdef fname indents
decoratedSymbol fname "|]"
pure t
pure (PIdiom (boundToFC fname b) b.val)
<|> do b <- bounds (pragma "runElab" *> expr pdef fname indents)
pure (PRunElab (boundToFC fname b) b.val)
<|> do b <- bounds $ do pragma "logging"
topic <- optional (split (('.') ==) <$> simpleStr)
lvl <- intLit
@ -529,6 +551,8 @@ mutual
pure (MkPair (mkLogLevel' topic (integerToNat lvl)) e)
(lvl, e) <- pure b.val
pure (PUnifyLog (boundToFC fname b) lvl e)
<|> withWarning "DEPRECATED: trailing lambda. Use a $ or parens"
(lam fname indents)
multiplicity : OriginDesc -> EmptyRule RigCount
multiplicity fname
@ -564,14 +588,14 @@ mutual
let ns = forget $ map (map UN) ns
decorateBoundedNames fname Bound ns
decoratedSymbol fname ":"
ty <- expr pdef fname indents
ty <- typeExpr pdef fname indents
atEnd indents
pure (map (\n => (rig, n, ty)) ns)
<|> forget <$> sepBy1 (decoratedSymbol fname ",")
(do rig <- multiplicity fname
n <- bounds (decorate fname Bound binderName)
decoratedSymbol fname ":"
ty <- expr pdef fname indents
ty <- typeExpr pdef fname indents
pure (rig, map UN n, ty))
where
-- _ gets treated specially here, it means "I don't care about the name"
@ -591,58 +615,73 @@ mutual
explicitPi : OriginDesc -> IndentInfo -> Rule PTerm
explicitPi fname indents
= do decoratedSymbol fname "("
binders <- pibindList fname indents
decoratedSymbol fname ")"
exp <- bindSymbol fname
= do b <- bounds $ do
decoratedSymbol fname "("
binders <- pibindList fname indents
decoratedSymbol fname ")"
pure binders
exp <- mustWorkBecause b.bounds "Cannot return a named argument"
$ bindSymbol fname
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname exp binders scope)
pure (pibindAll fname exp b.val scope)
autoImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm
autoImplicitPi fname indents
= do decoratedSymbol fname "{"
decoratedKeyword fname "auto"
commit
binders <- pibindList fname indents
decoratedSymbol fname "}"
decoratedSymbol fname "->"
= do b <- bounds $ do
decoratedSymbol fname "{"
decoratedKeyword fname "auto"
commit
binders <- pibindList fname indents
decoratedSymbol fname "}"
pure binders
mustWorkBecause b.bounds "Cannot return an auto implicit argument"
$ decoratedSymbol fname "->"
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname AutoImplicit binders scope)
pure (pibindAll fname AutoImplicit b.val scope)
defaultImplicitPi : OriginDesc -> IndentInfo -> Rule PTerm
defaultImplicitPi fname indents
= do decoratedSymbol fname "{"
decoratedKeyword fname "default"
commit
t <- simpleExpr fname indents
binders <- pibindList fname indents
decoratedSymbol fname "}"
decoratedSymbol fname "->"
= do b <- bounds $ do
decoratedSymbol fname "{"
decoratedKeyword fname "default"
commit
t <- simpleExpr fname indents
binders <- pibindList fname indents
decoratedSymbol fname "}"
pure (t, binders)
mustWorkBecause b.bounds "Cannot return a default implicit argument"
$ decoratedSymbol fname "->"
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname (DefImplicit t) binders scope)
pure $ let (t, binders) = b.val in
pibindAll fname (DefImplicit t) binders scope
forall_ : OriginDesc -> IndentInfo -> Rule PTerm
forall_ fname indents
= do decoratedKeyword fname "forall"
commit
ns <- sepBy1 (decoratedSymbol fname ",")
(bounds (decoratedSimpleBinderName fname))
let binders = map (\n => ( erased {a=RigCount}
= do b <- bounds $ do
decoratedKeyword fname "forall"
commit
ns <- sepBy1 (decoratedSymbol fname ",")
(bounds (decoratedSimpleBinderName fname))
pure $ map (\n => ( erased {a=RigCount}
, map (Just . UN) n
, PImplicit (boundToFC fname n))
) (forget ns)
decoratedSymbol fname "."
) (forget ns)
mustWorkBecause b.bounds "Cannot return a forall quantifier"
$ decoratedSymbol fname "."
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname Implicit binders scope)
pure (pibindAll fname Implicit b.val scope)
implicitPi : OriginDesc -> IndentInfo -> Rule PTerm
implicitPi fname indents
= do decoratedSymbol fname "{"
binders <- pibindList fname indents
decoratedSymbol fname "}"
decoratedSymbol fname "->"
= do b <- bounds $ do
decoratedSymbol fname "{"
binders <- pibindList fname indents
decoratedSymbol fname "}"
pure binders
mustWorkBecause b.bounds "Cannot return an implicit argument"
$ decoratedSymbol fname "->"
scope <- mustWork $ typeExpr pdef fname indents
pure (pibindAll fname Implicit binders scope)
pure (pibindAll fname Implicit b.val scope)
lam : OriginDesc -> IndentInfo -> Rule PTerm
lam fname indents
@ -666,7 +705,7 @@ mutual
binders <- bindList fname indents
decoratedSymbol fname "=>"
mustContinue indents Nothing
scope <- expr pdef fname indents
scope <- typeExpr pdef fname indents
pure (bindAll binders scope)
continueLamCase : WithBounds () -> Rule PTerm
@ -688,7 +727,7 @@ mutual
ty <- option (PImplicit (virtualiseFC $ boundToFC fname s))
(decoratedSymbol fname ":" *> typeExpr (pnoeq pdef) fname indents)
(decoratedSymbol fname "=" <|> decoratedSymbol fname ":=")
val <- expr pnowith fname indents
val <- typeExpr pnowith fname indents
alts <- block (patAlt fname)
pure (MkLetBinder rig pat ty val alts)
@ -722,7 +761,7 @@ mutual
caseRHS : OriginDesc -> WithBounds t -> IndentInfo -> PTerm -> Rule PClause
caseRHS fname start indents lhs
= do rhs <- bounds (decoratedSymbol fname "=>" *> mustContinue indents Nothing *> expr pdef fname indents)
= do rhs <- bounds (decoratedSymbol fname "=>" *> mustContinue indents Nothing *> typeExpr pdef fname indents)
atEnd indents
let fc = boundToFC fname (mergeBounds start rhs)
pure (MkPatClause fc lhs rhs.val [])
@ -735,9 +774,9 @@ mutual
= do b <- bounds (do decoratedKeyword fname "if"
x <- expr pdef fname indents
commitKeyword fname indents "then"
t <- expr pdef fname indents
t <- typeExpr pdef fname indents
commitKeyword fname indents "else"
e <- expr pdef fname indents
e <- typeExpr pdef fname indents
pure (x, t, e))
atEnd indents
(x, t, e) <- pure b.val
@ -761,7 +800,7 @@ mutual
upd <- (ifThenElse kw (decoratedSymbol fname "=") (decoratedSymbol fname ":=") $> PSetField)
<|>
(decoratedSymbol fname "$=" $> PSetFieldApp)
val <- opExpr plhs fname indents
val <- typeExpr plhs fname indents
pure (upd path val)
where
fieldName : Name -> String
@ -781,7 +820,7 @@ mutual
= do b <- bounds (do decoratedKeyword fname "rewrite"
rule <- expr pdef fname indents
commitKeyword fname indents "in"
tm <- expr pdef fname indents
tm <- typeExpr pdef fname indents
pure (rule, tm))
(rule, tm) <- pure b.val
pure (PRewrite (boundToFC fname b) rule tm)
@ -858,8 +897,7 @@ mutual
binder : OriginDesc -> IndentInfo -> Rule PTerm
binder fname indents
= let_ fname indents
<|> autoImplicitPi fname indents
= autoImplicitPi fname indents
<|> defaultImplicitPi fname indents
<|> forall_ fname indents
<|> implicitPi fname indents
@ -868,21 +906,29 @@ mutual
typeExpr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
typeExpr q fname indents
= do arg <- bounds (opExpr q fname indents)
(do continue indents
rest <- some [| (bindSymbol fname, bounds $ opExpr pdef fname indents) |]
pure (mkPi arg (forget rest)))
<|> pure arg.val
= binder fname indents
<|> do arg <- bounds (expr q fname indents)
mscope <- optional $ do
continue indents
bd <- bindSymbol fname
scope <- mustWork $ typeExpr q fname indents
pure (bd, scope)
pure (mkPi arg mscope)
where
mkPi : WithBounds PTerm -> List (PiInfo PTerm, WithBounds PTerm) -> PTerm
mkPi arg [] = arg.val
mkPi arg ((exp, a) :: as)
= PPi (boundToFC fname arg) top exp Nothing arg.val
(mkPi a as)
mkPi : WithBounds PTerm -> Maybe (PiInfo PTerm, PTerm) -> PTerm
mkPi arg Nothing = arg.val
mkPi arg (Just (exp, a))
= PPi (boundToFC fname arg) top exp Nothing arg.val a
export
expr : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
expr = typeExpr
expr q fname indents
= let_ fname indents
<|> rewrite_ fname indents
<|> do b <- bounds (pragma "runElab" *> expr pdef fname indents)
pure (PRunElab (boundToFC fname b) b.val)
<|> opExpr q fname indents
interpBlock : ParseOpts -> OriginDesc -> IndentInfo -> Rule PTerm
interpBlock q fname idents = interpBegin *> (mustWork $ expr q fname idents) <* interpEnd
@ -943,11 +989,12 @@ visibility fname
tyDecls : Rule Name -> String -> OriginDesc -> IndentInfo -> Rule (List1 PTypeDecl)
tyDecls declName predoc fname indents
= do bs <- do docns <- sepBy1 (decoratedSymbol fname ",") [| (option "" documentation, bounds declName) |]
decoratedSymbol fname ":"
mustWork $ do ty <- expr pdef fname indents
pure $ map (\(doc, n) => (doc, n.val, boundToFC fname n, ty))
docns
= do bs <- do docns <- sepBy1 (decoratedSymbol fname ",")
[| (option "" documentation, bounds declName) |]
b <- bounds $ decoratedSymbol fname ":"
mustWorkBecause b.bounds "Expected a type declaration" $ do
ty <- typeExpr pdef fname indents
pure $ map (\(doc, n) => (doc, n.val, boundToFC fname n, ty)) docns
atEnd indents
pure $ map (\(doc, n, nFC, ty) => (MkPTy nFC nFC n (predoc ++ doc) ty))
bs
@ -962,7 +1009,12 @@ mutual
OriginDesc -> WithBounds t -> Int ->
IndentInfo -> (lhs : PTerm) -> Rule PClause
parseRHS withArgs fname start col indents lhs
= do b <- bounds $ decoratedSymbol fname "=" *> mustWork [| (expr pdef fname indents, option [] $ whereBlock fname col) |]
= do b <- bounds $ do
decoratedSymbol fname "="
mustWork $ do
rhs <- typeExpr pdef fname indents
ws <- option [] $ whereBlock fname col
pure (rhs, ws)
atEnd indents
(rhs, ws) <- pure b.val
let fc = boundToFC fname (mergeBounds start b)
@ -986,7 +1038,7 @@ mutual
clause : Nat -> OriginDesc -> IndentInfo -> Rule PClause
clause withArgs fname indents
= do b <- bounds (do col <- column
lhs <- expr plhs fname indents
lhs <- opExpr plhs fname indents
extra <- many parseWithArg
pure (col, lhs, extra))
(col, lhs, extra) <- pure b.val
@ -1076,7 +1128,7 @@ gadtData : OriginDesc -> Int -> WithBounds t ->
gadtData fname mincol start tyName indents
= do decoratedSymbol fname ":"
commit
ty <- expr pdef fname indents
ty <- typeExpr pdef fname indents
dataBody fname mincol start tyName.val indents ty
dataDeclBody : OriginDesc -> IndentInfo -> Rule PDataDecl
@ -1352,7 +1404,7 @@ constraints fname indents
<|> do decoratedSymbol fname "("
n <- decorate fname Bound name
decoratedSymbol fname ":"
tm <- expr pdef fname indents
tm <- typeExpr pdef fname indents
decoratedSymbol fname ")"
decoratedSymbol fname "=>"
more <- constraints fname indents
@ -1365,7 +1417,7 @@ implBinds fname indents
rig <- multiplicity fname
n <- decorate fname Bound name
decoratedSymbol fname ":"
tm <- expr pdef fname indents
tm <- typeExpr pdef fname indents
decoratedSymbol fname "}"
decoratedSymbol fname "->"
more <- implBinds fname indents
@ -1378,7 +1430,7 @@ ifaceParam fname indents
rig <- multiplicity fname
ns <- sepBy1 (decoratedSymbol fname ",") (decorate fname Bound name)
decoratedSymbol fname ":"
tm <- expr pdef fname indents
tm <- typeExpr pdef fname indents
decoratedSymbol fname ")"
pure (forget ns, (rig, tm))
<|> do n <- bounds (decorate fname Bound name)
@ -1449,7 +1501,7 @@ fieldDecl fname indents
n <- decorate fname Function name
pure n)
decoratedSymbol fname ":"
ty <- expr pdef fname indents
ty <- typeExpr pdef fname indents
pure (\fc : FC => map (\n => MkField fc doc rig p n ty) (forget ns)))
pure (b.val (boundToFC fname b))
@ -1898,7 +1950,7 @@ exprArgCmd parseCmd command doc = (names, ExprArg, doc, parse)
parse = do
symbol ":"
runParseCmd parseCmd
tm <- mustWork $ expr pdef (Virtual Interactive) init
tm <- mustWork $ typeExpr pdef (Virtual Interactive) init
pure (command tm)
declsArgCmd : ParseCmd -> (List PDecl -> REPLCmd) -> String -> CommandDefinition
@ -2045,7 +2097,7 @@ nonEmptyCommand =
eval : Rule REPLCmd
eval = do
tm <- expr pdef (Virtual Interactive) init
tm <- typeExpr pdef (Virtual Interactive) init
pure (Eval tm)
export

View File

@ -206,8 +206,12 @@ readHeader path origin
-- Stop at the first :, that's definitely not part of the header, to
-- save lexing the whole file unnecessarily
setCurrentElabSource res -- for error printing purposes
let Right (decor, mod) = runParserTo (PhysicalIdrSrc origin) (isLitFile path) (is ':') res (progHdr (PhysicalIdrSrc origin))
let Right (ws, decor, mod)
= runParserTo (PhysicalIdrSrc origin)
(isLitFile path) (is ':') res
(progHdr (PhysicalIdrSrc origin))
| Left err => throw err
traverse_ recordWarning ws
pure mod
%foreign "scheme:collect"
@ -310,9 +314,14 @@ processMod sourceFileName ttcFileName msg sourcecode origin
pure Nothing
else -- needs rebuilding
do iputStrLn msg
Right (decor, mod) <- logTime ("++ Parsing " ++ sourceFileName) $
pure (runParser (PhysicalIdrSrc origin) (isLitFile sourceFileName) sourcecode (do p <- prog (PhysicalIdrSrc origin); eoi; pure p))
| Left err => pure (Just [err])
Right (ws, decor, mod) <-
logTime ("++ Parsing " ++ sourceFileName) $
pure $ runParser (PhysicalIdrSrc origin)
(isLitFile sourceFileName)
sourcecode
(do p <- prog (PhysicalIdrSrc origin); eoi; pure p)
| Left err => pure (Just [err])
traverse_ recordWarning ws
addSemanticDecorations decor
initHash
traverse_ addPublicHash (sort importMetas)

View File

@ -925,7 +925,7 @@ process (ImportPackage package) = do
let packageDirPath = parse packageDir
tree <- coreLift $ explore packageDirPath
fentries <- coreLift $ toPaths (toRelative tree)
errs <- for fentries \entry => do
errs <- for fentries $ \entry => do
let entry' = dropExtension entry
let sp = forget $ split (== dirSeparator) entry'
let ns = concat $ intersperse "." sp
@ -977,7 +977,7 @@ parseRepl : String -> Either Error (Maybe REPLCmd)
parseRepl inp
= case runParser (Virtual Interactive) Nothing inp (parseEmptyCmd <|> parseCmd) of
Left err => Left err
Right (decor, result) => Right result
Right (_, _, result) => Right result
export
interpret : {auto c : Ref Ctxt Defs} ->

View File

@ -114,5 +114,5 @@ export
parseVersion : String -> Maybe Version
parseVersion str =
case parse versionParser (lexVersion str) of
Right (version, []) => Just version
Right (_, version, []) => Just version
_ => Nothing

View File

@ -9,6 +9,8 @@ import public Libraries.Text.Bounded
%default total
%hide Prelude.(>>)
-- TODO: Add some primitives for helping with error messages.
-- e.g. perhaps set a string to state what we're currently trying to
-- parse, or to say what the next expected token is in words
@ -25,6 +27,8 @@ data Grammar : (state : Type) -> (tok : Type) -> (consumes : Bool) -> Type -> Ty
EOF : Grammar state tok False ()
Fail : (location : Maybe Bounds) -> (fatal : Bool) -> String -> Grammar state tok c ty
Warning : (location : Maybe Bounds) -> String -> Grammar state tok False ()
Try : Grammar state tok c ty -> Grammar state tok c ty
Commit : Grammar state tok False ()
@ -179,6 +183,7 @@ mapToken f (Empty val) = Empty val
mapToken f (Terminal msg g) = Terminal msg (g . f)
mapToken f (NextIs msg g) = SeqEmpty (NextIs msg (g . f)) (Empty . f)
mapToken f EOF = EOF
mapToken f (Warning bd msg) = Warning bd msg
mapToken f (Fail bd fatal msg) = Fail bd fatal msg
mapToken f (Try g) = Try (mapToken f g)
mapToken f (MustWork g) = MustWork (mapToken f g)
@ -260,6 +265,15 @@ export %inline
mustWork : {c : Bool} -> Grammar state tok c ty -> Grammar state tok c ty
mustWork = MustWork
||| If the parser fails, treat it as a fatal error and explain why
export
mustWorkBecause :
{c : Bool} -> Bounds -> String ->
Grammar state tok c ty -> Grammar state tok c ty
mustWorkBecause {c} loc msg p
= rewrite sym (andSameNeutral c) in
p <|> fatalLoc loc msg
export %inline
bounds : Grammar state tok c ty -> Grammar state tok c (WithBounds ty)
bounds = Bounds
@ -268,107 +282,130 @@ export %inline
position : Grammar state tok False Bounds
position = Position
||| Warn the user
export %inline
withWarning : {c : _} -> String -> Grammar state tok c ty -> Grammar state tok c ty
withWarning warn p
= rewrite sym $ orFalseNeutral c in
SeqEmpty (bounds p) $ \ res =>
do Warning (Just res.bounds) warn
pure res.val
public export
data ParsingError tok = Error String (Maybe Bounds)
public export
ParsingWarnings : Type
ParsingWarnings = List (Maybe Bounds, String)
data ParseResult : Type -> Type -> Type -> Type where
Failure : (committed : Bool) -> (fatal : Bool) ->
List1 (ParsingError tok) -> ParseResult state tok ty
Res : state -> (committed : Bool) ->
(val : WithBounds ty) -> (more : List (WithBounds tok)) -> ParseResult state tok ty
Res : state ->
(ws : ParsingWarnings) ->
(committed : Bool) ->
(val : WithBounds ty) ->
(more : List (WithBounds tok)) ->
ParseResult state tok ty
mergeWith : WithBounds ty -> ParseResult state tok sy -> ParseResult state tok sy
mergeWith x (Res s committed val more) = Res s committed (mergeBounds x val) more
mergeWith x (Res s ws committed val more) = Res s ws committed (mergeBounds x val) more
mergeWith x v = v
doParse : Semigroup state => state -> (commit : Bool) ->
doParse : Semigroup state =>
state -> (ws : ParsingWarnings) ->
(commit : Bool) ->
(act : Grammar state tok c ty) ->
(xs : List (WithBounds tok)) ->
ParseResult state tok ty
doParse s com (Empty val) xs = Res s com (irrelevantBounds val) xs
doParse s com (Fail location fatal str) xs
doParse s ws com (Empty val) xs = Res s ws com (irrelevantBounds val) xs
doParse s ws com (Warning mb msg) xs = Res s ((mb, msg) :: ws) com (irrelevantBounds ()) xs
doParse s ws com (Fail location fatal str) xs
= Failure com fatal (Error str (location <|> (bounds <$> head' xs)) ::: Nil)
doParse s com (Try g) xs = case doParse s com g xs of
doParse s ws com (Try g) xs = case doParse s ws com g xs of
-- recover from fatal match but still propagate the 'commit'
Failure com _ errs => Failure com False errs
res => res
doParse s com Commit xs = Res s True (irrelevantBounds ()) xs
doParse s com (MustWork g) xs =
case assert_total (doParse s com g xs) of
doParse s ws com Commit xs = Res s ws True (irrelevantBounds ()) xs
doParse s ws com (MustWork g) xs =
case assert_total (doParse s ws com g xs) of
Failure com' _ errs => Failure com' True errs
res => res
doParse s com (Terminal err f) [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s com (Terminal err f) (x :: xs) =
doParse s ws com (Terminal err f) [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s ws com (Terminal err f) (x :: xs) =
case f x.val of
Nothing => Failure com False (Error err (Just x.bounds) ::: Nil)
Just a => Res s com (const a <$> x) xs
doParse s com EOF [] = Res s com (irrelevantBounds ()) []
doParse s com EOF (x :: xs) = Failure com False (Error "Expected end of input" (Just x.bounds) ::: Nil)
doParse s com (NextIs err f) [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s com (NextIs err f) (x :: xs)
Just a => Res s ws com (const a <$> x) xs
doParse s ws com EOF [] = Res s ws com (irrelevantBounds ()) []
doParse s ws com EOF (x :: xs) = Failure com False (Error "Expected end of input" (Just x.bounds) ::: Nil)
doParse s ws com (NextIs err f) [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s ws com (NextIs err f) (x :: xs)
= if f x.val
then Res s com (removeIrrelevance x) (x :: xs)
then Res s ws com (removeIrrelevance x) (x :: xs)
else Failure com False (Error err (Just x.bounds) ::: Nil)
doParse s com (Alt {c1} {c2} x y) xs
= case doParse s False x xs of
doParse s ws com (Alt {c1} {c2} x y) xs
= case doParse s ws False x xs of
Failure com' fatal errs
=> if com' || fatal
-- If the alternative had committed, don't try the
-- other branch (and reset commit flag)
then Failure com fatal errs
else case (assert_total doParse s False y xs) of
else case (assert_total doParse s ws False y xs) of
(Failure com'' fatal' errs') => if com'' || fatal'
-- Only add the errors together if the second branch
-- is also non-committed and non-fatal.
then Failure com'' fatal' errs'
else Failure False False (errs ++ errs')
(Res s _ val xs) => Res s com val xs
(Res s ws _ val xs) => Res s ws com val xs
-- Successfully parsed the first option, so use the outer commit flag
Res s _ val xs => Res s com val xs
doParse s com (SeqEmpty act next) xs
= case assert_total (doParse s com act xs) of
Res s ws _ val xs => Res s ws com val xs
doParse s ws com (SeqEmpty act next) xs
= case assert_total (doParse s ws com act xs) of
Failure com fatal errs => Failure com fatal errs
Res s com v xs =>
mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse s com (SeqEat act next) xs
= case assert_total (doParse s com act xs) of
Res s ws com v xs =>
mergeWith v (assert_total $ doParse s ws com (next v.val) xs)
doParse s ws com (SeqEat act next) xs
= case assert_total (doParse s ws com act xs) of
Failure com fatal errs => Failure com fatal errs
Res s com v xs =>
mergeWith v (assert_total $ doParse s com (next v.val) xs)
doParse s com (ThenEmpty act next) xs
= case assert_total (doParse s com act xs) of
Res s ws com v xs =>
mergeWith v (assert_total $ doParse s ws com (next v.val) xs)
doParse s ws com (ThenEmpty act next) xs
= case assert_total (doParse s ws com act xs) of
Failure com fatal errs => Failure com fatal errs
Res s com v xs =>
mergeWith v (assert_total $ doParse s com next xs)
doParse s com (ThenEat act next) xs
= case assert_total (doParse s com act xs) of
Res s ws com v xs =>
mergeWith v (assert_total $ doParse s ws com next xs)
doParse s ws com (ThenEat act next) xs
= case assert_total (doParse s ws com act xs) of
Failure com fatal errs => Failure com fatal errs
Res s com v xs =>
mergeWith v (assert_total $ doParse s com next xs)
doParse s com (Bounds act) xs
= case assert_total (doParse s com act xs) of
Res s ws com v xs =>
mergeWith v (assert_total $ doParse s ws com next xs)
doParse s ws com (Bounds act) xs
= case assert_total (doParse s ws com act xs) of
Failure com fatal errs => Failure com fatal errs
Res s com v xs => Res s com (const v <$> v) xs
doParse s com Position [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s com Position (x :: xs)
= Res s com (irrelevantBounds x.bounds) (x :: xs)
doParse s com (Act action) xs = Res (s <+> action) com (irrelevantBounds ()) xs
Res s ws com v xs => Res s ws com (const v <$> v) xs
doParse s ws com Position [] = Failure com False (Error "End of input" Nothing ::: Nil)
doParse s ws com Position (x :: xs)
= Res s ws com (irrelevantBounds x.bounds) (x :: xs)
doParse s ws com (Act action) xs
= Res (s <+> action) ws com (irrelevantBounds ()) xs
||| Parse a list of tokens according to the given grammar. If successful,
||| returns a pair of the parse result and the unparsed tokens (the remaining
||| input).
export
parse : {c : Bool} -> (act : Grammar () tok c ty) -> (xs : List (WithBounds tok)) ->
Either (List1 (ParsingError tok)) (ty, List (WithBounds tok))
Either (List1 (ParsingError tok))
(ParsingWarnings, ty, List (WithBounds tok))
parse act xs
= case doParse neutral False act xs of
= case doParse neutral [] False act xs of
Failure _ _ errs => Left errs
Res _ _ v rest => Right (v.val, rest)
Res _ ws _ v rest => Right (ws, v.val, rest)
export
parseWith : Monoid state => {c : Bool} -> (act : Grammar state tok c ty) -> (xs : List (WithBounds tok)) ->
Either (List1 (ParsingError tok)) (state, ty, List (WithBounds tok))
Either (List1 (ParsingError tok))
(state, ParsingWarnings, ty, List (WithBounds tok))
parseWith act xs
= case doParse neutral False act xs of
= case doParse neutral [] False act xs of
Failure _ _ errs => Left errs
Res s _ v rest => Right (s, v.val, rest)
Res s ws _ v rest => Right (s, ws, v.val, rest)

View File

@ -276,7 +276,7 @@ export
parse : String -> Path
parse str =
case parse parsePath (lexPath str) of
Right (path, _) => path
Right (_, path, _) => path
_ => emptyPath
--------------------------------------------------------------------------------

View File

@ -18,8 +18,8 @@ runParser : (fname : String) -> (str : String) -> Rule ty -> Either Error ty
runParser fname str p
= do toks <- mapFst (\err => fromLexError
(PhysicalPkgSrc fname) (NoRuleApply, err)) $ lex str
parsed <- mapFst (fromParsingErrors (PhysicalPkgSrc fname)) $ parse p toks
Right (fst parsed)
(_, val, _) <- mapFst (fromParsingErrors (PhysicalPkgSrc fname)) $ parse p toks
Right val
export
covering

View File

@ -19,56 +19,56 @@ EmptyRule = Grammar () Token False
export
equals : Rule ()
equals = terminal "Expected equals"
equals = terminal "Expected equals" $
\case
Equals => Just ()
_ => Nothing
export
lte : Rule ()
lte = terminal "Expected <="
lte = terminal "Expected <=" $
\case
LTE => Just ()
_ => Nothing
export
gte : Rule ()
gte = terminal "Expected >="
gte = terminal "Expected >=" $
\case
GTE => Just ()
_ => Nothing
export
lt : Rule ()
lt = terminal "Expected <="
lt = terminal "Expected <=" $
\case
LT => Just ()
_ => Nothing
export
gt : Rule ()
gt = terminal "Expected >="
gt = terminal "Expected >=" $
\case
GT => Just ()
_ => Nothing
export
eqop : Rule ()
eqop = terminal "Expected =="
eqop = terminal "Expected ==" $
\case
EqOp => Just ()
_ => Nothing
export
andop : Rule ()
andop = terminal "Expected &&"
andop = terminal "Expected &&" $
\case
AndOp => Just ()
_ => Nothing
export
eoi : Rule ()
eoi = terminal "Expected end of input"
eoi = terminal "Expected end of input" $
\case
EndOfInput => Just ()
_ => Nothing
@ -83,28 +83,28 @@ exactProperty p = terminal ("Expected property " ++ p) $
export
stringLit : Rule String
stringLit = terminal "Expected string"
stringLit = terminal "Expected string" $
\case
StringLit str => Just str
_ => Nothing
export
integerLit : Rule Integer
integerLit = terminal "Expected integer"
integerLit = terminal "Expected integer" $
\case
IntegerLit i => Just i
_ => Nothing
export
namespacedIdent : Rule (Maybe Namespace, String)
namespacedIdent = terminal "Expected namespaced identifier"
namespacedIdent = terminal "Expected namespaced identifier" $
\case
DotSepIdent ns n => Just (ns, n)
_ => Nothing
export
moduleIdent : Rule ModuleIdent
moduleIdent = terminal "Expected module identifier"
moduleIdent = terminal "Expected module identifier" $
\case
DotSepIdent ns m =>
Just $ nsAsModuleIdent $
@ -113,7 +113,7 @@ moduleIdent = terminal "Expected module identifier"
export
packageName : Rule String
packageName = terminal "Expected package name"
packageName = terminal "Expected package name" $
\case
DotSepIdent Nothing str =>
if isIdent AllowDashes str
@ -123,13 +123,13 @@ packageName = terminal "Expected package name"
export
dot' : Rule ()
dot' = terminal "Expected dot"
dot' = terminal "Expected dot" $
\case
Dot => Just ()
_ => Nothing
sep' : Rule ()
sep' = terminal "Expected separator"
sep' = terminal "Expected separator" $
\case
Separator => Just ()
_ => Nothing

View File

@ -37,18 +37,17 @@ eoi = ignore $ nextIs "Expected end of input" isEOI
export
constant : Rule Constant
constant
= terminal "Expected constant"
\case
CharLit c => Ch <$> getCharLit c
DoubleLit d => Just (Db d)
IntegerLit i => Just (BI i)
Ident s => isConstantType (UN s) >>=
\case WorldType => Nothing
c => Just c
_ => Nothing
= terminal "Expected constant" $ \case
CharLit c => Ch <$> getCharLit c
DoubleLit d => Just (Db d)
IntegerLit i => Just (BI i)
Ident s => isConstantType (UN s) >>=
\case WorldType => Nothing
c => Just c
_ => Nothing
documentation' : Rule String
documentation' = terminal "Expected documentation comment"
documentation' = terminal "Expected documentation comment" $
\case
DocComment d => Just d
_ => Nothing
@ -60,7 +59,7 @@ documentation = (unlines . forget) <$> some documentation'
export
intLit : Rule Integer
intLit
= terminal "Expected integer literal"
= terminal "Expected integer literal" $
\case
IntegerLit i => Just i
_ => Nothing
@ -68,7 +67,7 @@ intLit
export
onOffLit : Rule Bool
onOffLit
= terminal "Expected on or off"
= terminal "Expected on or off" $
\case
Ident "on" => Just True
Ident "off" => Just False
@ -77,7 +76,7 @@ onOffLit
export
strLit : Rule String
strLit
= terminal "Expected string literal"
= terminal "Expected string literal" $
\case
StringLit n s => escape n s
_ => Nothing
@ -86,7 +85,7 @@ strLit
export
strLitLines : Rule (List1 String)
strLitLines
= terminal "Expected string literal"
= terminal "Expected string literal" $
\case
StringLit n s =>
traverse (escape n . fastPack)
@ -95,35 +94,35 @@ strLitLines
export
strBegin : Rule ()
strBegin = terminal "Expected string begin"
strBegin = terminal "Expected string begin" $
\case
StringBegin Single => Just ()
_ => Nothing
export
multilineBegin : Rule ()
multilineBegin = terminal "Expected multiline string begin"
multilineBegin = terminal "Expected multiline string begin" $
\case
StringBegin Multi => Just ()
_ => Nothing
export
strEnd : Rule ()
strEnd = terminal "Expected string end"
strEnd = terminal "Expected string end" $
\case
StringEnd => Just ()
_ => Nothing
export
interpBegin : Rule ()
interpBegin = terminal "Expected string interp begin"
interpBegin = terminal "Expected string interp begin" $
\case
InterpBegin => Just ()
_ => Nothing
export
interpEnd : Rule ()
interpEnd = terminal "Expected string interp end"
interpEnd = terminal "Expected string interp end" $
\case
InterpEnd => Just ()
_ => Nothing
@ -134,7 +133,7 @@ simpleStr = strBegin *> commit *> (option "" strLit) <* strEnd
export
aDotIdent : Rule String
aDotIdent = terminal "Expected dot+identifier"
aDotIdent = terminal "Expected dot+identifier" $
\case
DotIdent s => Just s
_ => Nothing
@ -146,7 +145,7 @@ postfixProj = RF <$> aDotIdent
export
symbol : String -> Rule ()
symbol req
= terminal ("Expected '" ++ req ++ "'")
= terminal ("Expected '" ++ req ++ "'") $
\case
Symbol s => if s == req then Just () else Nothing
_ => Nothing
@ -154,7 +153,7 @@ symbol req
export
keyword : String -> Rule ()
keyword req
= terminal ("Expected '" ++ req ++ "'")
= terminal ("Expected '" ++ req ++ "'") $
\case
Keyword s => if s == req then Just () else Nothing
_ => Nothing
@ -162,7 +161,7 @@ keyword req
export
exactIdent : String -> Rule ()
exactIdent req
= terminal ("Expected " ++ req)
= terminal ("Expected " ++ req) $
\case
Ident s => if s == req then Just () else Nothing
_ => Nothing
@ -170,7 +169,7 @@ exactIdent req
export
pragma : String -> Rule ()
pragma n =
terminal ("Expected pragma " ++ n)
terminal ("Expected pragma " ++ n) $
\case
Pragma s =>
if s == n
@ -187,7 +186,7 @@ builtinType =
operatorCandidate : Rule Name
operatorCandidate
= terminal "Expected operator"
= terminal "Expected operator" $
\case
Symbol s => Just (UN s)
_ => Nothing
@ -195,7 +194,7 @@ operatorCandidate
export
operator : Rule Name
operator
= terminal "Expected operator"
= terminal "Expected operator" $
\case
Symbol s =>
if s `elem` reservedSymbols
@ -205,7 +204,7 @@ operator
identPart : Rule String
identPart
= terminal "Expected name"
= terminal "Expected name" $
\case
Ident str => Just str
_ => Nothing
@ -213,7 +212,7 @@ identPart
export
namespacedIdent : Rule (Maybe Namespace, String)
namespacedIdent
= terminal "Expected namespaced name"
= terminal "Expected namespaced name" $
\case
DotSepIdent ns n => Just (Just ns, n)
Ident i => Just (Nothing, i)
@ -248,7 +247,7 @@ unqualifiedName = identPart
export
holeName : Rule String
holeName
= terminal "Expected hole name"
= terminal "Expected hole name" $
\case
HoleIdent str => Just str
_ => Nothing

View File

@ -17,24 +17,29 @@ export
runParserTo : {e : _} ->
(origin : OriginDesc) ->
Maybe LiterateStyle -> Lexer ->
String -> Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty)
String -> Grammar SemanticDecorations Token e ty ->
Either Error (List Warning, SemanticDecorations, ty)
runParserTo origin lit reject str p
= do str <- mapFst (fromLitError origin) $ unlit lit str
toks <- mapFst (fromLexError origin) $ lexTo reject str
(decs, (parsed, _)) <- mapFst (fromParsingErrors origin) $ parseWith p toks
Right (decs, parsed)
(decs, ws, (parsed, _)) <- mapFst (fromParsingErrors origin) $ parseWith p toks
let ws = ws <&> \ (mb, warn) =>
let mkFC = \ b => MkFC origin (startBounds b) (endBounds b)
in ParserWarning (maybe EmptyFC mkFC mb) warn
Right (ws, decs, parsed)
export
runParser : {e : _} ->
(origin : OriginDesc) -> Maybe LiterateStyle -> String ->
Grammar SemanticDecorations Token e ty -> Either Error (SemanticDecorations, ty)
Grammar SemanticDecorations Token e ty ->
Either Error (List Warning, SemanticDecorations, ty)
runParser origin lit = runParserTo origin lit (pred $ const False)
export covering
parseFile : (fname : String)
-> (origin : OriginDesc)
-> Rule ty
-> IO (Either Error (SemanticDecorations, ty))
-> IO (Either Error (List Warning, SemanticDecorations, ty))
parseFile fname origin p
= do Right str <- readFile fname
| Left err => pure (Left (FileErr fname err))

View File

@ -241,7 +241,7 @@ checkTermSub defining mode opts nest env env' sub tm ty
catch {t = Error}
(elabTermSub defining mode opts nest
env env' sub tm (Just ty))
\case
$ \case
TryWithImplicits loc benv ns
=> do put Ctxt defs
put UST ust

View File

@ -59,7 +59,7 @@ getNameType rigc env fc x
$ "getNameType is trying to add Bound: "
++ show x ++ " (" ++ show fc ++ ")"
when (isSourceName x) $
whenJust (isConcreteFC fc) \nfc => do
whenJust (isConcreteFC fc) $ \nfc => do
log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show x
addSemanticDecorations [(nfc, Bound, Just x)]
@ -82,7 +82,7 @@ getNameType rigc env fc x
++ show def.fullname ++ " (" ++ show fc ++ ")"
when (isSourceName def.fullname) $
whenJust (isConcreteFC fc) \nfc => do
whenJust (isConcreteFC fc) $ \nfc => do
let decor = nameTypeDecoration nt
log "ide-mode.highlight" 7
$ "getNameType is adding " ++ show decor ++ ": " ++ show def.fullname
@ -131,7 +131,7 @@ getVarType rigc nest env fc x
addNameType fc x env tyenv
when (isSourceName ndef.fullname) $
whenJust (isConcreteFC fc) \nfc => do
whenJust (isConcreteFC fc) $ \nfc => do
let decor = nameTypeDecoration nt
log "ide-mode.highlight" 7
$ "getNameType is adding "++ show decor ++": "

View File

@ -392,7 +392,7 @@ checkCase rig elabinfo nest env fc scr scrty_in alts exp
(scrtm_in, gscrty, caseRig) <- handle
(do c <- runDelays (const True) $ check chrig elabinfo nest env scr (Just (gnf env scrtyv))
pure (fst c, snd c, chrig))
\case
$ \case
e@(LinearMisuse _ _ r _)
=> branchOne
(do c <- runDelays (const True) $ check linear elabinfo nest env scr

View File

@ -418,7 +418,7 @@ checkBindVar rig elabinfo nest env fc str topexp
notePatVar n
est <- get EST
whenJust (isConcreteFC fc) \nfc => do
whenJust (isConcreteFC fc) $ \nfc => do
log "ide-mode.highlight" 7 $ "getNameType is adding Bound: " ++ show n
addSemanticDecorations [(nfc, Bound, Just n)]

View File

@ -49,7 +49,7 @@ checkRetType env nf chk = chk nf
checkIsType : {auto c : Ref Ctxt Defs} ->
FC -> Name -> Env Term vars -> NF vars -> Core ()
checkIsType loc n env nf
= checkRetType env nf
= checkRetType env nf $
\case
NType _ => pure ()
_ => throw $ BadTypeConType loc n
@ -57,7 +57,7 @@ checkIsType loc n env nf
checkFamily : {auto c : Ref Ctxt Defs} ->
FC -> Name -> Name -> Env Term vars -> NF vars -> Core ()
checkFamily loc cn tn env nf
= checkRetType env nf
= checkRetType env nf $
\case
NType _ => throw $ BadDataConType loc cn tn
NTCon _ n' _ _ _ =>
@ -235,7 +235,7 @@ findNewtype [con]
= do defs <- get Ctxt
Just arg <- getRelevantArg defs 0 Nothing True !(nf defs [] (type con))
| Nothing => pure ()
updateDef (name con)
updateDef (name con) $
\case
DCon t a _ => Just $ DCon t a $ Just arg
_ => Nothing

View File

@ -170,13 +170,14 @@ processTTImpFile : {auto c : Ref Ctxt Defs} ->
String -> Core Bool
processTTImpFile fname
= do modIdent <- ctxtPathToNS fname
Right (decor, tti) <- logTime "Parsing" $
Right (ws, decor, tti) <- logTime "Parsing" $
coreLift $ parseFile fname (PhysicalIdrSrc modIdent)
(do decls <- prog (PhysicalIdrSrc modIdent)
eoi
pure decls)
| Left err => do coreLift (putStrLn (show err))
pure False
traverse_ recordWarning ws
logTime "Elaboration" $
catch (do ignore $ processTTImpDecls (MkNested []) [] tti
Nothing <- checkDelayedHoles

View File

@ -153,4 +153,4 @@ repl
case runParser (Virtual Interactive) Nothing inp command of
Left err => do coreLift_ (printLn err)
repl
Right (decor, cmd) => when !(processCatch cmd) repl
Right (_, _, cmd) => when !(processCatch cmd) repl

View File

@ -70,7 +70,7 @@ idrisTestsError = MkTestPool "Error messages" [] Nothing
"error016", "error017", "error018", "error019", "error020",
-- Parse errors
"perror001", "perror002", "perror003", "perror004", "perror005",
"perror006", "perror007", "perror008", "perror009"]
"perror006", "perror007", "perror008", "perror009", "perror010"]
idrisTestsInteractive : TestPool
idrisTestsInteractive = MkTestPool "Interactive editing" [] Nothing

View File

@ -6,10 +6,10 @@ f p = case (p 0, p 1) of
(True , True) => 4
il : Int
il = f \x => x > 0
il = f $ \x => x > 0
lc : Int
lc = f $ \case 0 => True ; _ => False
ilc : Int
ilc = f \case 0 => True; _ => False
ilc = f (\case 0 => True; _ => False)

View File

@ -21,15 +21,14 @@ Bad1:3:1--3:5
^^^^
1/1: Building Bad2 (Bad2.idr)
Error: Couldn't parse any alternatives:
1: Expected 'case', 'if', 'do', application or operator expression.
Error: Cannot return a named argument.
Bad2:3:13--3:14
Bad2:3:13--3:29
1 | module Bad2
2 |
3 | badReturn : (whatever : Int)
^
... (53 others)
^^^^^^^^^^^^^^^^
1/1: Building Bad3 (Bad3.idr)
Error: Couldn't parse declaration.

View File

@ -6,7 +6,7 @@ StrError1:2:24--2:25
1 | foo : String
2 | foo = "empty interp: \{}"
^
... (41 others)
... (40 others)
1/1: Building StrError2 (StrError2.idr)
Error: Couldn't parse any alternatives:
1: Expected 'case', 'if', 'do', application or operator expression.
@ -24,7 +24,7 @@ StrError3:2:7--2:8
1 | foo : String
2 | foo = "incomplete: \{ a <+> }"
^
... (28 others)
... (29 others)
1/1: Building StrError4 (StrError4.idr)
Error: Bracket is not properly closed.

View File

@ -0,0 +1,2 @@
not : (arg1 : Bool) -> (arg2 : Bool)
not = ?a

View File

@ -0,0 +1,2 @@
not : (arg1 : Bool) -> {arg2 : Bool}
not = ?a

View File

@ -0,0 +1,2 @@
not : (arg1 : Bool) -> {auto arg2 : Bool}
not = ?a

View File

@ -0,0 +1,2 @@
not : (arg1 : Bool) -> {default arg2 True : Bool}
not = ?a

View File

@ -0,0 +1,6 @@
f : (Nat -> Nat) -> Nat
f k = f \n => k (S n)
g : Nat -> Nat
g x = f ?a

View File

@ -0,0 +1,40 @@
1/1: Building NamedReturn1 (NamedReturn1.idr)
Error: Cannot return a named argument.
NamedReturn1:1:24--1:37
1 | not : (arg1 : Bool) -> (arg2 : Bool)
^^^^^^^^^^^^^
1/1: Building NamedReturn2 (NamedReturn2.idr)
Error: Cannot return an implicit argument.
NamedReturn2:1:24--1:37
1 | not : (arg1 : Bool) -> {arg2 : Bool}
^^^^^^^^^^^^^
1/1: Building NamedReturn3 (NamedReturn3.idr)
Error: Cannot return an auto implicit argument.
NamedReturn3:1:24--1:42
1 | not : (arg1 : Bool) -> {auto arg2 : Bool}
^^^^^^^^^^^^^^^^^^
1/1: Building NamedReturn4 (NamedReturn4.idr)
Error: Cannot return a default implicit argument.
NamedReturn4:1:24--1:50
1 | not : (arg1 : Bool) -> {default arg2 True : Bool}
^^^^^^^^^^^^^^^^^^^^^^^^^^
1/1: Building TrailingLam (TrailingLam.idr)
Warning: DEPRECATED: trailing lambda. Use a $ or parens
TrailingLam:3:9--3:22
1 |
2 | f : (Nat -> Nat) -> Nat
3 | f k = f \n => k (S n)
^^^^^^^^^^^^^
Main> 1 hole: Main.a
Main>
Bye for now!

View File

@ -0,0 +1,2 @@
:m
:q

7
tests/idris2/perror010/run Executable file
View File

@ -0,0 +1,7 @@
rm -rf build
$1 --no-color --console-width 0 --check NamedReturn1.idr || true
$1 --no-color --console-width 0 --check NamedReturn2.idr || true
$1 --no-color --console-width 0 --check NamedReturn3.idr || true
$1 --no-color --console-width 0 --check NamedReturn4.idr || true
$1 --no-color --console-width 0 --no-banner TrailingLam.idr < input