@@ -369,7 +369,7 @@ mutual
369
369
toPTerm p (IAlternative fc _ _ ) = pure (PImplicit fc)
370
370
toPTerm p (IRewrite fc rule tm)
371
371
= pure (PRewrite fc ! (toPTerm startPrec rule)
372
- ! (toPTerm startPrec tm))
372
+ ! (toPTerm startPrec tm))
373
373
toPTerm p (ICoerced fc tm) = toPTerm p tm
374
374
toPTerm p (IPrimVal fc c) = pure (PPrimVal fc c)
375
375
toPTerm p (IHole fc str) = pure (PHole fc False str)
@@ -378,19 +378,18 @@ mutual
378
378
= let nm = UN (Basic v) in
379
379
pure (PRef fc (MkKindedName (Just Bound ) nm nm))
380
380
toPTerm p (IBindHere fc _ tm) = toPTerm p tm
381
- toPTerm p (IAs fc nameFC _ n pat) = pure ( PAs fc nameFC n ! ( toPTerm argPrec pat))
382
- toPTerm p (IMustUnify fc r pat) = pure ( PDotted fc ! ( toPTerm argPrec pat))
381
+ toPTerm p (IAs fc nameFC _ n pat) = PAs fc nameFC n <$> toPTerm argPrec pat
382
+ toPTerm p (IMustUnify fc r pat) = PDotted fc <$> toPTerm argPrec pat
383
383
384
- toPTerm p (IDelayed fc r ty) = pure ( PDelayed fc r ! ( toPTerm argPrec ty))
385
- toPTerm p (IDelay fc tm) = pure ( PDelay fc ! ( toPTerm argPrec tm))
386
- toPTerm p (IForce fc tm) = pure ( PForce fc ! ( toPTerm argPrec tm))
387
- toPTerm p (IQuote fc tm) = pure ( PQuote fc ! ( toPTerm argPrec tm))
384
+ toPTerm p (IDelayed fc r ty) = PDelayed fc r <$> toPTerm argPrec ty
385
+ toPTerm p (IDelay fc tm) = PDelay fc <$> toPTerm argPrec tm
386
+ toPTerm p (IForce fc tm) = PForce fc <$> toPTerm argPrec tm
387
+ toPTerm p (IQuote fc tm) = PQuote fc <$> toPTerm argPrec tm
388
388
toPTerm p (IQuoteName fc n) = pure (PQuoteName fc n)
389
389
toPTerm p (IQuoteDecl fc ds)
390
- = do ds' <- traverse toPDecl ds
391
- pure $ PQuoteDecl fc (catMaybes ds')
392
- toPTerm p (IUnquote fc tm) = pure (PUnquote fc ! (toPTerm argPrec tm))
393
- toPTerm p (IRunElab fc _ tm) = pure (PRunElab fc ! (toPTerm argPrec tm))
390
+ = PQuoteDecl fc . catMaybes <$> traverse toPDecl ds
391
+ toPTerm p (IUnquote fc tm) = PUnquote fc <$> toPTerm argPrec tm
392
+ toPTerm p (IRunElab fc _ tm) = PRunElab fc <$> toPTerm argPrec tm
394
393
395
394
toPTerm p (IUnifyLog fc _ tm) = toPTerm p tm
396
395
toPTerm p (Implicit fc True ) = pure (PImplicit fc)
@@ -468,22 +467,22 @@ mutual
468
467
flags
469
468
! (traverse toPClause cs)
470
469
toPClause (ImpossibleClause fc lhs)
471
- = pure ( MkImpossible fc ! ( toPTerm startPrec lhs))
470
+ = MkImpossible fc <$> toPTerm startPrec lhs
472
471
473
472
toPTypeDecl : {auto c : Ref Ctxt Defs} ->
474
473
{auto s : Ref Syn SyntaxInfo} ->
475
474
ImpTy' KindedName -> Core (PTypeDecl' KindedName)
476
475
toPTypeDecl (MkImpTy fc n ty)
477
- = pure ( MkFCVal fc $ MkPTy (pure (" " , n)) " " ! ( toPTerm startPrec ty))
476
+ = MkFCVal fc . MkPTy (pure (" " , n)) " " <$> toPTerm startPrec ty
478
477
479
478
toPData : {auto c : Ref Ctxt Defs} ->
480
479
{auto s : Ref Syn SyntaxInfo} ->
481
480
ImpData' KindedName -> Core (PDataDecl' KindedName)
482
481
toPData (MkImpData fc n ty opts cs)
483
482
= pure (MkPData fc n ! (traverseOpt (toPTerm startPrec) ty) opts
484
- ! (traverse toPTypeDecl cs))
483
+ ! (traverse toPTypeDecl cs))
485
484
toPData (MkImpLater fc n ty)
486
- = pure ( MkPLater fc n ! ( toPTerm startPrec ty))
485
+ = MkPLater fc n <$> toPTerm startPrec ty
487
486
488
487
toPField : {auto c : Ref Ctxt Defs} ->
489
488
{auto s : Ref Syn SyntaxInfo} ->
@@ -528,11 +527,11 @@ mutual
528
527
ImpDecl' KindedName -> Core (Maybe (PDecl' KindedName))
529
528
toPDecl (IClaim (MkFCVal fc $ MkIClaimData rig vis opts ty))
530
529
= do opts' <- traverse toPFnOpt opts
531
- pure ( Just ( PClaim ( MkFCVal fc $ MkPClaim rig vis opts' ! ( toPTypeDecl ty))))
530
+ Just . PClaim . MkFCVal fc . MkPClaim rig vis opts' <$> toPTypeDecl ty
532
531
toPDecl (IData fc vis mbtot d)
533
- = pure ( Just ( PData fc " " vis mbtot ! ( toPData d)))
532
+ = Just . PData fc " " vis mbtot <$> toPData d
534
533
toPDecl (IDef fc n cs)
535
- = pure ( Just ( PDef $ MkFCVal fc ! ( traverse toPClause cs)))
534
+ = Just . PDef . MkFCVal fc <$> traverse toPClause cs
536
535
toPDecl (IParameters fc ps ds)
537
536
= do ds' <- traverse toPDecl ds
538
537
pure (Just (PParameters fc
@@ -545,17 +544,15 @@ mutual
545
544
= do (n, ps, opts, con, fs) <- toPRecord r
546
545
pure (Just (PRecord fc " " vis mbtot (MkPRecord n ps opts con fs)))
547
546
toPDecl (IFail fc msg ds)
548
- = do ds' <- traverse toPDecl ds
549
- pure (Just (PFail fc msg (catMaybes ds')))
547
+ = Just . PFail fc msg . catMaybes <$> traverse toPDecl ds
550
548
toPDecl (INamespace fc ns ds)
551
- = do ds' <- traverse toPDecl ds
552
- pure (Just (PNamespace fc ns (catMaybes ds')))
549
+ = Just . PNamespace fc ns . catMaybes <$> traverse toPDecl ds
553
550
toPDecl (ITransform fc n lhs rhs)
554
551
= pure (Just (PTransform fc (show n)
555
552
! (toPTerm startPrec lhs)
556
553
! (toPTerm startPrec rhs)))
557
554
toPDecl (IRunElabDecl fc tm)
558
- = pure ( Just ( PRunElabDecl fc ! ( toPTerm startPrec tm)))
555
+ = Just . PRunElabDecl fc <$> toPTerm startPrec tm
559
556
toPDecl (IPragma _ _ _ ) = pure Nothing
560
557
toPDecl (ILog _ ) = pure Nothing
561
558
toPDecl (IBuiltin fc type name) = pure $ Just $ PBuiltin fc type name
0 commit comments