Emit warning for fixities with no export modifiers (#3234)

* Emit warning for fixities with no export modifiers

This is to help update all the existing code to program with explicit
fixity export directives in preparation for the behavioral change where
they will become private by default.
This commit is contained in:
André Videla 2024-04-03 15:41:57 +01:00 committed by GitHub
parent aa3f67cd11
commit 75032a7164
No known key found for this signature in database
GPG Key ID: B5690EEEBB952194
128 changed files with 417 additions and 303 deletions

View File

@ -30,6 +30,9 @@ This CHANGELOG describes the merged but unreleased changes. Please see [CHANGELO
customise the syntax of operator to look more like a binder.
See [#3113](https://github.com/idris-lang/Idris2/issues/3113).
* Fixity declarations without an export modifier now emit a warning in peparation
for a future version where they will become private by default.
* Elaborator scripts were made to be able to access the visibility modifier of a
definition, via `getVis`.

View File

@ -5,10 +5,10 @@ import Data.Vect
%default total
infixl 8 `shiftL`, `shiftR`
infixl 7 .&.
infixl 6 `xor`
infixl 5 .|.
export infixl 8 `shiftL`, `shiftR`
export infixl 7 .&.
export infixl 6 `xor`
export infixl 5 .|.
--------------------------------------------------------------------------------
-- Interface Bits

View File

@ -2,7 +2,7 @@ module Data.Contravariant
%default total
infixl 4 >$, >$<, >&<, $<
export infixl 4 >$, >$<, >&<, $<
||| Contravariant functors
public export

View File

@ -304,7 +304,7 @@ namespace Equality
FZ : Pointwise FZ FZ
FS : Pointwise k l -> Pointwise (FS k) (FS l)
infix 6 ~~~
export infix 6 ~~~
||| Convenient infix notation for the notion of pointwise equality of Fins
public export
(~~~) : Fin m -> Fin n -> Type

View File

@ -270,7 +270,7 @@ public export
deleteFirstsBy : (p : a -> b -> Bool) -> (source : List b) -> (undesirables : List a) -> List b
deleteFirstsBy p = foldl (flip (deleteBy p))
infix 7 \\
export infix 7 \\
||| The non-associative list-difference.
||| A specialized form of @deleteFirstsBy@ where the predicate is equality under the @Eq@

View File

@ -5,7 +5,7 @@ import public Control.Function
%default total
infixr 7 :::
export infixr 7 :::
||| Non-empty lists.
public export

View File

@ -9,7 +9,7 @@ record Morphism a b where
constructor Mor
applyMor : a -> b
infixr 1 ~>
export infixr 1 ~>
public export
(~>) : Type -> Type -> Type

View File

@ -21,7 +21,7 @@ interface Zippable z where
zip : z a -> z b -> z (a, b)
zip = zipWith (,)
infixr 6 `zip`
export infixr 6 `zip`
||| Combine three parameterised types by applying a function.
||| @ z the parameterised type

View File

@ -2,9 +2,7 @@
||| poor-man's equational reasoning
module Syntax.PreorderReasoning
infixl 0 ~~,~=
prefix 1 |~
infix 1 ...,..<,..>,.=.,.=<,.=>
import public Syntax.PreorderReasoning.Ops
|||Slightly nicer syntax for justifying equations:
|||```

View File

@ -2,10 +2,7 @@ module Syntax.PreorderReasoning.Generic
import Control.Relation
import Control.Order
infixl 0 ~~, ~=
infixl 0 <~
prefix 1 |~
infix 1 ...,..<,..>,.=.,.=<,.=>
import public Syntax.PreorderReasoning.Ops
public export
data Step : (leq : a -> a -> Type) -> a -> a -> Type where

View File

@ -0,0 +1,6 @@
module Syntax.PreorderReasoning.Ops
export infixl 0 ~~, ~=
export infixl 0 <~
export prefix 1 |~
export infix 1 ...,..<,..>,.=.,.=<,.=>

View File

@ -108,6 +108,7 @@ modules = Control.App,
Language.Reflection.TTImp,
Syntax.PreorderReasoning,
Syntax.PreorderReasoning.Ops,
Syntax.PreorderReasoning.Generic,
System,

View File

@ -1,7 +1,7 @@
module Control.Algebra
infixl 6 <->
infixl 7 <.>
export infixl 6 <->
export infixl 7 <.>
public export
interface Semigroup ty => SemigroupV ty where

View File

@ -5,11 +5,11 @@ import Data.Either
import Data.Morphisms
infixr 5 <++>
infixr 3 ***
infixr 3 &&&
infixr 2 +++
infixr 2 \|/
export infixr 5 <++>
export infixr 3 ***
export infixr 3 &&&
export infixr 2 +++
export infixr 2 \|/
public export
interface Category arr => Arrow (0 arr : Type -> Type -> Type) where

View File

@ -20,7 +20,7 @@ Monad m => Category (Kleislimorphism m) where
id = Kleisli (pure . id)
(Kleisli f) . (Kleisli g) = Kleisli $ \a => g a >>= f
infixr 1 >>>
export infixr 1 >>>
public export
(>>>) : Category cat => cat a b -> cat b c -> cat a c

View File

@ -96,7 +96,7 @@ export
fail : Applicative m => String -> ValidatorT m a b
fail s = MkValidator $ \_ => left s
infixl 2 >>>
export infixl 2 >>>
||| Compose two validators so that the second validates the output of the first.
export

View File

@ -3,8 +3,8 @@ module Data.List.Alternating
import Data.Bifoldable
import Data.List
infixl 5 +>
infixr 5 <+
export infixl 5 +>
export infixr 5 <+
%default total

View File

@ -23,7 +23,7 @@ public export
modular : Monoid a => a -> Nat -> a
modular v n = modularRec v (halfRec n)
infixr 10 ^
export infixr 10 ^
public export
(^) : Num a => a -> Nat -> a
(^) = modular @{Multiplicative}

View File

@ -342,7 +342,7 @@ viewRTree (Deep s pr m (Four w x y z)) =
-- Construction
infixr 5 `consTree`
export infixr 5 `consTree`
export
consTree : Sized e => (r : e) -> FingerTree e -> FingerTree e
a `consTree` Empty = Single a
@ -352,7 +352,7 @@ a `consTree` Deep s (Two b c) st d2 = Deep (size a + s) (Three a b c) st d2
a `consTree` Deep s (Three b c d) st d2 = Deep (size a + s) (Four a b c d) st d2
a `consTree` Deep s (Four b c d f) st d2 = Deep (size a + s) (Two a b) (node3 c d f `consTree` st) d2
infixl 5 `snocTree`
export infixl 5 `snocTree`
export
snocTree : Sized e => FingerTree e -> (r : e) -> FingerTree e
Empty `snocTree` a = Single a

View File

@ -48,13 +48,13 @@ export
reverse : Seq n a -> Seq n a
reverse (MkSeq tr) = MkSeq (reverseTree id tr)
infixr 5 `cons`
export infixr 5 `cons`
||| O(1). Add an element to the left end of a sequence.
export
cons : e -> Seq n e -> Seq (S n) e
a `cons` MkSeq tr = MkSeq (MkElem a `consTree` tr)
infixl 5 `snoc`
export infixl 5 `snoc`
||| O(1). Add an element to the right end of a sequence.
export
snoc : Seq n e -> e -> Seq (S n) e

View File

@ -40,13 +40,13 @@ export
reverse : Seq a -> Seq a
reverse (MkSeq tr) = MkSeq (reverseTree id tr)
infixr 5 `cons`
export infixr 5 `cons`
||| O(1). Add an element to the left end of a sequence.
export
cons : e -> Seq e -> Seq e
a `cons` MkSeq tr = MkSeq (MkElem a `consTree` tr)
infixl 5 `snoc`
export infixl 5 `snoc`
||| O(1). Add an element to the right end of a sequence.
export
snoc : Seq e -> e -> Seq e

View File

@ -5,8 +5,8 @@ import Data.String
%default total
infixl 5 +>
infixr 5 <+
export infixl 5 +>
export infixr 5 <+
||| Adds a character to the end of the specified string.
|||

View File

@ -106,7 +106,7 @@ export
Fail i _ => Fail i msg)
(p.runParser s)
infixl 0 <?>
export infixl 0 <?>
||| Discards the result of a parser
export

View File

@ -59,7 +59,7 @@ toTelescope seg = untabulate seg ()
%name Segment delta,delta',delta1,delta2
infixl 3 |++, :++
export infixl 3 |++, :++
||| This lemma comes up all the time when mixing induction on Nat with
||| indexing modulo addition. An alternative is to use something like
@ -141,7 +141,7 @@ projection {n = S n} {delta = ty :: delta} env
$ rewrite succLemma n k in env
in env'
infixl 4 .=
export infixl 4 .=
public export
data Environment : (env : Left.Environment gamma)

View File

@ -30,8 +30,8 @@ plusAccZeroRightNeutral m =
Refl
infixl 4 -.
infixr 4 .-
export infixl 4 -.
export infixr 4 .-
namespace Left
@ -158,7 +158,7 @@ namespace Right
namespace Tree
infixl 4 ><
export infixl 4 ><
mutual
||| A tree of dependent types
@ -189,14 +189,14 @@ namespace Tree
(transpL env1 ** snd (concat (delta (transpL env1))) env2)
)
infix 5 <++>
export infix 5 <++>
public export
(<++>) : (gamma : Left.Telescope m) -> (Environment gamma -> Right.Telescope n) -> Right.Telescope (plusAcc m n)
[] <++> delta = delta ()
(gamma -. sigma ) <++> delta = gamma <++> (\ env => sigma env .- \ v => delta (env ** v))
infix 5 >++<
export infix 5 >++<
(>++<) : {m, n : Nat} -> (gamma : Right.Telescope m) -> (Environment gamma -> Left.Telescope n) ->
Left.Telescope (plusAcc m n)

View File

@ -1,6 +1,6 @@
module Syntax.WithProof
prefix 10 @@
export prefix 10 @@
||| Until Idris2 supports the 'with (...) proof p' construct, here's a
||| poor-man's replacement.

View File

@ -13,8 +13,8 @@ import Text.Quantity
import System.Info
infixl 5 </>, />
infixr 7 <.>
export infixl 5 </>, />
export infixr 7 <.>
||| The character that separates directories in the path.

View File

@ -128,7 +128,7 @@ export %inline
Grammar state tok (c1 && c2) ty
(<|>) = Alt
infixr 2 <||>
export infixr 2 <||>
||| Take the tagged disjunction of two grammars. If both consume, the
||| combination is guaranteed to consume.
export

View File

@ -122,7 +122,7 @@ export
fill : Int -> Doc ann -> Doc ann
fill n doc = width doc (\w => spaces $ n - w)
infixr 6 <++>
export infixr 6 <++>
||| Concatenates two documents with a space in between.
export
(<++>) : Doc ann -> Doc ann -> Doc ann

View File

@ -6,7 +6,7 @@ import Data.Nat
%default total
infix 1 `Copies`
export infix 1 `Copies`
||| Carries multiple linear copies of the same value. Usage: m `Copies` x
||| reads as "m copies of value x". This data structure is necessary to implement

View File

@ -35,7 +35,7 @@ Consumable Int where
-- But crucially we don't have `Consumable World` or `Consumable Handle`.
infixr 5 `seq`
export infixr 5 `seq`
||| We can sequentially compose a computation returning a value that is
||| consumable with another computation. This is done by first consuming
||| the result of the first computation and then returning the second one.

View File

@ -2,7 +2,7 @@ module Data.Linear.Notation
%default total
infixr 0 -@
export infixr 0 -@
||| Infix notation for linear implication
public export
(-@) : Type -> Type -> Type
@ -18,7 +18,7 @@ public export
(.) : (b -@ c) -@ (a -@ b) -@ (a -@ c)
(.) f g v = f (g v)
prefix 5 !*
export prefix 5 !*
||| Prefix notation for the linear unrestricted modality
public export
record (!*) (a : Type) where

View File

@ -88,7 +88,7 @@ Memo (Const s prop) x b = (v : s) -> b v
Memo (d1 * d2) x b = Memo d1 x $ \ v1 => Memo d2 x $ \ v2 => b (v1, v2)
Memo (d1 + d2) x b = (Memo d1 x (b . Left), Memo d2 x (b . Right))
infixr 0 ~>
export infixr 0 ~>
||| A memo trie is a coinductive structure
export

View File

@ -66,12 +66,12 @@ namespace Finitary
Arr (AUnit nm) r = One nm r
Arr (d || e) r = (Arr d r, Arr e r)
infixr 0 ~>
export infixr 0 ~>
record (~>) (d : Fin) (r : Type) where
constructor MkArr
runArr : Arr d r
infix 5 .=
export infix 5 .=
(.=) : (nm : String) -> s -> One nm s
nm .= v = MkOne v
@ -98,7 +98,7 @@ namespace Finitary
appArr (d || e) f (Left x) = appArr d (fst f) x
appArr (d || e) f (Right x) = appArr e (snd f) x
infixl 0 $$
export infixl 0 $$
($$) : {d : Fin} -> (d ~> r) -> (Elem d -> r)
MkArr f $$ x = appArr d f x

View File

@ -76,7 +76,7 @@ data Infer : Scoped where
||| The application of a function to its argument
App : Infer f g -> Check f g -> Infer f g
infixl 3 `App`
export infixl 3 `App`
%name Infer e
@ -229,7 +229,7 @@ namespace Value
vfree : Level nm f -> Value f
vfree x = VEmb (NVar x)
infixl 5 `vapp`
export infixl 5 `vapp`
||| We can easily apply a value standing for a function
||| to a value standing for its argument by either deploying

View File

@ -9,7 +9,7 @@ import Data.List.Quantifiers
%default total
infixr 0 -#
export infixr 0 -#
public export
(-#) : Type -> Type -> Type
a -# b = (0 _ : a) -> b

View File

@ -15,6 +15,8 @@ import Data.List
%default covering
private infixl 3 `App`
||| We use this wrapper to mark places where binding occurs.
||| This is a major footgun and we hope the type constructor
||| forces users to think carefully about what they are doing
@ -55,7 +57,7 @@ namespace Section2
Quote m == Quote n = m == n
_ == _ = False
infixr 0 ~>
private infixr 0 ~>
total
data Ty : Type where
||| A family of base types
@ -91,8 +93,6 @@ namespace Section2
||| The application of a function to its argument
App : Infer -> Check -> Infer
infixl 3 `App`
%name Infer e, f
total
@ -359,7 +359,6 @@ namespace Section3
||| The application of a function to its argument
App : Infer -> Check -> Infer
infixl 3 `App`
%name Infer e, f
@ -600,8 +599,6 @@ namespace Section4
||| The application of a function to its argument
App : Infer -> Check -> Infer
infixl 3 `App`
%name Infer e, f
total
@ -695,7 +692,7 @@ namespace Section4
vfree : Name -> Value
vfree x = VEmb (NVar x)
infixl 5 `vapp`
private infixl 5 `vapp`
||| We can easily apply a value standing for a function
||| to a value standing for its argument by either deploying

View File

@ -114,7 +114,7 @@ not : (AnHDec l, Negates na a) => l na -> HDec (Not a)
not p = [| toNegation (toHDec p) |]
infixr 3 ==>
export infixr 3 ==>
||| Half deciders are closed under implication
public export

View File

@ -40,7 +40,7 @@ HilbertEpsilon p = (v : x ** (v0 : x) -> p v0 -> p v)
0 IsSearchable : Type -> Type
IsSearchable x = (0 p : Pred x) -> Decidable p -> HilbertEpsilon p
infix 0 <->
private infix 0 <->
record (<->) (a, b : Type) where
constructor MkIso
forwards : a -> b

View File

@ -74,7 +74,7 @@ swap (x, y) = (y, x)
-- This directive tells auto implicit search what to use to look inside pairs
%pair Pair fst snd
infixr 5 #
export infixr 5 #
||| A pair type where each component is linear
public export
@ -126,7 +126,7 @@ data Equal : forall a, b . a -> b -> Type where
%name Equal prf
infix 6 ===, ~=~
export infix 6 ===, ~=~
-- An equality type for when you want to assert that each side of the
-- equality has the same type, but there's not other evidence available

View File

@ -69,7 +69,7 @@ public export %tcinline
on : (b -> b -> c) -> (a -> b) -> a -> a -> c
on f g = \x, y => g x `f` g y
infixl 0 `on`
export infixl 0 `on`
||| Takes in the first two arguments in reverse order.
||| @ f the function to flip

View File

@ -1,30 +1,30 @@
module Prelude.Ops
-- Numerical operators
infix 6 ==, /=, <, <=, >, >=
infixl 8 +, -
infixl 9 *, /
export infix 6 ==, /=, <, <=, >, >=
export infixl 8 +, -
export infixl 9 *, /
-- Boolean operators
infixr 5 &&
infixr 4 ||
export infixr 5 &&
export infixr 4 ||
-- List and String operators
infixr 7 ::, ++
infixl 7 :<
export infixr 7 ::, ++
export infixl 7 :<
-- Equivalence
infix 0 <=>
export infix 0 <=>
-- Functor/Applicative/Monad/Algebra operators
infixl 1 >>=, =<<, >>, >=>, <=<, <&>
infixr 2 <|>
infixl 3 <*>, *>, <*
infixr 4 <$>, $>, <$
infixl 8 <+>
export infixl 1 >>=, =<<, >>, >=>, <=<, <&>
export infixr 2 <|>
export infixl 3 <*>, *>, <*
export infixr 4 <$>, $>, <$
export infixl 8 <+>
-- Utility operators
infixr 9 ., .:
infixr 0 $
export infixr 9 ., .:
export infixr 0 $
infixl 9 `div`, `mod`
export infixl 9 `div`, `mod`

View File

@ -406,8 +406,8 @@ Ord a => Ord (List a) where
namespace SnocList
infixl 7 <><
infixr 6 <>>
export infixl 7 <><
export infixr 6 <>>
||| 'fish': Action of lists on snoc-lists
public export

View File

@ -2,8 +2,8 @@ module Algebra.Semiring
%default total
infixl 8 |+|
infixl 9 |*|
export infixl 8 |+|
export infixl 9 |*|
||| A Semiring has two binary operations and an identity for each
public export

View File

@ -29,7 +29,7 @@ import public Libraries.Utils.Binary
||| version number if you're changing the version more than once in the same day.
export
ttcVersion : Int
ttcVersion = 2024_01_23_00
ttcVersion = 2024_03_29_00
export
checkTTCVersion : String -> Int -> Int -> Core ()

View File

@ -821,9 +821,12 @@ HasNames Error where
full gam (InRHS fc n err) = InRHS fc <$> full gam n <*> full gam err
full gam (MaybeMisspelling err xs) = MaybeMisspelling <$> full gam err <*> pure xs
full gam (WarningAsError wrn) = WarningAsError <$> full gam wrn
full gam (OperatorBindingMismatch {print} fc expected actual opName rhs candidates)
full gam (OperatorBindingMismatch {print} fc expected actual (Left opName) rhs candidates)
= OperatorBindingMismatch {print} fc expected actual
<$> full gam opName <*> pure rhs <*> pure candidates
<$> (Left <$> full gam opName) <*> pure rhs <*> pure candidates
full gam (OperatorBindingMismatch {print} fc expected actual (Right opName) rhs candidates)
= OperatorBindingMismatch {print} fc expected actual
<$> (Right <$> full gam opName) <*> pure rhs <*> pure candidates
resolved gam (Fatal err) = Fatal <$> resolved gam err
resolved _ (CantConvert fc gam rho s t)
@ -916,9 +919,12 @@ HasNames Error where
resolved gam (InRHS fc n err) = InRHS fc <$> resolved gam n <*> resolved gam err
resolved gam (MaybeMisspelling err xs) = MaybeMisspelling <$> resolved gam err <*> pure xs
resolved gam (WarningAsError wrn) = WarningAsError <$> resolved gam wrn
resolved gam (OperatorBindingMismatch {print} fc expected actual opName rhs candidates)
resolved gam (OperatorBindingMismatch {print} fc expected actual (Left opName) rhs candidates)
= OperatorBindingMismatch {print} fc expected actual
<$> resolved gam opName <*> pure rhs <*> pure candidates
<$> (Left <$> resolved gam opName) <*> pure rhs <*> pure candidates
resolved gam (OperatorBindingMismatch {print} fc expected actual (Right opName) rhs candidates)
= OperatorBindingMismatch {print} fc expected actual
<$> (Right <$> resolved gam opName) <*> pure rhs <*> pure candidates
export
HasNames Totality where

View File

@ -165,8 +165,9 @@ data Error : Type where
GenericMsg : FC -> String -> Error
GenericMsgSol : FC -> (message : String) -> (solutions : List String) -> Error
OperatorBindingMismatch : {a : Type} -> {print : a -> Doc ()} ->
FC -> (expectedFixity : BacktickOrOperatorFixity) -> (use_site : OperatorLHSInfo a) ->
(opName : Name) -> (rhs : a) -> (candidates : List String) -> Error
FC -> (expectedFixity : FixityDeclarationInfo) -> (use_site : OperatorLHSInfo a) ->
-- left: backticked, right: op symbolds
(opName : Either Name Name) -> (rhs : a) -> (candidates : List String) -> Error
TTCError : TTCErrorMsg -> Error
FileErr : String -> FileError -> Error
CantFindPackage : String -> Error
@ -403,7 +404,7 @@ Show Error where
show (OperatorBindingMismatch fc (DeclaredFixity expected) actual opName rhs _)
= show fc ++ ": Operator " ++ show opName ++ " is " ++ show expected
++ " but used as a " ++ show actual ++ " operator"
show (OperatorBindingMismatch fc Backticked actual opName rhs _)
show (OperatorBindingMismatch fc UndeclaredFixity actual opName rhs _)
= show fc ++ ": Operator " ++ show opName ++ " has no declared fixity"
++ " but used as a " ++ show actual ++ " operator"

View File

@ -23,7 +23,7 @@ interface Hashable a where
hash = hashWithSalt 5381
hashWithSalt h i = h * 33 + hash i
infixl 5 `hashWithSalt`
export infixl 5 `hashWithSalt`
export
Hashable Int where

View File

@ -170,11 +170,11 @@ Eq FixityInfo where
||| Whenever we read an operator from the parser, we don't know if it's a backticked expression with no fixity
||| declaration, or if it has a fixity declaration. If it does not have a declaration, we represent this state
||| with `Backticked`.
||| with `UndeclaredFixity`.
||| Note that a backticked expression can have a fixity declaration, in which case it is represented with
||| `DeclaredFixity`.
public export
data BacktickOrOperatorFixity = Backticked | DeclaredFixity FixityInfo
data FixityDeclarationInfo = UndeclaredFixity | DeclaredFixity FixityInfo
-- Left-hand-side information for operators, carries autobind information
-- an operator can either be

View File

@ -35,6 +35,7 @@ import TTImp.TTImp
import TTImp.Utils
import Libraries.Data.IMaybe
import Libraries.Data.WithDefault
import Libraries.Utils.Shunting
import Libraries.Text.PrettyPrint.Prettyprinter
@ -106,12 +107,12 @@ mkPrec Infix = NonAssoc
mkPrec Prefix = Prefix
-- This is used to print the error message for fixities
[interpName] Interpolation ((Name, BacktickOrOperatorFixity), b) where
interpolate ((x, _), _) = show x
[interpName] Interpolation ((OpStr' Name, FixityDeclarationInfo), b) where
interpolate ((x, _), _) = show x.toName
[showWithLoc] Show ((Name, BacktickOrOperatorFixity), b) where
[showWithLoc] Show ((OpStr' Name, FixityDeclarationInfo), b) where
show ((x, DeclaredFixity y), _) = show x ++ " at " ++ show y.fc
show ((x, Backticked), _) = show x
show ((x, UndeclaredFixity), _) = show x
-- Check that an operator does not have any conflicting fixities in scope.
-- Each operator can have its fixity defined multiple times across multiple
@ -121,19 +122,19 @@ mkPrec Prefix = Prefix
checkConflictingFixities : {auto s : Ref Syn SyntaxInfo} ->
{auto c : Ref Ctxt Defs} ->
(isPrefix : Bool) ->
FC -> Name -> Core (OpPrec, BacktickOrOperatorFixity)
FC -> OpStr' Name -> Core (OpPrec, FixityDeclarationInfo)
checkConflictingFixities isPrefix exprFC opn
= do let op = nameRoot opn
= do let op = nameRoot opn.toName
foundFixities <- getFixityInfo op
let (pre, inf) = partition ((== Prefix) . fix . snd) foundFixities
case (isPrefix, pre, inf) of
-- If we do not find any fixity for this operator we check that it uses operator
-- characters, if not, it must be a backticked expression.
(_, [], []) => if any isOpChar (fastUnpack op)
then throw (GenericMsg exprFC "Unknown operator '\{op}'")
else pure (NonAssoc 1, Backticked) -- Backticks are non associative by default
-- If we do not find any fixity, and it is a backticked operator, then we
-- return the default fixity and associativity for backticked operators
-- Otherwise, it's an unknown operator.
(_, [], []) => case opn of
OpSymbols _ => throw (GenericMsg exprFC "Unknown operator '\{op}'")
Backticked _ => pure (NonAssoc 1, UndeclaredFixity) -- Backticks are non associative by default
--
(True, ((fxName, fx) :: _), _) => do
-- in the prefix case, remove conflicts with infix (-)
let extraFixities = pre ++ (filter (\(nm, _) => not $ nameRoot nm == "-") inf)
@ -171,17 +172,18 @@ checkConflictingFixities isPrefix exprFC opn
checkConflictingBinding : Ref Ctxt Defs =>
Ref Syn SyntaxInfo =>
FC -> Name -> (foundFixity : BacktickOrOperatorFixity) ->
FC -> OpStr -> (foundFixity : FixityDeclarationInfo) ->
(usage : OperatorLHSInfo PTerm) -> (rhs : PTerm) -> Core ()
checkConflictingBinding fc opName foundFixity use_site rhs
= if isCompatible foundFixity use_site
then pure ()
else throw $ OperatorBindingMismatch
{print = byShow} fc foundFixity use_site opName rhs !candidates
{print = byShow} fc foundFixity use_site (opNameToEither opName) rhs !candidates
where
isCompatible : BacktickOrOperatorFixity -> OperatorLHSInfo PTerm -> Bool
isCompatible Backticked (NoBinder lhs) = True
isCompatible Backticked _ = False
isCompatible : FixityDeclarationInfo -> OperatorLHSInfo PTerm -> Bool
isCompatible UndeclaredFixity (NoBinder lhs) = True
isCompatible UndeclaredFixity _ = False
isCompatible (DeclaredFixity fixInfo) (NoBinder lhs) = fixInfo.bindingInfo == NotBinding
isCompatible (DeclaredFixity fixInfo) (BindType name ty) = fixInfo.bindingInfo == Typebind
isCompatible (DeclaredFixity fixInfo) (BindExpr name expr) = fixInfo.bindingInfo == Autobind
@ -198,10 +200,10 @@ checkConflictingBinding fc opName foundFixity use_site rhs
candidates = do let DeclaredFixity fxInfo = foundFixity
| _ => pure [] -- if there is no declared fixity we can't know what's
-- supposed to go there.
Just (nm, cs) <- getSimilarNames {keepPredicate = Just (keepCompatibleBinding fxInfo.bindingInfo)} opName
Just (nm, cs) <- getSimilarNames {keepPredicate = Just (keepCompatibleBinding fxInfo.bindingInfo)} opName.toName
| Nothing => pure []
ns <- currentNS <$> get Ctxt
pure (showSimilarNames ns opName nm cs)
pure (showSimilarNames ns opName.toName nm cs)
checkValidFixity : BindingModifier -> Fixity -> Nat -> Bool
@ -225,7 +227,7 @@ checkValidFixity _ _ _ = False
parameters (side : Side)
toTokList : {auto s : Ref Syn SyntaxInfo} ->
{auto c : Ref Ctxt Defs} ->
PTerm -> Core (List (Tok ((OpStr, BacktickOrOperatorFixity), Maybe (OperatorLHSInfo PTerm)) PTerm))
PTerm -> Core (List (Tok ((OpStr, FixityDeclarationInfo), Maybe (OperatorLHSInfo PTerm)) PTerm))
toTokList (POp fc opFC l opn r)
= do (precInfo, fixInfo) <- checkConflictingFixities False fc opn
unless (side == LHS) -- do not check for conflicting fixity on the LHS
@ -393,7 +395,7 @@ mutual
= do syn <- get Syn
-- It might actually be a prefix argument rather than a section
-- so check that first, otherwise desugar as a lambda
case lookupName op (prefixes syn) of
case lookupName op.toName (prefixes syn) of
[] =>
desugarB side ps
(PLam fc top Explicit (PRef fc (MN "arg" 0)) (PImplicit fc)
@ -798,10 +800,10 @@ mutual
{auto o : Ref ROpts REPLOpts} ->
Side -> List Name -> Tree (OpStr, Maybe $ OperatorLHSInfo PTerm) PTerm ->
Core PTerm
desugarTree side ps (Infix loc eqFC (UN $ Basic "=", _) l r) -- special case since '=' is special syntax
desugarTree side ps (Infix loc eqFC (OpSymbols $ UN $ Basic "=", _) l r) -- special case since '=' is special syntax
= pure $ PEq eqFC !(desugarTree side ps l) !(desugarTree side ps r)
desugarTree side ps (Infix loc _ (UN $ Basic "$", _) l r) -- special case since '$' is special syntax
desugarTree side ps (Infix loc _ (OpSymbols $ UN $ Basic "$", _) l r) -- special case since '$' is special syntax
= do l' <- desugarTree side ps l
r' <- desugarTree side ps r
pure (PApp loc l' r')
@ -809,25 +811,25 @@ mutual
desugarTree side ps (Infix loc opFC (op, Just (NoBinder lhs)) l r)
= do l' <- desugarTree side ps l
r' <- desugarTree side ps r
pure (PApp loc (PApp loc (PRef opFC op) l') r')
pure (PApp loc (PApp loc (PRef opFC op.toName) l') r')
-- (x : ty) =@ f x ==>> (=@) ty (\x : ty => f x)
desugarTree side ps (Infix loc opFC (op, Just (BindType pat lhs)) l r)
= do l' <- desugarTree side ps l
body <- desugarTree side ps r
pure $ PApp loc (PApp loc (PRef opFC op) l')
pure $ PApp loc (PApp loc (PRef opFC op.toName) l')
(PLam loc top Explicit pat l' body)
-- (x := exp) =@ f x ==>> (=@) exp (\x : ? => f x)
desugarTree side ps (Infix loc opFC (op, Just (BindExpr pat lhs)) l r)
= do l' <- desugarTree side ps l
body <- desugarTree side ps r
pure $ PApp loc (PApp loc (PRef opFC op) l')
pure $ PApp loc (PApp loc (PRef opFC op.toName) l')
(PLam loc top Explicit pat (PInfer opFC) body)
-- (x : ty := exp) =@ f x ==>> (=@) exp (\x : ty => f x)
desugarTree side ps (Infix loc opFC (op, Just (BindExplicitType pat ty expr)) l r)
= do l' <- desugarTree side ps l
body <- desugarTree side ps r
pure $ PApp loc (PApp loc (PRef opFC op) l')
pure $ PApp loc (PApp loc (PRef opFC op.toName) l')
(PLam loc top Explicit pat ty body)
desugarTree side ps (Infix loc opFC (op, Nothing) _ r)
= throw $ InternalError "illegal fixity: Parsed as infix but no binding information"
@ -838,7 +840,7 @@ mutual
-- Note: In case of negated signed integer literals, we apply the
-- negation directly. Otherwise, the literal might be
-- truncated to 0 before being passed on to `negate`.
desugarTree side ps (Pre loc opFC (UN $ Basic "-", _) $ Leaf $ PPrimVal fc c)
desugarTree side ps (Pre loc opFC (OpSymbols $ UN $ Basic "-", _) $ Leaf $ PPrimVal fc c)
= let newFC = fromMaybe EmptyFC (mergeFC loc fc)
continue = desugarTree side ps . Leaf . PPrimVal newFC
in case c of
@ -854,13 +856,13 @@ mutual
_ => do arg' <- desugarTree side ps (Leaf $ PPrimVal fc c)
pure (PApp loc (PRef opFC (UN $ Basic "negate")) arg')
desugarTree side ps (Pre loc opFC (UN $ Basic "-", _) arg)
desugarTree side ps (Pre loc opFC (OpSymbols $ UN $ Basic "-", _) arg)
= do arg' <- desugarTree side ps arg
pure (PApp loc (PRef opFC (UN $ Basic "negate")) arg')
desugarTree side ps (Pre loc opFC (op, _) arg)
= do arg' <- desugarTree side ps arg
pure (PApp loc (PRef opFC op) arg')
pure (PApp loc (PRef opFC op.toName) arg')
desugarTree side ps (Leaf t) = pure t
desugarType : {auto s : Ref Syn SyntaxInfo} ->
@ -1016,6 +1018,11 @@ mutual
List Name -> PiInfo PTerm -> Core (PiInfo RawImp)
mapDesugarPiInfo ps = PiInfo.traverse (desugar AnyExpr ps)
displayFixity : Maybe Visibility -> BindingModifier -> Fixity -> Nat -> OpStr -> String
displayFixity Nothing NotBinding fix prec op = "\{show fix} \{show prec} \{show op}"
displayFixity Nothing bind fix prec op = "\{show bind} \{show fix} \{show prec} \{show op}"
displayFixity (Just vis) bind fix prec op = "\{show vis} \{show bind} \{show fix} \{show prec} \{show op}"
-- Given a high level declaration, return a list of TTImp declarations
-- which process it, and update any necessary state on the way.
export
@ -1205,22 +1212,35 @@ mutual
NS ns (DN str (MN ("__mk" ++ str) 0))
mkConName n = DN (show n) (MN ("__mk" ++ show n) 0)
desugarDecl ps (PFixity fc vis binding fix prec opName)
desugarDecl ps fx@(PFixity fc vis binding fix prec opName)
= do unless (checkValidFixity binding fix prec)
(throw $ GenericMsgSol fc
"Invalid fixity, \{binding} operator must be infixr 0."
[ "Make it `infixr 0`: `\{binding} infixr 0 \{show opName}`"
, "Remove the binding keyword: `\{fix} \{show prec} \{show opName}`"
])
when (isDefaulted vis) $
let adjustedExport = displayFixity (Just Export) binding fix prec opName
adjustedPrivate = displayFixity (Just Private) binding fix prec opName
originalFixity = displayFixity Nothing binding fix prec opName
in recordWarning $ GenericWarn fc """
Fixity declaration '\{originalFixity}' does not have an export modifier, and
will become private by default in a future version.
To expose it outside of its module, write '\{adjustedExport}'. If you
intend to keep it private, write '\{adjustedPrivate}'.
"""
ctx <- get Ctxt
-- We update the context of fixities by adding a namespaced fixity
-- given by the current namespace and its fixity name.
-- This allows fixities to be stored along with the namespace at their
-- declaration site and detect and handle ambiguous fixities
let updatedNS = NS (mkNestedNamespace (Just ctx.currentNS) (show fix))
(UN $ Basic $ nameRoot opName)
(UN $ Basic $ nameRoot opName.toName)
update Syn { fixities $= addName updatedNS (MkFixityInfo fc vis binding fix prec) }
update Syn
{ fixities $=
addName updatedNS
(MkFixityInfo fc (collapseDefault vis) binding fix prec) }
pure []
desugarDecl ps d@(PFail fc mmsg ds)
= do -- save the state: the content of a failing block should be discarded

View File

@ -12,7 +12,8 @@ import Libraries.Data.List.Quantifiers.Extra
import Libraries.Data.String.Extra
infix 10 ::=
private infix 10 ::=
data DocFor : String -> Type where
(::=) : (0 str : String) -> (doc : Doc IdrisDocAnn) -> DocFor str

View File

@ -14,6 +14,7 @@ import Idris.Pretty
import Parser.Source
import Data.List
import Data.Either
import Data.List1
import Data.String
@ -626,14 +627,14 @@ perrorRaw (GenericMsgSol fc header solutions)
<+> "Possible solutions:" <+> line
<+> indent 1 (vsep (map (\s => "-" <++> pretty0 s) solutions))
perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candidates)
= pure $ "Operator" <++> pretty0 !(getFullName opName) <++> "is"
<++> printBindingInfo expected
= pure $ "Operator" <++> pretty0 !(getFullName (fromEither opName)) <++> "is"
<++> printBindingInfo expected-- .bindingInfo
<++> "operator, but is used as" <++> printBindingModifier actual.getBinder
<++> "operator."
<+> line <+> !(ploc fc)
<+> "Explanation: regular, typebind and autobind operators all use a slightly different"
<++> "syntax, typebind looks like this: '(name : type)" <++> pretty0 opName
<++> "expr', autobind looks like this: '(name := expr)" <++> pretty0 opName
<++> "syntax, typebind looks like this: '(name : type)" <++> infixOpName
<++> "expr', autobind looks like this: '(name := expr)" <++> infixOpName
<++> "expr'."
<+> line <+> line
<+> "Possible solutions:" <+> line
@ -650,27 +651,28 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candi
moduleDiagnostic : Doc IdrisAnn
moduleDiagnostic = case expected of
Backticked => "Import a module that exports a suitable fixity."
(DeclaredFixity a) => "Hide or remove the fixity at" <++> byShow a.fc
UndeclaredFixity => "Import a module that exports a suitable fixity."
(DeclaredFixity e) => "Hide or remove the fixity at" <++> byShow e.fc
<++> "and import a module that exports a compatible fixity."
infixOpName : Doc IdrisAnn
infixOpName = case expected of
Backticked => enclose "`" "`" (byShow opName)
_ => byShow opName
infixOpName = case opName of
Left backtickedOp => enclose "`" "`" (byShow backtickedOp)
Right symbolOp => byShow symbolOp
displayFixityInfo : FixityInfo -> BindingModifier -> Doc IdrisAnn
displayFixityInfo (MkFixityInfo fc1 vis _ fix precedence) NotBinding
= byShow vis <++> byShow fix <++> byShow precedence <++> pretty0 opName
= byShow vis <++> byShow fix <++> byShow precedence <++> pretty0 (fromEither opName)
displayFixityInfo (MkFixityInfo _ vis _ fix precedence) usedBinder
= byShow vis <++> byShow usedBinder <++> byShow fix <++> byShow precedence <++> pretty0 opName
= byShow vis <++> byShow usedBinder <++> byShow fix <++> byShow precedence <++> pretty0 (fromEither opName)
printE : ? -> Doc IdrisAnn
printE x = reAnnotate (const Code) (p x)
expressionDiagnositc : List (Doc IdrisAnn)
expressionDiagnositc = case expected of
Backticked => []
(DeclaredFixity e) => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:"
UndeclaredFixity => []
DeclaredFixity e => let sentence = "Write the expression using" <++> byShow e.bindingInfo <++> "syntax:"
in pure $ sentence <++> enclose "'" "'" (case e.bindingInfo of
NotBinding =>
printE actual.getLhs <++> infixOpName <++> printE rhs
@ -687,7 +689,7 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candi
fixityDiagnostic : Doc IdrisAnn
fixityDiagnostic = case expected of
Backticked => "Define a new fixity:" <++> "infixr 0" <++> infixOpName
UndeclaredFixity => "Define a new fixity:" <++> "infixr 0" <++> infixOpName
(DeclaredFixity fix) =>
"Change the fixity defined at" <++> pretty0 fix.fc <++> "to"
<++> enclose "'" "'" (displayFixityInfo fix actual.getBinder)
@ -698,8 +700,8 @@ perrorRaw (OperatorBindingMismatch fc {print=p} expected actual opName rhs candi
printBindingModifier Typebind = "a type-binding (typebind)"
printBindingModifier Autobind = "an automatically-binding (autobind)"
printBindingInfo : BacktickOrOperatorFixity -> Doc IdrisAnn
printBindingInfo Backticked = "a regular"
printBindingInfo : FixityDeclarationInfo -> Doc IdrisAnn
printBindingInfo UndeclaredFixity = "a regular"
printBindingInfo (DeclaredFixity x) = printBindingModifier x.bindingInfo

View File

@ -176,9 +176,10 @@ continueWith : IndentInfo -> String -> Rule ()
continueWith indents req
= mustContinue indents (Just req) *> symbol req
iOperator : Rule Name
iOperator : Rule OpStr
iOperator
= operator <|> (symbol "`" *> name <* symbol "`")
= OpSymbols <$> operator
<|> Backticked <$> (symbol "`" *> name <* symbol "`")
data ArgType
= UnnamedExpArg PTerm
@ -369,14 +370,14 @@ mutual
pure $
let fc = boundToFC fname (mergeBounds l r)
opFC = virtualiseFC fc -- already been highlighted: we don't care
in POp fc opFC (NoBinder l.val) (UN $ Basic "=") r.val
in POp fc opFC (NoBinder l.val) (OpSymbols $ UN $ Basic "=") r.val
else fail "= not allowed")
<|>
(do b <- bounds $ do
continue indents
op <- bounds iOperator
e <- case op.val of
UN (Basic "$") => typeExpr q fname indents
OpSymbols (UN (Basic "$")) => typeExpr q fname indents
_ => expr q fname indents
pure (op, e)
(op, r) <- pure b.val
@ -1175,10 +1176,10 @@ visibility fname
= (specified <$> visOption fname)
<|> pure defaulted
exportVisibility : OriginDesc -> EmptyRule Visibility
exportVisibility : OriginDesc -> EmptyRule (WithDefault Visibility Export)
exportVisibility fname
= visOption fname
<|> pure Export
= (specified <$> visOption fname)
<|> pure defaulted
tyDecls : Rule Name -> String -> OriginDesc -> IndentInfo -> Rule (List1 PTypeDecl)
tyDecls declName predoc fname indents

View File

@ -334,24 +334,24 @@ mutual
prettyPrec d (PInfer _) = annotate Hole $ "?"
prettyPrec d (POp _ _ (BindType nm left) op right) =
group $ parens (prettyPrec d nm <++> ":" <++> pretty left)
<++> prettyOp op
<++> prettyOp op.toName
<++> pretty right
prettyPrec d (POp _ _ (BindExpr nm left) op right) =
group $ parens (prettyPrec d nm <++> ":=" <++> pretty left)
<++> prettyOp op
<++> prettyOp op.toName
<++> pretty right
prettyPrec d (POp _ _ (BindExplicitType nm ty left) op right) =
group $ parens (prettyPrec d nm <++> ":" <++> pretty ty <++> ":=" <++> pretty left)
<++> prettyOp op
<++> prettyOp op.toName
<++> pretty right
prettyPrec d (POp _ _ (NoBinder x) op y) =
parenthesise (d >= App) $
group $ pretty x
<++> prettyOp op
<++> prettyOp op.toName
<++> pretty y
prettyPrec d (PPrefixOp _ _ op x) = parenthesise (d > startPrec) $ prettyOp op <+> pretty x
prettyPrec d (PSectionL _ _ op x) = parens (prettyOp op <++> pretty x)
prettyPrec d (PSectionR _ _ x op) = parens (pretty x <++> prettyOp op)
prettyPrec d (PPrefixOp _ _ op x) = parenthesise (d > startPrec) $ prettyOp op.toName <+> pretty x
prettyPrec d (PSectionL _ _ op x) = parens (prettyOp op.toName <++> pretty x)
prettyPrec d (PSectionR _ _ x op) = parens (pretty x <++> prettyOp op.toName)
prettyPrec d (PEq fc l r) = parenthesise (d > startPrec) $ prettyPrec Equal l <++> equals <++> prettyPrec Equal r
prettyPrec d (PBracketed _ tm) = parens (pretty tm)
prettyPrec d (PString _ _ xs) = parenthesise (d > startPrec) $ hsep $ punctuate "++" (prettyPStr <$> xs)

View File

@ -46,7 +46,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)) kn (unbracketApp y)
let asOp = POp fc opFC (NoBinder (unbracketApp x)) (OpSymbols kn) (unbracketApp y)
if not (null (lookupName rootName (infixes syn)))
then pure asOp
else case dropNS raw of
@ -55,7 +55,7 @@ mkOp tm@(PApp fc (PApp _ (PRef opFC kn) x) y)
mkOp tm@(PApp fc (PRef opFC kn) x)
= do syn <- get Syn
let n = rawName kn
let asOp = PSectionR fc opFC (unbracketApp x) kn
let asOp = PSectionR fc opFC (unbracketApp x) (OpSymbols kn)
if not (null $ lookupName (UN $ Basic (nameRoot n)) (infixes syn))
then pure asOp
else case dropNS n of
@ -73,7 +73,7 @@ mkSectionL tm@(PLam fc rig info (PRef _ bd) ty
| _ => pure tm
syn <- get Syn
let n = rawName kn
let asOp = PSectionL fc opFC kn (unbracketApp x)
let asOp = PSectionL fc opFC (OpSymbols kn) (unbracketApp x)
if not (null $ lookupName (UN $ Basic (nameRoot n)) (fixities syn))
then pure asOp
else case dropNS n of
@ -591,15 +591,15 @@ cleanPTerm ptm
cleanNode (PRef fc nm) =
PRef fc <$> cleanKindedName nm
cleanNode (POp fc opFC abi op y) =
(\ op => POp fc opFC abi op y) <$> cleanKindedName op
(\ op => POp fc opFC abi op y) <$> traverseOp @{Functor.CORE} cleanKindedName op
cleanNode (PPrefixOp fc opFC op x) =
(\ op => PPrefixOp fc opFC op x) <$> cleanKindedName op
(\ op => PPrefixOp fc opFC op x) <$> traverseOp @{Functor.CORE} cleanKindedName op
cleanNode (PSectionL fc opFC op x) =
(\ op => PSectionL fc opFC op x) <$> cleanKindedName op
(\ op => PSectionL fc opFC op x) <$> traverseOp @{Functor.CORE} cleanKindedName op
cleanNode (PSectionR fc opFC x op) =
PSectionR fc opFC x <$> cleanKindedName op
PSectionR fc opFC x <$> traverseOp @{Functor.CORE} cleanKindedName op
cleanNode (PPi fc rig vis (Just n) arg ret) =
(\ n => PPi fc rig vis n arg ret) <$> cleanBinderName vis n
(\ n => PPi fc rig vis n arg ret) <$> (cleanBinderName vis n)
cleanNode tm = pure tm
toCleanPTerm : {auto c : Ref Ctxt Defs} ->

View File

@ -30,8 +30,39 @@ import Parser.Lexer.Source
%default covering
public export
OpStr' : Type -> Type
OpStr' nm = nm
data OpStr' nm = OpSymbols nm | Backticked nm
-- Left: backticked, Right: operator symbols
export
opNameToEither : OpStr' nm -> Either nm nm
opNameToEither (Backticked nm) = Left nm
opNameToEither (OpSymbols nm) = Right nm
export
Functor OpStr' where
map f (OpSymbols x) = OpSymbols (f x)
map f (Backticked x) = Backticked (f x)
export
Foldable OpStr' where
foldr f i (OpSymbols x) = f x i
foldr f i (Backticked x) = f x i
export
traverseOp : (fn : Functor f) => (a -> f b) -> OpStr' a -> f (OpStr' b)
traverseOp f (OpSymbols x) = map OpSymbols (f x)
traverseOp f (Backticked x) = map Backticked (f x)
public export
Show nm => Show (OpStr' nm) where
show (OpSymbols nm) = show nm
show (Backticked nm) = "`\{show nm}`"
public export
(.toName) : OpStr' nm -> nm
(.toName) (OpSymbols nm) = nm
(.toName) (Backticked nm) = nm
public export
OpStr : Type
@ -434,7 +465,7 @@ mutual
PFail : FC -> Maybe String -> List (PDecl' nm) -> PDecl' nm
PMutual : FC -> List (PDecl' nm) -> PDecl' nm
PFixity : FC -> Visibility -> BindingModifier -> Fixity -> Nat -> OpStr -> PDecl' nm
PFixity : FC -> WithDefault Visibility Export -> BindingModifier -> Fixity -> Nat -> OpStr -> PDecl' nm
PNamespace : FC -> Namespace -> List (PDecl' nm) -> PDecl' nm
PTransform : FC -> String -> PTerm' nm -> PTerm' nm -> PDecl' nm
PRunElabDecl : FC -> PTerm' nm -> PDecl' nm
@ -832,10 +863,8 @@ parameters {0 nm : Type} (toName : nm -> Name)
showPTermPrec d (PWithUnambigNames fc ns rhs)
= "with " ++ show ns ++ " " ++ showPTermPrec d rhs
showOpPrec d op = let op = toName op in
if isOpName op
then showPrec d op
else "`" ++ showPrec d op ++ "`"
showOpPrec d (OpSymbols op) = showPrec d (toName op)
showOpPrec d (Backticked op) = "`\{showPrec d (toName op)}`"
export
covering
@ -1045,3 +1074,4 @@ Show PDecl where
show (PRunElabDecl{}) = "PRunElabDecl"
show (PDirective{}) = "PDirective"
show (PBuiltin{}) = "PBuiltin"

View File

@ -27,7 +27,7 @@ getFnArgs embed fts = go fts [] where
go (PNamedApp fc f n t) = go f . (Named fc n t ::)
go (PBracketed fc f) = go f
-- we don't care about the binder info here
go (POp fc opFC leftSide op r) = (PRef opFC op,) . (Explicit fc leftSide.getLhs ::) . (Explicit fc r ::)
go (POp fc opFC leftSide op r) = (PRef opFC op.toName,) . (Explicit fc leftSide.getLhs ::) . (Explicit fc r ::)
go (PEq fc l r) = (PRef fc $ embed eqName,) . (Explicit fc l ::) . (Explicit fc r ::)
-- ambiguous, picking the type constructor here
go (PPair fc l r) = (PRef fc $ embed pairname,) . (Explicit fc l ::) . (Explicit fc r ::)

View File

@ -2,7 +2,7 @@ module Libraries.Data.Ordering.Extra
%default total
infixl 5 `thenCmp`
export infixl 5 `thenCmp`
export
thenCmp : Ordering -> Lazy Ordering -> Ordering

View File

@ -12,8 +12,8 @@ import Data.List
%default total
infixr 5 <|, <:
infixl 5 |>, :>
export infixr 5 <|, <:
export infixl 5 |>, :>
public export
FileRange : Type

View File

@ -7,8 +7,8 @@ import Data.String
%default total
infixl 5 +>
infixr 5 <+
export infixl 5 +>
export infixr 5 <+
||| Adds a character to the end of the specified string.
|||

View File

@ -132,7 +132,7 @@ export %inline
Grammar state tok (c1 && c2) ty
(<|>) = Alt
infixr 2 <||>
export infixr 2 <||>
||| Take the tagged disjunction of two grammars. If both consume, the
||| combination is guaranteed to consume.
export

View File

@ -134,7 +134,7 @@ export
fill : Int -> Doc ann -> Doc ann
fill n doc = width doc (\w => spaces $ n - w)
infixr 6 <++>
export infixr 6 <++>
||| Concatenates two documents with a space in between.
export
(<++>) : Doc ann -> Doc ann -> Doc ann

View File

@ -19,9 +19,9 @@ import System.File
%default total
infixl 5 </>, />
infixr 7 <.>
infixr 7 <..>
export infixl 5 </>, />
export infixr 7 <.>
export infixr 7 <..>
||| The character that separates directories in the path.

View File

@ -3,7 +3,7 @@ import Data.SortedMap
(&~) : a -> (a -> b) -> b
(&~) x f = f x
infixl 2 &~
private infixl 2 &~
testLookupBetween : List (Maybe (Maybe (Int,Int),Maybe (Int,Int)))
testLookupBetween =

View File

@ -1,6 +1,6 @@
module LetBinders
infix 1 :::
private infix 1 :::
record List1 (a : Type) where
constructor (:::)
head : a

View File

@ -2,8 +2,9 @@
00003e(:write-string "1/1: Building LetBinders (LetBinders.idr)" 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 0 0) (:end 0 6)) ((:decor :keyword)))))) 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 0 7) (:end 0 17)) ((:decor :module)))))) 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 0) (:end 2 5)) ((:decor :keyword)))))) 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 6) (:end 2 7)) ((:decor :keyword)))))) 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 0) (:end 2 7)) ((:decor :keyword)))))) 1)
000076(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 8) (:end 2 13)) ((:decor :keyword)))))) 1)
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 2 14) (:end 2 15)) ((:decor :keyword)))))) 1)
000075(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 0) (:end 3 6)) ((:decor :keyword)))))) 1)
000073(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 7) (:end 3 12)) ((:decor :type)))))) 1)
000077(:output (:ok (:highlight-source ((((:filename "LetBinders.idr") (:start 3 13) (:end 3 14)) ((:decor :keyword)))))) 1)

View File

@ -1,7 +1,7 @@
data Maybe a = Nothing
| Just a
infixl 1 >>=
private infixl 1 >>=
(>>=) : Maybe a -> (a -> Maybe b) -> Maybe b
(>>=) Nothing k = Nothing

View File

@ -1,4 +1,4 @@
infixr 5 ::
private infixr 5 ::
export
data List a = Nil | (::) a (List a)

View File

@ -36,7 +36,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -36,7 +36,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -36,7 +36,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -13,7 +13,7 @@ not False = True
public export
data Maybe a = Nothing | Just a
infixl 4 &&
export infixl 4 &&
public export
(&&) : Bool -> Lazy Bool -> Bool
@ -43,7 +43,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -1,7 +1,7 @@
data Maybe a = Nothing
| Just a
infixl 1 >>=
private infixl 1 >>=
(>>=) : Maybe a -> (a -> Maybe b) -> Maybe b
(>>=) Nothing k = Nothing

View File

@ -1,19 +1,19 @@
--- Data declarations ---
infix 0 =%=
private infix 0 =%=
public export
data (=%=) : (a -> b) -> (a -> b) -> Type where
ExtEq : {0 f, g : a -> b} -> ((x : a) -> f x = g x) -> f =%= g
infix 0 %%
private infix 0 %%
public export
data (%%) a b = Equs (a = b) (b = a)
--- Records ---
infix 0 =%%=
private infix 0 =%%=
public export
record (=%%=) {a : Type} {b : Type} (f : a -> b) (g : a -> b) where
@ -22,7 +22,7 @@ record (=%%=) {a : Type} {b : Type} (f : a -> b) (g : a -> b) where
--- Interfaces ---
infix 0 =%%%=
private infix 0 =%%%=
public export
interface (=%%%=) (x : a) (y : a) (b : Type) (i : Type) where

View File

@ -4,7 +4,7 @@ import Stuff
%default partial
infixl 5 ==, /=
private infixl 5 ==, /=
interface Eq b where
(==) : b -> b -> Bool

View File

@ -4,7 +4,7 @@ import Stuff
%default partial
infixl 5 ==, /=
private infixl 5 ==, /=
interface Eq b where
(==) : b -> b -> Bool

View File

@ -13,7 +13,7 @@ not False = True
public export
data Maybe a = Nothing | Just a
infixl 4 &&
export infixl 4 &&
public export
(&&) : Bool -> Bool -> Bool
@ -43,7 +43,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)
@ -52,7 +52,7 @@ public export
data Equal : a -> b -> Type where
Refl : {0 x : a} -> Equal x x
infix 9 ===, ~=~
export infix 9 ===, ~=~
public export
(===) : (x : a) -> (y : a) -> Type

View File

@ -36,7 +36,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -2,7 +2,7 @@ module Apply
-- These are not Biapplicatives. Those are in Data.Biapplicative
infixl 4 <<$>>, <<&>>, <<.>>, <<., .>>, <<..>>
export infixl 4 <<$>>, <<&>>, <<.>>, <<., .>>, <<..>>
||| Primarily used to make the definitions of bilift2 and bilift3 pretty
|||

View File

@ -2,7 +2,7 @@ module Biapplicative
import Apply
infixl 4 <<*>>, <<*, *>>, <<**>>
export infixl 4 <<*>>, <<*, *>>, <<**>>
||| Biapplicatives
||| @p the action of the Biapplicative on pairs of objects

View File

@ -2,7 +2,7 @@ module Bimonad
import Biapplicative
infixl 4 >>==
export infixl 4 >>==
||| Bimonads
||| @p the action of the first Bifunctor component on pairs of objects

View File

@ -1,8 +1,8 @@
import Syntax.PreorderReasoning
-------- some notation ----------
infixr 6 .+.,:+:
infixr 7 .*.,:*:
private infixr 6 .+.,:+:
private infixr 7 .*.,:*:
interface Additive a where
constructor MkAdditive

View File

@ -43,7 +43,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -43,7 +43,7 @@ plus : Nat -> Nat -> Nat
plus Z y = y
plus (S k) y = S (plus k y)
infixr 5 ::
export infixr 5 ::
public export
data List a = Nil | (::) a (List a)

View File

@ -1,4 +1,4 @@
infixr 6 -*
private infixr 6 -*
(-*) : Type -> Type -> Type
(-*) a b = (1 _ : a) -> b

View File

@ -2,8 +2,8 @@ module List
data List a = Nil | (::) a (List a)
infixr 5 ::
infixr 5 ++
export infixr 5 ::
export infixr 5 ++
interface Monoid ty where
||| Users can hand-craft their own monoid implementations

View File

@ -5,6 +5,7 @@ infixr 5 ~:>
public export
infixr 5 ~>
export
infixl 5 |>
public export

View File

@ -1,4 +1,4 @@
infixr 5 ::
private infixr 5 ::
namespace List
public export

View File

@ -1,6 +1,6 @@
module Lib1
infixr 5 %%%
export infixr 5 %%%
export
(%%%) : Nat -> Nat -> Nat

View File

@ -1,3 +1,3 @@
module Lib2
infixl 5 %%%
export infixl 5 %%%

View File

@ -1,6 +1,6 @@
module LibPre1
prefix 6 *!
export prefix 6 *!
export
(*!) : Nat -> Nat

View File

@ -1,3 +1,3 @@
module LibPre2
prefix 2 *!
export prefix 2 *!

View File

@ -2,8 +2,8 @@ module Main3
import Lib1
prefix 4 %%%
infixr 8 -
private prefix 4 %%%
private infixr 8 -
(%%%) : Nat -> Nat
(%%%) = S

View File

@ -1,6 +1,6 @@
module NonConflict1
infixr 5 &&&
export infixr 5 &&&
export
(&&&) : Nat -> Nat -> Nat

View File

@ -1,3 +1,3 @@
module NonConflict2
infixr 5 &&&
export infixr 5 &&&

View File

@ -1,7 +1,7 @@
module Issue1328A
%default total
infix 4 `InCtx`
private infix 4 `InCtx`
data InCtx : (a, b : Type) -> Type where
prop0 : a `InCtx` b
@ -22,7 +22,7 @@ def1 m n = (`div` n) m
def2 : (m, n : Integer) -> Integer
def2 m n = (m `div`) n
infix 4 |-
private infix 4 |-
data (|-) : (a, b : Type) -> Type where
inProp0 : a |- b

View File

@ -3,8 +3,8 @@ module Test
import Data.Vect
typebind infixr 0 =@
infixr 0 -@
private typebind infixr 0 =@
private infixr 0 -@
-- typebind infixr 1 =@@
@ -39,7 +39,7 @@ map4 : (x : a) =@ (b -@ (y : List a) =@ List b)
test3 : Test.map3 === Test.map4
typebind infixr 0 *>
private typebind infixr 0 *>
-- (*>) : (ty : Type) -> (ty -> Type) -> Type
-- (*>) = DPair
@ -47,7 +47,7 @@ typebind infixr 0 *>
-- testCompose : (x : Nat) *> (y : Nat) *> Vect (n + m) String
-- testCompose = (1 ** 2 ** ["hello", "world", "!"])
autobind infixr 0 `MyLet`
private autobind infixr 0 `MyLet`
MyLet : (val) -> (val -> rest) -> rest
MyLet arg fn = fn arg
@ -58,14 +58,14 @@ program = (n := 3) `MyLet` 2 + n
program2 : Nat
program2 = (n : Nat := 3) `MyLet` 2 + n
typebind infixr 0 |>
private typebind infixr 0 |>
record Container where
constructor (|>)
shape : Type
position : shape -> Type
typebind infixr 0 @@
private typebind infixr 0 @@
record (@@) (x : Type) (y : x -> Type) where
constructor PairUp

View File

@ -1,5 +1,5 @@
typebind infixr 0 =@
private typebind infixr 0 =@
0 (=@) : (a : Type) -> (a -> Type) -> Type
(=@) a f = (1 x : a) -> f x

View File

@ -1,5 +1,5 @@
typebind infixr 0 =@
private typebind infixr 0 =@
0 (=@) : (a : Type) -> (a -> Type) -> Type
(=@) a f = (1 x : a) -> f x

View File

@ -1,5 +1,5 @@
autobind infixr 0 =@
private autobind infixr 0 =@
0 (=@) : (a : Type) -> (a -> Type) -> Type
(=@) a f = (1 x : a) -> f x

Some files were not shown because too many files have changed in this diff Show More