aboutsummaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--plugins/romega/ReflOmegaCore.v2
-rw-r--r--plugins/setoid_ring/Field_theory.v3
-rw-r--r--plugins/setoid_ring/Ncring_initial.v5
-rw-r--r--plugins/ssr/ssreflect.v3
-rw-r--r--plugins/ssr/ssrfun.v2
-rw-r--r--plugins/ssrmatching/ssrmatching.v4
-rw-r--r--test-suite/output/Arguments.v2
-rw-r--r--test-suite/output/Notations.v1
-rw-r--r--theories/Bool/Bool.v2
-rw-r--r--theories/Classes/CEquivalence.v2
-rw-r--r--theories/Classes/CMorphisms.v1
-rw-r--r--theories/Classes/Equivalence.v2
-rw-r--r--theories/Classes/Morphisms.v1
-rw-r--r--theories/Classes/RelationClasses.v2
-rw-r--r--theories/FSets/FMapAVL.v1
-rw-r--r--theories/Init/Datatypes.v8
-rw-r--r--theories/Init/Decimal.v3
-rw-r--r--theories/Init/Notations.v10
-rw-r--r--theories/Init/Specif.v2
-rw-r--r--theories/MSets/MSetAVL.v2
-rw-r--r--theories/Numbers/BinNums.v3
-rw-r--r--theories/Numbers/Cyclic/Int31/Int31.v1
-rw-r--r--theories/Numbers/Integer/Abstract/ZDivEucl.v1
-rw-r--r--theories/Numbers/Integer/NatPairs/ZNatPairs.v6
-rw-r--r--theories/Numbers/NatInt/NZDomain.v4
-rw-r--r--theories/Program/Basics.v2
-rw-r--r--theories/Program/Utils.v5
-rw-r--r--theories/QArith/QArith_base.v1
-rw-r--r--theories/QArith/Qcanon.v1
-rw-r--r--theories/Reals/Ranalysis1.v1
-rw-r--r--theories/Reals/Raxioms.v1
-rw-r--r--theories/Structures/OrdersFacts.v1
-rw-r--r--theories/Vectors/VectorDef.v1
-rw-r--r--theories/ZArith/Int.v1
34 files changed, 74 insertions, 13 deletions
diff --git a/plugins/romega/ReflOmegaCore.v b/plugins/romega/ReflOmegaCore.v
index 51b99b9935..da86f4274d 100644
--- a/plugins/romega/ReflOmegaCore.v
+++ b/plugins/romega/ReflOmegaCore.v
@@ -8,6 +8,7 @@
*************************************************************************)
Require Import List Bool Sumbool EqNat Setoid Ring_theory Decidable ZArith_base.
+Declare Scope Int_scope.
Delimit Scope Int_scope with I.
(** * Abstract Integers. *)
@@ -716,6 +717,7 @@ Inductive term : Set :=
| Topp : term -> term
| Tvar : N -> term.
+Declare Scope romega_scope.
Bind Scope romega_scope with term.
Delimit Scope romega_scope with term.
Arguments Tint _%I.
diff --git a/plugins/setoid_ring/Field_theory.v b/plugins/setoid_ring/Field_theory.v
index d9e32dbbf8..ce115f564f 100644
--- a/plugins/setoid_ring/Field_theory.v
+++ b/plugins/setoid_ring/Field_theory.v
@@ -19,6 +19,7 @@ Section MakeFieldPol.
(* Field elements : R *)
Variable R:Type.
+Declare Scope R_scope.
Bind Scope R_scope with R.
Delimit Scope R_scope with ring.
Local Open Scope R_scope.
@@ -94,6 +95,7 @@ Let rdistr_r := ARdistr_r Rsth Reqe ARth.
(* Coefficients : C *)
Variable C: Type.
+Declare Scope C_scope.
Bind Scope C_scope with C.
Delimit Scope C_scope with coef.
@@ -139,6 +141,7 @@ Let rpow_pow := pow_th.(rpow_pow_N).
(* Polynomial expressions : (PExpr C) *)
+Declare Scope PE_scope.
Bind Scope PE_scope with PExpr.
Delimit Scope PE_scope with poly.
diff --git a/plugins/setoid_ring/Ncring_initial.v b/plugins/setoid_ring/Ncring_initial.v
index 523c7b02eb..1ca6227f25 100644
--- a/plugins/setoid_ring/Ncring_initial.v
+++ b/plugins/setoid_ring/Ncring_initial.v
@@ -79,8 +79,9 @@ Context {R:Type}`{Ring R}.
| Z0 => 0
| Zneg p => -(gen_phiPOS p)
end.
- Local Notation "[ x ]" := (gen_phiZ x) : ZMORPHISM.
- Local Open Scope ZMORPHISM.
+ Declare Scope ZMORPHISM.
+ Notation "[ x ]" := (gen_phiZ x) : ZMORPHISM.
+ Open Scope ZMORPHISM.
Definition get_signZ z :=
match z with
diff --git a/plugins/ssr/ssreflect.v b/plugins/ssr/ssreflect.v
index b4144aa45e..460bdc6d23 100644
--- a/plugins/ssr/ssreflect.v
+++ b/plugins/ssr/ssreflect.v
@@ -86,6 +86,7 @@ Export SsrSyntax.
(* recognize the expansion of the boolean if; using the default printer *)
(* avoids a spurrious trailing %GEN_IF. *)
+Declare Scope general_if_scope.
Delimit Scope general_if_scope with GEN_IF.
Notation "'if' c 'then' v1 'else' v2" :=
@@ -103,6 +104,7 @@ Notation "'if' c 'as' x 'return' t 'then' v1 'else' v2" :=
(* Force boolean interpretation of simple if expressions. *)
+Declare Scope boolean_if_scope.
Delimit Scope boolean_if_scope with BOOL_IF.
Notation "'if' c 'return' t 'then' v1 'else' v2" :=
@@ -125,6 +127,7 @@ Open Scope boolean_if_scope.
(* Non-ssreflect libraries that do not respect the form syntax (e.g., the Coq *)
(* Lists library) should be loaded before ssreflect so that their notations *)
(* do not mask all ssreflect forms. *)
+Declare Scope form_scope.
Delimit Scope form_scope with FORM.
Open Scope form_scope.
diff --git a/plugins/ssr/ssrfun.v b/plugins/ssr/ssrfun.v
index b2d5143e36..99ff943e61 100644
--- a/plugins/ssr/ssrfun.v
+++ b/plugins/ssr/ssrfun.v
@@ -216,6 +216,7 @@ Set Implicit Arguments.
Unset Strict Implicit.
Unset Printing Implicit Defensive.
+Declare Scope fun_scope.
Delimit Scope fun_scope with FUN.
Open Scope fun_scope.
@@ -225,6 +226,7 @@ Notation "f ^~ y" := (fun x => f x y)
Notation "@^~ x" := (fun f => f x)
(at level 10, x at level 8, no associativity, format "@^~ x") : fun_scope.
+Declare Scope pair_scope.
Delimit Scope pair_scope with PAIR.
Open Scope pair_scope.
diff --git a/plugins/ssrmatching/ssrmatching.v b/plugins/ssrmatching/ssrmatching.v
index 829ee05e11..9a53e1dd1a 100644
--- a/plugins/ssrmatching/ssrmatching.v
+++ b/plugins/ssrmatching/ssrmatching.v
@@ -11,9 +11,11 @@ Reserved Notation "( a 'as' b )" (at level 0).
Reserved Notation "( a 'in' b 'in' c )" (at level 0).
Reserved Notation "( a 'as' b 'in' c )" (at level 0).
+Declare Scope ssrpatternscope.
+Delimit Scope ssrpatternscope with pattern.
+
(* Notation to define shortcuts for the "X in t" part of a pattern. *)
Notation "( X 'in' t )" := (_ : fun X => t) : ssrpatternscope.
-Delimit Scope ssrpatternscope with pattern.
(* Some shortcuts for recurrent "X in t" parts. *)
Notation RHS := (X in _ = X)%pattern.
diff --git a/test-suite/output/Arguments.v b/test-suite/output/Arguments.v
index bd9240476f..b67ac4f0df 100644
--- a/test-suite/output/Arguments.v
+++ b/test-suite/output/Arguments.v
@@ -10,6 +10,8 @@ Arguments Nat.sub !n !m.
About Nat.sub.
Definition pf (D1 C1 : Type) (f : D1 -> C1) (D2 C2 : Type) (g : D2 -> C2) :=
fun x => (f (fst x), g (snd x)).
+Declare Scope foo_scope.
+Declare Scope bar_scope.
Delimit Scope foo_scope with F.
Arguments pf {D1%F C1%type} f [D2 C2] g x : simpl never.
About pf.
diff --git a/test-suite/output/Notations.v b/test-suite/output/Notations.v
index fe6c05c39e..adab324cf0 100644
--- a/test-suite/output/Notations.v
+++ b/test-suite/output/Notations.v
@@ -76,6 +76,7 @@ Open Scope nat_scope.
Inductive znat : Set := Zpos (n : nat) | Zneg (m : nat).
Coercion Zpos: nat >-> znat.
+Declare Scope znat_scope.
Delimit Scope znat_scope with znat.
Open Scope znat_scope.
diff --git a/theories/Bool/Bool.v b/theories/Bool/Bool.v
index 66a82008d8..42af3583d4 100644
--- a/theories/Bool/Bool.v
+++ b/theories/Bool/Bool.v
@@ -757,6 +757,8 @@ Qed.
with lazy behavior (for vm_compute) *)
(*****************************************)
+Declare Scope lazy_bool_scope.
+
Notation "a &&& b" := (if a then b else false)
(at level 40, left associativity) : lazy_bool_scope.
Notation "a ||| b" := (if a then true else b)
diff --git a/theories/Classes/CEquivalence.v b/theories/Classes/CEquivalence.v
index 03e611f549..c376efef2e 100644
--- a/theories/Classes/CEquivalence.v
+++ b/theories/Classes/CEquivalence.v
@@ -35,6 +35,8 @@ Definition equiv `{Equivalence A R} : crelation A := R.
(** Overloaded notations for setoid equivalence and inequivalence.
Not to be confused with [eq] and [=]. *)
+Declare Scope equiv_scope.
+
Notation " x === y " := (equiv x y) (at level 70, no associativity) : equiv_scope.
Notation " x =/= y " := (complement equiv x y) (at level 70, no associativity) : equiv_scope.
diff --git a/theories/Classes/CMorphisms.v b/theories/Classes/CMorphisms.v
index 09b35ca75d..97510578ae 100644
--- a/theories/Classes/CMorphisms.v
+++ b/theories/Classes/CMorphisms.v
@@ -87,6 +87,7 @@ Hint Extern 2 (ProperProxy ?R _) =>
not_evar R; class_apply @proper_proper_proxy : typeclass_instances.
(** Notations reminiscent of the old syntax for declaring morphisms. *)
+Declare Scope signature_scope.
Delimit Scope signature_scope with signature.
Module ProperNotations.
diff --git a/theories/Classes/Equivalence.v b/theories/Classes/Equivalence.v
index 5217aedb88..516ea12099 100644
--- a/theories/Classes/Equivalence.v
+++ b/theories/Classes/Equivalence.v
@@ -35,6 +35,8 @@ Definition equiv `{Equivalence A R} : relation A := R.
(** Overloaded notations for setoid equivalence and inequivalence.
Not to be confused with [eq] and [=]. *)
+Declare Scope equiv_scope.
+
Notation " x === y " := (equiv x y) (at level 70, no associativity) : equiv_scope.
Notation " x =/= y " := (complement equiv x y) (at level 70, no associativity) : equiv_scope.
diff --git a/theories/Classes/Morphisms.v b/theories/Classes/Morphisms.v
index 1858ba76ae..001b7dfdfd 100644
--- a/theories/Classes/Morphisms.v
+++ b/theories/Classes/Morphisms.v
@@ -88,6 +88,7 @@ Hint Extern 2 (ProperProxy ?R _) =>
not_evar R; class_apply @proper_proper_proxy : typeclass_instances.
(** Notations reminiscent of the old syntax for declaring morphisms. *)
+Declare Scope signature_scope.
Delimit Scope signature_scope with signature.
Module ProperNotations.
diff --git a/theories/Classes/RelationClasses.v b/theories/Classes/RelationClasses.v
index 2ab3af2029..86a3a88be9 100644
--- a/theories/Classes/RelationClasses.v
+++ b/theories/Classes/RelationClasses.v
@@ -357,6 +357,8 @@ Definition predicate_implication {l : Tlist} :=
(** Notations for pointwise equivalence and implication of predicates. *)
+Declare Scope predicate_scope.
+
Infix "<∙>" := predicate_equivalence (at level 95, no associativity) : predicate_scope.
Infix "-∙>" := predicate_implication (at level 70, right associativity) : predicate_scope.
diff --git a/theories/FSets/FMapAVL.v b/theories/FSets/FMapAVL.v
index 3485b9c68d..b0d1824827 100644
--- a/theories/FSets/FMapAVL.v
+++ b/theories/FSets/FMapAVL.v
@@ -25,6 +25,7 @@ Unset Strict Implicit.
(** Notations and helper lemma about pairs *)
+Declare Scope pair_scope.
Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope.
Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope.
diff --git a/theories/Init/Datatypes.v b/theories/Init/Datatypes.v
index 1e6843d511..76c39f275d 100644
--- a/theories/Init/Datatypes.v
+++ b/theories/Init/Datatypes.v
@@ -37,8 +37,8 @@ Inductive bool : Set :=
Add Printing If bool.
+Declare Scope bool_scope.
Delimit Scope bool_scope with bool.
-
Bind Scope bool_scope with bool.
(** Basic boolean operators *)
@@ -136,6 +136,7 @@ Inductive nat : Set :=
| O : nat
| S : nat -> nat.
+Declare Scope nat_scope.
Delimit Scope nat_scope with nat.
Bind Scope nat_scope with nat.
Arguments S _%nat.
@@ -228,10 +229,13 @@ Inductive list (A : Type) : Type :=
Arguments nil {A}.
Arguments cons {A} a l.
-Infix "::" := cons (at level 60, right associativity) : list_scope.
+
+Declare Scope list_scope.
Delimit Scope list_scope with list.
Bind Scope list_scope with list.
+Infix "::" := cons (at level 60, right associativity) : list_scope.
+
Local Open Scope list_scope.
Definition length (A : Type) : list A -> nat :=
diff --git a/theories/Init/Decimal.v b/theories/Init/Decimal.v
index 1ff00ec11c..537400fb05 100644
--- a/theories/Init/Decimal.v
+++ b/theories/Init/Decimal.v
@@ -42,8 +42,11 @@ Notation zero := (D0 Nil).
Inductive int := Pos (d:uint) | Neg (d:uint).
+Declare Scope dec_uint_scope.
Delimit Scope dec_uint_scope with uint.
Bind Scope dec_uint_scope with uint.
+
+Declare Scope dec_int_scope.
Delimit Scope dec_int_scope with int.
Bind Scope dec_int_scope with int.
diff --git a/theories/Init/Notations.v b/theories/Init/Notations.v
index 72073bb4f6..8f8e639187 100644
--- a/theories/Init/Notations.v
+++ b/theories/Init/Notations.v
@@ -108,13 +108,17 @@ End IfNotations.
(** Scopes *)
-Delimit Scope type_scope with type.
-Delimit Scope function_scope with function.
+Declare Scope core_scope.
Delimit Scope core_scope with core.
-Bind Scope type_scope with Sortclass.
+Declare Scope function_scope.
+Delimit Scope function_scope with function.
Bind Scope function_scope with Funclass.
+Declare Scope type_scope.
+Delimit Scope type_scope with type.
+Bind Scope type_scope with Sortclass.
+
Open Scope core_scope.
Open Scope function_scope.
Open Scope type_scope.
diff --git a/theories/Init/Specif.v b/theories/Init/Specif.v
index b6afba29a0..db8857df64 100644
--- a/theories/Init/Specif.v
+++ b/theories/Init/Specif.v
@@ -697,7 +697,7 @@ End Choice_lemmas.
Section Dependent_choice_lemmas.
- Variables X : Set.
+ Variable X : Set.
Variable R : X -> X -> Prop.
Lemma dependent_choice :
diff --git a/theories/MSets/MSetAVL.v b/theories/MSets/MSetAVL.v
index b966f217aa..aec88f93bf 100644
--- a/theories/MSets/MSetAVL.v
+++ b/theories/MSets/MSetAVL.v
@@ -340,6 +340,8 @@ Functional Scheme union_ind := Induction for union Sort Prop.
(** Notations and helper lemma about pairs and triples *)
+Declare Scope pair_scope.
+
Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope.
Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope.
Notation "t #l" := (t_left t) (at level 9, format "t '#l'") : pair_scope.
diff --git a/theories/Numbers/BinNums.v b/theories/Numbers/BinNums.v
index 3ba9d1f5ed..7b6740e94b 100644
--- a/theories/Numbers/BinNums.v
+++ b/theories/Numbers/BinNums.v
@@ -23,6 +23,7 @@ Inductive positive : Set :=
| xO : positive -> positive
| xH : positive.
+Declare Scope positive_scope.
Delimit Scope positive_scope with positive.
Bind Scope positive_scope with positive.
Arguments xO _%positive.
@@ -37,6 +38,7 @@ Inductive N : Set :=
| N0 : N
| Npos : positive -> N.
+Declare Scope N_scope.
Delimit Scope N_scope with N.
Bind Scope N_scope with N.
Arguments Npos _%positive.
@@ -53,6 +55,7 @@ Inductive Z : Set :=
| Zpos : positive -> Z
| Zneg : positive -> Z.
+Declare Scope Z_scope.
Delimit Scope Z_scope with Z.
Bind Scope Z_scope with Z.
Arguments Zpos _%positive.
diff --git a/theories/Numbers/Cyclic/Int31/Int31.v b/theories/Numbers/Cyclic/Int31/Int31.v
index 39af62c32f..d91bfd4e2c 100644
--- a/theories/Numbers/Cyclic/Int31/Int31.v
+++ b/theories/Numbers/Cyclic/Int31/Int31.v
@@ -50,6 +50,7 @@ Inductive int31 : Type := I31 : digits31 int31.
Register digits as int31 bits in "coq_int31" by True.
Register int31 as int31 type in "coq_int31" by True.
+Declare Scope int31_scope.
Delimit Scope int31_scope with int31.
Bind Scope int31_scope with int31.
Local Open Scope int31_scope.
diff --git a/theories/Numbers/Integer/Abstract/ZDivEucl.v b/theories/Numbers/Integer/Abstract/ZDivEucl.v
index d7f25a6613..ab17bb6e1a 100644
--- a/theories/Numbers/Integer/Abstract/ZDivEucl.v
+++ b/theories/Numbers/Integer/Abstract/ZDivEucl.v
@@ -46,6 +46,7 @@ Module ZEuclidProp
(** We put notations in a scope, to avoid warnings about
redefinitions of notations *)
+ Declare Scope euclid.
Infix "/" := D.div : euclid.
Infix "mod" := D.modulo : euclid.
Local Open Scope euclid.
diff --git a/theories/Numbers/Integer/NatPairs/ZNatPairs.v b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
index 4b2d5c13b5..995d96b314 100644
--- a/theories/Numbers/Integer/NatPairs/ZNatPairs.v
+++ b/theories/Numbers/Integer/NatPairs/ZNatPairs.v
@@ -13,15 +13,18 @@
Require Import NSub ZAxioms.
Require Export Ring.
+Declare Scope pair_scope.
+Local Open Scope pair_scope.
+
Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope.
Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope.
-Local Open Scope pair_scope.
Module ZPairsAxiomsMod (Import N : NAxiomsMiniSig) <: ZAxiomsMiniSig.
Module Import NProp.
Include NSubProp N.
End NProp.
+Declare Scope NScope.
Delimit Scope NScope with N.
Bind Scope NScope with N.t.
Infix "==" := N.eq (at level 70) : NScope.
@@ -73,6 +76,7 @@ Definition max (n m : t) : t := (max (n#1 + m#2) (m#1 + n#2), n#2 + m#2).
End Z.
+Declare Scope ZScope.
Delimit Scope ZScope with Z.
Bind Scope ZScope with Z.t.
Infix "==" := Z.eq (at level 70) : ZScope.
diff --git a/theories/Numbers/NatInt/NZDomain.v b/theories/Numbers/NatInt/NZDomain.v
index 3d0c005fd1..acebfcf1d2 100644
--- a/theories/Numbers/NatInt/NZDomain.v
+++ b/theories/Numbers/NatInt/NZDomain.v
@@ -220,8 +220,10 @@ End NZDomainProp.
Module NZOfNat (Import NZ:NZDomainSig').
Definition ofnat (n : nat) : t := (S^n) 0.
-Notation "[ n ]" := (ofnat n) (at level 7) : ofnat.
+
+Declare Scope ofnat.
Local Open Scope ofnat.
+Notation "[ n ]" := (ofnat n) (at level 7) : ofnat.
Lemma ofnat_zero : [O] == 0.
Proof.
diff --git a/theories/Program/Basics.v b/theories/Program/Basics.v
index f55093ed48..c2316689fc 100644
--- a/theories/Program/Basics.v
+++ b/theories/Program/Basics.v
@@ -28,6 +28,8 @@ Definition compose {A B C} (g : B -> C) (f : A -> B) :=
Hint Unfold compose.
+Declare Scope program_scope.
+
Notation " g ∘ f " := (compose g f)
(at level 40, left associativity) : program_scope.
diff --git a/theories/Program/Utils.v b/theories/Program/Utils.v
index 78c36dc7d1..c51cacac68 100644
--- a/theories/Program/Utils.v
+++ b/theories/Program/Utils.v
@@ -20,12 +20,13 @@ Notation "{ ( x , y ) : A | P }" :=
(sig (fun anonymous : A => let (x,y) := anonymous in P))
(x ident, y ident, at level 10) : type_scope.
+Declare Scope program_scope.
+Delimit Scope program_scope with prg.
+
(** Generates an obligation to prove False. *)
Notation " ! " := (False_rect _ _) : program_scope.
-Delimit Scope program_scope with prg.
-
(** Abbreviation for first projection and hiding of proofs of subset objects. *)
Notation " ` t " := (proj1_sig t) (at level 10, t at next level) : program_scope.
diff --git a/theories/QArith/QArith_base.v b/theories/QArith/QArith_base.v
index 35706e7fa2..139c4bf432 100644
--- a/theories/QArith/QArith_base.v
+++ b/theories/QArith/QArith_base.v
@@ -18,6 +18,7 @@ Require Export Morphisms Setoid Bool.
Record Q : Set := Qmake {Qnum : Z; Qden : positive}.
+Declare Scope Q_scope.
Delimit Scope Q_scope with Q.
Bind Scope Q_scope with Q.
Arguments Qmake _%Z _%positive.
diff --git a/theories/QArith/Qcanon.v b/theories/QArith/Qcanon.v
index 1510a7b825..81c318138e 100644
--- a/theories/QArith/Qcanon.v
+++ b/theories/QArith/Qcanon.v
@@ -18,6 +18,7 @@ Require Import Eqdep_dec.
Record Qc : Set := Qcmake { this :> Q ; canon : Qred this = this }.
+Declare Scope Qc_scope.
Delimit Scope Qc_scope with Qc.
Bind Scope Qc_scope with Qc.
Arguments Qcmake this%Q _.
diff --git a/theories/Reals/Ranalysis1.v b/theories/Reals/Ranalysis1.v
index 36ac738ca6..9f8039ec9d 100644
--- a/theories/Reals/Ranalysis1.v
+++ b/theories/Reals/Ranalysis1.v
@@ -28,6 +28,7 @@ Definition div_real_fct (a:R) f (x:R) : R := a / f x.
Definition comp f1 f2 (x:R) : R := f1 (f2 x).
Definition inv_fct f (x:R) : R := / f x.
+Declare Scope Rfun_scope.
Delimit Scope Rfun_scope with F.
Arguments plus_fct (f1 f2)%F x%R.
diff --git a/theories/Reals/Raxioms.v b/theories/Reals/Raxioms.v
index 6019d4faf1..a2818371e9 100644
--- a/theories/Reals/Raxioms.v
+++ b/theories/Reals/Raxioms.v
@@ -14,6 +14,7 @@
Require Export ZArith_base.
Require Export Rdefinitions.
+Declare Scope R_scope.
Local Open Scope R_scope.
(*********************************************************)
diff --git a/theories/Structures/OrdersFacts.v b/theories/Structures/OrdersFacts.v
index 87df6b479d..60c64d306b 100644
--- a/theories/Structures/OrdersFacts.v
+++ b/theories/Structures/OrdersFacts.v
@@ -132,6 +132,7 @@ Module OrderedTypeFacts (Import O: OrderedType').
Module OrderTac := OT_to_OrderTac O.
Ltac order := OrderTac.order.
+ Declare Scope order.
Notation "x <= y" := (~lt y x) : order.
Infix "?=" := compare (at level 70, no associativity) : order.
diff --git a/theories/Vectors/VectorDef.v b/theories/Vectors/VectorDef.v
index ba3e411091..390ca78c0e 100644
--- a/theories/Vectors/VectorDef.v
+++ b/theories/Vectors/VectorDef.v
@@ -305,6 +305,7 @@ Eval cbv delta beta in fold_right (fun h H => Datatypes.cons h H) v Datatypes.ni
End VECTORLIST.
Module VectorNotations.
+Declare Scope vector_scope.
Delimit Scope vector_scope with vector.
Notation "[ ]" := [] (format "[ ]") : vector_scope.
Notation "h :: t" := (h :: t) (at level 60, right associativity)
diff --git a/theories/ZArith/Int.v b/theories/ZArith/Int.v
index 2f3bf9a32a..1e35370d29 100644
--- a/theories/ZArith/Int.v
+++ b/theories/ZArith/Int.v
@@ -17,6 +17,7 @@
*)
Require Import BinInt.
+Declare Scope Int_scope.
Delimit Scope Int_scope with I.
Local Open Scope Int_scope.