@@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Yakov Pechersky
5
5
-/
6
6
import data.fintype.card
7
+ import data.buffer.parser.basic
7
8
8
9
/-!
9
10
# Numeral parsers
@@ -43,15 +44,15 @@ variables (α : Type) [has_zero α] [has_one α] [has_add α]
43
44
/--
44
45
Parse a string of digits as a numeral while casting it to target type `α`.
45
46
-/
46
- def numeral : parser α :=
47
+ @[derive mono] def numeral : parser α :=
47
48
nat.bin_cast <$> nat
48
49
49
50
/--
50
51
Parse a string of digits as a numeral while casting it to target type `α`,
51
52
which has a `[fintype α]` constraint. The parser ensures that the numeral parsed in
52
53
is within the cardinality of the type `α`.
53
54
-/
54
- def numeral.of_fintype [fintype α] : parser α :=
55
+ @[derive mono] def numeral.of_fintype [fintype α] : parser α :=
55
56
do
56
57
c ← nat,
57
58
decorate_error (sformat!" <numeral less than {to_string (fintype.card α)}>" )
62
63
Parse a string of digits as a numeral while casting it to target type `α`. The parsing starts
63
64
at "1", so `"1"` is parsed in as `nat.cast 0`. Providing `"0"` to the parser causes a failure.
64
65
-/
65
- def numeral.from_one : parser α :=
66
+ @[derive mono] def numeral.from_one : parser α :=
66
67
do
67
68
c ← nat,
68
69
decorate_error (" <positive numeral>" )
@@ -75,7 +76,7 @@ which has a `[fintype α]` constraint. The parser ensures that the numeral parse
75
76
is within the cardinality of the type `α`. The parsing starts
76
77
at "1", so `"1"` is parsed in as `nat.cast 0`. Providing `"0"` to the parser causes a failure.
77
78
-/
78
- def numeral.from_one.of_fintype [fintype α] : parser α :=
79
+ @[derive mono] def numeral.from_one.of_fintype [fintype α] : parser α :=
79
80
do
80
81
c ← nat,
81
82
decorate_error (sformat!" <positive numeral less than or equal to {to_string (fintype.card α)}>" )
@@ -87,7 +88,7 @@ Parse a character as a numeral while casting it to target type `α`,
87
88
The parser ensures that the character parsed in is within the bounds set by `fromc` and `toc`,
88
89
and subtracts the value of `fromc` from the parsed in character.
89
90
-/
90
- def numeral.char (fromc toc : char) : parser α :=
91
+ @[derive mono] def numeral.char (fromc toc : char) : parser α :=
91
92
do
92
93
c ← decorate_error
93
94
(sformat!" <char between '{fromc.to_string}' to '{toc.to_string}' inclusively>" )
@@ -101,7 +102,7 @@ The parser ensures that the character parsed in is greater or equal to `fromc` a
101
102
and subtracts the value of `fromc` from the parsed in character. There is also a check
102
103
that the resulting value is within the cardinality of the type `α`.
103
104
-/
104
- def numeral.char.of_fintype [fintype α] (fromc : char) : parser α :=
105
+ @[derive mono] def numeral.char.of_fintype [fintype α] (fromc : char) : parser α :=
105
106
do
106
107
c ← decorate_error
107
108
(sformat!" <char from '{fromc.to_string}' to '{
0 commit comments