aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--checker/univ.mli2
-rw-r--r--kernel/names.mli10
-rw-r--r--lib/cMap.ml2
-rw-r--r--lib/cMap.mli2
-rw-r--r--lib/cSig.mli31
-rw-r--r--library/goptions.mli2
-rw-r--r--library/libnames.mli2
-rw-r--r--plugins/cc/ccalgo.mli6
-rw-r--r--plugins/extraction/table.mli2
-rw-r--r--plugins/firstorder/sequent.mli2
10 files changed, 46 insertions, 15 deletions
diff --git a/checker/univ.mli b/checker/univ.mli
index 02c1bbdb9..f3216feac 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -130,7 +130,7 @@ val check_constraints : constraints -> universes -> bool
(** {6 Support for universe polymorphism } *)
(** Polymorphic maps from universe levels to 'a *)
-module LMap : Map.S with type key = universe_level
+module LMap : CSig.MapS with type key = universe_level
module LSet : CSig.SetS with type elt = universe_level
type 'a universe_map = 'a LMap.t
diff --git a/kernel/names.mli b/kernel/names.mli
index 7cc444375..59419af2e 100644
--- a/kernel/names.mli
+++ b/kernel/names.mli
@@ -395,7 +395,7 @@ end
module Mindset : CSig.SetS with type elt = MutInd.t
module Mindmap : Map.ExtS with type key = MutInd.t and module Set := Mindset
-module Mindmap_env : Map.S with type key = MutInd.t
+module Mindmap_env : CSig.MapS with type key = MutInd.t
(** Beware: first inductive has index 0 *)
type inductive = MutInd.t * int
@@ -403,10 +403,10 @@ type inductive = MutInd.t * int
(** Beware: first constructor has index 1 *)
type constructor = inductive * int
-module Indmap : Map.S with type key = inductive
-module Constrmap : Map.S with type key = constructor
-module Indmap_env : Map.S with type key = inductive
-module Constrmap_env : Map.S with type key = constructor
+module Indmap : CSig.MapS with type key = inductive
+module Constrmap : CSig.MapS with type key = constructor
+module Indmap_env : CSig.MapS with type key = inductive
+module Constrmap_env : CSig.MapS with type key = constructor
val ind_modpath : inductive -> ModPath.t
val constr_modpath : constructor -> ModPath.t
diff --git a/lib/cMap.ml b/lib/cMap.ml
index cf590d96c..048a69081 100644
--- a/lib/cMap.ml
+++ b/lib/cMap.ml
@@ -16,7 +16,7 @@ module type S = Map.S
module type ExtS =
sig
- include Map.S
+ include CSig.MapS
module Set : CSig.SetS with type elt = key
val update : key -> 'a -> 'a t -> 'a t
val modify : key -> (key -> 'a -> 'a) -> 'a t -> 'a t
diff --git a/lib/cMap.mli b/lib/cMap.mli
index 23d3801e0..9d0fbbad2 100644
--- a/lib/cMap.mli
+++ b/lib/cMap.mli
@@ -18,7 +18,7 @@ module type S = Map.S
module type ExtS =
sig
- include Map.S
+ include CSig.MapS
(** The underlying Map library *)
module Set : CSig.SetS with type elt = key
diff --git a/lib/cSig.mli b/lib/cSig.mli
index 2a8bda293..e095c82cb 100644
--- a/lib/cSig.mli
+++ b/lib/cSig.mli
@@ -45,3 +45,34 @@ sig
end
(** Redeclaration of OCaml set signature, to preserve compatibility. See OCaml
documentation for more information. *)
+
+module type MapS =
+sig
+ type key
+ type (+'a) t
+ val empty: 'a t
+ val is_empty: 'a t -> bool
+ val mem: key -> 'a t -> bool
+ val add: key -> 'a -> 'a t -> 'a t
+ val singleton: key -> 'a -> 'a t
+ val remove: key -> 'a t -> 'a t
+ val merge:
+ (key -> 'a option -> 'b option -> 'c option) -> 'a t -> 'b t -> 'c t
+ val compare: ('a -> 'a -> int) -> 'a t -> 'a t -> int
+ val equal: ('a -> 'a -> bool) -> 'a t -> 'a t -> bool
+ val iter: (key -> 'a -> unit) -> 'a t -> unit
+ val fold: (key -> 'a -> 'b -> 'b) -> 'a t -> 'b -> 'b
+ val for_all: (key -> 'a -> bool) -> 'a t -> bool
+ val exists: (key -> 'a -> bool) -> 'a t -> bool
+ val filter: (key -> 'a -> bool) -> 'a t -> 'a t
+ val partition: (key -> 'a -> bool) -> 'a t -> 'a t * 'a t
+ val cardinal: 'a t -> int
+ val bindings: 'a t -> (key * 'a) list
+ val min_binding: 'a t -> (key * 'a)
+ val max_binding: 'a t -> (key * 'a)
+ val choose: 'a t -> (key * 'a)
+ val split: key -> 'a t -> 'a t * 'a option * 'a t
+ val find: key -> 'a t -> 'a
+ val map: ('a -> 'b) -> 'a t -> 'b t
+ val mapi: (key -> 'a -> 'b) -> 'a t -> 'b t
+end
diff --git a/library/goptions.mli b/library/goptions.mli
index 9d87c14c5..25b5315c2 100644
--- a/library/goptions.mli
+++ b/library/goptions.mli
@@ -133,7 +133,7 @@ val declare_stringopt_option: string option option_sig -> string option write_fu
(** {6 Special functions supposed to be used only in vernacentries.ml } *)
-module OptionMap : Map.S with type key = option_name
+module OptionMap : CSig.MapS with type key = option_name
val get_string_table :
option_name ->
diff --git a/library/libnames.mli b/library/libnames.mli
index b95c08871..c72f51753 100644
--- a/library/libnames.mli
+++ b/library/libnames.mli
@@ -60,7 +60,7 @@ val path_of_string : string -> full_path
val string_of_path : full_path -> string
val pr_path : full_path -> std_ppcmds
-module Spmap : Map.S with type key = full_path
+module Spmap : CSig.MapS with type key = full_path
val restrict_path : int -> full_path -> full_path
diff --git a/plugins/cc/ccalgo.mli b/plugins/cc/ccalgo.mli
index 0dcf3a870..34c19958a 100644
--- a/plugins/cc/ccalgo.mli
+++ b/plugins/cc/ccalgo.mli
@@ -20,8 +20,8 @@ type pa_fun=
fnargs:int}
-module PafMap : Map.S with type key = pa_fun
-module PacMap : Map.S with type key = pa_constructor
+module PafMap : CSig.MapS with type key = pa_fun
+module PacMap : CSig.MapS with type key = pa_constructor
type cinfo =
{ci_constr: pconstructor; (* inductive type *)
@@ -185,7 +185,7 @@ val empty_forest: unit -> forest
(*type pa_constructor
-module PacMap:Map.S with type key=pa_constructor
+module PacMap:CSig.MapS with type key=pa_constructor
type term =
Symb of Term.constr
diff --git a/plugins/extraction/table.mli b/plugins/extraction/table.mli
index 916cf3ad6..4e638a0ac 100644
--- a/plugins/extraction/table.mli
+++ b/plugins/extraction/table.mli
@@ -13,7 +13,7 @@ open Miniml
open Declarations
module Refset' : CSig.SetS with type elt = global_reference
-module Refmap' : Map.S with type key = global_reference
+module Refmap' : CSig.MapS with type key = global_reference
val safe_basename_of_global : global_reference -> Id.t
diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli
index dc3f05be6..760168c9f 100644
--- a/plugins/firstorder/sequent.mli
+++ b/plugins/firstorder/sequent.mli
@@ -13,7 +13,7 @@ open Globnames
module OrderedConstr: Set.OrderedType with type t=constr
-module CM: Map.S with type key=constr
+module CM: CSig.MapS with type key=constr
type h_item = global_reference * (int*constr) option