Skip to content

Commit

Permalink
chore(data/buffer/parser/numeral): derive mono for numeral (#5463)
Browse files Browse the repository at this point in the history
Thanks to Rob Lewis, using classes, instances, and the `delta_instance_handler`, transferring over instances becomes very easy.



Co-authored-by: Yakov Pechersky <pechersky@users.noreply.github.com>
  • Loading branch information
pechersky and pechersky committed Feb 19, 2021
1 parent e5c7789 commit 114f8ef
Showing 1 changed file with 7 additions and 6 deletions.
13 changes: 7 additions & 6 deletions src/data/buffer/parser/numeral.lean
Expand Up @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Yakov Pechersky
-/
import data.fintype.card
import data.buffer.parser.basic

/-!
# Numeral parsers
Expand Down Expand Up @@ -43,15 +44,15 @@ variables (α : Type) [has_zero α] [has_one α] [has_add α]
/--
Parse a string of digits as a numeral while casting it to target type `α`.
-/
def numeral : parser α :=
@[derive mono] def numeral : parser α :=
nat.bin_cast <$> nat

/--
Parse a string of digits as a numeral while casting it to target type `α`,
which has a `[fintype α]` constraint. The parser ensures that the numeral parsed in
is within the cardinality of the type `α`.
-/
def numeral.of_fintype [fintype α] : parser α :=
@[derive mono] def numeral.of_fintype [fintype α] : parser α :=
do
c ← nat,
decorate_error (sformat!"<numeral less than {to_string (fintype.card α)}>")
Expand All @@ -62,7 +63,7 @@ do
Parse a string of digits as a numeral while casting it to target type `α`. The parsing starts
at "1", so `"1"` is parsed in as `nat.cast 0`. Providing `"0"` to the parser causes a failure.
-/
def numeral.from_one : parser α :=
@[derive mono] def numeral.from_one : parser α :=
do
c ← nat,
decorate_error ("<positive numeral>")
Expand All @@ -75,7 +76,7 @@ which has a `[fintype α]` constraint. The parser ensures that the numeral parse
is within the cardinality of the type `α`. The parsing starts
at "1", so `"1"` is parsed in as `nat.cast 0`. Providing `"0"` to the parser causes a failure.
-/
def numeral.from_one.of_fintype [fintype α] : parser α :=
@[derive mono] def numeral.from_one.of_fintype [fintype α] : parser α :=
do
c ← nat,
decorate_error (sformat!"<positive numeral less than or equal to {to_string (fintype.card α)}>")
Expand All @@ -87,7 +88,7 @@ Parse a character as a numeral while casting it to target type `α`,
The parser ensures that the character parsed in is within the bounds set by `fromc` and `toc`,
and subtracts the value of `fromc` from the parsed in character.
-/
def numeral.char (fromc toc : char) : parser α :=
@[derive mono] def numeral.char (fromc toc : char) : parser α :=
do
c ← decorate_error
(sformat!"<char between '{fromc.to_string}' to '{toc.to_string}' inclusively>")
Expand All @@ -101,7 +102,7 @@ The parser ensures that the character parsed in is greater or equal to `fromc` a
and subtracts the value of `fromc` from the parsed in character. There is also a check
that the resulting value is within the cardinality of the type `α`.
-/
def numeral.char.of_fintype [fintype α] (fromc : char) : parser α :=
@[derive mono] def numeral.char.of_fintype [fintype α] (fromc : char) : parser α :=
do
c ← decorate_error
(sformat!"<char from '{fromc.to_string}' to '{
Expand Down

0 comments on commit 114f8ef

Please sign in to comment.