Library Coq.Numbers.DecimalString
Conversion between decimal numbers and Coq strings
Local Open Scope string_scope.
Parsing one char
Definition uint_of_char (a:ascii)(d:option uint) :=
match d with
| None => None
| Some d =>
match a with
| "0" => Some (D0 d)
| "1" => Some (D1 d)
| "2" => Some (D2 d)
| "3" => Some (D3 d)
| "4" => Some (D4 d)
| "5" => Some (D5 d)
| "6" => Some (D6 d)
| "7" => Some (D7 d)
| "8" => Some (D8 d)
| "9" => Some (D9 d)
| _ => None
end
end%char.
Lemma uint_of_char_spec c d d' :
uint_of_char c (Some d) = Some d' ->
(c = "0" /\ d' = D0 d \/
c = "1" /\ d' = D1 d \/
c = "2" /\ d' = D2 d \/
c = "3" /\ d' = D3 d \/
c = "4" /\ d' = D4 d \/
c = "5" /\ d' = D5 d \/
c = "6" /\ d' = D6 d \/
c = "7" /\ d' = D7 d \/
c = "8" /\ d' = D8 d \/
c = "9" /\ d' = D9 d)%char.
Decimal/String conversion where Nil is ""
Module NilEmpty.
Fixpoint string_of_uint (d:uint) :=
match d with
| Nil => EmptyString
| D0 d => String "0" (string_of_uint d)
| D1 d => String "1" (string_of_uint d)
| D2 d => String "2" (string_of_uint d)
| D3 d => String "3" (string_of_uint d)
| D4 d => String "4" (string_of_uint d)
| D5 d => String "5" (string_of_uint d)
| D6 d => String "6" (string_of_uint d)
| D7 d => String "7" (string_of_uint d)
| D8 d => String "8" (string_of_uint d)
| D9 d => String "9" (string_of_uint d)
end.
Fixpoint uint_of_string s :=
match s with
| EmptyString => Some Nil
| String a s => uint_of_char a (uint_of_string s)
end.
Definition string_of_int (d:int) :=
match d with
| Pos d => string_of_uint d
| Neg d => String "-" (string_of_uint d)
end.
Definition int_of_string s :=
match s with
| EmptyString => Some (Pos Nil)
| String a s' =>
if Ascii.eqb a "-" then option_map Neg (uint_of_string s')
else option_map Pos (uint_of_string s)
end.
Corresponding proofs
Lemma usu d :
uint_of_string (string_of_uint d) = Some d.
Lemma sus s d :
uint_of_string s = Some d -> string_of_uint d = s.
Lemma isi d : int_of_string (string_of_int d) = Some d.
Lemma sis s d :
int_of_string s = Some d -> string_of_int d = s.
End NilEmpty.
Decimal/String conversions where Nil is "0"
Module NilZero.
Definition string_of_uint (d:uint) :=
match d with
| Nil => "0"
| _ => NilEmpty.string_of_uint d
end.
Definition uint_of_string s :=
match s with
| EmptyString => None
| _ => NilEmpty.uint_of_string s
end.
Definition string_of_int (d:int) :=
match d with
| Pos d => string_of_uint d
| Neg d => String "-" (string_of_uint d)
end.
Definition int_of_string s :=
match s with
| EmptyString => None
| String a s' =>
if Ascii.eqb a "-" then option_map Neg (uint_of_string s')
else option_map Pos (uint_of_string s)
end.
Corresponding proofs
Lemma uint_of_string_nonnil s : uint_of_string s <> Some Nil.
Lemma sus s d :
uint_of_string s = Some d -> string_of_uint d = s.
Lemma usu d :
d<>Nil -> uint_of_string (string_of_uint d) = Some d.
Lemma usu_nil :
uint_of_string (string_of_uint Nil) = Some Decimal.zero.
Lemma usu_gen d :
uint_of_string (string_of_uint d) = Some d \/
uint_of_string (string_of_uint d) = Some Decimal.zero.
Lemma isi d :
d<>Pos Nil -> d<>Neg Nil ->
int_of_string (string_of_int d) = Some d.
Lemma isi_posnil :
int_of_string (string_of_int (Pos Nil)) = Some (Pos Decimal.zero).
Warning! (-0) won't parse (compatibility with the behavior of Z).
Lemma isi_negnil :
int_of_string (string_of_int (Neg Nil)) = Some (Neg (D0 Nil)).
Lemma sis s d :
int_of_string s = Some d -> string_of_int d = s.
End NilZero.