2020-05-20 13:23:04 +03:00
Dot syntax for records
======================
.. role :: idris(code)
:language: idris
Long story short, `` .field `` is a postfix projection operator that binds
tighter than function application.
Lexical structure
-----------------
2020-06-21 00:49:34 +03:00
* `` .foo `` is a valid lexeme, which stands for postfix application of `` foo ``
2020-05-20 13:23:04 +03:00
* `` Foo.bar.baz `` starting with uppercase `` F `` is one lexeme, a namespaced
2020-05-21 06:25:25 +03:00
identifier: `` DotSepIdent ["baz", "bar", "Foo"] ``
2020-05-20 13:23:04 +03:00
* `` foo.bar.baz `` starting with lowercase `` f `` is three lexemes: `` foo `` ,
`` .bar `` , `` .baz ``
* `` .foo.bar.baz `` is three lexemes: `` .foo `` , `` .bar `` , `` .baz ``
* If you want `` Constructor.field `` , you have to write `` (Constructor).field `` .
* All module names must start with an uppercase letter.
New syntax of `` simpleExpr ``
----------------------------
Expressions binding tighter than application (`` simpleExpr `` ), such as variables or parenthesised expressions, have been renamed to `` simplerExpr `` , and an extra layer of syntax has been inserted.
.. code-block :: idris
2020-06-21 00:49:34 +03:00
simpleExpr ::= (.field)+ -- parses as PPostfixAppPartial
| simplerExpr (.field)+ -- parses as PPostfixApp
2020-05-20 13:23:04 +03:00
| simplerExpr -- (parses as whatever it used to)
2020-06-21 00:49:34 +03:00
* `` (.foo.bar) `` is a valid parenthesised expression
2020-05-20 13:23:04 +03:00
Desugaring rules
----------------
* `` (.field1 .field2 .field3) `` desugars to `` (\x => .field3 (.field2 (.field1
x)))``
* `` (simpleExpr .field1 .field2 .field3) `` desugars to `` ((.field .field2
.field3) simpleExpr)``
Example code
------------
.. code-block :: idris
record Point where
constructor MkPoint
x : Double
y : Double
This record creates two projections:
* `` x : Point -> Double ``
* `` y : Point -> Double ``
2020-06-21 00:49:34 +03:00
We add another record:
2020-05-20 13:23:04 +03:00
.. code-block :: idris
record Rect where
constructor MkRect
topLeft : Point
bottomRight : Point
Let's define some constants:
.. code-block :: idris
pt : Point
pt = MkPoint 4.2 6.6
rect : Rect
rect =
MkRect
(MkPoint 1.1 2.5)
(MkPoint 4.3 6.3)
Finally, the examples:
.. code-block :: idris
main : IO ()
main = do
2020-06-21 00:49:34 +03:00
-- desugars to (x pt)
2020-05-20 13:23:04 +03:00
-- prints 4.2
printLn $ pt.x
-- prints 4.2, too
-- maybe we want to make this a parse error?
printLn $ pt .x
-- prints 10.8
printLn $ pt.x + pt.y
-- works fine with namespacing
-- prints 4.2
printLn $ (Main.pt).x
-- the LHS can be an arbitrary expression
-- prints 4.2
printLn $ (MkPoint pt.y pt.x).y
-- user-defined projection
-- prints 17.64
printLn $ pt.x.squared
-- prints [1.0, 3.0]
printLn $ map (.x) [MkPoint 1 2, MkPoint 3 4]
2020-06-21 00:49:34 +03:00
-- .topLeft.y desugars to (\x => y (topLeft x))
2020-05-20 13:23:04 +03:00
-- prints [2.5, 2.5]
printLn $ map (.topLeft.y) [rect, rect]
-- desugars to (.topLeft.x rect + .bottomRight.y rect)
-- prints 7.4
printLn $ rect.topLeft.x + rect.bottomRight.y
-- qualified names work, too
-- all these print 4.2
printLn $ Main.Point.(.x) pt
printLn $ Point.(.x) pt
printLn $ (.x) pt
printLn $ .x pt
-- haskell-style projections work, too
printLn $ Main.Point.x pt
printLn $ Point.x pt
printLn $ (x) pt
printLn $ x pt
-- record update syntax uses dots now
-- prints 3.0
printLn $ (record { topLeft.x = 3 } rect).topLeft.x
-- but for compatibility, we support the old syntax, too
printLn $ (record { topLeft->x = 3 } rect).topLeft.x
-- prints 2.1
printLn $ (record { topLeft.x $= (+1) } rect).topLeft.x
printLn $ (record { topLeft->x $= (+1) } rect).topLeft.x
Parses but does not typecheck:
.. code-block :: idris
-- parses as: map.x [MkPoint 1 2, MkPoint 3 4]
-- maybe we should disallow spaces before dots?
--
printLn $ map .x [MkPoint 1 2, MkPoint 3 4]