aboutsummaryrefslogtreecommitdiffhomepage
path: root/theories
diff options
context:
space:
mode:
Diffstat (limited to 'theories')
-rw-r--r--theories/.dir-locals.el4
-rw-r--r--theories/Arith/Between.v4
-rw-r--r--theories/Arith/Lt.v5
-rw-r--r--theories/Arith/PeanoNat.v20
-rw-r--r--theories/Arith/Peano_dec.v2
-rw-r--r--theories/Compat/Coq87.v11
-rw-r--r--theories/FSets/FSetCompat.v44
-rw-r--r--theories/FSets/FSetProperties.v2
-rw-r--r--theories/FSets/FSets.v2
-rw-r--r--theories/Init/Decimal.v161
-rw-r--r--theories/Init/Logic.v7
-rw-r--r--theories/Init/Nat.v58
-rw-r--r--theories/Init/Notations.v27
-rw-r--r--theories/Init/Prelude.v1
-rw-r--r--theories/Init/Specif.v9
-rw-r--r--theories/Init/Tactics.v7
-rw-r--r--theories/Init/Tauto.v2
-rw-r--r--theories/Lists/List.v27
-rw-r--r--theories/Logic/Classical_Prop.v6
-rw-r--r--theories/Logic/FunctionalExtensionality.v3
-rw-r--r--theories/MSets/MSetGenTree.v2
-rw-r--r--theories/MSets/MSets.v2
-rw-r--r--theories/NArith/BinNatDef.v20
-rw-r--r--theories/Numbers/DecimalFacts.v141
-rw-r--r--theories/Numbers/DecimalN.v105
-rw-r--r--theories/Numbers/DecimalNat.v300
-rw-r--r--theories/Numbers/DecimalPos.v381
-rw-r--r--theories/Numbers/DecimalString.v263
-rw-r--r--theories/Numbers/DecimalZ.v73
-rw-r--r--theories/Numbers/NatInt/NZParity.v2
-rw-r--r--theories/PArith/BinPosDef.v55
-rw-r--r--theories/Program/Combinators.v12
-rw-r--r--theories/Program/Tactics.v2
-rw-r--r--theories/QArith/QArith_base.v20
-rw-r--r--theories/QArith/Qabs.v7
-rw-r--r--theories/QArith/Qcabs.v2
-rw-r--r--theories/QArith/Qreduction.v8
-rw-r--r--theories/Reals/Ranalysis.v2
-rw-r--r--theories/Sets/Powerset_facts.v91
-rw-r--r--theories/Unicode/Utf8_core.v15
-rw-r--r--theories/Vectors/Vector.v2
-rw-r--r--theories/ZArith/BinIntDef.v19
-rw-r--r--theories/ZArith/Zsqrt_compat.v2
43 files changed, 1854 insertions, 74 deletions
diff --git a/theories/.dir-locals.el b/theories/.dir-locals.el
deleted file mode 100644
index 4e8830f6c..000000000
--- a/theories/.dir-locals.el
+++ /dev/null
@@ -1,4 +0,0 @@
-((coq-mode . ((eval . (let ((default-directory (locate-dominating-file
- buffer-file-name ".dir-locals.el")))
- (setq-local coq-prog-args `("-coqlib" ,(expand-file-name "..") "-R" ,(expand-file-name ".") "Coq"))
- (setq-local coq-prog-name (expand-file-name "../bin/coqtop")))))))
diff --git a/theories/Arith/Between.v b/theories/Arith/Between.v
index 9b4071085..ead08b3eb 100644
--- a/theories/Arith/Between.v
+++ b/theories/Arith/Between.v
@@ -16,6 +16,8 @@ Implicit Types k l p q r : nat.
Section Between.
Variables P Q : nat -> Prop.
+ (** The [between] type expresses the concept
+ [forall i: nat, k <= i < l -> P i.]. *)
Inductive between k : nat -> Prop :=
| bet_emp : between k k
| bet_S : forall l, between k l -> P l -> between k (S l).
@@ -47,6 +49,8 @@ Section Between.
induction 1; auto with arith.
Qed.
+ (** The [exists_between] type expresses the concept
+ [exists i: nat, k <= i < l /\ Q i]. *)
Inductive exists_between k : nat -> Prop :=
| exists_S : forall l, exists_between k l -> exists_between k (S l)
| exists_le : forall l, k <= l -> Q l -> exists_between k (S l).
diff --git a/theories/Arith/Lt.v b/theories/Arith/Lt.v
index 035c4e466..2c2bea4a6 100644
--- a/theories/Arith/Lt.v
+++ b/theories/Arith/Lt.v
@@ -107,6 +107,11 @@ Proof.
intros. symmetry. now apply Nat.lt_succ_pred with m.
Qed.
+Lemma S_pred_pos n: O < n -> n = S (pred n).
+Proof.
+ apply S_pred.
+Qed.
+
Lemma lt_pred n m : S n < m -> n < pred m.
Proof.
apply Nat.lt_succ_lt_pred.
diff --git a/theories/Arith/PeanoNat.v b/theories/Arith/PeanoNat.v
index bde6f1bb4..68060900c 100644
--- a/theories/Arith/PeanoNat.v
+++ b/theories/Arith/PeanoNat.v
@@ -724,6 +724,26 @@ Definition shiftr_spec a n m (_:0<=m) := shiftr_specif a n m.
Include NExtraProp.
+(** Properties of tail-recursive addition and multiplication *)
+
+Lemma tail_add_spec n m : tail_add n m = n + m.
+Proof.
+ revert m. induction n as [|n IH]; simpl; trivial.
+ intros. now rewrite IH, add_succ_r.
+Qed.
+
+Lemma tail_addmul_spec r n m : tail_addmul r n m = r + n * m.
+Proof.
+ revert m r. induction n as [| n IH]; simpl; trivial.
+ intros. rewrite IH, tail_add_spec.
+ rewrite add_assoc. f_equal. apply add_comm.
+Qed.
+
+Lemma tail_mul_spec n m : tail_mul n m = n * m.
+Proof.
+ unfold tail_mul. now rewrite tail_addmul_spec.
+Qed.
+
End Nat.
(** Re-export notations that should be available even when
diff --git a/theories/Arith/Peano_dec.v b/theories/Arith/Peano_dec.v
index 88cda79d8..247ea20a8 100644
--- a/theories/Arith/Peano_dec.v
+++ b/theories/Arith/Peano_dec.v
@@ -57,4 +57,4 @@ now rewrite H.
Qed.
(** For compatibility *)
-Require Import Le Lt. \ No newline at end of file
+Require Import Le Lt.
diff --git a/theories/Compat/Coq87.v b/theories/Compat/Coq87.v
index 61e911678..aeef9595d 100644
--- a/theories/Compat/Coq87.v
+++ b/theories/Compat/Coq87.v
@@ -7,3 +7,14 @@
(************************************************************************)
(** Compatibility file for making Coq act similar to Coq v8.7 *)
+
+(* In 8.7, omega wasn't taking advantage of local abbreviations,
+ see bug 148 and PR#768. For adjusting this flag, we're forced to
+ first dynlink the omega plugin, but we should avoid doing a full
+ "Require Omega", since it has some undesired effects (at least on hints)
+ and breaks at least fiat-crypto. *)
+Declare ML Module "omega_plugin".
+Unset Omega UseLocalDefs.
+
+
+Set Typeclasses Axioms Are Instances.
diff --git a/theories/FSets/FSetCompat.v b/theories/FSets/FSetCompat.v
index b1769da3d..31bc1cc31 100644
--- a/theories/FSets/FSetCompat.v
+++ b/theories/FSets/FSetCompat.v
@@ -165,13 +165,13 @@ End Backport_WSets.
(** * From new Sets to new ones *)
Module Backport_Sets
- (E:OrderedType.OrderedType)
- (M:MSetInterface.Sets with Definition E.t := E.t
- with Definition E.eq := E.eq
- with Definition E.lt := E.lt)
- <: FSetInterface.S with Module E:=E.
+ (O:OrderedType.OrderedType)
+ (M:MSetInterface.Sets with Definition E.t := O.t
+ with Definition E.eq := O.eq
+ with Definition E.lt := O.lt)
+ <: FSetInterface.S with Module E:=O.
- Include Backport_WSets E M.
+ Include Backport_WSets O M.
Implicit Type s : t.
Implicit Type x y : elt.
@@ -182,21 +182,21 @@ Module Backport_Sets
Definition min_elt_1 : forall s x, min_elt s = Some x -> In x s
:= M.min_elt_spec1.
Definition min_elt_2 : forall s x y,
- min_elt s = Some x -> In y s -> ~ E.lt y x
+ min_elt s = Some x -> In y s -> ~ O.lt y x
:= M.min_elt_spec2.
Definition min_elt_3 : forall s, min_elt s = None -> Empty s
:= M.min_elt_spec3.
Definition max_elt_1 : forall s x, max_elt s = Some x -> In x s
:= M.max_elt_spec1.
Definition max_elt_2 : forall s x y,
- max_elt s = Some x -> In y s -> ~ E.lt x y
+ max_elt s = Some x -> In y s -> ~ O.lt x y
:= M.max_elt_spec2.
Definition max_elt_3 : forall s, max_elt s = None -> Empty s
:= M.max_elt_spec3.
- Definition elements_3 : forall s, sort E.lt (elements s)
+ Definition elements_3 : forall s, sort O.lt (elements s)
:= M.elements_spec2.
Definition choose_3 : forall s s' x y,
- choose s = Some x -> choose s' = Some y -> Equal s s' -> E.eq x y
+ choose s = Some x -> choose s' = Some y -> Equal s s' -> O.eq x y
:= M.choose_spec3.
Definition lt_trans : forall s s' s'', lt s s' -> lt s' s'' -> lt s s''
:= @StrictOrder_Transitive _ _ M.lt_strorder.
@@ -211,7 +211,7 @@ Module Backport_Sets
[ apply EQ | apply LT | apply GT ]; auto.
Defined.
- Module E := E.
+ Module E := O.
End Backport_Sets.
@@ -342,13 +342,13 @@ End Update_WSets.
(** * From old Sets to new ones. *)
Module Update_Sets
- (E:Orders.OrderedType)
- (M:FSetInterface.S with Definition E.t := E.t
- with Definition E.eq := E.eq
- with Definition E.lt := E.lt)
- <: MSetInterface.Sets with Module E:=E.
+ (O:Orders.OrderedType)
+ (M:FSetInterface.S with Definition E.t := O.t
+ with Definition E.eq := O.eq
+ with Definition E.lt := O.lt)
+ <: MSetInterface.Sets with Module E:=O.
- Include Update_WSets E M.
+ Include Update_WSets O M.
Implicit Type s : t.
Implicit Type x y : elt.
@@ -359,21 +359,21 @@ Module Update_Sets
Definition min_elt_spec1 : forall s x, min_elt s = Some x -> In x s
:= M.min_elt_1.
Definition min_elt_spec2 : forall s x y,
- min_elt s = Some x -> In y s -> ~ E.lt y x
+ min_elt s = Some x -> In y s -> ~ O.lt y x
:= M.min_elt_2.
Definition min_elt_spec3 : forall s, min_elt s = None -> Empty s
:= M.min_elt_3.
Definition max_elt_spec1 : forall s x, max_elt s = Some x -> In x s
:= M.max_elt_1.
Definition max_elt_spec2 : forall s x y,
- max_elt s = Some x -> In y s -> ~ E.lt x y
+ max_elt s = Some x -> In y s -> ~ O.lt x y
:= M.max_elt_2.
Definition max_elt_spec3 : forall s, max_elt s = None -> Empty s
:= M.max_elt_3.
- Definition elements_spec2 : forall s, sort E.lt (elements s)
+ Definition elements_spec2 : forall s, sort O.lt (elements s)
:= M.elements_3.
Definition choose_spec3 : forall s s' x y,
- choose s = Some x -> choose s' = Some y -> Equal s s' -> E.eq x y
+ choose s = Some x -> choose s' = Some y -> Equal s s' -> O.eq x y
:= M.choose_3.
Instance lt_strorder : StrictOrder lt.
@@ -407,6 +407,6 @@ Module Update_Sets
Lemma compare_spec : forall s s', CompSpec eq lt s s' (compare s s').
Proof. intros; unfold compare; destruct M.compare; auto. Qed.
- Module E := E.
+ Module E := O.
End Update_Sets.
diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v
index 25b042ca9..0041bfa1c 100644
--- a/theories/FSets/FSetProperties.v
+++ b/theories/FSets/FSetProperties.v
@@ -762,7 +762,7 @@ Module WProperties_fun (Import E : DecidableType)(M : WSfun E).
rewrite (cardinal_2 (s:=remove x s') (s':=s') (x:=x)); eauto with set.
Qed.
- Add Morphism cardinal : cardinal_m.
+ Add Morphism cardinal with signature (Equal ==> Logic.eq) as cardinal_m.
Proof.
exact Equal_cardinal.
Qed.
diff --git a/theories/FSets/FSets.v b/theories/FSets/FSets.v
index 572f28654..e03fb2236 100644
--- a/theories/FSets/FSets.v
+++ b/theories/FSets/FSets.v
@@ -20,4 +20,4 @@ Require Export FSetEqProperties.
Require Export FSetWeakList.
Require Export FSetList.
Require Export FSetPositive.
-Require Export FSetAVL. \ No newline at end of file
+Require Export FSetAVL.
diff --git a/theories/Init/Decimal.v b/theories/Init/Decimal.v
new file mode 100644
index 000000000..fa462f147
--- /dev/null
+++ b/theories/Init/Decimal.v
@@ -0,0 +1,161 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * Decimal numbers *)
+
+(** These numbers coded in base 10 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 10 as lists of decimal digits,
+ in big-endian order (most significant digit comes first). *)
+
+(** Unsigned integers are just lists of digits.
+ For instance, ten 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).
+
+(** [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]. *)
+
+Inductive int := Pos (d:uint) | Neg (d:uint).
+
+Delimit Scope uint_scope with uint.
+Bind Scope uint_scope with uint.
+Delimit Scope int_scope with int.
+Bind Scope int_scope with int.
+
+(** 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')
+ end.
+
+Definition rev d := revapp d Nil.
+
+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 => 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 => D0 (succ_double d)
+ | D6 d => D2 (succ_double d)
+ | D7 d => D4 (succ_double d)
+ | D8 d => D6 (succ_double d)
+ | D9 d => D8 (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 => D1 (succ_double d)
+ | D6 d => D3 (succ_double d)
+ | D7 d => D5 (succ_double d)
+ | D8 d => D7 (succ_double d)
+ | D9 d => D9 (succ_double d)
+ end.
+
+End Little.
diff --git a/theories/Init/Logic.v b/theories/Init/Logic.v
index 037d37daf..053ed601f 100644
--- a/theories/Init/Logic.v
+++ b/theories/Init/Logic.v
@@ -267,6 +267,13 @@ Notation "'exists2' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
format "'[' 'exists2' '/ ' x : A , '/ ' '[' p & '/' q ']' ']'")
: type_scope.
+Notation "'exists2' ' x , p & q" := (ex2 (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, p at level 200, right associativity) : type_scope.
+Notation "'exists2' ' x : A , p & q" := (ex2 (A:=A) (fun x => p) (fun x => q))
+ (at level 200, x strict pattern, A at level 200, p at level 200, right associativity,
+ format "'[' 'exists2' '/ ' ' x : A , '/ ' '[' p & '/' q ']' ']'")
+ : type_scope.
+
(** Derived rules for universal quantification *)
Section universal_quantification.
diff --git a/theories/Init/Nat.v b/theories/Init/Nat.v
index e942ca562..8f2648269 100644
--- a/theories/Init/Nat.v
+++ b/theories/Init/Nat.v
@@ -7,7 +7,7 @@
(************************************************************************)
Require Import Notations Logic Datatypes.
-
+Require Decimal.
Local Open Scope nat_scope.
(**********************************************************************)
@@ -134,6 +134,62 @@ Fixpoint pow n m :=
where "n ^ m" := (pow n m) : nat_scope.
+(** ** Tail-recursive versions of [add] and [mul] *)
+
+Fixpoint tail_add n m :=
+ match n with
+ | O => m
+ | S n => tail_add n (S m)
+ end.
+
+(** [tail_addmul r n m] is [r + n * m]. *)
+
+Fixpoint tail_addmul r n m :=
+ match n with
+ | O => r
+ | S n => tail_addmul (tail_add m r) n m
+ end.
+
+Definition tail_mul n m := tail_addmul 0 n m.
+
+(** ** Conversion with a decimal representation for printing/parsing *)
+
+Local Notation ten := (S (S (S (S (S (S (S (S (S (S O)))))))))).
+
+Fixpoint of_uint_acc (d:Decimal.uint)(acc:nat) :=
+ match d with
+ | Decimal.Nil => acc
+ | Decimal.D0 d => of_uint_acc d (tail_mul ten acc)
+ | Decimal.D1 d => of_uint_acc d (S (tail_mul ten acc))
+ | Decimal.D2 d => of_uint_acc d (S (S (tail_mul ten acc)))
+ | Decimal.D3 d => of_uint_acc d (S (S (S (tail_mul ten acc))))
+ | Decimal.D4 d => of_uint_acc d (S (S (S (S (tail_mul ten acc)))))
+ | Decimal.D5 d => of_uint_acc d (S (S (S (S (S (tail_mul ten acc))))))
+ | Decimal.D6 d => of_uint_acc d (S (S (S (S (S (S (tail_mul ten acc)))))))
+ | Decimal.D7 d => of_uint_acc d (S (S (S (S (S (S (S (tail_mul ten acc))))))))
+ | Decimal.D8 d => of_uint_acc d (S (S (S (S (S (S (S (S (tail_mul ten acc)))))))))
+ | Decimal.D9 d => of_uint_acc d (S (S (S (S (S (S (S (S (S (tail_mul ten acc))))))))))
+ end.
+
+Definition of_uint (d:Decimal.uint) := of_uint_acc d O.
+
+Fixpoint to_little_uint n acc :=
+ match n with
+ | O => acc
+ | S n => to_little_uint n (Decimal.Little.succ acc)
+ end.
+
+Definition to_uint n :=
+ Decimal.rev (to_little_uint n Decimal.zero).
+
+Definition of_int (d:Decimal.int) : option nat :=
+ match Decimal.norm d with
+ | Decimal.Pos u => Some (of_uint u)
+ | _ => None
+ end.
+
+Definition to_int n := Decimal.Pos (to_uint n).
+
(** ** Euclidean division *)
(** This division is linear and tail-recursive.
diff --git a/theories/Init/Notations.v b/theories/Init/Notations.v
index 5e8d2faa5..a9051e761 100644
--- a/theories/Init/Notations.v
+++ b/theories/Init/Notations.v
@@ -78,6 +78,33 @@ Reserved Notation "{ x : A | P & Q }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P }" (at level 0, x at level 99).
Reserved Notation "{ x : A & P & Q }" (at level 0, x at level 99).
+Reserved Notation "{ ' pat | P }"
+ (at level 0, pat strict pattern, format "{ ' pat | P }").
+Reserved Notation "{ ' pat | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat | P & Q }").
+
+Reserved Notation "{ ' pat : A | P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P }").
+Reserved Notation "{ ' pat : A | P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A | P & Q }").
+
+Reserved Notation "{ ' pat : A & P }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P }").
+Reserved Notation "{ ' pat : A & P & Q }"
+ (at level 0, pat strict pattern, format "{ ' pat : A & P & Q }").
+
+(** Support for Gonthier-Ssreflect's "if c is pat then u else v" *)
+
+Module IfNotations.
+
+Notation "'if' c 'is' p 'then' u 'else' v" :=
+ (match c with p => u | _ => v end)
+ (at level 200, p pattern at level 100).
+
+End IfNotations.
+
+(** Scopes *)
+
Delimit Scope type_scope with type.
Delimit Scope function_scope with function.
Delimit Scope core_scope with core.
diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v
index f0867a034..63c431e8e 100644
--- a/theories/Init/Prelude.v
+++ b/theories/Init/Prelude.v
@@ -11,6 +11,7 @@ Require Export Logic.
Require Export Logic_Type.
Require Export Datatypes.
Require Export Specif.
+Require Coq.Init.Decimal.
Require Coq.Init.Nat.
Require Export Peano.
Require Export Coq.Init.Wf.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index 3b4f833a3..47e8a7558 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -53,6 +53,15 @@ Notation "{ x : A & P }" := (sigT (A:=A) (fun x => P)) : type_scope.
Notation "{ x : A & P & Q }" := (sigT2 (A:=A) (fun x => P) (fun x => Q)) :
type_scope.
+Notation "{ ' pat | P }" := (sig (fun pat => P)) : type_scope.
+Notation "{ ' pat | P & Q }" := (sig2 (fun pat => P) (fun pat => Q)) : type_scope.
+Notation "{ ' pat : A | P }" := (sig (A:=A) (fun pat => P)) : type_scope.
+Notation "{ ' pat : A | P & Q }" := (sig2 (A:=A) (fun pat => P) (fun pat => Q)) :
+ type_scope.
+Notation "{ ' pat : A & P }" := (sigT (A:=A) (fun pat => P)) : type_scope.
+Notation "{ ' pat : A & P & Q }" := (sigT2 (A:=A) (fun pat => P) (fun pat => Q)) :
+ type_scope.
+
Add Printing Let sig.
Add Printing Let sig2.
Add Printing Let sigT.
diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v
index 5d0e7602a..47a971ef0 100644
--- a/theories/Init/Tactics.v
+++ b/theories/Init/Tactics.v
@@ -306,3 +306,10 @@ Ltac inversion_sigma_step :=
=> induction_sigma_in_using H @eq_sigT2_rect
end.
Ltac inversion_sigma := repeat inversion_sigma_step.
+
+(** A version of [time] that works for constrs *)
+Ltac time_constr tac :=
+ let eval_early := match goal with _ => restart_timer end in
+ let ret := tac () in
+ let eval_early := match goal with _ => finish_timing ( "Tactic evaluation" ) end in
+ ret.
diff --git a/theories/Init/Tauto.v b/theories/Init/Tauto.v
index 886533586..87b7a9a3b 100644
--- a/theories/Init/Tauto.v
+++ b/theories/Init/Tauto.v
@@ -27,7 +27,7 @@ Local Ltac simplif flags :=
| id: ?X1 |- _ => is_disj flags X1; elim id; intro; clear id
| id0: (forall (_: ?X1), ?X2), id1: ?X1|- _ =>
(* generalize (id0 id1); intro; clear id0 does not work
- (see Marco Maggiesi's bug PR#301)
+ (see Marco Maggiesi's BZ#301)
so we instead use Assert and exact. *)
assert X2; [exact (id0 id1) | clear id0]
| id: forall (_ : ?X1), ?X2|- _ =>
diff --git a/theories/Lists/List.v b/theories/Lists/List.v
index fbf992dbf..eae2c52de 100644
--- a/theories/Lists/List.v
+++ b/theories/Lists/List.v
@@ -2110,13 +2110,13 @@ Section Exists_Forall.
{Exists l} + {~ Exists l}.
Proof.
intro Pdec. induction l as [|a l' Hrec].
- - right. now rewrite Exists_nil.
+ - right. abstract now rewrite Exists_nil.
- destruct Hrec as [Hl'|Hl'].
* left. now apply Exists_cons_tl.
* destruct (Pdec a) as [Ha|Ha].
+ left. now apply Exists_cons_hd.
- + right. now inversion_clear 1.
- Qed.
+ + right. abstract now inversion 1.
+ Defined.
Inductive Forall : list A -> Prop :=
| Forall_nil : Forall nil
@@ -2152,9 +2152,9 @@ Section Exists_Forall.
- destruct Hrec as [Hl'|Hl'].
+ destruct (Pdec a) as [Ha|Ha].
* left. now apply Forall_cons.
- * right. now inversion_clear 1.
- + right. now inversion_clear 1.
- Qed.
+ * right. abstract now inversion 1.
+ + right. abstract now inversion 1.
+ Defined.
End One_predicate.
@@ -2179,6 +2179,16 @@ Section Exists_Forall.
* now apply Exists_cons_hd.
Qed.
+ Lemma neg_Forall_Exists_neg (P:A->Prop) (l:list A) :
+ (forall x:A, {P x} + { ~ P x }) ->
+ ~ Forall P l ->
+ Exists (fun x => ~ P x) l.
+ Proof.
+ intro Dec.
+ apply Exists_Forall_neg; intros.
+ destruct (Dec x); auto.
+ Qed.
+
Lemma Forall_Exists_dec (P:A->Prop) :
(forall x:A, {P x} + { ~ P x }) ->
forall l:list A,
@@ -2186,9 +2196,8 @@ Section Exists_Forall.
Proof.
intros Pdec l.
destruct (Forall_dec P Pdec l); [left|right]; trivial.
- apply Exists_Forall_neg; trivial.
- intro x. destruct (Pdec x); [now left|now right].
- Qed.
+ now apply neg_Forall_Exists_neg.
+ Defined.
Lemma Forall_impl : forall (P Q : A -> Prop), (forall a, P a -> Q a) ->
forall l, Forall P l -> Forall Q l.
diff --git a/theories/Logic/Classical_Prop.v b/theories/Logic/Classical_Prop.v
index f7b53f1dc..a5ae07b64 100644
--- a/theories/Logic/Classical_Prop.v
+++ b/theories/Logic/Classical_Prop.v
@@ -97,12 +97,12 @@ Proof proof_irrelevance_cci classic.
(* classical_left transforms |- A \/ B into ~B |- A *)
(* classical_right transforms |- A \/ B into ~A |- B *)
-Ltac classical_right := match goal with
- | _:_ |-?X1 \/ _ => (elim (classic X1);intro;[left;trivial|right])
+Ltac classical_right := match goal with
+|- ?X \/ _ => (elim (classic X);intro;[left;trivial|right])
end.
Ltac classical_left := match goal with
-| _:_ |- _ \/?X1 => (elim (classic X1);intro;[right;trivial|left])
+|- _ \/ ?X => (elim (classic X);intro;[right;trivial|left])
end.
Require Export EqdepFacts.
diff --git a/theories/Logic/FunctionalExtensionality.v b/theories/Logic/FunctionalExtensionality.v
index ac95ddd0c..82b04d132 100644
--- a/theories/Logic/FunctionalExtensionality.v
+++ b/theories/Logic/FunctionalExtensionality.v
@@ -221,13 +221,12 @@ Tactic Notation "extensionality" "in" hyp(H) :=
(* If we [subst H], things break if we already have another equation of the form [_ = H] *)
destruct Heq; rename H_out into H.
-(** Eta expansion follows from extensionality. *)
+(** Eta expansion is built into Coq. *)
Lemma eta_expansion_dep {A} {B : A -> Type} (f : forall x : A, B x) :
f = fun x => f x.
Proof.
intros.
- extensionality x.
reflexivity.
Qed.
diff --git a/theories/MSets/MSetGenTree.v b/theories/MSets/MSetGenTree.v
index 036ff1aa4..9fb8e499b 100644
--- a/theories/MSets/MSetGenTree.v
+++ b/theories/MSets/MSetGenTree.v
@@ -1144,4 +1144,4 @@ Proof.
apply mindepth_cardinal.
Qed.
-End Props. \ No newline at end of file
+End Props.
diff --git a/theories/MSets/MSets.v b/theories/MSets/MSets.v
index f179bcd1d..1ee485cc1 100644
--- a/theories/MSets/MSets.v
+++ b/theories/MSets/MSets.v
@@ -18,4 +18,4 @@ Require Export MSetEqProperties.
Require Export MSetWeakList.
Require Export MSetList.
Require Export MSetPositive.
-Require Export MSetAVL. \ No newline at end of file
+Require Export MSetAVL.
diff --git a/theories/NArith/BinNatDef.v b/theories/NArith/BinNatDef.v
index ba923d062..e07758914 100644
--- a/theories/NArith/BinNatDef.v
+++ b/theories/NArith/BinNatDef.v
@@ -378,4 +378,22 @@ Definition iter (n:N) {A} (f:A->A) (x:A) : A :=
| pos p => Pos.iter f x p
end.
-End N. \ No newline at end of file
+(** Conversion with a decimal representation for printing/parsing *)
+
+Definition of_uint (d:Decimal.uint) := Pos.of_uint d.
+
+Definition of_int (d:Decimal.int) :=
+ match Decimal.norm d with
+ | Decimal.Pos d => Some (Pos.of_uint d)
+ | Decimal.Neg _ => None
+ end.
+
+Definition to_uint n :=
+ match n with
+ | 0 => Decimal.zero
+ | pos p => Pos.to_uint p
+ end.
+
+Definition to_int n := Decimal.Pos (to_uint n).
+
+End N.
diff --git a/theories/Numbers/DecimalFacts.v b/theories/Numbers/DecimalFacts.v
new file mode 100644
index 000000000..3eef63c7f
--- /dev/null
+++ b/theories/Numbers/DecimalFacts.v
@@ -0,0 +1,141 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * DecimalFacts : some facts about Decimal numbers *)
+
+Require Import Decimal.
+
+Lemma uint_dec (d d' : uint) : { d = d' } + { d <> d' }.
+Proof.
+ decide equality.
+Defined.
+
+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.
+
+(** 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.
+
+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_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 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 nzhead_invol d : nzhead (nzhead d) = nzhead d.
+Proof.
+ now induction d.
+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.
diff --git a/theories/Numbers/DecimalN.v b/theories/Numbers/DecimalN.v
new file mode 100644
index 000000000..998f009a7
--- /dev/null
+++ b/theories/Numbers/DecimalN.v
@@ -0,0 +1,105 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * DecimalN
+
+ Proofs that conversions between decimal numbers and [N]
+ are bijections *)
+
+Require Import Decimal DecimalFacts DecimalPos PArith NArith.
+
+Module Unsigned.
+
+Lemma of_to (n:N) : N.of_uint (N.to_uint n) = n.
+Proof.
+ 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).
+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.
+Qed.
+
+Lemma to_uint_surj d : exists p, N.to_uint p = unorm d.
+Proof.
+ 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.
+Qed.
+
+Lemma of_inj d d' :
+ N.of_uint d = N.of_uint d' -> unorm d = unorm d'.
+Proof.
+ 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.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed decimal numbers *)
+
+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.
+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 [= <-].
+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'.
+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.
+Qed.
+
+Lemma of_int_norm d : N.of_int (norm d) = N.of_int d.
+Proof.
+ 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.
+Qed.
+
+End Signed.
diff --git a/theories/Numbers/DecimalNat.v b/theories/Numbers/DecimalNat.v
new file mode 100644
index 000000000..4aa189e24
--- /dev/null
+++ b/theories/Numbers/DecimalNat.v
@@ -0,0 +1,300 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * DecimalNat
+
+ Proofs that conversions between decimal numbers and [nat]
+ are bijections. *)
+
+Require Import Decimal DecimalFacts Arith.
+
+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.
+
+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.
+
+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)
+ end.
+
+(** 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.
+
+(** A direct version of [of_little_uint] *)
+Fixpoint of_lu (d:uint) : nat :=
+ match d with
+ | Nil => 0
+ | D0 d => 10 * of_lu d
+ | D1 d => 1 + 10 * of_lu d
+ | D2 d => 2 + 10 * of_lu d
+ | D3 d => 3 + 10 * of_lu d
+ | D4 d => 4 + 10 * of_lu d
+ | D5 d => 5 + 10 * of_lu d
+ | D6 d => 6 + 10 * of_lu d
+ | D7 d => 7 + 10 * of_lu d
+ | D8 d => 8 + 10 * of_lu d
+ | D9 d => 9 + 10 * 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_uint n (Little.succ d) =
+ Little.succ (Nat.to_little_uint n d).
+Proof.
+ 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.
+Qed.
+
+Lemma to_uint_alt n :
+ Nat.to_uint n = rev (to_lu n).
+Proof.
+ 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).
+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 10).
+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' * 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.
+Qed.
+
+Lemma of_uint_acc_spec n 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.
+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.
+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.
+Qed.
+
+(** The other direction *)
+
+Lemma to_lu_tenfold n : n<>0 ->
+ 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).
+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_tenfold 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).
+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.
+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.
+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.
+Qed.
+
+Lemma to_uint_surj d : exists n, Nat.to_uint n = unorm d.
+Proof.
+ 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.
+Qed.
+
+Lemma of_inj d d' :
+ Nat.of_uint d = Nat.of_uint d' -> unorm d = unorm d'.
+Proof.
+ 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.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed decimal numbers *)
+
+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.
+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 [= <-].
+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'.
+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.
+Qed.
+
+Lemma of_int_norm d : Nat.of_int (norm d) = Nat.of_int d.
+Proof.
+ 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'.
+Proof.
+ 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
new file mode 100644
index 000000000..40c8f5a5a
--- /dev/null
+++ b/theories/Numbers/DecimalPos.v
@@ -0,0 +1,381 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * DecimalPos
+
+ Proofs that conversions between decimal numbers and [positive]
+ are bijections. *)
+
+Require Import Decimal DecimalFacts 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 => 10 * of_lu d
+ | D1 d => 1 + 10 * of_lu d
+ | D2 d => 2 + 10 * of_lu d
+ | D3 d => 3 + 10 * of_lu d
+ | D4 d => 4 + 10 * of_lu d
+ | D5 d => 5 + 10 * of_lu d
+ | D6 d => 6 + 10 * of_lu d
+ | D7 d => 7 + 10 * of_lu d
+ | D8 d => 8 + 10 * of_lu d
+ | D9 d => 9 + 10 * of_lu d
+ 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.
+
+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.
+
+Lemma of_lu_eqn d :
+ of_lu d = hd d + 10 * (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)
+ end.
+
+Lemma of_lu_revapp d 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.
+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_uint_acc d acc = Pos.of_uint_acc (tl d) (Nadd (hd d) (10*acc)).
+Proof.
+ 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.
+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_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_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_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.
+Qed.
+
+(** The other direction *)
+
+Definition to_lu n :=
+ match n with
+ | N0 => Decimal.zero
+ | Npos p => Pos.to_little_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_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.
+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 (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).
+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_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.
+Qed.
+
+(** Some consequences *)
+
+Lemma to_uint_nonzero p : Pos.to_uint p <> zero.
+Proof.
+ 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.
+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'.
+Qed.
+
+Lemma to_uint_pos_surj 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.
+Qed.
+
+Lemma of_uint_norm d : Pos.of_uint (unorm d) = Pos.of_uint d.
+Proof.
+ now induction d.
+Qed.
+
+Lemma of_inj d d' :
+ Pos.of_uint d = Pos.of_uint d' -> unorm d = unorm d'.
+Proof.
+ 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.
+Qed.
+
+End Unsigned.
+
+(** Conversion from/to signed decimal numbers *)
+
+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.
+Qed.
+
+Lemma to_of (d:int)(p:positive) :
+ 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 [= <-].
+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'.
+Qed.
+
+Lemma to_int_pos_surj 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.
+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.
+Qed.
+
+Lemma of_inj_pos d 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.
+Qed.
+
+End Signed.
diff --git a/theories/Numbers/DecimalString.v b/theories/Numbers/DecimalString.v
new file mode 100644
index 000000000..7e937f481
--- /dev/null
+++ b/theories/Numbers/DecimalString.v
@@ -0,0 +1,263 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+Require Import Decimal Ascii String.
+
+(** * Conversion between decimal numbers and Coq strings *)
+
+(** Pretty straightforward, which is precisely the point of the
+ [Decimal.int] datatype. The only catch is [Decimal.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. *)
+
+Local Open Scope string_scope.
+
+(** Parsing one char *)
+
+Definition uint_of_char (a:ascii)(d:option uint) :=
+ match d with
+ | None => None
+ | Some d =>
+ match a with
+ | "0" => Some (D0 d)
+ | "1" => Some (D1 d)
+ | "2" => Some (D2 d)
+ | "3" => Some (D3 d)
+ | "4" => Some (D4 d)
+ | "5" => Some (D5 d)
+ | "6" => Some (D6 d)
+ | "7" => Some (D7 d)
+ | "8" => Some (D8 d)
+ | "9" => Some (D9 d)
+ | _ => None
+ end
+ end%char.
+
+Lemma uint_of_char_spec c d d' :
+ uint_of_char c (Some d) = Some d' ->
+ (c = "0" /\ d' = D0 d \/
+ c = "1" /\ d' = D1 d \/
+ c = "2" /\ d' = D2 d \/
+ c = "3" /\ d' = D3 d \/
+ c = "4" /\ d' = D4 d \/
+ c = "5" /\ d' = D5 d \/
+ c = "6" /\ d' = D6 d \/
+ c = "7" /\ d' = D7 d \/
+ c = "8" /\ d' = D8 d \/
+ c = "9" /\ d' = D9 d)%char.
+Proof.
+ destruct c as [[|] [|] [|] [|] [|] [|] [|] [|]];
+ intros [= <-]; intuition.
+Qed.
+
+(** Decimal/String conversion where [Nil] is [""] *)
+
+Module NilEmpty.
+
+Fixpoint string_of_uint (d:uint) :=
+ match d with
+ | Nil => EmptyString
+ | D0 d => String "0" (string_of_uint d)
+ | D1 d => String "1" (string_of_uint d)
+ | D2 d => String "2" (string_of_uint d)
+ | D3 d => String "3" (string_of_uint d)
+ | D4 d => String "4" (string_of_uint d)
+ | D5 d => String "5" (string_of_uint d)
+ | D6 d => String "6" (string_of_uint d)
+ | D7 d => String "7" (string_of_uint d)
+ | D8 d => String "8" (string_of_uint d)
+ | D9 d => String "9" (string_of_uint d)
+ end.
+
+Fixpoint uint_of_string s :=
+ match s with
+ | EmptyString => Some Nil
+ | String a s => uint_of_char a (uint_of_string s)
+ end.
+
+Definition string_of_int (d:int) :=
+ match d with
+ | Pos d => string_of_uint d
+ | Neg d => String "-" (string_of_uint d)
+ end.
+
+Definition int_of_string s :=
+ match s with
+ | EmptyString => Some (Pos Nil)
+ | String a s' =>
+ if ascii_dec 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.
+ + destruct ascii_dec; subst.
+ * 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.
+ destruct ascii_dec; subst; simpl.
+ - 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.
+
+(** Decimal/String conversions where [Nil] is ["0"] *)
+
+Module NilZero.
+
+Definition string_of_uint (d:uint) :=
+ match d with
+ | Nil => "0"
+ | _ => NilEmpty.string_of_uint d
+ end.
+
+Definition uint_of_string s :=
+ match s with
+ | EmptyString => None
+ | _ => NilEmpty.uint_of_string s
+ end.
+
+Definition string_of_int (d:int) :=
+ match d with
+ | Pos d => string_of_uint d
+ | Neg d => String "-" (string_of_uint d)
+ end.
+
+Definition int_of_string s :=
+ match s with
+ | EmptyString => None
+ | String a s' =>
+ if ascii_dec 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 Decimal.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 Decimal.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.
+ + destruct ascii_dec; subst.
+ * 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).
+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.
+ destruct ascii_dec; subst; simpl.
+ - 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
new file mode 100644
index 000000000..92d66ecfb
--- /dev/null
+++ b/theories/Numbers/DecimalZ.v
@@ -0,0 +1,73 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(** * DecimalZ
+
+ Proofs that conversions between decimal numbers and [Z]
+ are bijections. *)
+
+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.
+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.
+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.
+Qed.
+
+Lemma to_int_surj d : exists n, Z.to_int n = norm d.
+Proof.
+ 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.
+Qed.
+
+Lemma of_inj d d' :
+ Z.of_int d = Z.of_int d' -> norm d = norm d'.
+Proof.
+ 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.
+Qed.
diff --git a/theories/Numbers/NatInt/NZParity.v b/theories/Numbers/NatInt/NZParity.v
index de3bbbca7..626d59d73 100644
--- a/theories/Numbers/NatInt/NZParity.v
+++ b/theories/Numbers/NatInt/NZParity.v
@@ -260,4 +260,4 @@ Proof.
intros. apply odd_add_mul_even. apply even_spec, even_2.
Qed.
-End NZParityProp. \ No newline at end of file
+End NZParityProp.
diff --git a/theories/PArith/BinPosDef.v b/theories/PArith/BinPosDef.v
index 2b647555c..a77c26e5a 100644
--- a/theories/PArith/BinPosDef.v
+++ b/theories/PArith/BinPosDef.v
@@ -557,4 +557,59 @@ Fixpoint of_succ_nat (n:nat) : positive :=
| S x => succ (of_succ_nat x)
end.
+(** ** Conversion with a decimal representation for printing/parsing *)
+
+Local Notation ten := 1~0~1~0.
+
+Fixpoint of_uint_acc (d:Decimal.uint)(acc:positive) :=
+ match d with
+ | Decimal.Nil => acc
+ | Decimal.D0 l => of_uint_acc l (mul ten acc)
+ | Decimal.D1 l => of_uint_acc l (add 1 (mul ten acc))
+ | Decimal.D2 l => of_uint_acc l (add 1~0 (mul ten acc))
+ | Decimal.D3 l => of_uint_acc l (add 1~1 (mul ten acc))
+ | Decimal.D4 l => of_uint_acc l (add 1~0~0 (mul ten acc))
+ | Decimal.D5 l => of_uint_acc l (add 1~0~1 (mul ten acc))
+ | Decimal.D6 l => of_uint_acc l (add 1~1~0 (mul ten acc))
+ | Decimal.D7 l => of_uint_acc l (add 1~1~1 (mul ten acc))
+ | Decimal.D8 l => of_uint_acc l (add 1~0~0~0 (mul ten acc))
+ | Decimal.D9 l => of_uint_acc l (add 1~0~0~1 (mul ten acc))
+ end.
+
+Fixpoint of_uint (d:Decimal.uint) : N :=
+ match d with
+ | Decimal.Nil => N0
+ | Decimal.D0 l => of_uint l
+ | Decimal.D1 l => Npos (of_uint_acc l 1)
+ | Decimal.D2 l => Npos (of_uint_acc l 1~0)
+ | Decimal.D3 l => Npos (of_uint_acc l 1~1)
+ | Decimal.D4 l => Npos (of_uint_acc l 1~0~0)
+ | Decimal.D5 l => Npos (of_uint_acc l 1~0~1)
+ | Decimal.D6 l => Npos (of_uint_acc l 1~1~0)
+ | Decimal.D7 l => Npos (of_uint_acc l 1~1~1)
+ | Decimal.D8 l => Npos (of_uint_acc l 1~0~0~0)
+ | Decimal.D9 l => Npos (of_uint_acc l 1~0~0~1)
+ end.
+
+Definition of_int (d:Decimal.int) : option positive :=
+ match d with
+ | Decimal.Pos d =>
+ match of_uint d with
+ | N0 => None
+ | Npos p => Some p
+ end
+ | Decimal.Neg _ => None
+ end.
+
+Fixpoint to_little_uint p :=
+ match p with
+ | 1 => Decimal.D1 Decimal.Nil
+ | p~1 => Decimal.Little.succ_double (to_little_uint p)
+ | p~0 => Decimal.Little.double (to_little_uint p)
+ end.
+
+Definition to_uint p := Decimal.rev (to_little_uint p).
+
+Definition to_int n := Decimal.Pos (to_uint n).
+
End Pos.
diff --git a/theories/Program/Combinators.v b/theories/Program/Combinators.v
index 90db10ef1..237d878bf 100644
--- a/theories/Program/Combinators.v
+++ b/theories/Program/Combinators.v
@@ -22,15 +22,13 @@ Open Scope program_scope.
Lemma compose_id_left : forall A B (f : A -> B), id ∘ f = f.
Proof.
intros.
- unfold id, compose.
- symmetry. apply eta_expansion.
+ reflexivity.
Qed.
Lemma compose_id_right : forall A B (f : A -> B), f ∘ id = f.
Proof.
intros.
- unfold id, compose.
- symmetry ; apply eta_expansion.
+ reflexivity.
Qed.
Lemma compose_assoc : forall A B C D (f : A -> B) (g : B -> C) (h : C -> D),
@@ -47,9 +45,7 @@ Hint Rewrite <- @compose_assoc : core.
Lemma flip_flip : forall A B C, @flip A B C ∘ flip = id.
Proof.
- unfold flip, compose.
intros.
- extensionality x ; extensionality y ; extensionality z.
reflexivity.
Qed.
@@ -57,9 +53,7 @@ Qed.
Lemma prod_uncurry_curry : forall A B C, @prod_uncurry A B C ∘ prod_curry = id.
Proof.
- simpl ; intros.
- unfold prod_uncurry, prod_curry, compose.
- extensionality x ; extensionality y ; extensionality z.
+ intros.
reflexivity.
Qed.
diff --git a/theories/Program/Tactics.v b/theories/Program/Tactics.v
index 9aca56f47..b06562fc4 100644
--- a/theories/Program/Tactics.v
+++ b/theories/Program/Tactics.v
@@ -328,4 +328,4 @@ Ltac program_simpl := program_simplify ; try typeclasses eauto with program ; tr
Obligation Tactic := program_simpl.
-Definition obligation (A : Type) {a : A} := a. \ No newline at end of file
+Definition obligation (A : Type) {a : A} := a.
diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v
index a19f9f902..5996d30f2 100644
--- a/theories/QArith/QArith_base.v
+++ b/theories/QArith/QArith_base.v
@@ -171,6 +171,26 @@ Proof.
auto with qarith.
Qed.
+Lemma Qeq_bool_comm x y: Qeq_bool x y = Qeq_bool y x.
+Proof.
+ apply eq_true_iff_eq. rewrite !Qeq_bool_iff. now symmetry.
+Qed.
+
+Lemma Qeq_bool_refl x: Qeq_bool x x = true.
+Proof.
+ rewrite Qeq_bool_iff. now reflexivity.
+Qed.
+
+Lemma Qeq_bool_sym x y: Qeq_bool x y = true -> Qeq_bool y x = true.
+Proof.
+ rewrite !Qeq_bool_iff. now symmetry.
+Qed.
+
+Lemma Qeq_bool_trans x y z: Qeq_bool x y = true -> Qeq_bool y z = true -> Qeq_bool x z = true.
+Proof.
+ rewrite !Qeq_bool_iff; apply Qeq_trans.
+Qed.
+
Hint Resolve Qnot_eq_sym : qarith.
(** * Addition, multiplication and opposite *)
diff --git a/theories/QArith/Qabs.v b/theories/QArith/Qabs.v
index 116aa0d42..ec2ac7832 100644
--- a/theories/QArith/Qabs.v
+++ b/theories/QArith/Qabs.v
@@ -100,6 +100,13 @@ rewrite Z.abs_mul.
reflexivity.
Qed.
+Lemma Qabs_Qinv : forall q, Qabs (/ q) == / (Qabs q).
+Proof.
+ intros [n d]; simpl.
+ unfold Qinv.
+ case_eq n; intros; simpl in *; apply Qeq_refl.
+Qed.
+
Lemma Qabs_Qminus x y: Qabs (x - y) = Qabs (y - x).
Proof.
unfold Qminus, Qopp. simpl.
diff --git a/theories/QArith/Qcabs.v b/theories/QArith/Qcabs.v
index 1883c77be..09908665e 100644
--- a/theories/QArith/Qcabs.v
+++ b/theories/QArith/Qcabs.v
@@ -126,4 +126,4 @@ Proof.
destruct (proj1 (Qcabs_Qcle_condition x 0)) as [A B].
+ rewrite H; apply Qcle_refl.
+ apply Qcle_antisym; auto.
-Qed. \ No newline at end of file
+Qed.
diff --git a/theories/QArith/Qreduction.v b/theories/QArith/Qreduction.v
index 88e1298fb..5d055b547 100644
--- a/theories/QArith/Qreduction.v
+++ b/theories/QArith/Qreduction.v
@@ -101,7 +101,7 @@ Proof.
- apply Qred_complete.
Qed.
-Add Morphism Qred : Qred_comp.
+Add Morphism Qred with signature (Qeq ==> Qeq) as Qred_comp.
Proof.
intros. now rewrite !Qred_correct.
Qed.
@@ -125,19 +125,19 @@ Proof.
intros; unfold Qminus'; apply Qred_correct; auto.
Qed.
-Add Morphism Qplus' : Qplus'_comp.
+Add Morphism Qplus' with signature (Qeq ==> Qeq ==> Qeq) as Qplus'_comp.
Proof.
intros; unfold Qplus'.
rewrite H, H0; auto with qarith.
Qed.
-Add Morphism Qmult' : Qmult'_comp.
+Add Morphism Qmult' with signature (Qeq ==> Qeq ==> Qeq) as Qmult'_comp.
Proof.
intros; unfold Qmult'.
rewrite H, H0; auto with qarith.
Qed.
-Add Morphism Qminus' : Qminus'_comp.
+Add Morphism Qminus' with signature (Qeq ==> Qeq ==> Qeq) as Qminus'_comp.
Proof.
intros; unfold Qminus'.
rewrite H, H0; auto with qarith.
diff --git a/theories/Reals/Ranalysis.v b/theories/Reals/Ranalysis.v
index 66e37e867..9b0357f03 100644
--- a/theories/Reals/Ranalysis.v
+++ b/theories/Reals/Ranalysis.v
@@ -26,4 +26,4 @@ Require Export RList.
Require Export Sqrt_reg.
Require Export Ranalysis4.
Require Export Rpower.
-Require Export Ranalysis_reg. \ No newline at end of file
+Require Export Ranalysis_reg.
diff --git a/theories/Sets/Powerset_facts.v b/theories/Sets/Powerset_facts.v
index 2dd559a95..209c22f71 100644
--- a/theories/Sets/Powerset_facts.v
+++ b/theories/Sets/Powerset_facts.v
@@ -40,6 +40,11 @@ Section Sets_as_an_algebra.
auto 6 with sets.
Qed.
+ Theorem Empty_set_zero_right : forall X:Ensemble U, Union U X (Empty_set U) = X.
+ Proof.
+ auto 6 with sets.
+ Qed.
+
Theorem Empty_set_zero' : forall x:U, Add U (Empty_set U) x = Singleton U x.
Proof.
unfold Add at 1; auto using Empty_set_zero with sets.
@@ -131,6 +136,17 @@ Section Sets_as_an_algebra.
elim H'; intros x0 H'0; elim H'0; auto with sets.
Qed.
+ Lemma Distributivity_l
+ : forall (A B C : Ensemble U),
+ Intersection U (Union U A B) C =
+ Union U (Intersection U A C) (Intersection U B C).
+ Proof.
+ intros A B C.
+ rewrite Intersection_commutative.
+ rewrite Distributivity.
+ f_equal; apply Intersection_commutative.
+ Qed.
+
Theorem Distributivity' :
forall A B C:Ensemble U,
Union U A (Intersection U B C) =
@@ -251,6 +267,81 @@ Section Sets_as_an_algebra.
intros; apply Definition_of_covers; auto with sets.
Qed.
+ Lemma Disjoint_Intersection:
+ forall A s1 s2, Disjoint A s1 s2 -> Intersection A s1 s2 = Empty_set A.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * destruct H.
+ intros x H1. unfold In in *. exfalso. intuition. apply (H _ H1).
+ * intuition.
+ Qed.
+
+ Lemma Intersection_Empty_set_l:
+ forall A s, Intersection A (Empty_set A) s = Empty_set A.
+ Proof.
+ intros. auto with sets.
+ Qed.
+
+ Lemma Intersection_Empty_set_r:
+ forall A s, Intersection A s (Empty_set A) = Empty_set A.
+ Proof.
+ intros. auto with sets.
+ Qed.
+
+ Lemma Seminus_Empty_set_l:
+ forall A s, Setminus A (Empty_set A) s = Empty_set A.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H1. destruct H1. unfold In in *. assumption.
+ * intuition.
+ Qed.
+
+ Lemma Seminus_Empty_set_r:
+ forall A s, Setminus A s (Empty_set A) = s.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H1. destruct H1. unfold In in *. assumption.
+ * intuition.
+ Qed.
+
+ Lemma Setminus_Union_l:
+ forall A s1 s2 s3,
+ Setminus A (Union A s1 s2) s3 = Union A (Setminus A s1 s3) (Setminus A s2 s3).
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H. inversion H. inversion H0; intuition.
+ * intros x H. constructor; inversion H; inversion H0; intuition.
+ Qed.
+
+ Lemma Setminus_Union_r:
+ forall A s1 s2 s3,
+ Setminus A s1 (Union A s2 s3) = Setminus A (Setminus A s1 s2) s3.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H. inversion H. constructor. intuition. contradict H1. intuition.
+ * intros x H. inversion H. inversion H0. constructor; intuition. inversion H4; intuition.
+ Qed.
+
+ Lemma Setminus_Disjoint_noop:
+ forall A s1 s2,
+ Intersection A s1 s2 = Empty_set A -> Setminus A s1 s2 = s1.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H1. inversion_clear H1. intuition.
+ * intros x H1. constructor; intuition. contradict H.
+ apply Inhabited_not_empty.
+ exists x. intuition.
+ Qed.
+
+ Lemma Setminus_Included_empty:
+ forall A s1 s2,
+ Included A s1 s2 -> Setminus A s1 s2 = Empty_set A.
+ Proof.
+ intros. apply Extensionality_Ensembles. split.
+ * intros x H1. inversion_clear H1. contradiction H2. intuition.
+ * intuition.
+ Qed.
+
End Sets_as_an_algebra.
Hint Resolve Empty_set_zero Empty_set_zero' Union_associative Union_add
diff --git a/theories/Unicode/Utf8_core.v b/theories/Unicode/Utf8_core.v
index a0545c0a4..95c8336d2 100644
--- a/theories/Unicode/Utf8_core.v
+++ b/theories/Unicode/Utf8_core.v
@@ -10,10 +10,12 @@
(* Logic *)
-Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
- (at level 200, x binder, y binder, right associativity) : type_scope.
-Notation "∃ x .. y , P" := (exists x, .. (exists y, P) ..)
- (at level 200, x binder, y binder, right associativity) : type_scope.
+Notation "∀ x .. y , P" := (forall x, .. (forall y, P) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' ∀ x .. y ']' , P") : type_scope.
+Notation "∃ x .. y , P" := (exists x, .. (exists y, P) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' ∃ x .. y ']' , P") : type_scope.
Notation "x ∨ y" := (x \/ y) (at level 85, right associativity) : type_scope.
Notation "x ∧ y" := (x /\ y) (at level 80, right associativity) : type_scope.
@@ -25,5 +27,6 @@ Notation "¬ x" := (~x) (at level 75, right associativity) : type_scope.
Notation "x ≠ y" := (x <> y) (at level 70) : type_scope.
(* Abstraction *)
-Notation "'λ' x .. y , t" := (fun x => .. (fun y => t) ..)
- (at level 200, x binder, y binder, right associativity).
+Notation "'λ' x .. y , t" := (fun x => .. (fun y => t) ..)
+ (at level 200, x binder, y binder, right associativity,
+ format "'[ ' 'λ' x .. y ']' , t").
diff --git a/theories/Vectors/Vector.v b/theories/Vectors/Vector.v
index 672858fa5..19d749fc8 100644
--- a/theories/Vectors/Vector.v
+++ b/theories/Vectors/Vector.v
@@ -21,4 +21,4 @@ Require VectorSpec.
Require VectorEq.
Include VectorDef.
Include VectorSpec.
-Include VectorEq. \ No newline at end of file
+Include VectorEq.
diff --git a/theories/ZArith/BinIntDef.v b/theories/ZArith/BinIntDef.v
index 7686fbae8..a0393f318 100644
--- a/theories/ZArith/BinIntDef.v
+++ b/theories/ZArith/BinIntDef.v
@@ -299,6 +299,23 @@ Definition to_pos (z:Z) : positive :=
| _ => 1%positive
end.
+(** Conversion with a decimal representation for printing/parsing *)
+
+Definition of_uint (d:Decimal.uint) := of_N (Pos.of_uint d).
+
+Definition of_int (d:Decimal.int) :=
+ match d with
+ | Decimal.Pos d => of_uint d
+ | Decimal.Neg d => opp (of_uint d)
+ end.
+
+Definition to_int n :=
+ match n with
+ | 0 => Decimal.Pos Decimal.zero
+ | pos p => Decimal.Pos (Pos.to_uint p)
+ | neg p => Decimal.Neg (Pos.to_uint p)
+ end.
+
(** ** Iteration of a function
By convention, iterating a negative number of times is identity.
@@ -616,4 +633,4 @@ Definition lxor a b :=
| neg a, neg b => of_N (N.lxor (Pos.pred_N a) (Pos.pred_N b))
end.
-End Z. \ No newline at end of file
+End Z.
diff --git a/theories/ZArith/Zsqrt_compat.v b/theories/ZArith/Zsqrt_compat.v
index fb7f71b4b..cccd970da 100644
--- a/theories/ZArith/Zsqrt_compat.v
+++ b/theories/ZArith/Zsqrt_compat.v
@@ -229,4 +229,4 @@ Proof.
symmetry. apply Z.sqrt_unique; trivial.
now apply Zsqrt_interval.
now destruct n.
-Qed. \ No newline at end of file
+Qed.