diff --git a/src/data/buffer/parser/numeral.lean b/src/data/buffer/parser/numeral.lean index ab7b0a4ea1ecc..0e8bd460a16f4 100644 --- a/src/data/buffer/parser/numeral.lean +++ b/src/data/buffer/parser/numeral.lean @@ -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 @@ -43,7 +44,7 @@ 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 /-- @@ -51,7 +52,7 @@ 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!"") @@ -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 ("") @@ -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!"") @@ -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!"") @@ -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!"