aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
authorHugo Herbelin2020-05-15 16:51:29 +0200
committerHugo Herbelin2020-05-15 16:51:29 +0200
commita5c9ad83071c99110fed464a0b9a0f5e73f1be9b (patch)
tree4e436ada97fc8e74311e8c77312e164772957ac9
parentb5b6e2d4c8347cb25da6f827a6b6f06cb0f566e5 (diff)
parent31f5e89eaefcff04a04850d77b0c83cb24602f98 (diff)
Merge PR #11948: Hexadecimal numerals
Reviewed-by: JasonGross Ack-by: Zimmi48 Ack-by: herbelin
-rw-r--r--coqpp/coqpp_main.ml2
-rw-r--r--dev/ci/user-overlays/11948-proux01-hexadecimal.sh12
-rw-r--r--doc/changelog/03-notations/11948-hexadecimal.rst12
-rw-r--r--doc/sphinx/language/core/basic.rst19
-rw-r--r--doc/sphinx/user-extensions/syntax-extensions.rst39
-rw-r--r--doc/stdlib/index-list.html.template10
-rw-r--r--interp/constrexpr_ops.ml2
-rw-r--r--interp/constrextern.ml2
-rw-r--r--interp/notation.ml180
-rw-r--r--interp/notation.mli23
-rw-r--r--interp/numTok.ml188
-rw-r--r--interp/numTok.mli52
-rw-r--r--plugins/syntax/float_syntax.ml5
-rw-r--r--plugins/syntax/numeral.ml97
-rw-r--r--plugins/syntax/r_syntax.ml18
-rw-r--r--test-suite/output/FloatSyntax.out16
-rw-r--r--test-suite/output/FloatSyntax.v6
-rw-r--r--test-suite/output/Int63Syntax.out26
-rw-r--r--test-suite/output/Int63Syntax.v13
-rw-r--r--test-suite/output/NatSyntax.out34
-rw-r--r--test-suite/output/NatSyntax.v18
-rw-r--r--test-suite/output/NumeralNotations.out6
-rw-r--r--test-suite/output/NumeralNotations.v86
-rw-r--r--test-suite/output/QArithSyntax.out12
-rw-r--r--test-suite/output/QArithSyntax.v6
-rw-r--r--test-suite/output/RealSyntax.out15
-rw-r--r--test-suite/output/RealSyntax.v6
-rw-r--r--test-suite/output/Search.out196
-rw-r--r--test-suite/output/SearchHead.out22
-rw-r--r--test-suite/output/SearchPattern.out77
-rw-r--r--test-suite/output/ZSyntax.out8
-rw-r--r--test-suite/output/ZSyntax.v8
-rw-r--r--test-suite/output/bug_12159.v8
-rw-r--r--theories/Init/Datatypes.v3
-rw-r--r--theories/Init/Decimal.v10
-rw-r--r--theories/Init/Hexadecimal.v245
-rw-r--r--theories/Init/Nat.v62
-rw-r--r--theories/Init/Numeral.v33
-rw-r--r--theories/Init/Prelude.v17
-rw-r--r--theories/NArith/BinNat.v2
-rw-r--r--theories/NArith/BinNatDef.v36
-rw-r--r--theories/Numbers/AltBinNotations.v8
-rw-r--r--theories/Numbers/DecimalFacts.v432
-rw-r--r--theories/Numbers/DecimalN.v56
-rw-r--r--theories/Numbers/DecimalNat.v226
-rw-r--r--theories/Numbers/DecimalPos.v321
-rw-r--r--theories/Numbers/DecimalQ.v561
-rw-r--r--theories/Numbers/DecimalString.v244
-rw-r--r--theories/Numbers/DecimalZ.v85
-rw-r--r--theories/Numbers/HexadecimalFacts.v340
-rw-r--r--theories/Numbers/HexadecimalN.v107
-rw-r--r--theories/Numbers/HexadecimalNat.v321
-rw-r--r--theories/Numbers/HexadecimalPos.v446
-rw-r--r--theories/Numbers/HexadecimalQ.v533
-rw-r--r--theories/Numbers/HexadecimalString.v286
-rw-r--r--theories/Numbers/HexadecimalZ.v142
-rw-r--r--theories/Numbers/NaryFunctions.v54
-rw-r--r--theories/PArith/BinPos.v2
-rw-r--r--theories/PArith/BinPosDef.v85
-rw-r--r--theories/QArith/QArith_base.v46
-rw-r--r--theories/ZArith/BinInt.v2
-rw-r--r--theories/ZArith/BinIntDef.v33
62 files changed, 5132 insertions, 830 deletions
diff --git a/coqpp/coqpp_main.ml b/coqpp/coqpp_main.ml
index 43cd6f1784..39ca5413cc 100644
--- a/coqpp/coqpp_main.ml
+++ b/coqpp/coqpp_main.ml
@@ -202,7 +202,7 @@ function
| "PATTERNIDENT", s -> fprintf fmt "Tok.PPATTERNIDENT (%a)" print_pat s
| "FIELD", s -> fprintf fmt "Tok.PFIELD (%a)" print_pat s
| "NUMERAL", None -> fprintf fmt "Tok.PNUMERAL None"
-| "NUMERAL", Some s -> fprintf fmt "Tok.PNUMERAL (Some (NumTok.int %a))" print_string s
+| "NUMERAL", Some s -> fprintf fmt "Tok.PNUMERAL (Some (NumTok.Unsigned.of_string %a))" print_string s
| "STRING", s -> fprintf fmt "Tok.PSTRING (%a)" print_pat s
| "LEFTQMARK", None -> fprintf fmt "Tok.PLEFTQMARK"
| "BULLET", s -> fprintf fmt "Tok.PBULLET (%a)" print_pat s
diff --git a/dev/ci/user-overlays/11948-proux01-hexadecimal.sh b/dev/ci/user-overlays/11948-proux01-hexadecimal.sh
new file mode 100644
index 0000000000..0b3133d1f1
--- /dev/null
+++ b/dev/ci/user-overlays/11948-proux01-hexadecimal.sh
@@ -0,0 +1,12 @@
+if [ "$CI_PULL_REQUEST" = "11948" ] || [ "$CI_BRANCH" = "hexadecimal" ]; then
+
+ stdlib2_CI_REF=hexadecimal
+ stdlib2_CI_GITURL=https://github.com/proux01/stdlib2
+
+ paramcoq_CI_REF=hexadecimal
+ paramcoq_CI_GITURL=https://github.com/proux01/paramcoq
+
+ metacoq_CI_REF=hexadecimal
+ metacoq_CI_GITURL=https://github.com/proux01/metacoq
+
+fi
diff --git a/doc/changelog/03-notations/11948-hexadecimal.rst b/doc/changelog/03-notations/11948-hexadecimal.rst
new file mode 100644
index 0000000000..e4b76d238c
--- /dev/null
+++ b/doc/changelog/03-notations/11948-hexadecimal.rst
@@ -0,0 +1,12 @@
+- **Added:**
+ Numeral notations now parse hexadecimal constants such as ``0x2a``
+ or ``0xb.2ap-2``. Parsers added for :g:`nat`, :g:`positive`, :g:`Z`,
+ :g:`N`, :g:`Q`, :g:`R`, primitive integers and primitive floats
+ (`#11948 <https://github.com/coq/coq/pull/11948>`_,
+ by Pierre Roux).
+- **Deprecated:**
+ Numeral Notation on ``Decimal.uint``, ``Decimal.int`` and
+ ``Decimal.decimal`` are replaced respectively by numeral notations
+ on ``Numeral.uint``, ``Numeral.int`` and ``Numeral.numeral``
+ (`#11948 <https://github.com/coq/coq/pull/11948>`_,
+ by Pierre Roux).
diff --git a/doc/sphinx/language/core/basic.rst b/doc/sphinx/language/core/basic.rst
index aa93b4d21f..250a0f0326 100644
--- a/doc/sphinx/language/core/basic.rst
+++ b/doc/sphinx/language/core/basic.rst
@@ -111,18 +111,27 @@ Identifiers
Numerals
Numerals are sequences of digits with an optional fractional part
- and exponent, optionally preceded by a minus sign. :n:`@int` is an integer;
- a numeral without fractional or exponent parts. :n:`@num` is a non-negative
+ and exponent, optionally preceded by a minus sign. Hexadecimal numerals
+ start with ``0x`` or ``0X``. :n:`@int` is an integer;
+ a numeral without fractional nor exponent parts. :n:`@num` is a non-negative
integer. Underscores embedded in the digits are ignored, for example
``1_000_000`` is the same as ``1000000``.
.. insertprodn numeral digit
.. prodn::
- numeral ::= {+ @digit } {? . {+ @digit } } {? {| e | E } {? {| + | - } } {+ @digit } }
- int ::= {? - } {+ @digit }
- num ::= {+ @digit }
digit ::= 0 .. 9
+ digitu ::= {| @digit | _ }
+ hexdigit ::= {| 0 .. 9 | a..f | A..F }
+ hexdigitu ::= {| @hexdigit | _ }
+ decnum ::= @digit {* @digitu }
+ hexnum ::= {| 0x | 0X } @hexdigit {* @hexdigitu }
+ num ::= {| @decnum | @hexnum }
+ sign ::= {| + | - }
+ int ::= {? - } @num
+ decimal ::= @decnum {? . {+ @digitu } } {? {| e | E } {? @sign } @decnum }
+ hexadecimal ::= @hexnum {? . {+ @hexdigitu } } {? {| p | P } {? @sign } @decnum }
+ numeral ::= {? - } {| @decimal | @hexadecimal }
Strings
Strings begin and end with ``"`` (double quote). Use ``""`` to represent
diff --git a/doc/sphinx/user-extensions/syntax-extensions.rst b/doc/sphinx/user-extensions/syntax-extensions.rst
index c5ec636d5f..60cd4c4ad8 100644
--- a/doc/sphinx/user-extensions/syntax-extensions.rst
+++ b/doc/sphinx/user-extensions/syntax-extensions.rst
@@ -1452,26 +1452,32 @@ Numeral notations
parsing and printing functions, respectively. The parsing function
:n:`@qualid__parse` should have one of the following types:
- * :n:`Decimal.int -> @qualid`
- * :n:`Decimal.int -> option @qualid`
- * :n:`Decimal.uint -> @qualid`
- * :n:`Decimal.uint -> option @qualid`
+ * :n:`Numeral.int -> @qualid`
+ * :n:`Numeral.int -> option @qualid`
+ * :n:`Numeral.uint -> @qualid`
+ * :n:`Numeral.uint -> option @qualid`
* :n:`Z -> @qualid`
* :n:`Z -> option @qualid`
- * :n:`Decimal.decimal -> @qualid`
- * :n:`Decimal.decimal -> option @qualid`
+ * :n:`Numeral.numeral -> @qualid`
+ * :n:`Numeral.numeral -> option @qualid`
And the printing function :n:`@qualid__print` should have one of the
following types:
- * :n:`@qualid -> Decimal.int`
- * :n:`@qualid -> option Decimal.int`
- * :n:`@qualid -> Decimal.uint`
- * :n:`@qualid -> option Decimal.uint`
+ * :n:`@qualid -> Numeral.int`
+ * :n:`@qualid -> option Numeral.int`
+ * :n:`@qualid -> Numeral.uint`
+ * :n:`@qualid -> option Numeral.uint`
* :n:`@qualid -> Z`
* :n:`@qualid -> option Z`
- * :n:`@qualid -> Decimal.decimal`
- * :n:`@qualid -> option Decimal.decimal`
+ * :n:`@qualid -> Numeral.numeral`
+ * :n:`@qualid -> option Numeral.numeral`
+
+ .. deprecated:: 8.12
+ Numeral notations on :g:`Decimal.uint`, :g:`Decimal.int` and
+ :g:`Decimal.decimal` are replaced respectively by numeral
+ notations on :g:`Numeral.uint`, :g:`Numeral.int` and
+ :g:`Numeral.numeral`.
When parsing, the application of the parsing function
:n:`@qualid__parse` to the number will be fully reduced, and universes
@@ -1495,7 +1501,7 @@ Numeral notations
returns :n:`(@qualid__parse m)` when parsing a literal
:n:`m` that's greater than :n:`@numeral` rather than reducing it to a normal form.
Here :g:`m` will be a
- :g:`Decimal.int` or :g:`Decimal.uint` or :g:`Z`, depending on the
+ :g:`Numeral.int` or :g:`Numeral.uint` or :g:`Z`, depending on the
type of the parsing function :n:`@qualid__parse`. This allows for a
more compact representation of literals in types such as :g:`nat`,
and limits parse failures due to stack overflow. Note that a
@@ -1525,12 +1531,13 @@ Numeral notations
only for integers or non-negative integers, and the given numeral
has a fractional or exponent part or is negative.
- .. exn:: @qualid__parse should go from Decimal.int to @type or (option @type). Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).
+
+ .. exn:: @qualid__parse should go from Numeral.int to @type or (option @type). Instead of Numeral.int, the types Numeral.uint or Z or Int63.int or Numeral.numeral could be used (you may need to require BinNums or Numeral or Int63 first).
The parsing function given to the :cmd:`Numeral Notation`
vernacular is not of the right type.
- .. exn:: @qualid__print should go from @type to Decimal.int or (option Decimal.int). Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).
+ .. exn:: @qualid__print should go from @type to Numeral.int or (option Numeral.int). Instead of Numeral.int, the types Numeral.uint or Z or Int63.int or Numeral.numeral could be used (you may need to require BinNums or Numeral or Int63 first).
The printing function given to the :cmd:`Numeral Notation`
vernacular is not of the right type.
@@ -1546,7 +1553,7 @@ Numeral notations
Parsing functions expected to return an :g:`option` must always
return a concrete :g:`Some` or :g:`None` when applied to a
- concrete numeral expressed as a decimal. They may not return
+ concrete numeral expressed as a (hexa)decimal. They may not return
opaque constants.
String notations
diff --git a/doc/stdlib/index-list.html.template b/doc/stdlib/index-list.html.template
index 4a62888552..426f67eb53 100644
--- a/doc/stdlib/index-list.html.template
+++ b/doc/stdlib/index-list.html.template
@@ -21,6 +21,8 @@ through the <tt>Require Import</tt> command.</p>
theories/Init/Byte.v
theories/Init/Nat.v
theories/Init/Decimal.v
+ theories/Init/Hexadecimal.v
+ theories/Init/Numeral.v
theories/Init/Peano.v
theories/Init/Specif.v
theories/Init/Tactics.v
@@ -235,7 +237,15 @@ through the <tt>Require Import</tt> command.</p>
theories/Numbers/DecimalPos.v
theories/Numbers/DecimalN.v
theories/Numbers/DecimalZ.v
+ theories/Numbers/DecimalQ.v
theories/Numbers/DecimalString.v
+ theories/Numbers/HexadecimalFacts.v
+ theories/Numbers/HexadecimalNat.v
+ theories/Numbers/HexadecimalPos.v
+ theories/Numbers/HexadecimalN.v
+ theories/Numbers/HexadecimalZ.v
+ theories/Numbers/HexadecimalQ.v
+ theories/Numbers/HexadecimalString.v
</dd>
<dt> <b>&nbsp;&nbsp;NatInt</b>:
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index d6097304ec..a8fb5a3f45 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -15,8 +15,8 @@ open Nameops
open Libnames
open Namegen
open Glob_term
-open Constrexpr
open Notation
+open Constrexpr
(***********************)
(* For binders parsing *)
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index d5a5bde616..3f7bb6e330 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -19,13 +19,13 @@ open Libnames
open Namegen
open Impargs
open CAst
+open Notation
open Constrexpr
open Constrexpr_ops
open Notation_ops
open Glob_term
open Glob_ops
open Pattern
-open Notation
open Detyping
module NamedDecl = Context.Named.Declaration
diff --git a/interp/notation.ml b/interp/notation.ml
index fb3cefd624..3f13476355 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -382,7 +382,7 @@ module InnerPrimToken = struct
| _ -> false
let mkNumeral n =
- Numeral (NumTok.Signed.of_bigint n)
+ Numeral (NumTok.Signed.of_bigint CDec n)
let mkString = function
| None -> None
@@ -423,23 +423,32 @@ type numnot_option =
| Abstract of NumTok.UnsignedNat.t
type int_ty =
- { uint : Names.inductive;
+ { dec_uint : Names.inductive;
+ dec_int : Names.inductive;
+ hex_uint : Names.inductive;
+ hex_int : Names.inductive;
+ uint : Names.inductive;
int : Names.inductive }
type z_pos_ty =
{ z_ty : Names.inductive;
pos_ty : Names.inductive }
-type decimal_ty =
+type numeral_ty =
{ int : int_ty;
- decimal : Names.inductive }
+ decimal : Names.inductive;
+ hexadecimal : Names.inductive;
+ numeral : Names.inductive }
type target_kind =
- | Int of int_ty (* Coq.Init.Decimal.int + uint *)
- | UInt of Names.inductive (* Coq.Init.Decimal.uint *)
+ | Int of int_ty (* Coq.Init.Numeral.int + uint *)
+ | UInt of int_ty (* Coq.Init.Numeral.uint *)
| Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *)
| Int63 (* Coq.Numbers.Cyclic.Int63.Int63.int *)
- | Decimal of decimal_ty (* Coq.Init.Decimal.decimal + uint + int *)
+ | Numeral of numeral_ty (* Coq.Init.Numeral.numeral + uint + int *)
+ | DecimalInt of int_ty (* Coq.Init.Decimal.int + uint (deprecated) *)
+ | DecimalUInt of int_ty (* Coq.Init.Decimal.uint (deprecated) *)
+ | Decimal of numeral_ty (* Coq.Init.Decimal.Decimal + uint + int (deprecated) *)
type string_target_kind =
| ListByte
@@ -601,17 +610,23 @@ let warn_abstract_large_num =
(** Decimal.Nil has index 1, then Decimal.D0 has index 2 .. Decimal.D9 is 11 *)
let digit_of_char c =
- assert ('0' <= c && c <= '9');
- Char.code c - Char.code '0' + 2
+ assert ('0' <= c && c <= '9' || 'a' <= c && c <= 'f');
+ if c <= '9' then Char.code c - Char.code '0' + 2
+ else Char.code c - Char.code 'a' + 12
let char_of_digit n =
- assert (2<=n && n<=11);
- Char.chr (n-2 + Char.code '0')
+ assert (2<=n && n<=17);
+ if n <= 11 then Char.chr (n-2 + Char.code '0')
+ else Char.chr (n-12 + Char.code 'a')
-let coquint_of_rawnum uint n =
+let coquint_of_rawnum inds c n =
+ let uint = match c with CDec -> inds.dec_uint | CHex -> inds.hex_uint in
let nil = mkConstruct (uint,1) in
match n with None -> nil | Some n ->
let str = NumTok.UnsignedNat.to_string n in
+ let str = match c with
+ | CDec -> str
+ | CHex -> String.sub str 2 (String.length str - 2) in (* cut "0x" *)
let rec do_chars s i acc =
if i < 0 then acc
else
@@ -620,61 +635,126 @@ let coquint_of_rawnum uint n =
in
do_chars str (String.length str - 1) nil
-let coqint_of_rawnum inds (sign,n) =
- let uint = coquint_of_rawnum inds.uint (Some n) in
+let coqint_of_rawnum inds c (sign,n) =
+ let ind = match c with CDec -> inds.dec_int | CHex -> inds.hex_int in
+ let uint = coquint_of_rawnum inds c (Some n) in
let pos_neg = match sign with SPlus -> 1 | SMinus -> 2 in
- mkApp (mkConstruct (inds.int, pos_neg), [|uint|])
+ mkApp (mkConstruct (ind, pos_neg), [|uint|])
-let coqdecimal_of_rawnum inds n =
- let i, f, e = NumTok.Signed.to_decimal_and_exponent n in
- let i = coqint_of_rawnum inds.int i in
- let f = coquint_of_rawnum inds.int.uint f in
+let coqnumeral_of_rawnum inds c n =
+ let ind = match c with CDec -> inds.decimal | CHex -> inds.hexadecimal in
+ let i, f, e = NumTok.Signed.to_int_frac_and_exponent n in
+ let i = coqint_of_rawnum inds.int c i in
+ let f = coquint_of_rawnum inds.int c f in
match e with
- | None -> mkApp (mkConstruct (inds.decimal, 1), [|i; f|]) (* Decimal *)
+ | None -> mkApp (mkConstruct (ind, 1), [|i; f|]) (* (D|Hexad)ecimal *)
| Some e ->
- let e = coqint_of_rawnum inds.int e in
- mkApp (mkConstruct (inds.decimal, 2), [|i; f; e|]) (* DecimalExp *)
+ let e = coqint_of_rawnum inds.int CDec e in
+ mkApp (mkConstruct (ind, 2), [|i; f; e|]) (* (D|Hexad)ecimalExp *)
-let rawnum_of_coquint c =
+let mkDecHex ind c n = match c with
+ | CDec -> mkApp (mkConstruct (ind, 1), [|n|]) (* (UInt|Int|)Dec *)
+ | CHex -> mkApp (mkConstruct (ind, 2), [|n|]) (* (UInt|Int|)Hex *)
+
+exception NonDecimal
+
+let decimal_coqnumeral_of_rawnum inds n =
+ if NumTok.Signed.classify n <> CDec then raise NonDecimal;
+ coqnumeral_of_rawnum inds CDec n
+
+let coqnumeral_of_rawnum inds n =
+ let c = NumTok.Signed.classify n in
+ let n = coqnumeral_of_rawnum inds c n in
+ mkDecHex inds.numeral c n
+
+let decimal_coquint_of_rawnum inds n =
+ if NumTok.UnsignedNat.classify n <> CDec then raise NonDecimal;
+ coquint_of_rawnum inds CDec (Some n)
+
+let coquint_of_rawnum inds n =
+ let c = NumTok.UnsignedNat.classify n in
+ let n = coquint_of_rawnum inds c (Some n) in
+ mkDecHex inds.uint c n
+
+let decimal_coqint_of_rawnum inds n =
+ if NumTok.SignedNat.classify n <> CDec then raise NonDecimal;
+ coqint_of_rawnum inds CDec n
+
+let coqint_of_rawnum inds n =
+ let c = NumTok.SignedNat.classify n in
+ let n = coqint_of_rawnum inds c n in
+ mkDecHex inds.int c n
+
+let rawnum_of_coquint cl c =
let rec of_uint_loop c buf =
match Constr.kind c with
| Construct ((_,1), _) (* Nil *) -> ()
| App (c, [|a|]) ->
(match Constr.kind c with
- | Construct ((_,n), _) (* D0 to D9 *) ->
+ | Construct ((_,n), _) (* D0 to Df *) ->
let () = Buffer.add_char buf (char_of_digit n) in
of_uint_loop a buf
| _ -> raise NotAValidPrimToken)
| _ -> raise NotAValidPrimToken
in
let buf = Buffer.create 64 in
+ if cl = CHex then (Buffer.add_char buf '0'; Buffer.add_char buf 'x');
let () = of_uint_loop c buf in
- if Int.equal (Buffer.length buf) 0 then
+ if Int.equal (Buffer.length buf) (match cl with CDec -> 0 | CHex -> 2) then
(* To avoid ambiguities between Nil and (D0 Nil), we choose
to not display Nil alone as "0" *)
raise NotAValidPrimToken
else NumTok.UnsignedNat.of_string (Buffer.contents buf)
-let rawnum_of_coqint c =
+let rawnum_of_coqint cl c =
match Constr.kind c with
| App (c,[|c'|]) ->
(match Constr.kind c with
- | Construct ((_,1), _) (* Pos *) -> (SPlus,rawnum_of_coquint c')
- | Construct ((_,2), _) (* Neg *) -> (SMinus,rawnum_of_coquint c')
+ | Construct ((_,1), _) (* Pos *) -> (SPlus,rawnum_of_coquint cl c')
+ | Construct ((_,2), _) (* Neg *) -> (SMinus,rawnum_of_coquint cl c')
| _ -> raise NotAValidPrimToken)
| _ -> raise NotAValidPrimToken
-let rawnum_of_decimal c =
+let rawnum_of_coqnumeral cl c =
let of_ife i f e =
- let n = rawnum_of_coqint i in
- let f = try Some (rawnum_of_coquint f) with NotAValidPrimToken -> None in
- let e = match e with None -> None | Some e -> Some (rawnum_of_coqint e) in
- NumTok.Signed.of_decimal_and_exponent n f e in
+ let n = rawnum_of_coqint cl i in
+ let f = try Some (rawnum_of_coquint cl f) with NotAValidPrimToken -> None in
+ let e = match e with None -> None | Some e -> Some (rawnum_of_coqint CDec e) in
+ NumTok.Signed.of_int_frac_and_exponent n f e in
match Constr.kind c with
| App (_,[|i; f|]) -> of_ife i f None
| App (_,[|i; f; e|]) -> of_ife i f (Some e)
| _ -> raise NotAValidPrimToken
+let destDecHex c = match Constr.kind c with
+ | App (c,[|c'|]) ->
+ (match Constr.kind c with
+ | Construct ((_,1), _) (* (UInt|Int|)Dec *) -> CDec, c'
+ | Construct ((_,2), _) (* (UInt|Int|)Hex *) -> CHex, c'
+ | _ -> raise NotAValidPrimToken)
+ | _ -> raise NotAValidPrimToken
+
+let decimal_rawnum_of_coqnumeral c =
+ rawnum_of_coqnumeral CDec c
+
+let rawnum_of_coqnumeral c =
+ let cl, c = destDecHex c in
+ rawnum_of_coqnumeral cl c
+
+let decimal_rawnum_of_coquint c =
+ rawnum_of_coquint CDec c
+
+let rawnum_of_coquint c =
+ let cl, c = destDecHex c in
+ rawnum_of_coquint cl c
+
+let decimal_rawnum_of_coqint c =
+ rawnum_of_coqint CDec c
+
+let rawnum_of_coqint c =
+ let cl, c = destDecHex c in
+ rawnum_of_coqint cl c
+
(***********************************************************************)
(** ** Conversion between Coq [Z] and internal bigint *)
@@ -768,15 +848,24 @@ let interp o ?loc n =
let c = match fst o.to_kind, NumTok.Signed.to_int n with
| Int int_ty, Some n ->
coqint_of_rawnum int_ty n
- | UInt uint_ty, Some (SPlus, n) ->
- coquint_of_rawnum uint_ty (Some n)
+ | UInt int_ty, Some (SPlus, n) ->
+ coquint_of_rawnum int_ty n
+ | DecimalInt int_ty, Some n ->
+ (try decimal_coqint_of_rawnum int_ty n
+ with NonDecimal -> no_such_prim_token "number" ?loc o.ty_name)
+ | DecimalUInt int_ty, Some (SPlus, n) ->
+ (try decimal_coquint_of_rawnum int_ty n
+ with NonDecimal -> no_such_prim_token "number" ?loc o.ty_name)
| Z z_pos_ty, Some n ->
z_of_bigint z_pos_ty (NumTok.SignedNat.to_bigint n)
| Int63, Some n ->
interp_int63 ?loc (NumTok.SignedNat.to_bigint n)
- | (Int _ | UInt _ | Z _ | Int63), _ ->
+ | (Int _ | UInt _ | DecimalInt _ | DecimalUInt _ | Z _ | Int63), _ ->
no_such_prim_token "number" ?loc o.ty_name
- | Decimal decimal_ty, _ -> coqdecimal_of_rawnum decimal_ty n
+ | Numeral numeral_ty, _ -> coqnumeral_of_rawnum numeral_ty n
+ | Decimal numeral_ty, _ ->
+ (try decimal_coqnumeral_of_rawnum numeral_ty n
+ with NonDecimal -> no_such_prim_token "number" ?loc o.ty_name)
in
let env = Global.env () in
let sigma = Evd.from_env env in
@@ -797,9 +886,12 @@ let uninterp o n =
begin function
| (Int _, c) -> NumTok.Signed.of_int (rawnum_of_coqint c)
| (UInt _, c) -> NumTok.Signed.of_nat (rawnum_of_coquint c)
- | (Z _, c) -> NumTok.Signed.of_bigint (bigint_of_z c)
- | (Int63, c) -> NumTok.Signed.of_bigint (bigint_of_int63 c)
- | (Decimal _, c) -> rawnum_of_decimal c
+ | (Z _, c) -> NumTok.Signed.of_bigint CDec (bigint_of_z c)
+ | (Int63, c) -> NumTok.Signed.of_bigint CDec (bigint_of_int63 c)
+ | (Numeral _, c) -> rawnum_of_coqnumeral c
+ | (DecimalInt _, c) -> NumTok.Signed.of_int (decimal_rawnum_of_coqint c)
+ | (DecimalUInt _, c) -> NumTok.Signed.of_nat (decimal_rawnum_of_coquint c)
+ | (Decimal _, c) -> decimal_rawnum_of_coqnumeral c
end o n
end
@@ -1126,8 +1218,8 @@ let find_notation ntn sc =
NotationMap.find ntn (find_scope sc).notations
let notation_of_prim_token = function
- | Numeral (SPlus,n) -> InConstrEntrySomeLevel, NumTok.Unsigned.sprint n
- | Numeral (SMinus,n) -> InConstrEntrySomeLevel, "- "^NumTok.Unsigned.sprint n
+ | Constrexpr.Numeral (SPlus,n) -> InConstrEntrySomeLevel, NumTok.Unsigned.sprint n
+ | Constrexpr.Numeral (SMinus,n) -> InConstrEntrySomeLevel, "- "^NumTok.Unsigned.sprint n
| String _ -> raise Not_found
let find_prim_token check_allowed ?loc p sc =
@@ -1333,14 +1425,14 @@ let availability_of_prim_token n printer_scope local_scopes =
let uid = snd (String.Map.find scope !prim_token_interp_infos) in
let open InnerPrimToken in
match n, uid with
- | Numeral _, NumeralNotation _ -> true
+ | Constrexpr.Numeral _, NumeralNotation _ -> true
| _, NumeralNotation _ -> false
| String _, StringNotation _ -> true
| _, StringNotation _ -> false
| _, Uid uid ->
let interp = Hashtbl.find prim_token_interpreters uid in
match n, interp with
- | Numeral _, (RawNumInterp _ | BigNumInterp _) -> true
+ | Constrexpr.Numeral _, (RawNumInterp _ | BigNumInterp _) -> true
| String _, StringInterp _ -> true
| _ -> false
with Not_found -> false
diff --git a/interp/notation.mli b/interp/notation.mli
index 96a76c4de6..e7e917463b 100644
--- a/interp/notation.mli
+++ b/interp/notation.mli
@@ -74,7 +74,7 @@ val find_delimiters_scope : ?loc:Loc.t -> delimiters -> scope_name
(** {6 Declare and uses back and forth an interpretation of primitive token } *)
-(** A numeral interpreter is the pair of an interpreter for **decimal**
+(** A numeral interpreter is the pair of an interpreter for **(hexa)decimal**
numbers in terms and an optional interpreter in pattern, if
non integer or negative numbers are not supported, the interpreter
must fail with an appropriate error message *)
@@ -120,23 +120,32 @@ type numnot_option =
| Abstract of NumTok.UnsignedNat.t
type int_ty =
- { uint : Names.inductive;
+ { dec_uint : Names.inductive;
+ dec_int : Names.inductive;
+ hex_uint : Names.inductive;
+ hex_int : Names.inductive;
+ uint : Names.inductive;
int : Names.inductive }
type z_pos_ty =
{ z_ty : Names.inductive;
pos_ty : Names.inductive }
-type decimal_ty =
+type numeral_ty =
{ int : int_ty;
- decimal : Names.inductive }
+ decimal : Names.inductive;
+ hexadecimal : Names.inductive;
+ numeral : Names.inductive }
type target_kind =
- | Int of int_ty (* Coq.Init.Decimal.int + uint *)
- | UInt of Names.inductive (* Coq.Init.Decimal.uint *)
+ | Int of int_ty (* Coq.Init.Numeral.int + uint *)
+ | UInt of int_ty (* Coq.Init.Numeral.uint *)
| Z of z_pos_ty (* Coq.Numbers.BinNums.Z and positive *)
| Int63 (* Coq.Numbers.Cyclic.Int63.Int63.int *)
- | Decimal of decimal_ty (* Coq.Init.Decimal.decimal + uint + int *)
+ | Numeral of numeral_ty (* Coq.Init.Numeral.numeral + uint + int *)
+ | DecimalInt of int_ty (* Coq.Init.Decimal.int + uint (deprecated) *)
+ | DecimalUInt of int_ty (* Coq.Init.Decimal.uint (deprecated) *)
+ | Decimal of numeral_ty (* Coq.Init.Decimal.Decimal + uint + int (deprecated) *)
type string_target_kind =
| ListByte
diff --git a/interp/numTok.ml b/interp/numTok.ml
index e254e9e972..bb14649b91 100644
--- a/interp/numTok.ml
+++ b/interp/numTok.ml
@@ -12,6 +12,10 @@
e.g. "e"/"E" or the presence of leading 0s, or the presence of a +
in the exponent *)
+type num_class = CDec | CHex
+
+let string_del_head n s = String.sub s n (String.length s - n)
+
module UnsignedNat =
struct
type t = string
@@ -20,12 +24,15 @@ struct
assert (s.[0] >= '0' && s.[0] <= '9');
s
let to_string s =
- String.(concat "" (split_on_char '_' s))
+ String.(map Char.lowercase_ascii (concat "" (split_on_char '_' s)))
let sprint s = s
let print s = Pp.str (sprint s)
- (** Comparing two raw numbers (base 10, big-endian, non-negative).
+ let classify s =
+ if String.length s >= 2 && (s.[1] = 'x' || s.[1] = 'X') then CHex else CDec
+
+ (** Comparing two raw numbers (base 10 or 16, big-endian, non-negative).
A bit nasty, but not critical: used e.g. to decide when a number
is considered as large (see threshold warnings in notation.ml). *)
@@ -45,12 +52,20 @@ struct
0
with Comp c -> c
- let is_zero s =
- compare s "0" = 0
+ let compare n d =
+ assert (classify d = CDec);
+ match classify n with
+ | CDec -> compare (to_string n) (to_string d)
+ | CHex -> compare (string_del_head 2 (to_string n)) (to_string d)
+
+ let is_zero s =
+ compare s "0" = 0
end
type sign = SPlus | SMinus
+type 'a exp = EDec of 'a | EBin of 'a
+
module SignedNat =
struct
type t = sign * UnsignedNat.t
@@ -58,26 +73,89 @@ struct
assert (String.length s > 0);
let sign,n =
match s.[0] with
- | '-' -> (SMinus,String.sub s 1 (String.length s - 1))
- | '+' -> (SPlus,String.sub s 1 (String.length s - 1))
+ | '-' -> (SMinus,string_del_head 1 s)
+ | '+' -> (SPlus,string_del_head 1 s)
| _ -> (SPlus,s) in
(sign,UnsignedNat.of_string n)
let to_string (sign,n) =
(match sign with SPlus -> "" | SMinus -> "-") ^ UnsignedNat.to_string n
- let to_bigint n = Bigint.of_string (to_string n)
- let of_bigint n =
+ let classify (_,n) = UnsignedNat.classify n
+ let bigint_of_string (sign,n) =
+ (* nasty code to remove when switching to zarith
+ since zarith's of_string handles hexadecimal *)
+ match UnsignedNat.classify n with
+ | CDec -> Bigint.of_string (to_string (sign,n))
+ | CHex ->
+ let int_of_char c = match c with
+ | 'a'..'f' -> 10 + int_of_char c - int_of_char 'a'
+ | _ -> int_of_char c - int_of_char '0' in
+ let c16 = Bigint.of_int 16 in
+ let s = UnsignedNat.to_string n in
+ let n = ref Bigint.zero in
+ let len = String.length s in
+ for d = 2 to len - 1 do
+ n := Bigint.(add (mult !n c16) (of_int (int_of_char s.[d])))
+ done;
+ match sign with SPlus -> !n | SMinus -> Bigint.neg !n
+ let to_bigint n = bigint_of_string n
+ let string_of_nonneg_bigint c n =
+ (* nasty code to remove when switching to zarith
+ since zarith's format handles hexadecimal *)
+ match c with
+ | CDec -> Bigint.to_string n
+ | CHex ->
+ let div16 n =
+ let n, r0 = Bigint.div2_with_rest n in
+ let n, r1 = Bigint.div2_with_rest n in
+ let n, r2 = Bigint.div2_with_rest n in
+ let n, r3 = Bigint.div2_with_rest n in
+ let r = match r3, r2, r1, r0 with
+ | false, false, false, false -> "0"
+ | false, false, false, true -> "1"
+ | false, false, true, false -> "2"
+ | false, false, true, true -> "3"
+ | false, true, false, false -> "4"
+ | false, true, false, true -> "5"
+ | false, true, true, false -> "6"
+ | false, true, true, true -> "7"
+ | true, false, false, false -> "8"
+ | true, false, false, true -> "9"
+ | true, false, true, false -> "a"
+ | true, false, true, true -> "b"
+ | true, true, false, false -> "c"
+ | true, true, false, true -> "d"
+ | true, true, true, false -> "e"
+ | true, true, true, true -> "f" in
+ n, r in
+ let n = ref n in
+ let l = ref [] in
+ while Bigint.is_strictly_pos !n do
+ let n', r = div16 !n in
+ n := n';
+ l := r :: !l
+ done;
+ "0x" ^ String.concat "" (List.rev !l)
+ let of_bigint c n =
let sign, n = if Bigint.is_strictly_neg n then (SMinus, Bigint.neg n) else (SPlus, n) in
- (sign, Bigint.to_string n)
+ (sign, string_of_nonneg_bigint c n)
end
module Unsigned =
struct
type t = {
- int : string; (** \[0-9\]\[0-9_\]* *)
- frac : string; (** empty or \[0-9_\]+ *)
- exp : string (** empty or \[eE\]\[+-\]?\[0-9\]\[0-9_\]* *)
+ int : string;
+ frac : string;
+ exp : string
}
+ (**
+ - int: \[0-9\]\[0-9_\]*
+ - frac: empty or \[0-9_\]+
+ - exp: empty or \[eE\]\[+-\]?\[0-9\]\[0-9_\]*
+ or
+ - int: 0\[xX\]\[0-9a-fA-F\]\[0-9a-fA-F_\]*
+ - frac: empty or \[0-9a-fA-F_\]+
+ - exp: empty or \[pP\]\[+-\]?\[0-9\]\[0-9_\]* *)
let equal n1 n2 =
String.(equal n1.int n2.int && equal n1.frac n2.frac && equal n1.exp n2.exp)
@@ -96,19 +174,35 @@ struct
| Some ('0'..'9' as c) -> Stream.junk s; number (store len c) s
| Some ('_' as c) when len > 0 -> Stream.junk s; number (store len c) s
| _ -> len in
+ (* reads [0-9a-fA-F_]* *)
+ let rec hex_number len s = match Stream.peek s with
+ | Some (('0'..'9' | 'a'..'f' | 'A'..'F') as c) ->
+ Stream.junk s; hex_number (store len c) s
+ | Some ('_' as c) when len > 0 ->
+ Stream.junk s; hex_number (store len c) s
+ | _ -> len in
fun s ->
- let i = get_buff (number 0 s) in
+ let hex, i =
+ match Stream.npeek 3 s with
+ | '0' :: (('x' | 'X') as x) :: (('0'..'9' | 'a'..'f' | 'A'..'F') as c) :: _ ->
+ Stream.junk s; Stream.junk s; Stream.junk s;
+ true, get_buff (hex_number (store (store (store 0 '0') x) c) s)
+ | _ -> false, get_buff (number 0 s) in
assert (i <> "");
let f =
- match Stream.npeek 2 s with
- | '.' :: (('0'..'9' | '_') as c) :: _ ->
+ match hex, Stream.npeek 2 s with
+ | true, '.' :: (('0'..'9' | 'a'..'f' | 'A'..'F' | '_') as c) :: _ ->
+ Stream.junk s; Stream.junk s; get_buff (hex_number (store 0 c) s)
+ | false, '.' :: (('0'..'9' | '_') as c) :: _ ->
Stream.junk s; Stream.junk s; get_buff (number (store 0 c) s)
| _ -> "" in
let e =
- match (Stream.npeek 2 s) with
- | (('e'|'E') as e) :: ('0'..'9' as c) :: _ ->
+ match hex, Stream.npeek 2 s with
+ | true, (('p'|'P') as e) :: ('0'..'9' as c) :: _
+ | false, (('e'|'E') as e) :: ('0'..'9' as c) :: _ ->
Stream.junk s; Stream.junk s; get_buff (number (store (store 0 e) c) s)
- | (('e'|'E') as e) :: (('+'|'-') as sign) :: _ ->
+ | true, (('p'|'P') as e) :: (('+'|'-') as sign) :: _
+ | false, (('e'|'E') as e) :: (('+'|'-') as sign) :: _ ->
begin match Stream.npeek 3 s with
| _ :: _ :: ('0'..'9' as c) :: _ ->
Stream.junk s; Stream.junk s; Stream.junk s;
@@ -146,6 +240,7 @@ struct
| { int = _; frac = ""; exp = "" } -> true
| _ -> false
+ let classify n = UnsignedNat.classify n.int
end
open Unsigned
@@ -162,19 +257,30 @@ struct
| (SPlus,{int;frac;exp}) -> UnsignedNat.is_zero int && UnsignedNat.is_zero frac
| _ -> false
- let of_decimal_and_exponent (sign,int) f e =
- let exp = match e with Some e -> "e" ^ SignedNat.to_string e | None -> "" in
- let frac = match f with Some f -> UnsignedNat.to_string f | None -> "" in
+ let of_int_frac_and_exponent (sign,int) f e =
+ assert (match e with None -> true | Some e -> SignedNat.classify e = CDec);
+ let c = UnsignedNat.classify int in
+ let exp = match e with None -> "" | Some e ->
+ let e = SignedNat.to_string e in
+ match c with CDec -> "e" ^ e | CHex -> "p" ^ e in
+ let frac = match f with None -> "" | Some f ->
+ assert (c = UnsignedNat.classify f);
+ let f = UnsignedNat.to_string f in
+ match c with CDec -> f | CHex -> string_del_head 2 f in
sign, { int; frac; exp }
- let to_decimal_and_exponent (sign, { int; frac; exp }) =
+ let to_int_frac_and_exponent (sign, { int; frac; exp }) =
let e =
if exp = "" then None else
Some (match exp.[1] with
- | '-' -> SMinus, String.sub exp 2 (String.length exp - 2)
- | '+' -> SPlus, String.sub exp 2 (String.length exp - 2)
- | _ -> SPlus, String.sub exp 1 (String.length exp - 1)) in
- let f = if frac = "" then None else Some frac in
+ | '-' -> SMinus, string_del_head 2 exp
+ | '+' -> SPlus, string_del_head 2 exp
+ | _ -> SPlus, string_del_head 1 exp) in
+ let f =
+ if frac = "" then None else
+ Some (match UnsignedNat.classify int with
+ | CDec -> frac
+ | CHex -> "0x" ^ frac) in
(sign, int), f, e
let of_nat i =
@@ -211,8 +317,8 @@ struct
let of_string s =
assert (s <> "");
let sign,u = match s.[0] with
- | '-' -> (SMinus, String.sub s 1 (String.length s - 1))
- | '+' -> (SPlus, String.sub s 1 (String.length s - 1))
+ | '-' -> (SMinus, string_del_head 1 s)
+ | '+' -> (SPlus, string_del_head 1 s)
| _ -> (SPlus, s) in
(sign, Unsigned.of_string u)
@@ -224,27 +330,31 @@ struct
Some (SignedNat.to_bigint (sign,UnsignedNat.to_string n))
| _ -> None
- let of_bigint n =
- let sign, int = SignedNat.of_bigint n in
- (sign, { int; frac = ""; exp = "" })
+ let of_bigint c n =
+ of_int (SignedNat.of_bigint c n)
let to_bigint_and_exponent (s, { int; frac; exp }) =
- let s = match s with SPlus -> "" | SMinus -> "-" in
+ let c = UnsignedNat.classify int in
let int = UnsignedNat.to_string int in
let frac = UnsignedNat.to_string frac in
- let i = Bigint.of_string (s ^ int ^ frac) in
+ let i = SignedNat.to_bigint (s, int ^ frac) in
let e =
let e = if exp = "" then Bigint.zero else match exp.[1] with
- | '+' -> Bigint.of_string (UnsignedNat.to_string (String.sub exp 2 (String.length exp - 2)))
- | '-' -> Bigint.(neg (of_string (UnsignedNat.to_string (String.sub exp 2 (String.length exp - 2)))))
- | _ -> Bigint.of_string (UnsignedNat.to_string (String.sub exp 1 (String.length exp - 1))) in
- Bigint.(sub e (of_int (String.length (String.concat "" (String.split_on_char '_' frac))))) in
- (i,e)
+ | '+' -> Bigint.of_string (UnsignedNat.to_string (string_del_head 2 exp))
+ | '-' -> Bigint.(neg (of_string (UnsignedNat.to_string (string_del_head 2 exp))))
+ | _ -> Bigint.of_string (UnsignedNat.to_string (string_del_head 1 exp)) in
+ let l = String.length frac in
+ let l = match c with CDec -> l | CHex -> 4 * l in
+ Bigint.(sub e (of_int l)) in
+ (i, match c with CDec -> EDec e | CHex -> EBin e)
let of_bigint_and_exponent i e =
- of_decimal_and_exponent (SignedNat.of_bigint i) None (Some (SignedNat.of_bigint e))
+ let c = match e with EDec _ -> CDec | EBin _ -> CHex in
+ let e = match e with EDec e | EBin e -> Some (SignedNat.of_bigint CDec e) in
+ of_int_frac_and_exponent (SignedNat.of_bigint c i) None e
let is_bigger_int_than (s, { int; frac; exp }) i =
frac = "" && exp = "" && UnsignedNat.compare int i >= 0
+ let classify (_, n) = Unsigned.classify n
end
diff --git a/interp/numTok.mli b/interp/numTok.mli
index ea289df237..11d5a0f980 100644
--- a/interp/numTok.mli
+++ b/interp/numTok.mli
@@ -11,7 +11,7 @@
(** Numerals in different forms: signed or unsigned, possibly with
fractional part and exponent.
- Numerals are represented using raw strings of decimal
+ Numerals are represented using raw strings of (hexa)decimal
literals and a separate sign flag.
Note that this representation is not unique, due to possible
@@ -25,21 +25,29 @@
type sign = SPlus | SMinus
+type num_class = CDec | CHex
+
+type 'a exp = EDec of 'a | EBin of 'a
+
(** {6 String representation of a natural number } *)
module UnsignedNat :
sig
type t
val of_string : string -> t
- (** Convert from a non-empty sequence of digits (which may contain "_") *)
+ (** Convert from a non-empty sequence of digits (which may contain "_")
+ (or hexdigits when starting with "0x" or "0X") *)
val to_string : t -> string
- (** Convert to a non-empty sequence of digit that does not contain "_" *)
+ (** Convert to a non-empty sequence of digit that does not contain "_"
+ (or hexdigits, starting with "0x", all hexdigits are lower case) *)
val sprint : t -> string
val print : t -> Pp.t
(** [sprint] and [print] returns the numeral as it was parsed, for printing *)
+ val classify : t -> num_class
+
val compare : t -> t -> int
end
@@ -49,13 +57,16 @@ module SignedNat :
sig
type t = sign * UnsignedNat.t
val of_string : string -> t
- (** Convert from a non-empty sequence of digits which may contain "_" *)
+ (** Convert from a non-empty sequence of (hex)digits which may contain "_" *)
val to_string : t -> string
- (** Convert to a non-empty sequence of digit that does not contain "_" *)
+ (** Convert to a non-empty sequence of (hex)digit that does not contain "_"
+ (hexadecimals start with "0x" and all hexdigits are lower case) *)
+ val classify : t -> num_class
+
+ val of_bigint : num_class -> Bigint.bigint -> t
val to_bigint : t -> Bigint.bigint
- val of_bigint : Bigint.bigint -> t
end
(** {6 Unsigned decimal numerals } *)
@@ -78,11 +89,17 @@ sig
The recognized syntax is:
- integer part: \[0-9\]\[0-9_\]*
- - decimal part: empty or .\[0-9_\]+
- - exponent part: empty or \[eE\]\[+-\]?\[0-9\]\[0-9_\]* *)
+ - fractional part: empty or .\[0-9_\]+
+ - exponent part: empty or \[eE\]\[+-\]?\[0-9\]\[0-9_\]*
+ or
+ - integer part: 0\[xX\]\[0-9a-fA-F\]\[0-9a-fA-F_\]*
+ - fractional part: empty or .\[0-9a-fA-F_\]+
+ - exponent part: empty or \[pP\]\[+-\]?\[0-9\]\[0-9_\]* *)
val parse_string : string -> t option
- (** Parse the string as a positive Coq numeral, if possible *)
+ (** Parse the string as a non negative Coq numeral, if possible *)
+
+ val classify : t -> num_class
end
@@ -114,17 +131,20 @@ sig
val to_string : t -> string
(** Returns a string in the syntax of OCaml's float_of_string *)
- val of_bigint : Bigint.bigint -> t
+ val of_bigint : num_class -> Bigint.bigint -> t
val to_bigint : t -> Bigint.bigint option
(** Convert from and to bigint when the denotation of a bigint *)
- val of_decimal_and_exponent : SignedNat.t -> UnsignedNat.t option -> SignedNat.t option -> t
- val to_decimal_and_exponent : t -> SignedNat.t * UnsignedNat.t option * SignedNat.t option
- (** n, p and q such that the number is n.p*10^q *)
+ val of_int_frac_and_exponent : SignedNat.t -> UnsignedNat.t option -> SignedNat.t option -> t
+ val to_int_frac_and_exponent : t -> SignedNat.t * UnsignedNat.t option * SignedNat.t option
+ (** n, p and q such that the number is n.p*10^q or n.p*2^q
+ pre/postcondition: classify n = classify p, classify q = CDec *)
+
+ val of_bigint_and_exponent : Bigint.bigint -> Bigint.bigint exp -> t
+ val to_bigint_and_exponent : t -> Bigint.bigint * Bigint.bigint exp
+ (** n and p such that the number is n*10^p or n*2^p *)
- val to_bigint_and_exponent : t -> Bigint.bigint * Bigint.bigint
- val of_bigint_and_exponent : Bigint.bigint -> Bigint.bigint -> t
- (** n and p such that the number is n*10^p *)
+ val classify : t -> num_class
val is_bigger_int_than : t -> UnsignedNat.t -> bool
(** Test if an integer whose absolute value is bounded *)
diff --git a/plugins/syntax/float_syntax.ml b/plugins/syntax/float_syntax.ml
index e0a9906689..9861a060ab 100644
--- a/plugins/syntax/float_syntax.ml
+++ b/plugins/syntax/float_syntax.ml
@@ -42,7 +42,7 @@ let interp_float ?loc n =
| Float64.(PZero | NZero) -> not (NumTok.Signed.is_zero n)
| Float64.(PNormal | NNormal | PSubn | NSubn) ->
let m, e =
- let (_, i), f, e = NumTok.Signed.to_decimal_and_exponent n in
+ let (_, i), f, e = NumTok.Signed.to_int_frac_and_exponent n in
let i = NumTok.UnsignedNat.to_string i in
let f = match f with
| None -> "" | Some f -> NumTok.UnsignedNat.to_string f in
@@ -70,7 +70,8 @@ let interp_float ?loc n =
else (* e < 0 *)
if e' <= e then check m' (-e) m (e - e')
else check' m' (-e) (e' - e) m in
- if inexact () then warn_inexact_float ?loc (sn, f);
+ if NumTok.(Signed.classify n = CDec) && inexact () then
+ warn_inexact_float ?loc (sn, f);
DAst.make ?loc (GFloat f)
(* Pretty printing is already handled in constrextern.ml *)
diff --git a/plugins/syntax/numeral.ml b/plugins/syntax/numeral.ml
index 2250d57809..2db76719b8 100644
--- a/plugins/syntax/numeral.ml
+++ b/plugins/syntax/numeral.ml
@@ -55,24 +55,45 @@ let locate_z () =
}, mkRefC q_z)
else None
-let locate_decimal () =
- let int = "num.int.type" in
- let uint = "num.uint.type" in
+let locate_numeral () =
+ let dint = "num.int.type" in
+ let duint = "num.uint.type" in
let dec = "num.decimal.type" in
- if Coqlib.has_ref int && Coqlib.has_ref uint && Coqlib.has_ref dec
+ let hint = "num.hexadecimal_int.type" in
+ let huint = "num.hexadecimal_uint.type" in
+ let hex = "num.hexadecimal.type" in
+ let int = "num.num_int.type" in
+ let uint = "num.num_uint.type" in
+ let num = "num.numeral.type" in
+ if Coqlib.has_ref dint && Coqlib.has_ref duint && Coqlib.has_ref dec
+ && Coqlib.has_ref hint && Coqlib.has_ref huint && Coqlib.has_ref hex
+ && Coqlib.has_ref int && Coqlib.has_ref uint && Coqlib.has_ref num
then
+ let q_dint = qualid_of_ref dint in
+ let q_duint = qualid_of_ref duint in
+ let q_dec = qualid_of_ref dec in
+ let q_hint = qualid_of_ref hint in
+ let q_huint = qualid_of_ref huint in
+ let q_hex = qualid_of_ref hex in
let q_int = qualid_of_ref int in
let q_uint = qualid_of_ref uint in
- let q_dec = qualid_of_ref dec in
+ let q_num = qualid_of_ref num in
let int_ty = {
+ dec_int = unsafe_locate_ind q_dint;
+ dec_uint = unsafe_locate_ind q_duint;
+ hex_int = unsafe_locate_ind q_hint;
+ hex_uint = unsafe_locate_ind q_huint;
int = unsafe_locate_ind q_int;
uint = unsafe_locate_ind q_uint;
} in
- let dec_ty = {
+ let num_ty = {
int = int_ty;
decimal = unsafe_locate_ind q_dec;
+ hexadecimal = unsafe_locate_ind q_hex;
+ numeral = unsafe_locate_ind q_num;
} in
- Some (int_ty, mkRefC q_int, mkRefC q_uint, dec_ty, mkRefC q_dec)
+ Some (int_ty, mkRefC q_int, mkRefC q_uint, mkRefC q_dint, mkRefC q_duint,
+ num_ty, mkRefC q_num, mkRefC q_dec)
else None
let locate_int63 () =
@@ -90,20 +111,27 @@ let has_type env sigma f ty =
let type_error_to f ty =
CErrors.user_err
- (pr_qualid f ++ str " should go from Decimal.int to " ++
+ (pr_qualid f ++ str " should go from Numeral.int to " ++
pr_qualid ty ++ str " or (option " ++ pr_qualid ty ++ str ")." ++
- fnl () ++ str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).")
+ fnl () ++ str "Instead of Numeral.int, the types Numeral.uint or Z or Int63.int or Numeral.numeral could be used (you may need to require BinNums or Numeral or Int63 first).")
let type_error_of g ty =
CErrors.user_err
(pr_qualid g ++ str " should go from " ++ pr_qualid ty ++
- str " to Decimal.int or (option Decimal.int)." ++ fnl () ++
- str "Instead of Decimal.int, the types Decimal.uint or Z or Int63.int or Decimal.decimal could be used (you may need to require BinNums or Decimal or Int63 first).")
+ str " to Numeral.int or (option Numeral.int)." ++ fnl () ++
+ str "Instead of Numeral.int, the types Numeral.uint or Z or Int63.int or Numeral.numeral could be used (you may need to require BinNums or Numeral or Int63 first).")
+
+let warn_deprecated_decimal =
+ CWarnings.create ~name:"decimal-numeral-notation" ~category:"deprecated"
+ (fun () ->
+ strbrk "Deprecated Numeral Notation for Decimal.uint, \
+ Decimal.int or Decimal.decimal. Use Numeral.uint, \
+ Numeral.int or Numeral.numeral respectively.")
let vernac_numeral_notation local ty f g scope opts =
let env = Global.env () in
let sigma = Evd.from_env env in
- let dec_ty = locate_decimal () in
+ let num_ty = locate_numeral () in
let z_pos_ty = locate_z () in
let int63_ty = locate_int63 () in
let tyc = Smartlocate.global_inductive_with_alias ty in
@@ -118,13 +146,19 @@ let vernac_numeral_notation local ty f g scope opts =
let constructors = get_constructors tyc in
(* Check the type of f *)
let to_kind =
- match dec_ty with
- | Some (int_ty, cint, _, _, _) when has_type env sigma f (arrow cint cty) -> Int int_ty, Direct
- | Some (int_ty, cint, _, _, _) when has_type env sigma f (arrow cint (opt cty)) -> Int int_ty, Option
- | Some (int_ty, _, cuint, _, _) when has_type env sigma f (arrow cuint cty) -> UInt int_ty.uint, Direct
- | Some (int_ty, _, cuint, _, _) when has_type env sigma f (arrow cuint (opt cty)) -> UInt int_ty.uint, Option
- | Some (_, _, _, dec_ty, cdec) when has_type env sigma f (arrow cdec cty) -> Decimal dec_ty, Direct
- | Some (_, _, _, dec_ty, cdec) when has_type env sigma f (arrow cdec (opt cty)) -> Decimal dec_ty, Option
+ match num_ty with
+ | Some (int_ty, cint, _, _, _, _, _, _) when has_type env sigma f (arrow cint cty) -> Int int_ty, Direct
+ | Some (int_ty, cint, _, _, _, _, _, _) when has_type env sigma f (arrow cint (opt cty)) -> Int int_ty, Option
+ | Some (int_ty, _, cuint, _, _, _, _, _) when has_type env sigma f (arrow cuint cty) -> UInt int_ty, Direct
+ | Some (int_ty, _, cuint, _, _, _, _, _) when has_type env sigma f (arrow cuint (opt cty)) -> UInt int_ty, Option
+ | Some (_, _, _, _, _, num_ty, cnum, _) when has_type env sigma f (arrow cnum cty) -> Numeral num_ty, Direct
+ | Some (_, _, _, _, _, num_ty, cnum, _) when has_type env sigma f (arrow cnum (opt cty)) -> Numeral num_ty, Option
+ | Some (int_ty, _, _, cint, _, _, _, _) when has_type env sigma f (arrow cint cty) -> DecimalInt int_ty, Direct
+ | Some (int_ty, _, _, cint, _, _, _, _) when has_type env sigma f (arrow cint (opt cty)) -> DecimalInt int_ty, Option
+ | Some (int_ty, _, _, _, cuint, _, _, _) when has_type env sigma f (arrow cuint cty) -> DecimalUInt int_ty, Direct
+ | Some (int_ty, _, _, _, cuint, _, _, _) when has_type env sigma f (arrow cuint (opt cty)) -> DecimalUInt int_ty, Option
+ | Some (_, _, _, _, _, num_ty, _, cdec) when has_type env sigma f (arrow cdec cty) -> Decimal num_ty, Direct
+ | Some (_, _, _, _, _, num_ty, _, cdec) when has_type env sigma f (arrow cdec (opt cty)) -> Decimal num_ty, Option
| _ ->
match z_pos_ty with
| Some (z_pos_ty, cZ) when has_type env sigma f (arrow cZ cty) -> Z z_pos_ty, Direct
@@ -137,13 +171,19 @@ let vernac_numeral_notation local ty f g scope opts =
in
(* Check the type of g *)
let of_kind =
- match dec_ty with
- | Some (int_ty, cint, _, _, _) when has_type env sigma g (arrow cty cint) -> Int int_ty, Direct
- | Some (int_ty, cint, _, _, _) when has_type env sigma g (arrow cty (opt cint)) -> Int int_ty, Option
- | Some (int_ty, _, cuint, _, _) when has_type env sigma g (arrow cty cuint) -> UInt int_ty.uint, Direct
- | Some (int_ty, _, cuint, _, _) when has_type env sigma g (arrow cty (opt cuint)) -> UInt int_ty.uint, Option
- | Some (_, _, _, dec_ty, cdec) when has_type env sigma g (arrow cty cdec) -> Decimal dec_ty, Direct
- | Some (_, _, _, dec_ty, cdec) when has_type env sigma g (arrow cty (opt cdec)) -> Decimal dec_ty, Option
+ match num_ty with
+ | Some (int_ty, cint, _, _, _, _, _, _) when has_type env sigma g (arrow cty cint) -> Int int_ty, Direct
+ | Some (int_ty, cint, _, _, _, _, _, _) when has_type env sigma g (arrow cty (opt cint)) -> Int int_ty, Option
+ | Some (int_ty, _, cuint, _, _, _, _, _) when has_type env sigma g (arrow cty cuint) -> UInt int_ty, Direct
+ | Some (int_ty, _, cuint, _, _, _, _, _) when has_type env sigma g (arrow cty (opt cuint)) -> UInt int_ty, Option
+ | Some (_, _, _, _, _, num_ty, cnum, _) when has_type env sigma g (arrow cty cnum) -> Numeral num_ty, Direct
+ | Some (_, _, _, _, _, num_ty, cnum, _) when has_type env sigma g (arrow cty (opt cnum)) -> Numeral num_ty, Option
+ | Some (int_ty, _, _, cint, _, _, _, _) when has_type env sigma g (arrow cty cint) -> DecimalInt int_ty, Direct
+ | Some (int_ty, _, _, cint, _, _, _, _) when has_type env sigma g (arrow cty (opt cint)) -> DecimalInt int_ty, Option
+ | Some (int_ty, _, _, _, cuint, _, _, _) when has_type env sigma g (arrow cty cuint) -> DecimalUInt int_ty, Direct
+ | Some (int_ty, _, _, _, cuint, _, _, _) when has_type env sigma g (arrow cty (opt cuint)) -> DecimalUInt int_ty, Option
+ | Some (_, _, _, _, _, num_ty, _, cdec) when has_type env sigma g (arrow cty cdec) -> Decimal num_ty, Direct
+ | Some (_, _, _, _, _, num_ty, _, cdec) when has_type env sigma g (arrow cty (opt cdec)) -> Decimal num_ty, Option
| _ ->
match z_pos_ty with
| Some (z_pos_ty, cZ) when has_type env sigma g (arrow cty cZ) -> Z z_pos_ty, Direct
@@ -154,6 +194,11 @@ let vernac_numeral_notation local ty f g scope opts =
| Some cint63 when has_type env sigma g (arrow cty (opt cint63)) -> Int63, Option
| _ -> type_error_of g ty
in
+ (match to_kind, of_kind with
+ | ((DecimalInt _ | DecimalUInt _ | Decimal _), _), _
+ | _, ((DecimalInt _ | DecimalUInt _ | Decimal _), _) ->
+ warn_deprecated_decimal ()
+ | _ -> ());
let o = { to_kind; to_ty; of_kind; of_ty;
ty_name = ty;
warning = opts }
diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml
index c4e9c8b73d..23a7cc07c5 100644
--- a/plugins/syntax/r_syntax.ml
+++ b/plugins/syntax/r_syntax.ml
@@ -114,20 +114,21 @@ let glob_pow_pos = GlobRef.ConstRef (Constant.make2 z_modpath @@ Label.make "pow
let r_of_rawnum ?loc n =
let n,e = NumTok.Signed.to_bigint_and_exponent n in
+ let e,p = NumTok.(match e with EDec e -> e, 10 | EBin e -> e, 2) in
let izr z =
DAst.make @@ GApp (DAst.make @@ GRef(glob_IZR,None), [z]) in
let rmult r r' =
DAst.make @@ GApp (DAst.make @@ GRef(glob_Rmult,None), [r; r']) in
let rdiv r r' =
DAst.make @@ GApp (DAst.make @@ GRef(glob_Rdiv,None), [r; r']) in
- let pow10 e =
- let ten = z_of_int ?loc (Bigint.of_int 10) in
+ let pow p e =
+ let p = z_of_int ?loc (Bigint.of_int p) in
let e = pos_of_bignat e in
- DAst.make @@ GApp (DAst.make @@ GRef(glob_pow_pos,None), [ten; e]) in
+ DAst.make @@ GApp (DAst.make @@ GRef(glob_pow_pos,None), [p; e]) in
let n =
izr (z_of_int ?loc n) in
- if Bigint.is_strictly_pos e then rmult n (izr (pow10 e))
- else if Bigint.is_strictly_neg e then rdiv n (izr (pow10 (neg e)))
+ if Bigint.is_strictly_pos e then rmult n (izr (pow p e))
+ else if Bigint.is_strictly_neg e then rdiv n (izr (pow p (neg e)))
else n (* e = 0 *)
(**********************************************************************)
@@ -149,7 +150,8 @@ let rawnum_of_r c =
let le = String.length (Bigint.to_string e) in
Bigint.(less_than (add (of_int li) (of_int le)) e) in
(* print 123 * 10^-2 as 123e-2 *)
- let numTok_exponent () = NumTok.Signed.of_bigint_and_exponent i e in
+ let numTok_exponent () =
+ NumTok.Signed.of_bigint_and_exponent i (NumTok.EDec e) in
(* print 123 * 10^-2 as 1.23, precondition e < 0 *)
let numTok_dot () =
let s, i =
@@ -163,12 +165,12 @@ let rawnum_of_r c =
else "0", String.make (e - ni) '0' ^ i in
let i = s, NumTok.UnsignedNat.of_string i in
let f = NumTok.UnsignedNat.of_string f in
- NumTok.Signed.of_decimal_and_exponent i (Some f) None in
+ NumTok.Signed.of_int_frac_and_exponent i (Some f) None in
if choose_exponent then numTok_exponent () else numTok_dot () in
match DAst.get c with
| GApp (r, [a]) when is_gr r glob_IZR ->
let n = bigint_of_z a in
- NumTok.Signed.of_bigint n
+ NumTok.(Signed.of_bigint CDec n)
| GApp (md, [l; r]) when is_gr md glob_Rmult || is_gr md glob_Rdiv ->
begin match DAst.get l, DAst.get r with
| GApp (i, [l]), GApp (i', [r])
diff --git a/test-suite/output/FloatSyntax.out b/test-suite/output/FloatSyntax.out
index 7941d2e647..4a5a700879 100644
--- a/test-suite/output/FloatSyntax.out
+++ b/test-suite/output/FloatSyntax.out
@@ -42,13 +42,25 @@ closest value will be used and unambiguously printed
: float
2.5 + 2.5
: float
-File "stdin", line 24, characters 6-11:
+-26
+ : float
+11.171875
+ : float
+-6882
+ : float
+44.6875
+ : float
+2860
+ : float
+-2.79296875
+ : float
+File "stdin", line 30, characters 6-11:
Warning: The constant 1e309 is not a binary64 floating-point value. A closest
value will be used and unambiguously printed infinity.
[inexact-float,parsing]
infinity
: float
-File "stdin", line 25, characters 6-12:
+File "stdin", line 31, characters 6-12:
Warning: The constant -1e309 is not a binary64 floating-point value. A
closest value will be used and unambiguously printed neg_infinity.
[inexact-float,parsing]
diff --git a/test-suite/output/FloatSyntax.v b/test-suite/output/FloatSyntax.v
index eca712db10..36ffc27239 100644
--- a/test-suite/output/FloatSyntax.v
+++ b/test-suite/output/FloatSyntax.v
@@ -20,6 +20,12 @@ Check 2.5e123.
Check (-2.5e-123).
Check (2 + 2).
Check (2.5 + 2.5).
+Check -0x1a.
+Check 0xb.2c.
+Check -0x1ae2.
+Check 0xb.2cp2.
+Check 0xb.2cp8.
+Check -0xb.2cp-2.
Check 1e309.
Check -1e309.
diff --git a/test-suite/output/Int63Syntax.out b/test-suite/output/Int63Syntax.out
index 4d76f1210b..eefa338f0d 100644
--- a/test-suite/output/Int63Syntax.out
+++ b/test-suite/output/Int63Syntax.out
@@ -6,6 +6,32 @@
: int
9223372036854775807
: int
+427
+ : int
+427
+ : int
+427
+ : int
+427
+ : int
+427
+ : int
+The command has indeed failed with message:
+Cannot interpret this number as a value of type Coq.Numbers.Cyclic.Int63.Int63.int
+The command has indeed failed with message:
+Cannot interpret this number as a value of type Coq.Numbers.Cyclic.Int63.Int63.int
+0
+ : int
+0
+ : int
+The command has indeed failed with message:
+The reference xg was not found in the current environment.
+The command has indeed failed with message:
+The reference xG was not found in the current environment.
+The command has indeed failed with message:
+The reference x1 was not found in the current environment.
+The command has indeed failed with message:
+The reference x was not found in the current environment.
2 + 2
: int
2 + 2
diff --git a/test-suite/output/Int63Syntax.v b/test-suite/output/Int63Syntax.v
index 0385e529bf..c49616d918 100644
--- a/test-suite/output/Int63Syntax.v
+++ b/test-suite/output/Int63Syntax.v
@@ -5,6 +5,19 @@ Check (2 + 2)%int63.
Open Scope int63_scope.
Check 2.
Check 9223372036854775807.
+Check 0x1ab.
+Check 0X1ab.
+Check 0x1Ab.
+Check 0x1aB.
+Check 0x1AB.
+Fail Check 0x1ap5. (* exponents not implemented (yet?) *)
+Fail Check 0x1aP5.
+Check 0x0.
+Check 0x000.
+Fail Check 0xg.
+Fail Check 0xG.
+Fail Check 00x1.
+Fail Check 0x.
Check (Int63.add 2 2).
Check (2+2).
Eval vm_compute in 2+2.
diff --git a/test-suite/output/NatSyntax.out b/test-suite/output/NatSyntax.out
new file mode 100644
index 0000000000..3067896d7a
--- /dev/null
+++ b/test-suite/output/NatSyntax.out
@@ -0,0 +1,34 @@
+42
+ : nat
+0
+ : nat
+0
+ : nat
+427
+ : nat
+427
+ : nat
+427
+ : nat
+427
+ : nat
+427
+ : nat
+The command has indeed failed with message:
+Cannot interpret this number as a value of type nat
+The command has indeed failed with message:
+Cannot interpret this number as a value of type nat
+0
+ : nat
+0
+ : nat
+The command has indeed failed with message:
+The reference xg was not found in the current environment.
+The command has indeed failed with message:
+The reference xG was not found in the current environment.
+The command has indeed failed with message:
+The reference x1 was not found in the current environment.
+The command has indeed failed with message:
+The reference x was not found in the current environment.
+0x2a
+ : nat
diff --git a/test-suite/output/NatSyntax.v b/test-suite/output/NatSyntax.v
new file mode 100644
index 0000000000..66f697d412
--- /dev/null
+++ b/test-suite/output/NatSyntax.v
@@ -0,0 +1,18 @@
+Check 42.
+Check 0.
+Check 00.
+Check 0x1ab.
+Check 0X1ab.
+Check 0x1Ab.
+Check 0x1aB.
+Check 0x1AB.
+Fail Check 0x1ap1. (* exponents not implemented (yet?) *)
+Fail Check 0x1aP1.
+Check 0x0.
+Check 0x000.
+Fail Check 0xg.
+Fail Check 0xG.
+Fail Check 00x1.
+Fail Check 0x.
+Open Scope hex_nat_scope.
+Check 42.
diff --git a/test-suite/output/NumeralNotations.out b/test-suite/output/NumeralNotations.out
index 060877707b..34c31ff8a6 100644
--- a/test-suite/output/NumeralNotations.out
+++ b/test-suite/output/NumeralNotations.out
@@ -32,7 +32,7 @@ Warning: To avoid stack overflow, large numbers in punit are interpreted as
applications of pto_punits. [abstract-large-number,numbers]
The command has indeed failed with message:
In environment
-v := pto_punits (Decimal.D1 Decimal.Nil) : punit
+v := pto_punits (Numeral.UIntDec (Decimal.D1 Decimal.Nil)) : punit
The term "v" has type "punit@{Set}" while it is expected to have type
"punit@{u}".
S
@@ -61,7 +61,7 @@ The command has indeed failed with message:
In environment
v := 0 : nat
The term "v" has type "nat" while it is expected to have type "wuint".
- = {| unwrap := Decimal.D0 Decimal.Nil |}
+ = {| unwrap := Numeral.UIntDec (Decimal.D0 Decimal.Nil) |}
: wuint
let v := 0%wuint8' in v : wuint
: wuint
@@ -82,7 +82,7 @@ function (of_uint) targets an option type.
The command has indeed failed with message:
The 'abstract after' directive has no effect when the parsing function
(of_uint) targets an option type. [abstract-large-number-no-op,numbers]
-let v := of_uint (Decimal.D1 Decimal.Nil) in v : unit
+let v := of_uint (Numeral.UIntDec (Decimal.D1 Decimal.Nil)) in v : unit
: unit
let v := 0%test13 in v : unit
: unit
diff --git a/test-suite/output/NumeralNotations.v b/test-suite/output/NumeralNotations.v
index 47e1b127cb..ca8a14cce1 100644
--- a/test-suite/output/NumeralNotations.v
+++ b/test-suite/output/NumeralNotations.v
@@ -5,17 +5,17 @@ Declare Scope opaque_scope.
(* https://github.com/coq/coq/pull/8064#discussion_r202497516 *)
Module Test1.
Axiom hold : forall {A B C}, A -> B -> C.
- Definition opaque3 (x : Decimal.int) : Decimal.int := hold x (fix f (x : nat) : nat := match x with O => O | S n => S (f n) end).
- Numeral Notation Decimal.int opaque3 opaque3 : opaque_scope.
+ Definition opaque3 (x : Numeral.int) : Numeral.int := hold x (fix f (x : nat) : nat := match x with O => O | S n => S (f n) end).
+ Numeral Notation Numeral.int opaque3 opaque3 : opaque_scope.
Delimit Scope opaque_scope with opaque.
Fail Check 1%opaque.
End Test1.
(* https://github.com/coq/coq/pull/8064#discussion_r202497990 *)
Module Test2.
- Axiom opaque4 : option Decimal.int.
- Definition opaque6 (x : Decimal.int) : option Decimal.int := opaque4.
- Numeral Notation Decimal.int opaque6 opaque6 : opaque_scope.
+ Axiom opaque4 : option Numeral.int.
+ Definition opaque6 (x : Numeral.int) : option Numeral.int := opaque4.
+ Numeral Notation Numeral.int opaque6 opaque6 : opaque_scope.
Delimit Scope opaque_scope with opaque.
Open Scope opaque_scope.
Fail Check 1%opaque.
@@ -24,8 +24,8 @@ End Test2.
Declare Scope silly_scope.
Module Test3.
- Inductive silly := SILLY (v : Decimal.uint) (f : forall A, A -> A).
- Definition to_silly (v : Decimal.uint) := SILLY v (fun _ x => x).
+ Inductive silly := SILLY (v : Numeral.uint) (f : forall A, A -> A).
+ Definition to_silly (v : Numeral.uint) := SILLY v (fun _ x => x).
Definition of_silly (v : silly) := match v with SILLY v _ => v end.
Numeral Notation silly to_silly of_silly : silly_scope.
Delimit Scope silly_scope with silly.
@@ -45,15 +45,15 @@ Module Test4.
Declare Scope upp.
Declare Scope ppps.
Polymorphic NonCumulative Inductive punit := ptt.
- Polymorphic Definition pto_punit (v : Decimal.uint) : option punit := match Nat.of_uint v with O => Some ptt | _ => None end.
- Polymorphic Definition pto_punit_all (v : Decimal.uint) : punit := ptt.
- Polymorphic Definition pof_punit (v : punit) : Decimal.uint := Nat.to_uint 0.
- Definition to_punit (v : Decimal.uint) : option punit := match Nat.of_uint v with O => Some ptt | _ => None end.
- Definition of_punit (v : punit) : Decimal.uint := Nat.to_uint 0.
- Polymorphic Definition pto_unit (v : Decimal.uint) : option unit := match Nat.of_uint v with O => Some tt | _ => None end.
- Polymorphic Definition pof_unit (v : unit) : Decimal.uint := Nat.to_uint 0.
- Definition to_unit (v : Decimal.uint) : option unit := match Nat.of_uint v with O => Some tt | _ => None end.
- Definition of_unit (v : unit) : Decimal.uint := Nat.to_uint 0.
+ Polymorphic Definition pto_punit (v : Numeral.uint) : option punit := match Nat.of_num_uint v with O => Some ptt | _ => None end.
+ Polymorphic Definition pto_punit_all (v : Numeral.uint) : punit := ptt.
+ Polymorphic Definition pof_punit (v : punit) : Numeral.uint := Nat.to_num_uint 0.
+ Definition to_punit (v : Numeral.uint) : option punit := match Nat.of_num_uint v with O => Some ptt | _ => None end.
+ Definition of_punit (v : punit) : Numeral.uint := Nat.to_num_uint 0.
+ Polymorphic Definition pto_unit (v : Numeral.uint) : option unit := match Nat.of_num_uint v with O => Some tt | _ => None end.
+ Polymorphic Definition pof_unit (v : unit) : Numeral.uint := Nat.to_num_uint 0.
+ Definition to_unit (v : Numeral.uint) : option unit := match Nat.of_num_uint v with O => Some tt | _ => None end.
+ Definition of_unit (v : unit) : Numeral.uint := Nat.to_num_uint 0.
Numeral Notation punit to_punit of_punit : pto.
Numeral Notation punit pto_punit of_punit : ppo.
Numeral Notation punit to_punit pof_punit : ptp.
@@ -113,8 +113,8 @@ Module Test6.
Local Set Primitive Projections.
Record > wnat := wrap { unwrap :> nat }.
- Definition to_uint (x : wnat) : Decimal.uint := Nat.to_uint x.
- Definition of_uint (x : Decimal.uint) : wnat := Nat.of_uint x.
+ Definition to_uint (x : wnat) : Numeral.uint := Nat.to_num_uint x.
+ Definition of_uint (x : Numeral.uint) : wnat := Nat.of_num_uint x.
Module Export Scopes.
Declare Scope wnat_scope.
Delimit Scope wnat_scope with wnat.
@@ -138,7 +138,7 @@ End Test6_2.
Module Test7.
Local Set Primitive Projections.
- Record wuint := wrap { unwrap : Decimal.uint }.
+ Record wuint := wrap { unwrap : Numeral.uint }.
Declare Scope wuint_scope.
Delimit Scope wuint_scope with wuint.
Numeral Notation wuint wrap unwrap : wuint_scope.
@@ -148,7 +148,7 @@ End Test7.
Module Test8.
Local Set Primitive Projections.
- Record wuint := wrap { unwrap : Decimal.uint }.
+ Record wuint := wrap { unwrap : Numeral.uint }.
Declare Scope wuint8_scope.
Declare Scope wuint8'_scope.
Delimit Scope wuint8_scope with wuint8.
@@ -162,7 +162,7 @@ Module Test8.
End with_var.
Check let v := 0%wuint8 in v : nat.
Fail Check let v := 0%wuint8 in v : wuint.
- Compute wrap (Nat.to_uint 0).
+ Compute wrap (Nat.to_num_uint 0).
Notation wrap'' := wrap.
Notation unwrap'' := unwrap.
@@ -177,7 +177,7 @@ Module Test9.
Delimit Scope wuint9'_scope with wuint9'.
Section with_let.
Local Set Primitive Projections.
- Record wuint := wrap { unwrap : Decimal.uint }.
+ Record wuint := wrap { unwrap : Numeral.uint }.
Let wrap' := wrap.
Let unwrap' := unwrap.
Local Notation wrap'' := wrap.
@@ -193,9 +193,9 @@ End Test9.
Module Test10.
(* Test that it is only a warning to add abstract after to an optional parsing function *)
- Definition to_uint (v : unit) := Nat.to_uint 0.
- Definition of_uint (v : Decimal.uint) := match Nat.of_uint v with O => Some tt | _ => None end.
- Definition of_any_uint (v : Decimal.uint) := tt.
+ Definition to_uint (v : unit) := Nat.to_num_uint 0.
+ Definition of_uint (v : Numeral.uint) := match Nat.of_num_uint v with O => Some tt | _ => None end.
+ Definition of_any_uint (v : Numeral.uint) := tt.
Declare Scope unit_scope.
Declare Scope unit2_scope.
Delimit Scope unit_scope with unit.
@@ -213,7 +213,7 @@ Module Test12.
Declare Scope test12_scope.
Delimit Scope test12_scope with test12.
Section test12.
- Context (to_uint : unit -> Decimal.uint) (of_uint : Decimal.uint -> unit).
+ Context (to_uint : unit -> Numeral.uint) (of_uint : Numeral.uint -> unit).
Numeral Notation unit of_uint to_uint : test12_scope.
Check let v := 1%test12 in v : unit.
@@ -228,8 +228,8 @@ Module Test13.
Delimit Scope test13_scope with test13.
Delimit Scope test13'_scope with test13'.
Delimit Scope test13''_scope with test13''.
- Definition to_uint (x y : unit) : Decimal.uint := Nat.to_uint O.
- Definition of_uint (x : Decimal.uint) : unit := tt.
+ Definition to_uint (x y : unit) : Numeral.uint := Nat.to_num_uint O.
+ Definition of_uint (x : Numeral.uint) : unit := tt.
Definition to_uint_good := to_uint tt.
Notation to_uint' := (to_uint tt).
Notation to_uint'' := (to_uint _).
@@ -254,8 +254,8 @@ Module Test14.
Delimit Scope test14''_scope with test14''.
Delimit Scope test14'''_scope with test14'''.
Module Inner.
- Definition to_uint (x : unit) : Decimal.uint := Nat.to_uint O.
- Definition of_uint (x : Decimal.uint) : unit := tt.
+ Definition to_uint (x : unit) : Numeral.uint := Nat.to_num_uint O.
+ Definition of_uint (x : Numeral.uint) : unit := tt.
Local Numeral Notation unit of_uint to_uint : test14_scope.
Global Numeral Notation unit of_uint to_uint : test14'_scope.
Check let v := 0%test14 in v : unit.
@@ -267,8 +267,8 @@ Module Test14.
Fail Check let v := 0%test14 in v : unit.
Check let v := 0%test14' in v : unit.
Section InnerSection.
- Definition to_uint (x : unit) : Decimal.uint := Nat.to_uint O.
- Definition of_uint (x : Decimal.uint) : unit := tt.
+ Definition to_uint (x : unit) : Numeral.uint := Nat.to_num_uint O.
+ Definition of_uint (x : Numeral.uint) : unit := tt.
Local Numeral Notation unit of_uint to_uint : test14''_scope.
Fail Global Numeral Notation unit of_uint to_uint : test14'''_scope.
Check let v := 0%test14'' in v : unit.
@@ -283,8 +283,8 @@ Module Test15.
Declare Scope test15_scope.
Delimit Scope test15_scope with test15.
Module Inner.
- Definition to_uint (x : unit) : Decimal.uint := Nat.to_uint O.
- Definition of_uint (x : Decimal.uint) : unit := tt.
+ Definition to_uint (x : unit) : Numeral.uint := Nat.to_num_uint O.
+ Definition of_uint (x : Numeral.uint) : unit := tt.
Numeral Notation unit of_uint to_uint : test15_scope.
Check let v := 0%test15 in v : unit.
End Inner.
@@ -306,8 +306,8 @@ Module Test16.
End A.
Module F (a : A).
Inductive Foo := foo (_ : a.T).
- Definition to_uint (x : Foo) : Decimal.uint := Nat.to_uint O.
- Definition of_uint (x : Decimal.uint) : Foo := foo a.t.
+ Definition to_uint (x : Foo) : Numeral.uint := Nat.to_num_uint O.
+ Definition of_uint (x : Numeral.uint) : Foo := foo a.t.
Global Numeral Notation Foo of_uint to_uint : test16_scope.
Check let v := 0%test16 in v : Foo.
End F.
@@ -352,9 +352,9 @@ Module Test18.
Definition Q_of_nat (x : nat) : Q := {| num := x ; den := 1 ; reduced := transparentify (nat_eq_dec _ _) (gcd_good _) |}.
Definition nat_of_Q (x : Q) : option nat
:= if Nat.eqb x.(den) 1 then Some (x.(num)) else None.
- Definition Q_of_uint (x : Decimal.uint) : Q := Q_of_nat (Nat.of_uint x).
- Definition uint_of_Q (x : Q) : option Decimal.uint
- := option_map Nat.to_uint (nat_of_Q x).
+ Definition Q_of_uint (x : Numeral.uint) : Q := Q_of_nat (Nat.of_num_uint x).
+ Definition uint_of_Q (x : Q) : option Numeral.uint
+ := option_map Nat.to_num_uint (nat_of_Q x).
Numeral Notation Q Q_of_uint uint_of_Q : Q_scope.
@@ -411,8 +411,8 @@ Module Test20.
Record > ty := { t : Type ; kt : known_type t }.
- Definition ty_of_uint (x : Decimal.uint) : option ty
- := match Nat.of_uint x with
+ Definition ty_of_uint (x : Numeral.uint) : option ty
+ := match Nat.of_num_uint x with
| 0 => @Some ty zero
| 1 => @Some ty one
| 2 => @Some ty two
@@ -421,8 +421,8 @@ Module Test20.
| 5 => @Some ty type
| _ => None
end.
- Definition uint_of_ty (x : ty) : Decimal.uint
- := Nat.to_uint match kt x with
+ Definition uint_of_ty (x : ty) : Numeral.uint
+ := Nat.to_num_uint match kt x with
| prop => 3
| set => 4
| type => 5
diff --git a/test-suite/output/QArithSyntax.out b/test-suite/output/QArithSyntax.out
index fe6a1d25c6..9b5c076cb4 100644
--- a/test-suite/output/QArithSyntax.out
+++ b/test-suite/output/QArithSyntax.out
@@ -12,3 +12,15 @@ eq_refl : -1e-4 = -1e-4
: -1e-4 = -1e-4
eq_refl : -0.50 = -0.50
: -0.50 = -0.50
+eq_refl : -26 = -26
+ : -26 = -26
+eq_refl : 2860 # 256 = 2860 # 256
+ : 2860 # 256 = 2860 # 256
+eq_refl : -6882 = -6882
+ : -6882 = -6882
+eq_refl : 2860 # 64 = 2860 # 64
+ : 2860 # 64 = 2860 # 64
+eq_refl : 2860 = 2860
+ : 2860 = 2860
+eq_refl : -2860 # 1024 = -2860 # 1024
+ : -2860 # 1024 = -2860 # 1024
diff --git a/test-suite/output/QArithSyntax.v b/test-suite/output/QArithSyntax.v
index 2f2ee0134a..b5c6222bba 100644
--- a/test-suite/output/QArithSyntax.v
+++ b/test-suite/output/QArithSyntax.v
@@ -7,3 +7,9 @@ Check (eq_refl : 1.02e+02 = 102 # 1).
Check (eq_refl : 10.2e-1 = 1.02).
Check (eq_refl : -0.0001 = -1 # 10000).
Check (eq_refl : -0.50 = - 50 # 100).
+Check (eq_refl : -0x1a = - 26 # 1).
+Check (eq_refl : 0xb.2c = 2860 # 256).
+Check (eq_refl : -0x1ae2 = -6882).
+Check (eq_refl : 0xb.2cp2 = 2860 # 64).
+Check (eq_refl : 0xb.2cp8 = 2860).
+Check (eq_refl : -0xb.2cp-2 = -2860 # 1024).
diff --git a/test-suite/output/RealSyntax.out b/test-suite/output/RealSyntax.out
index 1685964b0f..a9386b2781 100644
--- a/test-suite/output/RealSyntax.out
+++ b/test-suite/output/RealSyntax.out
@@ -20,3 +20,18 @@ eq_refl : -1e-4 = -1e-4
: -1e-4 = -1e-4
eq_refl : -0.50 = -0.50
: -0.50 = -0.50
+eq_refl : -26 = -26
+ : -26 = -26
+eq_refl : 2860 / IZR (BinIntDef.Z.pow_pos 2 8) = 2860 / IZR (Z.pow_pos 2 8)
+ : 2860 / IZR (BinIntDef.Z.pow_pos 2 8) = 2860 / IZR (Z.pow_pos 2 8)
+eq_refl : -6882 = -6882
+ : -6882 = -6882
+eq_refl : 2860 / IZR (BinIntDef.Z.pow_pos 2 6) = 2860 / IZR (Z.pow_pos 2 6)
+ : 2860 / IZR (BinIntDef.Z.pow_pos 2 6) = 2860 / IZR (Z.pow_pos 2 6)
+eq_refl : 2860 = 2860
+ : 2860 = 2860
+eq_refl
+:
+-2860 / IZR (BinIntDef.Z.pow_pos 2 10) = - (2860) / IZR (Z.pow_pos 2 10)
+ : -2860 / IZR (BinIntDef.Z.pow_pos 2 10) =
+ - (2860) / IZR (Z.pow_pos 2 10)
diff --git a/test-suite/output/RealSyntax.v b/test-suite/output/RealSyntax.v
index e5f9d06316..69ce3ef5f9 100644
--- a/test-suite/output/RealSyntax.v
+++ b/test-suite/output/RealSyntax.v
@@ -14,6 +14,12 @@ Check (eq_refl : 1.02e+02 = IZR 102).
Check (eq_refl : 10.2e-1 = 1.02).
Check (eq_refl : -0.0001 = IZR (-1) / IZR (Z.pow_pos 10 4)).
Check (eq_refl : -0.50 = IZR (-50) / IZR (Z.pow_pos 10 2)).
+Check (eq_refl : -0x1a = - 26).
+Check (eq_refl : 0xb.2c = IZR 2860 / IZR (Z.pow_pos 2 8)).
+Check (eq_refl : -0x1ae2 = -6882).
+Check (eq_refl : 0xb.2cp2 = IZR 2860 / IZR (Z.pow_pos 2 6)).
+Check (eq_refl : 0xb.2cp8 = 2860).
+Check (eq_refl : -0xb.2cp-2 = - IZR 2860 / IZR (Z.pow_pos 2 10)).
Require Import Reals.
diff --git a/test-suite/output/Search.out b/test-suite/output/Search.out
index 593d0c7f67..c01f4b2e19 100644
--- a/test-suite/output/Search.out
+++ b/test-suite/output/Search.out
@@ -21,57 +21,113 @@ le_sind:
(use "About" for full details on implicit arguments)
false: bool
true: bool
-eq_true: bool -> Prop
is_true: bool -> Prop
+eq_true: bool -> Prop
negb: bool -> bool
+xorb: bool -> bool -> bool
andb: bool -> bool -> bool
orb: bool -> bool -> bool
implb: bool -> bool -> bool
-xorb: bool -> bool -> bool
-Nat.even: nat -> bool
Nat.odd: nat -> bool
+Nat.even: nat -> bool
BoolSpec: Prop -> Prop -> bool -> Prop
-Nat.ltb: nat -> nat -> bool
-Nat.testbit: nat -> nat -> bool
+Numeral.numeral_beq: Numeral.numeral -> Numeral.numeral -> bool
Nat.eqb: nat -> nat -> bool
+Nat.testbit: nat -> nat -> bool
+Decimal.decimal_beq: Decimal.decimal -> Decimal.decimal -> bool
+Numeral.uint_beq: Numeral.uint -> Numeral.uint -> bool
+Decimal.uint_beq: Decimal.uint -> Decimal.uint -> bool
+Hexadecimal.hexadecimal_beq:
+ Hexadecimal.hexadecimal -> Hexadecimal.hexadecimal -> bool
+Numeral.int_beq: Numeral.int -> Numeral.int -> bool
+Hexadecimal.uint_beq: Hexadecimal.uint -> Hexadecimal.uint -> bool
+Nat.ltb: nat -> nat -> bool
Nat.leb: nat -> nat -> bool
+Hexadecimal.int_beq: Hexadecimal.int -> Hexadecimal.int -> bool
+Decimal.int_beq: Decimal.int -> Decimal.int -> bool
Nat.bitwise: (bool -> bool -> bool) -> nat -> nat -> nat -> nat
+eq_true_rec_r:
+ forall (P : bool -> Set) [b : bool], P b -> eq_true b -> P true
+eq_true_ind:
+ forall P : bool -> Prop, P true -> forall b : bool, eq_true b -> P b
bool_rec: forall P : bool -> Set, P true -> P false -> forall b : bool, P b
+eq_true_rect_r:
+ forall (P : bool -> Type) [b : bool], P b -> eq_true b -> P true
+bool_rect: forall P : bool -> Type, P true -> P false -> forall b : bool, P b
eq_true_ind_r:
forall (P : bool -> Prop) [b : bool], P b -> eq_true b -> P true
+eq_true_rect:
+ forall P : bool -> Type, P true -> forall b : bool, eq_true b -> P b
+bool_sind:
+ forall P : bool -> SProp, P true -> P false -> forall b : bool, P b
eq_true_rec:
forall P : bool -> Set, P true -> forall b : bool, eq_true b -> P b
bool_ind: forall P : bool -> Prop, P true -> P false -> forall b : bool, P b
-eq_true_rect:
- forall P : bool -> Type, P true -> forall b : bool, eq_true b -> P b
eq_true_sind:
forall P : bool -> SProp, P true -> forall b : bool, eq_true b -> P b
-bool_rect: forall P : bool -> Type, P true -> P false -> forall b : bool, P b
-eq_true_ind:
- forall P : bool -> Prop, P true -> forall b : bool, eq_true b -> P b
-eq_true_rec_r:
- forall (P : bool -> Set) [b : bool], P b -> eq_true b -> P true
-eq_true_rect_r:
- forall (P : bool -> Type) [b : bool], P b -> eq_true b -> P true
-bool_sind:
- forall P : bool -> SProp, P true -> P false -> forall b : bool, P b
-Byte.to_bits:
- Byte.byte ->
- bool * (bool * (bool * (bool * (bool * (bool * (bool * bool))))))
+Numeral.internal_uint_dec_bl1:
+ forall x y : Numeral.uint, Numeral.uint_beq x y = true -> x = y
+Hexadecimal.internal_hexadecimal_dec_lb:
+ forall x y : Hexadecimal.hexadecimal,
+ x = y -> Hexadecimal.hexadecimal_beq x y = true
+Hexadecimal.internal_int_dec_lb0:
+ forall x y : Hexadecimal.int, x = y -> Hexadecimal.int_beq x y = true
+Numeral.internal_numeral_dec_lb:
+ forall x y : Numeral.numeral, x = y -> Numeral.numeral_beq x y = true
+Decimal.internal_decimal_dec_lb:
+ forall x y : Decimal.decimal, x = y -> Decimal.decimal_beq x y = true
+Hexadecimal.internal_int_dec_bl0:
+ forall x y : Hexadecimal.int, Hexadecimal.int_beq x y = true -> x = y
+Numeral.internal_int_dec_lb1:
+ forall x y : Numeral.int, x = y -> Numeral.int_beq x y = true
+Numeral.internal_int_dec_bl1:
+ forall x y : Numeral.int, Numeral.int_beq x y = true -> x = y
+Hexadecimal.internal_hexadecimal_dec_bl:
+ forall x y : Hexadecimal.hexadecimal,
+ Hexadecimal.hexadecimal_beq x y = true -> x = y
+Numeral.internal_uint_dec_lb1:
+ forall x y : Numeral.uint, x = y -> Numeral.uint_beq x y = true
+Decimal.internal_int_dec_bl:
+ forall x y : Decimal.int, Decimal.int_beq x y = true -> x = y
+Decimal.internal_int_dec_lb:
+ forall x y : Decimal.int, x = y -> Decimal.int_beq x y = true
+Numeral.internal_numeral_dec_bl:
+ forall x y : Numeral.numeral, Numeral.numeral_beq x y = true -> x = y
Byte.of_bits:
bool * (bool * (bool * (bool * (bool * (bool * (bool * bool)))))) ->
Byte.byte
+Byte.to_bits:
+ Byte.byte ->
+ bool * (bool * (bool * (bool * (bool * (bool * (bool * bool))))))
+Decimal.internal_decimal_dec_bl:
+ forall x y : Decimal.decimal, Decimal.decimal_beq x y = true -> x = y
+Hexadecimal.internal_uint_dec_bl0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, Hexadecimal.uint_beq x0 y = true -> x0 = y) x
+Decimal.internal_uint_dec_bl:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, Decimal.uint_beq x0 y = true -> x0 = y) x
+Decimal.internal_uint_dec_lb:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, x0 = y -> Decimal.uint_beq x0 y = true) x
+Hexadecimal.internal_uint_dec_lb0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, x0 = y -> Hexadecimal.uint_beq x0 y = true) x
andb_prop: forall a b : bool, (a && b)%bool = true -> a = true /\ b = true
andb_true_intro:
forall [b1 b2 : bool], b1 = true /\ b2 = true -> (b1 && b2)%bool = true
-BoolSpec_sind:
- forall [P Q : Prop] (P0 : bool -> SProp),
- (P -> P0 true) ->
- (Q -> P0 false) -> forall b : bool, BoolSpec P Q b -> P0 b
BoolSpec_ind:
forall [P Q : Prop] (P0 : bool -> Prop),
(P -> P0 true) ->
(Q -> P0 false) -> forall b : bool, BoolSpec P Q b -> P0 b
+BoolSpec_sind:
+ forall [P Q : Prop] (P0 : bool -> SProp),
+ (P -> P0 true) ->
+ (Q -> P0 false) -> forall b : bool, BoolSpec P Q b -> P0 b
Byte.to_bits_of_bits:
forall
b : bool * (bool * (bool * (bool * (bool * (bool * (bool * bool)))))),
@@ -107,14 +163,106 @@ f_equal2_nat:
forall (B : Type) (f : nat -> nat -> B) (x1 y1 x2 y2 : nat),
x1 = y1 -> x2 = y2 -> f x1 x2 = f y1 y2
(use "About" for full details on implicit arguments)
+Numeral.internal_numeral_dec_lb:
+ forall x y : Numeral.numeral, x = y -> Numeral.numeral_beq x y = true
+Numeral.internal_int_dec_lb1:
+ forall x y : Numeral.int, x = y -> Numeral.int_beq x y = true
+Numeral.internal_numeral_dec_bl:
+ forall x y : Numeral.numeral, Numeral.numeral_beq x y = true -> x = y
+Hexadecimal.internal_hexadecimal_dec_lb:
+ forall x y : Hexadecimal.hexadecimal,
+ x = y -> Hexadecimal.hexadecimal_beq x y = true
+Numeral.internal_int_dec_bl1:
+ forall x y : Numeral.int, Numeral.int_beq x y = true -> x = y
+Numeral.internal_uint_dec_lb1:
+ forall x y : Numeral.uint, x = y -> Numeral.uint_beq x y = true
+Numeral.internal_uint_dec_bl1:
+ forall x y : Numeral.uint, Numeral.uint_beq x y = true -> x = y
+Decimal.internal_decimal_dec_lb:
+ forall x y : Decimal.decimal, x = y -> Decimal.decimal_beq x y = true
+Hexadecimal.internal_hexadecimal_dec_bl:
+ forall x y : Hexadecimal.hexadecimal,
+ Hexadecimal.hexadecimal_beq x y = true -> x = y
+Hexadecimal.internal_int_dec_lb0:
+ forall x y : Hexadecimal.int, x = y -> Hexadecimal.int_beq x y = true
+Hexadecimal.internal_int_dec_bl0:
+ forall x y : Hexadecimal.int, Hexadecimal.int_beq x y = true -> x = y
+Decimal.internal_int_dec_lb:
+ forall x y : Decimal.int, x = y -> Decimal.int_beq x y = true
+Decimal.internal_decimal_dec_bl:
+ forall x y : Decimal.decimal, Decimal.decimal_beq x y = true -> x = y
+Decimal.internal_int_dec_bl:
+ forall x y : Decimal.int, Decimal.int_beq x y = true -> x = y
+Decimal.internal_uint_dec_bl:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, Decimal.uint_beq x0 y = true -> x0 = y) x
+Decimal.internal_uint_dec_lb:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, x0 = y -> Decimal.uint_beq x0 y = true) x
+Hexadecimal.internal_uint_dec_lb0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, x0 = y -> Hexadecimal.uint_beq x0 y = true) x
+Hexadecimal.internal_uint_dec_bl0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, Hexadecimal.uint_beq x0 y = true -> x0 = y) x
+andb_prop: forall a b : bool, (a && b)%bool = true -> a = true /\ b = true
andb_true_intro:
forall [b1 b2 : bool], b1 = true /\ b2 = true -> (b1 && b2)%bool = true
-andb_prop: forall a b : bool, (a && b)%bool = true -> a = true /\ b = true
bool_choice:
forall [S : Set] [R1 R2 : S -> Prop],
(forall x : S, {R1 x} + {R2 x}) ->
{f : S -> bool | forall x : S, f x = true /\ R1 x \/ f x = false /\ R2 x}
(use "About" for full details on implicit arguments)
+Numeral.internal_numeral_dec_lb:
+ forall x y : Numeral.numeral, x = y -> Numeral.numeral_beq x y = true
+Numeral.internal_numeral_dec_bl:
+ forall x y : Numeral.numeral, Numeral.numeral_beq x y = true -> x = y
+Numeral.internal_int_dec_lb1:
+ forall x y : Numeral.int, x = y -> Numeral.int_beq x y = true
+Numeral.internal_int_dec_bl1:
+ forall x y : Numeral.int, Numeral.int_beq x y = true -> x = y
+Numeral.internal_uint_dec_lb1:
+ forall x y : Numeral.uint, x = y -> Numeral.uint_beq x y = true
+Numeral.internal_uint_dec_bl1:
+ forall x y : Numeral.uint, Numeral.uint_beq x y = true -> x = y
+Hexadecimal.internal_hexadecimal_dec_lb:
+ forall x y : Hexadecimal.hexadecimal,
+ x = y -> Hexadecimal.hexadecimal_beq x y = true
+Hexadecimal.internal_hexadecimal_dec_bl:
+ forall x y : Hexadecimal.hexadecimal,
+ Hexadecimal.hexadecimal_beq x y = true -> x = y
+Hexadecimal.internal_int_dec_lb0:
+ forall x y : Hexadecimal.int, x = y -> Hexadecimal.int_beq x y = true
+Hexadecimal.internal_int_dec_bl0:
+ forall x y : Hexadecimal.int, Hexadecimal.int_beq x y = true -> x = y
+Decimal.internal_decimal_dec_lb:
+ forall x y : Decimal.decimal, x = y -> Decimal.decimal_beq x y = true
+Decimal.internal_decimal_dec_bl:
+ forall x y : Decimal.decimal, Decimal.decimal_beq x y = true -> x = y
+Decimal.internal_int_dec_lb:
+ forall x y : Decimal.int, x = y -> Decimal.int_beq x y = true
+Decimal.internal_int_dec_bl:
+ forall x y : Decimal.int, Decimal.int_beq x y = true -> x = y
+Hexadecimal.internal_uint_dec_bl0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, Hexadecimal.uint_beq x0 y = true -> x0 = y) x
+Decimal.internal_uint_dec_bl:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, Decimal.uint_beq x0 y = true -> x0 = y) x
+Decimal.internal_uint_dec_lb:
+ forall x : Decimal.uint,
+ (fun x0 : Decimal.uint =>
+ forall y : Decimal.uint, x0 = y -> Decimal.uint_beq x0 y = true) x
+Hexadecimal.internal_uint_dec_lb0:
+ forall x : Hexadecimal.uint,
+ (fun x0 : Hexadecimal.uint =>
+ forall y : Hexadecimal.uint, x0 = y -> Hexadecimal.uint_beq x0 y = true) x
andb_true_intro:
forall [b1 b2 : bool], b1 = true /\ b2 = true -> (b1 && b2)%bool = true
andb_prop: forall a b : bool, (a && b)%bool = true -> a = true /\ b = true
diff --git a/test-suite/output/SearchHead.out b/test-suite/output/SearchHead.out
index 5627e4bd3c..d42dc575c2 100644
--- a/test-suite/output/SearchHead.out
+++ b/test-suite/output/SearchHead.out
@@ -8,16 +8,26 @@ le_S_n: forall n m : nat, S n <= S m -> n <= m
false: bool
true: bool
negb: bool -> bool
-implb: bool -> bool -> bool
-orb: bool -> bool -> bool
-andb: bool -> bool -> bool
xorb: bool -> bool -> bool
-Nat.even: nat -> bool
+andb: bool -> bool -> bool
+orb: bool -> bool -> bool
+implb: bool -> bool -> bool
Nat.odd: nat -> bool
-Nat.leb: nat -> nat -> bool
-Nat.ltb: nat -> nat -> bool
+Nat.even: nat -> bool
+Numeral.uint_beq: Numeral.uint -> Numeral.uint -> bool
Nat.testbit: nat -> nat -> bool
Nat.eqb: nat -> nat -> bool
+Hexadecimal.hexadecimal_beq:
+ Hexadecimal.hexadecimal -> Hexadecimal.hexadecimal -> bool
+Nat.ltb: nat -> nat -> bool
+Nat.leb: nat -> nat -> bool
+Numeral.numeral_beq: Numeral.numeral -> Numeral.numeral -> bool
+Numeral.int_beq: Numeral.int -> Numeral.int -> bool
+Hexadecimal.int_beq: Hexadecimal.int -> Hexadecimal.int -> bool
+Hexadecimal.uint_beq: Hexadecimal.uint -> Hexadecimal.uint -> bool
+Decimal.decimal_beq: Decimal.decimal -> Decimal.decimal -> bool
+Decimal.int_beq: Decimal.int -> Decimal.int -> bool
+Decimal.uint_beq: Decimal.uint -> Decimal.uint -> bool
(use "About" for full details on implicit arguments)
mult_n_O: forall n : nat, 0 = n * 0
plus_O_n: forall n : nat, 0 + n = n
diff --git a/test-suite/output/SearchPattern.out b/test-suite/output/SearchPattern.out
index 36fc1a5914..13d0a9e55b 100644
--- a/test-suite/output/SearchPattern.out
+++ b/test-suite/output/SearchPattern.out
@@ -1,47 +1,61 @@
false: bool
true: bool
negb: bool -> bool
-implb: bool -> bool -> bool
-orb: bool -> bool -> bool
-andb: bool -> bool -> bool
xorb: bool -> bool -> bool
-Nat.even: nat -> bool
+andb: bool -> bool -> bool
+orb: bool -> bool -> bool
+implb: bool -> bool -> bool
Nat.odd: nat -> bool
-Nat.leb: nat -> nat -> bool
-Nat.ltb: nat -> nat -> bool
+Nat.even: nat -> bool
+Numeral.uint_beq: Numeral.uint -> Numeral.uint -> bool
Nat.testbit: nat -> nat -> bool
Nat.eqb: nat -> nat -> bool
+Hexadecimal.hexadecimal_beq:
+ Hexadecimal.hexadecimal -> Hexadecimal.hexadecimal -> bool
+Nat.ltb: nat -> nat -> bool
+Nat.leb: nat -> nat -> bool
+Numeral.numeral_beq: Numeral.numeral -> Numeral.numeral -> bool
+Numeral.int_beq: Numeral.int -> Numeral.int -> bool
+Hexadecimal.int_beq: Hexadecimal.int -> Hexadecimal.int -> bool
+Hexadecimal.uint_beq: Hexadecimal.uint -> Hexadecimal.uint -> bool
+Decimal.decimal_beq: Decimal.decimal -> Decimal.decimal -> bool
+Decimal.int_beq: Decimal.int -> Decimal.int -> bool
+Decimal.uint_beq: Decimal.uint -> Decimal.uint -> bool
(use "About" for full details on implicit arguments)
Nat.two: nat
-Nat.one: nat
Nat.zero: nat
+Nat.one: nat
O: nat
+Nat.double: nat -> nat
+Nat.sqrt: nat -> nat
Nat.div2: nat -> nat
Nat.log2: nat -> nat
-Nat.succ: nat -> nat
-Nat.sqrt: nat -> nat
-S: nat -> nat
Nat.pred: nat -> nat
-Nat.double: nat -> nat
Nat.square: nat -> nat
+S: nat -> nat
+Nat.succ: nat -> nat
+Nat.ldiff: nat -> nat -> nat
+Nat.add: nat -> nat -> nat
Nat.land: nat -> nat -> nat
-Nat.lor: nat -> nat -> nat
+Nat.lxor: nat -> nat -> nat
+Nat.sub: nat -> nat -> nat
Nat.mul: nat -> nat -> nat
Nat.tail_mul: nat -> nat -> nat
-Nat.div: nat -> nat -> nat
-Nat.tail_add: nat -> nat -> nat
-Nat.gcd: nat -> nat -> nat
-Nat.modulo: nat -> nat -> nat
Nat.max: nat -> nat -> nat
-Nat.sub: nat -> nat -> nat
+Nat.tail_add: nat -> nat -> nat
Nat.pow: nat -> nat -> nat
-Nat.lxor: nat -> nat -> nat
-Nat.ldiff: nat -> nat -> nat
Nat.min: nat -> nat -> nat
-Nat.add: nat -> nat -> nat
+Nat.modulo: nat -> nat -> nat
+Nat.div: nat -> nat -> nat
+Nat.lor: nat -> nat -> nat
+Nat.gcd: nat -> nat -> nat
+Hexadecimal.nb_digits: Hexadecimal.uint -> nat
+Nat.of_hex_uint: Hexadecimal.uint -> nat
+Nat.of_num_uint: Numeral.uint -> nat
Nat.of_uint: Decimal.uint -> nat
Decimal.nb_digits: Decimal.uint -> nat
Nat.tail_addmul: nat -> nat -> nat -> nat
+Nat.of_hex_uint_acc: Hexadecimal.uint -> nat -> nat
Nat.of_uint_acc: Decimal.uint -> nat -> nat
Nat.sqrt_iter: nat -> nat -> nat -> nat -> nat
Nat.log2_iter: nat -> nat -> nat -> nat -> nat
@@ -53,29 +67,30 @@ Nat.div2: nat -> nat
Nat.sqrt: nat -> nat
Nat.log2: nat -> nat
Nat.double: nat -> nat
-Nat.pred: nat -> nat
+S: nat -> nat
Nat.square: nat -> nat
Nat.succ: nat -> nat
-S: nat -> nat
+Nat.pred: nat -> nat
+Nat.land: nat -> nat -> nat
+Nat.max: nat -> nat -> nat
+Nat.gcd: nat -> nat -> nat
+Nat.modulo: nat -> nat -> nat
Nat.ldiff: nat -> nat -> nat
+Nat.tail_add: nat -> nat -> nat
Nat.pow: nat -> nat -> nat
-Nat.land: nat -> nat -> nat
Nat.lxor: nat -> nat -> nat
Nat.div: nat -> nat -> nat
Nat.lor: nat -> nat -> nat
-Nat.tail_mul: nat -> nat -> nat
-Nat.modulo: nat -> nat -> nat
-Nat.sub: nat -> nat -> nat
Nat.mul: nat -> nat -> nat
-Nat.gcd: nat -> nat -> nat
-Nat.max: nat -> nat -> nat
-Nat.tail_add: nat -> nat -> nat
-Nat.add: nat -> nat -> nat
Nat.min: nat -> nat -> nat
+Nat.add: nat -> nat -> nat
+Nat.sub: nat -> nat -> nat
+Nat.tail_mul: nat -> nat -> nat
Nat.tail_addmul: nat -> nat -> nat -> nat
Nat.of_uint_acc: Decimal.uint -> nat -> nat
-Nat.log2_iter: nat -> nat -> nat -> nat -> nat
+Nat.of_hex_uint_acc: Hexadecimal.uint -> nat -> nat
Nat.sqrt_iter: nat -> nat -> nat -> nat -> nat
+Nat.log2_iter: nat -> nat -> nat -> nat -> nat
Nat.bitwise: (bool -> bool -> bool) -> nat -> nat -> nat -> nat
(use "About" for full details on implicit arguments)
mult_n_Sm: forall n m : nat, n * m + n = n * S m
diff --git a/test-suite/output/ZSyntax.out b/test-suite/output/ZSyntax.out
index dc41b0aa48..c3f2fb5b42 100644
--- a/test-suite/output/ZSyntax.out
+++ b/test-suite/output/ZSyntax.out
@@ -1,5 +1,7 @@
32%Z
: Z
+eq_refl : 42%Z = 42%Z
+ : 42%Z = 42%Z
fun f : nat -> Z => (f 0%nat + 0)%Z
: (nat -> Z) -> Z
fun x : positive => Z.pos x~0
@@ -24,3 +26,9 @@ Z.of_nat 0 = 0%Z
: Prop
(0 + Z.of_nat 11)%Z
: Z
+0x2a%Z
+ : Z
+(-0x2a)%Z
+ : Z
+0x0%Z
+ : Z
diff --git a/test-suite/output/ZSyntax.v b/test-suite/output/ZSyntax.v
index d3640cae44..be9dc543d6 100644
--- a/test-suite/output/ZSyntax.v
+++ b/test-suite/output/ZSyntax.v
@@ -1,5 +1,6 @@
Require Import ZArith.
Check 32%Z.
+Check (eq_refl : 0x2a%Z = 42%Z).
Check (fun f : nat -> Z => (f 0%nat + 0)%Z).
Check (fun x : positive => Zpos (xO x)).
Check (fun x : positive => (Zpos x + 1)%Z).
@@ -15,3 +16,10 @@ Check (Z.of_nat 0 = 0%Z).
(* Submitted by Pierre Casteran *)
Require Import Arith.
Check (0 + Z.of_nat 11)%Z.
+
+(* Check hexadecimal printing *)
+Definition to_num_int n := Numeral.IntHex (Z.to_hex_int n).
+Numeral Notation Z Z.of_num_int to_num_int : Z_scope.
+Check 42%Z.
+Check (-42)%Z.
+Check 0%Z.
diff --git a/test-suite/output/bug_12159.v b/test-suite/output/bug_12159.v
index 91d66f7f4c..6ea90eab29 100644
--- a/test-suite/output/bug_12159.v
+++ b/test-suite/output/bug_12159.v
@@ -2,10 +2,10 @@ Declare Scope A.
Declare Scope B.
Delimit Scope A with A.
Delimit Scope B with B.
-Definition to_unit (v : Decimal.uint) : option unit
- := match Nat.of_uint v with O => Some tt | _ => None end.
-Definition of_unit (v : unit) : Decimal.uint := Nat.to_uint 0.
-Definition of_unit' (v : unit) : Decimal.uint := Nat.to_uint 1.
+Definition to_unit (v : Numeral.uint) : option unit
+ := match Nat.of_num_uint v with O => Some tt | _ => None end.
+Definition of_unit (v : unit) : Numeral.uint := Nat.to_num_uint 0.
+Definition of_unit' (v : unit) : Numeral.uint := Nat.to_num_uint 1.
Numeral Notation unit to_unit of_unit : A.
Numeral Notation unit to_unit of_unit' : B.
Definition f x : unit := x.
diff --git a/theories/Init/Datatypes.v b/theories/Init/Datatypes.v
index b094f81dc6..cba90043d5 100644
--- a/theories/Init/Datatypes.v
+++ b/theories/Init/Datatypes.v
@@ -159,6 +159,9 @@ Inductive nat : Set :=
| O : nat
| S : nat -> nat.
+Declare Scope hex_nat_scope.
+Delimit Scope hex_nat_scope with xnat.
+
Declare Scope nat_scope.
Delimit Scope nat_scope with nat.
Bind Scope nat_scope with nat.
diff --git a/theories/Init/Decimal.v b/theories/Init/Decimal.v
index 2a84456500..5eae5567d7 100644
--- a/theories/Init/Decimal.v
+++ b/theories/Init/Decimal.v
@@ -16,7 +16,7 @@
We represent numbers in base 10 as lists of decimal digits,
in big-endian order (most significant digit comes first). *)
-Require Import Datatypes.
+Require Import Datatypes Specif.
(** Unsigned integers are just lists of digits.
For instance, ten is (D1 (D0 Nil)) *)
@@ -53,6 +53,10 @@ Variant decimal :=
| Decimal (i:int) (f:uint)
| DecimalExp (i:int) (f:uint) (e:int).
+Scheme Equality for uint.
+Scheme Equality for int.
+Scheme Equality for decimal.
+
Declare Scope dec_uint_scope.
Delimit Scope dec_uint_scope with uint.
Bind Scope dec_uint_scope with uint.
@@ -172,8 +176,8 @@ Fixpoint del_head n d :=
Definition del_head_int n d :=
match d with
- | Pos d => Pos (del_head n d)
- | Neg d => Neg (del_head n d)
+ | Pos d => del_head n d
+ | Neg d => del_head n d
end.
(** [del_tail n d] removes [n] digits at end of [d]
diff --git a/theories/Init/Hexadecimal.v b/theories/Init/Hexadecimal.v
new file mode 100644
index 0000000000..a4ddad2875
--- /dev/null
+++ b/theories/Init/Hexadecimal.v
@@ -0,0 +1,245 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * Hexadecimal numbers *)
+
+(** These numbers coded in base 16 will be used for parsing and printing
+ other Coq numeral datatypes in an human-readable way.
+ See the [Numeral Notation] command.
+ We represent numbers in base 16 as lists of hexadecimal digits,
+ in big-endian order (most significant digit comes first). *)
+
+Require Import Datatypes Specif Decimal.
+
+(** Unsigned integers are just lists of digits.
+ For instance, sixteen is (D1 (D0 Nil)) *)
+
+Inductive uint :=
+ | Nil
+ | D0 (_:uint)
+ | D1 (_:uint)
+ | D2 (_:uint)
+ | D3 (_:uint)
+ | D4 (_:uint)
+ | D5 (_:uint)
+ | D6 (_:uint)
+ | D7 (_:uint)
+ | D8 (_:uint)
+ | D9 (_:uint)
+ | Da (_:uint)
+ | Db (_:uint)
+ | Dc (_:uint)
+ | Dd (_:uint)
+ | De (_:uint)
+ | Df (_:uint).
+
+(** [Nil] is the number terminator. Taken alone, it behaves as zero,
+ but rather use [D0 Nil] instead, since this form will be denoted
+ as [0], while [Nil] will be printed as [Nil]. *)
+
+Notation zero := (D0 Nil).
+
+(** For signed integers, we use two constructors [Pos] and [Neg]. *)
+
+Variant int := Pos (d:uint) | Neg (d:uint).
+
+(** For decimal numbers, we use two constructors [Hexadecimal] and
+ [HexadecimalExp], depending on whether or not they are given with an
+ exponent (e.g., 0x1.a2p+01). [i] is the integral part while [f] is
+ the fractional part (beware that leading zeroes do matter). *)
+
+Variant hexadecimal :=
+ | Hexadecimal (i:int) (f:uint)
+ | HexadecimalExp (i:int) (f:uint) (e:Decimal.int).
+
+Scheme Equality for uint.
+Scheme Equality for int.
+Scheme Equality for hexadecimal.
+
+Declare Scope hex_uint_scope.
+Delimit Scope hex_uint_scope with huint.
+Bind Scope hex_uint_scope with uint.
+
+Declare Scope hex_int_scope.
+Delimit Scope hex_int_scope with hint.
+Bind Scope hex_int_scope with int.
+
+Register uint as num.hexadecimal_uint.type.
+Register int as num.hexadecimal_int.type.
+Register hexadecimal as num.hexadecimal.type.
+
+Fixpoint nb_digits d :=
+ match d with
+ | Nil => O
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d
+ | Da d | Db d | Dc d | Dd d | De d | Df d =>
+ S (nb_digits d)
+ end.
+
+(** This representation favors simplicity over canonicity.
+ For normalizing numbers, we need to remove head zero digits,
+ and choose our canonical representation of 0 (here [D0 Nil]
+ for unsigned numbers and [Pos (D0 Nil)] for signed numbers). *)
+
+(** [nzhead] removes all head zero digits *)
+
+Fixpoint nzhead d :=
+ match d with
+ | D0 d => nzhead d
+ | _ => d
+ end.
+
+(** [unorm] : normalization of unsigned integers *)
+
+Definition unorm d :=
+ match nzhead d with
+ | Nil => zero
+ | d => d
+ end.
+
+(** [norm] : normalization of signed integers *)
+
+Definition norm d :=
+ match d with
+ | Pos d => Pos (unorm d)
+ | Neg d =>
+ match nzhead d with
+ | Nil => Pos zero
+ | d => Neg d
+ end
+ end.
+
+(** A few easy operations. For more advanced computations, use the conversions
+ with other Coq numeral datatypes (e.g. Z) and the operations on them. *)
+
+Definition opp (d:int) :=
+ match d with
+ | Pos d => Neg d
+ | Neg d => Pos d
+ end.
+
+(** For conversions with binary numbers, it is easier to operate
+ on little-endian numbers. *)
+
+Fixpoint revapp (d d' : uint) :=
+ match d with
+ | Nil => d'
+ | D0 d => revapp d (D0 d')
+ | D1 d => revapp d (D1 d')
+ | D2 d => revapp d (D2 d')
+ | D3 d => revapp d (D3 d')
+ | D4 d => revapp d (D4 d')
+ | D5 d => revapp d (D5 d')
+ | D6 d => revapp d (D6 d')
+ | D7 d => revapp d (D7 d')
+ | D8 d => revapp d (D8 d')
+ | D9 d => revapp d (D9 d')
+ | Da d => revapp d (Da d')
+ | Db d => revapp d (Db d')
+ | Dc d => revapp d (Dc d')
+ | Dd d => revapp d (Dd d')
+ | De d => revapp d (De d')
+ | Df d => revapp d (Df d')
+ end.
+
+Definition rev d := revapp d Nil.
+
+Definition app d d' := revapp (rev d) d'.
+
+Definition app_int d1 d2 :=
+ match d1 with Pos d1 => Pos (app d1 d2) | Neg d1 => Neg (app d1 d2) end.
+
+(** [nztail] removes all trailing zero digits and return both the
+ result and the number of removed digits. *)
+
+Definition nztail d :=
+ let fix aux d_rev :=
+ match d_rev with
+ | D0 d_rev => let (r, n) := aux d_rev in pair r (S n)
+ | _ => pair d_rev O
+ end in
+ let (r, n) := aux (rev d) in pair (rev r) n.
+
+Definition nztail_int d :=
+ match d with
+ | Pos d => let (r, n) := nztail d in pair (Pos r) n
+ | Neg d => let (r, n) := nztail d in pair (Neg r) n
+ end.
+
+Module Little.
+
+(** Successor of little-endian numbers *)
+
+Fixpoint succ d :=
+ match d with
+ | Nil => D1 Nil
+ | D0 d => D1 d
+ | D1 d => D2 d
+ | D2 d => D3 d
+ | D3 d => D4 d
+ | D4 d => D5 d
+ | D5 d => D6 d
+ | D6 d => D7 d
+ | D7 d => D8 d
+ | D8 d => D9 d
+ | D9 d => Da d
+ | Da d => Db d
+ | Db d => Dc d
+ | Dc d => Dd d
+ | Dd d => De d
+ | De d => Df d
+ | Df d => D0 (succ d)
+ end.
+
+(** Doubling little-endian numbers *)
+
+Fixpoint double d :=
+ match d with
+ | Nil => Nil
+ | D0 d => D0 (double d)
+ | D1 d => D2 (double d)
+ | D2 d => D4 (double d)
+ | D3 d => D6 (double d)
+ | D4 d => D8 (double d)
+ | D5 d => Da (double d)
+ | D6 d => Dc (double d)
+ | D7 d => De (double d)
+ | D8 d => D0 (succ_double d)
+ | D9 d => D2 (succ_double d)
+ | Da d => D4 (succ_double d)
+ | Db d => D6 (succ_double d)
+ | Dc d => D8 (succ_double d)
+ | Dd d => Da (succ_double d)
+ | De d => Dc (succ_double d)
+ | Df d => De (succ_double d)
+ end
+
+with succ_double d :=
+ match d with
+ | Nil => D1 Nil
+ | D0 d => D1 (double d)
+ | D1 d => D3 (double d)
+ | D2 d => D5 (double d)
+ | D3 d => D7 (double d)
+ | D4 d => D9 (double d)
+ | D5 d => Db (double d)
+ | D6 d => Dd (double d)
+ | D7 d => Df (double d)
+ | D8 d => D1 (succ_double d)
+ | D9 d => D3 (succ_double d)
+ | Da d => D5 (succ_double d)
+ | Db d => D7 (succ_double d)
+ | Dc d => D9 (succ_double d)
+ | Dd d => Db (succ_double d)
+ | De d => Dd (succ_double d)
+ | Df d => Df (succ_double d)
+ end.
+
+End Little.
diff --git a/theories/Init/Nat.v b/theories/Init/Nat.v
index bfd41c5017..7c8cf0b536 100644
--- a/theories/Init/Nat.v
+++ b/theories/Init/Nat.v
@@ -9,7 +9,7 @@
(************************************************************************)
Require Import Notations Logic Datatypes.
-Require Decimal.
+Require Decimal Hexadecimal Numeral.
Local Open Scope nat_scope.
(**********************************************************************)
@@ -187,6 +187,37 @@ Fixpoint of_uint_acc (d:Decimal.uint)(acc:nat) :=
Definition of_uint (d:Decimal.uint) := of_uint_acc d O.
+Local Notation sixteen := (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S O)))))))))))))))).
+
+Fixpoint of_hex_uint_acc (d:Hexadecimal.uint)(acc:nat) :=
+ match d with
+ | Hexadecimal.Nil => acc
+ | Hexadecimal.D0 d => of_hex_uint_acc d (tail_mul sixteen acc)
+ | Hexadecimal.D1 d => of_hex_uint_acc d (S (tail_mul sixteen acc))
+ | Hexadecimal.D2 d => of_hex_uint_acc d (S (S (tail_mul sixteen acc)))
+ | Hexadecimal.D3 d => of_hex_uint_acc d (S (S (S (tail_mul sixteen acc))))
+ | Hexadecimal.D4 d => of_hex_uint_acc d (S (S (S (S (tail_mul sixteen acc)))))
+ | Hexadecimal.D5 d => of_hex_uint_acc d (S (S (S (S (S (tail_mul sixteen acc))))))
+ | Hexadecimal.D6 d => of_hex_uint_acc d (S (S (S (S (S (S (tail_mul sixteen acc)))))))
+ | Hexadecimal.D7 d => of_hex_uint_acc d (S (S (S (S (S (S (S (tail_mul sixteen acc))))))))
+ | Hexadecimal.D8 d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (tail_mul sixteen acc)))))))))
+ | Hexadecimal.D9 d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc))))))))))
+ | Hexadecimal.Da d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc)))))))))))
+ | Hexadecimal.Db d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc))))))))))))
+ | Hexadecimal.Dc d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc)))))))))))))
+ | Hexadecimal.Dd d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc))))))))))))))
+ | Hexadecimal.De d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc)))))))))))))))
+ | Hexadecimal.Df d => of_hex_uint_acc d (S (S (S (S (S (S (S (S (S (S (S (S (S (S (S (tail_mul sixteen acc))))))))))))))))
+ end.
+
+Definition of_hex_uint (d:Hexadecimal.uint) := of_hex_uint_acc d O.
+
+Definition of_num_uint (d:Numeral.uint) :=
+ match d with
+ | Numeral.UIntDec d => of_uint d
+ | Numeral.UIntHex d => of_hex_uint d
+ end.
+
Fixpoint to_little_uint n acc :=
match n with
| O => acc
@@ -196,14 +227,43 @@ Fixpoint to_little_uint n acc :=
Definition to_uint n :=
Decimal.rev (to_little_uint n Decimal.zero).
+Fixpoint to_little_hex_uint n acc :=
+ match n with
+ | O => acc
+ | S n => to_little_hex_uint n (Hexadecimal.Little.succ acc)
+ end.
+
+Definition to_hex_uint n :=
+ Hexadecimal.rev (to_little_hex_uint n Hexadecimal.zero).
+
+Definition to_num_uint n := Numeral.UIntDec (to_uint n).
+
+Definition to_num_hex_uint n := Numeral.UIntHex (to_hex_uint n).
+
Definition of_int (d:Decimal.int) : option nat :=
match Decimal.norm d with
| Decimal.Pos u => Some (of_uint u)
| _ => None
end.
+Definition of_hex_int (d:Hexadecimal.int) : option nat :=
+ match Hexadecimal.norm d with
+ | Hexadecimal.Pos u => Some (of_hex_uint u)
+ | _ => None
+ end.
+
+Definition of_num_int (d:Numeral.int) : option nat :=
+ match d with
+ | Numeral.IntDec d => of_int d
+ | Numeral.IntHex d => of_hex_int d
+ end.
+
Definition to_int n := Decimal.Pos (to_uint n).
+Definition to_hex_int n := Hexadecimal.Pos (to_hex_uint n).
+
+Definition to_num_int n := Numeral.IntDec (to_int n).
+
(** ** Euclidean division *)
(** This division is linear and tail-recursive.
diff --git a/theories/Init/Numeral.v b/theories/Init/Numeral.v
new file mode 100644
index 0000000000..8a0531e004
--- /dev/null
+++ b/theories/Init/Numeral.v
@@ -0,0 +1,33 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * Decimal or Hexadecimal numbers *)
+
+Require Import Decimal Hexadecimal.
+
+Variant uint := UIntDec (u:Decimal.uint) | UIntHex (u:Hexadecimal.uint).
+
+Variant int := IntDec (i:Decimal.int) | IntHex (i:Hexadecimal.int).
+
+Variant numeral := Dec (d:Decimal.decimal) | Hex (h:Hexadecimal.hexadecimal).
+
+Scheme Equality for uint.
+Scheme Equality for int.
+Scheme Equality for numeral.
+
+Register uint as num.num_uint.type.
+Register int as num.num_int.type.
+Register numeral as num.numeral.type.
+
+(** Pseudo-conversion functions used when declaring
+ Numeral Notations on [uint] and [int]. *)
+
+Definition uint_of_uint (i:uint) := i.
+Definition int_of_int (i:int) := i.
diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v
index 6a81517d7e..8f862e8cec 100644
--- a/theories/Init/Prelude.v
+++ b/theories/Init/Prelude.v
@@ -15,6 +15,8 @@ Require Export Datatypes.
Require Export Specif.
Require Coq.Init.Byte.
Require Coq.Init.Decimal.
+Require Coq.Init.Hexadecimal.
+Require Coq.Init.Numeral.
Require Coq.Init.Nat.
Require Export Peano.
Require Export Coq.Init.Wf.
@@ -30,16 +32,25 @@ Declare ML Module "ground_plugin".
Declare ML Module "numeral_notation_plugin".
Declare ML Module "string_notation_plugin".
+(* Parsing / printing of hexadecimal numbers *)
+Arguments Nat.of_hex_uint d%hex_uint_scope.
+Arguments Nat.of_hex_int d%hex_int_scope.
+Numeral Notation Numeral.uint Numeral.uint_of_uint Numeral.uint_of_uint
+ : hex_uint_scope.
+Numeral Notation Numeral.int Numeral.int_of_int Numeral.int_of_int
+ : hex_int_scope.
+
(* Parsing / printing of decimal numbers *)
Arguments Nat.of_uint d%dec_uint_scope.
Arguments Nat.of_int d%dec_int_scope.
-Numeral Notation Decimal.uint Decimal.uint_of_uint Decimal.uint_of_uint
+Numeral Notation Numeral.uint Numeral.uint_of_uint Numeral.uint_of_uint
: dec_uint_scope.
-Numeral Notation Decimal.int Decimal.int_of_int Decimal.int_of_int
+Numeral Notation Numeral.int Numeral.int_of_int Numeral.int_of_int
: dec_int_scope.
(* Parsing / printing of [nat] numbers *)
-Numeral Notation nat Nat.of_uint Nat.to_uint : nat_scope (abstract after 5001).
+Numeral Notation nat Nat.of_num_uint Nat.to_num_hex_uint : hex_nat_scope (abstract after 5001).
+Numeral Notation nat Nat.of_num_uint Nat.to_num_uint : nat_scope (abstract after 5001).
(* Printing/Parsing of bytes *)
Export Byte.ByteSyntaxNotations.
diff --git a/theories/NArith/BinNat.v b/theories/NArith/BinNat.v
index d68c32b371..1881e387a2 100644
--- a/theories/NArith/BinNat.v
+++ b/theories/NArith/BinNat.v
@@ -1007,7 +1007,7 @@ Bind Scope N_scope with N.t N.
(** Exportation of notations *)
-Numeral Notation N N.of_uint N.to_uint : N_scope.
+Numeral Notation N N.of_num_uint N.to_num_uint : N_scope.
Infix "+" := N.add : N_scope.
Infix "-" := N.sub : N_scope.
diff --git a/theories/NArith/BinNatDef.v b/theories/NArith/BinNatDef.v
index 04685cc3eb..8a0aee9cf4 100644
--- a/theories/NArith/BinNatDef.v
+++ b/theories/NArith/BinNatDef.v
@@ -388,23 +388,55 @@ Definition iter (n:N) {A} (f:A->A) (x:A) : A :=
Definition of_uint (d:Decimal.uint) := Pos.of_uint d.
+Definition of_hex_uint (d:Hexadecimal.uint) := Pos.of_hex_uint d.
+
+Definition of_num_uint (d:Numeral.uint) :=
+ match d with
+ | Numeral.UIntDec d => of_uint d
+ | Numeral.UIntHex d => of_hex_uint d
+ end.
+
Definition of_int (d:Decimal.int) :=
match Decimal.norm d with
| Decimal.Pos d => Some (Pos.of_uint d)
| Decimal.Neg _ => None
end.
+Definition of_hex_int (d:Hexadecimal.int) :=
+ match Hexadecimal.norm d with
+ | Hexadecimal.Pos d => Some (Pos.of_hex_uint d)
+ | Hexadecimal.Neg _ => None
+ end.
+
+Definition of_num_int (d:Numeral.int) :=
+ match d with
+ | Numeral.IntDec d => of_int d
+ | Numeral.IntHex d => of_hex_int d
+ end.
+
Definition to_uint n :=
match n with
| 0 => Decimal.zero
| pos p => Pos.to_uint p
end.
+Definition to_hex_uint n :=
+ match n with
+ | 0 => Hexadecimal.zero
+ | pos p => Pos.to_hex_uint p
+ end.
+
+Definition to_num_uint n := Numeral.UIntDec (to_uint n).
+
Definition to_int n := Decimal.Pos (to_uint n).
-Numeral Notation N of_uint to_uint : N_scope.
+Definition to_hex_int n := Hexadecimal.Pos (to_hex_uint n).
+
+Definition to_num_int n := Numeral.IntDec (to_int n).
+
+Numeral Notation N of_num_uint to_num_uint : N_scope.
End N.
(** Re-export the notation for those who just [Import NatIntDef] *)
-Numeral Notation N N.of_uint N.to_uint : N_scope.
+Numeral Notation N N.of_num_uint N.to_num_uint : N_scope.
diff --git a/theories/Numbers/AltBinNotations.v b/theories/Numbers/AltBinNotations.v
index a074ec6d2a..5585f478b3 100644
--- a/theories/Numbers/AltBinNotations.v
+++ b/theories/Numbers/AltBinNotations.v
@@ -55,10 +55,10 @@ Definition n_of_z z :=
end.
Definition n_to_z n :=
- match n with
- | N0 => Z0
- | Npos p => Zpos p
- end.
+ match n with
+ | N0 => Z0
+ | Npos p => Zpos p
+ end.
Numeral Notation N n_of_z n_to_z : N_scope.
diff --git a/theories/Numbers/DecimalFacts.v b/theories/Numbers/DecimalFacts.v
index a3bdcc579f..dd361562ba 100644
--- a/theories/Numbers/DecimalFacts.v
+++ b/theories/Numbers/DecimalFacts.v
@@ -10,134 +10,450 @@
(** * DecimalFacts : some facts about Decimal numbers *)
-Require Import Decimal.
+Require Import Decimal Arith.
Lemma uint_dec (d d' : uint) : { d = d' } + { d <> d' }.
Proof.
- decide equality.
+ decide equality.
Defined.
Lemma rev_revapp d d' :
- rev (revapp d d') = revapp d' d.
+ rev (revapp d d') = revapp d' d.
Proof.
- revert d'. induction d; simpl; intros; now rewrite ?IHd.
+ revert d'. induction d; simpl; intros; now rewrite ?IHd.
Qed.
Lemma rev_rev d : rev (rev d) = d.
Proof.
- apply rev_revapp.
+ apply rev_revapp.
Qed.
+Lemma revapp_rev_nil d : revapp (rev d) Nil = d.
+Proof. now fold (rev (rev d)); rewrite rev_rev. Qed.
+
+Lemma app_nil_r d : app d Nil = d.
+Proof. now unfold app; rewrite revapp_rev_nil. Qed.
+
+Lemma app_int_nil_r d : app_int d Nil = d.
+Proof. now case d; intro d'; simpl; rewrite app_nil_r. Qed.
+
+Lemma revapp_revapp_1 d d' d'' :
+ nb_digits d <= 1 ->
+ revapp (revapp d d') d'' = revapp d' (revapp d d'').
+Proof.
+ now case d; clear d; intro d;
+ [|case d; clear d; intro d;
+ [|simpl; case nb_digits; [|intros n]; intros Hn; exfalso;
+ [apply (Nat.nle_succ_diag_l _ Hn)|
+ apply (Nat.nle_succ_0 _ (le_S_n _ _ Hn))]..]..].
+Qed.
+
+Lemma nb_digits_pos d : d <> Nil -> 0 < nb_digits d.
+Proof. now case d; [|intros d' _; apply Nat.lt_0_succ..]. Qed.
+
+Lemma nb_digits_revapp d d' :
+ nb_digits (revapp d d') = nb_digits d + nb_digits d'.
+Proof.
+ now revert d'; induction d; [|intro d'; simpl; rewrite IHd; simpl..].
+Qed.
+
+Lemma nb_digits_rev u : nb_digits (rev u) = nb_digits u.
+Proof. now unfold rev; rewrite nb_digits_revapp. Qed.
+
+Lemma nb_digits_nzhead u : nb_digits (nzhead u) <= nb_digits u.
+Proof. now induction u; [|apply le_S|..]. Qed.
+
+Lemma del_head_nb_digits (u:uint) : del_head (nb_digits u) u = Nil.
+Proof. now induction u. Qed.
+
+Lemma nb_digits_del_head n u :
+ n <= nb_digits u -> nb_digits (del_head n u) = nb_digits u - n.
+Proof.
+ revert u; induction n; intros u; [now rewrite Nat.sub_0_r|].
+ now case u; clear u; intro u; [|intro Hn; apply IHn, le_S_n..].
+Qed.
+
+Lemma nb_digits_iter_D0 n d :
+ nb_digits (Nat.iter n D0 d) = n + nb_digits d.
+Proof. now induction n; simpl; [|rewrite IHn]. Qed.
+
+Fixpoint nth n u :=
+ match n with
+ | O =>
+ match u with
+ | Nil => Nil
+ | D0 d => D0 Nil
+ | D1 d => D1 Nil
+ | D2 d => D2 Nil
+ | D3 d => D3 Nil
+ | D4 d => D4 Nil
+ | D5 d => D5 Nil
+ | D6 d => D6 Nil
+ | D7 d => D7 Nil
+ | D8 d => D8 Nil
+ | D9 d => D9 Nil
+ end
+ | S n =>
+ match u with
+ | Nil => Nil
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d =>
+ nth n d
+ end
+ end.
+
+Lemma nb_digits_nth n u : nb_digits (nth n u) <= 1.
+Proof.
+ revert u; induction n.
+ - now intro u; case u; [apply Nat.le_0_1|..].
+ - intro u; case u; [apply Nat.le_0_1|intro u'; apply IHn..].
+Qed.
+
+Lemma del_head_nth n u :
+ n < nb_digits u ->
+ del_head n u = revapp (nth n u) (del_head (S n) u).
+Proof.
+ revert u; induction n; intro u; [now case u|].
+ now case u; [|intro u'; intro H; apply IHn, le_S_n..].
+Qed.
+
+Lemma nth_revapp_r n d d' :
+ nb_digits d <= n ->
+ nth n (revapp d d') = nth (n - nb_digits d) d'.
+Proof.
+ revert d d'; induction n; intro d.
+ - now case d; intro d';
+ [case d'|intros d'' H; exfalso; revert H; apply Nat.nle_succ_0..].
+ - now induction d;
+ [intro d'; case d'|
+ intros d' H;
+ simpl revapp; rewrite IHd; [|now apply le_Sn_le];
+ rewrite Nat.sub_succ_l; [|now apply le_S_n];
+ simpl; rewrite <-(IHn _ _ (le_S_n _ _ H))..].
+Qed.
+
+Lemma nth_revapp_l n d d' :
+ n < nb_digits d ->
+ nth n (revapp d d') = nth (nb_digits d - n - 1) d.
+Proof.
+ revert d d'; induction n; intro d.
+ - rewrite Nat.sub_0_r.
+ now induction d;
+ [|intros d' _; simpl revapp;
+ revert IHd; case d; clear d; [|intro d..]; intro IHd;
+ [|rewrite IHd; [simpl nb_digits; rewrite (Nat.sub_succ_l _ (S _))|];
+ [|apply le_n_S, Nat.le_0_l..]..]..].
+ - now induction d;
+ [|intros d' H;
+ simpl revapp; simpl nb_digits;
+ simpl in H; generalize (lt_S_n _ _ H); clear H; intro H;
+ case (le_lt_eq_dec _ _ H); clear H; intro H;
+ [rewrite (IHd _ H), Nat.sub_succ_l;
+ [rewrite Nat.sub_succ_l; [|apply Nat.le_add_le_sub_r]|
+ apply le_Sn_le]|
+ rewrite nth_revapp_r; rewrite <-H;
+ [rewrite Nat.sub_succ, Nat.sub_succ_l; [rewrite !Nat.sub_diag|]|]]..].
+Qed.
+
+Lemma app_del_tail_head (u:uint) n :
+ n <= nb_digits u ->
+ app (del_tail n u) (del_head (nb_digits u - n) u) = u.
+Proof.
+ unfold app, del_tail; rewrite rev_rev.
+ induction n.
+ - intros _; rewrite Nat.sub_0_r, del_head_nb_digits; simpl.
+ now rewrite revapp_rev_nil.
+ - intro Hn.
+ rewrite (del_head_nth (_ - _));
+ [|now apply Nat.sub_lt; [|apply Nat.lt_0_succ]].
+ rewrite Nat.sub_succ_r, <-Nat.sub_1_r.
+ rewrite <-(nth_revapp_l _ _ Nil Hn); fold (rev u).
+ rewrite <-revapp_revapp_1; [|now apply nb_digits_nth].
+ rewrite <-(del_head_nth _ _); [|now rewrite nb_digits_rev].
+ rewrite Nat.sub_1_r, Nat.succ_pred_pos; [|now apply Nat.lt_add_lt_sub_r].
+ apply (IHn (le_Sn_le _ _ Hn)).
+Qed.
+
+Lemma app_int_del_tail_head n (d:int) :
+ let ad := match d with Pos d | Neg d => d end in
+ n <= nb_digits ad ->
+ app_int (del_tail_int n d) (del_head (nb_digits ad - n) ad) = d.
+Proof. now case d; clear d; simpl; intros u Hu; rewrite app_del_tail_head. Qed.
+
(** Normalization on little-endian numbers *)
Fixpoint nztail d :=
- match d with
- | Nil => Nil
- | D0 d => match nztail d with Nil => Nil | d' => D0 d' end
- | D1 d => D1 (nztail d)
- | D2 d => D2 (nztail d)
- | D3 d => D3 (nztail d)
- | D4 d => D4 (nztail d)
- | D5 d => D5 (nztail d)
- | D6 d => D6 (nztail d)
- | D7 d => D7 (nztail d)
- | D8 d => D8 (nztail d)
- | D9 d => D9 (nztail d)
- end.
+ match d with
+ | Nil => Nil
+ | D0 d => match nztail d with Nil => Nil | d' => D0 d' end
+ | D1 d => D1 (nztail d)
+ | D2 d => D2 (nztail d)
+ | D3 d => D3 (nztail d)
+ | D4 d => D4 (nztail d)
+ | D5 d => D5 (nztail d)
+ | D6 d => D6 (nztail d)
+ | D7 d => D7 (nztail d)
+ | D8 d => D8 (nztail d)
+ | D9 d => D9 (nztail d)
+ end.
Definition lnorm d :=
- match nztail d with
- | Nil => zero
- | d => d
- end.
+ match nztail d with
+ | Nil => zero
+ | d => d
+ end.
Lemma nzhead_revapp_0 d d' : nztail d = Nil ->
- nzhead (revapp d d') = nzhead d'.
+ nzhead (revapp d d') = nzhead d'.
Proof.
- revert d'. induction d; intros d' [=]; simpl; trivial.
- destruct (nztail d); now rewrite IHd.
+ revert d'. induction d; intros d' [=]; simpl; trivial.
+ destruct (nztail d); now rewrite IHd.
Qed.
Lemma nzhead_revapp d d' : nztail d <> Nil ->
- nzhead (revapp d d') = revapp (nztail d) d'.
+ nzhead (revapp d d') = revapp (nztail d) d'.
Proof.
- revert d'.
- induction d; intros d' H; simpl in *;
- try destruct (nztail d) eqn:E;
+ revert d'.
+ induction d; intros d' H; simpl in *;
+ try destruct (nztail d) eqn:E;
(now rewrite ?nzhead_revapp_0) || (now rewrite IHd).
Qed.
Lemma nzhead_rev d : nztail d <> Nil ->
- nzhead (rev d) = rev (nztail d).
+ nzhead (rev d) = rev (nztail d).
Proof.
- apply nzhead_revapp.
+ apply nzhead_revapp.
Qed.
Lemma rev_nztail_rev d :
rev (nztail (rev d)) = nzhead d.
Proof.
- destruct (uint_dec (nztail (rev d)) Nil) as [H|H].
- - rewrite H. unfold rev; simpl.
- rewrite <- (rev_rev d). symmetry.
- now apply nzhead_revapp_0.
- - now rewrite <- nzhead_rev, rev_rev.
+ destruct (uint_dec (nztail (rev d)) Nil) as [H|H].
+ - rewrite H. unfold rev; simpl.
+ rewrite <- (rev_rev d). symmetry.
+ now apply nzhead_revapp_0.
+ - now rewrite <- nzhead_rev, rev_rev.
Qed.
+Lemma nzhead_D0 u : nzhead (D0 u) = nzhead u.
+Proof. reflexivity. Qed.
+
+Lemma nzhead_iter_D0 n u : nzhead (Nat.iter n D0 u) = nzhead u.
+Proof. now induction n. Qed.
+
Lemma revapp_nil_inv d d' : revapp d d' = Nil -> d = Nil /\ d' = Nil.
Proof.
- revert d'.
- induction d; simpl; intros d' H; auto; now apply IHd in H.
+ revert d'.
+ induction d; simpl; intros d' H; auto; now apply IHd in H.
Qed.
Lemma rev_nil_inv d : rev d = Nil -> d = Nil.
Proof.
- apply revapp_nil_inv.
+ apply revapp_nil_inv.
Qed.
Lemma rev_lnorm_rev d :
rev (lnorm (rev d)) = unorm d.
Proof.
- unfold unorm, lnorm.
- rewrite <- rev_nztail_rev.
- destruct nztail; simpl; trivial;
- destruct rev eqn:E; trivial; now apply rev_nil_inv in E.
+ unfold unorm, lnorm.
+ rewrite <- rev_nztail_rev.
+ destruct nztail; simpl; trivial;
+ destruct rev eqn:E; trivial; now apply rev_nil_inv in E.
Qed.
Lemma nzhead_nonzero d d' : nzhead d <> D0 d'.
Proof.
- induction d; easy.
+ induction d; easy.
Qed.
Lemma unorm_0 d : unorm d = zero <-> nzhead d = Nil.
Proof.
- unfold unorm. split.
- - generalize (nzhead_nonzero d).
- destruct nzhead; intros H [=]; trivial. now destruct (H u).
- - now intros ->.
+ unfold unorm. split.
+ - generalize (nzhead_nonzero d).
+ destruct nzhead; intros H [=]; trivial. now destruct (H u).
+ - now intros ->.
Qed.
Lemma unorm_nonnil d : unorm d <> Nil.
Proof.
- unfold unorm. now destruct nzhead.
+ unfold unorm. now destruct nzhead.
+Qed.
+
+Lemma unorm_D0 u : unorm (D0 u) = unorm u.
+Proof. reflexivity. Qed.
+
+Lemma unorm_iter_D0 n u : unorm (Nat.iter n D0 u) = unorm u.
+Proof. now induction n. Qed.
+
+Lemma nb_digits_unorm u :
+ u <> Nil -> nb_digits (unorm u) <= nb_digits u.
+Proof.
+ case u; clear u; [now simpl|intro u..]; [|now simpl..].
+ intros _; unfold unorm.
+ case_eq (nzhead (D0 u)); [|now intros u' <-; apply nb_digits_nzhead..].
+ intros _; apply le_n_S, Nat.le_0_l.
+Qed.
+
+Lemma del_head_nonnil n u :
+ n < nb_digits u -> del_head n u <> Nil.
+Proof.
+ now revert n; induction u; intro n;
+ [|case n; [|intro n'; simpl; intro H; apply IHu, lt_S_n]..].
+Qed.
+
+Lemma del_tail_nonnil n u :
+ n < nb_digits u -> del_tail n u <> Nil.
+Proof.
+ unfold del_tail.
+ rewrite <-nb_digits_rev.
+ generalize (rev u); clear u; intro u.
+ intros Hu H.
+ generalize (rev_nil_inv _ H); clear H.
+ now apply del_head_nonnil.
Qed.
Lemma nzhead_invol d : nzhead (nzhead d) = nzhead d.
Proof.
- now induction d.
+ now induction d.
+Qed.
+
+Lemma nztail_invol d : nztail (nztail d) = nztail d.
+Proof.
+ rewrite <-(rev_rev (nztail _)), <-(rev_rev (nztail d)), <-(rev_rev d).
+ now rewrite !rev_nztail_rev, nzhead_invol.
Qed.
Lemma unorm_invol d : unorm (unorm d) = unorm d.
Proof.
- unfold unorm.
- destruct (nzhead d) eqn:E; trivial.
- destruct (nzhead_nonzero _ _ E).
+ unfold unorm.
+ destruct (nzhead d) eqn:E; trivial.
+ destruct (nzhead_nonzero _ _ E).
Qed.
Lemma norm_invol d : norm (norm d) = norm d.
Proof.
- unfold norm.
- destruct d.
- - f_equal. apply unorm_invol.
- - destruct (nzhead d) eqn:E; auto.
- destruct (nzhead_nonzero _ _ E).
+ unfold norm.
+ destruct d.
+ - f_equal. apply unorm_invol.
+ - destruct (nzhead d) eqn:E; auto.
+ destruct (nzhead_nonzero _ _ E).
+Qed.
+
+Lemma nzhead_del_tail_nzhead_eq n u :
+ nzhead u = u ->
+ n < nb_digits u ->
+ nzhead (del_tail n u) = del_tail n u.
+Proof.
+ intros Hu Hn.
+ assert (Hhd : forall u,
+ nzhead u = u <-> match nth 0 u with D0 _ => False | _ => True end).
+ { clear n u Hu Hn; intro u.
+ case u; clear u; [|intro u..]; [now simpl| |now simpl..]; simpl.
+ split; [|now simpl].
+ apply nzhead_nonzero. }
+ assert (Hhd' : nth 0 (del_tail n u) = nth 0 u).
+ { rewrite <-(app_del_tail_head _ _ (le_Sn_le _ _ Hn)) at 2.
+ unfold app.
+ rewrite nth_revapp_l.
+ - rewrite <-(nth_revapp_l _ _ Nil).
+ + now fold (rev (rev (del_tail n u))); rewrite rev_rev.
+ + unfold del_tail; rewrite rev_rev.
+ rewrite nb_digits_del_head; rewrite nb_digits_rev.
+ * now rewrite <-Nat.lt_add_lt_sub_r.
+ * now apply Nat.lt_le_incl.
+ - unfold del_tail; rewrite rev_rev.
+ rewrite nb_digits_del_head; rewrite nb_digits_rev.
+ + now rewrite <-Nat.lt_add_lt_sub_r.
+ + now apply Nat.lt_le_incl. }
+ revert Hu; rewrite Hhd; intro Hu.
+ now rewrite Hhd, Hhd'.
+Qed.
+
+Lemma nzhead_del_tail_nzhead n u :
+ n < nb_digits (nzhead u) ->
+ nzhead (del_tail n (nzhead u)) = del_tail n (nzhead u).
+Proof. apply nzhead_del_tail_nzhead_eq, nzhead_invol. Qed.
+
+Lemma unorm_del_tail_unorm n u :
+ n < nb_digits (unorm u) ->
+ unorm (del_tail n (unorm u)) = del_tail n (unorm u).
+Proof.
+ case (uint_dec (nzhead u) Nil).
+ - unfold unorm; intros->; case n; [now simpl|]; intro n'.
+ now simpl; intro H; exfalso; generalize (lt_S_n _ _ H).
+ - unfold unorm.
+ set (m := match nzhead u with Nil => zero | _ => _ end).
+ intros H.
+ replace m with (nzhead u).
+ + intros H'.
+ rewrite (nzhead_del_tail_nzhead _ _ H').
+ now generalize (del_tail_nonnil _ _ H'); case del_tail.
+ + now unfold m; revert H; case nzhead.
+Qed.
+
+Lemma norm_del_tail_int_norm n d :
+ n < nb_digits (match norm d with Pos d | Neg d => d end) ->
+ norm (del_tail_int n (norm d)) = del_tail_int n (norm d).
+Proof.
+ case d; clear d; intros u; simpl.
+ - now intro H; simpl; rewrite unorm_del_tail_unorm.
+ - case (uint_dec (nzhead u) Nil); intro Hu.
+ + now rewrite Hu; case n; [|intros n' Hn'; generalize (lt_S_n _ _ Hn')].
+ + set (m := match nzhead u with Nil => Pos zero | _ => _ end).
+ replace m with (Neg (nzhead u)); [|now unfold m; revert Hu; case nzhead].
+ unfold del_tail_int.
+ clear m Hu.
+ simpl.
+ intro H; generalize (del_tail_nonnil _ _ H).
+ rewrite (nzhead_del_tail_nzhead _ _ H).
+ now case del_tail.
+Qed.
+
+Lemma nzhead_app_nzhead d d' :
+ nzhead (app (nzhead d) d') = nzhead (app d d').
+Proof.
+ unfold app.
+ rewrite <-(rev_nztail_rev d), rev_rev.
+ generalize (rev d); clear d; intro d.
+ generalize (nzhead_revapp_0 d d').
+ generalize (nzhead_revapp d d').
+ generalize (nzhead_revapp_0 (nztail d) d').
+ generalize (nzhead_revapp (nztail d) d').
+ rewrite nztail_invol.
+ now case nztail;
+ [intros _ H _ H'; rewrite (H eq_refl), (H' eq_refl)
+ |intros d'' H _ H' _; rewrite H; [rewrite H'|]..].
+Qed.
+
+Lemma unorm_app_unorm d d' :
+ unorm (app (unorm d) d') = unorm (app d d').
+Proof.
+ unfold unorm.
+ rewrite <-(nzhead_app_nzhead d d').
+ now case (nzhead d).
+Qed.
+
+Lemma norm_app_int_norm d d' :
+ unorm d' = zero ->
+ norm (app_int (norm d) d') = norm (app_int d d').
+Proof.
+ case d; clear d; intro d; simpl.
+ - now rewrite unorm_app_unorm.
+ - unfold app_int, app.
+ rewrite unorm_0; intro Hd'.
+ rewrite <-rev_nztail_rev.
+ generalize (nzhead_revapp (rev d) d').
+ generalize (nzhead_revapp_0 (rev d) d').
+ now case_eq (nztail (rev d));
+ [intros Hd'' H _; rewrite (H eq_refl); simpl;
+ unfold unorm; simpl; rewrite Hd'
+ |intros d'' Hd'' _ H; rewrite H; clear H; [|now simpl];
+ set (r := rev _);
+ set (m := match r with Nil => Pos zero | _ => _ end);
+ assert (H' : m = Neg r);
+ [now unfold m; case_eq r; unfold r;
+ [intro H''; generalize (rev_nil_inv _ H'')|..]
+ |rewrite H'; unfold r; clear m r H'];
+ unfold norm;
+ rewrite rev_rev, <-Hd'';
+ rewrite nzhead_revapp; rewrite nztail_invol; [|rewrite Hd'']..].
Qed.
diff --git a/theories/Numbers/DecimalN.v b/theories/Numbers/DecimalN.v
index ac6d9c663f..8bc5c38fb5 100644
--- a/theories/Numbers/DecimalN.v
+++ b/theories/Numbers/DecimalN.v
@@ -19,41 +19,41 @@ Module Unsigned.
Lemma of_to (n:N) : N.of_uint (N.to_uint n) = n.
Proof.
- destruct n.
- - reflexivity.
- - apply DecimalPos.Unsigned.of_to.
+ destruct n.
+ - reflexivity.
+ - apply DecimalPos.Unsigned.of_to.
Qed.
Lemma to_of (d:uint) : N.to_uint (N.of_uint d) = unorm d.
Proof.
- exact (DecimalPos.Unsigned.to_of d).
+ exact (DecimalPos.Unsigned.to_of d).
Qed.
Lemma to_uint_inj n n' : N.to_uint n = N.to_uint n' -> n = n'.
Proof.
- intros E. now rewrite <- (of_to n), <- (of_to n'), E.
+ intros E. now rewrite <- (of_to n), <- (of_to n'), E.
Qed.
Lemma to_uint_surj d : exists p, N.to_uint p = unorm d.
Proof.
- exists (N.of_uint d). apply to_of.
+ exists (N.of_uint d). apply to_of.
Qed.
Lemma of_uint_norm d : N.of_uint (unorm d) = N.of_uint d.
Proof.
- now induction d.
+ now induction d.
Qed.
Lemma of_inj d d' :
- N.of_uint d = N.of_uint d' -> unorm d = unorm d'.
+ N.of_uint d = N.of_uint d' -> unorm d = unorm d'.
Proof.
- intros. rewrite <- !to_of. now f_equal.
+ intros. rewrite <- !to_of. now f_equal.
Qed.
Lemma of_iff d d' : N.of_uint d = N.of_uint d' <-> unorm d = unorm d'.
Proof.
- split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
- apply of_uint_norm.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
Qed.
End Unsigned.
@@ -64,44 +64,44 @@ Module Signed.
Lemma of_to (n:N) : N.of_int (N.to_int n) = Some n.
Proof.
- unfold N.to_int, N.of_int, norm. f_equal.
- rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
+ unfold N.to_int, N.of_int, norm. f_equal.
+ rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
Qed.
Lemma to_of (d:int)(n:N) : N.of_int d = Some n -> N.to_int n = norm d.
Proof.
- unfold N.of_int.
- destruct (norm d) eqn:Hd; intros [= <-].
- unfold N.to_int. rewrite Unsigned.to_of. f_equal.
- revert Hd; destruct d; simpl.
- - intros [= <-]. apply unorm_invol.
- - destruct (nzhead d); now intros [= <-].
+ unfold N.of_int.
+ destruct (norm d) eqn:Hd; intros [= <-].
+ unfold N.to_int. rewrite Unsigned.to_of. f_equal.
+ revert Hd; destruct d; simpl.
+ - intros [= <-]. apply unorm_invol.
+ - destruct (nzhead d); now intros [= <-].
Qed.
Lemma to_int_inj n n' : N.to_int n = N.to_int n' -> n = n'.
Proof.
- intro E.
- assert (E' : Some n = Some n').
- { now rewrite <- (of_to n), <- (of_to n'), E. }
- now injection E'.
+ intro E.
+ assert (E' : Some n = Some n').
+ { now rewrite <- (of_to n), <- (of_to n'), E. }
+ now injection E'.
Qed.
Lemma to_int_pos_surj d : exists n, N.to_int n = norm (Pos d).
Proof.
- exists (N.of_uint d). unfold N.to_int. now rewrite Unsigned.to_of.
+ exists (N.of_uint d). unfold N.to_int. now rewrite Unsigned.to_of.
Qed.
Lemma of_int_norm d : N.of_int (norm d) = N.of_int d.
Proof.
- unfold N.of_int. now rewrite norm_invol.
+ unfold N.of_int. now rewrite norm_invol.
Qed.
Lemma of_inj_pos d d' :
N.of_int (Pos d) = N.of_int (Pos d') -> unorm d = unorm d'.
Proof.
- unfold N.of_int. simpl. intros [= H]. apply Unsigned.of_inj.
- change Pos.of_uint with N.of_uint in H.
- now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
+ unfold N.of_int. simpl. intros [= H]. apply Unsigned.of_inj.
+ change Pos.of_uint with N.of_uint in H.
+ now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
Qed.
End Signed.
diff --git a/theories/Numbers/DecimalNat.v b/theories/Numbers/DecimalNat.v
index e0ba8dc631..1962ac5d9d 100644
--- a/theories/Numbers/DecimalNat.v
+++ b/theories/Numbers/DecimalNat.v
@@ -20,25 +20,25 @@ Module Unsigned.
(** A few helper functions used during proofs *)
Definition hd d :=
- match d with
- | Nil => 0
- | D0 _ => 0
- | D1 _ => 1
- | D2 _ => 2
- | D3 _ => 3
- | D4 _ => 4
- | D5 _ => 5
- | D6 _ => 6
- | D7 _ => 7
- | D8 _ => 8
- | D9 _ => 9
-end.
+ match d with
+ | Nil => 0
+ | D0 _ => 0
+ | D1 _ => 1
+ | D2 _ => 2
+ | D3 _ => 3
+ | D4 _ => 4
+ | D5 _ => 5
+ | D6 _ => 6
+ | D7 _ => 7
+ | D8 _ => 8
+ | D9 _ => 9
+ end.
Definition tl d :=
- match d with
- | Nil => d
- | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d => d
-end.
+ match d with
+ | Nil => d
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d => d
+ end.
Fixpoint usize (d:uint) : nat :=
match d with
@@ -57,10 +57,10 @@ Fixpoint usize (d:uint) : nat :=
(** A direct version of [to_little_uint], not tail-recursive *)
Fixpoint to_lu n :=
- match n with
- | 0 => Decimal.zero
- | S n => Little.succ (to_lu n)
- end.
+ match n with
+ | 0 => Decimal.zero
+ | S n => Little.succ (to_lu n)
+ end.
(** A direct version of [of_little_uint] *)
Fixpoint of_lu (d:uint) : nat :=
@@ -82,174 +82,174 @@ Fixpoint of_lu (d:uint) : nat :=
Lemma to_lu_succ n : to_lu (S n) = Little.succ (to_lu n).
Proof.
- reflexivity.
+ reflexivity.
Qed.
Lemma to_little_uint_succ n d :
- Nat.to_little_uint n (Little.succ d) =
- Little.succ (Nat.to_little_uint n d).
+ Nat.to_little_uint n (Little.succ d) =
+ Little.succ (Nat.to_little_uint n d).
Proof.
- revert d; induction n; simpl; trivial.
+ revert d; induction n; simpl; trivial.
Qed.
Lemma to_lu_equiv n :
to_lu n = Nat.to_little_uint n zero.
Proof.
- induction n; simpl; trivial.
- now rewrite IHn, <- to_little_uint_succ.
+ induction n; simpl; trivial.
+ now rewrite IHn, <- to_little_uint_succ.
Qed.
Lemma to_uint_alt n :
- Nat.to_uint n = rev (to_lu n).
+ Nat.to_uint n = rev (to_lu n).
Proof.
- unfold Nat.to_uint. f_equal. symmetry. apply to_lu_equiv.
+ unfold Nat.to_uint. f_equal. symmetry. apply to_lu_equiv.
Qed.
(** Properties of [of_lu] *)
Lemma of_lu_eqn d :
- of_lu d = hd d + 10 * of_lu (tl d).
+ of_lu d = hd d + 10 * of_lu (tl d).
Proof.
- induction d; simpl; trivial.
+ induction d; simpl; trivial.
Qed.
Ltac simpl_of_lu :=
- match goal with
- | |- context [ of_lu (?f ?x) ] =>
- rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
- end.
+ match goal with
+ | |- context [ of_lu (?f ?x) ] =>
+ rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
+ end.
Lemma of_lu_succ d :
- of_lu (Little.succ d) = S (of_lu d).
+ of_lu (Little.succ d) = S (of_lu d).
Proof.
- induction d; trivial.
- simpl_of_lu. rewrite IHd. simpl_of_lu.
- now rewrite Nat.mul_succ_r, <- (Nat.add_comm 10).
+ induction d; trivial.
+ simpl_of_lu. rewrite IHd. simpl_of_lu.
+ now rewrite Nat.mul_succ_r, <- (Nat.add_comm 10).
Qed.
Lemma of_to_lu n :
- of_lu (to_lu n) = n.
+ of_lu (to_lu n) = n.
Proof.
- induction n; simpl; trivial. rewrite of_lu_succ. now f_equal.
+ induction n; simpl; trivial. rewrite of_lu_succ. now f_equal.
Qed.
Lemma of_lu_revapp d d' :
-of_lu (revapp d d') =
- of_lu (rev d) + of_lu d' * 10^usize d.
+ of_lu (revapp d d') =
+ of_lu (rev d) + of_lu d' * 10^usize d.
Proof.
- revert d'.
- induction d; intro d'; simpl usize;
- [ simpl; now rewrite Nat.mul_1_r | .. ];
- unfold rev; simpl revapp; rewrite 2 IHd;
- rewrite <- Nat.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
- rewrite Nat.pow_succ_r'; ring.
+ revert d'.
+ induction d; intro d'; simpl usize;
+ [ simpl; now rewrite Nat.mul_1_r | .. ];
+ unfold rev; simpl revapp; rewrite 2 IHd;
+ rewrite <- Nat.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
+ rewrite Nat.pow_succ_r'; ring.
Qed.
Lemma of_uint_acc_spec n d :
- Nat.of_uint_acc d n = of_lu (rev d) + n * 10^usize d.
+ Nat.of_uint_acc d n = of_lu (rev d) + n * 10^usize d.
Proof.
- revert n. induction d; intros;
- simpl Nat.of_uint_acc; rewrite ?Nat.tail_mul_spec, ?IHd;
- simpl rev; simpl usize; rewrite ?Nat.pow_succ_r';
- [ simpl; now rewrite Nat.mul_1_r | .. ];
- unfold rev at 2; simpl revapp; rewrite of_lu_revapp;
- simpl of_lu; ring.
+ revert n. induction d; intros;
+ simpl Nat.of_uint_acc; rewrite ?Nat.tail_mul_spec, ?IHd;
+ simpl rev; simpl usize; rewrite ?Nat.pow_succ_r';
+ [ simpl; now rewrite Nat.mul_1_r | .. ];
+ unfold rev at 2; simpl revapp; rewrite of_lu_revapp;
+ simpl of_lu; ring.
Qed.
Lemma of_uint_alt d : Nat.of_uint d = of_lu (rev d).
Proof.
- unfold Nat.of_uint. now rewrite of_uint_acc_spec.
+ unfold Nat.of_uint. now rewrite of_uint_acc_spec.
Qed.
(** First main bijection result *)
Lemma of_to (n:nat) : Nat.of_uint (Nat.to_uint n) = n.
Proof.
- rewrite to_uint_alt, of_uint_alt, rev_rev. apply of_to_lu.
+ rewrite to_uint_alt, of_uint_alt, rev_rev. apply of_to_lu.
Qed.
(** The other direction *)
Lemma to_lu_tenfold n : n<>0 ->
- to_lu (10 * n) = D0 (to_lu n).
+ to_lu (10 * n) = D0 (to_lu n).
Proof.
- induction n.
- - simpl. now destruct 1.
- - intros _.
- destruct (Nat.eq_dec n 0) as [->|H]; simpl; trivial.
- rewrite !Nat.add_succ_r.
- simpl in *. rewrite (IHn H). now destruct (to_lu n).
+ induction n.
+ - simpl. now destruct 1.
+ - intros _.
+ destruct (Nat.eq_dec n 0) as [->|H]; simpl; trivial.
+ rewrite !Nat.add_succ_r.
+ simpl in *. rewrite (IHn H). now destruct (to_lu n).
Qed.
Lemma of_lu_0 d : of_lu d = 0 <-> nztail d = Nil.
Proof.
- induction d; try simpl_of_lu; try easy.
- rewrite Nat.add_0_l.
- split; intros H.
- - apply Nat.eq_mul_0_r in H; auto.
- rewrite IHd in H. simpl. now rewrite H.
- - simpl in H. destruct (nztail d); try discriminate.
- now destruct IHd as [_ ->].
+ induction d; try simpl_of_lu; try easy.
+ rewrite Nat.add_0_l.
+ split; intros H.
+ - apply Nat.eq_mul_0_r in H; auto.
+ rewrite IHd in H. simpl. now rewrite H.
+ - simpl in H. destruct (nztail d); try discriminate.
+ now destruct IHd as [_ ->].
Qed.
Lemma to_of_lu_tenfold d :
- to_lu (of_lu d) = lnorm d ->
- to_lu (10 * of_lu d) = lnorm (D0 d).
+ to_lu (of_lu d) = lnorm d ->
+ to_lu (10 * of_lu d) = lnorm (D0 d).
Proof.
- intro IH.
- destruct (Nat.eq_dec (of_lu d) 0) as [H|H].
- - rewrite H. simpl. rewrite of_lu_0 in H.
- unfold lnorm. simpl. now rewrite H.
- - rewrite (to_lu_tenfold _ H), IH.
- rewrite of_lu_0 in H.
- unfold lnorm. simpl. now destruct (nztail d).
+ intro IH.
+ destruct (Nat.eq_dec (of_lu d) 0) as [H|H].
+ - rewrite H. simpl. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now rewrite H.
+ - rewrite (to_lu_tenfold _ H), IH.
+ rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now destruct (nztail d).
Qed.
Lemma to_of_lu d : to_lu (of_lu d) = lnorm d.
Proof.
- induction d; [ reflexivity | .. ];
- simpl_of_lu;
- rewrite ?Nat.add_succ_l, Nat.add_0_l, ?to_lu_succ, to_of_lu_tenfold
- by assumption;
- unfold lnorm; simpl; now destruct nztail.
+ induction d; [ reflexivity | .. ];
+ simpl_of_lu;
+ rewrite ?Nat.add_succ_l, Nat.add_0_l, ?to_lu_succ, to_of_lu_tenfold
+ by assumption;
+ unfold lnorm; simpl; now destruct nztail.
Qed.
(** Second bijection result *)
Lemma to_of (d:uint) : Nat.to_uint (Nat.of_uint d) = unorm d.
Proof.
- rewrite to_uint_alt, of_uint_alt, to_of_lu.
- apply rev_lnorm_rev.
+ rewrite to_uint_alt, of_uint_alt, to_of_lu.
+ apply rev_lnorm_rev.
Qed.
(** Some consequences *)
Lemma to_uint_inj n n' : Nat.to_uint n = Nat.to_uint n' -> n = n'.
Proof.
- intro EQ.
- now rewrite <- (of_to n), <- (of_to n'), EQ.
+ intro EQ.
+ now rewrite <- (of_to n), <- (of_to n'), EQ.
Qed.
Lemma to_uint_surj d : exists n, Nat.to_uint n = unorm d.
Proof.
- exists (Nat.of_uint d). apply to_of.
+ exists (Nat.of_uint d). apply to_of.
Qed.
Lemma of_uint_norm d : Nat.of_uint (unorm d) = Nat.of_uint d.
Proof.
- unfold Nat.of_uint. now induction d.
+ unfold Nat.of_uint. now induction d.
Qed.
Lemma of_inj d d' :
- Nat.of_uint d = Nat.of_uint d' -> unorm d = unorm d'.
+ Nat.of_uint d = Nat.of_uint d' -> unorm d = unorm d'.
Proof.
- intros. rewrite <- !to_of. now f_equal.
+ intros. rewrite <- !to_of. now f_equal.
Qed.
Lemma of_iff d d' : Nat.of_uint d = Nat.of_uint d' <-> unorm d = unorm d'.
Proof.
- split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
- apply of_uint_norm.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
Qed.
End Unsigned.
@@ -260,43 +260,43 @@ Module Signed.
Lemma of_to (n:nat) : Nat.of_int (Nat.to_int n) = Some n.
Proof.
- unfold Nat.to_int, Nat.of_int, norm. f_equal.
- rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
+ unfold Nat.to_int, Nat.of_int, norm. f_equal.
+ rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
Qed.
Lemma to_of (d:int)(n:nat) : Nat.of_int d = Some n -> Nat.to_int n = norm d.
Proof.
- unfold Nat.of_int.
- destruct (norm d) eqn:Hd; intros [= <-].
- unfold Nat.to_int. rewrite Unsigned.to_of. f_equal.
- revert Hd; destruct d; simpl.
- - intros [= <-]. apply unorm_invol.
- - destruct (nzhead d); now intros [= <-].
+ unfold Nat.of_int.
+ destruct (norm d) eqn:Hd; intros [= <-].
+ unfold Nat.to_int. rewrite Unsigned.to_of. f_equal.
+ revert Hd; destruct d; simpl.
+ - intros [= <-]. apply unorm_invol.
+ - destruct (nzhead d); now intros [= <-].
Qed.
Lemma to_int_inj n n' : Nat.to_int n = Nat.to_int n' -> n = n'.
Proof.
- intro E.
- assert (E' : Some n = Some n').
- { now rewrite <- (of_to n), <- (of_to n'), E. }
- now injection E'.
+ intro E.
+ assert (E' : Some n = Some n').
+ { now rewrite <- (of_to n), <- (of_to n'), E. }
+ now injection E'.
Qed.
Lemma to_int_pos_surj d : exists n, Nat.to_int n = norm (Pos d).
Proof.
- exists (Nat.of_uint d). unfold Nat.to_int. now rewrite Unsigned.to_of.
+ exists (Nat.of_uint d). unfold Nat.to_int. now rewrite Unsigned.to_of.
Qed.
Lemma of_int_norm d : Nat.of_int (norm d) = Nat.of_int d.
Proof.
- unfold Nat.of_int. now rewrite norm_invol.
+ unfold Nat.of_int. now rewrite norm_invol.
Qed.
Lemma of_inj_pos d d' :
- Nat.of_int (Pos d) = Nat.of_int (Pos d') -> unorm d = unorm d'.
+ Nat.of_int (Pos d) = Nat.of_int (Pos d') -> unorm d = unorm d'.
Proof.
- unfold Nat.of_int. simpl. intros [= H]. apply Unsigned.of_inj.
- now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
+ unfold Nat.of_int. simpl. intros [= H]. apply Unsigned.of_inj.
+ now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
Qed.
End Signed.
diff --git a/theories/Numbers/DecimalPos.v b/theories/Numbers/DecimalPos.v
index 803688f476..5611329b12 100644
--- a/theories/Numbers/DecimalPos.v
+++ b/theories/Numbers/DecimalPos.v
@@ -36,37 +36,37 @@ Fixpoint of_lu (d:uint) : N :=
end.
Definition hd d :=
-match d with
- | Nil => 0
- | D0 _ => 0
- | D1 _ => 1
- | D2 _ => 2
- | D3 _ => 3
- | D4 _ => 4
- | D5 _ => 5
- | D6 _ => 6
- | D7 _ => 7
- | D8 _ => 8
- | D9 _ => 9
-end.
+ match d with
+ | Nil => 0
+ | D0 _ => 0
+ | D1 _ => 1
+ | D2 _ => 2
+ | D3 _ => 3
+ | D4 _ => 4
+ | D5 _ => 5
+ | D6 _ => 6
+ | D7 _ => 7
+ | D8 _ => 8
+ | D9 _ => 9
+ end.
Definition tl d :=
- match d with
- | Nil => d
- | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d => d
-end.
+ match d with
+ | Nil => d
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d => d
+ end.
Lemma of_lu_eqn d :
- of_lu d = hd d + 10 * (of_lu (tl d)).
+ of_lu d = hd d + 10 * (of_lu (tl d)).
Proof.
- induction d; simpl; trivial.
+ induction d; simpl; trivial.
Qed.
Ltac simpl_of_lu :=
- match goal with
- | |- context [ of_lu (?f ?x) ] =>
- rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
- end.
+ match goal with
+ | |- context [ of_lu (?f ?x) ] =>
+ rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
+ end.
Fixpoint usize (d:uint) : N :=
match d with
@@ -84,89 +84,89 @@ Fixpoint usize (d:uint) : N :=
end.
Lemma of_lu_revapp d d' :
- of_lu (revapp d d') =
- of_lu (rev d) + of_lu d' * 10^usize d.
+ of_lu (revapp d d') =
+ of_lu (rev d) + of_lu d' * 10^usize d.
Proof.
- revert d'.
- induction d; simpl; intro d'; [ now rewrite N.mul_1_r | .. ];
- unfold rev; simpl revapp; rewrite 2 IHd;
- rewrite <- N.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
- rewrite N.pow_succ_r'; ring.
+ revert d'.
+ induction d; simpl; intro d'; [ now rewrite N.mul_1_r | .. ];
+ unfold rev; simpl revapp; rewrite 2 IHd;
+ rewrite <- N.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
+ rewrite N.pow_succ_r'; ring.
Qed.
Definition Nadd n p :=
- match n with
- | N0 => p
- | Npos p0 => (p0+p)%positive
- end.
+ match n with
+ | N0 => p
+ | Npos p0 => (p0+p)%positive
+ end.
Lemma Nadd_simpl n p q : Npos (Nadd n (p * q)) = n + Npos p * Npos q.
Proof.
- now destruct n.
+ now destruct n.
Qed.
Lemma of_uint_acc_eqn d acc : d<>Nil ->
- Pos.of_uint_acc d acc = Pos.of_uint_acc (tl d) (Nadd (hd d) (10*acc)).
+ Pos.of_uint_acc d acc = Pos.of_uint_acc (tl d) (Nadd (hd d) (10*acc)).
Proof.
- destruct d; simpl; trivial. now destruct 1.
+ destruct d; simpl; trivial. now destruct 1.
Qed.
Lemma of_uint_acc_rev d acc :
- Npos (Pos.of_uint_acc d acc) =
- of_lu (rev d) + (Npos acc) * 10^usize d.
+ Npos (Pos.of_uint_acc d acc) =
+ of_lu (rev d) + (Npos acc) * 10^usize d.
Proof.
- revert acc.
- induction d; intros; simpl usize;
- [ simpl; now rewrite Pos.mul_1_r | .. ];
- rewrite N.pow_succ_r';
- unfold rev; simpl revapp; try rewrite of_lu_revapp; simpl of_lu;
- rewrite of_uint_acc_eqn by easy; simpl tl; simpl hd;
- rewrite IHd, Nadd_simpl; ring.
+ revert acc.
+ induction d; intros; simpl usize;
+ [ simpl; now rewrite Pos.mul_1_r | .. ];
+ rewrite N.pow_succ_r';
+ unfold rev; simpl revapp; try rewrite of_lu_revapp; simpl of_lu;
+ rewrite of_uint_acc_eqn by easy; simpl tl; simpl hd;
+ rewrite IHd, Nadd_simpl; ring.
Qed.
Lemma of_uint_alt d : Pos.of_uint d = of_lu (rev d).
Proof.
- induction d; simpl; trivial; unfold rev; simpl revapp;
- rewrite of_lu_revapp; simpl of_lu; try apply of_uint_acc_rev.
- rewrite IHd. ring.
+ induction d; simpl; trivial; unfold rev; simpl revapp;
+ rewrite of_lu_revapp; simpl of_lu; try apply of_uint_acc_rev.
+ rewrite IHd. ring.
Qed.
Lemma of_lu_rev d : Pos.of_uint (rev d) = of_lu d.
Proof.
- rewrite of_uint_alt. now rewrite rev_rev.
+ rewrite of_uint_alt. now rewrite rev_rev.
Qed.
Lemma of_lu_double_gen d :
of_lu (Little.double d) = N.double (of_lu d) /\
of_lu (Little.succ_double d) = N.succ_double (of_lu d).
Proof.
- rewrite N.double_spec, N.succ_double_spec.
- induction d; try destruct IHd as (IH1,IH2);
- simpl Little.double; simpl Little.succ_double;
- repeat (simpl_of_lu; rewrite ?IH1, ?IH2); split; reflexivity || ring.
+ rewrite N.double_spec, N.succ_double_spec.
+ induction d; try destruct IHd as (IH1,IH2);
+ simpl Little.double; simpl Little.succ_double;
+ repeat (simpl_of_lu; rewrite ?IH1, ?IH2); split; reflexivity || ring.
Qed.
Lemma of_lu_double d :
of_lu (Little.double d) = N.double (of_lu d).
Proof.
- apply of_lu_double_gen.
+ apply of_lu_double_gen.
Qed.
Lemma of_lu_succ_double d :
of_lu (Little.succ_double d) = N.succ_double (of_lu d).
Proof.
- apply of_lu_double_gen.
+ apply of_lu_double_gen.
Qed.
(** First bijection result *)
Lemma of_to (p:positive) : Pos.of_uint (Pos.to_uint p) = Npos p.
Proof.
- unfold Pos.to_uint.
- rewrite of_lu_rev.
- induction p; simpl; trivial.
- - now rewrite of_lu_succ_double, IHp.
- - now rewrite of_lu_double, IHp.
+ unfold Pos.to_uint.
+ rewrite of_lu_rev.
+ induction p; simpl; trivial.
+ - now rewrite of_lu_succ_double, IHp.
+ - now rewrite of_lu_double, IHp.
Qed.
(** The other direction *)
@@ -180,149 +180,164 @@ Definition to_lu n :=
Lemma succ_double_alt d :
Little.succ_double d = Little.succ (Little.double d).
Proof.
- now induction d.
+ now induction d.
Qed.
Lemma double_succ d :
Little.double (Little.succ d) =
Little.succ (Little.succ_double d).
Proof.
- induction d; simpl; f_equal; auto using succ_double_alt.
+ induction d; simpl; f_equal; auto using succ_double_alt.
Qed.
Lemma to_lu_succ n :
- to_lu (N.succ n) = Little.succ (to_lu n).
+ to_lu (N.succ n) = Little.succ (to_lu n).
Proof.
- destruct n; simpl; trivial.
- induction p; simpl; rewrite ?IHp;
- auto using succ_double_alt, double_succ.
+ destruct n; simpl; trivial.
+ induction p; simpl; rewrite ?IHp;
+ auto using succ_double_alt, double_succ.
Qed.
Lemma nat_iter_S n {A} (f:A->A) i :
- Nat.iter (S n) f i = f (Nat.iter n f i).
+ Nat.iter (S n) f i = f (Nat.iter n f i).
Proof.
- reflexivity.
+ reflexivity.
Qed.
Lemma nat_iter_0 {A} (f:A->A) i : Nat.iter 0 f i = i.
Proof.
- reflexivity.
+ reflexivity.
Qed.
Lemma to_ldec_tenfold p :
to_lu (10 * Npos p) = D0 (to_lu (Npos p)).
Proof.
- induction p using Pos.peano_rect.
- - trivial.
- - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
- rewrite N.mul_succ_r.
- change 10 at 2 with (Nat.iter 10%nat N.succ 0).
- rewrite ?nat_iter_S, nat_iter_0.
- rewrite !N.add_succ_r, N.add_0_r, !to_lu_succ, IHp.
- destruct (to_lu (N.pos p)); simpl; auto.
+ induction p using Pos.peano_rect.
+ - trivial.
+ - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
+ rewrite N.mul_succ_r.
+ change 10 at 2 with (Nat.iter 10%nat N.succ 0).
+ rewrite ?nat_iter_S, nat_iter_0.
+ rewrite !N.add_succ_r, N.add_0_r, !to_lu_succ, IHp.
+ destruct (to_lu (N.pos p)); simpl; auto.
Qed.
Lemma of_lu_0 d : of_lu d = 0 <-> nztail d = Nil.
Proof.
- induction d; try simpl_of_lu; split; trivial; try discriminate;
- try (intros H; now apply N.eq_add_0 in H).
- - rewrite N.add_0_l. intros H.
- apply N.eq_mul_0_r in H; [|easy]. rewrite IHd in H.
- simpl. now rewrite H.
- - simpl. destruct (nztail d); try discriminate.
- now destruct IHd as [_ ->].
-Qed.
+ induction d; try simpl_of_lu; split; trivial; try discriminate;
+ try (intros H; now apply N.eq_add_0 in H).
+ - rewrite N.add_0_l. intros H.
+ apply N.eq_mul_0_r in H; [|easy]. rewrite IHd in H.
+ simpl. now rewrite H.
+ - simpl. destruct (nztail d); try discriminate.
+ now destruct IHd as [_ ->].
+ Qed.
Lemma to_of_lu_tenfold d :
- to_lu (of_lu d) = lnorm d ->
- to_lu (10 * of_lu d) = lnorm (D0 d).
+ to_lu (of_lu d) = lnorm d ->
+ to_lu (10 * of_lu d) = lnorm (D0 d).
Proof.
- intro IH.
- destruct (N.eq_dec (of_lu d) 0) as [H|H].
- - rewrite H. simpl. rewrite of_lu_0 in H.
- unfold lnorm. simpl. now rewrite H.
- - destruct (of_lu d) eqn:Eq; [easy| ].
- rewrite to_ldec_tenfold; auto. rewrite IH.
- rewrite <- Eq in H. rewrite of_lu_0 in H.
- unfold lnorm. simpl. now destruct (nztail d).
+ intro IH.
+ destruct (N.eq_dec (of_lu d) 0) as [H|H].
+ - rewrite H. simpl. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now rewrite H.
+ - destruct (of_lu d) eqn:Eq; [easy| ].
+ rewrite to_ldec_tenfold; auto. rewrite IH.
+ rewrite <- Eq in H. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now destruct (nztail d).
Qed.
Lemma Nadd_alt n m : n + m = Nat.iter (N.to_nat n) N.succ m.
Proof.
- destruct n. trivial.
- induction p using Pos.peano_rect.
- - now rewrite N.add_1_l.
- - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
- now rewrite N.add_succ_l, IHp, N2Nat.inj_succ.
+ destruct n. trivial.
+ induction p using Pos.peano_rect.
+ - now rewrite N.add_1_l.
+ - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
+ now rewrite N.add_succ_l, IHp, N2Nat.inj_succ.
Qed.
Ltac simpl_to_nat := simpl N.to_nat; unfold Pos.to_nat; simpl Pos.iter_op.
Lemma to_of_lu d : to_lu (of_lu d) = lnorm d.
Proof.
- induction d; [reflexivity|..];
- simpl_of_lu; rewrite Nadd_alt; simpl_to_nat;
- rewrite ?nat_iter_S, nat_iter_0, ?to_lu_succ, to_of_lu_tenfold by assumption;
- unfold lnorm; simpl; destruct nztail; auto.
+ induction d; [reflexivity|..];
+ simpl_of_lu; rewrite Nadd_alt; simpl_to_nat;
+ rewrite ?nat_iter_S, nat_iter_0, ?to_lu_succ, to_of_lu_tenfold by assumption;
+ unfold lnorm; simpl; destruct nztail; auto.
Qed.
(** Second bijection result *)
Lemma to_of (d:uint) : N.to_uint (Pos.of_uint d) = unorm d.
Proof.
- rewrite of_uint_alt.
- unfold N.to_uint, Pos.to_uint.
- destruct (of_lu (rev d)) eqn:H.
- - rewrite of_lu_0 in H. rewrite <- rev_lnorm_rev.
- unfold lnorm. now rewrite H.
- - change (Pos.to_little_uint p) with (to_lu (N.pos p)).
- rewrite <- H. rewrite to_of_lu. apply rev_lnorm_rev.
+ rewrite of_uint_alt.
+ unfold N.to_uint, Pos.to_uint.
+ destruct (of_lu (rev d)) eqn:H.
+ - rewrite of_lu_0 in H. rewrite <- rev_lnorm_rev.
+ unfold lnorm. now rewrite H.
+ - change (Pos.to_little_uint p) with (to_lu (N.pos p)).
+ rewrite <- H. rewrite to_of_lu. apply rev_lnorm_rev.
Qed.
(** Some consequences *)
Lemma to_uint_nonzero p : Pos.to_uint p <> zero.
Proof.
- intro E. generalize (of_to p). now rewrite E.
+ intro E. generalize (of_to p). now rewrite E.
Qed.
Lemma to_uint_nonnil p : Pos.to_uint p <> Nil.
Proof.
- intros E. generalize (of_to p). now rewrite E.
+ intros E. generalize (of_to p). now rewrite E.
Qed.
Lemma to_uint_inj p p' : Pos.to_uint p = Pos.to_uint p' -> p = p'.
Proof.
- intro E.
- assert (E' : N.pos p = N.pos p').
- { now rewrite <- (of_to p), <- (of_to p'), E. }
- now injection E'.
+ intro E.
+ assert (E' : N.pos p = N.pos p').
+ { now rewrite <- (of_to p), <- (of_to p'), E. }
+ now injection E'.
Qed.
Lemma to_uint_pos_surj d :
- unorm d<>zero -> exists p, Pos.to_uint p = unorm d.
+ unorm d<>zero -> exists p, Pos.to_uint p = unorm d.
Proof.
- intros.
- destruct (Pos.of_uint d) eqn:E.
- - destruct H. generalize (to_of d). now rewrite E.
- - exists p. generalize (to_of d). now rewrite E.
+ intros.
+ destruct (Pos.of_uint d) eqn:E.
+ - destruct H. generalize (to_of d). now rewrite E.
+ - exists p. generalize (to_of d). now rewrite E.
Qed.
Lemma of_uint_norm d : Pos.of_uint (unorm d) = Pos.of_uint d.
Proof.
- now induction d.
+ now induction d.
Qed.
Lemma of_inj d d' :
- Pos.of_uint d = Pos.of_uint d' -> unorm d = unorm d'.
+ Pos.of_uint d = Pos.of_uint d' -> unorm d = unorm d'.
Proof.
- intros. rewrite <- !to_of. now f_equal.
+ intros. rewrite <- !to_of. now f_equal.
Qed.
Lemma of_iff d d' : Pos.of_uint d = Pos.of_uint d' <-> unorm d = unorm d'.
Proof.
- split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
- apply of_uint_norm.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
+Qed.
+
+Lemma nztail_to_uint p :
+ let (h, n) := Decimal.nztail (Pos.to_uint p) in
+ Npos p = Pos.of_uint h * 10^(N.of_nat n).
+Proof.
+ rewrite <-(of_to p), <-(rev_rev (Pos.to_uint p)), of_lu_rev.
+ unfold Decimal.nztail.
+ rewrite rev_rev.
+ induction (rev (Pos.to_uint p)); [reflexivity| |
+ now simpl N.of_nat; simpl N.pow; rewrite N.mul_1_r, of_lu_rev..].
+ revert IHu.
+ set (t := _ u); case t; clear t; intros u0 n H.
+ rewrite of_lu_eqn; unfold hd, tl.
+ rewrite N.add_0_l, H, Nat2N.inj_succ, N.pow_succ_r'; ring.
Qed.
End Unsigned.
@@ -333,51 +348,51 @@ Module Signed.
Lemma of_to (p:positive) : Pos.of_int (Pos.to_int p) = Some p.
Proof.
- unfold Pos.to_int, Pos.of_int, norm.
- now rewrite Unsigned.of_to.
+ unfold Pos.to_int, Pos.of_int, norm.
+ now rewrite Unsigned.of_to.
Qed.
Lemma to_of (d:int)(p:positive) :
- Pos.of_int d = Some p -> Pos.to_int p = norm d.
+ Pos.of_int d = Some p -> Pos.to_int p = norm d.
Proof.
- unfold Pos.of_int.
- destruct d; [ | intros [=]].
- simpl norm. rewrite <- Unsigned.to_of.
- destruct (Pos.of_uint d); now intros [= <-].
+ unfold Pos.of_int.
+ destruct d; [ | intros [=]].
+ simpl norm. rewrite <- Unsigned.to_of.
+ destruct (Pos.of_uint d); now intros [= <-].
Qed.
Lemma to_int_inj p p' : Pos.to_int p = Pos.to_int p' -> p = p'.
Proof.
- intro E.
- assert (E' : Some p = Some p').
- { now rewrite <- (of_to p), <- (of_to p'), E. }
- now injection E'.
+ intro E.
+ assert (E' : Some p = Some p').
+ { now rewrite <- (of_to p), <- (of_to p'), E. }
+ now injection E'.
Qed.
Lemma to_int_pos_surj d :
- unorm d <> zero -> exists p, Pos.to_int p = norm (Pos d).
+ unorm d <> zero -> exists p, Pos.to_int p = norm (Pos d).
Proof.
- simpl. unfold Pos.to_int. intros H.
- destruct (Unsigned.to_uint_pos_surj d H) as (p,Hp).
- exists p. now f_equal.
+ simpl. unfold Pos.to_int. intros H.
+ destruct (Unsigned.to_uint_pos_surj d H) as (p,Hp).
+ exists p. now f_equal.
Qed.
Lemma of_int_norm d : Pos.of_int (norm d) = Pos.of_int d.
Proof.
- unfold Pos.of_int.
- destruct d.
- - simpl. now rewrite Unsigned.of_uint_norm.
- - simpl. now destruct (nzhead d) eqn:H.
+ unfold Pos.of_int.
+ destruct d.
+ - simpl. now rewrite Unsigned.of_uint_norm.
+ - simpl. now destruct (nzhead d) eqn:H.
Qed.
Lemma of_inj_pos d d' :
- Pos.of_int (Pos d) = Pos.of_int (Pos d') -> unorm d = unorm d'.
+ Pos.of_int (Pos d) = Pos.of_int (Pos d') -> unorm d = unorm d'.
Proof.
- unfold Pos.of_int.
- destruct (Pos.of_uint d) eqn:Hd, (Pos.of_uint d') eqn:Hd';
- intros [=].
- - apply Unsigned.of_inj; now rewrite Hd, Hd'.
- - apply Unsigned.of_inj; rewrite Hd, Hd'; now f_equal.
+ unfold Pos.of_int.
+ destruct (Pos.of_uint d) eqn:Hd, (Pos.of_uint d') eqn:Hd';
+ intros [=].
+ - apply Unsigned.of_inj; now rewrite Hd, Hd'.
+ - apply Unsigned.of_inj; rewrite Hd, Hd'; now f_equal.
Qed.
End Signed.
diff --git a/theories/Numbers/DecimalQ.v b/theories/Numbers/DecimalQ.v
new file mode 100644
index 0000000000..c51cced024
--- /dev/null
+++ b/theories/Numbers/DecimalQ.v
@@ -0,0 +1,561 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * DecimalQ
+
+ Proofs that conversions between decimal numbers and [Q]
+ are bijections. *)
+
+Require Import Decimal DecimalFacts DecimalPos DecimalN DecimalZ QArith.
+
+Lemma of_to (q:Q) : forall d, to_decimal q = Some d -> of_decimal d = q.
+Proof.
+ cut (match to_decimal q with None => True | Some d => of_decimal d = q end).
+ { now case to_decimal; [intros d <- d' Hd'; injection Hd'; intros ->|]. }
+ destruct q as (num, den).
+ unfold to_decimal; simpl.
+ generalize (DecimalPos.Unsigned.nztail_to_uint den).
+ case Decimal.nztail; intros u n.
+ case u; clear u; [intros; exact I|intros; exact I|intro u|intros; exact I..].
+ case u; clear u; [|intros; exact I..].
+ unfold Pos.of_uint, Pos.of_uint_acc; rewrite N.mul_1_l.
+ case n.
+ - unfold of_decimal, app_int, app, Z.to_int; simpl.
+ intro H; inversion H as (H1); clear H H1.
+ case num; [reflexivity|intro pnum; fold (rev (rev (Pos.to_uint pnum)))..].
+ + rewrite rev_rev; simpl.
+ now unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to.
+ + rewrite rev_rev; simpl.
+ now unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to.
+ - clear n; intros n H.
+ injection H; clear H; intros ->.
+ case Nat.ltb.
+ + unfold of_decimal.
+ rewrite of_to.
+ apply f_equal2; [|now simpl].
+ unfold app_int, app, Z.to_int; simpl.
+ now case num;
+ [|intro pnum; fold (rev (rev (Pos.to_uint pnum)));
+ rewrite rev_rev; unfold Z.of_int, Z.of_uint;
+ rewrite DecimalPos.Unsigned.of_to..].
+ + unfold of_decimal; case Nat.ltb_spec; intro Hn; simpl.
+ * rewrite nb_digits_del_head; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply le_Sn_le].
+ rewrite Z.sub_sub_distr, Z.sub_diag; simpl.
+ rewrite <-(of_to num) at 4.
+ now revert Hn; case Z.to_int; clear num; intros pnum Hn; simpl;
+ (rewrite app_del_tail_head; [|now apply le_Sn_le]).
+ * revert Hn.
+ set (anum := match Z.to_int num with Pos i => i | _ => _ end).
+ intro Hn.
+ assert (H : exists l, nb_digits anum = S l).
+ { exists (Nat.pred (nb_digits anum)); apply S_pred_pos.
+ now unfold anum; case num;
+ [apply Nat.lt_0_1|
+ intro pnum; apply nb_digits_pos, Unsigned.to_uint_nonnil..]. }
+ destruct H as (l, Hl); rewrite Hl.
+ assert (H : forall n d, (nb_digits (Nat.iter n D0 d) = n + nb_digits d)%nat).
+ { now intros n'; induction n'; intro d; [|simpl; rewrite IHn']. }
+ rewrite H, Hl.
+ rewrite Nat.add_succ_r, Nat.sub_add; [|now apply le_S_n; rewrite <-Hl].
+ assert (H' : forall n d, Pos.of_uint (Nat.iter n D0 d) = Pos.of_uint d).
+ { now intro n'; induction n'; intro d; [|simpl; rewrite IHn']. }
+ now unfold anum; case num; simpl; [|intro pnum..];
+ unfold app, Z.of_uint; simpl;
+ rewrite H', ?DecimalPos.Unsigned.of_to.
+Qed.
+
+(* normalize without fractional part, for instance norme 12.3e-1 is 123e-2 *)
+Definition dnorme (d:decimal) : decimal :=
+ let '(i, f, e) :=
+ match d with
+ | Decimal i f => (i, f, Pos Nil)
+ | DecimalExp i f e => (i, f, e)
+ end in
+ let i := norm (app_int i f) in
+ let e := norm (Z.to_int (Z.of_int e - Z.of_nat (nb_digits f))) in
+ match e with
+ | Pos zero => Decimal i Nil
+ | _ => DecimalExp i Nil e
+ end.
+
+(* normalize without exponent part, for instance norme 12.3e-1 is 1.23 *)
+Definition dnormf (d:decimal) : decimal :=
+ match dnorme d with
+ | Decimal i _ => Decimal i Nil
+ | DecimalExp i _ e =>
+ match Z.of_int e with
+ | Z0 => Decimal i Nil
+ | Zpos e => Decimal (norm (app_int i (Pos.iter D0 Nil e))) Nil
+ | Zneg e =>
+ let ne := Pos.to_nat e in
+ let ai := match i with Pos d | Neg d => d end in
+ let ni := nb_digits ai in
+ if ne <? ni then
+ Decimal (del_tail_int ne i) (del_head (ni - ne) ai)
+ else
+ let z := match i with Pos _ => Pos zero | Neg _ => Neg zero end in
+ Decimal z (Nat.iter (ne - ni) D0 ai)
+ end
+ end.
+
+Lemma dnorme_spec d :
+ match dnorme d with
+ | Decimal i Nil => i = norm i
+ | DecimalExp i Nil e => i = norm i /\ e = norm e /\ e <> Pos zero
+ | _ => False
+ end.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold dnorme; simpl.
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); [intros->|intro Hne'].
+ + now rewrite norm_invol.
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ replace m with r; [now unfold r; rewrite !norm_invol|].
+ unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..].
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); [intros->|intro Hne'].
+ + now rewrite norm_invol.
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ replace m with r; [now unfold r; rewrite !norm_invol|].
+ unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..].
+Qed.
+
+Lemma dnormf_spec d :
+ match dnormf d with
+ | Decimal i f => i = Neg zero \/ i = norm i
+ | _ => False
+ end.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold dnormf, dnorme; simpl.
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); [intros->|intro Hne'].
+ + now right; rewrite norm_invol.
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ case_eq (Z.of_int e'); [|intros pe'..]; intro Hpe';
+ [now right; rewrite norm_invol..|].
+ case Nat.ltb_spec.
+ * now intro H; rewrite (norm_del_tail_int_norm _ _ H); right.
+ * now intros _; case norm; intros _; [right|left].
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); [intros->|intro Hne'].
+ + now right; rewrite norm_invol.
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ case_eq (Z.of_int e'); [|intros pe'..]; intro Hpe';
+ [now right; rewrite norm_invol..|].
+ case Nat.ltb_spec.
+ * now intro H; rewrite (norm_del_tail_int_norm _ _ H); right.
+ * now intros _; case norm; intros _; [right|left].
+Qed.
+
+Lemma dnorme_invol d : dnorme (dnorme d) = dnorme d.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold dnorme; simpl.
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); intro Hne'.
+ + rewrite Hne'; simpl; rewrite app_int_nil_r, norm_invol.
+ revert Hne'.
+ rewrite <-to_of.
+ change (Pos zero) with (Z.to_int 0).
+ intro H; generalize (to_int_inj _ _ H); clear H.
+ unfold e'; rewrite DecimalZ.of_to.
+ now case f; [rewrite app_int_nil_r|..].
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ unfold nb_digits, Z.of_nat; rewrite Z.sub_0_r, to_of, norm_invol.
+ rewrite app_int_nil_r, norm_invol.
+ set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..].
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); intro Hne'.
+ + rewrite Hne'; simpl; rewrite app_int_nil_r, norm_invol.
+ revert Hne'.
+ rewrite <-to_of.
+ change (Pos zero) with (Z.to_int 0).
+ intro H; generalize (to_int_inj _ _ H); clear H.
+ unfold e'; rewrite DecimalZ.of_to.
+ now case f; [rewrite app_int_nil_r|..].
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ unfold nb_digits, Z.of_nat; rewrite Z.sub_0_r, to_of, norm_invol.
+ rewrite app_int_nil_r, norm_invol.
+ set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..].
+Qed.
+
+Lemma dnormf_invol d : dnormf (dnormf d) = dnormf d.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold dnormf, dnorme; simpl.
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); intro Hne'.
+ + rewrite Hne'; simpl; rewrite app_int_nil_r, norm_invol.
+ revert Hne'.
+ rewrite <-to_of.
+ change (Pos zero) with (Z.to_int 0).
+ intro H; generalize (to_int_inj _ _ H); clear H.
+ unfold e'; rewrite DecimalZ.of_to.
+ now case f; [rewrite app_int_nil_r|..].
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite of_int_norm.
+ case_eq (Z.of_int e'); [|intro pe'..]; intro Hnpe';
+ [now simpl; rewrite app_int_nil_r, norm_invol..|].
+ case Nat.ltb_spec; intro Hpe'.
+ * rewrite nb_digits_del_head; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.lt_le_incl].
+ simpl.
+ rewrite Z.sub_sub_distr, Z.sub_diag, Z.add_0_l.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ rewrite positive_nat_Z; simpl.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ rewrite app_int_del_tail_head; [|now apply Nat.lt_le_incl].
+ now rewrite norm_invol, (proj2 (Nat.ltb_lt _ _) Hpe').
+ * simpl.
+ rewrite nb_digits_iter_D0.
+ rewrite (Nat.sub_add _ _ Hpe').
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ rewrite positive_nat_Z; simpl.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ revert Hpe'.
+ set (i' := norm (app_int i f)).
+ case_eq i'; intros u Hu Hpe'.
+ ++ simpl; unfold app; simpl.
+ rewrite unorm_D0, unorm_iter_D0.
+ assert (Hu' : unorm u = u).
+ { generalize (f_equal norm Hu).
+ unfold i'; rewrite norm_invol; fold i'.
+ now simpl; rewrite Hu; intro H; injection H. }
+ now rewrite Hu', (proj2 (Nat.ltb_ge _ _) Hpe').
+ ++ simpl; rewrite nzhead_iter_D0.
+ assert (Hu' : nzhead u = u).
+ { generalize (f_equal norm Hu).
+ unfold i'; rewrite norm_invol; fold i'.
+ now rewrite Hu; simpl; case (nzhead u); [|intros u' H; injection H..]. }
+ rewrite Hu'.
+ assert (Hu'' : u <> Nil).
+ { intro H; revert Hu; rewrite H; unfold i'.
+ now case app_int; intro u'; [|simpl; case nzhead]. }
+ set (m := match u with Nil => Pos zero | _ => _ end).
+ assert (H : m = Neg u); [|rewrite H; clear m H].
+ { now revert Hu''; unfold m; case u. }
+ now rewrite (proj2 (Nat.ltb_ge _ _) Hpe').
+ - set (e' := Z.to_int _).
+ case (int_eq_dec (norm e') (Pos zero)); intro Hne'.
+ + rewrite Hne'; simpl; rewrite app_int_nil_r, norm_invol.
+ revert Hne'.
+ rewrite <-to_of.
+ change (Pos zero) with (Z.to_int 0).
+ intro H; generalize (to_int_inj _ _ H); clear H.
+ unfold e'; rewrite DecimalZ.of_to.
+ now case f; [rewrite app_int_nil_r|..].
+ + set (r := DecimalExp _ _ _).
+ set (m := match norm e' with Pos zero => _ | _ => _ end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { unfold m; revert Hne'; case (norm e'); intro e''; [|now simpl].
+ now case e''; [|intro e'''; case e'''..]. }
+ rewrite of_int_norm.
+ case_eq (Z.of_int e'); [|intro pe'..]; intro Hnpe';
+ [now simpl; rewrite app_int_nil_r, norm_invol..|].
+ case Nat.ltb_spec; intro Hpe'.
+ * rewrite nb_digits_del_head; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.lt_le_incl].
+ simpl.
+ rewrite Z.sub_sub_distr, Z.sub_diag, Z.add_0_l.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ rewrite positive_nat_Z; simpl.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ rewrite app_int_del_tail_head; [|now apply Nat.lt_le_incl].
+ now rewrite norm_invol, (proj2 (Nat.ltb_lt _ _) Hpe').
+ * simpl.
+ rewrite nb_digits_iter_D0.
+ rewrite (Nat.sub_add _ _ Hpe').
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ rewrite positive_nat_Z; simpl.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ revert Hpe'.
+ set (i' := norm (app_int i f)).
+ case_eq i'; intros u Hu Hpe'.
+ ++ simpl; unfold app; simpl.
+ rewrite unorm_D0, unorm_iter_D0.
+ assert (Hu' : unorm u = u).
+ { generalize (f_equal norm Hu).
+ unfold i'; rewrite norm_invol; fold i'.
+ now simpl; rewrite Hu; intro H; injection H. }
+ now rewrite Hu', (proj2 (Nat.ltb_ge _ _) Hpe').
+ ++ simpl; rewrite nzhead_iter_D0.
+ assert (Hu' : nzhead u = u).
+ { generalize (f_equal norm Hu).
+ unfold i'; rewrite norm_invol; fold i'.
+ now rewrite Hu; simpl; case (nzhead u); [|intros u' H; injection H..]. }
+ rewrite Hu'.
+ assert (Hu'' : u <> Nil).
+ { intro H; revert Hu; rewrite H; unfold i'.
+ now case app_int; intro u'; [|simpl; case nzhead]. }
+ set (m := match u with Nil => Pos zero | _ => _ end).
+ assert (H : m = Neg u); [|rewrite H; clear m H].
+ { now revert Hu''; unfold m; case u. }
+ now rewrite (proj2 (Nat.ltb_ge _ _) Hpe').
+Qed.
+
+Lemma to_of (d:decimal) :
+ to_decimal (of_decimal d) = Some (dnorme d)
+ \/ to_decimal (of_decimal d) = Some (dnormf d).
+Proof.
+ unfold to_decimal.
+ pose (t10 := fun y => ((y + y~0~0)~0)%positive).
+ assert (H : exists e_den,
+ Decimal.nztail (Pos.to_uint (Qden (of_decimal d))) = (D1 Nil, e_den)).
+ { assert (H : forall p,
+ Decimal.nztail (Pos.to_uint (Pos.iter t10 1%positive p))
+ = (D1 Nil, Pos.to_nat p)).
+ { intro p; rewrite Pos2Nat.inj_iter.
+ fold (Nat.iter (Pos.to_nat p) t10 1%positive).
+ induction (Pos.to_nat p); [now simpl|].
+ rewrite DecimalPos.Unsigned.nat_iter_S.
+ unfold Pos.to_uint.
+ change (Pos.to_little_uint _)
+ with (Unsigned.to_lu (10 * N.pos (Nat.iter n t10 1%positive))).
+ rewrite Unsigned.to_ldec_tenfold.
+ revert IHn; unfold Pos.to_uint.
+ unfold Decimal.nztail; rewrite !rev_rev; simpl.
+ set (f'' := _ (Pos.to_little_uint _)).
+ now case f''; intros r n' H; inversion H. }
+ case d; intros i f; [|intro e]; unfold of_decimal; simpl.
+ - case (- Z.of_nat _)%Z; [|intro p..]; simpl; [now exists O..|].
+ exists (Pos.to_nat p); apply H.
+ - case (_ - _)%Z; [|intros p..]; simpl; [now exists O..|].
+ exists (Pos.to_nat p); apply H. }
+ generalize (DecimalPos.Unsigned.nztail_to_uint (Qden (of_decimal d))).
+ destruct H as (e, He); rewrite He; clear He; simpl.
+ assert (Hn1 : forall p, N.pos (Pos.iter t10 1%positive p) = 1%N -> False).
+ { intro p.
+ rewrite Pos2Nat.inj_iter.
+ case_eq (Pos.to_nat p); [|now simpl].
+ intro H; exfalso; apply (lt_irrefl O).
+ rewrite <-H at 2; apply Pos2Nat.is_pos. }
+ assert (Ht10inj : forall n m, t10 n = t10 m -> n = m).
+ { intros n m H; generalize (f_equal Z.pos H); clear H.
+ change (Z.pos (t10 n)) with (Z.mul 10 (Z.pos n)).
+ change (Z.pos (t10 m)) with (Z.mul 10 (Z.pos m)).
+ rewrite Z.mul_comm, (Z.mul_comm 10).
+ intro H; generalize (f_equal (fun z => Z.div z 10) H); clear H.
+ now rewrite !Z.div_mul; [|now simpl..]; intro H; inversion H. }
+ assert (Hinj : forall n m,
+ Nat.iter n t10 1%positive = Nat.iter m t10 1%positive -> n = m).
+ { induction n; [now intro m; case m|].
+ intro m; case m; [now simpl|]; clear m; intro m.
+ rewrite !Unsigned.nat_iter_S.
+ intro H; generalize (Ht10inj _ _ H); clear H; intro H.
+ now rewrite (IHn _ H). }
+ case e; clear e; [|intro e]; simpl; unfold of_decimal, dnormf, dnorme.
+ - case d; clear d; intros i f; [|intro e]; simpl.
+ + intro H; left; revert H.
+ generalize (nb_digits_pos f).
+ case f;
+ [|now clear f; intro f; intros H1 H2; exfalso; revert H1 H2;
+ case nb_digits; simpl;
+ [intros H _; apply (lt_irrefl O), H|intros n _; apply Hn1]..].
+ now intros _ _; simpl; rewrite to_of.
+ + intro H; right; revert H.
+ rewrite <-to_of, DecimalZ.of_to.
+ set (emf := (_ - _)%Z).
+ case_eq emf; [|intro pemf..].
+ * now simpl; rewrite to_of.
+ * set (r := DecimalExp _ _ _).
+ set (m := match _ with Pos _ => _ | _ => r end).
+ assert (H : m = r).
+ { unfold m, Z.to_int.
+ generalize (Unsigned.to_uint_nonzero pemf).
+ now case Pos.to_uint; [|intro u; case u..]. }
+ rewrite H; unfold r; clear H m r.
+ rewrite DecimalZ.of_to.
+ simpl Qnum.
+ intros Hpemf _.
+ apply f_equal; apply f_equal2; [|reflexivity].
+ rewrite !Pos2Nat.inj_iter.
+ set (n := _ pemf).
+ fold (Nat.iter n (Z.mul 10) (Z.of_int (app_int i f))).
+ fold (Nat.iter n D0 Nil).
+ rewrite <-of_int_iter_D0, to_of.
+ now rewrite norm_app_int_norm; [|induction n].
+ * simpl Qden; intros _ H; exfalso; revert H; apply Hn1.
+ - case d; clear d; intros i f; [|intro e']; simpl.
+ + case_eq (nb_digits f); [|intros nf' Hnf'];
+ [now simpl; intros _ H; exfalso; symmetry in H; revert H; apply Hn1|].
+ unfold Z.of_nat, Z.opp.
+ simpl Qden.
+ intro H; injection H; clear H; unfold Pos.pow.
+ rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (SuccNat2Pos.inj _ _ ((Pos2Nat.inj _ _ H))); clear H.
+ intro He; rewrite <-He; clear e He.
+ simpl Qnum.
+ case Nat.ltb; [left|right].
+ * now rewrite <-to_of, DecimalZ.of_to, to_of.
+ * rewrite to_of.
+ set (nif := norm _).
+ set (anif := match nif with Pos i0 => i0 | _ => _ end).
+ set (r := DecimalExp nif Nil _).
+ set (m := match _ with Pos _ => _ | _ => r end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { now unfold m; rewrite <-to_of, DecimalZ.of_to. }
+ rewrite <-to_of, !DecimalZ.of_to.
+ fold anif.
+ now rewrite SuccNat2Pos.id_succ.
+ + set (nemf := (_ - _)%Z); intro H.
+ assert (H' : exists pnemf, nemf = Z.neg pnemf); [|revert H].
+ { revert H; case nemf; [|intro pnemf..]; [..|now intros _; exists pnemf];
+ simpl Qden; intro H; exfalso; symmetry in H; revert H; apply Hn1. }
+ destruct H' as (pnemf,Hpnemf); rewrite Hpnemf.
+ simpl Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H.
+ intro H; revert Hpnemf; rewrite H; clear pnemf H; intro Hnemf.
+ simpl Qnum.
+ case Nat.ltb; [left|right].
+ * now rewrite <-to_of, DecimalZ.of_to, to_of.
+ * rewrite to_of.
+ set (nif := norm _).
+ set (anif := match nif with Pos i0 => i0 | _ => _ end).
+ set (r := DecimalExp nif Nil _).
+ set (m := match _ with Pos _ => _ | _ => r end).
+ assert (H : m = r); [|rewrite H; unfold m, r; clear m r H].
+ { now unfold m; rewrite <-to_of, DecimalZ.of_to. }
+ rewrite <-to_of, !DecimalZ.of_to.
+ fold anif.
+ now rewrite SuccNat2Pos.id_succ.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_decimal_inj q q' :
+ to_decimal q <> None -> to_decimal q = to_decimal q' -> q = q'.
+Proof.
+ intros Hnone EQ.
+ generalize (of_to q) (of_to q').
+ rewrite <-EQ.
+ revert Hnone; case to_decimal; [|now simpl].
+ now intros d _ H1 H2; rewrite <-(H1 d eq_refl), <-(H2 d eq_refl).
+Qed.
+
+Lemma to_decimal_surj d :
+ exists q, to_decimal q = Some (dnorme d) \/ to_decimal q = Some (dnormf d).
+Proof.
+ exists (of_decimal d). apply to_of.
+Qed.
+
+Lemma of_decimal_dnorme d : of_decimal (dnorme d) = of_decimal d.
+Proof.
+ unfold of_decimal, dnorme.
+ destruct d.
+ - rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ case_eq (nb_digits f); [|intro nf]; intro Hnf.
+ + now simpl; rewrite app_int_nil_r, <-DecimalZ.to_of, DecimalZ.of_to.
+ + simpl; rewrite Z.sub_0_r.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ rewrite app_int_nil_r.
+ now rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ - rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ set (emf := (_ - _)%Z).
+ case_eq emf; [|intro pemf..]; intro Hemf.
+ + now simpl; rewrite app_int_nil_r, <-DecimalZ.to_of, DecimalZ.of_to.
+ + simpl.
+ set (r := DecimalExp _ Nil _).
+ set (m := match Pos.to_uint pemf with zero => _ | _ => r end).
+ assert (H : m = r); [|rewrite H; unfold r; clear m r H].
+ { generalize (Unsigned.to_uint_nonzero pemf).
+ now unfold m; case Pos.to_uint; [|intro u; case u|..]. }
+ simpl; rewrite Z.sub_0_r.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ rewrite app_int_nil_r.
+ now rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ + simpl.
+ unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to; simpl.
+ rewrite app_int_nil_r.
+ now rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+Qed.
+
+Lemma of_decimal_dnormf d : of_decimal (dnormf d) = of_decimal d.
+Proof.
+ rewrite <-(of_decimal_dnorme d).
+ unfold of_decimal, dnormf.
+ assert (H : match dnorme d with Decimal _ f | DecimalExp _ f _ => f end = Nil).
+ { now unfold dnorme; destruct d;
+ (case norm; intro d; [case d; [|intro u; case u|..]|]). }
+ revert H; generalize (dnorme d); clear d; intro d.
+ destruct d; intro H; rewrite H; clear H; [now simpl|].
+ case (Z.of_int e); clear e; [|intro e..].
+ - now simpl.
+ - simpl.
+ rewrite app_int_nil_r.
+ apply f_equal2; [|reflexivity].
+ rewrite app_int_nil_r.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ rewrite !Pos2Nat.inj_iter.
+ fold (Nat.iter (Pos.to_nat e) D0 Nil).
+ now rewrite of_int_iter_D0.
+ - simpl.
+ set (ai := match i with Pos _ => _ | _ => _ end).
+ rewrite app_int_nil_r.
+ case Nat.ltb_spec; intro Hei; simpl.
+ + rewrite nb_digits_del_head; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply Nat.le_sub_l].
+ rewrite Nat2Z.inj_sub; [|now apply le_Sn_le].
+ rewrite Z.sub_sub_distr, Z.sub_diag; simpl.
+ rewrite positive_nat_Z; simpl.
+ now revert Hei; unfold ai; case i; clear i ai; intros i Hei; simpl;
+ (rewrite app_del_tail_head; [|now apply le_Sn_le]).
+ + set (n := nb_digits _).
+ assert (H : (n = Pos.to_nat e - nb_digits ai + nb_digits ai)%nat).
+ { unfold n; induction (_ - _)%nat; [now simpl|].
+ now rewrite Unsigned.nat_iter_S; simpl; rewrite IHn0. }
+ rewrite H; clear n H.
+ rewrite Nat2Z.inj_add, (Nat2Z.inj_sub _ _ Hei).
+ rewrite <-Z.sub_sub_distr, Z.sub_diag, Z.sub_0_r.
+ rewrite positive_nat_Z; simpl.
+ rewrite <-(DecimalZ.of_to (Z.of_int (app_int _ _))), DecimalZ.to_of.
+ rewrite <-(DecimalZ.of_to (Z.of_int i)), DecimalZ.to_of.
+ apply f_equal2; [|reflexivity]; apply f_equal.
+ now unfold ai; case i; clear i ai Hei; intro i;
+ (induction (_ - _)%nat; [|rewrite <-IHn]).
+Qed.
diff --git a/theories/Numbers/DecimalString.v b/theories/Numbers/DecimalString.v
index 7a4906a183..de577592e4 100644
--- a/theories/Numbers/DecimalString.v
+++ b/theories/Numbers/DecimalString.v
@@ -24,23 +24,23 @@ 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.
+ 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' ->
@@ -55,8 +55,8 @@ Lemma uint_of_char_spec c d d' :
c = "8" /\ d' = D8 d \/
c = "9" /\ d' = D9 d)%char.
Proof.
- destruct c as [[|] [|] [|] [|] [|] [|] [|] [|]];
- intros [= <-]; intuition.
+ destruct c as [[|] [|] [|] [|] [|] [|] [|] [|]];
+ intros [= <-]; intuition.
Qed.
(** Decimal/String conversion where [Nil] is [""] *)
@@ -64,39 +64,39 @@ Qed.
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.
+ 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.
+ 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.
+ 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.
+ 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.
(* NB: For the moment whitespace between - and digits are not accepted.
And in this variant [int_of_string "-" = Some (Neg Nil)].
@@ -110,44 +110,44 @@ Compute string_of_int (-123456890123456890123456890123456890).
Lemma usu d :
uint_of_string (string_of_uint d) = Some d.
Proof.
- induction d; simpl; rewrite ?IHd; simpl; auto.
+ induction d; simpl; rewrite ?IHd; simpl; auto.
Qed.
Lemma sus s d :
uint_of_string s = Some d -> string_of_uint d = s.
Proof.
- revert d.
- induction s; simpl.
- - now intros d [= <-].
- - intros d.
- destruct (uint_of_string s); [intros H | intros [=]].
- apply uint_of_char_spec in H.
- intuition subst; simpl; f_equal; auto.
+ revert d.
+ induction s; simpl.
+ - now intros d [= <-].
+ - intros d.
+ destruct (uint_of_string s); [intros H | intros [=]].
+ apply uint_of_char_spec in H.
+ intuition subst; simpl; f_equal; auto.
Qed.
Lemma isi d : int_of_string (string_of_int d) = Some d.
Proof.
- destruct d; simpl.
- - unfold int_of_string.
- destruct (string_of_uint d) eqn:Hd.
- + now destruct d.
- + case Ascii.eqb_spec.
- * intros ->. now destruct d.
- * rewrite <- Hd, usu; auto.
- - rewrite usu; auto.
+ destruct d; simpl.
+ - unfold int_of_string.
+ destruct (string_of_uint d) eqn:Hd.
+ + now destruct d.
+ + case Ascii.eqb_spec.
+ * intros ->. now destruct d.
+ * rewrite <- Hd, usu; auto.
+ - rewrite usu; auto.
Qed.
Lemma sis s d :
- int_of_string s = Some d -> string_of_int d = s.
+ int_of_string s = Some d -> string_of_int d = s.
Proof.
- destruct s; [intros [= <-]| ]; simpl; trivial.
- case Ascii.eqb_spec.
- - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
- simpl; f_equal. now apply sus.
- - destruct d; [ | now destruct uint_of_char].
- simpl string_of_int.
- intros. apply sus; simpl.
- destruct uint_of_char; simpl in *; congruence.
+ destruct s; [intros [= <-]| ]; simpl; trivial.
+ case Ascii.eqb_spec.
+ - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
+ simpl; f_equal. now apply sus.
+ - destruct d; [ | now destruct uint_of_char].
+ simpl string_of_int.
+ intros. apply sus; simpl.
+ destruct uint_of_char; simpl in *; congruence.
Qed.
End NilEmpty.
@@ -157,109 +157,109 @@ End NilEmpty.
Module NilZero.
Definition string_of_uint (d:uint) :=
- match d with
- | Nil => "0"
- | _ => NilEmpty.string_of_uint d
- end.
+ 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.
+ 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.
+ 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.
+ 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.
Proof.
- destruct s; simpl.
- - easy.
- - destruct (NilEmpty.uint_of_string s); [intros H | intros [=]].
- apply uint_of_char_spec in H.
- now intuition subst.
+ destruct s; simpl.
+ - easy.
+ - destruct (NilEmpty.uint_of_string s); [intros H | intros [=]].
+ apply uint_of_char_spec in H.
+ now intuition subst.
Qed.
Lemma sus s d :
uint_of_string s = Some d -> string_of_uint d = s.
Proof.
- destruct s; [intros [=] | intros H].
- apply NilEmpty.sus in H. now destruct d.
+ destruct s; [intros [=] | intros H].
+ apply NilEmpty.sus in H. now destruct d.
Qed.
Lemma usu d :
d<>Nil -> uint_of_string (string_of_uint d) = Some d.
Proof.
- destruct d; (now destruct 1) || (intros _; apply NilEmpty.usu).
+ destruct d; (now destruct 1) || (intros _; apply NilEmpty.usu).
Qed.
Lemma usu_nil :
uint_of_string (string_of_uint Nil) = Some Decimal.zero.
Proof.
- reflexivity.
+ reflexivity.
Qed.
Lemma usu_gen d :
uint_of_string (string_of_uint d) = Some d \/
uint_of_string (string_of_uint d) = Some Decimal.zero.
Proof.
- destruct d; (now right) || (left; now apply usu).
+ destruct d; (now right) || (left; now apply usu).
Qed.
Lemma isi d :
d<>Pos Nil -> d<>Neg Nil ->
int_of_string (string_of_int d) = Some d.
Proof.
- destruct d; simpl.
- - intros H _.
- unfold int_of_string.
- destruct (string_of_uint d) eqn:Hd.
- + now destruct d.
- + case Ascii.eqb_spec.
- * intros ->. now destruct d.
- * rewrite <- Hd, usu; auto. now intros ->.
- - intros _ H.
- rewrite usu; auto. now intros ->.
+ destruct d; simpl.
+ - intros H _.
+ unfold int_of_string.
+ destruct (string_of_uint d) eqn:Hd.
+ + now destruct d.
+ + case Ascii.eqb_spec.
+ * intros ->. now destruct d.
+ * rewrite <- Hd, usu; auto. now intros ->.
+ - intros _ H.
+ rewrite usu; auto. now intros ->.
Qed.
Lemma isi_posnil :
- int_of_string (string_of_int (Pos Nil)) = Some (Pos Decimal.zero).
+ int_of_string (string_of_int (Pos Nil)) = Some (Pos Decimal.zero).
Proof.
- reflexivity.
+ reflexivity.
Qed.
(** 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)).
+ int_of_string (string_of_int (Neg Nil)) = Some (Neg (D0 Nil)).
Proof.
- reflexivity.
+ reflexivity.
Qed.
Lemma sis s d :
- int_of_string s = Some d -> string_of_int d = s.
+ int_of_string s = Some d -> string_of_int d = s.
Proof.
- destruct s; [intros [=]| ]; simpl.
- case Ascii.eqb_spec.
- - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
- simpl; f_equal. now apply sus.
- - destruct d; [ | now destruct uint_of_char].
- simpl string_of_int.
- intros. apply sus; simpl.
- destruct uint_of_char; simpl in *; congruence.
+ destruct s; [intros [=]| ]; simpl.
+ case Ascii.eqb_spec.
+ - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
+ simpl; f_equal. now apply sus.
+ - destruct d; [ | now destruct uint_of_char].
+ simpl string_of_int.
+ intros. apply sus; simpl.
+ destruct uint_of_char; simpl in *; congruence.
Qed.
End NilZero.
diff --git a/theories/Numbers/DecimalZ.v b/theories/Numbers/DecimalZ.v
index 6055ebb5d3..69d8073fc7 100644
--- a/theories/Numbers/DecimalZ.v
+++ b/theories/Numbers/DecimalZ.v
@@ -17,59 +17,86 @@ Require Import Decimal DecimalFacts DecimalPos DecimalN ZArith.
Lemma of_to (z:Z) : Z.of_int (Z.to_int z) = z.
Proof.
- destruct z; simpl.
- - trivial.
- - unfold Z.of_uint. rewrite DecimalPos.Unsigned.of_to. now destruct p.
- - unfold Z.of_uint. rewrite DecimalPos.Unsigned.of_to. destruct p; auto.
+ destruct z; simpl.
+ - trivial.
+ - unfold Z.of_uint. rewrite DecimalPos.Unsigned.of_to. now destruct p.
+ - unfold Z.of_uint. rewrite DecimalPos.Unsigned.of_to. destruct p; auto.
Qed.
Lemma to_of (d:int) : Z.to_int (Z.of_int d) = norm d.
Proof.
- destruct d; simpl; unfold Z.to_int, Z.of_uint.
- - rewrite <- (DecimalN.Unsigned.to_of d). unfold N.of_uint.
- now destruct (Pos.of_uint d).
- - destruct (Pos.of_uint d) eqn:Hd; simpl; f_equal.
- + generalize (DecimalPos.Unsigned.to_of d). rewrite Hd. simpl.
- intros H. symmetry in H. apply unorm_0 in H. now rewrite H.
- + assert (Hp := DecimalPos.Unsigned.to_of d). rewrite Hd in Hp. simpl in *.
- rewrite Hp. unfold unorm in *.
- destruct (nzhead d); trivial.
- generalize (DecimalPos.Unsigned.of_to p). now rewrite Hp.
+ destruct d; simpl; unfold Z.to_int, Z.of_uint.
+ - rewrite <- (DecimalN.Unsigned.to_of d). unfold N.of_uint.
+ now destruct (Pos.of_uint d).
+ - destruct (Pos.of_uint d) eqn:Hd; simpl; f_equal.
+ + generalize (DecimalPos.Unsigned.to_of d). rewrite Hd. simpl.
+ intros H. symmetry in H. apply unorm_0 in H. now rewrite H.
+ + assert (Hp := DecimalPos.Unsigned.to_of d). rewrite Hd in Hp. simpl in *.
+ rewrite Hp. unfold unorm in *.
+ destruct (nzhead d); trivial.
+ generalize (DecimalPos.Unsigned.of_to p). now rewrite Hp.
Qed.
(** Some consequences *)
Lemma to_int_inj n n' : Z.to_int n = Z.to_int n' -> n = n'.
Proof.
- intro EQ.
- now rewrite <- (of_to n), <- (of_to n'), EQ.
+ intro EQ.
+ now rewrite <- (of_to n), <- (of_to n'), EQ.
Qed.
Lemma to_int_surj d : exists n, Z.to_int n = norm d.
Proof.
- exists (Z.of_int d). apply to_of.
+ exists (Z.of_int d). apply to_of.
Qed.
Lemma of_int_norm d : Z.of_int (norm d) = Z.of_int d.
Proof.
- unfold Z.of_int, Z.of_uint.
- destruct d.
- - simpl. now rewrite DecimalPos.Unsigned.of_uint_norm.
- - simpl. destruct (nzhead d) eqn:H;
- [ induction d; simpl; auto; discriminate |
- destruct (nzhead_nonzero _ _ H) | .. ];
- f_equal; f_equal; apply DecimalPos.Unsigned.of_iff;
- unfold unorm; now rewrite H.
+ unfold Z.of_int, Z.of_uint.
+ destruct d.
+ - simpl. now rewrite DecimalPos.Unsigned.of_uint_norm.
+ - simpl. destruct (nzhead d) eqn:H;
+ [ induction d; simpl; auto; discriminate |
+ destruct (nzhead_nonzero _ _ H) | .. ];
+ f_equal; f_equal; apply DecimalPos.Unsigned.of_iff;
+ unfold unorm; now rewrite H.
Qed.
Lemma of_inj d d' :
- Z.of_int d = Z.of_int d' -> norm d = norm d'.
+ Z.of_int d = Z.of_int d' -> norm d = norm d'.
Proof.
- intros. rewrite <- !to_of. now f_equal.
+ intros. rewrite <- !to_of. now f_equal.
Qed.
Lemma of_iff d d' : Z.of_int d = Z.of_int d' <-> norm d = norm d'.
Proof.
- split. apply of_inj. intros E. rewrite <- of_int_norm, E.
- apply of_int_norm.
+ split. apply of_inj. intros E. rewrite <- of_int_norm, E.
+ apply of_int_norm.
+Qed.
+
+(** Various lemmas *)
+
+Lemma of_uint_iter_D0 d n :
+ Z.of_uint (app d (Nat.iter n D0 Nil)) = Nat.iter n (Z.mul 10) (Z.of_uint d).
+Proof.
+ unfold Z.of_uint.
+ unfold app; rewrite <-rev_revapp.
+ rewrite Unsigned.of_lu_rev, Unsigned.of_lu_revapp.
+ rewrite <-!Unsigned.of_lu_rev, !rev_rev.
+ assert (H' : Pos.of_uint (Nat.iter n D0 Nil) = 0%N).
+ { now induction n; [|rewrite Unsigned.nat_iter_S]. }
+ rewrite H', N.add_0_l; clear H'.
+ induction n; [now simpl; rewrite N.mul_1_r|].
+ rewrite !Unsigned.nat_iter_S, <-IHn.
+ simpl Unsigned.usize; rewrite N.pow_succ_r'.
+ rewrite !N2Z.inj_mul; simpl Z.of_N; ring.
+Qed.
+
+Lemma of_int_iter_D0 d n :
+ Z.of_int (app_int d (Nat.iter n D0 Nil)) = Nat.iter n (Z.mul 10) (Z.of_int d).
+Proof.
+ case d; clear d; intro d; simpl.
+ - now rewrite of_uint_iter_D0.
+ - rewrite of_uint_iter_D0; induction n; [now simpl|].
+ rewrite !Unsigned.nat_iter_S, <-IHn; ring.
Qed.
diff --git a/theories/Numbers/HexadecimalFacts.v b/theories/Numbers/HexadecimalFacts.v
new file mode 100644
index 0000000000..7328b2303d
--- /dev/null
+++ b/theories/Numbers/HexadecimalFacts.v
@@ -0,0 +1,340 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalFacts : some facts about Hexadecimal numbers *)
+
+Require Import Hexadecimal Arith.
+
+Scheme Equality for uint.
+
+Scheme Equality for int.
+
+Lemma rev_revapp d d' :
+ rev (revapp d d') = revapp d' d.
+Proof.
+ revert d'. induction d; simpl; intros; now rewrite ?IHd.
+Qed.
+
+Lemma rev_rev d : rev (rev d) = d.
+Proof.
+ apply rev_revapp.
+Qed.
+
+Lemma revapp_rev_nil d : revapp (rev d) Nil = d.
+Proof. now fold (rev (rev d)); rewrite rev_rev. Qed.
+
+Lemma app_nil_r d : app d Nil = d.
+Proof. now unfold app; rewrite revapp_rev_nil. Qed.
+
+Lemma app_int_nil_r d : app_int d Nil = d.
+Proof. now case d; intro d'; simpl; rewrite app_nil_r. Qed.
+
+Lemma revapp_revapp_1 d d' d'' :
+ nb_digits d <= 1 ->
+ revapp (revapp d d') d'' = revapp d' (revapp d d'').
+Proof.
+ now case d; clear d; intro d;
+ [|case d; clear d; intro d;
+ [|simpl; case nb_digits; [|intros n]; intros Hn; exfalso;
+ [apply (Nat.nle_succ_diag_l _ Hn)|
+ apply (Nat.nle_succ_0 _ (le_S_n _ _ Hn))]..]..].
+Qed.
+
+Lemma nb_digits_pos d : d <> Nil -> 0 < nb_digits d.
+Proof. now case d; [|intros d' _; apply Nat.lt_0_succ..]. Qed.
+
+Lemma nb_digits_revapp d d' :
+ nb_digits (revapp d d') = nb_digits d + nb_digits d'.
+Proof.
+ now revert d'; induction d; [|intro d'; simpl; rewrite IHd; simpl..].
+Qed.
+
+Lemma nb_digits_rev u : nb_digits (rev u) = nb_digits u.
+Proof. now unfold rev; rewrite nb_digits_revapp. Qed.
+
+Lemma nb_digits_nzhead u : nb_digits (nzhead u) <= nb_digits u.
+Proof. now induction u; [|apply le_S|..]. Qed.
+
+Lemma nb_digits_iter_D0 n d :
+ nb_digits (Nat.iter n D0 d) = n + nb_digits d.
+Proof. now induction n; simpl; [|rewrite IHn]. Qed.
+
+Fixpoint nth n u :=
+ match n with
+ | O =>
+ match u with
+ | Nil => Nil
+ | D0 d => D0 Nil
+ | D1 d => D1 Nil
+ | D2 d => D2 Nil
+ | D3 d => D3 Nil
+ | D4 d => D4 Nil
+ | D5 d => D5 Nil
+ | D6 d => D6 Nil
+ | D7 d => D7 Nil
+ | D8 d => D8 Nil
+ | D9 d => D9 Nil
+ | Da d => Da Nil
+ | Db d => Db Nil
+ | Dc d => Dc Nil
+ | Dd d => Dd Nil
+ | De d => De Nil
+ | Df d => Df Nil
+ end
+ | S n =>
+ match u with
+ | Nil => Nil
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d
+ | Da d | Db d | Dc d | Dd d | De d | Df d =>
+ nth n d
+ end
+ end.
+
+Lemma nb_digits_nth n u : nb_digits (nth n u) <= 1.
+Proof.
+ revert u; induction n.
+ - now intro u; case u; [apply Nat.le_0_1|..].
+ - intro u; case u; [apply Nat.le_0_1|intro u'; apply IHn..].
+Qed.
+
+Lemma nth_revapp_r n d d' :
+ nb_digits d <= n ->
+ nth n (revapp d d') = nth (n - nb_digits d) d'.
+Proof.
+ revert d d'; induction n; intro d.
+ - now case d; intro d';
+ [case d'|intros d'' H; exfalso; revert H; apply Nat.nle_succ_0..].
+ - now induction d;
+ [intro d'; case d'|
+ intros d' H;
+ simpl revapp; rewrite IHd; [|now apply le_Sn_le];
+ rewrite Nat.sub_succ_l; [|now apply le_S_n];
+ simpl; rewrite <-(IHn _ _ (le_S_n _ _ H))..].
+Qed.
+
+Lemma nth_revapp_l n d d' :
+ n < nb_digits d ->
+ nth n (revapp d d') = nth (nb_digits d - n - 1) d.
+Proof.
+ revert d d'; induction n; intro d.
+ - rewrite Nat.sub_0_r.
+ now induction d;
+ [|intros d' _; simpl revapp;
+ revert IHd; case d; clear d; [|intro d..]; intro IHd;
+ [|rewrite IHd; [simpl nb_digits; rewrite (Nat.sub_succ_l _ (S _))|];
+ [|apply le_n_S, Nat.le_0_l..]..]..].
+ - now induction d;
+ [|intros d' H;
+ simpl revapp; simpl nb_digits;
+ simpl in H; generalize (lt_S_n _ _ H); clear H; intro H;
+ case (le_lt_eq_dec _ _ H); clear H; intro H;
+ [rewrite (IHd _ H), Nat.sub_succ_l;
+ [rewrite Nat.sub_succ_l; [|apply Nat.le_add_le_sub_r]|
+ apply le_Sn_le]|
+ rewrite nth_revapp_r; rewrite <-H;
+ [rewrite Nat.sub_succ, Nat.sub_succ_l; [rewrite !Nat.sub_diag|]|]]..].
+Qed.
+
+(** Normalization on little-endian numbers *)
+
+Fixpoint nztail d :=
+ match d with
+ | Nil => Nil
+ | D0 d => match nztail d with Nil => Nil | d' => D0 d' end
+ | D1 d => D1 (nztail d)
+ | D2 d => D2 (nztail d)
+ | D3 d => D3 (nztail d)
+ | D4 d => D4 (nztail d)
+ | D5 d => D5 (nztail d)
+ | D6 d => D6 (nztail d)
+ | D7 d => D7 (nztail d)
+ | D8 d => D8 (nztail d)
+ | D9 d => D9 (nztail d)
+ | Da d => Da (nztail d)
+ | Db d => Db (nztail d)
+ | Dc d => Dc (nztail d)
+ | Dd d => Dd (nztail d)
+ | De d => De (nztail d)
+ | Df d => Df (nztail d)
+ end.
+
+Definition lnorm d :=
+ match nztail d with
+ | Nil => zero
+ | d => d
+ end.
+
+Lemma nzhead_revapp_0 d d' : nztail d = Nil ->
+ nzhead (revapp d d') = nzhead d'.
+Proof.
+ revert d'. induction d; intros d' [=]; simpl; trivial.
+ destruct (nztail d); now rewrite IHd.
+Qed.
+
+Lemma nzhead_revapp d d' : nztail d <> Nil ->
+ nzhead (revapp d d') = revapp (nztail d) d'.
+Proof.
+ revert d'.
+ induction d; intros d' H; simpl in *;
+ try destruct (nztail d) eqn:E;
+ (now rewrite ?nzhead_revapp_0) || (now rewrite IHd).
+Qed.
+
+Lemma nzhead_rev d : nztail d <> Nil ->
+ nzhead (rev d) = rev (nztail d).
+Proof.
+ apply nzhead_revapp.
+Qed.
+
+Lemma rev_nztail_rev d :
+ rev (nztail (rev d)) = nzhead d.
+Proof.
+ destruct (uint_eq_dec (nztail (rev d)) Nil) as [H|H].
+ - rewrite H. unfold rev; simpl.
+ rewrite <- (rev_rev d). symmetry.
+ now apply nzhead_revapp_0.
+ - now rewrite <- nzhead_rev, rev_rev.
+Qed.
+
+Lemma nzhead_D0 u : nzhead (D0 u) = nzhead u.
+Proof. reflexivity. Qed.
+
+Lemma nzhead_iter_D0 n u : nzhead (Nat.iter n D0 u) = nzhead u.
+Proof. now induction n. Qed.
+
+Lemma revapp_nil_inv d d' : revapp d d' = Nil -> d = Nil /\ d' = Nil.
+Proof.
+ revert d'.
+ induction d; simpl; intros d' H; auto; now apply IHd in H.
+Qed.
+
+Lemma rev_nil_inv d : rev d = Nil -> d = Nil.
+Proof.
+ apply revapp_nil_inv.
+Qed.
+
+Lemma rev_lnorm_rev d :
+ rev (lnorm (rev d)) = unorm d.
+Proof.
+ unfold unorm, lnorm.
+ rewrite <- rev_nztail_rev.
+ destruct nztail; simpl; trivial;
+ destruct rev eqn:E; trivial; now apply rev_nil_inv in E.
+Qed.
+
+Lemma nzhead_nonzero d d' : nzhead d <> D0 d'.
+Proof.
+ induction d; easy.
+Qed.
+
+Lemma unorm_0 d : unorm d = zero <-> nzhead d = Nil.
+Proof.
+ unfold unorm. split.
+ - generalize (nzhead_nonzero d).
+ destruct nzhead; intros H [=]; trivial. now destruct (H u).
+ - now intros ->.
+Qed.
+
+Lemma unorm_nonnil d : unorm d <> Nil.
+Proof.
+ unfold unorm. now destruct nzhead.
+Qed.
+
+Lemma unorm_D0 u : unorm (D0 u) = unorm u.
+Proof. reflexivity. Qed.
+
+Lemma unorm_iter_D0 n u : unorm (Nat.iter n D0 u) = unorm u.
+Proof. now induction n. Qed.
+
+Lemma nb_digits_unorm u :
+ u <> Nil -> nb_digits (unorm u) <= nb_digits u.
+Proof.
+ case u; clear u; [now simpl|intro u..]; [|now simpl..].
+ intros _; unfold unorm.
+ case_eq (nzhead (D0 u)); [|now intros u' <-; apply nb_digits_nzhead..].
+ intros _; apply le_n_S, Nat.le_0_l.
+Qed.
+
+Lemma nzhead_invol d : nzhead (nzhead d) = nzhead d.
+Proof.
+ now induction d.
+Qed.
+
+Lemma nztail_invol d : nztail (nztail d) = nztail d.
+Proof.
+ rewrite <-(rev_rev (nztail _)), <-(rev_rev (nztail d)), <-(rev_rev d).
+ now rewrite !rev_nztail_rev, nzhead_invol.
+Qed.
+
+Lemma unorm_invol d : unorm (unorm d) = unorm d.
+Proof.
+ unfold unorm.
+ destruct (nzhead d) eqn:E; trivial.
+ destruct (nzhead_nonzero _ _ E).
+Qed.
+
+Lemma norm_invol d : norm (norm d) = norm d.
+Proof.
+ unfold norm.
+ destruct d.
+ - f_equal. apply unorm_invol.
+ - destruct (nzhead d) eqn:E; auto.
+ destruct (nzhead_nonzero _ _ E).
+Qed.
+
+Lemma nzhead_app_nzhead d d' :
+ nzhead (app (nzhead d) d') = nzhead (app d d').
+Proof.
+ unfold app.
+ rewrite <-(rev_nztail_rev d), rev_rev.
+ generalize (rev d); clear d; intro d.
+ generalize (nzhead_revapp_0 d d').
+ generalize (nzhead_revapp d d').
+ generalize (nzhead_revapp_0 (nztail d) d').
+ generalize (nzhead_revapp (nztail d) d').
+ rewrite nztail_invol.
+ now case nztail;
+ [intros _ H _ H'; rewrite (H eq_refl), (H' eq_refl)
+ |intros d'' H _ H' _; rewrite H; [rewrite H'|]..].
+Qed.
+
+Lemma unorm_app_unorm d d' :
+ unorm (app (unorm d) d') = unorm (app d d').
+Proof.
+ unfold unorm.
+ rewrite <-(nzhead_app_nzhead d d').
+ now case (nzhead d).
+Qed.
+
+Lemma norm_app_int_norm d d' :
+ unorm d' = zero ->
+ norm (app_int (norm d) d') = norm (app_int d d').
+Proof.
+ case d; clear d; intro d; simpl.
+ - now rewrite unorm_app_unorm.
+ - unfold app_int, app.
+ rewrite unorm_0; intro Hd'.
+ rewrite <-rev_nztail_rev.
+ generalize (nzhead_revapp (rev d) d').
+ generalize (nzhead_revapp_0 (rev d) d').
+ now case_eq (nztail (rev d));
+ [intros Hd'' H _; rewrite (H eq_refl); simpl;
+ unfold unorm; simpl; rewrite Hd'
+ |intros d'' Hd'' _ H; rewrite H; clear H; [|now simpl];
+ set (r := rev _);
+ set (m := match r with Nil => Pos zero | _ => _ end);
+ assert (H' : m = Neg r);
+ [now unfold m; case_eq r; unfold r;
+ [intro H''; generalize (rev_nil_inv _ H'')|..]
+ |rewrite H'; unfold r; clear m r H'];
+ unfold norm;
+ rewrite rev_rev, <-Hd'';
+ rewrite nzhead_revapp; rewrite nztail_invol; [|rewrite Hd'']..].
+Qed.
diff --git a/theories/Numbers/HexadecimalN.v b/theories/Numbers/HexadecimalN.v
new file mode 100644
index 0000000000..f333e2b7f6
--- /dev/null
+++ b/theories/Numbers/HexadecimalN.v
@@ -0,0 +1,107 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalN
+
+ Proofs that conversions between hexadecimal numbers and [N]
+ are bijections *)
+
+Require Import Hexadecimal HexadecimalFacts HexadecimalPos PArith NArith.
+
+Module Unsigned.
+
+Lemma of_to (n:N) : N.of_hex_uint (N.to_hex_uint n) = n.
+Proof.
+ destruct n.
+ - reflexivity.
+ - apply HexadecimalPos.Unsigned.of_to.
+Qed.
+
+Lemma to_of (d:uint) : N.to_hex_uint (N.of_hex_uint d) = unorm d.
+Proof.
+ exact (HexadecimalPos.Unsigned.to_of d).
+Qed.
+
+Lemma to_uint_inj n n' : N.to_hex_uint n = N.to_hex_uint n' -> n = n'.
+Proof.
+ intros E. now rewrite <- (of_to n), <- (of_to n'), E.
+Qed.
+
+Lemma to_uint_surj d : exists p, N.to_hex_uint p = unorm d.
+Proof.
+ exists (N.of_hex_uint d). apply to_of.
+Qed.
+
+Lemma of_uint_norm d : N.of_hex_uint (unorm d) = N.of_hex_uint d.
+Proof.
+ now induction d.
+Qed.
+
+Lemma of_inj d d' :
+ N.of_hex_uint d = N.of_hex_uint d' -> unorm d = unorm d'.
+Proof.
+ intros. rewrite <- !to_of. now f_equal.
+Qed.
+
+Lemma of_iff d d' : N.of_hex_uint d = N.of_hex_uint d' <-> unorm d = unorm d'.
+Proof.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed hexadecimal numbers *)
+
+Module Signed.
+
+Lemma of_to (n:N) : N.of_hex_int (N.to_hex_int n) = Some n.
+Proof.
+ unfold N.to_hex_int, N.of_hex_int, norm. f_equal.
+ rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
+Qed.
+
+Lemma to_of (d:int)(n:N) : N.of_hex_int d = Some n -> N.to_hex_int n = norm d.
+Proof.
+ unfold N.of_hex_int.
+ destruct (norm d) eqn:Hd; intros [= <-].
+ unfold N.to_hex_int. rewrite Unsigned.to_of. f_equal.
+ revert Hd; destruct d; simpl.
+ - intros [= <-]. apply unorm_invol.
+ - destruct (nzhead d); now intros [= <-].
+Qed.
+
+Lemma to_int_inj n n' : N.to_hex_int n = N.to_hex_int n' -> n = n'.
+Proof.
+ intro E.
+ assert (E' : Some n = Some n').
+ { now rewrite <- (of_to n), <- (of_to n'), E. }
+ now injection E'.
+Qed.
+
+Lemma to_int_pos_surj d : exists n, N.to_hex_int n = norm (Pos d).
+Proof.
+ exists (N.of_hex_uint d). unfold N.to_hex_int. now rewrite Unsigned.to_of.
+Qed.
+
+Lemma of_int_norm d : N.of_hex_int (norm d) = N.of_hex_int d.
+Proof.
+ unfold N.of_hex_int. now rewrite norm_invol.
+Qed.
+
+Lemma of_inj_pos d d' :
+ N.of_hex_int (Pos d) = N.of_hex_int (Pos d') -> unorm d = unorm d'.
+Proof.
+ unfold N.of_hex_int. simpl. intros [= H]. apply Unsigned.of_inj.
+ change Pos.of_hex_uint with N.of_hex_uint in H.
+ now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
+Qed.
+
+End Signed.
diff --git a/theories/Numbers/HexadecimalNat.v b/theories/Numbers/HexadecimalNat.v
new file mode 100644
index 0000000000..b05184e821
--- /dev/null
+++ b/theories/Numbers/HexadecimalNat.v
@@ -0,0 +1,321 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalNat
+
+ Proofs that conversions between hexadecimal numbers and [nat]
+ are bijections. *)
+
+Require Import Hexadecimal HexadecimalFacts Arith.
+
+Module Unsigned.
+
+(** A few helper functions used during proofs *)
+
+Definition hd d :=
+ match d with
+ | Nil => 0x0
+ | D0 _ => 0x0
+ | D1 _ => 0x1
+ | D2 _ => 0x2
+ | D3 _ => 0x3
+ | D4 _ => 0x4
+ | D5 _ => 0x5
+ | D6 _ => 0x6
+ | D7 _ => 0x7
+ | D8 _ => 0x8
+ | D9 _ => 0x9
+ | Da _ => 0xa
+ | Db _ => 0xb
+ | Dc _ => 0xc
+ | Dd _ => 0xd
+ | De _ => 0xe
+ | Df _ => 0xf
+ end.
+
+Definition tl d :=
+ match d with
+ | Nil => d
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d
+ | Da d | Db d | Dc d | Dd d | De d | Df d => d
+ end.
+
+Fixpoint usize (d:uint) : nat :=
+ match d with
+ | Nil => 0
+ | D0 d => S (usize d)
+ | D1 d => S (usize d)
+ | D2 d => S (usize d)
+ | D3 d => S (usize d)
+ | D4 d => S (usize d)
+ | D5 d => S (usize d)
+ | D6 d => S (usize d)
+ | D7 d => S (usize d)
+ | D8 d => S (usize d)
+ | D9 d => S (usize d)
+ | Da d => S (usize d)
+ | Db d => S (usize d)
+ | Dc d => S (usize d)
+ | Dd d => S (usize d)
+ | De d => S (usize d)
+ | Df d => S (usize d)
+ end.
+
+(** A direct version of [to_little_uint], not tail-recursive *)
+Fixpoint to_lu n :=
+ match n with
+ | 0 => Hexadecimal.zero
+ | S n => Little.succ (to_lu n)
+ end.
+
+(** A direct version of [of_little_uint] *)
+Fixpoint of_lu (d:uint) : nat :=
+ match d with
+ | Nil => 0x0
+ | D0 d => 0x10 * of_lu d
+ | D1 d => 0x1 + 0x10 * of_lu d
+ | D2 d => 0x2 + 0x10 * of_lu d
+ | D3 d => 0x3 + 0x10 * of_lu d
+ | D4 d => 0x4 + 0x10 * of_lu d
+ | D5 d => 0x5 + 0x10 * of_lu d
+ | D6 d => 0x6 + 0x10 * of_lu d
+ | D7 d => 0x7 + 0x10 * of_lu d
+ | D8 d => 0x8 + 0x10 * of_lu d
+ | D9 d => 0x9 + 0x10 * of_lu d
+ | Da d => 0xa + 0x10 * of_lu d
+ | Db d => 0xb + 0x10 * of_lu d
+ | Dc d => 0xc + 0x10 * of_lu d
+ | Dd d => 0xd + 0x10 * of_lu d
+ | De d => 0xe + 0x10 * of_lu d
+ | Df d => 0xf + 0x10 * of_lu d
+ end.
+
+(** Properties of [to_lu] *)
+
+Lemma to_lu_succ n : to_lu (S n) = Little.succ (to_lu n).
+Proof.
+ reflexivity.
+Qed.
+
+Lemma to_little_uint_succ n d :
+ Nat.to_little_hex_uint n (Little.succ d) =
+ Little.succ (Nat.to_little_hex_uint n d).
+Proof.
+ revert d; induction n; simpl; trivial.
+Qed.
+
+Lemma to_lu_equiv n :
+ to_lu n = Nat.to_little_hex_uint n zero.
+Proof.
+ induction n; simpl; trivial.
+ now rewrite IHn, <- to_little_uint_succ.
+Qed.
+
+Lemma to_uint_alt n :
+ Nat.to_hex_uint n = rev (to_lu n).
+Proof.
+ unfold Nat.to_hex_uint. f_equal. symmetry. apply to_lu_equiv.
+Qed.
+
+(** Properties of [of_lu] *)
+
+Lemma of_lu_eqn d :
+ of_lu d = hd d + 0x10 * of_lu (tl d).
+Proof.
+ induction d; simpl; trivial.
+Qed.
+
+Ltac simpl_of_lu :=
+ match goal with
+ | |- context [ of_lu (?f ?x) ] =>
+ rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
+ end.
+
+Lemma of_lu_succ d :
+ of_lu (Little.succ d) = S (of_lu d).
+Proof.
+ induction d; trivial.
+ simpl_of_lu. rewrite IHd. simpl_of_lu.
+ now rewrite Nat.mul_succ_r, <- (Nat.add_comm 0x10).
+Qed.
+
+Lemma of_to_lu n :
+ of_lu (to_lu n) = n.
+Proof.
+ induction n; simpl; trivial. rewrite of_lu_succ. now f_equal.
+Qed.
+
+Lemma of_lu_revapp d d' :
+ of_lu (revapp d d') =
+ of_lu (rev d) + of_lu d' * 0x10^usize d.
+Proof.
+ revert d'.
+ induction d; intro d'; simpl usize;
+ [ simpl; now rewrite Nat.mul_1_r | .. ];
+ unfold rev; simpl revapp; rewrite 2 IHd;
+ rewrite <- Nat.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
+ rewrite Nat.pow_succ_r'; ring.
+Qed.
+
+Lemma of_uint_acc_spec n d :
+ Nat.of_hex_uint_acc d n = of_lu (rev d) + n * 0x10^usize d.
+Proof.
+ revert n. induction d; intros;
+ simpl Nat.of_hex_uint_acc; rewrite ?Nat.tail_mul_spec, ?IHd;
+ simpl rev; simpl usize; rewrite ?Nat.pow_succ_r';
+ [ simpl; now rewrite Nat.mul_1_r | .. ];
+ unfold rev at 2; simpl revapp; rewrite of_lu_revapp;
+ simpl of_lu; ring.
+Qed.
+
+Lemma of_uint_alt d : Nat.of_hex_uint d = of_lu (rev d).
+Proof.
+ unfold Nat.of_hex_uint. now rewrite of_uint_acc_spec.
+Qed.
+
+(** First main bijection result *)
+
+Lemma of_to (n:nat) : Nat.of_hex_uint (Nat.to_hex_uint n) = n.
+Proof.
+ rewrite to_uint_alt, of_uint_alt, rev_rev. apply of_to_lu.
+Qed.
+
+(** The other direction *)
+
+Lemma to_lu_sixteenfold n : n<>0 ->
+ to_lu (0x10 * n) = D0 (to_lu n).
+Proof.
+ induction n.
+ - simpl. now destruct 1.
+ - intros _.
+ destruct (Nat.eq_dec n 0) as [->|H]; simpl; trivial.
+ rewrite !Nat.add_succ_r.
+ simpl in *. rewrite (IHn H). now destruct (to_lu n).
+Qed.
+
+Lemma of_lu_0 d : of_lu d = 0 <-> nztail d = Nil.
+Proof.
+ induction d; try simpl_of_lu; try easy.
+ rewrite Nat.add_0_l.
+ split; intros H.
+ - apply Nat.eq_mul_0_r in H; auto.
+ rewrite IHd in H. simpl. now rewrite H.
+ - simpl in H. destruct (nztail d); try discriminate.
+ now destruct IHd as [_ ->].
+Qed.
+
+Lemma to_of_lu_sixteenfold d :
+ to_lu (of_lu d) = lnorm d ->
+ to_lu (0x10 * of_lu d) = lnorm (D0 d).
+Proof.
+ intro IH.
+ destruct (Nat.eq_dec (of_lu d) 0) as [H|H].
+ - rewrite H. simpl. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now rewrite H.
+ - rewrite (to_lu_sixteenfold _ H), IH.
+ rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now destruct (nztail d).
+Qed.
+
+Lemma to_of_lu d : to_lu (of_lu d) = lnorm d.
+Proof.
+ induction d; [ reflexivity | .. ];
+ simpl_of_lu;
+ rewrite ?Nat.add_succ_l, Nat.add_0_l, ?to_lu_succ, to_of_lu_sixteenfold
+ by assumption;
+ unfold lnorm; simpl; now destruct nztail.
+Qed.
+
+(** Second bijection result *)
+
+Lemma to_of (d:uint) : Nat.to_hex_uint (Nat.of_hex_uint d) = unorm d.
+Proof.
+ rewrite to_uint_alt, of_uint_alt, to_of_lu.
+ apply rev_lnorm_rev.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_uint_inj n n' : Nat.to_hex_uint n = Nat.to_hex_uint n' -> n = n'.
+Proof.
+ intro EQ.
+ now rewrite <- (of_to n), <- (of_to n'), EQ.
+Qed.
+
+Lemma to_uint_surj d : exists n, Nat.to_hex_uint n = unorm d.
+Proof.
+ exists (Nat.of_hex_uint d). apply to_of.
+Qed.
+
+Lemma of_uint_norm d : Nat.of_hex_uint (unorm d) = Nat.of_hex_uint d.
+Proof.
+ unfold Nat.of_hex_uint. now induction d.
+Qed.
+
+Lemma of_inj d d' :
+ Nat.of_hex_uint d = Nat.of_hex_uint d' -> unorm d = unorm d'.
+Proof.
+ intros. rewrite <- !to_of. now f_equal.
+Qed.
+
+Lemma of_iff d d' : Nat.of_hex_uint d = Nat.of_hex_uint d' <-> unorm d = unorm d'.
+Proof.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed hexadecimal numbers *)
+
+Module Signed.
+
+Lemma of_to (n:nat) : Nat.of_hex_int (Nat.to_hex_int n) = Some n.
+Proof.
+ unfold Nat.to_hex_int, Nat.of_hex_int, norm. f_equal.
+ rewrite Unsigned.of_uint_norm. apply Unsigned.of_to.
+Qed.
+
+Lemma to_of (d:int)(n:nat) : Nat.of_hex_int d = Some n -> Nat.to_hex_int n = norm d.
+Proof.
+ unfold Nat.of_hex_int.
+ destruct (norm d) eqn:Hd; intros [= <-].
+ unfold Nat.to_hex_int. rewrite Unsigned.to_of. f_equal.
+ revert Hd; destruct d; simpl.
+ - intros [= <-]. apply unorm_invol.
+ - destruct (nzhead d); now intros [= <-].
+Qed.
+
+Lemma to_int_inj n n' : Nat.to_hex_int n = Nat.to_hex_int n' -> n = n'.
+Proof.
+ intro E.
+ assert (E' : Some n = Some n').
+ { now rewrite <- (of_to n), <- (of_to n'), E. }
+ now injection E'.
+Qed.
+
+Lemma to_int_pos_surj d : exists n, Nat.to_hex_int n = norm (Pos d).
+Proof.
+ exists (Nat.of_hex_uint d). unfold Nat.to_hex_int. now rewrite Unsigned.to_of.
+Qed.
+
+Lemma of_int_norm d : Nat.of_hex_int (norm d) = Nat.of_hex_int d.
+Proof.
+ unfold Nat.of_hex_int. now rewrite norm_invol.
+Qed.
+
+Lemma of_inj_pos d d' :
+ Nat.of_hex_int (Pos d) = Nat.of_hex_int (Pos d') -> unorm d = unorm d'.
+Proof.
+ unfold Nat.of_hex_int. simpl. intros [= H]. apply Unsigned.of_inj.
+ now rewrite <- Unsigned.of_uint_norm, H, Unsigned.of_uint_norm.
+Qed.
+
+End Signed.
diff --git a/theories/Numbers/HexadecimalPos.v b/theories/Numbers/HexadecimalPos.v
new file mode 100644
index 0000000000..47f6d983b7
--- /dev/null
+++ b/theories/Numbers/HexadecimalPos.v
@@ -0,0 +1,446 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalPos
+
+ Proofs that conversions between hexadecimal numbers and [positive]
+ are bijections. *)
+
+Require Import Hexadecimal HexadecimalFacts PArith NArith.
+
+Module Unsigned.
+
+Local Open Scope N.
+
+(** A direct version of [of_little_uint] *)
+Fixpoint of_lu (d:uint) : N :=
+ match d with
+ | Nil => 0
+ | D0 d => 0x10 * of_lu d
+ | D1 d => 0x1 + 0x10 * of_lu d
+ | D2 d => 0x2 + 0x10 * of_lu d
+ | D3 d => 0x3 + 0x10 * of_lu d
+ | D4 d => 0x4 + 0x10 * of_lu d
+ | D5 d => 0x5 + 0x10 * of_lu d
+ | D6 d => 0x6 + 0x10 * of_lu d
+ | D7 d => 0x7 + 0x10 * of_lu d
+ | D8 d => 0x8 + 0x10 * of_lu d
+ | D9 d => 0x9 + 0x10 * of_lu d
+ | Da d => 0xa + 0x10 * of_lu d
+ | Db d => 0xb + 0x10 * of_lu d
+ | Dc d => 0xc + 0x10 * of_lu d
+ | Dd d => 0xd + 0x10 * of_lu d
+ | De d => 0xe + 0x10 * of_lu d
+ | Df d => 0xf + 0x10 * of_lu d
+ end.
+
+Definition hd d :=
+ match d with
+ | Nil => 0x0
+ | D0 _ => 0x0
+ | D1 _ => 0x1
+ | D2 _ => 0x2
+ | D3 _ => 0x3
+ | D4 _ => 0x4
+ | D5 _ => 0x5
+ | D6 _ => 0x6
+ | D7 _ => 0x7
+ | D8 _ => 0x8
+ | D9 _ => 0x9
+ | Da _ => 0xa
+ | Db _ => 0xb
+ | Dc _ => 0xc
+ | Dd _ => 0xd
+ | De _ => 0xe
+ | Df _ => 0xf
+ end.
+
+Definition tl d :=
+ match d with
+ | Nil => d
+ | D0 d | D1 d | D2 d | D3 d | D4 d | D5 d | D6 d | D7 d | D8 d | D9 d
+ | Da d | Db d | Dc d | Dd d | De d | Df d => d
+ end.
+
+Lemma of_lu_eqn d :
+ of_lu d = hd d + 0x10 * (of_lu (tl d)).
+Proof.
+ induction d; simpl; trivial.
+Qed.
+
+Ltac simpl_of_lu :=
+ match goal with
+ | |- context [ of_lu (?f ?x) ] =>
+ rewrite (of_lu_eqn (f x)); simpl hd; simpl tl
+ end.
+
+Fixpoint usize (d:uint) : N :=
+ match d with
+ | Nil => 0
+ | D0 d => N.succ (usize d)
+ | D1 d => N.succ (usize d)
+ | D2 d => N.succ (usize d)
+ | D3 d => N.succ (usize d)
+ | D4 d => N.succ (usize d)
+ | D5 d => N.succ (usize d)
+ | D6 d => N.succ (usize d)
+ | D7 d => N.succ (usize d)
+ | D8 d => N.succ (usize d)
+ | D9 d => N.succ (usize d)
+ | Da d => N.succ (usize d)
+ | Db d => N.succ (usize d)
+ | Dc d => N.succ (usize d)
+ | Dd d => N.succ (usize d)
+ | De d => N.succ (usize d)
+ | Df d => N.succ (usize d)
+ end.
+
+Lemma of_lu_revapp d d' :
+ of_lu (revapp d d') =
+ of_lu (rev d) + of_lu d' * 0x10^usize d.
+Proof.
+ revert d'.
+ induction d; simpl; intro d'; [ now rewrite N.mul_1_r | .. ];
+ unfold rev; simpl revapp; rewrite 2 IHd;
+ rewrite <- N.add_assoc; f_equal; simpl_of_lu; simpl of_lu;
+ rewrite N.pow_succ_r'; ring.
+Qed.
+
+Definition Nadd n p :=
+ match n with
+ | N0 => p
+ | Npos p0 => (p0+p)%positive
+ end.
+
+Lemma Nadd_simpl n p q : Npos (Nadd n (p * q)) = n + Npos p * Npos q.
+Proof.
+ now destruct n.
+Qed.
+
+Lemma of_uint_acc_eqn d acc : d<>Nil ->
+ Pos.of_hex_uint_acc d acc = Pos.of_hex_uint_acc (tl d) (Nadd (hd d) (0x10*acc)).
+Proof.
+ destruct d; simpl; trivial. now destruct 1.
+Qed.
+
+Lemma of_uint_acc_rev d acc :
+ Npos (Pos.of_hex_uint_acc d acc) =
+ of_lu (rev d) + (Npos acc) * 0x10^usize d.
+Proof.
+ revert acc.
+ induction d; intros; simpl usize;
+ [ simpl; now rewrite Pos.mul_1_r | .. ];
+ rewrite N.pow_succ_r';
+ unfold rev; simpl revapp; try rewrite of_lu_revapp; simpl of_lu;
+ rewrite of_uint_acc_eqn by easy; simpl tl; simpl hd;
+ rewrite IHd, Nadd_simpl; ring.
+Qed.
+
+Lemma of_uint_alt d : Pos.of_hex_uint d = of_lu (rev d).
+Proof.
+ induction d; simpl; trivial; unfold rev; simpl revapp;
+ rewrite of_lu_revapp; simpl of_lu; try apply of_uint_acc_rev.
+ rewrite IHd. ring.
+Qed.
+
+Lemma of_lu_rev d : Pos.of_hex_uint (rev d) = of_lu d.
+Proof.
+ rewrite of_uint_alt. now rewrite rev_rev.
+Qed.
+
+Lemma of_lu_double_gen d :
+ of_lu (Little.double d) = N.double (of_lu d) /\
+ of_lu (Little.succ_double d) = N.succ_double (of_lu d).
+Proof.
+ rewrite N.double_spec, N.succ_double_spec.
+ induction d; try destruct IHd as (IH1,IH2);
+ simpl Little.double; simpl Little.succ_double;
+ repeat (simpl_of_lu; rewrite ?IH1, ?IH2); split; reflexivity || ring.
+Qed.
+
+Lemma of_lu_double d :
+ of_lu (Little.double d) = N.double (of_lu d).
+Proof.
+ apply of_lu_double_gen.
+Qed.
+
+Lemma of_lu_succ_double d :
+ of_lu (Little.succ_double d) = N.succ_double (of_lu d).
+Proof.
+ apply of_lu_double_gen.
+Qed.
+
+(** First bijection result *)
+
+Lemma of_to (p:positive) : Pos.of_hex_uint (Pos.to_hex_uint p) = Npos p.
+Proof.
+ unfold Pos.to_hex_uint.
+ rewrite of_lu_rev.
+ induction p; simpl; trivial.
+ - now rewrite of_lu_succ_double, IHp.
+ - now rewrite of_lu_double, IHp.
+Qed.
+
+(** The other direction *)
+
+Definition to_lu n :=
+ match n with
+ | N0 => Hexadecimal.zero
+ | Npos p => Pos.to_little_hex_uint p
+ end.
+
+Lemma succ_double_alt d :
+ Little.succ_double d = Little.succ (Little.double d).
+Proof.
+ now induction d.
+Qed.
+
+Lemma double_succ d :
+ Little.double (Little.succ d) =
+ Little.succ (Little.succ_double d).
+Proof.
+ induction d; simpl; f_equal; auto using succ_double_alt.
+Qed.
+
+Lemma to_lu_succ n :
+ to_lu (N.succ n) = Little.succ (to_lu n).
+Proof.
+ destruct n; simpl; trivial.
+ induction p; simpl; rewrite ?IHp;
+ auto using succ_double_alt, double_succ.
+Qed.
+
+Lemma nat_iter_S n {A} (f:A->A) i :
+ Nat.iter (S n) f i = f (Nat.iter n f i).
+Proof.
+ reflexivity.
+Qed.
+
+Lemma nat_iter_0 {A} (f:A->A) i : Nat.iter 0 f i = i.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma to_lhex_tenfold p :
+ to_lu (0x10 * Npos p) = D0 (to_lu (Npos p)).
+Proof.
+ induction p using Pos.peano_rect.
+ - trivial.
+ - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
+ rewrite N.mul_succ_r.
+ change 0x10 at 2 with (Nat.iter 0x10%nat N.succ 0).
+ rewrite ?nat_iter_S, nat_iter_0.
+ rewrite !N.add_succ_r, N.add_0_r, !to_lu_succ, IHp.
+ destruct (to_lu (N.pos p)); simpl; auto.
+Qed.
+
+Lemma of_lu_0 d : of_lu d = 0 <-> nztail d = Nil.
+Proof.
+ induction d; try simpl_of_lu; split; trivial; try discriminate;
+ try (intros H; now apply N.eq_add_0 in H).
+ - rewrite N.add_0_l. intros H.
+ apply N.eq_mul_0_r in H; [|easy]. rewrite IHd in H.
+ simpl. now rewrite H.
+ - simpl. destruct (nztail d); try discriminate.
+ now destruct IHd as [_ ->].
+Qed.
+
+Lemma to_of_lu_tenfold d :
+ to_lu (of_lu d) = lnorm d ->
+ to_lu (0x10 * of_lu d) = lnorm (D0 d).
+Proof.
+ intro IH.
+ destruct (N.eq_dec (of_lu d) 0) as [H|H].
+ - rewrite H. simpl. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now rewrite H.
+ - destruct (of_lu d) eqn:Eq; [easy| ].
+ rewrite to_lhex_tenfold; auto. rewrite IH.
+ rewrite <- Eq in H. rewrite of_lu_0 in H.
+ unfold lnorm. simpl. now destruct (nztail d).
+Qed.
+
+Lemma Nadd_alt n m : n + m = Nat.iter (N.to_nat n) N.succ m.
+Proof.
+ destruct n. trivial.
+ induction p using Pos.peano_rect.
+ - now rewrite N.add_1_l.
+ - change (N.pos (Pos.succ p)) with (N.succ (N.pos p)).
+ now rewrite N.add_succ_l, IHp, N2Nat.inj_succ.
+Qed.
+
+Ltac simpl_to_nat := simpl N.to_nat; unfold Pos.to_nat; simpl Pos.iter_op.
+
+Lemma to_of_lu d : to_lu (of_lu d) = lnorm d.
+Proof.
+ induction d; [reflexivity|..];
+ simpl_of_lu; rewrite Nadd_alt; simpl_to_nat;
+ rewrite ?nat_iter_S, nat_iter_0, ?to_lu_succ, to_of_lu_tenfold by assumption;
+ unfold lnorm; simpl; destruct nztail; auto.
+Qed.
+
+(** Second bijection result *)
+
+Lemma to_of (d:uint) : N.to_hex_uint (Pos.of_hex_uint d) = unorm d.
+Proof.
+ rewrite of_uint_alt.
+ unfold N.to_hex_uint, Pos.to_hex_uint.
+ destruct (of_lu (rev d)) eqn:H.
+ - rewrite of_lu_0 in H. rewrite <- rev_lnorm_rev.
+ unfold lnorm. now rewrite H.
+ - change (Pos.to_little_hex_uint p) with (to_lu (N.pos p)).
+ rewrite <- H. rewrite to_of_lu. apply rev_lnorm_rev.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_uint_nonzero p : Pos.to_hex_uint p <> zero.
+Proof.
+ intro E. generalize (of_to p). now rewrite E.
+Qed.
+
+Lemma to_uint_nonnil p : Pos.to_hex_uint p <> Nil.
+Proof.
+ intros E. generalize (of_to p). now rewrite E.
+Qed.
+
+Lemma to_uint_inj p p' : Pos.to_hex_uint p = Pos.to_hex_uint p' -> p = p'.
+Proof.
+ intro E.
+ assert (E' : N.pos p = N.pos p').
+ { now rewrite <- (of_to p), <- (of_to p'), E. }
+ now injection E'.
+Qed.
+
+Lemma to_uint_pos_surj d :
+ unorm d<>zero -> exists p, Pos.to_hex_uint p = unorm d.
+Proof.
+ intros.
+ destruct (Pos.of_hex_uint d) eqn:E.
+ - destruct H. generalize (to_of d). now rewrite E.
+ - exists p. generalize (to_of d). now rewrite E.
+Qed.
+
+Lemma of_uint_norm d : Pos.of_hex_uint (unorm d) = Pos.of_hex_uint d.
+Proof.
+ now induction d.
+Qed.
+
+Lemma of_inj d d' :
+ Pos.of_hex_uint d = Pos.of_hex_uint d' -> unorm d = unorm d'.
+Proof.
+ intros. rewrite <- !to_of. now f_equal.
+Qed.
+
+Lemma of_iff d d' : Pos.of_hex_uint d = Pos.of_hex_uint d' <-> unorm d = unorm d'.
+Proof.
+ split. apply of_inj. intros E. rewrite <- of_uint_norm, E.
+ apply of_uint_norm.
+Qed.
+
+(* various lemmas *)
+
+Lemma nztail_to_hex_uint p :
+ let (h, n) := Hexadecimal.nztail (Pos.to_hex_uint p) in
+ Npos p = Pos.of_hex_uint h * 0x10^(N.of_nat n).
+Proof.
+ rewrite <-(of_to p), <-(rev_rev (Pos.to_hex_uint p)), of_lu_rev.
+ unfold Hexadecimal.nztail.
+ rewrite rev_rev.
+ induction (rev (Pos.to_hex_uint p)); [reflexivity| |
+ now simpl N.of_nat; simpl N.pow; rewrite N.mul_1_r, of_lu_rev..].
+ revert IHu.
+ set (t := _ u); case t; clear t; intros u0 n H.
+ rewrite of_lu_eqn; unfold hd, tl.
+ rewrite N.add_0_l, H, Nat2N.inj_succ, N.pow_succ_r'; ring.
+Qed.
+
+Definition double d := rev (Little.double (rev d)).
+
+Lemma double_unorm d : double (unorm d) = unorm (double d).
+Proof.
+ unfold double.
+ rewrite <-!rev_lnorm_rev, !rev_rev, <-!to_of_lu, of_lu_double.
+ now case of_lu; [now simpl|]; intro p; induction p.
+Qed.
+
+Lemma double_nzhead d : double (nzhead d) = nzhead (double d).
+Proof.
+ unfold double.
+ rewrite <-!rev_nztail_rev, !rev_rev.
+ apply f_equal; generalize (rev d); clear d; intro d.
+ cut (Little.double (nztail d) = nztail (Little.double d)
+ /\ Little.succ_double (nztail d) = nztail (Little.succ_double d)).
+ { now simpl. }
+ now induction d;
+ [|split; simpl; rewrite <-?(proj1 IHd), <-?(proj2 IHd); case nztail..].
+Qed.
+
+Lemma of_hex_uint_double d :
+ Pos.of_hex_uint (double d) = N.double (Pos.of_hex_uint d).
+Proof.
+ now unfold double; rewrite of_lu_rev, of_lu_double, <-of_lu_rev, rev_rev.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed decimal numbers *)
+
+Module Signed.
+
+Lemma of_to (p:positive) : Pos.of_hex_int (Pos.to_hex_int p) = Some p.
+Proof.
+ unfold Pos.to_hex_int, Pos.of_hex_int, norm.
+ now rewrite Unsigned.of_to.
+Qed.
+
+Lemma to_of (d:int)(p:positive) :
+ Pos.of_hex_int d = Some p -> Pos.to_hex_int p = norm d.
+Proof.
+ unfold Pos.of_hex_int.
+ destruct d; [ | intros [=]].
+ simpl norm. rewrite <- Unsigned.to_of.
+ destruct (Pos.of_hex_uint d); now intros [= <-].
+Qed.
+
+Lemma to_int_inj p p' : Pos.to_hex_int p = Pos.to_hex_int p' -> p = p'.
+Proof.
+ intro E.
+ assert (E' : Some p = Some p').
+ { now rewrite <- (of_to p), <- (of_to p'), E. }
+ now injection E'.
+Qed.
+
+Lemma to_int_pos_surj d :
+ unorm d <> zero -> exists p, Pos.to_hex_int p = norm (Pos d).
+Proof.
+ simpl. unfold Pos.to_hex_int. intros H.
+ destruct (Unsigned.to_uint_pos_surj d H) as (p,Hp).
+ exists p. now f_equal.
+Qed.
+
+Lemma of_int_norm d : Pos.of_hex_int (norm d) = Pos.of_hex_int d.
+Proof.
+ unfold Pos.of_int.
+ destruct d.
+ - simpl. now rewrite Unsigned.of_uint_norm.
+ - simpl. now destruct (nzhead d) eqn:H.
+Qed.
+
+Lemma of_inj_pos d d' :
+ Pos.of_hex_int (Pos d) = Pos.of_hex_int (Pos d') -> unorm d = unorm d'.
+Proof.
+ unfold Pos.of_hex_int.
+ destruct (Pos.of_hex_uint d) eqn:Hd, (Pos.of_hex_uint d') eqn:Hd';
+ intros [=].
+ - apply Unsigned.of_inj; now rewrite Hd, Hd'.
+ - apply Unsigned.of_inj; rewrite Hd, Hd'; now f_equal.
+Qed.
+
+End Signed.
diff --git a/theories/Numbers/HexadecimalQ.v b/theories/Numbers/HexadecimalQ.v
new file mode 100644
index 0000000000..9bf43ceb88
--- /dev/null
+++ b/theories/Numbers/HexadecimalQ.v
@@ -0,0 +1,533 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalQ
+
+ Proofs that conversions between hexadecimal numbers and [Q]
+ are bijections. *)
+
+Require Import Decimal DecimalFacts DecimalPos DecimalN DecimalZ.
+Require Import Hexadecimal HexadecimalFacts HexadecimalPos HexadecimalN HexadecimalZ QArith.
+
+Lemma of_to (q:Q) : forall d, to_hexadecimal q = Some d -> of_hexadecimal d = q.
+Proof.
+ cut (match to_hexadecimal q with None => True | Some d => of_hexadecimal d = q end).
+ { now case to_hexadecimal; [intros d <- d' Hd'; injection Hd'; intros ->|]. }
+ destruct q as (num, den).
+ unfold to_hexadecimal; simpl Qnum; simpl Qden.
+ generalize (HexadecimalPos.Unsigned.nztail_to_hex_uint den).
+ case Hexadecimal.nztail; intros u n.
+ change 16%N with (2^4)%N; rewrite <-N.pow_mul_r.
+ change 4%N with (N.of_nat 4); rewrite <-Nnat.Nat2N.inj_mul.
+ change 4%Z with (Z.of_nat 4); rewrite <-Nat2Z.inj_mul.
+ case u; clear u; try (intros; exact I); [| | |]; intro u;
+ (case u; clear u; [|intros; exact I..]).
+ - unfold Pos.of_hex_uint, Pos.of_hex_uint_acc; rewrite N.mul_1_l.
+ case n.
+ + unfold of_hexadecimal, app_int, app, Z.to_hex_int; simpl.
+ intro H; inversion H as (H1); clear H H1.
+ case num; [reflexivity|intro pnum; fold (rev (rev (Pos.to_hex_uint pnum)))..].
+ * rewrite rev_rev; simpl.
+ now unfold Z.of_hex_uint; rewrite HexadecimalPos.Unsigned.of_to.
+ * rewrite rev_rev; simpl.
+ now unfold Z.of_hex_uint; rewrite HexadecimalPos.Unsigned.of_to.
+ + clear n; intros n.
+ intro H; injection H; intros ->; clear H.
+ unfold of_hexadecimal.
+ rewrite DecimalZ.of_to.
+ simpl nb_digits; rewrite Nat2Z.inj_0, Z.mul_0_r, Z.sub_0_r.
+ now apply f_equal2; [rewrite app_int_nil_r, of_to|].
+ - unfold Pos.of_hex_uint, Pos.of_hex_uint_acc.
+ rewrite <-N.pow_succ_r', <-Nnat.Nat2N.inj_succ.
+ intro H; injection H; intros ->; clear H.
+ fold (4 * n)%nat.
+ change 1%Z with (Z.of_nat 1); rewrite <-Znat.Nat2Z.inj_add.
+ unfold of_hexadecimal.
+ rewrite DecimalZ.of_to.
+ simpl nb_digits; rewrite Nat2Z.inj_0, Z.mul_0_r, Z.sub_0_r.
+ now apply f_equal2; [rewrite app_int_nil_r, of_to|].
+ - change 2%Z with (Z.of_nat 2); rewrite <-Znat.Nat2Z.inj_add.
+ unfold Pos.of_hex_uint, Pos.of_hex_uint_acc.
+ change 4%N with (2^2)%N; rewrite <-N.pow_add_r.
+ change 2%N with (N.of_nat 2); rewrite <-Nnat.Nat2N.inj_add.
+ intro H; injection H; intros ->; clear H.
+ fold (4 * n)%nat.
+ unfold of_hexadecimal.
+ rewrite DecimalZ.of_to.
+ simpl nb_digits; rewrite Nat2Z.inj_0, Z.mul_0_r, Z.sub_0_r.
+ now apply f_equal2; [rewrite app_int_nil_r, of_to|].
+ - change 3%Z with (Z.of_nat 3); rewrite <-Znat.Nat2Z.inj_add.
+ unfold Pos.of_hex_uint, Pos.of_hex_uint_acc.
+ change 8%N with (2^3)%N; rewrite <-N.pow_add_r.
+ change 3%N with (N.of_nat 3); rewrite <-Nnat.Nat2N.inj_add.
+ intro H; injection H; intros ->; clear H.
+ fold (4 * n)%nat.
+ unfold of_hexadecimal.
+ rewrite DecimalZ.of_to.
+ simpl nb_digits; rewrite Nat2Z.inj_0, Z.mul_0_r, Z.sub_0_r.
+ now apply f_equal2; [rewrite app_int_nil_r, of_to|].
+Qed.
+
+(* normalize without fractional part, for instance norme 0x1.2p-1 is 0x12e-5 *)
+Definition hnorme (d:hexadecimal) : hexadecimal :=
+ let '(i, f, e) :=
+ match d with
+ | Hexadecimal i f => (i, f, Decimal.Pos Decimal.Nil)
+ | HexadecimalExp i f e => (i, f, e)
+ end in
+ let i := norm (app_int i f) in
+ let e := (Z.of_int e - 4 * Z.of_nat (nb_digits f))%Z in
+ match e with
+ | Z0 => Hexadecimal i Nil
+ | Zpos e => Hexadecimal (Pos.iter double i e) Nil
+ | Zneg _ => HexadecimalExp i Nil (Decimal.norm (Z.to_int e))
+ end.
+
+Lemma hnorme_spec d :
+ match hnorme d with
+ | Hexadecimal i Nil => i = norm i
+ | HexadecimalExp i Nil e =>
+ i = norm i /\ e = Decimal.norm e /\ e <> Decimal.Pos Decimal.zero
+ | _ => False
+ end.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold hnorme; simpl.
+ - case_eq (nb_digits f); [now simpl; rewrite norm_invol|]; intros nf Hnf.
+ split; [now simpl; rewrite norm_invol|].
+ unfold Z.of_nat.
+ now rewrite <-!DecimalZ.to_of, !DecimalZ.of_to.
+ - set (e' := (_ - _)%Z).
+ case_eq e'; [|intro pe'..]; intro He'.
+ + now rewrite norm_invol.
+ + rewrite Pos2Nat.inj_iter.
+ set (ne' := Pos.to_nat pe').
+ fold (Nat.iter ne' double (norm (app_int i f))).
+ induction ne'; [now simpl; rewrite norm_invol|].
+ now rewrite Unsigned.nat_iter_S, <-double_norm, IHne', norm_invol.
+ + split; [now rewrite norm_invol|].
+ split; [now rewrite DecimalFacts.norm_invol|].
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ change (Decimal.Pos _) with (Z.to_int 0).
+ now intro H; generalize (DecimalZ.to_int_inj _ _ H).
+Qed.
+
+Lemma hnorme_invol d : hnorme (hnorme d) = hnorme d.
+Proof.
+ case d; clear d; intros i f; [|intro e]; unfold hnorme; simpl.
+ - case_eq (nb_digits f); [now simpl; rewrite app_int_nil_r, norm_invol|].
+ intros nf Hnf.
+ unfold Z.of_nat.
+ simpl.
+ set (pnf := Pos.to_uint _).
+ set (nz := Decimal.nzhead pnf).
+ assert (Hnz : nz <> Decimal.Nil).
+ { unfold nz, pnf.
+ rewrite <-DecimalFacts.unorm_0.
+ rewrite <-DecimalPos.Unsigned.to_of.
+ rewrite DecimalPos.Unsigned.of_to.
+ change Decimal.zero with (N.to_uint 0).
+ now intro H; generalize (DecimalN.Unsigned.to_uint_inj _ _ H). }
+ set (m := match nz with Decimal.Nil => _ | _ => _ end).
+ assert (Hm : m = (Decimal.Neg (Decimal.unorm pnf))).
+ { now revert Hnz; unfold m, nz, Decimal.unorm; fold nz; case nz. }
+ rewrite Hm; unfold pnf.
+ rewrite <-DecimalPos.Unsigned.to_of, DecimalPos.Unsigned.of_to.
+ simpl; unfold Z.of_uint; rewrite DecimalPos.Unsigned.of_to.
+ rewrite Z.sub_0_r; simpl.
+ fold pnf; fold nz; fold m; rewrite Hm; unfold pnf.
+ rewrite <-DecimalPos.Unsigned.to_of, DecimalPos.Unsigned.of_to.
+ now rewrite app_int_nil_r, norm_invol.
+ - set (e' := (_ - _)%Z).
+ case_eq e'; [|intro pe'..]; intro Hpe'.
+ + now simpl; rewrite app_int_nil_r, norm_invol.
+ + simpl; rewrite app_int_nil_r.
+ apply f_equal2; [|reflexivity].
+ rewrite Pos2Nat.inj_iter.
+ set (ne' := Pos.to_nat pe').
+ fold (Nat.iter ne' double (norm (app_int i f))).
+ induction ne'; [now simpl; rewrite norm_invol|].
+ now rewrite Unsigned.nat_iter_S, <-double_norm, IHne'.
+ + rewrite <-DecimalZ.to_of, !DecimalZ.of_to; simpl.
+ rewrite app_int_nil_r, norm_invol.
+ set (pnf := Pos.to_uint _).
+ set (nz := Decimal.nzhead pnf).
+ assert (Hnz : nz <> Decimal.Nil).
+ { unfold nz, pnf.
+ rewrite <-DecimalFacts.unorm_0.
+ rewrite <-DecimalPos.Unsigned.to_of.
+ rewrite DecimalPos.Unsigned.of_to.
+ change Decimal.zero with (N.to_uint 0).
+ now intro H; generalize (DecimalN.Unsigned.to_uint_inj _ _ H). }
+ set (m := match nz with Decimal.Nil => _ | _ => _ end).
+ assert (Hm : m = (Decimal.Neg (Decimal.unorm pnf))).
+ { now revert Hnz; unfold m, nz, Decimal.unorm; fold nz; case nz. }
+ rewrite Hm; unfold pnf.
+ now rewrite <-DecimalPos.Unsigned.to_of, DecimalPos.Unsigned.of_to.
+Qed.
+
+Lemma to_of (d:hexadecimal) :
+ to_hexadecimal (of_hexadecimal d) = Some (hnorme d).
+Proof.
+ unfold to_hexadecimal.
+ pose (t10 := fun y => (y~0~0~0~0)%positive).
+ assert (H : exists h e_den,
+ Hexadecimal.nztail (Pos.to_hex_uint (Qden (of_hexadecimal d)))
+ = (h, e_den)
+ /\ (h = D1 Nil \/ h = D2 Nil \/ h = D4 Nil \/ h = D8 Nil)).
+ { assert (H : forall p,
+ Hexadecimal.nztail (Pos.to_hex_uint (Pos.iter (Pos.mul 2) 1%positive p))
+ = ((match (Pos.to_nat p) mod 4 with 0%nat => D1 | 1 => D2 | 2 => D4 | _ => D8 end)%nat Nil,
+ (Pos.to_nat p / 4)%nat)).
+ { intro p; clear d; rewrite Pos2Nat.inj_iter.
+ fold (Nat.iter (Pos.to_nat p) (Pos.mul 2) 1%positive).
+ set (n := Pos.to_nat p).
+ fold (Nat.iter n t10 1%positive).
+ set (nm4 := (n mod 4)%nat); set (nd4 := (n / 4)%nat).
+ rewrite (Nat.div_mod n 4); [|now simpl].
+ unfold nm4, nd4; clear nm4 nd4.
+ generalize (Nat.mod_upper_bound n 4 ltac:(now simpl)).
+ generalize (n mod 4); generalize (n / 4)%nat.
+ intros d r Hr; clear p n.
+ induction d.
+ { simpl; revert Hr.
+ do 4 (case r; [now simpl|clear r; intro r]).
+ intro H; exfalso.
+ now do 4 (generalize (lt_S_n _ _ H); clear H; intro H). }
+ rewrite Nat.mul_succ_r, <-Nat.add_assoc, (Nat.add_comm 4), Nat.add_assoc.
+ rewrite (Nat.add_comm _ 4).
+ change (4 + _)%nat with (S (S (S (S (4 * d + r))))).
+ rewrite !Unsigned.nat_iter_S.
+ rewrite !Pos.mul_assoc.
+ unfold Pos.to_hex_uint.
+ change (2 * 2 * 2 * 2)%positive with 0x10%positive.
+ set (n := Nat.iter _ _ _).
+ change (Pos.to_little_hex_uint _) with (Unsigned.to_lu (16 * N.pos n)).
+ rewrite Unsigned.to_lhex_tenfold.
+ unfold Hexadecimal.nztail; rewrite rev_rev.
+ rewrite <-(rev_rev (Unsigned.to_lu _)).
+ set (m := _ (rev _)).
+ replace m with (let (r, n) := let (r, n) := m in (rev r, n) in (rev r, n)).
+ 2:{ now case m; intros r' n'; rewrite rev_rev. }
+ change (let (r, n) := m in (rev r, n))
+ with (Hexadecimal.nztail (Pos.to_hex_uint n)).
+ now unfold n; rewrite IHd, rev_rev; clear n m. }
+ unfold of_hexadecimal.
+ case d; intros i f; [|intro e]; unfold of_hexadecimal; simpl.
+ - case (Z.of_nat _)%Z; [|intro p..];
+ [now exists (D1 Nil), O; split; [|left]
+ | |now exists (D1 Nil), O; split; [|left]].
+ exists (D1 Nil), (Pos.to_nat p).
+ split; [|now left]; simpl.
+ change (Pos.iter _ _ _) with (Pos.iter (Pos.mul 2) 1%positive (4 * p)).
+ rewrite H.
+ rewrite Pos2Nat.inj_mul, Nat.mul_comm, Nat.div_mul; [|now simpl].
+ now rewrite Nat.mod_mul; [|now simpl].
+ - case (_ - _)%Z; [|intros p..]; [now exists (D1 Nil), O; split; [|left]..|].
+ simpl Qden; rewrite H.
+ eexists; eexists; split; [reflexivity|].
+ case (_ mod _); [now left|intro n].
+ case n; [now right; left|clear n; intro n].
+ case n; [now right; right; left|clear n; intro n].
+ now right; right; right. }
+ generalize (HexadecimalPos.Unsigned.nztail_to_hex_uint (Qden (of_hexadecimal d))).
+ destruct H as (h, (e, (He, Hh))); rewrite He; clear He.
+ assert (Hn1 : forall p, N.pos (Pos.iter (Pos.mul 2) 1%positive p) = 1%N -> False).
+ { intro p.
+ rewrite Pos2Nat.inj_iter.
+ case_eq (Pos.to_nat p); [|now simpl].
+ intro H; exfalso; apply (lt_irrefl O).
+ rewrite <-H at 2; apply Pos2Nat.is_pos. }
+ assert (H16_2 : forall p, (16^p = 2^(4 * p))%positive).
+ { intro p.
+ apply (@f_equal _ _ (fun z => match z with Z.pos p => p | _ => 1%positive end)
+ (Z.pos _) (Z.pos _)).
+ rewrite !Pos2Z.inj_pow_pos, !Z.pow_pos_fold, Pos2Z.inj_mul.
+ now change 16%Z with (2^4)%Z; rewrite <-Z.pow_mul_r. }
+ assert (HN16_2 : forall n, (16^n = 2^(4 * n))%N).
+ { intro n.
+ apply N2Z.inj; rewrite !N2Z.inj_pow, N2Z.inj_mul.
+ change (Z.of_N 16) with (2^4)%Z.
+ now rewrite <-Z.pow_mul_r; [| |apply N2Z.is_nonneg]. }
+ assert (Hn1' : forall p, N.pos (Pos.iter (Pos.mul 16) 1%positive p) = 1%N -> False).
+ { intro p; fold (16^p)%positive; rewrite H16_2; apply Hn1. }
+ assert (Ht10inj : forall n m, t10 n = t10 m -> n = m).
+ { intros n m H; generalize (f_equal Z.pos H); clear H.
+ change (Z.pos (t10 n)) with (Z.mul 0x10 (Z.pos n)).
+ change (Z.pos (t10 m)) with (Z.mul 0x10 (Z.pos m)).
+ rewrite Z.mul_comm, (Z.mul_comm 0x10).
+ intro H; generalize (f_equal (fun z => Z.div z 0x10) H); clear H.
+ now rewrite !Z.div_mul; [|now simpl..]; intro H; inversion H. }
+ assert (Ht2inj : forall n m, Pos.mul 2 n = Pos.mul 2 m -> n = m).
+ { intros n m H; generalize (f_equal Z.pos H); clear H.
+ change (Z.pos (Pos.mul 2 n)) with (Z.mul 2 (Z.pos n)).
+ change (Z.pos (Pos.mul 2 m)) with (Z.mul 2 (Z.pos m)).
+ rewrite Z.mul_comm, (Z.mul_comm 2).
+ intro H; generalize (f_equal (fun z => Z.div z 2) H); clear H.
+ now rewrite !Z.div_mul; [|now simpl..]; intro H; inversion H. }
+ assert (Hinj : forall n m,
+ Nat.iter n (Pos.mul 2) 1%positive = Nat.iter m (Pos.mul 2) 1%positive
+ -> n = m).
+ { induction n; [now intro m; case m|].
+ intro m; case m; [now simpl|]; clear m; intro m.
+ rewrite !Unsigned.nat_iter_S.
+ intro H; generalize (Ht2inj _ _ H); clear H; intro H.
+ now rewrite (IHn _ H). }
+ change 4%Z with (Z.of_nat 4); rewrite <-Nat2Z.inj_mul.
+ change 1%Z with (Z.of_nat 1); rewrite <-Nat2Z.inj_add.
+ change 2%Z with (Z.of_nat 2); rewrite <-Nat2Z.inj_add.
+ change 3%Z with (Z.of_nat 3); rewrite <-Nat2Z.inj_add.
+ destruct Hh as [Hh|[Hh|[Hh|Hh]]]; rewrite Hh; clear h Hh.
+ - case e; clear e; [|intro e]; simpl; unfold of_hexadecimal, hnorme.
+ + case d; clear d; intros i f; [|intro e].
+ * generalize (nb_digits_pos f).
+ case f;
+ [|now clear f; intro f; intros H1 H2; exfalso; revert H1 H2;
+ case nb_digits;
+ [intros H _; apply (lt_irrefl O), H|intros n _; apply Hn1]..].
+ now intros _ _; simpl; rewrite to_of.
+ * rewrite <-DecimalZ.to_of, DecimalZ.of_to.
+ set (emf := (_ - _)%Z).
+ case_eq emf; [|intro pemf..].
+ ++ now simpl; rewrite to_of.
+ ++ intros Hemf _; simpl.
+ apply f_equal, f_equal2; [|reflexivity].
+ rewrite !Pos2Nat.inj_iter.
+ fold (Nat.iter (Pos.to_nat pemf) (Z.mul 2) (Z.of_hex_int (app_int i f))).
+ fold (Nat.iter (Pos.to_nat pemf) double (norm (app_int i f))).
+ induction Pos.to_nat; [now simpl; rewrite HexadecimalZ.to_of|].
+ now rewrite !Unsigned.nat_iter_S, <-IHn, double_to_hex_int.
+ ++ simpl Qden; intros _ H; exfalso; revert H; apply Hn1.
+ + case d; clear d; intros i f; [|intro e'].
+ * simpl; case_eq (nb_digits f); [|intros nf' Hnf'];
+ [now simpl; intros _ H; exfalso; symmetry in H; revert H; apply Hn1'|].
+ unfold Z.of_nat, Z.opp, Qnum, Qden.
+ rewrite H16_2.
+ fold (Pos.mul 2); fold (2^(Pos.of_succ_nat nf')~0~0)%positive.
+ intro H; injection H; clear H.
+ unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intro H; injection H.
+ clear H; intro H; generalize (SuccNat2Pos.inj _ _ H); clear H.
+ intros <-.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite SuccNat2Pos.id_succ.
+ change (_~0)%positive with (4 * Pos.of_succ_nat nf')%positive.
+ now rewrite Pos2Nat.inj_mul, SuccNat2Pos.id_succ.
+ * set (nemf := (_ - _)%Z); intro H.
+ assert (H' : exists pnemf, nemf = Z.neg pnemf); [|revert H].
+ { revert H; case nemf; [|intro pnemf..]; [..|now intros _; exists pnemf];
+ simpl Qden; intro H; exfalso; symmetry in H; revert H; apply Hn1'. }
+ destruct H' as (pnemf,Hpnemf); rewrite Hpnemf.
+ unfold Qnum, Qden.
+ rewrite H16_2.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H.
+ intro H; revert Hpnemf; rewrite H; clear pnemf H; intro Hnemf.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite SuccNat2Pos.id_succ.
+ change (_~0)%positive with (4 * Pos.of_succ_nat e)%positive.
+ now rewrite Pos2Nat.inj_mul, SuccNat2Pos.id_succ.
+ - simpl Pos.of_hex_uint.
+ rewrite HN16_2.
+ rewrite <-N.pow_succ_r; [|now apply N.le_0_l].
+ rewrite <-N.succ_pos_spec.
+ case d; clear d; intros i f; [|intro e']; unfold of_hexadecimal, hnorme.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+ - simpl Pos.of_hex_uint.
+ rewrite HN16_2.
+ change 4%N with (2 * 2)%N at 1; rewrite <-!N.mul_assoc.
+ do 2 (rewrite <-N.pow_succ_r; [|now apply N.le_0_l]).
+ rewrite <-N.succ_pos_spec.
+ case d; clear d; intros i f; [|intro e']; unfold of_hexadecimal, hnorme.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite <-SuccNat2Pos.inj_succ.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite <-SuccNat2Pos.inj_succ.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+ - simpl Pos.of_hex_uint.
+ rewrite HN16_2.
+ change 8%N with (2 * 2 * 2)%N; rewrite <-!N.mul_assoc.
+ do 3 (rewrite <-N.pow_succ_r; [|now apply N.le_0_l]).
+ rewrite <-N.succ_pos_spec.
+ case d; clear d; intros i f; [|intro e']; unfold of_hexadecimal, hnorme.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite <-!SuccNat2Pos.inj_succ.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+ + set (em4f := (_ - _)%Z).
+ case_eq em4f; [|intros pem4f..]; intro Hpem4f;
+ [now simpl; intros H; exfalso; symmetry in H; revert H; apply Hn1..|].
+ unfold Qnum, Qden.
+ intro H; injection H; clear H; unfold Pos.pow; rewrite !Pos2Nat.inj_iter.
+ intro H; generalize (Hinj _ _ H); clear H; intro H.
+ generalize (Pos2Nat.inj _ _ H); clear H; intros ->.
+ rewrite to_of.
+ rewrite <-DecimalZ.to_of, DecimalZ.of_to; simpl.
+ do 4 apply f_equal.
+ apply Pos2Nat.inj.
+ rewrite <-!SuccNat2Pos.inj_succ.
+ rewrite SuccNat2Pos.id_succ.
+ case e; [now simpl|intro e'']; simpl.
+ unfold Pos.to_nat; simpl.
+ now rewrite Pmult_nat_mult, SuccNat2Pos.id_succ, Nat.mul_comm.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_hexadecimal_inj q q' :
+ to_hexadecimal q <> None -> to_hexadecimal q = to_hexadecimal q' -> q = q'.
+Proof.
+ intros Hnone EQ.
+ generalize (of_to q) (of_to q').
+ rewrite <-EQ.
+ revert Hnone; case to_hexadecimal; [|now simpl].
+ now intros d _ H1 H2; rewrite <-(H1 d eq_refl), <-(H2 d eq_refl).
+Qed.
+
+Lemma to_hexadecimal_surj d : exists q, to_hexadecimal q = Some (hnorme d).
+Proof.
+ exists (of_hexadecimal d). apply to_of.
+Qed.
+
+Lemma of_hexadecimal_hnorme d : of_hexadecimal (hnorme d) = of_hexadecimal d.
+Proof.
+ unfold of_hexadecimal, hnorme.
+ destruct d.
+ - simpl Z.of_int; unfold Z.of_uint, Z.of_N, Pos.of_uint.
+ rewrite Z.sub_0_l.
+ set (n4f := (- _)%Z).
+ case_eq n4f; [|intro pn4f..]; intro Hn4f.
+ + apply f_equal2; [|reflexivity].
+ rewrite app_int_nil_r.
+ now rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to.
+ + apply f_equal2; [|reflexivity].
+ rewrite app_int_nil_r.
+ generalize (app_int i f); intro i'.
+ rewrite !Pos2Nat.inj_iter.
+ generalize (Pos.to_nat pn4f); intro n.
+ fold (Nat.iter n double (norm i')).
+ fold (Nat.iter n (Z.mul 2) (Z.of_hex_int i')).
+ induction n; [now simpl; rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to|].
+ now rewrite !Unsigned.nat_iter_S, <-IHn, of_hex_int_double.
+ + unfold nb_digits, Z.of_nat.
+ rewrite Z.mul_0_r, Z.sub_0_r.
+ rewrite <-DecimalZ.to_of, !DecimalZ.of_to.
+ rewrite app_int_nil_r.
+ now rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to.
+ - set (nem4f := (_ - _)%Z).
+ case_eq nem4f; [|intro pnem4f..]; intro Hnem4f.
+ + apply f_equal2; [|reflexivity].
+ rewrite app_int_nil_r.
+ now rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to.
+ + apply f_equal2; [|reflexivity].
+ rewrite app_int_nil_r.
+ generalize (app_int i f); intro i'.
+ rewrite !Pos2Nat.inj_iter.
+ generalize (Pos.to_nat pnem4f); intro n.
+ fold (Nat.iter n double (norm i')).
+ fold (Nat.iter n (Z.mul 2) (Z.of_hex_int i')).
+ induction n; [now simpl; rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to|].
+ now rewrite !Unsigned.nat_iter_S, <-IHn, of_hex_int_double.
+ + unfold nb_digits, Z.of_nat.
+ rewrite Z.mul_0_r, Z.sub_0_r.
+ rewrite <-DecimalZ.to_of, !DecimalZ.of_to.
+ rewrite app_int_nil_r.
+ now rewrite <-HexadecimalZ.to_of, HexadecimalZ.of_to.
+Qed.
+
+Lemma of_inj d d' :
+ of_hexadecimal d = of_hexadecimal d' -> hnorme d = hnorme d'.
+Proof.
+ intros.
+ cut (Some (hnorme d) = Some (hnorme d')); [now intro H'; injection H'|].
+ rewrite <- !to_of. now f_equal.
+Qed.
+
+Lemma of_iff d d' :
+ of_hexadecimal d = of_hexadecimal d' <-> hnorme d = hnorme d'.
+Proof.
+ split. apply of_inj. intros E. rewrite <- of_hexadecimal_hnorme, E.
+ apply of_hexadecimal_hnorme.
+Qed.
diff --git a/theories/Numbers/HexadecimalString.v b/theories/Numbers/HexadecimalString.v
new file mode 100644
index 0000000000..7017bdf60f
--- /dev/null
+++ b/theories/Numbers/HexadecimalString.v
@@ -0,0 +1,286 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+Require Import Hexadecimal Ascii String.
+
+(** * Conversion between hexadecimal numbers and Coq strings *)
+
+(** Pretty straightforward, which is precisely the point of the
+ [Hexadecimal.int] datatype. The only catch is [Hexadecimal.Nil] : we could
+ choose to convert it as [""] or as ["0"]. In the first case, it is
+ awkward to consider "" (or "-") as a number, while in the second case
+ we don't have a perfect bijection. Since the second variant is implemented
+ thanks to the first one, we provide both.
+
+ Hexadecimal digits are lower case ('a'..'f'). We ignore upper case
+ digits ('A'..'F') for the sake of simplicity. *)
+
+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)
+ | "a" => Some (Da d)
+ | "b" => Some (Db d)
+ | "c" => Some (Dc d)
+ | "d" => Some (Dd d)
+ | "e" => Some (De d)
+ | "f" => Some (Df 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 \/
+ c = "a" /\ d' = Da d \/
+ c = "b" /\ d' = Db d \/
+ c = "c" /\ d' = Dc d \/
+ c = "d" /\ d' = Dd d \/
+ c = "e" /\ d' = De d \/
+ c = "f" /\ d' = Df d)%char.
+Proof.
+ destruct c as [[|] [|] [|] [|] [|] [|] [|] [|]];
+ intros [= <-]; intuition.
+Qed.
+
+(** Hexadecimal/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)
+ | Da d => String "a" (string_of_uint d)
+ | Db d => String "b" (string_of_uint d)
+ | Dc d => String "c" (string_of_uint d)
+ | Dd d => String "d" (string_of_uint d)
+ | De d => String "e" (string_of_uint d)
+ | Df d => String "f" (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.
+
+(* NB: For the moment whitespace between - and digits are not accepted.
+ And in this variant [int_of_string "-" = Some (Neg Nil)].
+
+Compute int_of_string "-123456890123456890123456890123456890".
+Compute string_of_int (-123456890123456890123456890123456890).
+*)
+
+(** Corresponding proofs *)
+
+Lemma usu d :
+ uint_of_string (string_of_uint d) = Some d.
+Proof.
+ induction d; simpl; rewrite ?IHd; simpl; auto.
+Qed.
+
+Lemma sus s d :
+ uint_of_string s = Some d -> string_of_uint d = s.
+Proof.
+ revert d.
+ induction s; simpl.
+ - now intros d [= <-].
+ - intros d.
+ destruct (uint_of_string s); [intros H | intros [=]].
+ apply uint_of_char_spec in H.
+ intuition subst; simpl; f_equal; auto.
+Qed.
+
+Lemma isi d : int_of_string (string_of_int d) = Some d.
+Proof.
+ destruct d; simpl.
+ - unfold int_of_string.
+ destruct (string_of_uint d) eqn:Hd.
+ + now destruct d.
+ + case Ascii.eqb_spec.
+ * intros ->. now destruct d.
+ * rewrite <- Hd, usu; auto.
+ - rewrite usu; auto.
+Qed.
+
+Lemma sis s d :
+ int_of_string s = Some d -> string_of_int d = s.
+Proof.
+ destruct s; [intros [= <-]| ]; simpl; trivial.
+ case Ascii.eqb_spec.
+ - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
+ simpl; f_equal. now apply sus.
+ - destruct d; [ | now destruct uint_of_char].
+ simpl string_of_int.
+ intros. apply sus; simpl.
+ destruct uint_of_char; simpl in *; congruence.
+Qed.
+
+End NilEmpty.
+
+(** Hexadecimal/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.
+Proof.
+ destruct s; simpl.
+ - easy.
+ - destruct (NilEmpty.uint_of_string s); [intros H | intros [=]].
+ apply uint_of_char_spec in H.
+ now intuition subst.
+Qed.
+
+Lemma sus s d :
+ uint_of_string s = Some d -> string_of_uint d = s.
+Proof.
+ destruct s; [intros [=] | intros H].
+ apply NilEmpty.sus in H. now destruct d.
+Qed.
+
+Lemma usu d :
+ d<>Nil -> uint_of_string (string_of_uint d) = Some d.
+Proof.
+ destruct d; (now destruct 1) || (intros _; apply NilEmpty.usu).
+Qed.
+
+Lemma usu_nil :
+ uint_of_string (string_of_uint Nil) = Some Hexadecimal.zero.
+Proof.
+ reflexivity.
+Qed.
+
+Lemma usu_gen d :
+ uint_of_string (string_of_uint d) = Some d \/
+ uint_of_string (string_of_uint d) = Some Hexadecimal.zero.
+Proof.
+ destruct d; (now right) || (left; now apply usu).
+Qed.
+
+Lemma isi d :
+ d<>Pos Nil -> d<>Neg Nil ->
+ int_of_string (string_of_int d) = Some d.
+Proof.
+ destruct d; simpl.
+ - intros H _.
+ unfold int_of_string.
+ destruct (string_of_uint d) eqn:Hd.
+ + now destruct d.
+ + case Ascii.eqb_spec.
+ * intros ->. now destruct d.
+ * rewrite <- Hd, usu; auto. now intros ->.
+ - intros _ H.
+ rewrite usu; auto. now intros ->.
+Qed.
+
+Lemma isi_posnil :
+ int_of_string (string_of_int (Pos Nil)) = Some (Pos Hexadecimal.zero).
+Proof.
+ reflexivity.
+Qed.
+
+(** 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)).
+Proof.
+ reflexivity.
+Qed.
+
+Lemma sis s d :
+ int_of_string s = Some d -> string_of_int d = s.
+Proof.
+ destruct s; [intros [=]| ]; simpl.
+ case Ascii.eqb_spec.
+ - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-].
+ simpl; f_equal. now apply sus.
+ - destruct d; [ | now destruct uint_of_char].
+ simpl string_of_int.
+ intros. apply sus; simpl.
+ destruct uint_of_char; simpl in *; congruence.
+Qed.
+
+End NilZero.
diff --git a/theories/Numbers/HexadecimalZ.v b/theories/Numbers/HexadecimalZ.v
new file mode 100644
index 0000000000..c5ed0b5b28
--- /dev/null
+++ b/theories/Numbers/HexadecimalZ.v
@@ -0,0 +1,142 @@
+(************************************************************************)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * Copyright INRIA, CNRS and contributors *)
+(* <O___,, * (see version control and CREDITS file for authors & dates) *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
+(************************************************************************)
+
+(** * HexadecimalZ
+
+ Proofs that conversions between hexadecimal numbers and [Z]
+ are bijections. *)
+
+Require Import Hexadecimal HexadecimalFacts HexadecimalPos HexadecimalN ZArith.
+
+Lemma of_to (z:Z) : Z.of_hex_int (Z.to_hex_int z) = z.
+Proof.
+ destruct z; simpl.
+ - trivial.
+ - unfold Z.of_hex_uint. rewrite HexadecimalPos.Unsigned.of_to. now destruct p.
+ - unfold Z.of_hex_uint. rewrite HexadecimalPos.Unsigned.of_to. destruct p; auto.
+Qed.
+
+Lemma to_of (d:int) : Z.to_hex_int (Z.of_hex_int d) = norm d.
+Proof.
+ destruct d; simpl; unfold Z.to_hex_int, Z.of_hex_uint.
+ - rewrite <- (HexadecimalN.Unsigned.to_of d). unfold N.of_hex_uint.
+ now destruct (Pos.of_hex_uint d).
+ - destruct (Pos.of_hex_uint d) eqn:Hd; simpl; f_equal.
+ + generalize (HexadecimalPos.Unsigned.to_of d). rewrite Hd. simpl.
+ intros H. symmetry in H. apply unorm_0 in H. now rewrite H.
+ + assert (Hp := HexadecimalPos.Unsigned.to_of d). rewrite Hd in Hp. simpl in *.
+ rewrite Hp. unfold unorm in *.
+ destruct (nzhead d); trivial.
+ generalize (HexadecimalPos.Unsigned.of_to p). now rewrite Hp.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_int_inj n n' : Z.to_hex_int n = Z.to_hex_int n' -> n = n'.
+Proof.
+ intro EQ.
+ now rewrite <- (of_to n), <- (of_to n'), EQ.
+Qed.
+
+Lemma to_int_surj d : exists n, Z.to_hex_int n = norm d.
+Proof.
+ exists (Z.of_hex_int d). apply to_of.
+Qed.
+
+Lemma of_int_norm d : Z.of_hex_int (norm d) = Z.of_hex_int d.
+Proof.
+ unfold Z.of_hex_int, Z.of_hex_uint.
+ destruct d.
+ - simpl. now rewrite HexadecimalPos.Unsigned.of_uint_norm.
+ - simpl. destruct (nzhead d) eqn:H;
+ [ induction d; simpl; auto; discriminate |
+ destruct (nzhead_nonzero _ _ H) | .. ];
+ f_equal; f_equal; apply HexadecimalPos.Unsigned.of_iff;
+ unfold unorm; now rewrite H.
+Qed.
+
+Lemma of_inj d d' :
+ Z.of_hex_int d = Z.of_hex_int d' -> norm d = norm d'.
+Proof.
+ intros. rewrite <- !to_of. now f_equal.
+Qed.
+
+Lemma of_iff d d' : Z.of_hex_int d = Z.of_hex_int d' <-> norm d = norm d'.
+Proof.
+ split. apply of_inj. intros E. rewrite <- of_int_norm, E.
+ apply of_int_norm.
+Qed.
+
+(** Various lemmas *)
+
+Lemma of_hex_uint_iter_D0 d n :
+ Z.of_hex_uint (app d (Nat.iter n D0 Nil))
+ = Nat.iter n (Z.mul 0x10) (Z.of_hex_uint d).
+Proof.
+ unfold Z.of_hex_uint.
+ unfold app; rewrite <-rev_revapp.
+ rewrite Unsigned.of_lu_rev, Unsigned.of_lu_revapp.
+ rewrite <-!Unsigned.of_lu_rev, !rev_rev.
+ assert (H' : Pos.of_hex_uint (Nat.iter n D0 Nil) = 0%N).
+ { now induction n; [|rewrite Unsigned.nat_iter_S]. }
+ rewrite H', N.add_0_l; clear H'.
+ induction n; [now simpl; rewrite N.mul_1_r|].
+ rewrite !Unsigned.nat_iter_S, <-IHn.
+ simpl Unsigned.usize; rewrite N.pow_succ_r'.
+ rewrite !N2Z.inj_mul; simpl Z.of_N; ring.
+Qed.
+
+Lemma of_hex_int_iter_D0 d n :
+ Z.of_hex_int (app_int d (Nat.iter n D0 Nil))
+ = Nat.iter n (Z.mul 0x10) (Z.of_hex_int d).
+Proof.
+ case d; clear d; intro d; simpl.
+ - now rewrite of_hex_uint_iter_D0.
+ - rewrite of_hex_uint_iter_D0; induction n; [now simpl|].
+ rewrite !Unsigned.nat_iter_S, <-IHn; ring.
+Qed.
+
+Definition double d :=
+ match d with
+ | Pos u => Pos (Unsigned.double u)
+ | Neg u => Neg (Unsigned.double u)
+ end.
+
+Lemma double_norm d : double (norm d) = norm (double d).
+Proof.
+ destruct d.
+ - now simpl; rewrite Unsigned.double_unorm.
+ - simpl; rewrite <-Unsigned.double_nzhead.
+ case (uint_eq_dec (nzhead d) Nil); intro Hnzd.
+ + now rewrite Hnzd.
+ + assert (H : Unsigned.double (nzhead d) <> Nil).
+ { unfold Unsigned.double.
+ intro H; apply Hnzd, rev_nil_inv.
+ now generalize (rev_nil_inv _ H); case rev. }
+ revert H.
+ set (r := Unsigned.double _).
+ set (m := match r with Nil => Pos zero | _ => _ end).
+ intro H.
+ assert (H' : m = Neg r).
+ { now unfold m; clear m; revert H; case r. }
+ rewrite H'; unfold r; clear m r H H'.
+ now revert Hnzd; case nzhead.
+Qed.
+
+Lemma of_hex_int_double d :
+ Z.of_hex_int (double d) = Z.double (Z.of_hex_int d).
+Proof.
+ now destruct d; simpl; unfold Z.of_hex_uint;
+ rewrite Unsigned.of_hex_uint_double; case Pos.of_hex_uint.
+Qed.
+
+Lemma double_to_hex_int n :
+ double (Z.to_hex_int n) = Z.to_hex_int (Z.double n).
+Proof. now rewrite <-(of_to n), <-of_hex_int_double, !to_of, double_norm. Qed.
diff --git a/theories/Numbers/NaryFunctions.v b/theories/Numbers/NaryFunctions.v
index 80ad41f7c0..270cd4d111 100644
--- a/theories/Numbers/NaryFunctions.v
+++ b/theories/Numbers/NaryFunctions.v
@@ -20,70 +20,70 @@ Require Import List.
[A -> ... -> A -> B] with [n] occurrences of [A] in this type. *)
Fixpoint nfun A n B :=
- match n with
+ match n with
| O => B
| S n => A -> (nfun A n B)
- end.
+ end.
Notation " A ^^ n --> B " := (nfun A n B)
- (at level 50, n at next level) : type_scope.
+ (at level 50, n at next level) : type_scope.
(** [napply_cst _ _ a n f] iterates [n] times the application of a
particular constant [a] to the [n]-ary function [f]. *)
Fixpoint napply_cst (A B:Type)(a:A) n : (A^^n-->B) -> B :=
- match n return (A^^n-->B) -> B with
+ match n return (A^^n-->B) -> B with
| O => fun x => x
| S n => fun x => napply_cst _ _ a n (x a)
- end.
+ end.
(** A generic transformation from an n-ary function to another one.*)
Fixpoint nfun_to_nfun (A B C:Type)(f:B -> C) n :
- (A^^n-->B) -> (A^^n-->C) :=
- match n return (A^^n-->B) -> (A^^n-->C) with
+ (A^^n-->B) -> (A^^n-->C) :=
+ match n return (A^^n-->B) -> (A^^n-->C) with
| O => f
| S n => fun g a => nfun_to_nfun _ _ _ f n (g a)
- end.
+ end.
(** [napply_except_last _ _ n f] expects [n] arguments of type [A],
applies [n-1] of them to [f] and discard the last one. *)
Definition napply_except_last (A B:Type) :=
- nfun_to_nfun A B (A->B) (fun b a => b).
+ nfun_to_nfun A B (A->B) (fun b a => b).
(** [napply_then_last _ _ a n f] expects [n] arguments of type [A],
applies them to [f] and then apply [a] to the result. *)
Definition napply_then_last (A B:Type)(a:A) :=
- nfun_to_nfun A (A->B) B (fun fab => fab a).
+ nfun_to_nfun A (A->B) B (fun fab => fab a).
(** [napply_discard _ b n] expects [n] arguments, discards then,
and returns [b]. *)
Fixpoint napply_discard (A B:Type)(b:B) n : A^^n-->B :=
- match n return A^^n-->B with
+ match n return A^^n-->B with
| O => b
| S n => fun _ => napply_discard _ _ b n
- end.
+ end.
(** A fold function *)
Fixpoint nfold A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
- match n return (A^^n-->B) with
+ match n return (A^^n-->B) with
| O => b
| S n => fun a => (nfold _ _ f (f a b) n)
- end.
+ end.
(** [n]-ary products : [nprod A n] is [A*...*A*unit],
with [n] occurrences of [A] in this type. *)
Fixpoint nprod A n : Type := match n with
- | O => unit
- | S n => (A * nprod A n)%type
-end.
+ | O => unit
+ | S n => (A * nprod A n)%type
+ end.
Notation "A ^ n" := (nprod A n) : type_scope.
@@ -96,44 +96,44 @@ Fixpoint ncurry (A B:Type) n : (A^n -> B) -> (A^^n-->B) :=
end.
Fixpoint nuncurry (A B:Type) n : (A^^n-->B) -> (A^n -> B) :=
- match n return (A^^n-->B) -> (A^n -> B) with
+ match n return (A^^n-->B) -> (A^n -> B) with
| O => fun x _ => x
| S n => fun f p => let (x,p) := p in nuncurry _ _ n (f x) p
- end.
+ end.
(** Earlier functions can also be defined via [ncurry/nuncurry].
For instance : *)
Definition nfun_to_nfun_bis A B C (f:B->C) n :
- (A^^n-->B) -> (A^^n-->C) :=
- fun anb => ncurry _ _ n (fun an => f ((nuncurry _ _ n anb) an)).
+ (A^^n-->B) -> (A^^n-->C) :=
+ fun anb => ncurry _ _ n (fun an => f ((nuncurry _ _ n anb) an)).
(** We can also us it to obtain another [fold] function,
equivalent to the previous one, but with a nicer expansion
(see for instance Int31.iszero). *)
Fixpoint nfold_bis A B (f:A->B->B)(b:B) n : (A^^n-->B) :=
- match n return (A^^n-->B) with
+ match n return (A^^n-->B) with
| O => b
| S n => fun a =>
nfun_to_nfun_bis _ _ _ (f a) n (nfold_bis _ _ f b n)
- end.
+ end.
(** From [nprod] to [list] *)
Fixpoint nprod_to_list (A:Type) n : A^n -> list A :=
- match n with
+ match n with
| O => fun _ => nil
| S n => fun p => let (x,p) := p in x::(nprod_to_list _ n p)
- end.
+ end.
(** From [list] to [nprod] *)
Fixpoint nprod_of_list (A:Type)(l:list A) : A^(length l) :=
- match l return A^(length l) with
+ match l return A^(length l) with
| nil => tt
| x::l => (x, nprod_of_list _ l)
- end.
+ end.
(** This gives an additional way to write the fold *)
diff --git a/theories/PArith/BinPos.v b/theories/PArith/BinPos.v
index 4179765dca..c8414c241d 100644
--- a/theories/PArith/BinPos.v
+++ b/theories/PArith/BinPos.v
@@ -1886,7 +1886,7 @@ Bind Scope positive_scope with Pos.t positive.
(** Exportation of notations *)
-Numeral Notation positive Pos.of_int Pos.to_uint : positive_scope.
+Numeral Notation positive Pos.of_num_int Pos.to_num_uint : positive_scope.
Infix "+" := Pos.add : positive_scope.
Infix "-" := Pos.sub : positive_scope.
diff --git a/theories/PArith/BinPosDef.v b/theories/PArith/BinPosDef.v
index 953068b010..cdb9af542c 100644
--- a/theories/PArith/BinPosDef.v
+++ b/theories/PArith/BinPosDef.v
@@ -595,6 +595,56 @@ Fixpoint of_uint (d:Decimal.uint) : N :=
| Decimal.D9 l => Npos (of_uint_acc l 1~0~0~1)
end.
+Local Notation sixteen := 1~0~0~0~0.
+
+Fixpoint of_hex_uint_acc (d:Hexadecimal.uint)(acc:positive) :=
+ match d with
+ | Hexadecimal.Nil => acc
+ | Hexadecimal.D0 l => of_hex_uint_acc l (mul sixteen acc)
+ | Hexadecimal.D1 l => of_hex_uint_acc l (add 1 (mul sixteen acc))
+ | Hexadecimal.D2 l => of_hex_uint_acc l (add 1~0 (mul sixteen acc))
+ | Hexadecimal.D3 l => of_hex_uint_acc l (add 1~1 (mul sixteen acc))
+ | Hexadecimal.D4 l => of_hex_uint_acc l (add 1~0~0 (mul sixteen acc))
+ | Hexadecimal.D5 l => of_hex_uint_acc l (add 1~0~1 (mul sixteen acc))
+ | Hexadecimal.D6 l => of_hex_uint_acc l (add 1~1~0 (mul sixteen acc))
+ | Hexadecimal.D7 l => of_hex_uint_acc l (add 1~1~1 (mul sixteen acc))
+ | Hexadecimal.D8 l => of_hex_uint_acc l (add 1~0~0~0 (mul sixteen acc))
+ | Hexadecimal.D9 l => of_hex_uint_acc l (add 1~0~0~1 (mul sixteen acc))
+ | Hexadecimal.Da l => of_hex_uint_acc l (add 1~0~1~0 (mul sixteen acc))
+ | Hexadecimal.Db l => of_hex_uint_acc l (add 1~0~1~1 (mul sixteen acc))
+ | Hexadecimal.Dc l => of_hex_uint_acc l (add 1~1~0~0 (mul sixteen acc))
+ | Hexadecimal.Dd l => of_hex_uint_acc l (add 1~1~0~1 (mul sixteen acc))
+ | Hexadecimal.De l => of_hex_uint_acc l (add 1~1~1~0 (mul sixteen acc))
+ | Hexadecimal.Df l => of_hex_uint_acc l (add 1~1~1~1 (mul sixteen acc))
+ end.
+
+Fixpoint of_hex_uint (d:Hexadecimal.uint) : N :=
+ match d with
+ | Hexadecimal.Nil => N0
+ | Hexadecimal.D0 l => of_hex_uint l
+ | Hexadecimal.D1 l => Npos (of_hex_uint_acc l 1)
+ | Hexadecimal.D2 l => Npos (of_hex_uint_acc l 1~0)
+ | Hexadecimal.D3 l => Npos (of_hex_uint_acc l 1~1)
+ | Hexadecimal.D4 l => Npos (of_hex_uint_acc l 1~0~0)
+ | Hexadecimal.D5 l => Npos (of_hex_uint_acc l 1~0~1)
+ | Hexadecimal.D6 l => Npos (of_hex_uint_acc l 1~1~0)
+ | Hexadecimal.D7 l => Npos (of_hex_uint_acc l 1~1~1)
+ | Hexadecimal.D8 l => Npos (of_hex_uint_acc l 1~0~0~0)
+ | Hexadecimal.D9 l => Npos (of_hex_uint_acc l 1~0~0~1)
+ | Hexadecimal.Da l => Npos (of_hex_uint_acc l 1~0~1~0)
+ | Hexadecimal.Db l => Npos (of_hex_uint_acc l 1~0~1~1)
+ | Hexadecimal.Dc l => Npos (of_hex_uint_acc l 1~1~0~0)
+ | Hexadecimal.Dd l => Npos (of_hex_uint_acc l 1~1~0~1)
+ | Hexadecimal.De l => Npos (of_hex_uint_acc l 1~1~1~0)
+ | Hexadecimal.Df l => Npos (of_hex_uint_acc l 1~1~1~1)
+ end.
+
+Definition of_num_uint (d:Numeral.uint) : N :=
+ match d with
+ | Numeral.UIntDec d => of_uint d
+ | Numeral.UIntHex d => of_hex_uint d
+ end.
+
Definition of_int (d:Decimal.int) : option positive :=
match d with
| Decimal.Pos d =>
@@ -605,6 +655,22 @@ Definition of_int (d:Decimal.int) : option positive :=
| Decimal.Neg _ => None
end.
+Definition of_hex_int (d:Hexadecimal.int) : option positive :=
+ match d with
+ | Hexadecimal.Pos d =>
+ match of_hex_uint d with
+ | N0 => None
+ | Npos p => Some p
+ end
+ | Hexadecimal.Neg _ => None
+ end.
+
+Definition of_num_int (d:Numeral.int) : option positive :=
+ match d with
+ | Numeral.IntDec d => of_int d
+ | Numeral.IntHex d => of_hex_int d
+ end.
+
Fixpoint to_little_uint p :=
match p with
| 1 => Decimal.D1 Decimal.Nil
@@ -614,11 +680,26 @@ Fixpoint to_little_uint p :=
Definition to_uint p := Decimal.rev (to_little_uint p).
+Fixpoint to_little_hex_uint p :=
+ match p with
+ | 1 => Hexadecimal.D1 Hexadecimal.Nil
+ | p~1 => Hexadecimal.Little.succ_double (to_little_hex_uint p)
+ | p~0 => Hexadecimal.Little.double (to_little_hex_uint p)
+ end.
+
+Definition to_hex_uint p := Hexadecimal.rev (to_little_hex_uint p).
+
+Definition to_num_uint p := Numeral.UIntDec (to_uint p).
+
Definition to_int n := Decimal.Pos (to_uint n).
-Numeral Notation positive of_int to_uint : positive_scope.
+Definition to_hex_int p := Hexadecimal.Pos (to_hex_uint p).
+
+Definition to_num_int n := Numeral.IntDec (to_int n).
+
+Numeral Notation positive of_num_int to_num_uint : positive_scope.
End Pos.
(** Re-export the notation for those who just [Import BinPosDef] *)
-Numeral Notation positive Pos.of_int Pos.to_uint : positive_scope.
+Numeral Notation positive Pos.of_num_int Pos.to_num_uint : positive_scope.
diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v
index 74cdd1797c..84d70e56de 100644
--- a/theories/QArith/QArith_base.v
+++ b/theories/QArith/QArith_base.v
@@ -85,7 +85,51 @@ Definition to_decimal (q:Q) : option Decimal.decimal :=
| _ => None
end.
-Numeral Notation Q of_decimal to_decimal : Q_scope.
+Definition of_hexadecimal (d:Hexadecimal.hexadecimal) : Q :=
+ let '(i, f, e) :=
+ match d with
+ | Hexadecimal.Hexadecimal i f => (i, f, Decimal.Pos Decimal.Nil)
+ | Hexadecimal.HexadecimalExp i f e => (i, f, e)
+ end in
+ let num := Z.of_hex_int (Hexadecimal.app_int i f) in
+ let e := Z.sub (Z.of_int e) (Z.mul 4 (Z.of_nat (Hexadecimal.nb_digits f))) in
+ match e with
+ | Z0 => Qmake num 1
+ | Zpos e => Qmake (Pos.iter (Z.mul 2) num e) 1
+ | Zneg e => Qmake num (Pos.iter (Pos.mul 2) 1%positive e)
+ end.
+
+Definition to_hexadecimal (q:Q) : option Hexadecimal.hexadecimal :=
+ let mk_exp i e :=
+ Hexadecimal.HexadecimalExp i Hexadecimal.Nil (Z.to_int (Z.opp e)) in
+ let num := Z.to_hex_int (Qnum q) in
+ let (den, e_den) := Hexadecimal.nztail (Pos.to_hex_uint (Qden q)) in
+ let e := Z.of_nat e_den in
+ match den with
+ | Hexadecimal.D1 Hexadecimal.Nil =>
+ match e_den with
+ | O => Some (Hexadecimal.Hexadecimal num Hexadecimal.Nil)
+ | _ => Some (mk_exp num (4 * e)%Z)
+ end
+ | Hexadecimal.D2 Hexadecimal.Nil => Some (mk_exp num (1 + 4 * e)%Z)
+ | Hexadecimal.D4 Hexadecimal.Nil => Some (mk_exp num (2 + 4 * e)%Z)
+ | Hexadecimal.D8 Hexadecimal.Nil => Some (mk_exp num (3 + 4 * e)%Z)
+ | _ => None
+ end.
+
+Definition of_numeral (d:Numeral.numeral) : option Q :=
+ match d with
+ | Numeral.Dec d => Some (of_decimal d)
+ | Numeral.Hex d => Some (of_hexadecimal d)
+ end.
+
+Definition to_numeral (q:Q) : option Numeral.numeral :=
+ match to_decimal q with
+ | None => None
+ | Some q => Some (Numeral.Dec q)
+ end.
+
+Numeral Notation Q of_numeral to_numeral : Q_scope.
Definition inject_Z (x : Z) := Qmake x 1.
Arguments inject_Z x%Z.
diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v
index 78b26c83ea..1729b9f85e 100644
--- a/theories/ZArith/BinInt.v
+++ b/theories/ZArith/BinInt.v
@@ -1290,7 +1290,7 @@ Bind Scope Z_scope with Z.t Z.
(** Re-export Notations *)
-Numeral Notation Z Z.of_int Z.to_int : Z_scope.
+Numeral Notation Z Z.of_num_int Z.to_num_int : Z_scope.
Infix "+" := Z.add : Z_scope.
Notation "- x" := (Z.opp x) : Z_scope.
diff --git a/theories/ZArith/BinIntDef.v b/theories/ZArith/BinIntDef.v
index c05ed9ebf4..8464ad1012 100644
--- a/theories/ZArith/BinIntDef.v
+++ b/theories/ZArith/BinIntDef.v
@@ -309,12 +309,32 @@ Definition to_pos (z:Z) : positive :=
Definition of_uint (d:Decimal.uint) := of_N (Pos.of_uint d).
+Definition of_hex_uint (d:Hexadecimal.uint) := of_N (Pos.of_hex_uint d).
+
+Definition of_num_uint (d:Numeral.uint) :=
+ match d with
+ | Numeral.UIntDec d => of_uint d
+ | Numeral.UIntHex d => of_hex_uint d
+ end.
+
Definition of_int (d:Decimal.int) :=
match d with
| Decimal.Pos d => of_uint d
| Decimal.Neg d => opp (of_uint d)
end.
+Definition of_hex_int (d:Hexadecimal.int) :=
+ match d with
+ | Hexadecimal.Pos d => of_hex_uint d
+ | Hexadecimal.Neg d => opp (of_hex_uint d)
+ end.
+
+Definition of_num_int (d:Numeral.int) :=
+ match d with
+ | Numeral.IntDec d => of_int d
+ | Numeral.IntHex d => of_hex_int d
+ end.
+
Definition to_int n :=
match n with
| 0 => Decimal.Pos Decimal.zero
@@ -322,6 +342,15 @@ Definition to_int n :=
| neg p => Decimal.Neg (Pos.to_uint p)
end.
+Definition to_hex_int n :=
+ match n with
+ | 0 => Hexadecimal.Pos Hexadecimal.zero
+ | pos p => Hexadecimal.Pos (Pos.to_hex_uint p)
+ | neg p => Hexadecimal.Neg (Pos.to_hex_uint p)
+ end.
+
+Definition to_num_int n := Numeral.IntDec (to_int n).
+
(** ** Iteration of a function
By convention, iterating a negative number of times is identity.
@@ -639,9 +668,9 @@ Definition lxor a b :=
| neg a, neg b => of_N (N.lxor (Pos.pred_N a) (Pos.pred_N b))
end.
-Numeral Notation Z of_int to_int : Z_scope.
+Numeral Notation Z of_num_int to_num_int : Z_scope.
End Z.
(** Re-export the notation for those who just [Import BinIntDef] *)
-Numeral Notation Z Z.of_int Z.to_int : Z_scope.
+Numeral Notation Z Z.of_num_int Z.to_num_int : Z_scope.