diff options
Diffstat (limited to 'theories')
409 files changed, 5177 insertions, 2756 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/Arith.v b/theories/Arith/Arith.v index 649819878..1cba8fafe 100644 --- a/theories/Arith/Arith.v +++ b/theories/Arith/Arith.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Arith_base. diff --git a/theories/Arith/Arith_base.v b/theories/Arith/Arith_base.v index 1493deb48..e3a033a4a 100644 --- a/theories/Arith/Arith_base.v +++ b/theories/Arith/Arith_base.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export PeanoNat. diff --git a/theories/Arith/Between.v b/theories/Arith/Between.v index ead08b3eb..25d84a621 100644 --- a/theories/Arith/Between.v +++ b/theories/Arith/Between.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Le. diff --git a/theories/Arith/Bool_nat.v b/theories/Arith/Bool_nat.v index a1eaf02f7..d892542e7 100644 --- a/theories/Arith/Bool_nat.v +++ b/theories/Arith/Bool_nat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Compare_dec. diff --git a/theories/Arith/Compare.v b/theories/Arith/Compare.v index 8381be5ce..6778d6a02 100644 --- a/theories/Arith/Compare.v +++ b/theories/Arith/Compare.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Equality is decidable on [nat] *) diff --git a/theories/Arith/Compare_dec.v b/theories/Arith/Compare_dec.v index 1e3237d10..713aef858 100644 --- a/theories/Arith/Compare_dec.v +++ b/theories/Arith/Compare_dec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Le Lt Gt Decidable PeanoNat. @@ -133,11 +135,11 @@ Qed. See now [Nat.compare] and its properties. In scope [nat_scope], the notation for [Nat.compare] is "?=" *) -Notation nat_compare := Nat.compare (compat "8.4"). +Notation nat_compare := Nat.compare (compat "8.6"). -Notation nat_compare_spec := Nat.compare_spec (compat "8.4"). -Notation nat_compare_eq_iff := Nat.compare_eq_iff (compat "8.4"). -Notation nat_compare_S := Nat.compare_succ (compat "8.4"). +Notation nat_compare_spec := Nat.compare_spec (compat "8.6"). +Notation nat_compare_eq_iff := Nat.compare_eq_iff (compat "8.6"). +Notation nat_compare_S := Nat.compare_succ (only parsing). Lemma nat_compare_lt n m : n<m <-> (n ?= m) = Lt. Proof. @@ -198,9 +200,9 @@ Qed. See now [Nat.leb] and its properties. In scope [nat_scope], the notation for [Nat.leb] is "<=?" *) -Notation leb := Nat.leb (compat "8.4"). +Notation leb := Nat.leb (only parsing). -Notation leb_iff := Nat.leb_le (compat "8.4"). +Notation leb_iff := Nat.leb_le (only parsing). Lemma leb_iff_conv m n : (n <=? m) = false <-> m < n. Proof. diff --git a/theories/Arith/Div2.v b/theories/Arith/Div2.v index ecb9a5706..42956c475 100644 --- a/theories/Arith/Div2.v +++ b/theories/Arith/Div2.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Nota : this file is OBSOLETE, and left only for compatibility. @@ -18,7 +20,7 @@ Implicit Type n : nat. (** Here we define [n/2] and prove some of its properties *) -Notation div2 := Nat.div2 (compat "8.4"). +Notation div2 := Nat.div2 (only parsing). (** Since [div2] is recursively defined on [0], [1] and [(S (S n))], it is useful to prove the corresponding induction principle *) @@ -84,7 +86,7 @@ Qed. (** Properties related to the double ([2n]) *) -Notation double := Nat.double (compat "8.4"). +Notation double := Nat.double (only parsing). Hint Unfold double Nat.double: arith. diff --git a/theories/Arith/EqNat.v b/theories/Arith/EqNat.v index 722615428..4b51dfc00 100644 --- a/theories/Arith/EqNat.v +++ b/theories/Arith/EqNat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 PeanoNat. @@ -69,10 +71,10 @@ Defined. We reuse the one already defined in module [Nat]. In scope [nat_scope], the notation "=?" can be used. *) -Notation beq_nat := Nat.eqb (compat "8.4"). +Notation beq_nat := Nat.eqb (only parsing). -Notation beq_nat_true_iff := Nat.eqb_eq (compat "8.4"). -Notation beq_nat_false_iff := Nat.eqb_neq (compat "8.4"). +Notation beq_nat_true_iff := Nat.eqb_eq (only parsing). +Notation beq_nat_false_iff := Nat.eqb_neq (only parsing). Lemma beq_nat_refl n : true = (n =? n). Proof. diff --git a/theories/Arith/Euclid.v b/theories/Arith/Euclid.v index 6c6bf7fef..29f4d3e23 100644 --- a/theories/Arith/Euclid.v +++ b/theories/Arith/Euclid.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Mult. diff --git a/theories/Arith/Even.v b/theories/Arith/Even.v index f30d05c7a..baf119732 100644 --- a/theories/Arith/Even.v +++ b/theories/Arith/Even.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Nota : this file is OBSOLETE, and left only for compatibility. diff --git a/theories/Arith/Factorial.v b/theories/Arith/Factorial.v index 0625c03da..22f586d7e 100644 --- a/theories/Arith/Factorial.v +++ b/theories/Arith/Factorial.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 PeanoNat Plus Mult Lt. diff --git a/theories/Arith/Gt.v b/theories/Arith/Gt.v index 2d783f9e2..52ecf131b 100644 --- a/theories/Arith/Gt.v +++ b/theories/Arith/Gt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Theorems about [gt] in [nat]. diff --git a/theories/Arith/Le.v b/theories/Arith/Le.v index d95b05770..69626cc10 100644 --- a/theories/Arith/Le.v +++ b/theories/Arith/Le.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Order on natural numbers. @@ -26,17 +28,17 @@ Local Open Scope nat_scope. (** * [le] is an order on [nat] *) -Notation le_refl := Nat.le_refl (compat "8.4"). -Notation le_trans := Nat.le_trans (compat "8.4"). -Notation le_antisym := Nat.le_antisymm (compat "8.4"). +Notation le_refl := Nat.le_refl (only parsing). +Notation le_trans := Nat.le_trans (only parsing). +Notation le_antisym := Nat.le_antisymm (only parsing). Hint Resolve le_trans: arith. Hint Immediate le_antisym: arith. (** * Properties of [le] w.r.t 0 *) -Notation le_0_n := Nat.le_0_l (compat "8.4"). (* 0 <= n *) -Notation le_Sn_0 := Nat.nle_succ_0 (compat "8.4"). (* ~ S n <= 0 *) +Notation le_0_n := Nat.le_0_l (only parsing). (* 0 <= n *) +Notation le_Sn_0 := Nat.nle_succ_0 (only parsing). (* ~ S n <= 0 *) Lemma le_n_0_eq n : n <= 0 -> 0 = n. Proof. @@ -53,8 +55,8 @@ Proof Peano.le_n_S. Theorem le_S_n : forall n m, S n <= S m -> n <= m. Proof Peano.le_S_n. -Notation le_n_Sn := Nat.le_succ_diag_r (compat "8.4"). (* n <= S n *) -Notation le_Sn_n := Nat.nle_succ_diag_l (compat "8.4"). (* ~ S n <= n *) +Notation le_n_Sn := Nat.le_succ_diag_r (only parsing). (* n <= S n *) +Notation le_Sn_n := Nat.nle_succ_diag_l (only parsing). (* ~ S n <= n *) Theorem le_Sn_le : forall n m, S n <= m -> n <= m. Proof Nat.lt_le_incl. @@ -65,8 +67,8 @@ Hint Immediate le_n_0_eq le_Sn_le le_S_n : arith. (** * Properties of [le] w.r.t predecessor *) -Notation le_pred_n := Nat.le_pred_l (compat "8.4"). (* pred n <= n *) -Notation le_pred := Nat.pred_le_mono (compat "8.4"). (* n<=m -> pred n <= pred m *) +Notation le_pred_n := Nat.le_pred_l (only parsing). (* pred n <= n *) +Notation le_pred := Nat.pred_le_mono (only parsing). (* n<=m -> pred n <= pred m *) Hint Resolve le_pred_n: arith. diff --git a/theories/Arith/Lt.v b/theories/Arith/Lt.v index 035c4e466..0c7515c6f 100644 --- a/theories/Arith/Lt.v +++ b/theories/Arith/Lt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Strict order on natural numbers. @@ -23,7 +25,7 @@ Local Open Scope nat_scope. (** * Irreflexivity *) -Notation lt_irrefl := Nat.lt_irrefl (compat "8.4"). (* ~ x < x *) +Notation lt_irrefl := Nat.lt_irrefl (only parsing). (* ~ x < x *) Hint Resolve lt_irrefl: arith. @@ -62,12 +64,12 @@ Hint Immediate le_not_lt lt_not_le: arith. (** * Asymmetry *) -Notation lt_asym := Nat.lt_asymm (compat "8.4"). (* n<m -> ~m<n *) +Notation lt_asym := Nat.lt_asymm (only parsing). (* n<m -> ~m<n *) (** * Order and 0 *) -Notation lt_0_Sn := Nat.lt_0_succ (compat "8.4"). (* 0 < S n *) -Notation lt_n_0 := Nat.nlt_0_r (compat "8.4"). (* ~ n < 0 *) +Notation lt_0_Sn := Nat.lt_0_succ (only parsing). (* 0 < S n *) +Notation lt_n_0 := Nat.nlt_0_r (only parsing). (* ~ n < 0 *) Theorem neq_0_lt n : 0 <> n -> 0 < n. Proof. @@ -84,8 +86,8 @@ Hint Immediate neq_0_lt lt_0_neq: arith. (** * Order and successor *) -Notation lt_n_Sn := Nat.lt_succ_diag_r (compat "8.4"). (* n < S n *) -Notation lt_S := Nat.lt_lt_succ_r (compat "8.4"). (* n < m -> n < S m *) +Notation lt_n_Sn := Nat.lt_succ_diag_r (only parsing). (* n < S n *) +Notation lt_S := Nat.lt_lt_succ_r (only parsing). (* n < m -> n < S m *) Theorem lt_n_S n m : n < m -> S n < S m. Proof. @@ -107,6 +109,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. @@ -122,28 +129,28 @@ Hint Resolve lt_pred_n_n: arith. (** * Transitivity properties *) -Notation lt_trans := Nat.lt_trans (compat "8.4"). -Notation lt_le_trans := Nat.lt_le_trans (compat "8.4"). -Notation le_lt_trans := Nat.le_lt_trans (compat "8.4"). +Notation lt_trans := Nat.lt_trans (only parsing). +Notation lt_le_trans := Nat.lt_le_trans (only parsing). +Notation le_lt_trans := Nat.le_lt_trans (only parsing). Hint Resolve lt_trans lt_le_trans le_lt_trans: arith. (** * Large = strict or equal *) -Notation le_lt_or_eq_iff := Nat.lt_eq_cases (compat "8.4"). +Notation le_lt_or_eq_iff := Nat.lt_eq_cases (only parsing). Theorem le_lt_or_eq n m : n <= m -> n < m \/ n = m. Proof. apply Nat.lt_eq_cases. Qed. -Notation lt_le_weak := Nat.lt_le_incl (compat "8.4"). +Notation lt_le_weak := Nat.lt_le_incl (only parsing). Hint Immediate lt_le_weak: arith. (** * Dichotomy *) -Notation le_or_lt := Nat.le_gt_cases (compat "8.4"). (* n <= m \/ m < n *) +Notation le_or_lt := Nat.le_gt_cases (only parsing). (* n <= m \/ m < n *) Theorem nat_total_order n m : n <> m -> n < m \/ m < n. Proof. diff --git a/theories/Arith/Max.v b/theories/Arith/Max.v index d8c6f6528..1727fa371 100644 --- a/theories/Arith/Max.v +++ b/theories/Arith/Max.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** THIS FILE IS DEPRECATED. Use [PeanoNat.Nat] instead. *) diff --git a/theories/Arith/Min.v b/theories/Arith/Min.v index 4826013f0..fcf0b14c8 100644 --- a/theories/Arith/Min.v +++ b/theories/Arith/Min.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** THIS FILE IS DEPRECATED. Use [PeanoNat.Nat] instead. *) diff --git a/theories/Arith/Minus.v b/theories/Arith/Minus.v index 950f985d4..3bf6cd952 100644 --- a/theories/Arith/Minus.v +++ b/theories/Arith/Minus.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of subtraction between natural numbers. @@ -46,7 +48,7 @@ Qed. (** * Diagonal *) -Notation minus_diag := Nat.sub_diag (compat "8.4"). (* n - n = 0 *) +Notation minus_diag := Nat.sub_diag (only parsing). (* n - n = 0 *) Lemma minus_diag_reverse n : 0 = n - n. Proof. @@ -87,13 +89,13 @@ Qed. (** * Relation with order *) Notation minus_le_compat_r := - Nat.sub_le_mono_r (compat "8.4"). (* n <= m -> n - p <= m - p. *) + Nat.sub_le_mono_r (only parsing). (* n <= m -> n - p <= m - p. *) Notation minus_le_compat_l := - Nat.sub_le_mono_l (compat "8.4"). (* n <= m -> p - m <= p - n. *) + Nat.sub_le_mono_l (only parsing). (* n <= m -> p - m <= p - n. *) -Notation le_minus := Nat.le_sub_l (compat "8.4"). (* n - m <= n *) -Notation lt_minus := Nat.sub_lt (compat "8.4"). (* m <= n -> 0 < m -> n-m < n *) +Notation le_minus := Nat.le_sub_l (only parsing). (* n - m <= n *) +Notation lt_minus := Nat.sub_lt (only parsing). (* m <= n -> 0 < m -> n-m < n *) Lemma lt_O_minus_lt n m : 0 < n - m -> m < n. Proof. diff --git a/theories/Arith/Mult.v b/theories/Arith/Mult.v index e4084ba47..4f4aa1837 100644 --- a/theories/Arith/Mult.v +++ b/theories/Arith/Mult.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Properties of multiplication. @@ -23,35 +25,35 @@ Local Open Scope nat_scope. (** ** Zero property *) -Notation mult_0_l := Nat.mul_0_l (compat "8.4"). (* 0 * n = 0 *) -Notation mult_0_r := Nat.mul_0_r (compat "8.4"). (* n * 0 = 0 *) +Notation mult_0_l := Nat.mul_0_l (only parsing). (* 0 * n = 0 *) +Notation mult_0_r := Nat.mul_0_r (only parsing). (* n * 0 = 0 *) (** ** 1 is neutral *) -Notation mult_1_l := Nat.mul_1_l (compat "8.4"). (* 1 * n = n *) -Notation mult_1_r := Nat.mul_1_r (compat "8.4"). (* n * 1 = n *) +Notation mult_1_l := Nat.mul_1_l (only parsing). (* 1 * n = n *) +Notation mult_1_r := Nat.mul_1_r (only parsing). (* n * 1 = n *) Hint Resolve mult_1_l mult_1_r: arith. (** ** Commutativity *) -Notation mult_comm := Nat.mul_comm (compat "8.4"). (* n * m = m * n *) +Notation mult_comm := Nat.mul_comm (only parsing). (* n * m = m * n *) Hint Resolve mult_comm: arith. (** ** Distributivity *) Notation mult_plus_distr_r := - Nat.mul_add_distr_r (compat "8.4"). (* (n+m)*p = n*p + m*p *) + Nat.mul_add_distr_r (only parsing). (* (n+m)*p = n*p + m*p *) Notation mult_plus_distr_l := - Nat.mul_add_distr_l (compat "8.4"). (* n*(m+p) = n*m + n*p *) + Nat.mul_add_distr_l (only parsing). (* n*(m+p) = n*m + n*p *) Notation mult_minus_distr_r := - Nat.mul_sub_distr_r (compat "8.4"). (* (n-m)*p = n*p - m*p *) + Nat.mul_sub_distr_r (only parsing). (* (n-m)*p = n*p - m*p *) Notation mult_minus_distr_l := - Nat.mul_sub_distr_l (compat "8.4"). (* n*(m-p) = n*m - n*p *) + Nat.mul_sub_distr_l (only parsing). (* n*(m-p) = n*m - n*p *) Hint Resolve mult_plus_distr_r: arith. Hint Resolve mult_minus_distr_r: arith. @@ -59,7 +61,7 @@ Hint Resolve mult_minus_distr_l: arith. (** ** Associativity *) -Notation mult_assoc := Nat.mul_assoc (compat "8.4"). (* n*(m*p)=n*m*p *) +Notation mult_assoc := Nat.mul_assoc (only parsing). (* n*(m*p)=n*m*p *) Lemma mult_assoc_reverse n m p : n * m * p = n * (m * p). Proof. @@ -83,8 +85,8 @@ Qed. (** ** Multiplication and successor *) -Notation mult_succ_l := Nat.mul_succ_l (compat "8.4"). (* S n * m = n * m + m *) -Notation mult_succ_r := Nat.mul_succ_r (compat "8.4"). (* n * S m = n * m + n *) +Notation mult_succ_l := Nat.mul_succ_l (only parsing). (* S n * m = n * m + m *) +Notation mult_succ_r := Nat.mul_succ_r (only parsing). (* n * S m = n * m + n *) (** * Compatibility with orders *) diff --git a/theories/Arith/PeanoNat.v b/theories/Arith/PeanoNat.v index bde6f1bb4..4e4938a99 100644 --- a/theories/Arith/PeanoNat.v +++ b/theories/Arith/PeanoNat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) @@ -724,6 +726,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 247ea20a8..9a24c804a 100644 --- a/theories/Arith/Peano_dec.v +++ b/theories/Arith/Peano_dec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Decidable PeanoNat. @@ -19,7 +21,7 @@ Proof. - left; exists n; auto. Defined. -Notation eq_nat_dec := Nat.eq_dec (compat "8.4"). +Notation eq_nat_dec := Nat.eq_dec (only parsing). Hint Resolve O_or_S eq_nat_dec: arith. diff --git a/theories/Arith/Plus.v b/theories/Arith/Plus.v index 600e5e518..b8297c2d8 100644 --- a/theories/Arith/Plus.v +++ b/theories/Arith/Plus.v @@ -1,9 +1,11 @@ - (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2010 *) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of addition. @@ -27,12 +29,12 @@ Local Open Scope nat_scope. (** * Neutrality of 0, commutativity, associativity *) -Notation plus_0_l := Nat.add_0_l (compat "8.4"). -Notation plus_0_r := Nat.add_0_r (compat "8.4"). -Notation plus_comm := Nat.add_comm (compat "8.4"). -Notation plus_assoc := Nat.add_assoc (compat "8.4"). +Notation plus_0_l := Nat.add_0_l (only parsing). +Notation plus_0_r := Nat.add_0_r (only parsing). +Notation plus_comm := Nat.add_comm (only parsing). +Notation plus_assoc := Nat.add_assoc (only parsing). -Notation plus_permute := Nat.add_shuffle3 (compat "8.4"). +Notation plus_permute := Nat.add_shuffle3 (only parsing). Definition plus_Snm_nSm : forall n m, S n + m = n + S m := Peano.plus_n_Sm. @@ -138,7 +140,7 @@ Defined. (** * Derived properties *) -Notation plus_permute_2_in_4 := Nat.add_shuffle1 (compat "8.4"). +Notation plus_permute_2_in_4 := Nat.add_shuffle1 (only parsing). (** * Tail-recursive plus *) diff --git a/theories/Arith/Wf_nat.v b/theories/Arith/Wf_nat.v index c9693dbb4..b02288986 100644 --- a/theories/Arith/Wf_nat.v +++ b/theories/Arith/Wf_nat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Well-founded relations and natural numbers *) diff --git a/theories/Bool/Bool.v b/theories/Bool/Bool.v index fe050ca26..edf78ed52 100644 --- a/theories/Bool/Bool.v +++ b/theories/Bool/Bool.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** The type [bool] is defined in the prelude as diff --git a/theories/Bool/BoolEq.v b/theories/Bool/BoolEq.v index aaa774904..106a79e8c 100644 --- a/theories/Bool/BoolEq.v +++ b/theories/Bool/BoolEq.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Cuihtlauac Alvarado - octobre 2000 *) diff --git a/theories/Bool/Bvector.v b/theories/Bool/Bvector.v index d697bd86e..aecdb59db 100644 --- a/theories/Bool/Bvector.v +++ b/theories/Bool/Bvector.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Bit vectors. Contribution by Jean Duprat (ENS Lyon). *) diff --git a/theories/Bool/DecBool.v b/theories/Bool/DecBool.v index 78bca1f5e..f84aed191 100644 --- a/theories/Bool/DecBool.v +++ b/theories/Bool/DecBool.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) Set Implicit Arguments. diff --git a/theories/Bool/IfProp.v b/theories/Bool/IfProp.v index fea952585..6f99ea1da 100644 --- a/theories/Bool/IfProp.v +++ b/theories/Bool/IfProp.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool. diff --git a/theories/Bool/Sumbool.v b/theories/Bool/Sumbool.v index 065c979c9..5333c7411 100644 --- a/theories/Bool/Sumbool.v +++ b/theories/Bool/Sumbool.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Here are collected some results about the type sumbool (see INIT/Specif.v) diff --git a/theories/Bool/Zerob.v b/theories/Bool/Zerob.v index b2fc58f30..2420c0fdc 100644 --- a/theories/Bool/Zerob.v +++ b/theories/Bool/Zerob.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Arith. diff --git a/theories/Classes/CEquivalence.v b/theories/Classes/CEquivalence.v index 6e0f9e014..03e611f54 100644 --- a/theories/Classes/CEquivalence.v +++ b/theories/Classes/CEquivalence.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based setoids. Definitions on [Equivalence]. diff --git a/theories/Classes/CMorphisms.v b/theories/Classes/CMorphisms.v index fae132fc6..09b35ca75 100644 --- a/theories/Classes/CMorphisms.v +++ b/theories/Classes/CMorphisms.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8; coq-prog-args: ("-coqlib" "../.." "-R" ".." "Coq" "-top" "Coq.Classes.CMorphisms") -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based morphism definition and standard, minimal instances diff --git a/theories/Classes/CRelationClasses.v b/theories/Classes/CRelationClasses.v index 02baa9114..bc821532f 100644 --- a/theories/Classes/CRelationClasses.v +++ b/theories/Classes/CRelationClasses.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based relations, tactics and standard instances diff --git a/theories/Classes/DecidableClass.v b/theories/Classes/DecidableClass.v index 34ed6f7a4..bffad2b38 100644 --- a/theories/Classes/DecidableClass.v +++ b/theories/Classes/DecidableClass.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * A typeclass to ease the handling of decidable properties. *) diff --git a/theories/Classes/EquivDec.v b/theories/Classes/EquivDec.v index 1278dac64..e9a9d6aff 100644 --- a/theories/Classes/EquivDec.v +++ b/theories/Classes/EquivDec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Decidable equivalences. diff --git a/theories/Classes/Equivalence.v b/theories/Classes/Equivalence.v index 76c3e768f..5217aedb8 100644 --- a/theories/Classes/Equivalence.v +++ b/theories/Classes/Equivalence.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based setoids. Definitions on [Equivalence]. diff --git a/theories/Classes/Init.v b/theories/Classes/Init.v index bddaa44b5..8a04206bb 100644 --- a/theories/Classes/Init.v +++ b/theories/Classes/Init.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Initialization code for typeclasses, setting up the default tactic diff --git a/theories/Classes/Morphisms.v b/theories/Classes/Morphisms.v index 9859fb52f..1858ba76a 100644 --- a/theories/Classes/Morphisms.v +++ b/theories/Classes/Morphisms.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8; coq-prog-args: ("-coqlib" "../.." "-R" ".." "Coq" "-top" "Coq.Classes.Morphisms") -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based morphism definition and standard, minimal instances diff --git a/theories/Classes/Morphisms_Prop.v b/theories/Classes/Morphisms_Prop.v index 24529ff37..8881fda57 100644 --- a/theories/Classes/Morphisms_Prop.v +++ b/theories/Classes/Morphisms_Prop.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * [Proper] instances for propositional connectives. diff --git a/theories/Classes/Morphisms_Relations.v b/theories/Classes/Morphisms_Relations.v index c12b6b7bd..a3e575010 100644 --- a/theories/Classes/Morphisms_Relations.v +++ b/theories/Classes/Morphisms_Relations.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Morphism instances for relations. diff --git a/theories/Classes/RelationClasses.v b/theories/Classes/RelationClasses.v index 63225853d..2ab3af202 100644 --- a/theories/Classes/RelationClasses.v +++ b/theories/Classes/RelationClasses.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based relations, tactics and standard instances diff --git a/theories/Classes/RelationPairs.v b/theories/Classes/RelationPairs.v index 8d1c49822..7af2b0fc4 100644 --- a/theories/Classes/RelationPairs.v +++ b/theories/Classes/RelationPairs.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Relations over pairs *) diff --git a/theories/Classes/SetoidClass.v b/theories/Classes/SetoidClass.v index 628160711..2673a1191 100644 --- a/theories/Classes/SetoidClass.v +++ b/theories/Classes/SetoidClass.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Typeclass-based setoids, tactics and standard instances. diff --git a/theories/Classes/SetoidDec.v b/theories/Classes/SetoidDec.v index 54e54cd89..f6b240bf2 100644 --- a/theories/Classes/SetoidDec.v +++ b/theories/Classes/SetoidDec.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Decidable setoid equality theory. diff --git a/theories/Classes/SetoidTactics.v b/theories/Classes/SetoidTactics.v index 292e8fc2e..3fab3c5a0 100644 --- a/theories/Classes/SetoidTactics.v +++ b/theories/Classes/SetoidTactics.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Tactics for typeclass-based setoids. diff --git a/theories/Compat/AdmitAxiom.v b/theories/Compat/AdmitAxiom.v index 7747b75ed..f6b751bd5 100644 --- a/theories/Compat/AdmitAxiom.v +++ b/theories/Compat/AdmitAxiom.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Compatibility file for making the admit tactic act similar to Coq v8.4. In diff --git a/theories/Compat/Coq85.v b/theories/Compat/Coq85.v deleted file mode 100644 index 56a9130d1..000000000 --- a/theories/Compat/Coq85.v +++ /dev/null @@ -1,36 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) -(* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(************************************************************************) - -(** Compatibility file for making Coq act similar to Coq v8.5 *) - -(** Any compatibility changes to make future versions of Coq behave like Coq 8.6 - are likely needed to make them behave like Coq 8.5. *) -Require Export Coq.Compat.Coq86. - -(** We use some deprecated options in this file, so we disable the - corresponding warning, to silence the build of this file. *) -Local Set Warnings "-deprecated-option". - -(* In 8.5, "intros [|]", taken e.g. on a goal "A\/B->C", does not - behave as "intros [H|H]" but leave instead hypotheses quantified in - the goal, here producing subgoals A->C and B->C. *) - -Global Unset Bracketing Last Introduction Pattern. -Global Unset Regular Subst Tactic. -Global Unset Structural Injection. -Global Unset Shrink Abstract. -Global Unset Shrink Obligations. -Global Set Refolding Reduction. - -(** The resolution algorithm for type classes has changed. *) -Global Set Typeclasses Legacy Resolution. -Global Set Typeclasses Limit Intros. -Global Unset Typeclasses Filtered Unification. - -(** Allow silently letting unification constraints float after a "." *) -Global Unset Solve Unification Constraints. diff --git a/theories/Compat/Coq86.v b/theories/Compat/Coq86.v index 34061ddd6..666be207e 100644 --- a/theories/Compat/Coq86.v +++ b/theories/Compat/Coq86.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Compatibility file for making Coq act similar to Coq v8.6 *) diff --git a/theories/Compat/Coq87.v b/theories/Compat/Coq87.v index aeef9595d..89556ee75 100644 --- a/theories/Compat/Coq87.v +++ b/theories/Compat/Coq87.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Compatibility file for making Coq act similar to Coq v8.7 *) diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v index 4a790296b..3485b9c68 100644 --- a/theories/FSets/FMapAVL.v +++ b/theories/FSets/FMapAVL.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (* Finite map library. *) diff --git a/theories/FSets/FMapFacts.v b/theories/FSets/FMapFacts.v index 3c5690a72..997059669 100644 --- a/theories/FSets/FMapFacts.v +++ b/theories/FSets/FMapFacts.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite maps library *) @@ -24,7 +26,7 @@ Hint Extern 1 (Equivalence _) => constructor; congruence. Module WFacts_fun (E:DecidableType)(Import M:WSfun E). -Notation option_map := option_map (compat "8.4"). +Notation option_map := option_map (compat "8.6"). Notation eq_dec := E.eq_dec. Definition eqb x y := if eq_dec x y then true else false. @@ -440,7 +442,7 @@ destruct (eq_dec x y); auto. Qed. Lemma map_o : forall m x (f:elt->elt'), - find x (map f m) = option_map f (find x m). + find x (map f m) = Datatypes.option_map f (find x m). Proof. intros. generalize (find_mapsto_iff (map f m) x) (find_mapsto_iff m x) @@ -473,7 +475,7 @@ Qed. Lemma mapi_o : forall m x (f:key->elt->elt'), (forall x y e, E.eq x y -> f x e = f y e) -> - find x (mapi f m) = option_map (f x) (find x m). + find x (mapi f m) = Datatypes.option_map (f x) (find x m). Proof. intros. generalize (find_mapsto_iff (mapi f m) x) (find_mapsto_iff m x) diff --git a/theories/FSets/FMapFullAVL.v b/theories/FSets/FMapFullAVL.v index b8e362f15..345296782 100644 --- a/theories/FSets/FMapFullAVL.v +++ b/theories/FSets/FMapFullAVL.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (* Finite map library. *) diff --git a/theories/FSets/FMapInterface.v b/theories/FSets/FMapInterface.v index 4d89b562d..38a96dc39 100644 --- a/theories/FSets/FMapInterface.v +++ b/theories/FSets/FMapInterface.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite map library *) diff --git a/theories/FSets/FMapList.v b/theories/FSets/FMapList.v index aadef476d..3e98d1197 100644 --- a/theories/FSets/FMapList.v +++ b/theories/FSets/FMapList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite map library *) diff --git a/theories/FSets/FMapPositive.v b/theories/FSets/FMapPositive.v index 3e7664929..0fc68b143 100644 --- a/theories/FSets/FMapPositive.v +++ b/theories/FSets/FMapPositive.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * FMapPositive : an implementation of FMapInterface for [positive] keys. *) diff --git a/theories/FSets/FMapWeakList.v b/theories/FSets/FMapWeakList.v index 812409702..673609650 100644 --- a/theories/FSets/FMapWeakList.v +++ b/theories/FSets/FMapWeakList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite map library *) diff --git a/theories/FSets/FMaps.v b/theories/FSets/FMaps.v index 19b25d95a..ec5076358 100644 --- a/theories/FSets/FMaps.v +++ b/theories/FSets/FMaps.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export OrderedType OrderedTypeEx OrderedTypeAlt. diff --git a/theories/FSets/FSetAVL.v b/theories/FSets/FSetAVL.v index df627a14b..3c9b6b428 100644 --- a/theories/FSets/FSetAVL.v +++ b/theories/FSets/FSetAVL.v @@ -1,11 +1,13 @@ (* -*- coding: utf-8 -*- *) -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * FSetAVL : Implementation of FSetInterface via AVL trees *) diff --git a/theories/FSets/FSetBridge.v b/theories/FSets/FSetBridge.v index 97f140b39..0c4ecb1f3 100644 --- a/theories/FSets/FSetBridge.v +++ b/theories/FSets/FSetBridge.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSetCompat.v b/theories/FSets/FSetCompat.v index 31bc1cc31..736c85dad 100644 --- a/theories/FSets/FSetCompat.v +++ b/theories/FSets/FSetCompat.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Compatibility functors between FSetInterface and MSetInterface. *) diff --git a/theories/FSets/FSetDecide.v b/theories/FSets/FSetDecide.v index 1db6a71e8..83bb07ffb 100644 --- a/theories/FSets/FSetDecide.v +++ b/theories/FSets/FSetDecide.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (**************************************************************) (* FSetDecide.v *) diff --git a/theories/FSets/FSetEqProperties.v b/theories/FSets/FSetEqProperties.v index f2f4cc2cc..56844f477 100644 --- a/theories/FSets/FSetEqProperties.v +++ b/theories/FSets/FSetEqProperties.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSetFacts.v b/theories/FSets/FSetFacts.v index b240ede4d..f4d281e93 100644 --- a/theories/FSets/FSetFacts.v +++ b/theories/FSets/FSetFacts.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSetInterface.v b/theories/FSets/FSetInterface.v index c791f49a6..0926d3ae9 100644 --- a/theories/FSets/FSetInterface.v +++ b/theories/FSets/FSetInterface.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite set library *) diff --git a/theories/FSets/FSetList.v b/theories/FSets/FSetList.v index 9c3ec71ae..2036d360a 100644 --- a/theories/FSets/FSetList.v +++ b/theories/FSets/FSetList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSetPositive.v b/theories/FSets/FSetPositive.v index 507f1cda6..8a93f3816 100644 --- a/theories/FSets/FSetPositive.v +++ b/theories/FSets/FSetPositive.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** Efficient implementation of [FSetInterface.S] for positive keys, inspired from the [FMapPositive] module. diff --git a/theories/FSets/FSetProperties.v b/theories/FSets/FSetProperties.v index 0041bfa1c..c9cfb94ac 100644 --- a/theories/FSets/FSetProperties.v +++ b/theories/FSets/FSetProperties.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSetToFiniteSet.v b/theories/FSets/FSetToFiniteSet.v index 3ac5d9e4a..f8d13ed2b 100644 --- a/theories/FSets/FSetToFiniteSet.v +++ b/theories/FSets/FSetToFiniteSet.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library : conversion to old [Finite_sets] *) diff --git a/theories/FSets/FSetWeakList.v b/theories/FSets/FSetWeakList.v index 9dbea8849..1dacd0568 100644 --- a/theories/FSets/FSetWeakList.v +++ b/theories/FSets/FSetWeakList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/FSets/FSets.v b/theories/FSets/FSets.v index e03fb2236..7e9e7aae7 100644 --- a/theories/FSets/FSets.v +++ b/theories/FSets/FSets.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export OrderedType. Require Export OrderedTypeEx. diff --git a/theories/Init/Datatypes.v b/theories/Init/Datatypes.v index 22e10e2e4..05b741f0a 100644 --- a/theories/Init/Datatypes.v +++ b/theories/Init/Datatypes.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) Set Implicit Arguments. @@ -359,14 +361,14 @@ Definition idProp : IDProp := fun A x => x. (* Compatibility *) -Notation prodT := prod (compat "8.2"). -Notation pairT := pair (compat "8.2"). -Notation prodT_rect := prod_rect (compat "8.2"). -Notation prodT_rec := prod_rec (compat "8.2"). -Notation prodT_ind := prod_ind (compat "8.2"). -Notation fstT := fst (compat "8.2"). -Notation sndT := snd (compat "8.2"). -Notation prodT_uncurry := prod_uncurry (compat "8.2"). -Notation prodT_curry := prod_curry (compat "8.2"). +Notation prodT := prod (only parsing). +Notation pairT := pair (only parsing). +Notation prodT_rect := prod_rect (only parsing). +Notation prodT_rec := prod_rec (only parsing). +Notation prodT_ind := prod_ind (only parsing). +Notation fstT := fst (only parsing). +Notation sndT := snd (only parsing). +Notation prodT_uncurry := prod_uncurry (only parsing). +Notation prodT_curry := prod_curry (only parsing). (* end hide *) diff --git a/theories/Init/Decimal.v b/theories/Init/Decimal.v new file mode 100644 index 000000000..57163b1b0 --- /dev/null +++ b/theories/Init/Decimal.v @@ -0,0 +1,163 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 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..10ca9ecc9 100644 --- a/theories/Init/Logic.v +++ b/theories/Init/Logic.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) Set Implicit Arguments. @@ -267,6 +269,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. @@ -542,14 +551,14 @@ Qed. (* Aliases *) -Notation sym_eq := eq_sym (compat "8.3"). -Notation trans_eq := eq_trans (compat "8.3"). -Notation sym_not_eq := not_eq_sym (compat "8.3"). +Notation sym_eq := eq_sym (only parsing). +Notation trans_eq := eq_trans (only parsing). +Notation sym_not_eq := not_eq_sym (only parsing). -Notation refl_equal := eq_refl (compat "8.3"). -Notation sym_equal := eq_sym (compat "8.3"). -Notation trans_equal := eq_trans (compat "8.3"). -Notation sym_not_equal := not_eq_sym (compat "8.3"). +Notation refl_equal := eq_refl (only parsing). +Notation sym_equal := eq_sym (only parsing). +Notation trans_equal := eq_trans (only parsing). +Notation sym_not_equal := not_eq_sym (only parsing). Hint Immediate eq_sym not_eq_sym: core. diff --git a/theories/Init/Logic_Type.v b/theories/Init/Logic_Type.v index 567d2c15c..6f10a9399 100644 --- a/theories/Init/Logic_Type.v +++ b/theories/Init/Logic_Type.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module defines type constructors for types in [Type] @@ -66,7 +68,7 @@ Defined. Hint Immediate identity_sym not_identity_sym: core. -Notation refl_id := identity_refl (compat "8.3"). -Notation sym_id := identity_sym (compat "8.3"). -Notation trans_id := identity_trans (compat "8.3"). -Notation sym_not_id := not_identity_sym (compat "8.3"). +Notation refl_id := identity_refl (only parsing). +Notation sym_id := identity_sym (only parsing). +Notation trans_id := identity_trans (only parsing). +Notation sym_not_id := not_identity_sym (only parsing). diff --git a/theories/Init/Nat.v b/theories/Init/Nat.v index e942ca562..ad1bc717c 100644 --- a/theories/Init/Nat.v +++ b/theories/Init/Nat.v @@ -1,13 +1,15 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Notations Logic Datatypes. - +Require Decimal. Local Open Scope nat_scope. (**********************************************************************) @@ -134,6 +136,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 70a6c61e7..72073bb4f 100644 --- a/theories/Init/Notations.v +++ b/theories/Init/Notations.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** These are the notations whose level and associativity are imposed by Coq *) @@ -79,6 +81,33 @@ Reserved Notation "{ x & P }" (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/Peano.v b/theories/Init/Peano.v index 571d2f2dd..73c8c5ef4 100644 --- a/theories/Init/Peano.v +++ b/theories/Init/Peano.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** The type [nat] of Peano natural numbers (built from [O] and [S]) @@ -37,7 +39,7 @@ Hint Resolve f_equal_nat: core. (** The predecessor function *) -Notation pred := Nat.pred (compat "8.4"). +Notation pred := Nat.pred (only parsing). Definition f_equal_pred := f_equal pred. @@ -79,7 +81,7 @@ Hint Resolve n_Sn: core. (** Addition *) -Notation plus := Nat.add (compat "8.4"). +Notation plus := Nat.add (only parsing). Infix "+" := Nat.add : nat_scope. Definition f_equal2_plus := f_equal2 plus. @@ -110,12 +112,12 @@ Qed. (** Standard associated names *) -Notation plus_0_r_reverse := plus_n_O (compat "8.2"). -Notation plus_succ_r_reverse := plus_n_Sm (compat "8.2"). +Notation plus_0_r_reverse := plus_n_O (only parsing). +Notation plus_succ_r_reverse := plus_n_Sm (only parsing). (** Multiplication *) -Notation mult := Nat.mul (compat "8.4"). +Notation mult := Nat.mul (only parsing). Infix "*" := Nat.mul : nat_scope. Definition f_equal2_mult := f_equal2 mult. @@ -137,12 +139,12 @@ Hint Resolve mult_n_Sm: core. (** Standard associated names *) -Notation mult_0_r_reverse := mult_n_O (compat "8.2"). -Notation mult_succ_r_reverse := mult_n_Sm (compat "8.2"). +Notation mult_0_r_reverse := mult_n_O (only parsing). +Notation mult_succ_r_reverse := mult_n_Sm (only parsing). (** Truncated subtraction: [m-n] is [0] if [n>=m] *) -Notation minus := Nat.sub (compat "8.4"). +Notation minus := Nat.sub (only parsing). Infix "-" := Nat.sub : nat_scope. (** Definition of the usual orders, the basic properties of [le] and [lt] @@ -219,8 +221,8 @@ Qed. (** Maximum and minimum : definitions and specifications *) -Notation max := Nat.max (compat "8.4"). -Notation min := Nat.min (compat "8.4"). +Notation max := Nat.max (only parsing). +Notation min := Nat.min (only parsing). Lemma max_l n m : m <= n -> Nat.max n m = n. Proof. diff --git a/theories/Init/Prelude.v b/theories/Init/Prelude.v index f0867a034..802f18c0f 100644 --- a/theories/Init/Prelude.v +++ b/theories/Init/Prelude.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Notations. @@ -11,6 +13,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 c42cc00b8..b6afba29a 100644 --- a/theories/Init/Specif.v +++ b/theories/Init/Specif.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Basic specifications : sets that may contain logical information *) @@ -54,6 +56,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. @@ -734,16 +745,16 @@ Hint Resolve exist exist2 existT existT2: core. (* Compatibility *) -Notation sigS := sigT (compat "8.2"). -Notation existS := existT (compat "8.2"). -Notation sigS_rect := sigT_rect (compat "8.2"). -Notation sigS_rec := sigT_rec (compat "8.2"). -Notation sigS_ind := sigT_ind (compat "8.2"). -Notation projS1 := projT1 (compat "8.2"). -Notation projS2 := projT2 (compat "8.2"). - -Notation sigS2 := sigT2 (compat "8.2"). -Notation existS2 := existT2 (compat "8.2"). -Notation sigS2_rect := sigT2_rect (compat "8.2"). -Notation sigS2_rec := sigT2_rec (compat "8.2"). -Notation sigS2_ind := sigT2_ind (compat "8.2"). +Notation sigS := sigT (compat "8.6"). +Notation existS := existT (compat "8.6"). +Notation sigS_rect := sigT_rect (compat "8.6"). +Notation sigS_rec := sigT_rec (compat "8.6"). +Notation sigS_ind := sigT_ind (compat "8.6"). +Notation projS1 := projT1 (compat "8.6"). +Notation projS2 := projT2 (compat "8.6"). + +Notation sigS2 := sigT2 (compat "8.6"). +Notation existS2 := existT2 (compat "8.6"). +Notation sigS2_rect := sigT2_rect (compat "8.6"). +Notation sigS2_rec := sigT2_rec (compat "8.6"). +Notation sigS2_ind := sigT2_ind (compat "8.6"). diff --git a/theories/Init/Tactics.v b/theories/Init/Tactics.v index 47a971ef0..9e6c26b10 100644 --- a/theories/Init/Tactics.v +++ b/theories/Init/Tactics.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Notations. diff --git a/theories/Init/Wf.v b/theories/Init/Wf.v index 690a3f644..c27ffa33f 100644 --- a/theories/Init/Wf.v +++ b/theories/Init/Wf.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * This module proves the validity of diff --git a/theories/Lists/List.v b/theories/Lists/List.v index eae2c52de..ca5f154e9 100644 --- a/theories/Lists/List.v +++ b/theories/Lists/List.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Setoid. @@ -27,7 +29,6 @@ Module ListNotations. Notation "[ ]" := nil (format "[ ]") : list_scope. Notation "[ x ]" := (cons x nil) : list_scope. Notation "[ x ; y ; .. ; z ]" := (cons x (cons y .. (cons z nil) ..)) : list_scope. -Notation "[ x ; .. ; y ]" := (cons x .. (cons y nil) ..) (compat "8.4") : list_scope. End ListNotations. Import ListNotations. diff --git a/theories/Lists/ListDec.v b/theories/Lists/ListDec.v index b03461e0b..e7e2cfc87 100644 --- a/theories/Lists/ListDec.v +++ b/theories/Lists/ListDec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Decidability results about lists *) diff --git a/theories/Lists/ListSet.v b/theories/Lists/ListSet.v index 31970da81..cc7d6f553 100644 --- a/theories/Lists/ListSet.v +++ b/theories/Lists/ListSet.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** A library for finite sets, implemented as lists *) diff --git a/theories/Lists/ListTactics.v b/theories/Lists/ListTactics.v index 0d42b7499..843e38352 100644 --- a/theories/Lists/ListTactics.v +++ b/theories/Lists/ListTactics.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinPos. diff --git a/theories/Lists/SetoidList.v b/theories/Lists/SetoidList.v index c95fb4d5c..0c5fe55b2 100644 --- a/theories/Lists/SetoidList.v +++ b/theories/Lists/SetoidList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export List. Require Export Sorted. diff --git a/theories/Lists/SetoidPermutation.v b/theories/Lists/SetoidPermutation.v index cea3e839f..24b96514f 100644 --- a/theories/Lists/SetoidPermutation.v +++ b/theories/Lists/SetoidPermutation.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Permutation SetoidList. (* Set Universe Polymorphism. *) diff --git a/theories/Lists/StreamMemo.v b/theories/Lists/StreamMemo.v index 5fccd0dec..57f558de5 100644 --- a/theories/Lists/StreamMemo.v +++ b/theories/Lists/StreamMemo.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Eqdep_dec. diff --git a/theories/Lists/Streams.v b/theories/Lists/Streams.v index 64e8dffaa..310c651e8 100644 --- a/theories/Lists/Streams.v +++ b/theories/Lists/Streams.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) Set Implicit Arguments. diff --git a/theories/Logic/Berardi.v b/theories/Logic/Berardi.v index dac43ad52..c6836a1c7 100644 --- a/theories/Logic/Berardi.v +++ b/theories/Logic/Berardi.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file formalizes Berardi's paradox which says that in diff --git a/theories/Logic/ChoiceFacts.v b/theories/Logic/ChoiceFacts.v index 78ec8ff24..238ac7df0 100644 --- a/theories/Logic/ChoiceFacts.v +++ b/theories/Logic/ChoiceFacts.v @@ -1,10 +1,12 @@ -(* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) +(************************************************************************) (************************************************************************) (** Some facts and definitions concerning choice and description in @@ -26,6 +28,8 @@ intentional type theory, Journal of Symbolic Logic 70(2):488-514, 2005. [[Werner97]] Benjamin Werner, Sets in Types, Types in Sets, TACS, 1997. *) +Require Import RelationClasses Logic. + Set Implicit Arguments. Local Unset Intuition Negation Unfolding. @@ -123,8 +127,6 @@ Definition DependentFunctionalRelReification_on (A:Type) (B:A -> Type) := formulation of choice); Note also a typo in its intended formulation in [[Werner97]]. *) -Require Import RelationClasses Logic. - Definition RepresentativeFunctionalChoice_on := forall R:A->A->Prop, (Equivalence R) -> @@ -1308,11 +1310,11 @@ Qed. (**********************************************************************) (** * Compatibility notations *) Notation description_rel_choice_imp_funct_choice := - functional_rel_reification_and_rel_choice_imp_fun_choice (compat "8.6"). + functional_rel_reification_and_rel_choice_imp_fun_choice (only parsing). -Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (compat "8.6"). +Notation funct_choice_imp_rel_choice := fun_choice_imp_rel_choice (only parsing). Notation FunChoice_Equiv_RelChoice_and_ParamDefinDescr := - fun_choice_iff_rel_choice_and_functional_rel_reification (compat "8.6"). + fun_choice_iff_rel_choice_and_functional_rel_reification (only parsing). -Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (compat "8.6"). +Notation funct_choice_imp_description := fun_choice_imp_functional_rel_reification (only parsing). diff --git a/theories/Logic/Classical.v b/theories/Logic/Classical.v index 044ca2122..72f53a46e 100644 --- a/theories/Logic/Classical.v +++ b/theories/Logic/Classical.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* File created for Coq V5.10.14b, Oct 1995 *) diff --git a/theories/Logic/ClassicalChoice.v b/theories/Logic/ClassicalChoice.v index 27cb59a8f..016fa72f9 100644 --- a/theories/Logic/ClassicalChoice.v +++ b/theories/Logic/ClassicalChoice.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides classical logic and functional choice; this diff --git a/theories/Logic/ClassicalDescription.v b/theories/Logic/ClassicalDescription.v index 304ae59f8..6867c76e2 100644 --- a/theories/Logic/ClassicalDescription.v +++ b/theories/Logic/ClassicalDescription.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides classical logic and definite description, which is diff --git a/theories/Logic/ClassicalEpsilon.v b/theories/Logic/ClassicalEpsilon.v index e660e9566..77af90481 100644 --- a/theories/Logic/ClassicalEpsilon.v +++ b/theories/Logic/ClassicalEpsilon.v @@ -1,10 +1,12 @@ -(* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) +(************************************************************************) (************************************************************************) (** This file provides classical logic and indefinite description under diff --git a/theories/Logic/ClassicalFacts.v b/theories/Logic/ClassicalFacts.v index c90a97612..b06384e99 100644 --- a/theories/Logic/ClassicalFacts.v +++ b/theories/Logic/ClassicalFacts.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Some facts and definitions about classical logic diff --git a/theories/Logic/ClassicalUniqueChoice.v b/theories/Logic/ClassicalUniqueChoice.v index 1309f91a3..841bd1bed 100644 --- a/theories/Logic/ClassicalUniqueChoice.v +++ b/theories/Logic/ClassicalUniqueChoice.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides classical logic and unique choice; this is diff --git a/theories/Logic/Classical_Pred_Type.v b/theories/Logic/Classical_Pred_Type.v index 06c298466..18820d3ba 100644 --- a/theories/Logic/Classical_Pred_Type.v +++ b/theories/Logic/Classical_Pred_Type.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* This file is a renaming for V5.10.14b, Oct 1995, of file Classical.v diff --git a/theories/Logic/Classical_Prop.v b/theories/Logic/Classical_Prop.v index a5ae07b64..9f5a29937 100644 --- a/theories/Logic/Classical_Prop.v +++ b/theories/Logic/Classical_Prop.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* File created for Coq V5.10.14b, Oct 1995 *) diff --git a/theories/Logic/ConstructiveEpsilon.v b/theories/Logic/ConstructiveEpsilon.v index 71ce2a40a..6e3da423e 100644 --- a/theories/Logic/ConstructiveEpsilon.v +++ b/theories/Logic/ConstructiveEpsilon.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*i $Id: ConstructiveEpsilon.v 12112 2009-04-28 15:47:34Z herbelin $ i*) diff --git a/theories/Logic/Decidable.v b/theories/Logic/Decidable.v index acb34771f..35920d913 100644 --- a/theories/Logic/Decidable.v +++ b/theories/Logic/Decidable.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of decidable propositions *) diff --git a/theories/Logic/Description.v b/theories/Logic/Description.v index 4ab52ae46..5c4f1960f 100644 --- a/theories/Logic/Description.v +++ b/theories/Logic/Description.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides a constructive form of definite description; it diff --git a/theories/Logic/Diaconescu.v b/theories/Logic/Diaconescu.v index c124e7141..3317766c9 100644 --- a/theories/Logic/Diaconescu.v +++ b/theories/Logic/Diaconescu.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Diaconescu showed that the Axiom of Choice entails Excluded-Middle @@ -40,6 +42,7 @@ [[Carlström04]] Jesper Carlström, EM + Ext_ + AC_int is equivalent to AC_ext, Mathematical Logic Quaterly, vol 50(3), pp 236-240, 2004. *) +Require ClassicalFacts ChoiceFacts. (**********************************************************************) (** * Pred. Ext. + Rel. Axiom of Choice -> Excluded-Middle *) @@ -54,7 +57,7 @@ Definition PredicateExtensionality := (** From predicate extensionality we get propositional extensionality hence proof-irrelevance *) -Require Import ClassicalFacts. +Import ClassicalFacts. Variable pred_extensionality : PredicateExtensionality. @@ -76,7 +79,7 @@ Qed. (** From proof-irrelevance and relational choice, we get guarded relational choice *) -Require Import ChoiceFacts. +Import ChoiceFacts. Variable rel_choice : RelationalChoice. @@ -89,7 +92,7 @@ Qed. (** The form of choice we need: there is a functional relation which chooses an element in any non empty subset of bool *) -Require Import Bool. +Import Bool. Lemma AC_bool_subset_to_bool : exists R : (bool -> Prop) -> bool -> Prop, @@ -161,6 +164,8 @@ End PredExt_RelChoice_imp_EM. Section ProofIrrel_RelChoice_imp_EqEM. +Import ChoiceFacts. + Variable rel_choice : RelationalChoice. Variable proof_irrelevance : forall P:Prop , forall x y:P, x=y. diff --git a/theories/Logic/Epsilon.v b/theories/Logic/Epsilon.v index 96f677e43..d8c527c61 100644 --- a/theories/Logic/Epsilon.v +++ b/theories/Logic/Epsilon.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides indefinite description under the form of diff --git a/theories/Logic/Eqdep.v b/theories/Logic/Eqdep.v index ac003bf1a..35bc42259 100644 --- a/theories/Logic/Eqdep.v +++ b/theories/Logic/Eqdep.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* File Eqdep.v created by Christine Paulin-Mohring in Coq V5.6, May 1992 *) diff --git a/theories/Logic/EqdepFacts.v b/theories/Logic/EqdepFacts.v index c9dca432a..d938b315f 100644 --- a/theories/Logic/EqdepFacts.v +++ b/theories/Logic/EqdepFacts.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* File Eqdep.v created by Christine Paulin-Mohring in Coq V5.6, May 1992 *) @@ -123,7 +125,7 @@ Proof. apply eq_dep_intro. Qed. -Notation eq_sigS_eq_dep := eq_sigT_eq_dep (compat "8.2"). (* Compatibility *) +Notation eq_sigS_eq_dep := eq_sigT_eq_dep (compat "8.6"). (* Compatibility *) Lemma eq_dep_eq_sigT : forall (U:Type) (P:U -> Type) (p q:U) (x:P p) (y:P q), diff --git a/theories/Logic/Eqdep_dec.v b/theories/Logic/Eqdep_dec.v index beb4d0b3d..0560d9ed4 100644 --- a/theories/Logic/Eqdep_dec.v +++ b/theories/Logic/Eqdep_dec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Created by Bruno Barras, Jan 1998 *) diff --git a/theories/Logic/ExtensionalFunctionRepresentative.v b/theories/Logic/ExtensionalFunctionRepresentative.v index becc9f147..0aac56bbc 100644 --- a/theories/Logic/ExtensionalFunctionRepresentative.v +++ b/theories/Logic/ExtensionalFunctionRepresentative.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module states a limited form axiom of functional diff --git a/theories/Logic/ExtensionalityFacts.v b/theories/Logic/ExtensionalityFacts.v index 86539544b..02c8998a8 100644 --- a/theories/Logic/ExtensionalityFacts.v +++ b/theories/Logic/ExtensionalityFacts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Some facts and definitions about extensionality diff --git a/theories/Logic/FinFun.v b/theories/Logic/FinFun.v index 8bdbfe85b..89f5a82a8 100644 --- a/theories/Logic/FinFun.v +++ b/theories/Logic/FinFun.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Functions on finite domains *) diff --git a/theories/Logic/FunctionalExtensionality.v b/theories/Logic/FunctionalExtensionality.v index 82b04d132..95e1af2ea 100644 --- a/theories/Logic/FunctionalExtensionality.v +++ b/theories/Logic/FunctionalExtensionality.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module states the axiom of (dependent) functional extensionality and (dependent) eta-expansion. diff --git a/theories/Logic/Hurkens.v b/theories/Logic/Hurkens.v index 4590d609d..6c4a8533a 100644 --- a/theories/Logic/Hurkens.v +++ b/theories/Logic/Hurkens.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Hurkens.v *) (************************************************************************) diff --git a/theories/Logic/IndefiniteDescription.v b/theories/Logic/IndefiniteDescription.v index 10ef72ade..86e81529d 100644 --- a/theories/Logic/IndefiniteDescription.v +++ b/theories/Logic/IndefiniteDescription.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file provides a constructive form of indefinite description that diff --git a/theories/Logic/JMeq.v b/theories/Logic/JMeq.v index aa6193bc7..9c56b60aa 100644 --- a/theories/Logic/JMeq.v +++ b/theories/Logic/JMeq.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** John Major's Equality as proposed by Conor McBride diff --git a/theories/Logic/ProofIrrelevance.v b/theories/Logic/ProofIrrelevance.v index d11e7c872..134bf649d 100644 --- a/theories/Logic/ProofIrrelevance.v +++ b/theories/Logic/ProofIrrelevance.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file axiomatizes proof-irrelevance and derives some consequences *) diff --git a/theories/Logic/ProofIrrelevanceFacts.v b/theories/Logic/ProofIrrelevanceFacts.v index f359a2109..10d9dbcda 100644 --- a/theories/Logic/ProofIrrelevanceFacts.v +++ b/theories/Logic/ProofIrrelevanceFacts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This defines the functor that build consequences of proof-irrelevance *) diff --git a/theories/Logic/PropExtensionality.v b/theories/Logic/PropExtensionality.v index b8bc9a52c..80dd4e850 100644 --- a/theories/Logic/PropExtensionality.v +++ b/theories/Logic/PropExtensionality.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module states propositional extensionality and draws diff --git a/theories/Logic/PropExtensionalityFacts.v b/theories/Logic/PropExtensionalityFacts.v index 8a1cf9708..2b3035173 100644 --- a/theories/Logic/PropExtensionalityFacts.v +++ b/theories/Logic/PropExtensionalityFacts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Some facts and definitions about propositional and predicate extensionality diff --git a/theories/Logic/PropFacts.v b/theories/Logic/PropFacts.v index 309539e5c..067870356 100644 --- a/theories/Logic/PropFacts.v +++ b/theories/Logic/PropFacts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Basic facts about Prop as a type *) diff --git a/theories/Logic/RelationalChoice.v b/theories/Logic/RelationalChoice.v index 668568b03..994f07856 100644 --- a/theories/Logic/RelationalChoice.v +++ b/theories/Logic/RelationalChoice.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file axiomatizes the relational form of the axiom of choice *) diff --git a/theories/Logic/SetIsType.v b/theories/Logic/SetIsType.v index 19b6f62dd..afa85514e 100644 --- a/theories/Logic/SetIsType.v +++ b/theories/Logic/SetIsType.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * The Set universe seen as a synonym for Type *) diff --git a/theories/Logic/SetoidChoice.v b/theories/Logic/SetoidChoice.v index 62979330d..21bf73356 100644 --- a/theories/Logic/SetoidChoice.v +++ b/theories/Logic/SetoidChoice.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module states the functional form of the axiom of choice over diff --git a/theories/Logic/WKL.v b/theories/Logic/WKL.v index 7f98d8e91..579800b80 100644 --- a/theories/Logic/WKL.v +++ b/theories/Logic/WKL.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** A constructive proof of a version of Weak König's Lemma over a diff --git a/theories/Logic/WeakFan.v b/theories/Logic/WeakFan.v index 0068f72eb..c9822f47d 100644 --- a/theories/Logic/WeakFan.v +++ b/theories/Logic/WeakFan.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** A constructive proof of a non-standard version of the weak Fan Theorem diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v index b30cb6b56..b966f217a 100644 --- a/theories/MSets/MSetAVL.v +++ b/theories/MSets/MSetAVL.v @@ -1,11 +1,13 @@ (* -*- coding: utf-8 -*- *) -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * MSetAVL : Implementation of MSetInterface via AVL trees *) diff --git a/theories/MSets/MSetDecide.v b/theories/MSets/MSetDecide.v index 9c622fd78..f228cbb3b 100644 --- a/theories/MSets/MSetDecide.v +++ b/theories/MSets/MSetDecide.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (**************************************************************) (* MSetDecide.v *) diff --git a/theories/MSets/MSetEqProperties.v b/theories/MSets/MSetEqProperties.v index ae20edc87..1ee098cb0 100644 --- a/theories/MSets/MSetEqProperties.v +++ b/theories/MSets/MSetEqProperties.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/MSets/MSetFacts.v b/theories/MSets/MSetFacts.v index 4e17618f7..d57a7741e 100644 --- a/theories/MSets/MSetFacts.v +++ b/theories/MSets/MSetFacts.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/MSets/MSetGenTree.v b/theories/MSets/MSetGenTree.v index 9fb8e499b..9d2a73ed0 100644 --- a/theories/MSets/MSetGenTree.v +++ b/theories/MSets/MSetGenTree.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * MSetGenTree : sets via generic trees diff --git a/theories/MSets/MSetInterface.v b/theories/MSets/MSetInterface.v index 74a7f6df8..f0e757157 100644 --- a/theories/MSets/MSetInterface.v +++ b/theories/MSets/MSetInterface.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite set library *) diff --git a/theories/MSets/MSetList.v b/theories/MSets/MSetList.v index 05c20eb8f..35fe4cee4 100644 --- a/theories/MSets/MSetList.v +++ b/theories/MSets/MSetList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/MSets/MSetPositive.v b/theories/MSets/MSetPositive.v index be95a0379..a726eebd3 100644 --- a/theories/MSets/MSetPositive.v +++ b/theories/MSets/MSetPositive.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** Efficient implementation of [MSetInterface.S] for positive keys, inspired from the [FMapPositive] module. diff --git a/theories/MSets/MSetProperties.v b/theories/MSets/MSetProperties.v index 396067b57..3c7dea736 100644 --- a/theories/MSets/MSetProperties.v +++ b/theories/MSets/MSetProperties.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/MSets/MSetRBT.v b/theories/MSets/MSetRBT.v index 83a2343dd..eab01a55b 100644 --- a/theories/MSets/MSetRBT.v +++ b/theories/MSets/MSetRBT.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * MSetRBT : Implementation of MSetInterface via Red-Black trees *) diff --git a/theories/MSets/MSetToFiniteSet.v b/theories/MSets/MSetToFiniteSet.v index e8087bc57..f456c407b 100644 --- a/theories/MSets/MSetToFiniteSet.v +++ b/theories/MSets/MSetToFiniteSet.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library : conversion to old [Finite_sets] *) diff --git a/theories/MSets/MSetWeakList.v b/theories/MSets/MSetWeakList.v index 2ac57a932..8df1ff1cd 100644 --- a/theories/MSets/MSetWeakList.v +++ b/theories/MSets/MSetWeakList.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * Finite sets library *) diff --git a/theories/MSets/MSets.v b/theories/MSets/MSets.v index 1ee485cc1..0c9bc20cc 100644 --- a/theories/MSets/MSets.v +++ b/theories/MSets/MSets.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export Orders. Require Export OrdersEx. diff --git a/theories/NArith/BinNat.v b/theories/NArith/BinNat.v index 75a8bfdb3..5d3ec5abc 100644 --- a/theories/NArith/BinNat.v +++ b/theories/NArith/BinNat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export BinNums. @@ -956,95 +958,94 @@ Notation "( p | q )" := (N.divide p q) (at level 0) : N_scope. (** Compatibility notations *) -(*Notation N := N (compat "8.3").*) (*hidden by module N above *) Notation N_rect := N_rect (only parsing). Notation N_rec := N_rec (only parsing). Notation N_ind := N_ind (only parsing). Notation N0 := N0 (only parsing). Notation Npos := N.pos (only parsing). -Notation Ndiscr := N.discr (compat "8.3"). -Notation Ndouble_plus_one := N.succ_double (compat "8.3"). -Notation Ndouble := N.double (compat "8.3"). -Notation Nsucc := N.succ (compat "8.3"). -Notation Npred := N.pred (compat "8.3"). -Notation Nsucc_pos := N.succ_pos (compat "8.3"). -Notation Ppred_N := Pos.pred_N (compat "8.3"). -Notation Nplus := N.add (compat "8.3"). -Notation Nminus := N.sub (compat "8.3"). -Notation Nmult := N.mul (compat "8.3"). -Notation Neqb := N.eqb (compat "8.3"). -Notation Ncompare := N.compare (compat "8.3"). -Notation Nlt := N.lt (compat "8.3"). -Notation Ngt := N.gt (compat "8.3"). -Notation Nle := N.le (compat "8.3"). -Notation Nge := N.ge (compat "8.3"). -Notation Nmin := N.min (compat "8.3"). -Notation Nmax := N.max (compat "8.3"). -Notation Ndiv2 := N.div2 (compat "8.3"). -Notation Neven := N.even (compat "8.3"). -Notation Nodd := N.odd (compat "8.3"). -Notation Npow := N.pow (compat "8.3"). -Notation Nlog2 := N.log2 (compat "8.3"). - -Notation nat_of_N := N.to_nat (compat "8.3"). -Notation N_of_nat := N.of_nat (compat "8.3"). -Notation N_eq_dec := N.eq_dec (compat "8.3"). -Notation Nrect := N.peano_rect (compat "8.3"). -Notation Nrect_base := N.peano_rect_base (compat "8.3"). -Notation Nrect_step := N.peano_rect_succ (compat "8.3"). -Notation Nind := N.peano_ind (compat "8.3"). -Notation Nrec := N.peano_rec (compat "8.3"). -Notation Nrec_base := N.peano_rec_base (compat "8.3"). -Notation Nrec_succ := N.peano_rec_succ (compat "8.3"). - -Notation Npred_succ := N.pred_succ (compat "8.3"). -Notation Npred_minus := N.pred_sub (compat "8.3"). -Notation Nsucc_pred := N.succ_pred (compat "8.3"). -Notation Ppred_N_spec := N.pos_pred_spec (compat "8.3"). -Notation Nsucc_pos_spec := N.succ_pos_spec (compat "8.3"). -Notation Ppred_Nsucc := N.pos_pred_succ (compat "8.3"). -Notation Nplus_0_l := N.add_0_l (compat "8.3"). -Notation Nplus_0_r := N.add_0_r (compat "8.3"). -Notation Nplus_comm := N.add_comm (compat "8.3"). -Notation Nplus_assoc := N.add_assoc (compat "8.3"). -Notation Nplus_succ := N.add_succ_l (compat "8.3"). -Notation Nsucc_0 := N.succ_0_discr (compat "8.3"). -Notation Nsucc_inj := N.succ_inj (compat "8.3"). -Notation Nminus_N0_Nle := N.sub_0_le (compat "8.3"). -Notation Nminus_0_r := N.sub_0_r (compat "8.3"). -Notation Nminus_succ_r:= N.sub_succ_r (compat "8.3"). -Notation Nmult_0_l := N.mul_0_l (compat "8.3"). -Notation Nmult_1_l := N.mul_1_l (compat "8.3"). -Notation Nmult_1_r := N.mul_1_r (compat "8.3"). -Notation Nmult_comm := N.mul_comm (compat "8.3"). -Notation Nmult_assoc := N.mul_assoc (compat "8.3"). -Notation Nmult_plus_distr_r := N.mul_add_distr_r (compat "8.3"). -Notation Neqb_eq := N.eqb_eq (compat "8.3"). -Notation Nle_0 := N.le_0_l (compat "8.3"). -Notation Ncompare_refl := N.compare_refl (compat "8.3"). -Notation Ncompare_Eq_eq := N.compare_eq (compat "8.3"). -Notation Ncompare_eq_correct := N.compare_eq_iff (compat "8.3"). -Notation Nlt_irrefl := N.lt_irrefl (compat "8.3"). -Notation Nlt_trans := N.lt_trans (compat "8.3"). -Notation Nle_lteq := N.lt_eq_cases (compat "8.3"). -Notation Nlt_succ_r := N.lt_succ_r (compat "8.3"). -Notation Nle_trans := N.le_trans (compat "8.3"). -Notation Nle_succ_l := N.le_succ_l (compat "8.3"). -Notation Ncompare_spec := N.compare_spec (compat "8.3"). -Notation Ncompare_0 := N.compare_0_r (compat "8.3"). -Notation Ndouble_div2 := N.div2_double (compat "8.3"). -Notation Ndouble_plus_one_div2 := N.div2_succ_double (compat "8.3"). -Notation Ndouble_inj := N.double_inj (compat "8.3"). -Notation Ndouble_plus_one_inj := N.succ_double_inj (compat "8.3"). -Notation Npow_0_r := N.pow_0_r (compat "8.3"). -Notation Npow_succ_r := N.pow_succ_r (compat "8.3"). -Notation Nlog2_spec := N.log2_spec (compat "8.3"). -Notation Nlog2_nonpos := N.log2_nonpos (compat "8.3"). -Notation Neven_spec := N.even_spec (compat "8.3"). -Notation Nodd_spec := N.odd_spec (compat "8.3"). -Notation Nlt_not_eq := N.lt_neq (compat "8.3"). -Notation Ngt_Nlt := N.gt_lt (compat "8.3"). +Notation Ndiscr := N.discr (compat "8.6"). +Notation Ndouble_plus_one := N.succ_double (only parsing). +Notation Ndouble := N.double (compat "8.6"). +Notation Nsucc := N.succ (compat "8.6"). +Notation Npred := N.pred (compat "8.6"). +Notation Nsucc_pos := N.succ_pos (compat "8.6"). +Notation Ppred_N := Pos.pred_N (compat "8.6"). +Notation Nplus := N.add (only parsing). +Notation Nminus := N.sub (only parsing). +Notation Nmult := N.mul (only parsing). +Notation Neqb := N.eqb (compat "8.6"). +Notation Ncompare := N.compare (compat "8.6"). +Notation Nlt := N.lt (compat "8.6"). +Notation Ngt := N.gt (compat "8.6"). +Notation Nle := N.le (compat "8.6"). +Notation Nge := N.ge (compat "8.6"). +Notation Nmin := N.min (compat "8.6"). +Notation Nmax := N.max (compat "8.6"). +Notation Ndiv2 := N.div2 (compat "8.6"). +Notation Neven := N.even (compat "8.6"). +Notation Nodd := N.odd (compat "8.6"). +Notation Npow := N.pow (compat "8.6"). +Notation Nlog2 := N.log2 (compat "8.6"). + +Notation nat_of_N := N.to_nat (only parsing). +Notation N_of_nat := N.of_nat (only parsing). +Notation N_eq_dec := N.eq_dec (compat "8.6"). +Notation Nrect := N.peano_rect (only parsing). +Notation Nrect_base := N.peano_rect_base (only parsing). +Notation Nrect_step := N.peano_rect_succ (only parsing). +Notation Nind := N.peano_ind (only parsing). +Notation Nrec := N.peano_rec (only parsing). +Notation Nrec_base := N.peano_rec_base (only parsing). +Notation Nrec_succ := N.peano_rec_succ (only parsing). + +Notation Npred_succ := N.pred_succ (compat "8.6"). +Notation Npred_minus := N.pred_sub (only parsing). +Notation Nsucc_pred := N.succ_pred (compat "8.6"). +Notation Ppred_N_spec := N.pos_pred_spec (only parsing). +Notation Nsucc_pos_spec := N.succ_pos_spec (compat "8.6"). +Notation Ppred_Nsucc := N.pos_pred_succ (only parsing). +Notation Nplus_0_l := N.add_0_l (only parsing). +Notation Nplus_0_r := N.add_0_r (only parsing). +Notation Nplus_comm := N.add_comm (only parsing). +Notation Nplus_assoc := N.add_assoc (only parsing). +Notation Nplus_succ := N.add_succ_l (only parsing). +Notation Nsucc_0 := N.succ_0_discr (only parsing). +Notation Nsucc_inj := N.succ_inj (compat "8.6"). +Notation Nminus_N0_Nle := N.sub_0_le (only parsing). +Notation Nminus_0_r := N.sub_0_r (only parsing). +Notation Nminus_succ_r:= N.sub_succ_r (only parsing). +Notation Nmult_0_l := N.mul_0_l (only parsing). +Notation Nmult_1_l := N.mul_1_l (only parsing). +Notation Nmult_1_r := N.mul_1_r (only parsing). +Notation Nmult_comm := N.mul_comm (only parsing). +Notation Nmult_assoc := N.mul_assoc (only parsing). +Notation Nmult_plus_distr_r := N.mul_add_distr_r (only parsing). +Notation Neqb_eq := N.eqb_eq (compat "8.6"). +Notation Nle_0 := N.le_0_l (only parsing). +Notation Ncompare_refl := N.compare_refl (compat "8.6"). +Notation Ncompare_Eq_eq := N.compare_eq (only parsing). +Notation Ncompare_eq_correct := N.compare_eq_iff (only parsing). +Notation Nlt_irrefl := N.lt_irrefl (compat "8.6"). +Notation Nlt_trans := N.lt_trans (compat "8.6"). +Notation Nle_lteq := N.lt_eq_cases (only parsing). +Notation Nlt_succ_r := N.lt_succ_r (compat "8.6"). +Notation Nle_trans := N.le_trans (compat "8.6"). +Notation Nle_succ_l := N.le_succ_l (compat "8.6"). +Notation Ncompare_spec := N.compare_spec (compat "8.6"). +Notation Ncompare_0 := N.compare_0_r (only parsing). +Notation Ndouble_div2 := N.div2_double (only parsing). +Notation Ndouble_plus_one_div2 := N.div2_succ_double (only parsing). +Notation Ndouble_inj := N.double_inj (compat "8.6"). +Notation Ndouble_plus_one_inj := N.succ_double_inj (only parsing). +Notation Npow_0_r := N.pow_0_r (compat "8.6"). +Notation Npow_succ_r := N.pow_succ_r (compat "8.6"). +Notation Nlog2_spec := N.log2_spec (compat "8.6"). +Notation Nlog2_nonpos := N.log2_nonpos (compat "8.6"). +Notation Neven_spec := N.even_spec (compat "8.6"). +Notation Nodd_spec := N.odd_spec (compat "8.6"). +Notation Nlt_not_eq := N.lt_neq (only parsing). +Notation Ngt_Nlt := N.gt_lt (only parsing). (** More complex compatibility facts, expressed as lemmas (to preserve scopes for instance) *) diff --git a/theories/NArith/BinNatDef.v b/theories/NArith/BinNatDef.v index 6771e57ad..5de75537c 100644 --- a/theories/NArith/BinNatDef.v +++ b/theories/NArith/BinNatDef.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export BinNums. @@ -378,4 +380,22 @@ Definition iter (n:N) {A} (f:A->A) (x:A) : A := | pos p => Pos.iter f x p end. +(** 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/NArith/NArith.v b/theories/NArith/NArith.v index 64eea4419..f3007970b 100644 --- a/theories/NArith/NArith.v +++ b/theories/NArith/NArith.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Library for binary natural numbers *) diff --git a/theories/NArith/Ndec.v b/theories/NArith/Ndec.v index 892bbe7cd..67c30f225 100644 --- a/theories/NArith/Ndec.v +++ b/theories/NArith/Ndec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool. @@ -20,11 +22,11 @@ Local Open Scope N_scope. (** Obsolete results about boolean comparisons over [N], kept for compatibility with IntMap and SMC. *) -Notation Peqb := Pos.eqb (compat "8.3"). -Notation Neqb := N.eqb (compat "8.3"). -Notation Peqb_correct := Pos.eqb_refl (compat "8.3"). -Notation Neqb_correct := N.eqb_refl (compat "8.3"). -Notation Neqb_comm := N.eqb_sym (compat "8.3"). +Notation Peqb := Pos.eqb (compat "8.6"). +Notation Neqb := N.eqb (compat "8.6"). +Notation Peqb_correct := Pos.eqb_refl (only parsing). +Notation Neqb_correct := N.eqb_refl (only parsing). +Notation Neqb_comm := N.eqb_sym (only parsing). Lemma Peqb_complete p p' : Pos.eqb p p' = true -> p = p'. Proof. now apply Pos.eqb_eq. Qed. @@ -274,7 +276,7 @@ Qed. (* Old results about [N.min] *) -Notation Nmin_choice := N.min_dec (compat "8.3"). +Notation Nmin_choice := N.min_dec (only parsing). Lemma Nmin_le_1 a b : Nleb (N.min a b) a = true. Proof. rewrite Nleb_Nle. apply N.le_min_l. Qed. diff --git a/theories/NArith/Ndigits.v b/theories/NArith/Ndigits.v index 9aadf985d..3ccaa7211 100644 --- a/theories/NArith/Ndigits.v +++ b/theories/NArith/Ndigits.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool Morphisms Setoid Bvector BinPos BinNat PeanoNat Pnat Nnat. @@ -14,17 +16,17 @@ Local Open Scope N_scope. (** Compatibility names for some bitwise operations *) -Notation Pxor := Pos.lxor (compat "8.3"). -Notation Nxor := N.lxor (compat "8.3"). -Notation Pbit := Pos.testbit_nat (compat "8.3"). -Notation Nbit := N.testbit_nat (compat "8.3"). +Notation Pxor := Pos.lxor (only parsing). +Notation Nxor := N.lxor (only parsing). +Notation Pbit := Pos.testbit_nat (only parsing). +Notation Nbit := N.testbit_nat (only parsing). -Notation Nxor_eq := N.lxor_eq (compat "8.3"). -Notation Nxor_comm := N.lxor_comm (compat "8.3"). -Notation Nxor_assoc := N.lxor_assoc (compat "8.3"). -Notation Nxor_neutral_left := N.lxor_0_l (compat "8.3"). -Notation Nxor_neutral_right := N.lxor_0_r (compat "8.3"). -Notation Nxor_nilpotent := N.lxor_nilpotent (compat "8.3"). +Notation Nxor_eq := N.lxor_eq (only parsing). +Notation Nxor_comm := N.lxor_comm (only parsing). +Notation Nxor_assoc := N.lxor_assoc (only parsing). +Notation Nxor_neutral_left := N.lxor_0_l (only parsing). +Notation Nxor_neutral_right := N.lxor_0_r (only parsing). +Notation Nxor_nilpotent := N.lxor_nilpotent (only parsing). (** Equivalence of bit-testing functions, either with index in [N] or in [nat]. *) @@ -249,7 +251,7 @@ Local Close Scope N_scope. (** Checking whether a number is odd, i.e. if its lower bit is set. *) -Notation Nbit0 := N.odd (compat "8.3"). +Notation Nbit0 := N.odd (only parsing). Definition Nodd (n:N) := N.odd n = true. Definition Neven (n:N) := N.odd n = false. @@ -498,7 +500,7 @@ Qed. (** Number of digits in a number *) -Notation Nsize := N.size_nat (compat "8.3"). +Notation Nsize := N.size_nat (only parsing). (** conversions between N and bit vectors. *) diff --git a/theories/NArith/Ndist.v b/theories/NArith/Ndist.v index d81c119d3..d9a58c057 100644 --- a/theories/NArith/Ndist.v +++ b/theories/NArith/Ndist.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Arith. Require Import Min. diff --git a/theories/NArith/Ndiv_def.v b/theories/NArith/Ndiv_def.v index 974e93994..7c9fd8695 100644 --- a/theories/NArith/Ndiv_def.v +++ b/theories/NArith/Ndiv_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinNat. @@ -22,10 +24,10 @@ Lemma Pdiv_eucl_remainder a b : snd (Pdiv_eucl a b) < Npos b. Proof. now apply (N.pos_div_eucl_remainder a (Npos b)). Qed. -Notation Ndiv_eucl := N.div_eucl (compat "8.3"). -Notation Ndiv := N.div (compat "8.3"). -Notation Nmod := N.modulo (compat "8.3"). +Notation Ndiv_eucl := N.div_eucl (compat "8.6"). +Notation Ndiv := N.div (compat "8.6"). +Notation Nmod := N.modulo (only parsing). -Notation Ndiv_eucl_correct := N.div_eucl_spec (compat "8.3"). -Notation Ndiv_mod_eq := N.div_mod' (compat "8.3"). -Notation Nmod_lt := N.mod_lt (compat "8.3"). +Notation Ndiv_eucl_correct := N.div_eucl_spec (only parsing). +Notation Ndiv_mod_eq := N.div_mod' (only parsing). +Notation Nmod_lt := N.mod_lt (compat "8.6"). diff --git a/theories/NArith/Ngcd_def.v b/theories/NArith/Ngcd_def.v index cfca82eb3..70784d9ce 100644 --- a/theories/NArith/Ngcd_def.v +++ b/theories/NArith/Ngcd_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinPos BinNat. diff --git a/theories/NArith/Nnat.v b/theories/NArith/Nnat.v index 798ab2828..3488c8b43 100644 --- a/theories/NArith/Nnat.v +++ b/theories/NArith/Nnat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinPos BinNat PeanoNat Pnat. @@ -208,30 +210,30 @@ Hint Rewrite Nat2N.id : Nnat. (** Compatibility notations *) -Notation nat_of_N_inj := N2Nat.inj (compat "8.3"). -Notation N_of_nat_of_N := N2Nat.id (compat "8.3"). -Notation nat_of_Ndouble := N2Nat.inj_double (compat "8.3"). -Notation nat_of_Ndouble_plus_one := N2Nat.inj_succ_double (compat "8.3"). -Notation nat_of_Nsucc := N2Nat.inj_succ (compat "8.3"). -Notation nat_of_Nplus := N2Nat.inj_add (compat "8.3"). -Notation nat_of_Nmult := N2Nat.inj_mul (compat "8.3"). -Notation nat_of_Nminus := N2Nat.inj_sub (compat "8.3"). -Notation nat_of_Npred := N2Nat.inj_pred (compat "8.3"). -Notation nat_of_Ndiv2 := N2Nat.inj_div2 (compat "8.3"). -Notation nat_of_Ncompare := N2Nat.inj_compare (compat "8.3"). -Notation nat_of_Nmax := N2Nat.inj_max (compat "8.3"). -Notation nat_of_Nmin := N2Nat.inj_min (compat "8.3"). - -Notation nat_of_N_of_nat := Nat2N.id (compat "8.3"). -Notation N_of_nat_inj := Nat2N.inj (compat "8.3"). -Notation N_of_double := Nat2N.inj_double (compat "8.3"). -Notation N_of_double_plus_one := Nat2N.inj_succ_double (compat "8.3"). -Notation N_of_S := Nat2N.inj_succ (compat "8.3"). -Notation N_of_pred := Nat2N.inj_pred (compat "8.3"). -Notation N_of_plus := Nat2N.inj_add (compat "8.3"). -Notation N_of_minus := Nat2N.inj_sub (compat "8.3"). -Notation N_of_mult := Nat2N.inj_mul (compat "8.3"). -Notation N_of_div2 := Nat2N.inj_div2 (compat "8.3"). -Notation N_of_nat_compare := Nat2N.inj_compare (compat "8.3"). -Notation N_of_min := Nat2N.inj_min (compat "8.3"). -Notation N_of_max := Nat2N.inj_max (compat "8.3"). +Notation nat_of_N_inj := N2Nat.inj (only parsing). +Notation N_of_nat_of_N := N2Nat.id (only parsing). +Notation nat_of_Ndouble := N2Nat.inj_double (only parsing). +Notation nat_of_Ndouble_plus_one := N2Nat.inj_succ_double (only parsing). +Notation nat_of_Nsucc := N2Nat.inj_succ (only parsing). +Notation nat_of_Nplus := N2Nat.inj_add (only parsing). +Notation nat_of_Nmult := N2Nat.inj_mul (only parsing). +Notation nat_of_Nminus := N2Nat.inj_sub (only parsing). +Notation nat_of_Npred := N2Nat.inj_pred (only parsing). +Notation nat_of_Ndiv2 := N2Nat.inj_div2 (only parsing). +Notation nat_of_Ncompare := N2Nat.inj_compare (only parsing). +Notation nat_of_Nmax := N2Nat.inj_max (only parsing). +Notation nat_of_Nmin := N2Nat.inj_min (only parsing). + +Notation nat_of_N_of_nat := Nat2N.id (only parsing). +Notation N_of_nat_inj := Nat2N.inj (only parsing). +Notation N_of_double := Nat2N.inj_double (only parsing). +Notation N_of_double_plus_one := Nat2N.inj_succ_double (only parsing). +Notation N_of_S := Nat2N.inj_succ (only parsing). +Notation N_of_pred := Nat2N.inj_pred (only parsing). +Notation N_of_plus := Nat2N.inj_add (only parsing). +Notation N_of_minus := Nat2N.inj_sub (only parsing). +Notation N_of_mult := Nat2N.inj_mul (only parsing). +Notation N_of_div2 := Nat2N.inj_div2 (only parsing). +Notation N_of_nat_compare := Nat2N.inj_compare (only parsing). +Notation N_of_min := Nat2N.inj_min (only parsing). +Notation N_of_max := Nat2N.inj_max (only parsing). diff --git a/theories/NArith/Nsqrt_def.v b/theories/NArith/Nsqrt_def.v index 97de41c20..e771fe916 100644 --- a/theories/NArith/Nsqrt_def.v +++ b/theories/NArith/Nsqrt_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinNat. @@ -11,8 +13,8 @@ Require Import BinNat. (** Obsolete file, see [BinNat] now, only compatibility notations remain here. *) -Notation Nsqrtrem := N.sqrtrem (compat "8.3"). -Notation Nsqrt := N.sqrt (compat "8.3"). -Notation Nsqrtrem_spec := N.sqrtrem_spec (compat "8.3"). -Notation Nsqrt_spec := (fun n => N.sqrt_spec n (N.le_0_l n)) (compat "8.3"). -Notation Nsqrtrem_sqrt := N.sqrtrem_sqrt (compat "8.3"). +Notation Nsqrtrem := N.sqrtrem (compat "8.6"). +Notation Nsqrt := N.sqrt (compat "8.6"). +Notation Nsqrtrem_spec := N.sqrtrem_spec (compat "8.6"). +Notation Nsqrt_spec := (fun n => N.sqrt_spec n (N.le_0_l n)) (only parsing). +Notation Nsqrtrem_sqrt := N.sqrtrem_sqrt (compat "8.6"). diff --git a/theories/Numbers/BinNums.v b/theories/Numbers/BinNums.v index d7e4185f7..f8b3d9e1d 100644 --- a/theories/Numbers/BinNums.v +++ b/theories/Numbers/BinNums.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Binary Numerical Datatypes *) diff --git a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v index 1777743fa..951a4ef2b 100644 --- a/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v +++ b/theories/Numbers/Cyclic/Abstract/CyclicAxioms.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Cyclic/Abstract/DoubleType.v b/theories/Numbers/Cyclic/Abstract/DoubleType.v index 73a064acd..fe0476e4d 100644 --- a/theories/Numbers/Cyclic/Abstract/DoubleType.v +++ b/theories/Numbers/Cyclic/Abstract/DoubleType.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Cyclic/Abstract/NZCyclic.v b/theories/Numbers/Cyclic/Abstract/NZCyclic.v index 94f860841..64935ffe1 100644 --- a/theories/Numbers/Cyclic/Abstract/NZCyclic.v +++ b/theories/Numbers/Cyclic/Abstract/NZCyclic.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Cyclic/Int31/Cyclic31.v b/theories/Numbers/Cyclic/Int31/Cyclic31.v index d2390e2ce..bd4f0279d 100644 --- a/theories/Numbers/Cyclic/Int31/Cyclic31.v +++ b/theories/Numbers/Cyclic/Int31/Cyclic31.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Int31 numbers defines indeed a cyclic structure : Z/(2^31)Z *) diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v index cc359bfdf..9f8da831d 100644 --- a/theories/Numbers/Cyclic/Int31/Int31.v +++ b/theories/Numbers/Cyclic/Int31/Int31.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Benjamin Gregoire, Laurent Thery, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Cyclic/Int31/Ring31.v b/theories/Numbers/Cyclic/Int31/Ring31.v index c076dbd9a..b69352945 100644 --- a/theories/Numbers/Cyclic/Int31/Ring31.v +++ b/theories/Numbers/Cyclic/Int31/Ring31.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Int31 numbers defines Z/(2^31)Z, and can hence be equipped diff --git a/theories/Numbers/Cyclic/ZModulo/ZModulo.v b/theories/Numbers/Cyclic/ZModulo/ZModulo.v index e7658841a..784e81758 100644 --- a/theories/Numbers/Cyclic/ZModulo/ZModulo.v +++ b/theories/Numbers/Cyclic/ZModulo/ZModulo.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Type [Z] viewed modulo a particular constant corresponds to [Z/nZ] diff --git a/theories/Numbers/DecimalFacts.v b/theories/Numbers/DecimalFacts.v new file mode 100644 index 000000000..0f4905277 --- /dev/null +++ b/theories/Numbers/DecimalFacts.v @@ -0,0 +1,143 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) + +(** * 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..ef00e2805 --- /dev/null +++ b/theories/Numbers/DecimalN.v @@ -0,0 +1,107 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) + +(** * 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..5ffe1688b --- /dev/null +++ b/theories/Numbers/DecimalNat.v @@ -0,0 +1,302 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) + +(** * 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..722e73d96 --- /dev/null +++ b/theories/Numbers/DecimalPos.v @@ -0,0 +1,383 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) + +(** * 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..1a3220f63 --- /dev/null +++ b/theories/Numbers/DecimalString.v @@ -0,0 +1,265 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 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..3a0837963 --- /dev/null +++ b/theories/Numbers/DecimalZ.v @@ -0,0 +1,75 @@ +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) + +(** * 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/Integer/Abstract/ZAdd.v b/theories/Numbers/Integer/Abstract/ZAdd.v index 9d94ff65d..c4c5174da 100644 --- a/theories/Numbers/Integer/Abstract/ZAdd.v +++ b/theories/Numbers/Integer/Abstract/ZAdd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZAddOrder.v b/theories/Numbers/Integer/Abstract/ZAddOrder.v index 84bb1a24a..7f5b0df68 100644 --- a/theories/Numbers/Integer/Abstract/ZAddOrder.v +++ b/theories/Numbers/Integer/Abstract/ZAddOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZAxioms.v b/theories/Numbers/Integer/Abstract/ZAxioms.v index 6efb7ea23..4f1ab7752 100644 --- a/theories/Numbers/Integer/Abstract/ZAxioms.v +++ b/theories/Numbers/Integer/Abstract/ZAxioms.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZBase.v b/theories/Numbers/Integer/Abstract/ZBase.v index b2014135f..7fdd018d3 100644 --- a/theories/Numbers/Integer/Abstract/ZBase.v +++ b/theories/Numbers/Integer/Abstract/ZBase.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZBits.v b/theories/Numbers/Integer/Abstract/ZBits.v index 68d27fec9..2da445281 100644 --- a/theories/Numbers/Integer/Abstract/ZBits.v +++ b/theories/Numbers/Integer/Abstract/ZBits.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v index 967b68d36..d7f25a661 100644 --- a/theories/Numbers/Integer/Abstract/ZDivEucl.v +++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZAxioms ZMulOrder ZSgnAbs NZDiv. diff --git a/theories/Numbers/Integer/Abstract/ZDivFloor.v b/theories/Numbers/Integer/Abstract/ZDivFloor.v index a9077127e..a0d1821b6 100644 --- a/theories/Numbers/Integer/Abstract/ZDivFloor.v +++ b/theories/Numbers/Integer/Abstract/ZDivFloor.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZAxioms ZMulOrder ZSgnAbs NZDiv. diff --git a/theories/Numbers/Integer/Abstract/ZDivTrunc.v b/theories/Numbers/Integer/Abstract/ZDivTrunc.v index bbb8ad5ae..31e427383 100644 --- a/theories/Numbers/Integer/Abstract/ZDivTrunc.v +++ b/theories/Numbers/Integer/Abstract/ZDivTrunc.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZAxioms ZMulOrder ZSgnAbs NZDiv. diff --git a/theories/Numbers/Integer/Abstract/ZGcd.v b/theories/Numbers/Integer/Abstract/ZGcd.v index 1144bd2bf..f0b7bf9d2 100644 --- a/theories/Numbers/Integer/Abstract/ZGcd.v +++ b/theories/Numbers/Integer/Abstract/ZGcd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of the greatest common divisor *) diff --git a/theories/Numbers/Integer/Abstract/ZLcm.v b/theories/Numbers/Integer/Abstract/ZLcm.v index 4b0f9f978..0ab528de8 100644 --- a/theories/Numbers/Integer/Abstract/ZLcm.v +++ b/theories/Numbers/Integer/Abstract/ZLcm.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZAxioms ZMulOrder ZSgnAbs ZGcd ZDivTrunc ZDivFloor. diff --git a/theories/Numbers/Integer/Abstract/ZLt.v b/theories/Numbers/Integer/Abstract/ZLt.v index f43d74d70..726b041c2 100644 --- a/theories/Numbers/Integer/Abstract/ZLt.v +++ b/theories/Numbers/Integer/Abstract/ZLt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZMaxMin.v b/theories/Numbers/Integer/Abstract/ZMaxMin.v index 40c3980a5..f3f3a861b 100644 --- a/theories/Numbers/Integer/Abstract/ZMaxMin.v +++ b/theories/Numbers/Integer/Abstract/ZMaxMin.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZAxioms ZMulOrder GenericMinMax. diff --git a/theories/Numbers/Integer/Abstract/ZMul.v b/theories/Numbers/Integer/Abstract/ZMul.v index 2550ae1c5..120647dcc 100644 --- a/theories/Numbers/Integer/Abstract/ZMul.v +++ b/theories/Numbers/Integer/Abstract/ZMul.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZMulOrder.v b/theories/Numbers/Integer/Abstract/ZMulOrder.v index adea36f0e..cd9523d34 100644 --- a/theories/Numbers/Integer/Abstract/ZMulOrder.v +++ b/theories/Numbers/Integer/Abstract/ZMulOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/Abstract/ZParity.v b/theories/Numbers/Integer/Abstract/ZParity.v index db379474b..a5e53b361 100644 --- a/theories/Numbers/Integer/Abstract/ZParity.v +++ b/theories/Numbers/Integer/Abstract/ZParity.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool ZMulOrder NZParity. diff --git a/theories/Numbers/Integer/Abstract/ZPow.v b/theories/Numbers/Integer/Abstract/ZPow.v index 478724e2e..a4b964e52 100644 --- a/theories/Numbers/Integer/Abstract/ZPow.v +++ b/theories/Numbers/Integer/Abstract/ZPow.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of the power function *) diff --git a/theories/Numbers/Integer/Abstract/ZProperties.v b/theories/Numbers/Integer/Abstract/ZProperties.v index 4a0ce5e5b..e4b997cfd 100644 --- a/theories/Numbers/Integer/Abstract/ZProperties.v +++ b/theories/Numbers/Integer/Abstract/ZProperties.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export ZAxioms ZMaxMin ZSgnAbs ZParity ZPow ZDivTrunc ZDivFloor diff --git a/theories/Numbers/Integer/Abstract/ZSgnAbs.v b/theories/Numbers/Integer/Abstract/ZSgnAbs.v index b2f77c3f5..dda128726 100644 --- a/theories/Numbers/Integer/Abstract/ZSgnAbs.v +++ b/theories/Numbers/Integer/Abstract/ZSgnAbs.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of [abs] and [sgn] *) diff --git a/theories/Numbers/Integer/Binary/ZBinary.v b/theories/Numbers/Integer/Binary/ZBinary.v index 553bd68ac..bed827fd0 100644 --- a/theories/Numbers/Integer/Binary/ZBinary.v +++ b/theories/Numbers/Integer/Binary/ZBinary.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Integer/NatPairs/ZNatPairs.v b/theories/Numbers/Integer/NatPairs/ZNatPairs.v index c640145ba..4b2d5c13b 100644 --- a/theories/Numbers/Integer/NatPairs/ZNatPairs.v +++ b/theories/Numbers/Integer/NatPairs/ZNatPairs.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NaryFunctions.v b/theories/Numbers/NaryFunctions.v index cf0d2e835..ee28628ed 100644 --- a/theories/Numbers/NaryFunctions.v +++ b/theories/Numbers/NaryFunctions.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Pierre Letouzey, Jerome Vouillon, PPS, Paris 7, 2008 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZAdd.v b/theories/Numbers/NatInt/NZAdd.v index 4beab1e1b..bc366c508 100644 --- a/theories/Numbers/NatInt/NZAdd.v +++ b/theories/Numbers/NatInt/NZAdd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZAddOrder.v b/theories/Numbers/NatInt/NZAddOrder.v index 82e500ca2..99812ee3f 100644 --- a/theories/Numbers/NatInt/NZAddOrder.v +++ b/theories/Numbers/NatInt/NZAddOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZAxioms.v b/theories/Numbers/NatInt/NZAxioms.v index c0851791d..8c364cde7 100644 --- a/theories/Numbers/NatInt/NZAxioms.v +++ b/theories/Numbers/NatInt/NZAxioms.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Initial Author : Evgeny Makarov, INRIA, 2007 *) diff --git a/theories/Numbers/NatInt/NZBase.v b/theories/Numbers/NatInt/NZBase.v index e0d55b893..595b2182a 100644 --- a/theories/Numbers/NatInt/NZBase.v +++ b/theories/Numbers/NatInt/NZBase.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZBits.v b/theories/Numbers/NatInt/NZBits.v index d731de62d..eefa51572 100644 --- a/theories/Numbers/NatInt/NZBits.v +++ b/theories/Numbers/NatInt/NZBits.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool NZAxioms NZMulOrder NZParity NZPow NZDiv NZLog. diff --git a/theories/Numbers/NatInt/NZDiv.v b/theories/Numbers/NatInt/NZDiv.v index 5610eccb7..550aa226a 100644 --- a/theories/Numbers/NatInt/NZDiv.v +++ b/theories/Numbers/NatInt/NZDiv.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Euclidean Division *) diff --git a/theories/Numbers/NatInt/NZDomain.v b/theories/Numbers/NatInt/NZDomain.v index f5ee1351d..3d0c005fd 100644 --- a/theories/Numbers/NatInt/NZDomain.v +++ b/theories/Numbers/NatInt/NZDomain.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export NumPrelude NZAxioms. diff --git a/theories/Numbers/NatInt/NZGcd.v b/theories/Numbers/NatInt/NZGcd.v index f8af97196..c38d1aac3 100644 --- a/theories/Numbers/NatInt/NZGcd.v +++ b/theories/Numbers/NatInt/NZGcd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Greatest Common Divisor *) diff --git a/theories/Numbers/NatInt/NZLog.v b/theories/Numbers/NatInt/NZLog.v index 239f20a42..794851a9d 100644 --- a/theories/Numbers/NatInt/NZLog.v +++ b/theories/Numbers/NatInt/NZLog.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Base-2 Logarithm *) diff --git a/theories/Numbers/NatInt/NZMul.v b/theories/Numbers/NatInt/NZMul.v index 145f56352..44cbc5171 100644 --- a/theories/Numbers/NatInt/NZMul.v +++ b/theories/Numbers/NatInt/NZMul.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZMulOrder.v b/theories/Numbers/NatInt/NZMulOrder.v index 430f8fd2e..292f0837c 100644 --- a/theories/Numbers/NatInt/NZMulOrder.v +++ b/theories/Numbers/NatInt/NZMulOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZOrder.v b/theories/Numbers/NatInt/NZOrder.v index 37ec46410..60e1123b3 100644 --- a/theories/Numbers/NatInt/NZOrder.v +++ b/theories/Numbers/NatInt/NZOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZParity.v b/theories/Numbers/NatInt/NZParity.v index 626d59d73..93d99f08f 100644 --- a/theories/Numbers/NatInt/NZParity.v +++ b/theories/Numbers/NatInt/NZParity.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool NZAxioms NZMulOrder. diff --git a/theories/Numbers/NatInt/NZPow.v b/theories/Numbers/NatInt/NZPow.v index 350047fa1..a1310667e 100644 --- a/theories/Numbers/NatInt/NZPow.v +++ b/theories/Numbers/NatInt/NZPow.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Power Function *) diff --git a/theories/Numbers/NatInt/NZProperties.v b/theories/Numbers/NatInt/NZProperties.v index aee3b6245..fbcf43e88 100644 --- a/theories/Numbers/NatInt/NZProperties.v +++ b/theories/Numbers/NatInt/NZProperties.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/NatInt/NZSqrt.v b/theories/Numbers/NatInt/NZSqrt.v index c03107299..c2d2c4ae1 100644 --- a/theories/Numbers/NatInt/NZSqrt.v +++ b/theories/Numbers/NatInt/NZSqrt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Square Root Function *) diff --git a/theories/Numbers/Natural/Abstract/NAdd.v b/theories/Numbers/Natural/Abstract/NAdd.v index 899fa3933..dc5f8e537 100644 --- a/theories/Numbers/Natural/Abstract/NAdd.v +++ b/theories/Numbers/Natural/Abstract/NAdd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NAddOrder.v b/theories/Numbers/Natural/Abstract/NAddOrder.v index d62903db6..2da3f0bfc 100644 --- a/theories/Numbers/Natural/Abstract/NAddOrder.v +++ b/theories/Numbers/Natural/Abstract/NAddOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NAxioms.v b/theories/Numbers/Natural/Abstract/NAxioms.v index d67689dbd..dd09ac5f3 100644 --- a/theories/Numbers/Natural/Abstract/NAxioms.v +++ b/theories/Numbers/Natural/Abstract/NAxioms.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NBase.v b/theories/Numbers/Natural/Abstract/NBase.v index 85a015cf2..ad0b3d3d2 100644 --- a/theories/Numbers/Natural/Abstract/NBase.v +++ b/theories/Numbers/Natural/Abstract/NBase.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NBits.v b/theories/Numbers/Natural/Abstract/NBits.v index 1e7644eda..e1391f599 100644 --- a/theories/Numbers/Natural/Abstract/NBits.v +++ b/theories/Numbers/Natural/Abstract/NBits.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool NAxioms NSub NPow NDiv NParity NLog. diff --git a/theories/Numbers/Natural/Abstract/NDefOps.v b/theories/Numbers/Natural/Abstract/NDefOps.v index e934eda9b..8e1be0d70 100644 --- a/theories/Numbers/Natural/Abstract/NDefOps.v +++ b/theories/Numbers/Natural/Abstract/NDefOps.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NDiv.v b/theories/Numbers/Natural/Abstract/NDiv.v index eff27abc5..4c26a071f 100644 --- a/theories/Numbers/Natural/Abstract/NDiv.v +++ b/theories/Numbers/Natural/Abstract/NDiv.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 NAxioms NSub NZDiv. diff --git a/theories/Numbers/Natural/Abstract/NGcd.v b/theories/Numbers/Natural/Abstract/NGcd.v index 943caef63..96fb4247c 100644 --- a/theories/Numbers/Natural/Abstract/NGcd.v +++ b/theories/Numbers/Natural/Abstract/NGcd.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of the greatest common divisor *) diff --git a/theories/Numbers/Natural/Abstract/NIso.v b/theories/Numbers/Natural/Abstract/NIso.v index 2b61f9784..d41d0aff5 100644 --- a/theories/Numbers/Natural/Abstract/NIso.v +++ b/theories/Numbers/Natural/Abstract/NIso.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NLcm.v b/theories/Numbers/Natural/Abstract/NLcm.v index 2e7bcf40c..47b74193e 100644 --- a/theories/Numbers/Natural/Abstract/NLcm.v +++ b/theories/Numbers/Natural/Abstract/NLcm.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 NAxioms NSub NDiv NGcd. diff --git a/theories/Numbers/Natural/Abstract/NLog.v b/theories/Numbers/Natural/Abstract/NLog.v index d22510ab9..fe6fcee56 100644 --- a/theories/Numbers/Natural/Abstract/NLog.v +++ b/theories/Numbers/Natural/Abstract/NLog.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Base-2 Logarithm Properties *) diff --git a/theories/Numbers/Natural/Abstract/NMaxMin.v b/theories/Numbers/Natural/Abstract/NMaxMin.v index 1020fe375..3cf4d3f9f 100644 --- a/theories/Numbers/Natural/Abstract/NMaxMin.v +++ b/theories/Numbers/Natural/Abstract/NMaxMin.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 NAxioms NSub GenericMinMax. diff --git a/theories/Numbers/Natural/Abstract/NMulOrder.v b/theories/Numbers/Natural/Abstract/NMulOrder.v index ace1e736e..b7f1c8e45 100644 --- a/theories/Numbers/Natural/Abstract/NMulOrder.v +++ b/theories/Numbers/Natural/Abstract/NMulOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NOrder.v b/theories/Numbers/Natural/Abstract/NOrder.v index f05d783ad..acaecad93 100644 --- a/theories/Numbers/Natural/Abstract/NOrder.v +++ b/theories/Numbers/Natural/Abstract/NOrder.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NParity.v b/theories/Numbers/Natural/Abstract/NParity.v index fd136ff93..cb89e1d72 100644 --- a/theories/Numbers/Natural/Abstract/NParity.v +++ b/theories/Numbers/Natural/Abstract/NParity.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Bool NSub NZParity. diff --git a/theories/Numbers/Natural/Abstract/NPow.v b/theories/Numbers/Natural/Abstract/NPow.v index d31d67a1c..fc1cc93b1 100644 --- a/theories/Numbers/Natural/Abstract/NPow.v +++ b/theories/Numbers/Natural/Abstract/NPow.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of the power function *) diff --git a/theories/Numbers/Natural/Abstract/NProperties.v b/theories/Numbers/Natural/Abstract/NProperties.v index b753d659a..bcf906cf9 100644 --- a/theories/Numbers/Natural/Abstract/NProperties.v +++ b/theories/Numbers/Natural/Abstract/NProperties.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export NAxioms. diff --git a/theories/Numbers/Natural/Abstract/NSqrt.v b/theories/Numbers/Natural/Abstract/NSqrt.v index 68c06775e..6bffe693e 100644 --- a/theories/Numbers/Natural/Abstract/NSqrt.v +++ b/theories/Numbers/Natural/Abstract/NSqrt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Properties of Square Root Function *) diff --git a/theories/Numbers/Natural/Abstract/NStrongRec.v b/theories/Numbers/Natural/Abstract/NStrongRec.v index 8e825ef7d..f76d8ae8a 100644 --- a/theories/Numbers/Natural/Abstract/NStrongRec.v +++ b/theories/Numbers/Natural/Abstract/NStrongRec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Abstract/NSub.v b/theories/Numbers/Natural/Abstract/NSub.v index dce78f610..453b0c0d4 100644 --- a/theories/Numbers/Natural/Abstract/NSub.v +++ b/theories/Numbers/Natural/Abstract/NSub.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Binary/NBinary.v b/theories/Numbers/Natural/Binary/NBinary.v index bdb715f33..c9e1c6401 100644 --- a/theories/Numbers/Natural/Binary/NBinary.v +++ b/theories/Numbers/Natural/Binary/NBinary.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/Numbers/Natural/Peano/NPeano.v b/theories/Numbers/Natural/Peano/NPeano.v index 787ef81dc..6000bdcf7 100644 --- a/theories/Numbers/Natural/Peano/NPeano.v +++ b/theories/Numbers/Natural/Peano/NPeano.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) @@ -18,74 +20,74 @@ Module Nat <: NAxiomsSig := Nat. (** Compat notations for stuff that used to be at the beginning of NPeano. *) -Notation leb := Nat.leb (compat "8.4"). -Notation ltb := Nat.ltb (compat "8.4"). -Notation leb_le := Nat.leb_le (compat "8.4"). -Notation ltb_lt := Nat.ltb_lt (compat "8.4"). -Notation pow := Nat.pow (compat "8.4"). -Notation pow_0_r := Nat.pow_0_r (compat "8.4"). -Notation pow_succ_r := Nat.pow_succ_r (compat "8.4"). -Notation square := Nat.square (compat "8.4"). -Notation square_spec := Nat.square_spec (compat "8.4"). -Notation Even := Nat.Even (compat "8.4"). -Notation Odd := Nat.Odd (compat "8.4"). -Notation even := Nat.even (compat "8.4"). -Notation odd := Nat.odd (compat "8.4"). -Notation even_spec := Nat.even_spec (compat "8.4"). -Notation odd_spec := Nat.odd_spec (compat "8.4"). +Notation leb := Nat.leb (only parsing). +Notation ltb := Nat.ltb (only parsing). +Notation leb_le := Nat.leb_le (only parsing). +Notation ltb_lt := Nat.ltb_lt (only parsing). +Notation pow := Nat.pow (only parsing). +Notation pow_0_r := Nat.pow_0_r (only parsing). +Notation pow_succ_r := Nat.pow_succ_r (only parsing). +Notation square := Nat.square (only parsing). +Notation square_spec := Nat.square_spec (only parsing). +Notation Even := Nat.Even (only parsing). +Notation Odd := Nat.Odd (only parsing). +Notation even := Nat.even (only parsing). +Notation odd := Nat.odd (only parsing). +Notation even_spec := Nat.even_spec (only parsing). +Notation odd_spec := Nat.odd_spec (only parsing). Lemma Even_equiv n : Even n <-> Even.even n. Proof. symmetry. apply Even.even_equiv. Qed. Lemma Odd_equiv n : Odd n <-> Even.odd n. Proof. symmetry. apply Even.odd_equiv. Qed. -Notation divmod := Nat.divmod (compat "8.4"). -Notation div := Nat.div (compat "8.4"). -Notation modulo := Nat.modulo (compat "8.4"). -Notation divmod_spec := Nat.divmod_spec (compat "8.4"). -Notation div_mod := Nat.div_mod (compat "8.4"). -Notation mod_bound_pos := Nat.mod_bound_pos (compat "8.4"). -Notation sqrt_iter := Nat.sqrt_iter (compat "8.4"). -Notation sqrt := Nat.sqrt (compat "8.4"). -Notation sqrt_iter_spec := Nat.sqrt_iter_spec (compat "8.4"). -Notation sqrt_spec := Nat.sqrt_spec (compat "8.4"). -Notation log2_iter := Nat.log2_iter (compat "8.4"). -Notation log2 := Nat.log2 (compat "8.4"). -Notation log2_iter_spec := Nat.log2_iter_spec (compat "8.4"). -Notation log2_spec := Nat.log2_spec (compat "8.4"). -Notation log2_nonpos := Nat.log2_nonpos (compat "8.4"). -Notation gcd := Nat.gcd (compat "8.4"). -Notation divide := Nat.divide (compat "8.4"). -Notation gcd_divide := Nat.gcd_divide (compat "8.4"). -Notation gcd_divide_l := Nat.gcd_divide_l (compat "8.4"). -Notation gcd_divide_r := Nat.gcd_divide_r (compat "8.4"). -Notation gcd_greatest := Nat.gcd_greatest (compat "8.4"). -Notation testbit := Nat.testbit (compat "8.4"). -Notation shiftl := Nat.shiftl (compat "8.4"). -Notation shiftr := Nat.shiftr (compat "8.4"). -Notation bitwise := Nat.bitwise (compat "8.4"). -Notation land := Nat.land (compat "8.4"). -Notation lor := Nat.lor (compat "8.4"). -Notation ldiff := Nat.ldiff (compat "8.4"). -Notation lxor := Nat.lxor (compat "8.4"). -Notation double_twice := Nat.double_twice (compat "8.4"). -Notation testbit_0_l := Nat.testbit_0_l (compat "8.4"). -Notation testbit_odd_0 := Nat.testbit_odd_0 (compat "8.4"). -Notation testbit_even_0 := Nat.testbit_even_0 (compat "8.4"). -Notation testbit_odd_succ := Nat.testbit_odd_succ (compat "8.4"). -Notation testbit_even_succ := Nat.testbit_even_succ (compat "8.4"). -Notation shiftr_spec := Nat.shiftr_spec (compat "8.4"). -Notation shiftl_spec_high := Nat.shiftl_spec_high (compat "8.4"). -Notation shiftl_spec_low := Nat.shiftl_spec_low (compat "8.4"). -Notation div2_bitwise := Nat.div2_bitwise (compat "8.4"). -Notation odd_bitwise := Nat.odd_bitwise (compat "8.4"). -Notation div2_decr := Nat.div2_decr (compat "8.4"). -Notation testbit_bitwise_1 := Nat.testbit_bitwise_1 (compat "8.4"). -Notation testbit_bitwise_2 := Nat.testbit_bitwise_2 (compat "8.4"). -Notation land_spec := Nat.land_spec (compat "8.4"). -Notation ldiff_spec := Nat.ldiff_spec (compat "8.4"). -Notation lor_spec := Nat.lor_spec (compat "8.4"). -Notation lxor_spec := Nat.lxor_spec (compat "8.4"). +Notation divmod := Nat.divmod (only parsing). +Notation div := Nat.div (only parsing). +Notation modulo := Nat.modulo (only parsing). +Notation divmod_spec := Nat.divmod_spec (only parsing). +Notation div_mod := Nat.div_mod (only parsing). +Notation mod_bound_pos := Nat.mod_bound_pos (only parsing). +Notation sqrt_iter := Nat.sqrt_iter (only parsing). +Notation sqrt := Nat.sqrt (only parsing). +Notation sqrt_iter_spec := Nat.sqrt_iter_spec (only parsing). +Notation sqrt_spec := Nat.sqrt_spec (only parsing). +Notation log2_iter := Nat.log2_iter (only parsing). +Notation log2 := Nat.log2 (only parsing). +Notation log2_iter_spec := Nat.log2_iter_spec (only parsing). +Notation log2_spec := Nat.log2_spec (only parsing). +Notation log2_nonpos := Nat.log2_nonpos (only parsing). +Notation gcd := Nat.gcd (only parsing). +Notation divide := Nat.divide (only parsing). +Notation gcd_divide := Nat.gcd_divide (only parsing). +Notation gcd_divide_l := Nat.gcd_divide_l (only parsing). +Notation gcd_divide_r := Nat.gcd_divide_r (only parsing). +Notation gcd_greatest := Nat.gcd_greatest (only parsing). +Notation testbit := Nat.testbit (only parsing). +Notation shiftl := Nat.shiftl (only parsing). +Notation shiftr := Nat.shiftr (only parsing). +Notation bitwise := Nat.bitwise (only parsing). +Notation land := Nat.land (only parsing). +Notation lor := Nat.lor (only parsing). +Notation ldiff := Nat.ldiff (only parsing). +Notation lxor := Nat.lxor (only parsing). +Notation double_twice := Nat.double_twice (only parsing). +Notation testbit_0_l := Nat.testbit_0_l (only parsing). +Notation testbit_odd_0 := Nat.testbit_odd_0 (only parsing). +Notation testbit_even_0 := Nat.testbit_even_0 (only parsing). +Notation testbit_odd_succ := Nat.testbit_odd_succ (only parsing). +Notation testbit_even_succ := Nat.testbit_even_succ (only parsing). +Notation shiftr_spec := Nat.shiftr_spec (only parsing). +Notation shiftl_spec_high := Nat.shiftl_spec_high (only parsing). +Notation shiftl_spec_low := Nat.shiftl_spec_low (only parsing). +Notation div2_bitwise := Nat.div2_bitwise (only parsing). +Notation odd_bitwise := Nat.odd_bitwise (only parsing). +Notation div2_decr := Nat.div2_decr (only parsing). +Notation testbit_bitwise_1 := Nat.testbit_bitwise_1 (only parsing). +Notation testbit_bitwise_2 := Nat.testbit_bitwise_2 (only parsing). +Notation land_spec := Nat.land_spec (only parsing). +Notation ldiff_spec := Nat.ldiff_spec (only parsing). +Notation lor_spec := Nat.lor_spec (only parsing). +Notation lxor_spec := Nat.lxor_spec (only parsing). Infix "<=?" := Nat.leb (at level 70) : nat_scope. Infix "<?" := Nat.ltb (at level 70) : nat_scope. diff --git a/theories/Numbers/NumPrelude.v b/theories/Numbers/NumPrelude.v index f01d5880e..7cf13feae 100644 --- a/theories/Numbers/NumPrelude.v +++ b/theories/Numbers/NumPrelude.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Evgeny Makarov, INRIA, 2007 *) (************************************************************************) diff --git a/theories/PArith/BinPos.v b/theories/PArith/BinPos.v index ff880eefa..8d0896db7 100644 --- a/theories/PArith/BinPos.v +++ b/theories/PArith/BinPos.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export BinNums. @@ -1903,180 +1905,180 @@ Notation IsNul := Pos.IsNul (only parsing). Notation IsPos := Pos.IsPos (only parsing). Notation IsNeg := Pos.IsNeg (only parsing). -Notation Psucc := Pos.succ (compat "8.3"). -Notation Pplus := Pos.add (compat "8.3"). -Notation Pplus_carry := Pos.add_carry (compat "8.3"). -Notation Ppred := Pos.pred (compat "8.3"). -Notation Piter_op := Pos.iter_op (compat "8.3"). -Notation Piter_op_succ := Pos.iter_op_succ (compat "8.3"). -Notation Pmult_nat := (Pos.iter_op plus) (compat "8.3"). -Notation nat_of_P := Pos.to_nat (compat "8.3"). -Notation P_of_succ_nat := Pos.of_succ_nat (compat "8.3"). -Notation Pdouble_minus_one := Pos.pred_double (compat "8.3"). -Notation positive_mask := Pos.mask (compat "8.3"). -Notation positive_mask_rect := Pos.mask_rect (compat "8.3"). -Notation positive_mask_ind := Pos.mask_ind (compat "8.3"). -Notation positive_mask_rec := Pos.mask_rec (compat "8.3"). -Notation Pdouble_plus_one_mask := Pos.succ_double_mask (compat "8.3"). -Notation Pdouble_mask := Pos.double_mask (compat "8.3"). -Notation Pdouble_minus_two := Pos.double_pred_mask (compat "8.3"). -Notation Pminus_mask := Pos.sub_mask (compat "8.3"). -Notation Pminus_mask_carry := Pos.sub_mask_carry (compat "8.3"). -Notation Pminus := Pos.sub (compat "8.3"). -Notation Pmult := Pos.mul (compat "8.3"). -Notation iter_pos := @Pos.iter (compat "8.3"). -Notation Ppow := Pos.pow (compat "8.3"). -Notation Pdiv2 := Pos.div2 (compat "8.3"). -Notation Pdiv2_up := Pos.div2_up (compat "8.3"). -Notation Psize := Pos.size_nat (compat "8.3"). -Notation Psize_pos := Pos.size (compat "8.3"). -Notation Pcompare x y m := (Pos.compare_cont m x y) (compat "8.3"). -Notation Plt := Pos.lt (compat "8.3"). -Notation Pgt := Pos.gt (compat "8.3"). -Notation Ple := Pos.le (compat "8.3"). -Notation Pge := Pos.ge (compat "8.3"). -Notation Pmin := Pos.min (compat "8.3"). -Notation Pmax := Pos.max (compat "8.3"). -Notation Peqb := Pos.eqb (compat "8.3"). -Notation positive_eq_dec := Pos.eq_dec (compat "8.3"). -Notation xI_succ_xO := Pos.xI_succ_xO (compat "8.3"). -Notation Psucc_discr := Pos.succ_discr (compat "8.3"). +Notation Psucc := Pos.succ (compat "8.6"). +Notation Pplus := Pos.add (only parsing). +Notation Pplus_carry := Pos.add_carry (only parsing). +Notation Ppred := Pos.pred (compat "8.6"). +Notation Piter_op := Pos.iter_op (compat "8.6"). +Notation Piter_op_succ := Pos.iter_op_succ (compat "8.6"). +Notation Pmult_nat := (Pos.iter_op plus) (only parsing). +Notation nat_of_P := Pos.to_nat (only parsing). +Notation P_of_succ_nat := Pos.of_succ_nat (only parsing). +Notation Pdouble_minus_one := Pos.pred_double (only parsing). +Notation positive_mask := Pos.mask (only parsing). +Notation positive_mask_rect := Pos.mask_rect (only parsing). +Notation positive_mask_ind := Pos.mask_ind (only parsing). +Notation positive_mask_rec := Pos.mask_rec (only parsing). +Notation Pdouble_plus_one_mask := Pos.succ_double_mask (only parsing). +Notation Pdouble_mask := Pos.double_mask (compat "8.6"). +Notation Pdouble_minus_two := Pos.double_pred_mask (only parsing). +Notation Pminus_mask := Pos.sub_mask (only parsing). +Notation Pminus_mask_carry := Pos.sub_mask_carry (only parsing). +Notation Pminus := Pos.sub (only parsing). +Notation Pmult := Pos.mul (only parsing). +Notation iter_pos := @Pos.iter (only parsing). +Notation Ppow := Pos.pow (compat "8.6"). +Notation Pdiv2 := Pos.div2 (compat "8.6"). +Notation Pdiv2_up := Pos.div2_up (compat "8.6"). +Notation Psize := Pos.size_nat (only parsing). +Notation Psize_pos := Pos.size (only parsing). +Notation Pcompare x y m := (Pos.compare_cont m x y) (only parsing). +Notation Plt := Pos.lt (compat "8.6"). +Notation Pgt := Pos.gt (compat "8.6"). +Notation Ple := Pos.le (compat "8.6"). +Notation Pge := Pos.ge (compat "8.6"). +Notation Pmin := Pos.min (compat "8.6"). +Notation Pmax := Pos.max (compat "8.6"). +Notation Peqb := Pos.eqb (compat "8.6"). +Notation positive_eq_dec := Pos.eq_dec (only parsing). +Notation xI_succ_xO := Pos.xI_succ_xO (only parsing). +Notation Psucc_discr := Pos.succ_discr (compat "8.6"). Notation Psucc_o_double_minus_one_eq_xO := - Pos.succ_pred_double (compat "8.3"). + Pos.succ_pred_double (only parsing). Notation Pdouble_minus_one_o_succ_eq_xI := - Pos.pred_double_succ (compat "8.3"). -Notation xO_succ_permute := Pos.double_succ (compat "8.3"). + Pos.pred_double_succ (only parsing). +Notation xO_succ_permute := Pos.double_succ (only parsing). Notation double_moins_un_xO_discr := - Pos.pred_double_xO_discr (compat "8.3"). -Notation Psucc_not_one := Pos.succ_not_1 (compat "8.3"). -Notation Ppred_succ := Pos.pred_succ (compat "8.3"). -Notation Psucc_pred := Pos.succ_pred_or (compat "8.3"). -Notation Psucc_inj := Pos.succ_inj (compat "8.3"). -Notation Pplus_carry_spec := Pos.add_carry_spec (compat "8.3"). -Notation Pplus_comm := Pos.add_comm (compat "8.3"). -Notation Pplus_succ_permute_r := Pos.add_succ_r (compat "8.3"). -Notation Pplus_succ_permute_l := Pos.add_succ_l (compat "8.3"). -Notation Pplus_no_neutral := Pos.add_no_neutral (compat "8.3"). -Notation Pplus_carry_plus := Pos.add_carry_add (compat "8.3"). -Notation Pplus_reg_r := Pos.add_reg_r (compat "8.3"). -Notation Pplus_reg_l := Pos.add_reg_l (compat "8.3"). -Notation Pplus_carry_reg_r := Pos.add_carry_reg_r (compat "8.3"). -Notation Pplus_carry_reg_l := Pos.add_carry_reg_l (compat "8.3"). -Notation Pplus_assoc := Pos.add_assoc (compat "8.3"). -Notation Pplus_xO := Pos.add_xO (compat "8.3"). -Notation Pplus_xI_double_minus_one := Pos.add_xI_pred_double (compat "8.3"). -Notation Pplus_xO_double_minus_one := Pos.add_xO_pred_double (compat "8.3"). -Notation Pplus_diag := Pos.add_diag (compat "8.3"). -Notation PeanoView := Pos.PeanoView (compat "8.3"). -Notation PeanoOne := Pos.PeanoOne (compat "8.3"). -Notation PeanoSucc := Pos.PeanoSucc (compat "8.3"). -Notation PeanoView_rect := Pos.PeanoView_rect (compat "8.3"). -Notation PeanoView_ind := Pos.PeanoView_ind (compat "8.3"). -Notation PeanoView_rec := Pos.PeanoView_rec (compat "8.3"). -Notation peanoView_xO := Pos.peanoView_xO (compat "8.3"). -Notation peanoView_xI := Pos.peanoView_xI (compat "8.3"). -Notation peanoView := Pos.peanoView (compat "8.3"). -Notation PeanoView_iter := Pos.PeanoView_iter (compat "8.3"). -Notation eq_dep_eq_positive := Pos.eq_dep_eq_positive (compat "8.3"). -Notation PeanoViewUnique := Pos.PeanoViewUnique (compat "8.3"). -Notation Prect := Pos.peano_rect (compat "8.3"). -Notation Prect_succ := Pos.peano_rect_succ (compat "8.3"). -Notation Prect_base := Pos.peano_rect_base (compat "8.3"). -Notation Prec := Pos.peano_rec (compat "8.3"). -Notation Pind := Pos.peano_ind (compat "8.3"). -Notation Pcase := Pos.peano_case (compat "8.3"). -Notation Pmult_1_r := Pos.mul_1_r (compat "8.3"). -Notation Pmult_Sn_m := Pos.mul_succ_l (compat "8.3"). -Notation Pmult_xO_permute_r := Pos.mul_xO_r (compat "8.3"). -Notation Pmult_xI_permute_r := Pos.mul_xI_r (compat "8.3"). -Notation Pmult_comm := Pos.mul_comm (compat "8.3"). -Notation Pmult_plus_distr_l := Pos.mul_add_distr_l (compat "8.3"). -Notation Pmult_plus_distr_r := Pos.mul_add_distr_r (compat "8.3"). -Notation Pmult_assoc := Pos.mul_assoc (compat "8.3"). -Notation Pmult_xI_mult_xO_discr := Pos.mul_xI_mul_xO_discr (compat "8.3"). -Notation Pmult_xO_discr := Pos.mul_xO_discr (compat "8.3"). -Notation Pmult_reg_r := Pos.mul_reg_r (compat "8.3"). -Notation Pmult_reg_l := Pos.mul_reg_l (compat "8.3"). -Notation Pmult_1_inversion_l := Pos.mul_eq_1_l (compat "8.3"). -Notation Psquare_xO := Pos.square_xO (compat "8.3"). -Notation Psquare_xI := Pos.square_xI (compat "8.3"). -Notation iter_pos_swap_gen := Pos.iter_swap_gen (compat "8.3"). -Notation iter_pos_swap := Pos.iter_swap (compat "8.3"). -Notation iter_pos_succ := Pos.iter_succ (compat "8.3"). -Notation iter_pos_plus := Pos.iter_add (compat "8.3"). -Notation iter_pos_invariant := Pos.iter_invariant (compat "8.3"). -Notation Ppow_1_r := Pos.pow_1_r (compat "8.3"). -Notation Ppow_succ_r := Pos.pow_succ_r (compat "8.3"). -Notation Peqb_refl := Pos.eqb_refl (compat "8.3"). -Notation Peqb_eq := Pos.eqb_eq (compat "8.3"). -Notation Pcompare_refl_id := Pos.compare_cont_refl (compat "8.3"). -Notation Pcompare_eq_iff := Pos.compare_eq_iff (compat "8.3"). -Notation Pcompare_Gt_Lt := Pos.compare_cont_Gt_Lt (compat "8.3"). -Notation Pcompare_eq_Lt := Pos.compare_lt_iff (compat "8.3"). -Notation Pcompare_Lt_Gt := Pos.compare_cont_Lt_Gt (compat "8.3"). - -Notation Pcompare_antisym := Pos.compare_cont_antisym (compat "8.3"). -Notation ZC1 := Pos.gt_lt (compat "8.3"). -Notation ZC2 := Pos.lt_gt (compat "8.3"). -Notation Pcompare_spec := Pos.compare_spec (compat "8.3"). -Notation Pcompare_p_Sp := Pos.lt_succ_diag_r (compat "8.3"). -Notation Pcompare_succ_succ := Pos.compare_succ_succ (compat "8.3"). -Notation Pcompare_1 := Pos.nlt_1_r (compat "8.3"). -Notation Plt_1 := Pos.nlt_1_r (compat "8.3"). -Notation Plt_1_succ := Pos.lt_1_succ (compat "8.3"). -Notation Plt_lt_succ := Pos.lt_lt_succ (compat "8.3"). -Notation Plt_irrefl := Pos.lt_irrefl (compat "8.3"). -Notation Plt_trans := Pos.lt_trans (compat "8.3"). -Notation Plt_ind := Pos.lt_ind (compat "8.3"). -Notation Ple_lteq := Pos.le_lteq (compat "8.3"). -Notation Ple_refl := Pos.le_refl (compat "8.3"). -Notation Ple_lt_trans := Pos.le_lt_trans (compat "8.3"). -Notation Plt_le_trans := Pos.lt_le_trans (compat "8.3"). -Notation Ple_trans := Pos.le_trans (compat "8.3"). -Notation Plt_succ_r := Pos.lt_succ_r (compat "8.3"). -Notation Ple_succ_l := Pos.le_succ_l (compat "8.3"). -Notation Pplus_compare_mono_l := Pos.add_compare_mono_l (compat "8.3"). -Notation Pplus_compare_mono_r := Pos.add_compare_mono_r (compat "8.3"). -Notation Pplus_lt_mono_l := Pos.add_lt_mono_l (compat "8.3"). -Notation Pplus_lt_mono_r := Pos.add_lt_mono_r (compat "8.3"). -Notation Pplus_lt_mono := Pos.add_lt_mono (compat "8.3"). -Notation Pplus_le_mono_l := Pos.add_le_mono_l (compat "8.3"). -Notation Pplus_le_mono_r := Pos.add_le_mono_r (compat "8.3"). -Notation Pplus_le_mono := Pos.add_le_mono (compat "8.3"). -Notation Pmult_compare_mono_l := Pos.mul_compare_mono_l (compat "8.3"). -Notation Pmult_compare_mono_r := Pos.mul_compare_mono_r (compat "8.3"). -Notation Pmult_lt_mono_l := Pos.mul_lt_mono_l (compat "8.3"). -Notation Pmult_lt_mono_r := Pos.mul_lt_mono_r (compat "8.3"). -Notation Pmult_lt_mono := Pos.mul_lt_mono (compat "8.3"). -Notation Pmult_le_mono_l := Pos.mul_le_mono_l (compat "8.3"). -Notation Pmult_le_mono_r := Pos.mul_le_mono_r (compat "8.3"). -Notation Pmult_le_mono := Pos.mul_le_mono (compat "8.3"). -Notation Plt_plus_r := Pos.lt_add_r (compat "8.3"). -Notation Plt_not_plus_l := Pos.lt_not_add_l (compat "8.3"). -Notation Ppow_gt_1 := Pos.pow_gt_1 (compat "8.3"). -Notation Ppred_mask := Pos.pred_mask (compat "8.3"). -Notation Pminus_mask_succ_r := Pos.sub_mask_succ_r (compat "8.3"). -Notation Pminus_mask_carry_spec := Pos.sub_mask_carry_spec (compat "8.3"). -Notation Pminus_succ_r := Pos.sub_succ_r (compat "8.3"). -Notation Pminus_mask_diag := Pos.sub_mask_diag (compat "8.3"). - -Notation Pplus_minus_eq := Pos.add_sub (compat "8.3"). -Notation Pmult_minus_distr_l := Pos.mul_sub_distr_l (compat "8.3"). -Notation Pminus_lt_mono_l := Pos.sub_lt_mono_l (compat "8.3"). -Notation Pminus_compare_mono_l := Pos.sub_compare_mono_l (compat "8.3"). -Notation Pminus_compare_mono_r := Pos.sub_compare_mono_r (compat "8.3"). -Notation Pminus_lt_mono_r := Pos.sub_lt_mono_r (compat "8.3"). -Notation Pminus_decr := Pos.sub_decr (compat "8.3"). -Notation Pminus_xI_xI := Pos.sub_xI_xI (compat "8.3"). -Notation Pplus_minus_assoc := Pos.add_sub_assoc (compat "8.3"). -Notation Pminus_plus_distr := Pos.sub_add_distr (compat "8.3"). -Notation Pminus_minus_distr := Pos.sub_sub_distr (compat "8.3"). -Notation Pminus_mask_Lt := Pos.sub_mask_neg (compat "8.3"). -Notation Pminus_Lt := Pos.sub_lt (compat "8.3"). -Notation Pminus_Eq := Pos.sub_diag (compat "8.3"). -Notation Psize_monotone := Pos.size_nat_monotone (compat "8.3"). -Notation Psize_pos_gt := Pos.size_gt (compat "8.3"). -Notation Psize_pos_le := Pos.size_le (compat "8.3"). + Pos.pred_double_xO_discr (only parsing). +Notation Psucc_not_one := Pos.succ_not_1 (only parsing). +Notation Ppred_succ := Pos.pred_succ (compat "8.6"). +Notation Psucc_pred := Pos.succ_pred_or (only parsing). +Notation Psucc_inj := Pos.succ_inj (compat "8.6"). +Notation Pplus_carry_spec := Pos.add_carry_spec (only parsing). +Notation Pplus_comm := Pos.add_comm (only parsing). +Notation Pplus_succ_permute_r := Pos.add_succ_r (only parsing). +Notation Pplus_succ_permute_l := Pos.add_succ_l (only parsing). +Notation Pplus_no_neutral := Pos.add_no_neutral (only parsing). +Notation Pplus_carry_plus := Pos.add_carry_add (only parsing). +Notation Pplus_reg_r := Pos.add_reg_r (only parsing). +Notation Pplus_reg_l := Pos.add_reg_l (only parsing). +Notation Pplus_carry_reg_r := Pos.add_carry_reg_r (only parsing). +Notation Pplus_carry_reg_l := Pos.add_carry_reg_l (only parsing). +Notation Pplus_assoc := Pos.add_assoc (only parsing). +Notation Pplus_xO := Pos.add_xO (only parsing). +Notation Pplus_xI_double_minus_one := Pos.add_xI_pred_double (only parsing). +Notation Pplus_xO_double_minus_one := Pos.add_xO_pred_double (only parsing). +Notation Pplus_diag := Pos.add_diag (only parsing). +Notation PeanoView := Pos.PeanoView (only parsing). +Notation PeanoOne := Pos.PeanoOne (only parsing). +Notation PeanoSucc := Pos.PeanoSucc (only parsing). +Notation PeanoView_rect := Pos.PeanoView_rect (only parsing). +Notation PeanoView_ind := Pos.PeanoView_ind (only parsing). +Notation PeanoView_rec := Pos.PeanoView_rec (only parsing). +Notation peanoView_xO := Pos.peanoView_xO (only parsing). +Notation peanoView_xI := Pos.peanoView_xI (only parsing). +Notation peanoView := Pos.peanoView (only parsing). +Notation PeanoView_iter := Pos.PeanoView_iter (only parsing). +Notation eq_dep_eq_positive := Pos.eq_dep_eq_positive (only parsing). +Notation PeanoViewUnique := Pos.PeanoViewUnique (only parsing). +Notation Prect := Pos.peano_rect (only parsing). +Notation Prect_succ := Pos.peano_rect_succ (only parsing). +Notation Prect_base := Pos.peano_rect_base (only parsing). +Notation Prec := Pos.peano_rec (only parsing). +Notation Pind := Pos.peano_ind (only parsing). +Notation Pcase := Pos.peano_case (only parsing). +Notation Pmult_1_r := Pos.mul_1_r (only parsing). +Notation Pmult_Sn_m := Pos.mul_succ_l (only parsing). +Notation Pmult_xO_permute_r := Pos.mul_xO_r (only parsing). +Notation Pmult_xI_permute_r := Pos.mul_xI_r (only parsing). +Notation Pmult_comm := Pos.mul_comm (only parsing). +Notation Pmult_plus_distr_l := Pos.mul_add_distr_l (only parsing). +Notation Pmult_plus_distr_r := Pos.mul_add_distr_r (only parsing). +Notation Pmult_assoc := Pos.mul_assoc (only parsing). +Notation Pmult_xI_mult_xO_discr := Pos.mul_xI_mul_xO_discr (only parsing). +Notation Pmult_xO_discr := Pos.mul_xO_discr (only parsing). +Notation Pmult_reg_r := Pos.mul_reg_r (only parsing). +Notation Pmult_reg_l := Pos.mul_reg_l (only parsing). +Notation Pmult_1_inversion_l := Pos.mul_eq_1_l (only parsing). +Notation Psquare_xO := Pos.square_xO (compat "8.6"). +Notation Psquare_xI := Pos.square_xI (compat "8.6"). +Notation iter_pos_swap_gen := Pos.iter_swap_gen (only parsing). +Notation iter_pos_swap := Pos.iter_swap (only parsing). +Notation iter_pos_succ := Pos.iter_succ (only parsing). +Notation iter_pos_plus := Pos.iter_add (only parsing). +Notation iter_pos_invariant := Pos.iter_invariant (only parsing). +Notation Ppow_1_r := Pos.pow_1_r (compat "8.6"). +Notation Ppow_succ_r := Pos.pow_succ_r (compat "8.6"). +Notation Peqb_refl := Pos.eqb_refl (compat "8.6"). +Notation Peqb_eq := Pos.eqb_eq (compat "8.6"). +Notation Pcompare_refl_id := Pos.compare_cont_refl (only parsing). +Notation Pcompare_eq_iff := Pos.compare_eq_iff (only parsing). +Notation Pcompare_Gt_Lt := Pos.compare_cont_Gt_Lt (only parsing). +Notation Pcompare_eq_Lt := Pos.compare_lt_iff (only parsing). +Notation Pcompare_Lt_Gt := Pos.compare_cont_Lt_Gt (only parsing). + +Notation Pcompare_antisym := Pos.compare_cont_antisym (only parsing). +Notation ZC1 := Pos.gt_lt (only parsing). +Notation ZC2 := Pos.lt_gt (only parsing). +Notation Pcompare_spec := Pos.compare_spec (compat "8.6"). +Notation Pcompare_p_Sp := Pos.lt_succ_diag_r (only parsing). +Notation Pcompare_succ_succ := Pos.compare_succ_succ (compat "8.6"). +Notation Pcompare_1 := Pos.nlt_1_r (only parsing). +Notation Plt_1 := Pos.nlt_1_r (only parsing). +Notation Plt_1_succ := Pos.lt_1_succ (compat "8.6"). +Notation Plt_lt_succ := Pos.lt_lt_succ (compat "8.6"). +Notation Plt_irrefl := Pos.lt_irrefl (compat "8.6"). +Notation Plt_trans := Pos.lt_trans (compat "8.6"). +Notation Plt_ind := Pos.lt_ind (compat "8.6"). +Notation Ple_lteq := Pos.le_lteq (compat "8.6"). +Notation Ple_refl := Pos.le_refl (compat "8.6"). +Notation Ple_lt_trans := Pos.le_lt_trans (compat "8.6"). +Notation Plt_le_trans := Pos.lt_le_trans (compat "8.6"). +Notation Ple_trans := Pos.le_trans (compat "8.6"). +Notation Plt_succ_r := Pos.lt_succ_r (compat "8.6"). +Notation Ple_succ_l := Pos.le_succ_l (compat "8.6"). +Notation Pplus_compare_mono_l := Pos.add_compare_mono_l (only parsing). +Notation Pplus_compare_mono_r := Pos.add_compare_mono_r (only parsing). +Notation Pplus_lt_mono_l := Pos.add_lt_mono_l (only parsing). +Notation Pplus_lt_mono_r := Pos.add_lt_mono_r (only parsing). +Notation Pplus_lt_mono := Pos.add_lt_mono (only parsing). +Notation Pplus_le_mono_l := Pos.add_le_mono_l (only parsing). +Notation Pplus_le_mono_r := Pos.add_le_mono_r (only parsing). +Notation Pplus_le_mono := Pos.add_le_mono (only parsing). +Notation Pmult_compare_mono_l := Pos.mul_compare_mono_l (only parsing). +Notation Pmult_compare_mono_r := Pos.mul_compare_mono_r (only parsing). +Notation Pmult_lt_mono_l := Pos.mul_lt_mono_l (only parsing). +Notation Pmult_lt_mono_r := Pos.mul_lt_mono_r (only parsing). +Notation Pmult_lt_mono := Pos.mul_lt_mono (only parsing). +Notation Pmult_le_mono_l := Pos.mul_le_mono_l (only parsing). +Notation Pmult_le_mono_r := Pos.mul_le_mono_r (only parsing). +Notation Pmult_le_mono := Pos.mul_le_mono (only parsing). +Notation Plt_plus_r := Pos.lt_add_r (only parsing). +Notation Plt_not_plus_l := Pos.lt_not_add_l (only parsing). +Notation Ppow_gt_1 := Pos.pow_gt_1 (compat "8.6"). +Notation Ppred_mask := Pos.pred_mask (compat "8.6"). +Notation Pminus_mask_succ_r := Pos.sub_mask_succ_r (only parsing). +Notation Pminus_mask_carry_spec := Pos.sub_mask_carry_spec (only parsing). +Notation Pminus_succ_r := Pos.sub_succ_r (only parsing). +Notation Pminus_mask_diag := Pos.sub_mask_diag (only parsing). + +Notation Pplus_minus_eq := Pos.add_sub (only parsing). +Notation Pmult_minus_distr_l := Pos.mul_sub_distr_l (only parsing). +Notation Pminus_lt_mono_l := Pos.sub_lt_mono_l (only parsing). +Notation Pminus_compare_mono_l := Pos.sub_compare_mono_l (only parsing). +Notation Pminus_compare_mono_r := Pos.sub_compare_mono_r (only parsing). +Notation Pminus_lt_mono_r := Pos.sub_lt_mono_r (only parsing). +Notation Pminus_decr := Pos.sub_decr (only parsing). +Notation Pminus_xI_xI := Pos.sub_xI_xI (only parsing). +Notation Pplus_minus_assoc := Pos.add_sub_assoc (only parsing). +Notation Pminus_plus_distr := Pos.sub_add_distr (only parsing). +Notation Pminus_minus_distr := Pos.sub_sub_distr (only parsing). +Notation Pminus_mask_Lt := Pos.sub_mask_neg (only parsing). +Notation Pminus_Lt := Pos.sub_lt (only parsing). +Notation Pminus_Eq := Pos.sub_diag (only parsing). +Notation Psize_monotone := Pos.size_nat_monotone (only parsing). +Notation Psize_pos_gt := Pos.size_gt (only parsing). +Notation Psize_pos_le := Pos.size_le (only parsing). (** More complex compatibility facts, expressed as lemmas (to preserve scopes for instance) *) diff --git a/theories/PArith/BinPosDef.v b/theories/PArith/BinPosDef.v index 2b647555c..070314746 100644 --- a/theories/PArith/BinPosDef.v +++ b/theories/PArith/BinPosDef.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (**********************************************************************) @@ -557,4 +559,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/PArith/PArith.v b/theories/PArith/PArith.v index 66e1ae152..2be3d07cc 100644 --- a/theories/PArith/PArith.v +++ b/theories/PArith/PArith.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Library for positive natural numbers *) diff --git a/theories/PArith/POrderedType.v b/theories/PArith/POrderedType.v index b73ddff82..c454e8afd 100644 --- a/theories/PArith/POrderedType.v +++ b/theories/PArith/POrderedType.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinPos Equalities Orders OrdersTac. diff --git a/theories/PArith/Pnat.v b/theories/PArith/Pnat.v index 461967de8..26aba87fb 100644 --- a/theories/PArith/Pnat.v +++ b/theories/PArith/Pnat.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinPos PeanoNat. @@ -382,36 +384,36 @@ End SuccNat2Pos. (** For compatibility, old names and old-style lemmas *) -Notation Psucc_S := Pos2Nat.inj_succ (compat "8.3"). -Notation Pplus_plus := Pos2Nat.inj_add (compat "8.3"). -Notation Pmult_mult := Pos2Nat.inj_mul (compat "8.3"). -Notation Pcompare_nat_compare := Pos2Nat.inj_compare (compat "8.3"). -Notation nat_of_P_xH := Pos2Nat.inj_1 (compat "8.3"). -Notation nat_of_P_xO := Pos2Nat.inj_xO (compat "8.3"). -Notation nat_of_P_xI := Pos2Nat.inj_xI (compat "8.3"). -Notation nat_of_P_is_S := Pos2Nat.is_succ (compat "8.3"). -Notation nat_of_P_pos := Pos2Nat.is_pos (compat "8.3"). -Notation nat_of_P_inj_iff := Pos2Nat.inj_iff (compat "8.3"). -Notation nat_of_P_inj := Pos2Nat.inj (compat "8.3"). -Notation Plt_lt := Pos2Nat.inj_lt (compat "8.3"). -Notation Pgt_gt := Pos2Nat.inj_gt (compat "8.3"). -Notation Ple_le := Pos2Nat.inj_le (compat "8.3"). -Notation Pge_ge := Pos2Nat.inj_ge (compat "8.3"). -Notation Pminus_minus := Pos2Nat.inj_sub (compat "8.3"). -Notation iter_nat_of_P := @Pos2Nat.inj_iter (compat "8.3"). - -Notation nat_of_P_of_succ_nat := SuccNat2Pos.id_succ (compat "8.3"). -Notation P_of_succ_nat_of_P := Pos2SuccNat.id_succ (compat "8.3"). - -Notation nat_of_P_succ_morphism := Pos2Nat.inj_succ (compat "8.3"). -Notation nat_of_P_plus_morphism := Pos2Nat.inj_add (compat "8.3"). -Notation nat_of_P_mult_morphism := Pos2Nat.inj_mul (compat "8.3"). -Notation nat_of_P_compare_morphism := Pos2Nat.inj_compare (compat "8.3"). -Notation lt_O_nat_of_P := Pos2Nat.is_pos (compat "8.3"). -Notation ZL4 := Pos2Nat.is_succ (compat "8.3"). -Notation nat_of_P_o_P_of_succ_nat_eq_succ := SuccNat2Pos.id_succ (compat "8.3"). -Notation P_of_succ_nat_o_nat_of_P_eq_succ := Pos2SuccNat.id_succ (compat "8.3"). -Notation pred_o_P_of_succ_nat_o_nat_of_P_eq_id := Pos2SuccNat.pred_id (compat "8.3"). +Notation Psucc_S := Pos2Nat.inj_succ (only parsing). +Notation Pplus_plus := Pos2Nat.inj_add (only parsing). +Notation Pmult_mult := Pos2Nat.inj_mul (only parsing). +Notation Pcompare_nat_compare := Pos2Nat.inj_compare (only parsing). +Notation nat_of_P_xH := Pos2Nat.inj_1 (only parsing). +Notation nat_of_P_xO := Pos2Nat.inj_xO (only parsing). +Notation nat_of_P_xI := Pos2Nat.inj_xI (only parsing). +Notation nat_of_P_is_S := Pos2Nat.is_succ (only parsing). +Notation nat_of_P_pos := Pos2Nat.is_pos (only parsing). +Notation nat_of_P_inj_iff := Pos2Nat.inj_iff (only parsing). +Notation nat_of_P_inj := Pos2Nat.inj (only parsing). +Notation Plt_lt := Pos2Nat.inj_lt (only parsing). +Notation Pgt_gt := Pos2Nat.inj_gt (only parsing). +Notation Ple_le := Pos2Nat.inj_le (only parsing). +Notation Pge_ge := Pos2Nat.inj_ge (only parsing). +Notation Pminus_minus := Pos2Nat.inj_sub (only parsing). +Notation iter_nat_of_P := @Pos2Nat.inj_iter (only parsing). + +Notation nat_of_P_of_succ_nat := SuccNat2Pos.id_succ (only parsing). +Notation P_of_succ_nat_of_P := Pos2SuccNat.id_succ (only parsing). + +Notation nat_of_P_succ_morphism := Pos2Nat.inj_succ (only parsing). +Notation nat_of_P_plus_morphism := Pos2Nat.inj_add (only parsing). +Notation nat_of_P_mult_morphism := Pos2Nat.inj_mul (only parsing). +Notation nat_of_P_compare_morphism := Pos2Nat.inj_compare (only parsing). +Notation lt_O_nat_of_P := Pos2Nat.is_pos (only parsing). +Notation ZL4 := Pos2Nat.is_succ (only parsing). +Notation nat_of_P_o_P_of_succ_nat_eq_succ := SuccNat2Pos.id_succ (only parsing). +Notation P_of_succ_nat_o_nat_of_P_eq_succ := Pos2SuccNat.id_succ (only parsing). +Notation pred_o_P_of_succ_nat_o_nat_of_P_eq_id := Pos2SuccNat.pred_id (only parsing). Lemma nat_of_P_minus_morphism p q : Pos.compare_cont Eq p q = Gt -> diff --git a/theories/Program/Basics.v b/theories/Program/Basics.v index ff0d5b91b..f55093ed4 100644 --- a/theories/Program/Basics.v +++ b/theories/Program/Basics.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Standard functions and combinators. diff --git a/theories/Program/Combinators.v b/theories/Program/Combinators.v index 237d878bf..f78d06b1d 100644 --- a/theories/Program/Combinators.v +++ b/theories/Program/Combinators.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Proofs about standard combinators, exports functional extensionality. diff --git a/theories/Program/Equality.v b/theories/Program/Equality.v index 5e3d0b1a7..cf42ed18d 100644 --- a/theories/Program/Equality.v +++ b/theories/Program/Equality.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Tactics related to (dependent) equality and proof irrelevance. *) diff --git a/theories/Program/Program.v b/theories/Program/Program.v index be8bb26d3..de0a6d5d6 100644 --- a/theories/Program/Program.v +++ b/theories/Program/Program.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Coq.Program.Utils. Require Export Coq.Program.Wf. diff --git a/theories/Program/Subset.v b/theories/Program/Subset.v index c68be0d22..1c89b6c3b 100644 --- a/theories/Program/Subset.v +++ b/theories/Program/Subset.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Tactics related to subsets and proof irrelevance. *) diff --git a/theories/Program/Syntax.v b/theories/Program/Syntax.v index a62348716..785b9437e 100644 --- a/theories/Program/Syntax.v +++ b/theories/Program/Syntax.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Custom notations and implicits for Coq prelude definitions. diff --git a/theories/Program/Tactics.v b/theories/Program/Tactics.v index b06562fc4..bc8388184 100644 --- a/theories/Program/Tactics.v +++ b/theories/Program/Tactics.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module implements various tactics used to simplify the goals produced by Program, diff --git a/theories/Program/Utils.v b/theories/Program/Utils.v index 8d7548803..78c36dc7d 100644 --- a/theories/Program/Utils.v +++ b/theories/Program/Utils.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Various syntactic shorthands that are useful with [Program]. *) diff --git a/theories/Program/Wf.v b/theories/Program/Wf.v index da9020bc1..627879854 100644 --- a/theories/Program/Wf.v +++ b/theories/Program/Wf.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Reformulation of the Wf module using subsets where possible, providing the support for [Program]'s treatment of well-founded definitions. *) diff --git a/theories/QArith/QArith.v b/theories/QArith/QArith.v index 439006abf..813900822 100644 --- a/theories/QArith/QArith.v +++ b/theories/QArith/QArith.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export QArith_base. diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v index 5996d30f2..467f263be 100644 --- a/theories/QArith/QArith_base.v +++ b/theories/QArith/QArith_base.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export ZArith. diff --git a/theories/QArith/QOrderedType.v b/theories/QArith/QOrderedType.v index cf18ed896..37b4b298a 100644 --- a/theories/QArith/QOrderedType.v +++ b/theories/QArith/QOrderedType.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 QArith_base Equalities Orders OrdersTac. diff --git a/theories/QArith/Qabs.v b/theories/QArith/Qabs.v index ec2ac7832..48be89417 100644 --- a/theories/QArith/Qabs.v +++ b/theories/QArith/Qabs.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export QArith. diff --git a/theories/QArith/Qcabs.v b/theories/QArith/Qcabs.v index 09908665e..f45868a77 100644 --- a/theories/QArith/Qcabs.v +++ b/theories/QArith/Qcabs.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * An absolute value for normalized rational numbers. *) diff --git a/theories/QArith/Qcanon.v b/theories/QArith/Qcanon.v index 0b399febe..e25f69c31 100644 --- a/theories/QArith/Qcanon.v +++ b/theories/QArith/Qcanon.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Field. diff --git a/theories/QArith/Qfield.v b/theories/QArith/Qfield.v index bb1bb2079..6cbb491b8 100644 --- a/theories/QArith/Qfield.v +++ b/theories/QArith/Qfield.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Field. diff --git a/theories/QArith/Qminmax.v b/theories/QArith/Qminmax.v index 254c5b57f..264b2f928 100644 --- a/theories/QArith/Qminmax.v +++ b/theories/QArith/Qminmax.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 QArith_base Orders QOrderedType GenericMinMax. diff --git a/theories/QArith/Qpower.v b/theories/QArith/Qpower.v index 187bef2ab..3fd78f092 100644 --- a/theories/QArith/Qpower.v +++ b/theories/QArith/Qpower.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Zpow_facts Qfield Qreduction. diff --git a/theories/QArith/Qreals.v b/theories/QArith/Qreals.v index e6f7e7ac9..14ab1700e 100644 --- a/theories/QArith/Qreals.v +++ b/theories/QArith/Qreals.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Rbase. diff --git a/theories/QArith/Qreduction.v b/theories/QArith/Qreduction.v index 5d055b547..7b08515d2 100644 --- a/theories/QArith/Qreduction.v +++ b/theories/QArith/Qreduction.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Normalisation functions for rational numbers. *) @@ -11,8 +13,8 @@ Require Export QArith_base. Require Import Znumtheory. -Notation Z2P := Z.to_pos (compat "8.3"). -Notation Z2P_correct := Z2Pos.id (compat "8.3"). +Notation Z2P := Z.to_pos (only parsing). +Notation Z2P_correct := Z2Pos.id (only parsing). (** Simplification of fractions using [Z.gcd]. This version can compute within Coq. *) diff --git a/theories/QArith/Qring.v b/theories/QArith/Qring.v index 9569348b9..7f972d568 100644 --- a/theories/QArith/Qring.v +++ b/theories/QArith/Qring.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Qfield. diff --git a/theories/QArith/Qround.v b/theories/QArith/Qround.v index 6e72dd2c2..e4e974972 100644 --- a/theories/QArith/Qround.v +++ b/theories/QArith/Qround.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 QArith. diff --git a/theories/Reals/Alembert.v b/theories/Reals/Alembert.v index 155bf977b..09aad1ecb 100644 --- a/theories/Reals/Alembert.v +++ b/theories/Reals/Alembert.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/AltSeries.v b/theories/Reals/AltSeries.v index 9e106f267..c17ad0cfa 100644 --- a/theories/Reals/AltSeries.v +++ b/theories/Reals/AltSeries.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/ArithProp.v b/theories/Reals/ArithProp.v index 4cdc035bb..37240eb74 100644 --- a/theories/Reals/ArithProp.v +++ b/theories/Reals/ArithProp.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Binomial.v b/theories/Reals/Binomial.v index d608a9359..271100a58 100644 --- a/theories/Reals/Binomial.v +++ b/theories/Reals/Binomial.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Cauchy_prod.v b/theories/Reals/Cauchy_prod.v index 61af31e34..306b09dc4 100644 --- a/theories/Reals/Cauchy_prod.v +++ b/theories/Reals/Cauchy_prod.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Cos_plus.v b/theories/Reals/Cos_plus.v index 194d6fda2..d046ecf1e 100644 --- a/theories/Reals/Cos_plus.v +++ b/theories/Reals/Cos_plus.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Cos_rel.v b/theories/Reals/Cos_rel.v index b5ae76939..f9919278d 100644 --- a/theories/Reals/Cos_rel.v +++ b/theories/Reals/Cos_rel.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/DiscrR.v b/theories/Reals/DiscrR.v index 5dc5269c7..f3bc2f22e 100644 --- a/theories/Reals/DiscrR.v +++ b/theories/Reals/DiscrR.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 RIneq. diff --git a/theories/Reals/Exp_prop.v b/theories/Reals/Exp_prop.v index 6666dc3a1..3de131eae 100644 --- a/theories/Reals/Exp_prop.v +++ b/theories/Reals/Exp_prop.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Integration.v b/theories/Reals/Integration.v index 3db55efed..1f4fd5764 100644 --- a/theories/Reals/Integration.v +++ b/theories/Reals/Integration.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export NewtonInt. diff --git a/theories/Reals/MVT.v b/theories/Reals/MVT.v index a4b3845e8..717df1b11 100644 --- a/theories/Reals/MVT.v +++ b/theories/Reals/MVT.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Machin.v b/theories/Reals/Machin.v index 6ed0658a0..cdf98cbde 100644 --- a/theories/Reals/Machin.v +++ b/theories/Reals/Machin.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Fourier. diff --git a/theories/Reals/NewtonInt.v b/theories/Reals/NewtonInt.v index 405296f72..66918eeed 100644 --- a/theories/Reals/NewtonInt.v +++ b/theories/Reals/NewtonInt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/PSeries_reg.v b/theories/Reals/PSeries_reg.v index dbef02e02..61d1b5afe 100644 --- a/theories/Reals/PSeries_reg.v +++ b/theories/Reals/PSeries_reg.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/PartSum.v b/theories/Reals/PartSum.v index 1376c5134..33feeac0c 100644 --- a/theories/Reals/PartSum.v +++ b/theories/Reals/PartSum.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/RIneq.v b/theories/Reals/RIneq.v index 7bcd2799a..59a104965 100644 --- a/theories/Reals/RIneq.v +++ b/theories/Reals/RIneq.v @@ -1,10 +1,12 @@ -(* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) +(************************************************************************) (************************************************************************) (*********************************************************) @@ -1611,6 +1613,9 @@ Proof. Qed. Hint Resolve mult_INR: real. +Lemma pow_INR (m n: nat) : INR (m ^ n) = pow (INR m) n. +Proof. now induction n as [|n IHn];[ | simpl; rewrite mult_INR, IHn]. Qed. + (*********) Lemma lt_0_INR : forall n:nat, (0 < n)%nat -> 0 < INR n. Proof. @@ -2024,7 +2029,7 @@ Qed. Lemma R_rm : ring_morph 0%R 1%R Rplus Rmult Rminus Ropp eq - 0%Z 1%Z Zplus Zmult Zminus Zopp Zeq_bool IZR. + 0%Z 1%Z Zplus Zmult Zminus Z.opp Zeq_bool IZR. Proof. constructor ; try easy. exact plus_IZR. diff --git a/theories/Reals/RList.v b/theories/Reals/RList.v index f739d1550..e12937c70 100644 --- a/theories/Reals/RList.v +++ b/theories/Reals/RList.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/ROrderedType.v b/theories/Reals/ROrderedType.v index e20652aad..ee65ee1d1 100644 --- a/theories/Reals/ROrderedType.v +++ b/theories/Reals/ROrderedType.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase Equalities Orders OrdersTac. diff --git a/theories/Reals/R_Ifp.v b/theories/Reals/R_Ifp.v index 5705eacbd..77e2a1e04 100644 --- a/theories/Reals/R_Ifp.v +++ b/theories/Reals/R_Ifp.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (**********************************************************) diff --git a/theories/Reals/R_sqr.v b/theories/Reals/R_sqr.v index 057a16976..a60bb7cf4 100644 --- a/theories/Reals/R_sqr.v +++ b/theories/Reals/R_sqr.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/R_sqrt.v b/theories/Reals/R_sqrt.v index 7a386bd2e..d4035fad6 100644 --- a/theories/Reals/R_sqrt.v +++ b/theories/Reals/R_sqrt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis.v b/theories/Reals/Ranalysis.v index 9b0357f03..4bde9b609 100644 --- a/theories/Reals/Ranalysis.v +++ b/theories/Reals/Ranalysis.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis1.v b/theories/Reals/Ranalysis1.v index 7f7344031..36ac738ca 100644 --- a/theories/Reals/Ranalysis1.v +++ b/theories/Reals/Ranalysis1.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis2.v b/theories/Reals/Ranalysis2.v index 04ee7a7bf..7a97ca63e 100644 --- a/theories/Reals/Ranalysis2.v +++ b/theories/Reals/Ranalysis2.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis3.v b/theories/Reals/Ranalysis3.v index 1fa2b6965..301d6d2c5 100644 --- a/theories/Reals/Ranalysis3.v +++ b/theories/Reals/Ranalysis3.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis4.v b/theories/Reals/Ranalysis4.v index 0ea698d8d..94f1757a8 100644 --- a/theories/Reals/Ranalysis4.v +++ b/theories/Reals/Ranalysis4.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ranalysis5.v b/theories/Reals/Ranalysis5.v index 61c0debb7..afb78e1c8 100644 --- a/theories/Reals/Ranalysis5.v +++ b/theories/Reals/Ranalysis5.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. @@ -27,46 +29,34 @@ Lemma f_incr_implies_g_incr_interv : forall f g:R->R, forall lb ub, (forall x , f lb <= x -> x <= f ub -> lb <= g x <= ub) -> (forall x y, f lb <= x -> x < y -> y <= f ub -> g x < g y). Proof. -intros f g lb ub lb_lt_ub f_incr f_eq_g g_ok x y lb_le_x x_lt_y y_le_ub. - assert (x_encad : f lb <= x <= f ub). - split ; [assumption | apply Rle_trans with (r2:=y) ; [apply Rlt_le|] ; assumption]. - assert (y_encad : f lb <= y <= f ub). - split ; [apply Rle_trans with (r2:=x) ; [|apply Rlt_le] ; assumption | assumption]. - assert (Temp1 : lb <= lb) by intuition ; assert (Temp2 : ub <= ub) by intuition. - assert (gx_encad := g_ok _ (proj1 x_encad) (proj2 x_encad)). - assert (gy_encad := g_ok _ (proj1 y_encad) (proj2 y_encad)). - clear Temp1 Temp2. - case (Rlt_dec (g x) (g y)). - intuition. + intros f g lb ub lb_lt_ub f_incr f_eq_g g_ok x y lb_le_x x_lt_y y_le_ub. + assert (x_encad : f lb <= x <= f ub) by lra. + assert (y_encad : f lb <= y <= f ub) by lra. + assert (gx_encad := g_ok _ (proj1 x_encad) (proj2 x_encad)). + assert (gy_encad := g_ok _ (proj1 y_encad) (proj2 y_encad)). + case (Rlt_dec (g x) (g y)); [ easy |]. intros Hfalse. - assert (Temp := Rnot_lt_le _ _ Hfalse). - assert (Hcontradiction : y <= x). - replace y with (id y) by intuition ; replace x with (id x) by intuition ; - rewrite <- f_eq_g. rewrite <- f_eq_g. - assert (f_incr2 : forall x y, lb <= x -> x <= y -> y < ub -> f x <= f y). + assert (Temp := Rnot_lt_le _ _ Hfalse). + enough (y <= x) by lra. + replace y with (id y) by easy. + replace x with (id x) by easy. + rewrite <- f_eq_g by easy. + rewrite <- f_eq_g by easy. + assert (f_incr2 : forall x y, lb <= x -> x <= y -> y < ub -> f x <= f y). { intros m n lb_le_m m_le_n n_lt_ub. case (m_le_n). - intros ; apply Rlt_le ; apply f_incr ; [| | apply Rlt_le] ; assumption. - intros Hyp ; rewrite Hyp ; apply Req_le ; reflexivity. - apply f_incr2. - intuition. intuition. - Focus 3. intuition. - Focus 2. intuition. - Focus 2. intuition. Focus 2. intuition. - assert (Temp2 : g x <> ub). - intro Hf. - assert (Htemp : (comp f g) x = f ub). - unfold comp ; rewrite Hf ; reflexivity. - rewrite f_eq_g in Htemp ; unfold id in Htemp. - assert (Htemp2 : x < f ub). - apply Rlt_le_trans with (r2:=y) ; intuition. - clear -Htemp Htemp2. fourier. - intuition. intuition. - clear -Temp2 gx_encad. - case (proj2 gx_encad). - intuition. - intro Hfalse ; apply False_ind ; apply Temp2 ; assumption. - apply False_ind. clear - Hcontradiction x_lt_y. fourier. + - intros; apply Rlt_le, f_incr, Rlt_le; assumption. + - intros Hyp; rewrite Hyp; apply Req_le; reflexivity. + } + apply f_incr2; intuition. + enough (g x <> ub) by lra. + intro Hf. + assert (Htemp : (comp f g) x = f ub). { + unfold comp; rewrite Hf; reflexivity. + } + rewrite f_eq_g in Htemp by easy. + unfold id in Htemp. + fourier. Qed. Lemma derivable_pt_id_interv : forall (lb ub x:R), diff --git a/theories/Reals/Ranalysis_reg.v b/theories/Reals/Ranalysis_reg.v index f5ebb4c53..e1d4781bf 100644 --- a/theories/Reals/Ranalysis_reg.v +++ b/theories/Reals/Ranalysis_reg.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Ratan.v b/theories/Reals/Ratan.v index 8c631dade..ce39d5ffe 100644 --- a/theories/Reals/Ratan.v +++ b/theories/Reals/Ratan.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Fourier. diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v index 947972bd5..6019d4faf 100644 --- a/theories/Reals/Raxioms.v +++ b/theories/Reals/Raxioms.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************) diff --git a/theories/Reals/Rbase.v b/theories/Reals/Rbase.v index 11d5a5b29..b63c8e1c6 100644 --- a/theories/Reals/Rbase.v +++ b/theories/Reals/Rbase.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Rdefinitions. diff --git a/theories/Reals/Rbasic_fun.v b/theories/Reals/Rbasic_fun.v index 17b3c5099..aa886cee0 100644 --- a/theories/Reals/Rbasic_fun.v +++ b/theories/Reals/Rbasic_fun.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************) @@ -609,7 +611,7 @@ Qed. Lemma Rabs_Zabs : forall z:Z, Rabs (IZR z) = IZR (Z.abs z). Proof. - intros z; case z; unfold Zabs. + intros z; case z; unfold Z.abs. apply Rabs_R0. now intros p0; apply Rabs_pos_eq, (IZR_le 0). unfold IZR at 1. diff --git a/theories/Reals/Rcomplete.v b/theories/Reals/Rcomplete.v index 783b81974..19cbbeca1 100644 --- a/theories/Reals/Rcomplete.v +++ b/theories/Reals/Rcomplete.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rdefinitions.v b/theories/Reals/Rdefinitions.v index c668a708e..857b4ec33 100644 --- a/theories/Reals/Rdefinitions.v +++ b/theories/Reals/Rdefinitions.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************) diff --git a/theories/Reals/Rderiv.v b/theories/Reals/Rderiv.v index 67a06e290..dfa5c7104 100644 --- a/theories/Reals/Rderiv.v +++ b/theories/Reals/Rderiv.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************) diff --git a/theories/Reals/Reals.v b/theories/Reals/Reals.v index 8c4a9727e..b249b519f 100644 --- a/theories/Reals/Reals.v +++ b/theories/Reals/Reals.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** The library REALS is divided in 6 parts : diff --git a/theories/Reals/Rfunctions.v b/theories/Reals/Rfunctions.v index c70ec42ef..77e531474 100644 --- a/theories/Reals/Rfunctions.v +++ b/theories/Reals/Rfunctions.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*i Some properties about pow and sum have been made with John Harrison i*) diff --git a/theories/Reals/Rgeom.v b/theories/Reals/Rgeom.v index a7002e959..6c2f3ac6e 100644 --- a/theories/Reals/Rgeom.v +++ b/theories/Reals/Rgeom.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/RiemannInt.v b/theories/Reals/RiemannInt.v index f8617b01e..f7d98fca8 100644 --- a/theories/Reals/RiemannInt.v +++ b/theories/Reals/RiemannInt.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rfunctions. diff --git a/theories/Reals/RiemannInt_SF.v b/theories/Reals/RiemannInt_SF.v index 0829dac52..ceac021ef 100644 --- a/theories/Reals/RiemannInt_SF.v +++ b/theories/Reals/RiemannInt_SF.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rlimit.v b/theories/Reals/Rlimit.v index d769593a4..b14fcc4d3 100644 --- a/theories/Reals/Rlimit.v +++ b/theories/Reals/Rlimit.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************) diff --git a/theories/Reals/Rlogic.v b/theories/Reals/Rlogic.v index 4ad3339ec..04f13477c 100644 --- a/theories/Reals/Rlogic.v +++ b/theories/Reals/Rlogic.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This module proves some logical properties of the axiomatic of Reals. @@ -63,7 +65,7 @@ destruct (Rle_lt_dec l 0) as [Hl|Hl]. now apply Rinv_0_lt_compat. now apply Hnp. left. -set (N := Zabs_nat (up (/l) - 2)). +set (N := Z.abs_nat (up (/l) - 2)). assert (H1l: (1 <= /l)%R). rewrite <- Rinv_1. apply Rinv_le_contravar with (1 := Hl). @@ -75,7 +77,7 @@ assert (HN: (INR N + 1 = IZR (up (/ l)) - 1)%R). rewrite inj_Zabs_nat. replace (IZR (up (/ l)) - 1)%R with (IZR (up (/ l) - 2) + 1)%R. apply (f_equal (fun v => IZR v + 1)%R). - apply Zabs_eq. + apply Z.abs_eq. apply Zle_minus_le_0. apply (Zlt_le_succ 1). apply lt_IZR. diff --git a/theories/Reals/Rminmax.v b/theories/Reals/Rminmax.v index 57e485cb7..7f73f7c18 100644 --- a/theories/Reals/Rminmax.v +++ b/theories/Reals/Rminmax.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Orders Rbase Rbasic_fun ROrderedType GenericMinMax. diff --git a/theories/Reals/Rpow_def.v b/theories/Reals/Rpow_def.v index 6279e1f16..0d9213031 100644 --- a/theories/Reals/Rpow_def.v +++ b/theories/Reals/Rpow_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rdefinitions. diff --git a/theories/Reals/Rpower.v b/theories/Reals/Rpower.v index a646104cd..c6fac951b 100644 --- a/theories/Reals/Rpower.v +++ b/theories/Reals/Rpower.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*i Due to L.Thery i*) @@ -431,9 +433,9 @@ Proof. Qed. Theorem Rpower_lt : - forall x y z:R, 1 < x -> 0 <= y -> y < z -> x ^R y < x ^R z. + forall x y z:R, 1 < x -> y < z -> x ^R y < x ^R z. Proof. - intros x y z H H0 H1. + intros x y z H H1. unfold Rpower. apply exp_increasing. apply Rmult_lt_compat_r. @@ -488,11 +490,13 @@ Proof. Qed. Theorem Rle_Rpower : - forall e n m:R, 1 < e -> 0 <= n -> n <= m -> e ^R n <= e ^R m. + forall e n m:R, 1 <= e -> n <= m -> e ^R n <= e ^R m. Proof. - intros e n m H H0 H1; case H1. - intros H2; left; apply Rpower_lt; assumption. - intros H2; rewrite H2; right; reflexivity. + intros e n m [H | H]; intros H1. + case H1. + intros H2; left; apply Rpower_lt; assumption. + intros H2; rewrite H2; right; reflexivity. + now rewrite <- H; unfold Rpower; rewrite ln_1, !Rmult_0_r; apply Rle_refl. Qed. Theorem ln_lt_2 : / 2 < ln 2. @@ -707,13 +711,18 @@ intros x y z x0 y0; unfold Rpower. rewrite <- exp_plus, ln_mult, Rmult_plus_distr_l; auto. Qed. -Lemma Rle_Rpower_l a b c: 0 <= c -> 0 < a <= b -> Rpower a c <= Rpower b c. +Lemma Rlt_Rpower_l a b c: 0 < c -> 0 < a < b -> a ^R c < b ^R c. +Proof. +intros c0 [a0 ab]; apply exp_increasing. +now apply Rmult_lt_compat_l; auto; apply ln_increasing; fourier. +Qed. + +Lemma Rle_Rpower_l a b c: 0 <= c -> 0 < a <= b -> a ^R c <= b ^R c. Proof. intros [c0 | c0]; [ | intros; rewrite <- c0, !Rpower_O; [apply Rle_refl | |] ]. intros [a0 [ab|ab]]. - left; apply exp_increasing. - now apply Rmult_lt_compat_l; auto; apply ln_increasing; fourier. + now apply Rlt_le, Rlt_Rpower_l;[ | split]; fourier. rewrite ab; apply Rle_refl. apply Rlt_le_trans with a; tauto. tauto. diff --git a/theories/Reals/Rprod.v b/theories/Reals/Rprod.v index 2cda84a58..17736af65 100644 --- a/theories/Reals/Rprod.v +++ b/theories/Reals/Rprod.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Compare. diff --git a/theories/Reals/Rseries.v b/theories/Reals/Rseries.v index 4ed943070..3521a476b 100644 --- a/theories/Reals/Rseries.v +++ b/theories/Reals/Rseries.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rsigma.v b/theories/Reals/Rsigma.v index 91b1a979a..83c60751d 100644 --- a/theories/Reals/Rsigma.v +++ b/theories/Reals/Rsigma.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rsqrt_def.v b/theories/Reals/Rsqrt_def.v index 584ba125d..6a3dd9765 100644 --- a/theories/Reals/Rsqrt_def.v +++ b/theories/Reals/Rsqrt_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Sumbool. diff --git a/theories/Reals/Rtopology.v b/theories/Reals/Rtopology.v index 91b742978..171dba552 100644 --- a/theories/Reals/Rtopology.v +++ b/theories/Reals/Rtopology.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo.v b/theories/Reals/Rtrigo.v index db1c46e9c..ffc0adf50 100644 --- a/theories/Reals/Rtrigo.v +++ b/theories/Reals/Rtrigo.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo1.v b/theories/Reals/Rtrigo1.v index 32b5cb694..bf00f736f 100644 --- a/theories/Reals/Rtrigo1.v +++ b/theories/Reals/Rtrigo1.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo_alt.v b/theories/Reals/Rtrigo_alt.v index f03ba549b..71b90fb45 100644 --- a/theories/Reals/Rtrigo_alt.v +++ b/theories/Reals/Rtrigo_alt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo_calc.v b/theories/Reals/Rtrigo_calc.v index 3cffaee6c..7cbfc6303 100644 --- a/theories/Reals/Rtrigo_calc.v +++ b/theories/Reals/Rtrigo_calc.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo_def.v b/theories/Reals/Rtrigo_def.v index 88b72f0df..d2faf95bc 100644 --- a/theories/Reals/Rtrigo_def.v +++ b/theories/Reals/Rtrigo_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase Rfunctions SeqSeries Rtrigo_fun Max. diff --git a/theories/Reals/Rtrigo_fun.v b/theories/Reals/Rtrigo_fun.v index 4b4423726..744a99a12 100644 --- a/theories/Reals/Rtrigo_fun.v +++ b/theories/Reals/Rtrigo_fun.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/Rtrigo_reg.v b/theories/Reals/Rtrigo_reg.v index a0ae73770..456fb6a71 100644 --- a/theories/Reals/Rtrigo_reg.v +++ b/theories/Reals/Rtrigo_reg.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/SeqProp.v b/theories/Reals/SeqProp.v index 62a954ce8..38b0b3c4b 100644 --- a/theories/Reals/SeqProp.v +++ b/theories/Reals/SeqProp.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/SeqSeries.v b/theories/Reals/SeqSeries.v index 667164d22..ccd205e23 100644 --- a/theories/Reals/SeqSeries.v +++ b/theories/Reals/SeqSeries.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Reals/SplitAbsolu.v b/theories/Reals/SplitAbsolu.v index 3dc6ca1f3..aa67b6774 100644 --- a/theories/Reals/SplitAbsolu.v +++ b/theories/Reals/SplitAbsolu.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbasic_fun. diff --git a/theories/Reals/SplitRmult.v b/theories/Reals/SplitRmult.v index ab656440f..a8ff60b07 100644 --- a/theories/Reals/SplitRmult.v +++ b/theories/Reals/SplitRmult.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*i Lemma mult_non_zero :(r1,r2:R)``r1<>0`` /\ ``r2<>0`` -> ``r1*r2<>0``. i*) diff --git a/theories/Reals/Sqrt_reg.v b/theories/Reals/Sqrt_reg.v index 04062fbbd..d6b386f10 100644 --- a/theories/Reals/Sqrt_reg.v +++ b/theories/Reals/Sqrt_reg.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Rbase. diff --git a/theories/Relations/Operators_Properties.v b/theories/Relations/Operators_Properties.v index e8d80f469..e82a67344 100644 --- a/theories/Relations/Operators_Properties.v +++ b/theories/Relations/Operators_Properties.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (************************************************************************) diff --git a/theories/Relations/Relation_Definitions.v b/theories/Relations/Relation_Definitions.v index 2198ab165..53def4741 100644 --- a/theories/Relations/Relation_Definitions.v +++ b/theories/Relations/Relation_Definitions.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) Section Relation_Definition. diff --git a/theories/Relations/Relation_Operators.v b/theories/Relations/Relation_Operators.v index fb6f11158..529e4d08e 100644 --- a/theories/Relations/Relation_Operators.v +++ b/theories/Relations/Relation_Operators.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (************************************************************************) diff --git a/theories/Relations/Relations.v b/theories/Relations/Relations.v index 403e69238..61344974e 100644 --- a/theories/Relations/Relations.v +++ b/theories/Relations/Relations.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Relation_Definitions. diff --git a/theories/Setoids/Setoid.v b/theories/Setoids/Setoid.v index 2bc0bc3b2..af06bcf47 100644 --- a/theories/Setoids/Setoid.v +++ b/theories/Setoids/Setoid.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Coq.Classes.SetoidTactics. diff --git a/theories/Sets/Classical_sets.v b/theories/Sets/Classical_sets.v index 68ef08e60..b68022f8f 100644 --- a/theories/Sets/Classical_sets.v +++ b/theories/Sets/Classical_sets.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Constructive_sets.v b/theories/Sets/Constructive_sets.v index 823b5cb84..f2475af12 100644 --- a/theories/Sets/Constructive_sets.v +++ b/theories/Sets/Constructive_sets.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Cpo.v b/theories/Sets/Cpo.v index a6208ebb2..3977097e8 100644 --- a/theories/Sets/Cpo.v +++ b/theories/Sets/Cpo.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Ensembles.v b/theories/Sets/Ensembles.v index cb340f260..c37132078 100644 --- a/theories/Sets/Ensembles.v +++ b/theories/Sets/Ensembles.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Finite_sets.v b/theories/Sets/Finite_sets.v index 8934a322a..5c42cbe67 100644 --- a/theories/Sets/Finite_sets.v +++ b/theories/Sets/Finite_sets.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Finite_sets_facts.v b/theories/Sets/Finite_sets_facts.v index 714de75ac..1c191613f 100644 --- a/theories/Sets/Finite_sets_facts.v +++ b/theories/Sets/Finite_sets_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Image.v b/theories/Sets/Image.v index 27ce2d4dc..3e28bbe91 100644 --- a/theories/Sets/Image.v +++ b/theories/Sets/Image.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Infinite_sets.v b/theories/Sets/Infinite_sets.v index 845861d70..bdeeb6a7c 100644 --- a/theories/Sets/Infinite_sets.v +++ b/theories/Sets/Infinite_sets.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Integers.v b/theories/Sets/Integers.v index 1471a44e1..225e388b1 100644 --- a/theories/Sets/Integers.v +++ b/theories/Sets/Integers.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Multiset.v b/theories/Sets/Multiset.v index 01805fe92..a79ddead2 100644 --- a/theories/Sets/Multiset.v +++ b/theories/Sets/Multiset.v @@ -1,14 +1,17 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* G. Huet 1-9-95 *) Require Import Permut Setoid. +Require Plus. (* comm. and ass. of plus *) Set Implicit Arguments. @@ -67,9 +70,6 @@ Section multiset_defs. unfold meq; unfold munion; simpl; auto. Qed. - - Require Plus. (* comm. and ass. of plus *) - Lemma munion_comm : forall x y:multiset, meq (munion x y) (munion y x). Proof. unfold meq; unfold multiplicity; unfold munion. diff --git a/theories/Sets/Partial_Order.v b/theories/Sets/Partial_Order.v index a9f136eff..17fc0ed25 100644 --- a/theories/Sets/Partial_Order.v +++ b/theories/Sets/Partial_Order.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Permut.v b/theories/Sets/Permut.v index 48d2248da..86a500dfd 100644 --- a/theories/Sets/Permut.v +++ b/theories/Sets/Permut.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* G. Huet 1-9-95 *) diff --git a/theories/Sets/Powerset.v b/theories/Sets/Powerset.v index a98b4aca2..88bcd6555 100644 --- a/theories/Sets/Powerset.v +++ b/theories/Sets/Powerset.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Powerset_Classical_facts.v b/theories/Sets/Powerset_Classical_facts.v index 66cc44fbd..7975a0269 100644 --- a/theories/Sets/Powerset_Classical_facts.v +++ b/theories/Sets/Powerset_Classical_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Powerset_facts.v b/theories/Sets/Powerset_facts.v index 209c22f71..81b475ac6 100644 --- a/theories/Sets/Powerset_facts.v +++ b/theories/Sets/Powerset_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_1.v b/theories/Sets/Relations_1.v index 7f1b03cf3..1ed745995 100644 --- a/theories/Sets/Relations_1.v +++ b/theories/Sets/Relations_1.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_1_facts.v b/theories/Sets/Relations_1_facts.v index 679525bf4..296ec42ad 100644 --- a/theories/Sets/Relations_1_facts.v +++ b/theories/Sets/Relations_1_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_2.v b/theories/Sets/Relations_2.v index 8d19eab8c..cc839506f 100644 --- a/theories/Sets/Relations_2.v +++ b/theories/Sets/Relations_2.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_2_facts.v b/theories/Sets/Relations_2_facts.v index 1fc5a2f10..36da36844 100644 --- a/theories/Sets/Relations_2_facts.v +++ b/theories/Sets/Relations_2_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_3.v b/theories/Sets/Relations_3.v index 292cdbb5d..de3212e2e 100644 --- a/theories/Sets/Relations_3.v +++ b/theories/Sets/Relations_3.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Relations_3_facts.v b/theories/Sets/Relations_3_facts.v index 59a987f15..0c1f670d0 100644 --- a/theories/Sets/Relations_3_facts.v +++ b/theories/Sets/Relations_3_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (****************************************************************************) (* *) diff --git a/theories/Sets/Uniset.v b/theories/Sets/Uniset.v index 8406c049e..7940bda1a 100644 --- a/theories/Sets/Uniset.v +++ b/theories/Sets/Uniset.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Sets as characteristic functions *) @@ -11,7 +13,7 @@ (* G. Huet 1-9-95 *) (* Updated Papageno 12/98 *) -Require Import Bool. +Require Import Bool Permut. Set Implicit Arguments. @@ -138,8 +140,6 @@ Hint Resolve seq_right. (** Here we should make uniset an abstract datatype, by hiding [Charac], [union], [charac]; all further properties are proved abstractly *) -Require Import Permut. - Lemma union_rotate : forall x y z:uniset, seq (union x (union y z)) (union z (union x y)). Proof. diff --git a/theories/Sorting/Heap.v b/theories/Sorting/Heap.v index eaadb440f..d9e5ad676 100644 --- a/theories/Sorting/Heap.v +++ b/theories/Sorting/Heap.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file is deprecated, for a tree on list, use [Mergesort.v]. *) @@ -134,7 +136,7 @@ Section defs. (munion (list_contents _ eqA_dec l1) (list_contents _ eqA_dec l2)) -> (forall a, HdRel leA a l1 -> HdRel leA a l2 -> HdRel leA a l) -> merge_lem l1 l2. - Require Import Morphisms. + Import Morphisms. Instance: Equivalence (@meq A). Proof. constructor; auto with datatypes. red. apply meq_trans. Defined. diff --git a/theories/Sorting/Mergesort.v b/theories/Sorting/Mergesort.v index 51e34f42d..824d000d2 100644 --- a/theories/Sorting/Mergesort.v +++ b/theories/Sorting/Mergesort.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** A modular implementation of mergesort (the complexity is O(n.log n) in diff --git a/theories/Sorting/PermutEq.v b/theories/Sorting/PermutEq.v index 90f64beec..97297f7ef 100644 --- a/theories/Sorting/PermutEq.v +++ b/theories/Sorting/PermutEq.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Relations Setoid SetoidList List Multiset PermutSetoid Permutation Omega. diff --git a/theories/Sorting/PermutSetoid.v b/theories/Sorting/PermutSetoid.v index 3ccf15d5e..08bc400f0 100644 --- a/theories/Sorting/PermutSetoid.v +++ b/theories/Sorting/PermutSetoid.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Omega Relations Multiset SetoidList. diff --git a/theories/Sorting/Permutation.v b/theories/Sorting/Permutation.v index 301548a1d..7b99b3626 100644 --- a/theories/Sorting/Permutation.v +++ b/theories/Sorting/Permutation.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (*********************************************************************) diff --git a/theories/Sorting/Sorted.v b/theories/Sorting/Sorted.v index 2e99f0d6b..89e9c7f3e 100644 --- a/theories/Sorting/Sorted.v +++ b/theories/Sorting/Sorted.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* Made by Hugo Herbelin *) diff --git a/theories/Sorting/Sorting.v b/theories/Sorting/Sorting.v index 1152c1118..c2be14616 100644 --- a/theories/Sorting/Sorting.v +++ b/theories/Sorting/Sorting.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Sorted. diff --git a/theories/Strings/Ascii.v b/theories/Strings/Ascii.v index 0f0e760c7..5154b75b3 100644 --- a/theories/Strings/Ascii.v +++ b/theories/Strings/Ascii.v @@ -1,10 +1,12 @@ -(* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) +(************************************************************************) (************************************************************************) (** Contributed by Laurent Théry (INRIA); diff --git a/theories/Strings/String.v b/theories/Strings/String.v index c39b47fb1..2be6618ad 100644 --- a/theories/Strings/String.v +++ b/theories/Strings/String.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Contributed by Laurent Théry (INRIA); @@ -163,6 +165,18 @@ intros n0 H; apply Rec; simpl; auto. apply Le.le_S_n; auto. Qed. +(** *** Concatenating lists of strings *) + +(** [concat sep sl] concatenates the list of strings [sl], inserting + the separator string [sep] between each. *) + +Fixpoint concat (sep : string) (ls : list string) := + match ls with + | nil => EmptyString + | cons x nil => x + | cons x xs => x ++ sep ++ concat sep xs + end. + (** *** Test functions *) (** Test if [s1] is a prefix of [s2] *) diff --git a/theories/Structures/DecidableType.v b/theories/Structures/DecidableType.v index d811f04ef..24333ad81 100644 --- a/theories/Structures/DecidableType.v +++ b/theories/Structures/DecidableType.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export SetoidList. Require Equalities. diff --git a/theories/Structures/DecidableTypeEx.v b/theories/Structures/DecidableTypeEx.v index 163a40f2e..8dd2e7103 100644 --- a/theories/Structures/DecidableTypeEx.v +++ b/theories/Structures/DecidableTypeEx.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 DecidableType OrderedType OrderedTypeEx. Set Implicit Arguments. diff --git a/theories/Structures/Equalities.v b/theories/Structures/Equalities.v index 747d03f8a..5f60a979c 100644 --- a/theories/Structures/Equalities.v +++ b/theories/Structures/Equalities.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export RelationClasses. Require Import Bool Morphisms Setoid. diff --git a/theories/Structures/EqualitiesFacts.v b/theories/Structures/EqualitiesFacts.v index cee3d63f0..7b6ee2eac 100644 --- a/theories/Structures/EqualitiesFacts.v +++ b/theories/Structures/EqualitiesFacts.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Equalities Bool SetoidList RelationPairs. diff --git a/theories/Structures/GenericMinMax.v b/theories/Structures/GenericMinMax.v index ac52d1bbb..05edc6ccd 100644 --- a/theories/Structures/GenericMinMax.v +++ b/theories/Structures/GenericMinMax.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Orders OrdersTac OrdersFacts Setoid Morphisms Basics. diff --git a/theories/Structures/OrderedType.v b/theories/Structures/OrderedType.v index 93ca383b2..f6fc247d5 100644 --- a/theories/Structures/OrderedType.v +++ b/theories/Structures/OrderedType.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export SetoidList Morphisms OrdersTac. Set Implicit Arguments. diff --git a/theories/Structures/OrderedTypeAlt.v b/theories/Structures/OrderedTypeAlt.v index b054496e9..278046a89 100644 --- a/theories/Structures/OrderedTypeAlt.v +++ b/theories/Structures/OrderedTypeAlt.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 OrderedType. (** * An alternative (but equivalent) presentation for an Ordered Type diff --git a/theories/Structures/OrderedTypeEx.v b/theories/Structures/OrderedTypeEx.v index 3c6afc7b2..f2a9a5691 100644 --- a/theories/Structures/OrderedTypeEx.v +++ b/theories/Structures/OrderedTypeEx.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 OrderedType. Require Import ZArith. @@ -55,7 +57,7 @@ Module Nat_as_OT <: UsualOrderedType. Definition compare x y : Compare lt eq x y. Proof. - case_eq (nat_compare x y); intro. + case_eq (Nat.compare x y); intro. - apply EQ. now apply nat_compare_eq. - apply LT. now apply nat_compare_Lt_lt. - apply GT. now apply nat_compare_Gt_gt. diff --git a/theories/Structures/Orders.v b/theories/Structures/Orders.v index 724690b42..42756ad33 100644 --- a/theories/Structures/Orders.v +++ b/theories/Structures/Orders.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export Relations Morphisms Setoid Equalities. Set Implicit Arguments. diff --git a/theories/Structures/OrdersAlt.v b/theories/Structures/OrdersAlt.v index 5dd917a71..ad6a38763 100644 --- a/theories/Structures/OrdersAlt.v +++ b/theories/Structures/OrdersAlt.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (* Finite sets library. * Authors: Pierre Letouzey and Jean-Christophe Filliâtre diff --git a/theories/Structures/OrdersEx.v b/theories/Structures/OrdersEx.v index 89c563882..93168f7df 100644 --- a/theories/Structures/OrdersEx.v +++ b/theories/Structures/OrdersEx.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (* Finite sets library. * Authors: Pierre Letouzey and Jean-Christophe Filliâtre diff --git a/theories/Structures/OrdersFacts.v b/theories/Structures/OrdersFacts.v index 0115d8a54..87df6b479 100644 --- a/theories/Structures/OrdersFacts.v +++ b/theories/Structures/OrdersFacts.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Bool Basics OrdersTac. Require Export Orders. diff --git a/theories/Structures/OrdersLists.v b/theories/Structures/OrdersLists.v index bf8529bc7..abdb9eff0 100644 --- a/theories/Structures/OrdersLists.v +++ b/theories/Structures/OrdersLists.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Export RelationPairs SetoidList Orders EqualitiesFacts. diff --git a/theories/Structures/OrdersTac.v b/theories/Structures/OrdersTac.v index 475a25a41..ebd8ee8fc 100644 --- a/theories/Structures/OrdersTac.v +++ b/theories/Structures/OrdersTac.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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 Setoid Morphisms Basics Equalities Orders. Set Implicit Arguments. diff --git a/theories/Unicode/Utf8.v b/theories/Unicode/Utf8.v index a8b512a8f..6cc9ea09d 100644 --- a/theories/Unicode/Utf8.v +++ b/theories/Unicode/Utf8.v @@ -1,10 +1,12 @@ (* -*- coding:utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Utf8_core. diff --git a/theories/Unicode/Utf8_core.v b/theories/Unicode/Utf8_core.v index a0545c0a4..5a8931a8c 100644 --- a/theories/Unicode/Utf8_core.v +++ b/theories/Unicode/Utf8_core.v @@ -1,19 +1,23 @@ (* -*- coding:utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (* 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 +29,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/Fin.v b/theories/Vectors/Fin.v index 2955184f6..4088843a1 100644 --- a/theories/Vectors/Fin.v +++ b/theories/Vectors/Fin.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Arith_base. diff --git a/theories/Vectors/Vector.v b/theories/Vectors/Vector.v index 19d749fc8..08158769f 100644 --- a/theories/Vectors/Vector.v +++ b/theories/Vectors/Vector.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Vectors. diff --git a/theories/Vectors/VectorDef.v b/theories/Vectors/VectorDef.v index c49451776..f6f3cafa2 100644 --- a/theories/Vectors/VectorDef.v +++ b/theories/Vectors/VectorDef.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Definitions of Vectors and functions to use them @@ -305,12 +307,10 @@ End VECTORLIST. Module VectorNotations. Delimit Scope vector_scope with vector. Notation "[ ]" := [] (format "[ ]") : vector_scope. -Notation "[]" := [] (compat "8.5") : vector_scope. Notation "h :: t" := (h :: t) (at level 60, right associativity) : vector_scope. Notation "[ x ]" := (x :: []) : vector_scope. Notation "[ x ; y ; .. ; z ]" := (cons _ x _ (cons _ y _ .. (cons _ z _ (nil _)) ..)) : vector_scope. -Notation "[ x ; .. ; y ]" := (cons _ x _ .. (cons _ y _ (nil _)) ..) (compat "8.4") : vector_scope. Notation "v [@ p ]" := (nth v p) (at level 1, format "v [@ p ]") : vector_scope. Open Scope vector_scope. End VectorNotations. diff --git a/theories/Vectors/VectorEq.v b/theories/Vectors/VectorEq.v index 04c570731..317f3f1c6 100644 --- a/theories/Vectors/VectorEq.v +++ b/theories/Vectors/VectorEq.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Equalities and Vector relations diff --git a/theories/Vectors/VectorSpec.v b/theories/Vectors/VectorSpec.v index 869d0fb5a..34dbaf36a 100644 --- a/theories/Vectors/VectorSpec.v +++ b/theories/Vectors/VectorSpec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Proofs of specification for functions defined over Vector diff --git a/theories/Wellfounded/Disjoint_Union.v b/theories/Wellfounded/Disjoint_Union.v index 693b0892a..10d902743 100644 --- a/theories/Wellfounded/Disjoint_Union.v +++ b/theories/Wellfounded/Disjoint_Union.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Cristina Cornes diff --git a/theories/Wellfounded/Inclusion.v b/theories/Wellfounded/Inclusion.v index 7fba5ef8d..ff233ef9c 100644 --- a/theories/Wellfounded/Inclusion.v +++ b/theories/Wellfounded/Inclusion.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Bruno Barras *) diff --git a/theories/Wellfounded/Inverse_Image.v b/theories/Wellfounded/Inverse_Image.v index b743edc2b..4a4ab87d9 100644 --- a/theories/Wellfounded/Inverse_Image.v +++ b/theories/Wellfounded/Inverse_Image.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Bruno Barras *) diff --git a/theories/Wellfounded/Lexicographic_Exponentiation.v b/theories/Wellfounded/Lexicographic_Exponentiation.v index 83a4926e4..684efeebe 100644 --- a/theories/Wellfounded/Lexicographic_Exponentiation.v +++ b/theories/Wellfounded/Lexicographic_Exponentiation.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Cristina Cornes diff --git a/theories/Wellfounded/Lexicographic_Product.v b/theories/Wellfounded/Lexicographic_Product.v index 0b09ada9f..37fd2fb23 100644 --- a/theories/Wellfounded/Lexicographic_Product.v +++ b/theories/Wellfounded/Lexicographic_Product.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Authors: Bruno Barras, Cristina Cornes *) diff --git a/theories/Wellfounded/Transitive_Closure.v b/theories/Wellfounded/Transitive_Closure.v index b7a8e63e5..59068623a 100644 --- a/theories/Wellfounded/Transitive_Closure.v +++ b/theories/Wellfounded/Transitive_Closure.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Bruno Barras *) diff --git a/theories/Wellfounded/Union.v b/theories/Wellfounded/Union.v index cf897afd7..9e671651f 100644 --- a/theories/Wellfounded/Union.v +++ b/theories/Wellfounded/Union.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Bruno Barras *) diff --git a/theories/Wellfounded/Well_Ordering.v b/theories/Wellfounded/Well_Ordering.v index aab864e83..fd363d02c 100644 --- a/theories/Wellfounded/Well_Ordering.v +++ b/theories/Wellfounded/Well_Ordering.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Author: Cristina Cornes. diff --git a/theories/Wellfounded/Wellfounded.v b/theories/Wellfounded/Wellfounded.v index 17b3138e4..bfe09e40b 100644 --- a/theories/Wellfounded/Wellfounded.v +++ b/theories/Wellfounded/Wellfounded.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export Disjoint_Union. diff --git a/theories/ZArith/BinInt.v b/theories/ZArith/BinInt.v index e6fd0f22e..cf7397b57 100644 --- a/theories/ZArith/BinInt.v +++ b/theories/ZArith/BinInt.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export BinNums BinPos Pnat. @@ -1565,94 +1567,94 @@ End Z2Pos. (** Compatibility Notations *) -Notation Zdouble_plus_one := Z.succ_double (compat "8.3"). -Notation Zdouble_minus_one := Z.pred_double (compat "8.3"). -Notation Zdouble := Z.double (compat "8.3"). -Notation ZPminus := Z.pos_sub (compat "8.3"). -Notation Zsucc' := Z.succ (compat "8.3"). -Notation Zpred' := Z.pred (compat "8.3"). -Notation Zplus' := Z.add (compat "8.3"). -Notation Zplus := Z.add (compat "8.3"). (* Slightly incompatible *) -Notation Zopp := Z.opp (compat "8.3"). -Notation Zsucc := Z.succ (compat "8.3"). -Notation Zpred := Z.pred (compat "8.3"). -Notation Zminus := Z.sub (compat "8.3"). -Notation Zmult := Z.mul (compat "8.3"). -Notation Zcompare := Z.compare (compat "8.3"). -Notation Zsgn := Z.sgn (compat "8.3"). -Notation Zle := Z.le (compat "8.3"). -Notation Zge := Z.ge (compat "8.3"). -Notation Zlt := Z.lt (compat "8.3"). -Notation Zgt := Z.gt (compat "8.3"). -Notation Zmax := Z.max (compat "8.3"). -Notation Zmin := Z.min (compat "8.3"). -Notation Zabs := Z.abs (compat "8.3"). -Notation Zabs_nat := Z.abs_nat (compat "8.3"). -Notation Zabs_N := Z.abs_N (compat "8.3"). -Notation Z_of_nat := Z.of_nat (compat "8.3"). -Notation Z_of_N := Z.of_N (compat "8.3"). - -Notation Zind := Z.peano_ind (compat "8.3"). -Notation Zopp_0 := Z.opp_0 (compat "8.3"). -Notation Zopp_involutive := Z.opp_involutive (compat "8.3"). -Notation Zopp_inj := Z.opp_inj (compat "8.3"). -Notation Zplus_0_l := Z.add_0_l (compat "8.3"). -Notation Zplus_0_r := Z.add_0_r (compat "8.3"). -Notation Zplus_comm := Z.add_comm (compat "8.3"). -Notation Zopp_plus_distr := Z.opp_add_distr (compat "8.3"). -Notation Zopp_succ := Z.opp_succ (compat "8.3"). -Notation Zplus_opp_r := Z.add_opp_diag_r (compat "8.3"). -Notation Zplus_opp_l := Z.add_opp_diag_l (compat "8.3"). -Notation Zplus_assoc := Z.add_assoc (compat "8.3"). -Notation Zplus_permute := Z.add_shuffle3 (compat "8.3"). -Notation Zplus_reg_l := Z.add_reg_l (compat "8.3"). -Notation Zplus_succ_l := Z.add_succ_l (compat "8.3"). -Notation Zplus_succ_comm := Z.add_succ_comm (compat "8.3"). -Notation Zsucc_discr := Z.neq_succ_diag_r (compat "8.3"). -Notation Zsucc_inj := Z.succ_inj (compat "8.3"). -Notation Zsucc'_inj := Z.succ_inj (compat "8.3"). -Notation Zsucc'_pred' := Z.succ_pred (compat "8.3"). -Notation Zpred'_succ' := Z.pred_succ (compat "8.3"). -Notation Zpred'_inj := Z.pred_inj (compat "8.3"). -Notation Zsucc'_discr := Z.neq_succ_diag_r (compat "8.3"). -Notation Zminus_0_r := Z.sub_0_r (compat "8.3"). -Notation Zminus_diag := Z.sub_diag (compat "8.3"). -Notation Zminus_plus_distr := Z.sub_add_distr (compat "8.3"). -Notation Zminus_succ_r := Z.sub_succ_r (compat "8.3"). -Notation Zminus_plus := Z.add_simpl_l (compat "8.3"). -Notation Zmult_0_l := Z.mul_0_l (compat "8.3"). -Notation Zmult_0_r := Z.mul_0_r (compat "8.3"). -Notation Zmult_1_l := Z.mul_1_l (compat "8.3"). -Notation Zmult_1_r := Z.mul_1_r (compat "8.3"). -Notation Zmult_comm := Z.mul_comm (compat "8.3"). -Notation Zmult_assoc := Z.mul_assoc (compat "8.3"). -Notation Zmult_permute := Z.mul_shuffle3 (compat "8.3"). -Notation Zmult_1_inversion_l := Z.mul_eq_1 (compat "8.3"). -Notation Zdouble_mult := Z.double_spec (compat "8.3"). -Notation Zdouble_plus_one_mult := Z.succ_double_spec (compat "8.3"). -Notation Zopp_mult_distr_l_reverse := Z.mul_opp_l (compat "8.3"). -Notation Zmult_opp_opp := Z.mul_opp_opp (compat "8.3"). -Notation Zmult_opp_comm := Z.mul_opp_comm (compat "8.3"). -Notation Zopp_eq_mult_neg_1 := Z.opp_eq_mul_m1 (compat "8.3"). -Notation Zmult_plus_distr_r := Z.mul_add_distr_l (compat "8.3"). -Notation Zmult_plus_distr_l := Z.mul_add_distr_r (compat "8.3"). -Notation Zmult_minus_distr_r := Z.mul_sub_distr_r (compat "8.3"). -Notation Zmult_reg_l := Z.mul_reg_l (compat "8.3"). -Notation Zmult_reg_r := Z.mul_reg_r (compat "8.3"). -Notation Zmult_succ_l := Z.mul_succ_l (compat "8.3"). -Notation Zmult_succ_r := Z.mul_succ_r (compat "8.3"). - -Notation Zpos_xI := Pos2Z.inj_xI (compat "8.3"). -Notation Zpos_xO := Pos2Z.inj_xO (compat "8.3"). -Notation Zneg_xI := Pos2Z.neg_xI (compat "8.3"). -Notation Zneg_xO := Pos2Z.neg_xO (compat "8.3"). -Notation Zopp_neg := Pos2Z.opp_neg (compat "8.3"). -Notation Zpos_succ_morphism := Pos2Z.inj_succ (compat "8.3"). -Notation Zpos_mult_morphism := Pos2Z.inj_mul (compat "8.3"). -Notation Zpos_minus_morphism := Pos2Z.inj_sub (compat "8.3"). -Notation Zpos_eq_rev := Pos2Z.inj (compat "8.3"). -Notation Zpos_plus_distr := Pos2Z.inj_add (compat "8.3"). -Notation Zneg_plus_distr := Pos2Z.add_neg_neg (compat "8.3"). +Notation Zdouble_plus_one := Z.succ_double (only parsing). +Notation Zdouble_minus_one := Z.pred_double (only parsing). +Notation Zdouble := Z.double (compat "8.6"). +Notation ZPminus := Z.pos_sub (only parsing). +Notation Zsucc' := Z.succ (compat "8.6"). +Notation Zpred' := Z.pred (compat "8.6"). +Notation Zplus' := Z.add (compat "8.6"). +Notation Zplus := Z.add (only parsing). (* Slightly incompatible *) +Notation Zopp := Z.opp (compat "8.6"). +Notation Zsucc := Z.succ (compat "8.6"). +Notation Zpred := Z.pred (compat "8.6"). +Notation Zminus := Z.sub (only parsing). +Notation Zmult := Z.mul (only parsing). +Notation Zcompare := Z.compare (compat "8.6"). +Notation Zsgn := Z.sgn (compat "8.6"). +Notation Zle := Z.le (compat "8.6"). +Notation Zge := Z.ge (compat "8.6"). +Notation Zlt := Z.lt (compat "8.6"). +Notation Zgt := Z.gt (compat "8.6"). +Notation Zmax := Z.max (compat "8.6"). +Notation Zmin := Z.min (compat "8.6"). +Notation Zabs := Z.abs (compat "8.6"). +Notation Zabs_nat := Z.abs_nat (compat "8.6"). +Notation Zabs_N := Z.abs_N (compat "8.6"). +Notation Z_of_nat := Z.of_nat (only parsing). +Notation Z_of_N := Z.of_N (only parsing). + +Notation Zind := Z.peano_ind (only parsing). +Notation Zopp_0 := Z.opp_0 (compat "8.6"). +Notation Zopp_involutive := Z.opp_involutive (compat "8.6"). +Notation Zopp_inj := Z.opp_inj (compat "8.6"). +Notation Zplus_0_l := Z.add_0_l (only parsing). +Notation Zplus_0_r := Z.add_0_r (only parsing). +Notation Zplus_comm := Z.add_comm (only parsing). +Notation Zopp_plus_distr := Z.opp_add_distr (only parsing). +Notation Zopp_succ := Z.opp_succ (compat "8.6"). +Notation Zplus_opp_r := Z.add_opp_diag_r (only parsing). +Notation Zplus_opp_l := Z.add_opp_diag_l (only parsing). +Notation Zplus_assoc := Z.add_assoc (only parsing). +Notation Zplus_permute := Z.add_shuffle3 (only parsing). +Notation Zplus_reg_l := Z.add_reg_l (only parsing). +Notation Zplus_succ_l := Z.add_succ_l (only parsing). +Notation Zplus_succ_comm := Z.add_succ_comm (only parsing). +Notation Zsucc_discr := Z.neq_succ_diag_r (only parsing). +Notation Zsucc_inj := Z.succ_inj (compat "8.6"). +Notation Zsucc'_inj := Z.succ_inj (compat "8.6"). +Notation Zsucc'_pred' := Z.succ_pred (compat "8.6"). +Notation Zpred'_succ' := Z.pred_succ (compat "8.6"). +Notation Zpred'_inj := Z.pred_inj (compat "8.6"). +Notation Zsucc'_discr := Z.neq_succ_diag_r (only parsing). +Notation Zminus_0_r := Z.sub_0_r (only parsing). +Notation Zminus_diag := Z.sub_diag (only parsing). +Notation Zminus_plus_distr := Z.sub_add_distr (only parsing). +Notation Zminus_succ_r := Z.sub_succ_r (only parsing). +Notation Zminus_plus := Z.add_simpl_l (only parsing). +Notation Zmult_0_l := Z.mul_0_l (only parsing). +Notation Zmult_0_r := Z.mul_0_r (only parsing). +Notation Zmult_1_l := Z.mul_1_l (only parsing). +Notation Zmult_1_r := Z.mul_1_r (only parsing). +Notation Zmult_comm := Z.mul_comm (only parsing). +Notation Zmult_assoc := Z.mul_assoc (only parsing). +Notation Zmult_permute := Z.mul_shuffle3 (only parsing). +Notation Zmult_1_inversion_l := Z.mul_eq_1 (only parsing). +Notation Zdouble_mult := Z.double_spec (only parsing). +Notation Zdouble_plus_one_mult := Z.succ_double_spec (only parsing). +Notation Zopp_mult_distr_l_reverse := Z.mul_opp_l (only parsing). +Notation Zmult_opp_opp := Z.mul_opp_opp (only parsing). +Notation Zmult_opp_comm := Z.mul_opp_comm (only parsing). +Notation Zopp_eq_mult_neg_1 := Z.opp_eq_mul_m1 (only parsing). +Notation Zmult_plus_distr_r := Z.mul_add_distr_l (only parsing). +Notation Zmult_plus_distr_l := Z.mul_add_distr_r (only parsing). +Notation Zmult_minus_distr_r := Z.mul_sub_distr_r (only parsing). +Notation Zmult_reg_l := Z.mul_reg_l (only parsing). +Notation Zmult_reg_r := Z.mul_reg_r (only parsing). +Notation Zmult_succ_l := Z.mul_succ_l (only parsing). +Notation Zmult_succ_r := Z.mul_succ_r (only parsing). + +Notation Zpos_xI := Pos2Z.inj_xI (only parsing). +Notation Zpos_xO := Pos2Z.inj_xO (only parsing). +Notation Zneg_xI := Pos2Z.neg_xI (only parsing). +Notation Zneg_xO := Pos2Z.neg_xO (only parsing). +Notation Zopp_neg := Pos2Z.opp_neg (only parsing). +Notation Zpos_succ_morphism := Pos2Z.inj_succ (only parsing). +Notation Zpos_mult_morphism := Pos2Z.inj_mul (only parsing). +Notation Zpos_minus_morphism := Pos2Z.inj_sub (only parsing). +Notation Zpos_eq_rev := Pos2Z.inj (only parsing). +Notation Zpos_plus_distr := Pos2Z.inj_add (only parsing). +Notation Zneg_plus_distr := Pos2Z.add_neg_neg (only parsing). Notation Z := Z (only parsing). Notation Z_rect := Z_rect (only parsing). diff --git a/theories/ZArith/BinIntDef.v b/theories/ZArith/BinIntDef.v index 443667f48..db4de0b90 100644 --- a/theories/ZArith/BinIntDef.v +++ b/theories/ZArith/BinIntDef.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Export BinNums. @@ -299,6 +301,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. diff --git a/theories/ZArith/Int.v b/theories/ZArith/Int.v index 72021f2e4..2f3bf9a32 100644 --- a/theories/ZArith/Int.v +++ b/theories/ZArith/Int.v @@ -1,10 +1,12 @@ -(***********************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA-Rocquencourt & LRI-CNRS-Orsay *) -(* \VV/ *************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) -(***********************************************************************) +(************************************************************************) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) +(* \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) *) +(************************************************************************) (** * An light axiomatization of integers (used in MSetAVL). *) @@ -452,7 +454,7 @@ Module Z_as_Int <: Int. Proof. reflexivity. Qed. (** Compatibility notations for Coq v8.4 *) - Notation plus := add (compat "8.4"). - Notation minus := sub (compat "8.4"). - Notation mult := mul (compat "8.4"). + Notation plus := add (only parsing). + Notation minus := sub (only parsing). + Notation mult := mul (only parsing). End Z_as_Int. diff --git a/theories/ZArith/Wf_Z.v b/theories/ZArith/Wf_Z.v index f6583a335..864342088 100644 --- a/theories/ZArith/Wf_Z.v +++ b/theories/ZArith/Wf_Z.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinInt. diff --git a/theories/ZArith/ZArith.v b/theories/ZArith/ZArith.v index 5f7b571d9..0842920bb 100644 --- a/theories/ZArith/ZArith.v +++ b/theories/ZArith/ZArith.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Library for manipulating integers based on binary encoding *) diff --git a/theories/ZArith/ZArith_base.v b/theories/ZArith/ZArith_base.v index 9135f5591..3d6bcddcd 100644 --- a/theories/ZArith/ZArith_base.v +++ b/theories/ZArith/ZArith_base.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Library for manipulating integers based on binary encoding. diff --git a/theories/ZArith/ZArith_dec.v b/theories/ZArith/ZArith_dec.v index 0ee233a35..9bcdb73af 100644 --- a/theories/ZArith/ZArith_dec.v +++ b/theories/ZArith/ZArith_dec.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Sumbool. @@ -32,7 +34,7 @@ Lemma Zcompare_rec (P:Set) (n m:Z) : ((n ?= m) = Eq -> P) -> ((n ?= m) = Lt -> P) -> ((n ?= m) = Gt -> P) -> P. Proof. apply Zcompare_rect. Defined. -Notation Z_eq_dec := Z.eq_dec (compat "8.3"). +Notation Z_eq_dec := Z.eq_dec (compat "8.6"). Section decidability. diff --git a/theories/ZArith/Zabs.v b/theories/ZArith/Zabs.v index d4a46930a..0d8450e36 100644 --- a/theories/ZArith/Zabs.v +++ b/theories/ZArith/Zabs.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers : properties of absolute value *) @@ -27,17 +29,17 @@ Local Open Scope Z_scope. (**********************************************************************) (** * Properties of absolute value *) -Notation Zabs_eq := Z.abs_eq (compat "8.3"). -Notation Zabs_non_eq := Z.abs_neq (compat "8.3"). -Notation Zabs_Zopp := Z.abs_opp (compat "8.3"). -Notation Zabs_pos := Z.abs_nonneg (compat "8.3"). -Notation Zabs_involutive := Z.abs_involutive (compat "8.3"). -Notation Zabs_eq_case := Z.abs_eq_cases (compat "8.3"). -Notation Zabs_triangle := Z.abs_triangle (compat "8.3"). -Notation Zsgn_Zabs := Z.sgn_abs (compat "8.3"). -Notation Zabs_Zsgn := Z.abs_sgn (compat "8.3"). -Notation Zabs_Zmult := Z.abs_mul (compat "8.3"). -Notation Zabs_square := Z.abs_square (compat "8.3"). +Notation Zabs_eq := Z.abs_eq (compat "8.6"). +Notation Zabs_non_eq := Z.abs_neq (only parsing). +Notation Zabs_Zopp := Z.abs_opp (only parsing). +Notation Zabs_pos := Z.abs_nonneg (only parsing). +Notation Zabs_involutive := Z.abs_involutive (compat "8.6"). +Notation Zabs_eq_case := Z.abs_eq_cases (only parsing). +Notation Zabs_triangle := Z.abs_triangle (compat "8.6"). +Notation Zsgn_Zabs := Z.sgn_abs (only parsing). +Notation Zabs_Zsgn := Z.abs_sgn (only parsing). +Notation Zabs_Zmult := Z.abs_mul (only parsing). +Notation Zabs_square := Z.abs_square (compat "8.6"). (** * Proving a property of the absolute value by cases *) @@ -68,11 +70,11 @@ Qed. (** * Some results about the sign function. *) -Notation Zsgn_Zmult := Z.sgn_mul (compat "8.3"). -Notation Zsgn_Zopp := Z.sgn_opp (compat "8.3"). -Notation Zsgn_pos := Z.sgn_pos_iff (compat "8.3"). -Notation Zsgn_neg := Z.sgn_neg_iff (compat "8.3"). -Notation Zsgn_null := Z.sgn_null_iff (compat "8.3"). +Notation Zsgn_Zmult := Z.sgn_mul (only parsing). +Notation Zsgn_Zopp := Z.sgn_opp (only parsing). +Notation Zsgn_pos := Z.sgn_pos_iff (only parsing). +Notation Zsgn_neg := Z.sgn_neg_iff (only parsing). +Notation Zsgn_null := Z.sgn_null_iff (only parsing). (** A characterization of the sign function: *) @@ -86,13 +88,13 @@ Qed. (** Compatibility *) -Notation inj_Zabs_nat := Zabs2Nat.id_abs (compat "8.3"). -Notation Zabs_nat_Z_of_nat := Zabs2Nat.id (compat "8.3"). -Notation Zabs_nat_mult := Zabs2Nat.inj_mul (compat "8.3"). -Notation Zabs_nat_Zsucc := Zabs2Nat.inj_succ (compat "8.3"). -Notation Zabs_nat_Zplus := Zabs2Nat.inj_add (compat "8.3"). -Notation Zabs_nat_Zminus := (fun n m => Zabs2Nat.inj_sub m n) (compat "8.3"). -Notation Zabs_nat_compare := Zabs2Nat.inj_compare (compat "8.3"). +Notation inj_Zabs_nat := Zabs2Nat.id_abs (only parsing). +Notation Zabs_nat_Z_of_nat := Zabs2Nat.id (only parsing). +Notation Zabs_nat_mult := Zabs2Nat.inj_mul (only parsing). +Notation Zabs_nat_Zsucc := Zabs2Nat.inj_succ (only parsing). +Notation Zabs_nat_Zplus := Zabs2Nat.inj_add (only parsing). +Notation Zabs_nat_Zminus := (fun n m => Zabs2Nat.inj_sub m n) (only parsing). +Notation Zabs_nat_compare := Zabs2Nat.inj_compare (only parsing). Lemma Zabs_nat_le n m : 0 <= n <= m -> (Z.abs_nat n <= Z.abs_nat m)%nat. Proof. diff --git a/theories/ZArith/Zbool.v b/theories/ZArith/Zbool.v index 407aef3b6..632d41b6a 100644 --- a/theories/ZArith/Zbool.v +++ b/theories/ZArith/Zbool.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinInt. @@ -33,10 +35,10 @@ Definition Zeven_odd_bool (x:Z) := bool_of_sumbool (Zeven_odd_dec x). (**********************************************************************) (** * Boolean comparisons of binary integers *) -Notation Zle_bool := Z.leb (compat "8.3"). -Notation Zge_bool := Z.geb (compat "8.3"). -Notation Zlt_bool := Z.ltb (compat "8.3"). -Notation Zgt_bool := Z.gtb (compat "8.3"). +Notation Zle_bool := Z.leb (only parsing). +Notation Zge_bool := Z.geb (only parsing). +Notation Zlt_bool := Z.ltb (only parsing). +Notation Zgt_bool := Z.gtb (only parsing). (** We now provide a direct [Z.eqb] that doesn't refer to [Z.compare]. The old [Zeq_bool] is kept for compatibility. *) @@ -87,7 +89,7 @@ Proof. apply Z.leb_le. Qed. -Notation Zle_bool_refl := Z.leb_refl (compat "8.3"). +Notation Zle_bool_refl := Z.leb_refl (only parsing). Lemma Zle_bool_antisym n m : (n <=? m) = true -> (m <=? n) = true -> n = m. diff --git a/theories/ZArith/Zcompare.v b/theories/ZArith/Zcompare.v index f823c41a2..c8432e27b 100644 --- a/theories/ZArith/Zcompare.v +++ b/theories/ZArith/Zcompare.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers : results about Z.compare *) @@ -181,18 +183,18 @@ Qed. (** Compatibility notations *) -Notation Zcompare_refl := Z.compare_refl (compat "8.3"). -Notation Zcompare_Eq_eq := Z.compare_eq (compat "8.3"). -Notation Zcompare_Eq_iff_eq := Z.compare_eq_iff (compat "8.3"). -Notation Zcompare_spec := Z.compare_spec (compat "8.3"). -Notation Zmin_l := Z.min_l (compat "8.3"). -Notation Zmin_r := Z.min_r (compat "8.3"). -Notation Zmax_l := Z.max_l (compat "8.3"). -Notation Zmax_r := Z.max_r (compat "8.3"). -Notation Zabs_eq := Z.abs_eq (compat "8.3"). -Notation Zabs_non_eq := Z.abs_neq (compat "8.3"). -Notation Zsgn_0 := Z.sgn_null (compat "8.3"). -Notation Zsgn_1 := Z.sgn_pos (compat "8.3"). -Notation Zsgn_m1 := Z.sgn_neg (compat "8.3"). +Notation Zcompare_refl := Z.compare_refl (compat "8.6"). +Notation Zcompare_Eq_eq := Z.compare_eq (only parsing). +Notation Zcompare_Eq_iff_eq := Z.compare_eq_iff (only parsing). +Notation Zcompare_spec := Z.compare_spec (compat "8.6"). +Notation Zmin_l := Z.min_l (compat "8.6"). +Notation Zmin_r := Z.min_r (compat "8.6"). +Notation Zmax_l := Z.max_l (compat "8.6"). +Notation Zmax_r := Z.max_r (compat "8.6"). +Notation Zabs_eq := Z.abs_eq (compat "8.6"). +Notation Zabs_non_eq := Z.abs_neq (only parsing). +Notation Zsgn_0 := Z.sgn_null (only parsing). +Notation Zsgn_1 := Z.sgn_pos (only parsing). +Notation Zsgn_m1 := Z.sgn_neg (only parsing). (** Not kept: Zcompare_egal_dec *) diff --git a/theories/ZArith/Zcomplements.v b/theories/ZArith/Zcomplements.v index eb1e17180..adf72a6ac 100644 --- a/theories/ZArith/Zcomplements.v +++ b/theories/ZArith/Zcomplements.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZArithRing. diff --git a/theories/ZArith/Zdigits.v b/theories/ZArith/Zdigits.v index f80494c5e..d1eef6131 100644 --- a/theories/ZArith/Zdigits.v +++ b/theories/ZArith/Zdigits.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Bit vectors interpreted as integers. diff --git a/theories/ZArith/Zdiv.v b/theories/ZArith/Zdiv.v index fa1ddf56f..15d0e4874 100644 --- a/theories/ZArith/Zdiv.v +++ b/theories/ZArith/Zdiv.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Euclidean Division *) @@ -18,16 +20,16 @@ Local Open Scope Z_scope. (** The definition of the division is now in [BinIntDef], the initial specifications and properties are in [BinInt]. *) -Notation Zdiv_eucl_POS := Z.pos_div_eucl (compat "8.3"). -Notation Zdiv_eucl := Z.div_eucl (compat "8.3"). -Notation Zdiv := Z.div (compat "8.3"). -Notation Zmod := Z.modulo (compat "8.3"). +Notation Zdiv_eucl_POS := Z.pos_div_eucl (only parsing). +Notation Zdiv_eucl := Z.div_eucl (compat "8.6"). +Notation Zdiv := Z.div (compat "8.6"). +Notation Zmod := Z.modulo (only parsing). -Notation Zdiv_eucl_eq := Z.div_eucl_eq (compat "8.3"). -Notation Z_div_mod_eq_full := Z.div_mod (compat "8.3"). -Notation Zmod_POS_bound := Z.pos_div_eucl_bound (compat "8.3"). -Notation Zmod_pos_bound := Z.mod_pos_bound (compat "8.3"). -Notation Zmod_neg_bound := Z.mod_neg_bound (compat "8.3"). +Notation Zdiv_eucl_eq := Z.div_eucl_eq (compat "8.6"). +Notation Z_div_mod_eq_full := Z.div_mod (only parsing). +Notation Zmod_POS_bound := Z.pos_div_eucl_bound (only parsing). +Notation Zmod_pos_bound := Z.mod_pos_bound (only parsing). +Notation Zmod_neg_bound := Z.mod_neg_bound (only parsing). (** * Main division theorems *) diff --git a/theories/ZArith/Zeuclid.v b/theories/ZArith/Zeuclid.v index 3839d1d77..dc75db131 100644 --- a/theories/ZArith/Zeuclid.v +++ b/theories/ZArith/Zeuclid.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Morphisms BinInt ZDivEucl. diff --git a/theories/ZArith/Zeven.v b/theories/ZArith/Zeven.v index 42a6a8ee3..00a58b517 100644 --- a/theories/ZArith/Zeven.v +++ b/theories/ZArith/Zeven.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers : Parity and Division by Two *) @@ -58,8 +60,8 @@ Proof (Zodd_equiv n). (** Boolean tests of parity (now in BinInt.Z) *) -Notation Zeven_bool := Z.even (compat "8.3"). -Notation Zodd_bool := Z.odd (compat "8.3"). +Notation Zeven_bool := Z.even (only parsing). +Notation Zodd_bool := Z.odd (only parsing). Lemma Zeven_bool_iff n : Z.even n = true <-> Zeven n. Proof. @@ -130,17 +132,17 @@ Qed. Hint Unfold Zeven Zodd: zarith. -Notation Zeven_bool_succ := Z.even_succ (compat "8.3"). -Notation Zeven_bool_pred := Z.even_pred (compat "8.3"). -Notation Zodd_bool_succ := Z.odd_succ (compat "8.3"). -Notation Zodd_bool_pred := Z.odd_pred (compat "8.3"). +Notation Zeven_bool_succ := Z.even_succ (only parsing). +Notation Zeven_bool_pred := Z.even_pred (only parsing). +Notation Zodd_bool_succ := Z.odd_succ (only parsing). +Notation Zodd_bool_pred := Z.odd_pred (only parsing). (******************************************************************) (** * Definition of [Z.quot2], [Z.div2] and properties wrt [Zeven] and [Zodd] *) -Notation Zdiv2 := Z.div2 (compat "8.3"). -Notation Zquot2 := Z.quot2 (compat "8.3"). +Notation Zdiv2 := Z.div2 (compat "8.6"). +Notation Zquot2 := Z.quot2 (compat "8.6"). (** Properties of [Z.div2] *) diff --git a/theories/ZArith/Zgcd_alt.v b/theories/ZArith/Zgcd_alt.v index 3a460a563..cced1190c 100644 --- a/theories/ZArith/Zgcd_alt.v +++ b/theories/ZArith/Zgcd_alt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** * Zgcd_alt : an alternate version of Z.gcd, based on Euclid's algorithm *) diff --git a/theories/ZArith/Zhints.v b/theories/ZArith/Zhints.v index 424b35a43..bfcc60edd 100644 --- a/theories/ZArith/Zhints.v +++ b/theories/ZArith/Zhints.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** This file centralizes the lemmas about [Z], classifying them diff --git a/theories/ZArith/Zlogarithm.v b/theories/ZArith/Zlogarithm.v index 28a2e287e..24412e943 100644 --- a/theories/ZArith/Zlogarithm.v +++ b/theories/ZArith/Zlogarithm.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (**********************************************************************) diff --git a/theories/ZArith/Zmax.v b/theories/ZArith/Zmax.v index b52da8563..7f595fcfd 100644 --- a/theories/ZArith/Zmax.v +++ b/theories/ZArith/Zmax.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** THIS FILE IS DEPRECATED. *) @@ -16,32 +18,32 @@ Local Open Scope Z_scope. (** Exact compatibility *) -Notation Zmax_case := Z.max_case (compat "8.3"). -Notation Zmax_case_strong := Z.max_case_strong (compat "8.3"). -Notation Zmax_right := Z.max_r (compat "8.3"). -Notation Zle_max_l := Z.le_max_l (compat "8.3"). -Notation Zle_max_r := Z.le_max_r (compat "8.3"). -Notation Zmax_lub := Z.max_lub (compat "8.3"). -Notation Zmax_lub_lt := Z.max_lub_lt (compat "8.3"). -Notation Zle_max_compat_r := Z.max_le_compat_r (compat "8.3"). -Notation Zle_max_compat_l := Z.max_le_compat_l (compat "8.3"). -Notation Zmax_idempotent := Z.max_id (compat "8.3"). -Notation Zmax_n_n := Z.max_id (compat "8.3"). -Notation Zmax_comm := Z.max_comm (compat "8.3"). -Notation Zmax_assoc := Z.max_assoc (compat "8.3"). -Notation Zmax_irreducible_dec := Z.max_dec (compat "8.3"). -Notation Zmax_le_prime := Z.max_le (compat "8.3"). -Notation Zsucc_max_distr := Z.succ_max_distr (compat "8.3"). -Notation Zmax_SS := Z.succ_max_distr (compat "8.3"). -Notation Zplus_max_distr_l := Z.add_max_distr_l (compat "8.3"). -Notation Zplus_max_distr_r := Z.add_max_distr_r (compat "8.3"). -Notation Zmax_plus := Z.add_max_distr_r (compat "8.3"). -Notation Zmax1 := Z.le_max_l (compat "8.3"). -Notation Zmax2 := Z.le_max_r (compat "8.3"). -Notation Zmax_irreducible_inf := Z.max_dec (compat "8.3"). -Notation Zmax_le_prime_inf := Z.max_le (compat "8.3"). -Notation Zpos_max := Pos2Z.inj_max (compat "8.3"). -Notation Zpos_minus := Pos2Z.inj_sub_max (compat "8.3"). +Notation Zmax_case := Z.max_case (compat "8.6"). +Notation Zmax_case_strong := Z.max_case_strong (compat "8.6"). +Notation Zmax_right := Z.max_r (only parsing). +Notation Zle_max_l := Z.le_max_l (compat "8.6"). +Notation Zle_max_r := Z.le_max_r (compat "8.6"). +Notation Zmax_lub := Z.max_lub (compat "8.6"). +Notation Zmax_lub_lt := Z.max_lub_lt (compat "8.6"). +Notation Zle_max_compat_r := Z.max_le_compat_r (only parsing). +Notation Zle_max_compat_l := Z.max_le_compat_l (only parsing). +Notation Zmax_idempotent := Z.max_id (only parsing). +Notation Zmax_n_n := Z.max_id (only parsing). +Notation Zmax_comm := Z.max_comm (compat "8.6"). +Notation Zmax_assoc := Z.max_assoc (compat "8.6"). +Notation Zmax_irreducible_dec := Z.max_dec (only parsing). +Notation Zmax_le_prime := Z.max_le (only parsing). +Notation Zsucc_max_distr := Z.succ_max_distr (compat "8.6"). +Notation Zmax_SS := Z.succ_max_distr (only parsing). +Notation Zplus_max_distr_l := Z.add_max_distr_l (only parsing). +Notation Zplus_max_distr_r := Z.add_max_distr_r (only parsing). +Notation Zmax_plus := Z.add_max_distr_r (only parsing). +Notation Zmax1 := Z.le_max_l (only parsing). +Notation Zmax2 := Z.le_max_r (only parsing). +Notation Zmax_irreducible_inf := Z.max_dec (only parsing). +Notation Zmax_le_prime_inf := Z.max_le (only parsing). +Notation Zpos_max := Pos2Z.inj_max (only parsing). +Notation Zpos_minus := Pos2Z.inj_sub_max (only parsing). (** Slightly different lemmas *) diff --git a/theories/ZArith/Zmin.v b/theories/ZArith/Zmin.v index d9e3ab19e..6bc72227b 100644 --- a/theories/ZArith/Zmin.v +++ b/theories/ZArith/Zmin.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** THIS FILE IS DEPRECATED. *) @@ -16,24 +18,24 @@ Local Open Scope Z_scope. (** Exact compatibility *) -Notation Zmin_case := Z.min_case (compat "8.3"). -Notation Zmin_case_strong := Z.min_case_strong (compat "8.3"). -Notation Zle_min_l := Z.le_min_l (compat "8.3"). -Notation Zle_min_r := Z.le_min_r (compat "8.3"). -Notation Zmin_glb := Z.min_glb (compat "8.3"). -Notation Zmin_glb_lt := Z.min_glb_lt (compat "8.3"). -Notation Zle_min_compat_r := Z.min_le_compat_r (compat "8.3"). -Notation Zle_min_compat_l := Z.min_le_compat_l (compat "8.3"). -Notation Zmin_idempotent := Z.min_id (compat "8.3"). -Notation Zmin_n_n := Z.min_id (compat "8.3"). -Notation Zmin_comm := Z.min_comm (compat "8.3"). -Notation Zmin_assoc := Z.min_assoc (compat "8.3"). -Notation Zmin_irreducible_inf := Z.min_dec (compat "8.3"). -Notation Zsucc_min_distr := Z.succ_min_distr (compat "8.3"). -Notation Zmin_SS := Z.succ_min_distr (compat "8.3"). -Notation Zplus_min_distr_r := Z.add_min_distr_r (compat "8.3"). -Notation Zmin_plus := Z.add_min_distr_r (compat "8.3"). -Notation Zpos_min := Pos2Z.inj_min (compat "8.3"). +Notation Zmin_case := Z.min_case (compat "8.6"). +Notation Zmin_case_strong := Z.min_case_strong (compat "8.6"). +Notation Zle_min_l := Z.le_min_l (compat "8.6"). +Notation Zle_min_r := Z.le_min_r (compat "8.6"). +Notation Zmin_glb := Z.min_glb (compat "8.6"). +Notation Zmin_glb_lt := Z.min_glb_lt (compat "8.6"). +Notation Zle_min_compat_r := Z.min_le_compat_r (only parsing). +Notation Zle_min_compat_l := Z.min_le_compat_l (only parsing). +Notation Zmin_idempotent := Z.min_id (only parsing). +Notation Zmin_n_n := Z.min_id (only parsing). +Notation Zmin_comm := Z.min_comm (compat "8.6"). +Notation Zmin_assoc := Z.min_assoc (compat "8.6"). +Notation Zmin_irreducible_inf := Z.min_dec (only parsing). +Notation Zsucc_min_distr := Z.succ_min_distr (compat "8.6"). +Notation Zmin_SS := Z.succ_min_distr (only parsing). +Notation Zplus_min_distr_r := Z.add_min_distr_r (only parsing). +Notation Zmin_plus := Z.add_min_distr_r (only parsing). +Notation Zpos_min := Pos2Z.inj_min (only parsing). (** Slightly different lemmas *) @@ -46,7 +48,7 @@ Qed. Lemma Zmin_irreducible n m : Z.min n m = n \/ Z.min n m = m. Proof. destruct (Z.min_dec n m); auto. Qed. -Notation Zmin_or := Zmin_irreducible (compat "8.3"). +Notation Zmin_or := Zmin_irreducible (only parsing). Lemma Zmin_le_prime_inf n m p : Z.min n m <= p -> {n <= p} + {m <= p}. Proof. apply Z.min_case; auto. Qed. diff --git a/theories/ZArith/Zminmax.v b/theories/ZArith/Zminmax.v index 7e62d6689..06919642f 100644 --- a/theories/ZArith/Zminmax.v +++ b/theories/ZArith/Zminmax.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Orders BinInt Zcompare Zorder. @@ -12,11 +14,11 @@ Require Import Orders BinInt Zcompare Zorder. (*begin hide*) (* Compatibility with names of the old Zminmax file *) -Notation Zmin_max_absorption_r_r := Z.min_max_absorption (compat "8.3"). -Notation Zmax_min_absorption_r_r := Z.max_min_absorption (compat "8.3"). -Notation Zmax_min_distr_r := Z.max_min_distr (compat "8.3"). -Notation Zmin_max_distr_r := Z.min_max_distr (compat "8.3"). -Notation Zmax_min_modular_r := Z.max_min_modular (compat "8.3"). -Notation Zmin_max_modular_r := Z.min_max_modular (compat "8.3"). -Notation max_min_disassoc := Z.max_min_disassoc (compat "8.3"). +Notation Zmin_max_absorption_r_r := Z.min_max_absorption (only parsing). +Notation Zmax_min_absorption_r_r := Z.max_min_absorption (only parsing). +Notation Zmax_min_distr_r := Z.max_min_distr (only parsing). +Notation Zmin_max_distr_r := Z.min_max_distr (only parsing). +Notation Zmax_min_modular_r := Z.max_min_modular (only parsing). +Notation Zmin_max_modular_r := Z.min_max_modular (only parsing). +Notation max_min_disassoc := Z.max_min_disassoc (only parsing). (*end hide*) diff --git a/theories/ZArith/Zmisc.v b/theories/ZArith/Zmisc.v index a6f29936b..c46241cee 100644 --- a/theories/ZArith/Zmisc.v +++ b/theories/ZArith/Zmisc.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Wf_nat. @@ -18,7 +20,7 @@ Local Open Scope Z_scope. (** [n]th iteration of the function [f] *) -Notation iter := @Z.iter (compat "8.3"). +Notation iter := @Z.iter (only parsing). Lemma iter_nat_of_Z : forall n A f x, 0 <= n -> Z.iter n f x = iter_nat (Z.abs_nat n) A f x. diff --git a/theories/ZArith/Znat.v b/theories/ZArith/Znat.v index be712db6b..5c960da1f 100644 --- a/theories/ZArith/Znat.v +++ b/theories/ZArith/Znat.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers (Pierre Crégut, CNET, Lannion, France) *) @@ -951,65 +953,65 @@ Definition inj_gt n m := proj1 (Nat2Z.inj_gt n m). (** For the others, a Notation is fine *) -Notation inj_0 := Nat2Z.inj_0 (compat "8.3"). -Notation inj_S := Nat2Z.inj_succ (compat "8.3"). -Notation inj_compare := Nat2Z.inj_compare (compat "8.3"). -Notation inj_eq_rev := Nat2Z.inj (compat "8.3"). -Notation inj_eq_iff := (fun n m => iff_sym (Nat2Z.inj_iff n m)) (compat "8.3"). -Notation inj_le_iff := Nat2Z.inj_le (compat "8.3"). -Notation inj_lt_iff := Nat2Z.inj_lt (compat "8.3"). -Notation inj_ge_iff := Nat2Z.inj_ge (compat "8.3"). -Notation inj_gt_iff := Nat2Z.inj_gt (compat "8.3"). -Notation inj_le_rev := (fun n m => proj2 (Nat2Z.inj_le n m)) (compat "8.3"). -Notation inj_lt_rev := (fun n m => proj2 (Nat2Z.inj_lt n m)) (compat "8.3"). -Notation inj_ge_rev := (fun n m => proj2 (Nat2Z.inj_ge n m)) (compat "8.3"). -Notation inj_gt_rev := (fun n m => proj2 (Nat2Z.inj_gt n m)) (compat "8.3"). -Notation inj_plus := Nat2Z.inj_add (compat "8.3"). -Notation inj_mult := Nat2Z.inj_mul (compat "8.3"). -Notation inj_minus1 := Nat2Z.inj_sub (compat "8.3"). -Notation inj_minus := Nat2Z.inj_sub_max (compat "8.3"). -Notation inj_min := Nat2Z.inj_min (compat "8.3"). -Notation inj_max := Nat2Z.inj_max (compat "8.3"). - -Notation Z_of_nat_of_P := positive_nat_Z (compat "8.3"). +Notation inj_0 := Nat2Z.inj_0 (only parsing). +Notation inj_S := Nat2Z.inj_succ (only parsing). +Notation inj_compare := Nat2Z.inj_compare (only parsing). +Notation inj_eq_rev := Nat2Z.inj (only parsing). +Notation inj_eq_iff := (fun n m => iff_sym (Nat2Z.inj_iff n m)) (only parsing). +Notation inj_le_iff := Nat2Z.inj_le (only parsing). +Notation inj_lt_iff := Nat2Z.inj_lt (only parsing). +Notation inj_ge_iff := Nat2Z.inj_ge (only parsing). +Notation inj_gt_iff := Nat2Z.inj_gt (only parsing). +Notation inj_le_rev := (fun n m => proj2 (Nat2Z.inj_le n m)) (only parsing). +Notation inj_lt_rev := (fun n m => proj2 (Nat2Z.inj_lt n m)) (only parsing). +Notation inj_ge_rev := (fun n m => proj2 (Nat2Z.inj_ge n m)) (only parsing). +Notation inj_gt_rev := (fun n m => proj2 (Nat2Z.inj_gt n m)) (only parsing). +Notation inj_plus := Nat2Z.inj_add (only parsing). +Notation inj_mult := Nat2Z.inj_mul (only parsing). +Notation inj_minus1 := Nat2Z.inj_sub (only parsing). +Notation inj_minus := Nat2Z.inj_sub_max (only parsing). +Notation inj_min := Nat2Z.inj_min (only parsing). +Notation inj_max := Nat2Z.inj_max (only parsing). + +Notation Z_of_nat_of_P := positive_nat_Z (only parsing). Notation Zpos_eq_Z_of_nat_o_nat_of_P := - (fun p => eq_sym (positive_nat_Z p)) (compat "8.3"). - -Notation Z_of_nat_of_N := N_nat_Z (compat "8.3"). -Notation Z_of_N_of_nat := nat_N_Z (compat "8.3"). - -Notation Z_of_N_eq := (f_equal Z.of_N) (compat "8.3"). -Notation Z_of_N_eq_rev := N2Z.inj (compat "8.3"). -Notation Z_of_N_eq_iff := (fun n m => iff_sym (N2Z.inj_iff n m)) (compat "8.3"). -Notation Z_of_N_compare := N2Z.inj_compare (compat "8.3"). -Notation Z_of_N_le_iff := N2Z.inj_le (compat "8.3"). -Notation Z_of_N_lt_iff := N2Z.inj_lt (compat "8.3"). -Notation Z_of_N_ge_iff := N2Z.inj_ge (compat "8.3"). -Notation Z_of_N_gt_iff := N2Z.inj_gt (compat "8.3"). -Notation Z_of_N_le := (fun n m => proj1 (N2Z.inj_le n m)) (compat "8.3"). -Notation Z_of_N_lt := (fun n m => proj1 (N2Z.inj_lt n m)) (compat "8.3"). -Notation Z_of_N_ge := (fun n m => proj1 (N2Z.inj_ge n m)) (compat "8.3"). -Notation Z_of_N_gt := (fun n m => proj1 (N2Z.inj_gt n m)) (compat "8.3"). -Notation Z_of_N_le_rev := (fun n m => proj2 (N2Z.inj_le n m)) (compat "8.3"). -Notation Z_of_N_lt_rev := (fun n m => proj2 (N2Z.inj_lt n m)) (compat "8.3"). -Notation Z_of_N_ge_rev := (fun n m => proj2 (N2Z.inj_ge n m)) (compat "8.3"). -Notation Z_of_N_gt_rev := (fun n m => proj2 (N2Z.inj_gt n m)) (compat "8.3"). -Notation Z_of_N_pos := N2Z.inj_pos (compat "8.3"). -Notation Z_of_N_abs := N2Z.inj_abs_N (compat "8.3"). -Notation Z_of_N_le_0 := N2Z.is_nonneg (compat "8.3"). -Notation Z_of_N_plus := N2Z.inj_add (compat "8.3"). -Notation Z_of_N_mult := N2Z.inj_mul (compat "8.3"). -Notation Z_of_N_minus := N2Z.inj_sub_max (compat "8.3"). -Notation Z_of_N_succ := N2Z.inj_succ (compat "8.3"). -Notation Z_of_N_min := N2Z.inj_min (compat "8.3"). -Notation Z_of_N_max := N2Z.inj_max (compat "8.3"). -Notation Zabs_of_N := Zabs2N.id (compat "8.3"). -Notation Zabs_N_succ_abs := Zabs2N.inj_succ_abs (compat "8.3"). -Notation Zabs_N_succ := Zabs2N.inj_succ (compat "8.3"). -Notation Zabs_N_plus_abs := Zabs2N.inj_add_abs (compat "8.3"). -Notation Zabs_N_plus := Zabs2N.inj_add (compat "8.3"). -Notation Zabs_N_mult_abs := Zabs2N.inj_mul_abs (compat "8.3"). -Notation Zabs_N_mult := Zabs2N.inj_mul (compat "8.3"). + (fun p => eq_sym (positive_nat_Z p)) (only parsing). + +Notation Z_of_nat_of_N := N_nat_Z (only parsing). +Notation Z_of_N_of_nat := nat_N_Z (only parsing). + +Notation Z_of_N_eq := (f_equal Z.of_N) (only parsing). +Notation Z_of_N_eq_rev := N2Z.inj (only parsing). +Notation Z_of_N_eq_iff := (fun n m => iff_sym (N2Z.inj_iff n m)) (only parsing). +Notation Z_of_N_compare := N2Z.inj_compare (only parsing). +Notation Z_of_N_le_iff := N2Z.inj_le (only parsing). +Notation Z_of_N_lt_iff := N2Z.inj_lt (only parsing). +Notation Z_of_N_ge_iff := N2Z.inj_ge (only parsing). +Notation Z_of_N_gt_iff := N2Z.inj_gt (only parsing). +Notation Z_of_N_le := (fun n m => proj1 (N2Z.inj_le n m)) (only parsing). +Notation Z_of_N_lt := (fun n m => proj1 (N2Z.inj_lt n m)) (only parsing). +Notation Z_of_N_ge := (fun n m => proj1 (N2Z.inj_ge n m)) (only parsing). +Notation Z_of_N_gt := (fun n m => proj1 (N2Z.inj_gt n m)) (only parsing). +Notation Z_of_N_le_rev := (fun n m => proj2 (N2Z.inj_le n m)) (only parsing). +Notation Z_of_N_lt_rev := (fun n m => proj2 (N2Z.inj_lt n m)) (only parsing). +Notation Z_of_N_ge_rev := (fun n m => proj2 (N2Z.inj_ge n m)) (only parsing). +Notation Z_of_N_gt_rev := (fun n m => proj2 (N2Z.inj_gt n m)) (only parsing). +Notation Z_of_N_pos := N2Z.inj_pos (only parsing). +Notation Z_of_N_abs := N2Z.inj_abs_N (only parsing). +Notation Z_of_N_le_0 := N2Z.is_nonneg (only parsing). +Notation Z_of_N_plus := N2Z.inj_add (only parsing). +Notation Z_of_N_mult := N2Z.inj_mul (only parsing). +Notation Z_of_N_minus := N2Z.inj_sub_max (only parsing). +Notation Z_of_N_succ := N2Z.inj_succ (only parsing). +Notation Z_of_N_min := N2Z.inj_min (only parsing). +Notation Z_of_N_max := N2Z.inj_max (only parsing). +Notation Zabs_of_N := Zabs2N.id (only parsing). +Notation Zabs_N_succ_abs := Zabs2N.inj_succ_abs (only parsing). +Notation Zabs_N_succ := Zabs2N.inj_succ (only parsing). +Notation Zabs_N_plus_abs := Zabs2N.inj_add_abs (only parsing). +Notation Zabs_N_plus := Zabs2N.inj_add (only parsing). +Notation Zabs_N_mult_abs := Zabs2N.inj_mul_abs (only parsing). +Notation Zabs_N_mult := Zabs2N.inj_mul (only parsing). Theorem inj_minus2 : forall n m:nat, (m > n)%nat -> Z.of_nat (n - m) = 0. Proof. diff --git a/theories/ZArith/Znumtheory.v b/theories/ZArith/Znumtheory.v index 5dfc37095..f5444c31d 100644 --- a/theories/ZArith/Znumtheory.v +++ b/theories/ZArith/Znumtheory.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZArith_base. @@ -25,20 +27,20 @@ Open Scope Z_scope. - properties of the efficient [Z.gcd] function *) -Notation Zgcd := Z.gcd (compat "8.3"). -Notation Zggcd := Z.ggcd (compat "8.3"). -Notation Zggcd_gcd := Z.ggcd_gcd (compat "8.3"). -Notation Zggcd_correct_divisors := Z.ggcd_correct_divisors (compat "8.3"). -Notation Zgcd_divide_l := Z.gcd_divide_l (compat "8.3"). -Notation Zgcd_divide_r := Z.gcd_divide_r (compat "8.3"). -Notation Zgcd_greatest := Z.gcd_greatest (compat "8.3"). -Notation Zgcd_nonneg := Z.gcd_nonneg (compat "8.3"). -Notation Zggcd_opp := Z.ggcd_opp (compat "8.3"). +Notation Zgcd := Z.gcd (compat "8.6"). +Notation Zggcd := Z.ggcd (compat "8.6"). +Notation Zggcd_gcd := Z.ggcd_gcd (compat "8.6"). +Notation Zggcd_correct_divisors := Z.ggcd_correct_divisors (compat "8.6"). +Notation Zgcd_divide_l := Z.gcd_divide_l (compat "8.6"). +Notation Zgcd_divide_r := Z.gcd_divide_r (compat "8.6"). +Notation Zgcd_greatest := Z.gcd_greatest (compat "8.6"). +Notation Zgcd_nonneg := Z.gcd_nonneg (compat "8.6"). +Notation Zggcd_opp := Z.ggcd_opp (compat "8.6"). (** The former specialized inductive predicate [Z.divide] is now a generic existential predicate. *) -Notation Zdivide := Z.divide (compat "8.3"). +Notation Zdivide := Z.divide (compat "8.6"). (** Its former constructor is now a pseudo-constructor. *) @@ -46,17 +48,17 @@ Definition Zdivide_intro a b q (H:b=q*a) : Z.divide a b := ex_intro _ q H. (** Results concerning divisibility*) -Notation Zdivide_refl := Z.divide_refl (compat "8.3"). -Notation Zone_divide := Z.divide_1_l (compat "8.3"). -Notation Zdivide_0 := Z.divide_0_r (compat "8.3"). -Notation Zmult_divide_compat_l := Z.mul_divide_mono_l (compat "8.3"). -Notation Zmult_divide_compat_r := Z.mul_divide_mono_r (compat "8.3"). -Notation Zdivide_plus_r := Z.divide_add_r (compat "8.3"). -Notation Zdivide_minus_l := Z.divide_sub_r (compat "8.3"). -Notation Zdivide_mult_l := Z.divide_mul_l (compat "8.3"). -Notation Zdivide_mult_r := Z.divide_mul_r (compat "8.3"). -Notation Zdivide_factor_r := Z.divide_factor_l (compat "8.3"). -Notation Zdivide_factor_l := Z.divide_factor_r (compat "8.3"). +Notation Zdivide_refl := Z.divide_refl (compat "8.6"). +Notation Zone_divide := Z.divide_1_l (only parsing). +Notation Zdivide_0 := Z.divide_0_r (only parsing). +Notation Zmult_divide_compat_l := Z.mul_divide_mono_l (only parsing). +Notation Zmult_divide_compat_r := Z.mul_divide_mono_r (only parsing). +Notation Zdivide_plus_r := Z.divide_add_r (only parsing). +Notation Zdivide_minus_l := Z.divide_sub_r (only parsing). +Notation Zdivide_mult_l := Z.divide_mul_l (only parsing). +Notation Zdivide_mult_r := Z.divide_mul_r (only parsing). +Notation Zdivide_factor_r := Z.divide_factor_l (only parsing). +Notation Zdivide_factor_l := Z.divide_factor_r (only parsing). Lemma Zdivide_opp_r a b : (a | b) -> (a | - b). Proof. apply Z.divide_opp_r. Qed. @@ -91,12 +93,12 @@ Qed. (** Only [1] and [-1] divide [1]. *) -Notation Zdivide_1 := Z.divide_1_r (compat "8.3"). +Notation Zdivide_1 := Z.divide_1_r (only parsing). (** If [a] divides [b] and [b] divides [a] then [a] is [b] or [-b]. *) -Notation Zdivide_antisym := Z.divide_antisym (compat "8.3"). -Notation Zdivide_trans := Z.divide_trans (compat "8.3"). +Notation Zdivide_antisym := Z.divide_antisym (compat "8.6"). +Notation Zdivide_trans := Z.divide_trans (compat "8.6"). (** If [a] divides [b] and [b<>0] then [|a| <= |b|]. *) @@ -734,7 +736,7 @@ Qed. (** we now prove that [Z.gcd] is indeed a gcd in the sense of [Zis_gcd]. *) -Notation Zgcd_is_pos := Z.gcd_nonneg (compat "8.3"). +Notation Zgcd_is_pos := Z.gcd_nonneg (only parsing). Lemma Zgcd_is_gcd : forall a b, Zis_gcd a b (Z.gcd a b). Proof. @@ -767,8 +769,8 @@ Proof. - subst. now case (Z.gcd a b). Qed. -Notation Zgcd_inv_0_l := Z.gcd_eq_0_l (compat "8.3"). -Notation Zgcd_inv_0_r := Z.gcd_eq_0_r (compat "8.3"). +Notation Zgcd_inv_0_l := Z.gcd_eq_0_l (only parsing). +Notation Zgcd_inv_0_r := Z.gcd_eq_0_r (only parsing). Theorem Zgcd_div_swap0 : forall a b : Z, 0 < Z.gcd a b -> @@ -798,16 +800,16 @@ Proof. rewrite <- Zdivide_Zdiv_eq; auto. Qed. -Notation Zgcd_comm := Z.gcd_comm (compat "8.3"). +Notation Zgcd_comm := Z.gcd_comm (compat "8.6"). Lemma Zgcd_ass a b c : Z.gcd (Z.gcd a b) c = Z.gcd a (Z.gcd b c). Proof. symmetry. apply Z.gcd_assoc. Qed. -Notation Zgcd_Zabs := Z.gcd_abs_l (compat "8.3"). -Notation Zgcd_0 := Z.gcd_0_r (compat "8.3"). -Notation Zgcd_1 := Z.gcd_1_r (compat "8.3"). +Notation Zgcd_Zabs := Z.gcd_abs_l (only parsing). +Notation Zgcd_0 := Z.gcd_0_r (only parsing). +Notation Zgcd_1 := Z.gcd_1_r (only parsing). Hint Resolve Z.gcd_0_r Z.gcd_1_r : zarith. diff --git a/theories/ZArith/Zorder.v b/theories/ZArith/Zorder.v index ff8e22029..a1ec4b35e 100644 --- a/theories/ZArith/Zorder.v +++ b/theories/ZArith/Zorder.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers : results about order predicates *) @@ -38,9 +40,9 @@ Qed. (**********************************************************************) (** * Decidability of equality and order on Z *) -Notation dec_eq := Z.eq_decidable (compat "8.3"). -Notation dec_Zle := Z.le_decidable (compat "8.3"). -Notation dec_Zlt := Z.lt_decidable (compat "8.3"). +Notation dec_eq := Z.eq_decidable (only parsing). +Notation dec_Zle := Z.le_decidable (only parsing). +Notation dec_Zlt := Z.lt_decidable (only parsing). Theorem dec_Zne n m : decidable (Zne n m). Proof. @@ -64,12 +66,12 @@ Qed. (** * Relating strict and large orders *) -Notation Zgt_lt := Z.gt_lt (compat "8.3"). -Notation Zlt_gt := Z.lt_gt (compat "8.3"). -Notation Zge_le := Z.ge_le (compat "8.3"). -Notation Zle_ge := Z.le_ge (compat "8.3"). -Notation Zgt_iff_lt := Z.gt_lt_iff (compat "8.3"). -Notation Zge_iff_le := Z.ge_le_iff (compat "8.3"). +Notation Zgt_lt := Z.gt_lt (compat "8.6"). +Notation Zlt_gt := Z.lt_gt (compat "8.6"). +Notation Zge_le := Z.ge_le (compat "8.6"). +Notation Zle_ge := Z.le_ge (compat "8.6"). +Notation Zgt_iff_lt := Z.gt_lt_iff (only parsing). +Notation Zge_iff_le := Z.ge_le_iff (only parsing). Lemma Zle_not_lt n m : n <= m -> ~ m < n. Proof. @@ -121,18 +123,18 @@ Qed. (** Reflexivity *) -Notation Zle_refl := Z.le_refl (compat "8.3"). -Notation Zeq_le := Z.eq_le_incl (compat "8.3"). +Notation Zle_refl := Z.le_refl (compat "8.6"). +Notation Zeq_le := Z.eq_le_incl (only parsing). Hint Resolve Z.le_refl: zarith. (** Antisymmetry *) -Notation Zle_antisym := Z.le_antisymm (compat "8.3"). +Notation Zle_antisym := Z.le_antisymm (only parsing). (** Asymmetry *) -Notation Zlt_asym := Z.lt_asymm (compat "8.3"). +Notation Zlt_asym := Z.lt_asymm (only parsing). Lemma Zgt_asym n m : n > m -> ~ m > n. Proof. @@ -141,8 +143,8 @@ Qed. (** Irreflexivity *) -Notation Zlt_irrefl := Z.lt_irrefl (compat "8.3"). -Notation Zlt_not_eq := Z.lt_neq (compat "8.3"). +Notation Zlt_irrefl := Z.lt_irrefl (compat "8.6"). +Notation Zlt_not_eq := Z.lt_neq (only parsing). Lemma Zgt_irrefl n : ~ n > n. Proof. @@ -151,8 +153,8 @@ Qed. (** Large = strict or equal *) -Notation Zlt_le_weak := Z.lt_le_incl (compat "8.3"). -Notation Zle_lt_or_eq_iff := Z.lt_eq_cases (compat "8.3"). +Notation Zlt_le_weak := Z.lt_le_incl (only parsing). +Notation Zle_lt_or_eq_iff := Z.lt_eq_cases (only parsing). Lemma Zle_lt_or_eq n m : n <= m -> n < m \/ n = m. Proof. @@ -161,11 +163,11 @@ Qed. (** Dichotomy *) -Notation Zle_or_lt := Z.le_gt_cases (compat "8.3"). +Notation Zle_or_lt := Z.le_gt_cases (only parsing). (** Transitivity of strict orders *) -Notation Zlt_trans := Z.lt_trans (compat "8.3"). +Notation Zlt_trans := Z.lt_trans (compat "8.6"). Lemma Zgt_trans n m p : n > m -> m > p -> n > p. Proof. @@ -174,8 +176,8 @@ Qed. (** Mixed transitivity *) -Notation Zlt_le_trans := Z.lt_le_trans (compat "8.3"). -Notation Zle_lt_trans := Z.le_lt_trans (compat "8.3"). +Notation Zlt_le_trans := Z.lt_le_trans (compat "8.6"). +Notation Zle_lt_trans := Z.le_lt_trans (compat "8.6"). Lemma Zle_gt_trans n m p : m <= n -> m > p -> n > p. Proof. @@ -189,7 +191,7 @@ Qed. (** Transitivity of large orders *) -Notation Zle_trans := Z.le_trans (compat "8.3"). +Notation Zle_trans := Z.le_trans (compat "8.6"). Lemma Zge_trans n m p : n >= m -> m >= p -> n >= p. Proof. @@ -240,8 +242,8 @@ Qed. (** Special base instances of order *) -Notation Zlt_succ := Z.lt_succ_diag_r (compat "8.3"). -Notation Zlt_pred := Z.lt_pred_l (compat "8.3"). +Notation Zlt_succ := Z.lt_succ_diag_r (only parsing). +Notation Zlt_pred := Z.lt_pred_l (only parsing). Lemma Zgt_succ n : Z.succ n > n. Proof. @@ -255,8 +257,8 @@ Qed. (** Relating strict and large order using successor or predecessor *) -Notation Zlt_succ_r := Z.lt_succ_r (compat "8.3"). -Notation Zle_succ_l := Z.le_succ_l (compat "8.3"). +Notation Zlt_succ_r := Z.lt_succ_r (compat "8.6"). +Notation Zle_succ_l := Z.le_succ_l (compat "8.6"). Lemma Zgt_le_succ n m : m > n -> Z.succ n <= m. Proof. @@ -295,10 +297,10 @@ Qed. (** Weakening order *) -Notation Zle_succ := Z.le_succ_diag_r (compat "8.3"). -Notation Zle_pred := Z.le_pred_l (compat "8.3"). -Notation Zlt_lt_succ := Z.lt_lt_succ_r (compat "8.3"). -Notation Zle_le_succ := Z.le_le_succ_r (compat "8.3"). +Notation Zle_succ := Z.le_succ_diag_r (only parsing). +Notation Zle_pred := Z.le_pred_l (only parsing). +Notation Zlt_lt_succ := Z.lt_lt_succ_r (only parsing). +Notation Zle_le_succ := Z.le_le_succ_r (only parsing). Lemma Zle_succ_le n m : Z.succ n <= m -> n <= m. Proof. @@ -334,8 +336,8 @@ Qed. (** Special cases of ordered integers *) -Notation Zlt_0_1 := Z.lt_0_1 (compat "8.3"). -Notation Zle_0_1 := Z.le_0_1 (compat "8.3"). +Notation Zlt_0_1 := Z.lt_0_1 (compat "8.6"). +Notation Zle_0_1 := Z.le_0_1 (compat "8.6"). Lemma Zle_neg_pos : forall p q:positive, Zneg p <= Zpos q. Proof. @@ -375,10 +377,10 @@ Qed. (** ** Addition *) (** Compatibility of addition wrt to order *) -Notation Zplus_lt_le_compat := Z.add_lt_le_mono (compat "8.3"). -Notation Zplus_le_lt_compat := Z.add_le_lt_mono (compat "8.3"). -Notation Zplus_le_compat := Z.add_le_mono (compat "8.3"). -Notation Zplus_lt_compat := Z.add_lt_mono (compat "8.3"). +Notation Zplus_lt_le_compat := Z.add_lt_le_mono (only parsing). +Notation Zplus_le_lt_compat := Z.add_le_lt_mono (only parsing). +Notation Zplus_le_compat := Z.add_le_mono (only parsing). +Notation Zplus_lt_compat := Z.add_lt_mono (only parsing). Lemma Zplus_gt_compat_l n m p : n > m -> p + n > p + m. Proof. @@ -412,7 +414,7 @@ Qed. (** Compatibility of addition wrt to being positive *) -Notation Zplus_le_0_compat := Z.add_nonneg_nonneg (compat "8.3"). +Notation Zplus_le_0_compat := Z.add_nonneg_nonneg (only parsing). (** Simplification of addition wrt to order *) @@ -570,9 +572,9 @@ Qed. (** Compatibility of multiplication by a positive wrt to being positive *) -Notation Zmult_le_0_compat := Z.mul_nonneg_nonneg (compat "8.3"). -Notation Zmult_lt_0_compat := Z.mul_pos_pos (compat "8.3"). -Notation Zmult_lt_O_compat := Z.mul_pos_pos (compat "8.3"). +Notation Zmult_le_0_compat := Z.mul_nonneg_nonneg (only parsing). +Notation Zmult_lt_0_compat := Z.mul_pos_pos (only parsing). +Notation Zmult_lt_O_compat := Z.mul_pos_pos (only parsing). Lemma Zmult_gt_0_compat n m : n > 0 -> m > 0 -> n * m > 0. Proof. @@ -624,9 +626,9 @@ Qed. (** * Equivalence between inequalities *) -Notation Zle_plus_swap := Z.le_add_le_sub_r (compat "8.3"). -Notation Zlt_plus_swap := Z.lt_add_lt_sub_r (compat "8.3"). -Notation Zlt_minus_simpl_swap := Z.lt_sub_pos (compat "8.3"). +Notation Zle_plus_swap := Z.le_add_le_sub_r (only parsing). +Notation Zlt_plus_swap := Z.lt_add_lt_sub_r (only parsing). +Notation Zlt_minus_simpl_swap := Z.lt_sub_pos (only parsing). Lemma Zeq_plus_swap n m p : n + p = m <-> n = m - p. Proof. diff --git a/theories/ZArith/Zpow_alt.v b/theories/ZArith/Zpow_alt.v index e1220dcee..983405acb 100644 --- a/theories/ZArith/Zpow_alt.v +++ b/theories/ZArith/Zpow_alt.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinInt. diff --git a/theories/ZArith/Zpow_def.v b/theories/ZArith/Zpow_def.v index a768868bd..2b099671f 100644 --- a/theories/ZArith/Zpow_def.v +++ b/theories/ZArith/Zpow_def.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 BinInt Ring_theory. @@ -14,12 +16,12 @@ Local Open Scope Z_scope. (** Nota : this file is mostly deprecated. The definition of [Z.pow] and its usual properties are now provided by module [BinInt.Z]. *) -Notation Zpower_pos := Z.pow_pos (compat "8.3"). -Notation Zpower := Z.pow (compat "8.3"). -Notation Zpower_0_r := Z.pow_0_r (compat "8.3"). -Notation Zpower_succ_r := Z.pow_succ_r (compat "8.3"). -Notation Zpower_neg_r := Z.pow_neg_r (compat "8.3"). -Notation Zpower_Ppow := Pos2Z.inj_pow (compat "8.3"). +Notation Zpower_pos := Z.pow_pos (only parsing). +Notation Zpower := Z.pow (only parsing). +Notation Zpower_0_r := Z.pow_0_r (only parsing). +Notation Zpower_succ_r := Z.pow_succ_r (only parsing). +Notation Zpower_neg_r := Z.pow_neg_r (only parsing). +Notation Zpower_Ppow := Pos2Z.inj_pow (only parsing). Lemma Zpower_theory : power_theory 1 Z.mul (@eq Z) Z.of_N Z.pow. Proof. diff --git a/theories/ZArith/Zpow_facts.v b/theories/ZArith/Zpow_facts.v index 3ea3ae4ab..a9bc5bd09 100644 --- a/theories/ZArith/Zpow_facts.v +++ b/theories/ZArith/Zpow_facts.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZArith_base ZArithRing Zcomplements Zdiv Znumtheory. @@ -29,17 +31,17 @@ Proof. now apply (Z.pow_0_l (Zpos p)). Qed. Lemma Zpower_pos_pos x p : 0 < x -> 0 < Z.pow_pos x p. Proof. intros. now apply (Z.pow_pos_nonneg x (Zpos p)). Qed. -Notation Zpower_1_r := Z.pow_1_r (compat "8.3"). -Notation Zpower_1_l := Z.pow_1_l (compat "8.3"). -Notation Zpower_0_l := Z.pow_0_l' (compat "8.3"). -Notation Zpower_0_r := Z.pow_0_r (compat "8.3"). -Notation Zpower_2 := Z.pow_2_r (compat "8.3"). -Notation Zpower_gt_0 := Z.pow_pos_nonneg (compat "8.3"). -Notation Zpower_ge_0 := Z.pow_nonneg (compat "8.3"). -Notation Zpower_Zabs := Z.abs_pow (compat "8.3"). -Notation Zpower_Zsucc := Z.pow_succ_r (compat "8.3"). -Notation Zpower_mult := Z.pow_mul_r (compat "8.3"). -Notation Zpower_le_monotone2 := Z.pow_le_mono_r (compat "8.3"). +Notation Zpower_1_r := Z.pow_1_r (only parsing). +Notation Zpower_1_l := Z.pow_1_l (only parsing). +Notation Zpower_0_l := Z.pow_0_l' (only parsing). +Notation Zpower_0_r := Z.pow_0_r (only parsing). +Notation Zpower_2 := Z.pow_2_r (only parsing). +Notation Zpower_gt_0 := Z.pow_pos_nonneg (only parsing). +Notation Zpower_ge_0 := Z.pow_nonneg (only parsing). +Notation Zpower_Zabs := Z.abs_pow (only parsing). +Notation Zpower_Zsucc := Z.pow_succ_r (only parsing). +Notation Zpower_mult := Z.pow_mul_r (only parsing). +Notation Zpower_le_monotone2 := Z.pow_le_mono_r (only parsing). Theorem Zpower_le_monotone a b c : 0 < a -> 0 <= b <= c -> a^b <= a^c. @@ -231,7 +233,7 @@ Qed. (** * Z.square: a direct definition of [z^2] *) -Notation Psquare := Pos.square (compat "8.3"). -Notation Zsquare := Z.square (compat "8.3"). -Notation Psquare_correct := Pos.square_spec (compat "8.3"). -Notation Zsquare_correct := Z.square_spec (compat "8.3"). +Notation Psquare := Pos.square (compat "8.6"). +Notation Zsquare := Z.square (compat "8.6"). +Notation Psquare_correct := Pos.square_spec (only parsing). +Notation Zsquare_correct := Z.square_spec (only parsing). diff --git a/theories/ZArith/Zpower.v b/theories/ZArith/Zpower.v index 6dcbdbdee..fa6905355 100644 --- a/theories/ZArith/Zpower.v +++ b/theories/ZArith/Zpower.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Wf_nat ZArith_base Omega Zcomplements. diff --git a/theories/ZArith/Zquot.v b/theories/ZArith/Zquot.v index efb56c469..e93ebb1ad 100644 --- a/theories/ZArith/Zquot.v +++ b/theories/ZArith/Zquot.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 Nnat ZArith_base ROmega ZArithRing Zdiv Morphisms. @@ -31,21 +33,21 @@ Local Open Scope Z_scope. exploiting the arbitrary value of division by 0). *) -Notation Ndiv_Zquot := N2Z.inj_quot (compat "8.3"). -Notation Nmod_Zrem := N2Z.inj_rem (compat "8.3"). -Notation Z_quot_rem_eq := Z.quot_rem' (compat "8.3"). -Notation Zrem_lt := Z.rem_bound_abs (compat "8.3"). -Notation Zquot_unique := Z.quot_unique (compat "8.3"). -Notation Zrem_unique := Z.rem_unique (compat "8.3"). -Notation Zrem_1_r := Z.rem_1_r (compat "8.3"). -Notation Zquot_1_r := Z.quot_1_r (compat "8.3"). -Notation Zrem_1_l := Z.rem_1_l (compat "8.3"). -Notation Zquot_1_l := Z.quot_1_l (compat "8.3"). -Notation Z_quot_same := Z.quot_same (compat "8.3"). -Notation Z_quot_mult := Z.quot_mul (compat "8.3"). -Notation Zquot_small := Z.quot_small (compat "8.3"). -Notation Zrem_small := Z.rem_small (compat "8.3"). -Notation Zquot2_quot := Zquot2_quot (compat "8.3"). +Notation Ndiv_Zquot := N2Z.inj_quot (only parsing). +Notation Nmod_Zrem := N2Z.inj_rem (only parsing). +Notation Z_quot_rem_eq := Z.quot_rem' (only parsing). +Notation Zrem_lt := Z.rem_bound_abs (only parsing). +Notation Zquot_unique := Z.quot_unique (compat "8.6"). +Notation Zrem_unique := Z.rem_unique (compat "8.6"). +Notation Zrem_1_r := Z.rem_1_r (compat "8.6"). +Notation Zquot_1_r := Z.quot_1_r (compat "8.6"). +Notation Zrem_1_l := Z.rem_1_l (compat "8.6"). +Notation Zquot_1_l := Z.quot_1_l (compat "8.6"). +Notation Z_quot_same := Z.quot_same (compat "8.6"). +Notation Z_quot_mult := Z.quot_mul (only parsing). +Notation Zquot_small := Z.quot_small (compat "8.6"). +Notation Zrem_small := Z.rem_small (compat "8.6"). +Notation Zquot2_quot := Zquot2_quot (compat "8.6"). (** Particular values taken for [a÷0] and [(Z.rem a 0)]. We avise to not rely on these arbitrary values. *) diff --git a/theories/ZArith/Zsqrt_compat.v b/theories/ZArith/Zsqrt_compat.v index cccd970da..bd0904540 100644 --- a/theories/ZArith/Zsqrt_compat.v +++ b/theories/ZArith/Zsqrt_compat.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZArithRing. diff --git a/theories/ZArith/Zwf.v b/theories/ZArith/Zwf.v index ca4b386dc..a71ea4f30 100644 --- a/theories/ZArith/Zwf.v +++ b/theories/ZArith/Zwf.v @@ -1,9 +1,11 @@ (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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 ZArith_base. diff --git a/theories/ZArith/auxiliary.v b/theories/ZArith/auxiliary.v index 494cb30dd..306a85638 100644 --- a/theories/ZArith/auxiliary.v +++ b/theories/ZArith/auxiliary.v @@ -1,10 +1,12 @@ (* -*- coding: utf-8 -*- *) (************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017 *) +(* * The Coq Proof Assistant / The Coq Development Team *) +(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *) +(* <O___,, * (see CREDITS file for the list of authors) *) (* \VV/ **************************************************************) -(* // * This file is distributed under the terms of the *) -(* * GNU Lesser General Public License Version 2.1 *) +(* // * 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) *) (************************************************************************) (** Binary Integers (Pierre Crégut, CNET, Lannion, France) *) |