[ fix #3097 ] Fix issues parsing %logging followed by named impls (#3098)

Co-authored-by: G. Allais <guillaume.allais@ens-lyon.org>
This commit is contained in:
Steve Dunham 2023-10-13 11:02:58 -07:00 committed by GitHub
parent 419a440dad
commit c04404a95b
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
7 changed files with 15 additions and 18 deletions

View File

@ -979,9 +979,6 @@ mutual
uds' <- traverse (desugarDecl ps) uds
update Syn { usingImpl := oldu }
pure (concat uds')
desugarDecl ps (PReflect fc tm)
= throw (GenericMsg fc "Reflection not implemented yet")
-- pure [IReflect fc !(desugar AnyExpr ps tm)]
desugarDecl ps (PInterface fc vis cons_in tn doc params det conname body)
= do addDocString tn doc
let paramNames = map fst params

View File

@ -1835,14 +1835,8 @@ fixDecl fname indents
directiveDecl : OriginDesc -> IndentInfo -> Rule PDecl
directiveDecl fname indents
= do b <- bounds ((do d <- directive fname indents
pure (\fc : FC => PDirective fc d))
<|>
(do decoratedPragma fname "runElab"
tm <- expr pdef fname indents
atEnd indents
pure (\fc : FC => PReflect fc tm)))
pure (b.val (boundToFC fname b))
= do b <- bounds (directive fname indents)
pure (PDirective (boundToFC fname b) b.val)
-- Declared at the top
-- topDecl : OriginDesc -> IndentInfo -> Rule (List PDecl)
@ -1855,6 +1849,8 @@ topDecl fname indents
pure [d]
<|> do ds <- claims fname indents
pure (forget ds)
<|> do d <- directiveDecl fname indents
pure [d]
<|> do d <- implDecl fname indents
pure [d]
<|> do d <- definition fname indents
@ -1880,8 +1876,6 @@ topDecl fname indents
pure [d]
<|> do d <- transformDecl fname indents
pure [d]
<|> do d <- directiveDecl fname indents
pure [d]
<|> do dstr <- bounds (terminal "Expected CG directive"
(\case
CGDirective d => Just d

View File

@ -434,7 +434,6 @@ mutual
List (PDecl' nm) -> PDecl' nm
PUsing : FC -> List (Maybe Name, PTerm' nm) ->
List (PDecl' nm) -> PDecl' nm
PReflect : FC -> PTerm' nm -> PDecl' nm
PInterface : FC ->
Visibility ->
(constraints : List (Maybe Name, PTerm' nm)) ->
@ -482,7 +481,6 @@ mutual
getPDeclLoc (PData fc _ _ _ _) = fc
getPDeclLoc (PParameters fc _ _) = fc
getPDeclLoc (PUsing fc _ _) = fc
getPDeclLoc (PReflect fc _) = fc
getPDeclLoc (PInterface fc _ _ _ _ _ _ _ _) = fc
getPDeclLoc (PImplementation fc _ _ _ _ _ _ _ _ _ _) = fc
getPDeclLoc (PRecord fc _ _ _ _) = fc
@ -1055,7 +1053,6 @@ Show PDecl where
show (PData{}) = "PData"
show (PParameters{}) = "PParameters"
show (PUsing{}) = "PUsing"
show (PReflect{}) = "PReflect"
show (PInterface{}) = "PInterface"
show (PImplementation{}) = "PImplementation"
show (PRecord{}) = "PRecord"

View File

@ -240,7 +240,6 @@ mapPTermM f = goPTerm where
goPDecl (PUsing fc mnts ps) =
PUsing fc <$> goPairedPTerms mnts
<*> goPDecls ps
goPDecl (PReflect fc t) = PReflect fc <$> goPTerm t
goPDecl (PInterface fc v mnts n doc nrts ns mn ps) =
PInterface fc v <$> goPairedPTerms mnts
<*> pure n
@ -523,7 +522,6 @@ mapPTerm f = goPTerm where
= PParameters fc (go4TupledPTerms nts) (goPDecl <$> ps)
goPDecl (PUsing fc mnts ps)
= PUsing fc (goPairedPTerms mnts) (goPDecl <$> ps)
goPDecl (PReflect fc t) = PReflect fc $ goPTerm t
goPDecl (PInterface fc v mnts n doc nrts ns mn ps)
= PInterface fc v (goPairedPTerms mnts) n doc (go3TupledPTerms nrts) ns mn (goPDecl <$> ps)
goPDecl (PImplementation fc v opts p is cs n ts mn ns mps)

View File

@ -0,0 +1,7 @@
import Data.Vect
%logging "elab" 1
[Named] {n : Nat} -> Show (Vect n a) where
show x = "whatever"

View File

@ -0,0 +1 @@
1/1: Building LoggingParse (LoggingParse.idr)

View File

@ -0,0 +1,3 @@
. ../../../testutils.sh
check LoggingParse.idr