fix printing of backticked operators

This commit is contained in:
André Videla 2024-04-04 16:08:23 +01:00
parent 2823281af6
commit 3b0d9d9f35
4 changed files with 17 additions and 1 deletions

View File

@ -38,6 +38,7 @@ mkOp : {auto s : Ref Syn SyntaxInfo} ->
mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y)
= do syn <- get Syn
let raw = rawName kn
let pop = if isOpName raw then OpSymbols else Backticked
-- to check if the name is an operator we use the root name as a basic
-- user name. This is because if the name is qualified with the namespace
-- looking the fixity context will fail. A qualified operator would look
@ -46,7 +47,7 @@ mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y)
-- to know if the name is an operator or not, it's enough to check
-- that the fixity context contains the name `(++)`
let rootName = UN (Basic (nameRoot raw))
let asOp = POp fc opFC (NoBinder (unbracketApp x)) (OpSymbols kn) (unbracketApp y)
let asOp = POp fc opFC (NoBinder (unbracketApp x)) (pop kn) (unbracketApp y)
if not (null (lookupName rootName (infixes syn)))
then pure asOp
else case dropNS raw of

View File

@ -0,0 +1,10 @@
import Language.Reflection
%default total
%language ElabReflection
export infix 6 `op`
%runElab logSugaredTerm "debug" 0 "term with infix" `(a `op` b)

View File

@ -0,0 +1,2 @@
1/1: Building Test (Test.idr)
LOG debug:0: term with infix: a `op` b

View File

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