move basic053->interpolation002, move basic061->basic053

This commit is contained in:
André Videla 2021-07-22 09:45:22 +00:00
parent 1276f47cf6
commit bab3897a9b
15 changed files with 167 additions and 39 deletions

View File

@ -41,7 +41,7 @@ idrisTestsBasic = MkTestPool "Fundamental language features" [] Nothing
"basic046", "basic047", "basic049", "basic050",
"basic051", "basic052", "basic053", "basic054", "basic055",
"basic056", "basic057", "basic058", "basic059", "basic060",
"basic061", "interpolation001"]
"interpolation001", "interpolation002"]
idrisTestsCoverage : TestPool
idrisTestsCoverage = MkTestPool "Coverage checking" [] Nothing

View File

@ -1,25 +1,8 @@
""""
Just 1 + Just 2 = Just 3
hello world.
Hello Idris2!
[project]
name: "project"
version: "0.1.0"
[deps]
"semver" = 0.2
a b\n
c
sticking together!
A very very
very very very
very long string.
a
name: #"foo"
version: "bar"
bzs: \#\'a\n\t\\'
"""
"contains\NULcharacter"
1/1: Building UnderscoredIntegerLiterals (UnderscoredIntegerLiterals.idr)
UnderscoredIntegerLiterals> 10000000000
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> 3405705229
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> Bye for now!

View File

@ -1,4 +1,4 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner --exec test StringLiteral.idr
$1 --no-color --console-width 0 --no-banner UnderscoredIntegerLiterals.idr < input

View File

@ -1,8 +0,0 @@
1/1: Building UnderscoredIntegerLiterals (UnderscoredIntegerLiterals.idr)
UnderscoredIntegerLiterals> 10000000000
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> 3405705229
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> True
UnderscoredIntegerLiterals> Bye for now!

View File

@ -1,4 +0,0 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner UnderscoredIntegerLiterals.idr < input

View File

@ -0,0 +1,69 @@
module StringLiteral
rawStr : String
rawStr = #""""""#
interp1 : String
interp1 = "Just 1 + Just 2 = \{
show $ do a <- Just 1
b <- Just 2
Just (a + b)
}"
interp2 : String
interp2 = "hello\{ " " ++ ##"world\##{#"."#}"## }"
multi1 : String
multi1 = """
[project]
name: "project"
version: "0.1.0"
[deps]
"semver" = 0.2
"""
multi2 : String
multi2 = #"""
a\#
b\n
\#{"c"}
"""#
multi3 : String
multi3 = """
\{"sticking"} \{"together"}\{
""}\{#"""
!
"""#}
"""
multi4 : String
multi4 = """
A very very \n\nvery very very \n
very long string. \
"""
test : IO ()
test = do
putStrLn rawStr
putStrLn interp1
putStrLn interp2
let idris = "Idris"
putStrLn "Hello \{idris ++ show 2}!"
putStrLn multi1
putStrLn multi2
putStrLn multi3
putStrLn multi4
putStrLn """
"""
putStrLn ##"""
a
"""##
putStrLn ##"""
name: #"foo"
version: "bar"
bzs: \#\'a\n\t\\'
"""
"""##
printLn "contains\NULcharacter"

View File

@ -0,0 +1,23 @@
module UnderscoredIntegerLiterals
-- grouping decimal numbers by thousands
amount : Integer
amount = 10_000_000_000
equalAmounts : Bool
equalAmounts = amount == 10000000000
-- grouping hexadecimal addresses by words
addr : Int
addr = 0xCAFE_F00D
equalAddrs : Bool
equalAddrs = addr == 0xCAFEF00D
-- grouping bits into nibbles in a binary literal
equalFlags : Bool
equalFlags = 0b0011_1111_0100_1110 == 0b0011111101001110
-- grouping octals
equalOctals : Bool
equalOctals = 0o455_777 == 0o455777

View File

@ -0,0 +1,25 @@
""""
Just 1 + Just 2 = Just 3
hello world.
Hello Idris2!
[project]
name: "project"
version: "0.1.0"
[deps]
"semver" = 0.2
a b\n
c
sticking together!
A very very
very very very
very long string.
a
name: #"foo"
version: "bar"
bzs: \#\'a\n\t\\'
"""
"contains\NULcharacter"

View File

@ -0,0 +1,7 @@
amount
equalAmounts
addr
equalAddrs
equalFlags
equalOctals
:q

View File

@ -0,0 +1,4 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner --exec test StringLiteral.idr

View File

@ -0,0 +1,25 @@
""""
Just 1 + Just 2 = Just 3
hello world.
Hello Idris2!
[project]
name: "project"
version: "0.1.0"
[deps]
"semver" = 0.2
a b\n
c
sticking together!
A very very
very very very
very long string.
a
name: #"foo"
version: "bar"
bzs: \#\'a\n\t\\'
"""
"contains\NULcharacter"

View File

@ -0,0 +1,4 @@
rm -rf build
$1 --no-color --console-width 0 --no-banner --exec test StringLiteral.idr