aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--Makefile.doc15
-rw-r--r--checker/univ.mli2
-rw-r--r--clib/option.ml2
-rw-r--r--engine/evd.mli4
-rw-r--r--engine/proofview.ml6
-rw-r--r--interp/constrintern.ml2
-rw-r--r--intf/constrexpr.ml4
-rw-r--r--kernel/cClosure.ml2
-rw-r--r--kernel/cClosure.mli2
-rw-r--r--kernel/environ.mli2
-rw-r--r--kernel/uGraph.mli2
-rw-r--r--lib/lib.mllib1
-rw-r--r--plugins/ssr/ssrcommon.ml8
-rw-r--r--plugins/ssr/ssrparser.ml46
-rw-r--r--plugins/ssr/ssrtacticals.ml2
-rw-r--r--plugins/ssr/ssrvernac.ml46
-rw-r--r--plugins/ssrmatching/ssrmatching.ml44
-rw-r--r--plugins/ssrmatching/ssrmatching.mli2
-rw-r--r--pretyping/reductionops.mli2
-rw-r--r--tactics/hipattern.ml4
20 files changed, 41 insertions, 37 deletions
diff --git a/Makefile.doc b/Makefile.doc
index 894ef9a99..9fd93651d 100644
--- a/Makefile.doc
+++ b/Makefile.doc
@@ -387,10 +387,10 @@ install-doc-index-urls:
OCAMLDOCDIR=dev/ocamldoc
-DOCMLIS=$(wildcard ./lib/*.mli ./intf/*.mli ./kernel/*.mli ./library/*.mli \
- ./engine/*.mli ./pretyping/*.mli ./interp/*.mli printing/*.mli \
- ./parsing/*.mli ./proofs/*.mli \
- ./tactics/*.mli ./stm/*.mli ./toplevel/*.mli ./ltac/*.mli)
+DOCMLLIBS= $(CORECMA:.cma=_MLLIB_DEPENDENCIES) $(PLUGINSCMO:.cmo=_MLPACK_DEPENDENCIES)
+DOCMLS=$(foreach lib,$(DOCMLLIBS),$(addsuffix .ml, $($(lib))))
+
+DOCMLIS=$(wildcard $(addsuffix /*.mli, $(SRCDIRS)))
# Defining options to generate dependencies graphs
DOT=dot
@@ -434,7 +434,12 @@ OCAMLDOC_MLLIBD = $(OCAMLFIND) ocamldoc -rectypes $(MLINCLUDES) $(ODOCDOTOPTS) -
$(OCAMLDOC_MLLIBD)
ml-doc:
- $(OCAMLFIND) ocamldoc -charset utf-8 -html -rectypes -I +threads $(MLINCLUDES) $(COQIDEFLAGS) -d $(OCAMLDOCDIR) $(MLSTATICFILES)
+ $(SHOW)'OCAMLDOC -html'
+ $(HIDE)mkdir -p $(OCAMLDOCDIR)/html/implementation
+ $(HIDE)$(OCAMLFIND) ocamldoc -charset utf-8 -html -rectypes -I +threads $(MLINCLUDES) $(COQIDEFLAGS) \
+ $(DOCMLS) -d $(OCAMLDOCDIR)/html/implementation -colorize-code \
+ -t "Coq mls documentation" \
+ -css-style ../style.css
parsing/parsing.dot : | parsing/parsing.mllib.d
$(OCAMLDOC_MLLIBD)
diff --git a/checker/univ.mli b/checker/univ.mli
index 3876e7bbc..935f0a2b8 100644
--- a/checker/univ.mli
+++ b/checker/univ.mli
@@ -84,7 +84,7 @@ val check_eq : universe check_function
val initial_universes : universes
(** Adds a universe to the graph, ensuring it is >= or > Set.
- @raises AlreadyDeclared if the level is already declared in the graph. *)
+ @raise AlreadyDeclared if the level is already declared in the graph. *)
exception AlreadyDeclared
diff --git a/clib/option.ml b/clib/option.ml
index c2e2e7097..21913e8f7 100644
--- a/clib/option.ml
+++ b/clib/option.ml
@@ -44,7 +44,7 @@ let hash f = function
exception IsNone
(** [get x] returns [y] where [x] is [Some y].
- @raise [IsNone] if [x] equals [None]. *)
+ @raise IsNone if [x] equals [None]. *)
let get = function
| Some y -> y
| _ -> raise IsNone
diff --git a/engine/evd.mli b/engine/evd.mli
index 55b8e3a83..ed3316c16 100644
--- a/engine/evd.mli
+++ b/engine/evd.mli
@@ -320,8 +320,8 @@ exception UniversesDiffer
val add_universe_constraints : evar_map -> Universes.Constraints.t -> evar_map
(** Add the given universe unification constraints to the evar map.
- @raises UniversesDiffer in case a first-order unification fails.
- @raises UniverseInconsistency
+ @raise UniversesDiffer in case a first-order unification fails.
+ @raise UniverseInconsistency .
*)
(** {5 Extra data}
diff --git a/engine/proofview.ml b/engine/proofview.ml
index 25c8e2d80..8a844bbf5 100644
--- a/engine/proofview.ml
+++ b/engine/proofview.ml
@@ -127,7 +127,7 @@ let focus_context (left,right) =
(** This (internal) function extracts a sublist between two indices,
and returns this sublist together with its context: if it returns
- [(a,(b,c))] then [a] is the sublist and (rev b)@a@c is the
+ [(a,(b,c))] then [a] is the sublist and [(rev b) @ a @ c] is the
original list. The focused list has lenght [j-i-1] and contains
the goals from number [i] to number [j] (both included) the first
goal of the list being numbered [1]. [focus_sublist i j l] raises
@@ -572,8 +572,8 @@ let tclDISPATCHL tacs = tclDISPATCHGEN CList.rev tacs
(** [extend_to_list startxs rx endxs l] builds a list
- [startxs@[rx,...,rx]@endxs] of the same length as [l]. Raises
- [SizeMismatch] if [startxs@endxs] is already longer than [l]. *)
+ [startxs @ [rx,...,rx] @ endxs] of the same length as [l]. Raises
+ [SizeMismatch] if [startxs @ endxs] is already longer than [l]. *)
let extend_to_list startxs rx endxs l =
(* spiwack: I use [l] essentially as a natural number *)
let rec duplicate acc = function
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 84162ca89..918e12e5c 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -1077,7 +1077,7 @@ type 'a raw_cases_pattern_expr_r =
| RCPatAlias of 'a raw_cases_pattern_expr * Misctypes.lname
| RCPatCstr of Globnames.global_reference
* 'a raw_cases_pattern_expr list * 'a raw_cases_pattern_expr list
- (** [RCPatCstr (loc, c, l1, l2)] represents ((@c l1) l2) *)
+ (** [RCPatCstr (loc, c, l1, l2)] represents [((@ c l1) l2)] *)
| RCPatAtom of (Misctypes.lident * (Notation_term.tmp_scope_name option * Notation_term.scope_name list)) option
| RCPatOr of 'a raw_cases_pattern_expr list
and 'a raw_cases_pattern_expr = ('a raw_cases_pattern_expr_r, 'a) DAst.t
diff --git a/intf/constrexpr.ml b/intf/constrexpr.ml
index b0a769c5a..31f811bc8 100644
--- a/intf/constrexpr.ml
+++ b/intf/constrexpr.ml
@@ -56,7 +56,7 @@ type cases_pattern_expr_r =
| CPatAlias of cases_pattern_expr * lname
| CPatCstr of reference
* cases_pattern_expr list option * cases_pattern_expr list
- (** [CPatCstr (_, c, Some l1, l2)] represents (@c l1) l2 *)
+ (** [CPatCstr (_, c, Some l1, l2)] represents [(@ c l1) l2] *)
| CPatAtom of reference option
| CPatOr of cases_pattern_expr list
| CPatNotation of notation * cases_pattern_notation_substitution
@@ -126,7 +126,7 @@ and recursion_order_expr =
| CWfRec of constr_expr
| CMeasureRec of constr_expr * constr_expr option (** measure, relation *)
-(** Anonymous defs allowed ?? *)
+(* Anonymous defs allowed ?? *)
and local_binder_expr =
| CLocalAssum of lname list * binder_kind * constr_expr
| CLocalDef of lname * constr_expr * constr_expr option
diff --git a/kernel/cClosure.ml b/kernel/cClosure.ml
index c5a8c7b23..11faef02c 100644
--- a/kernel/cClosure.ml
+++ b/kernel/cClosure.ml
@@ -798,7 +798,7 @@ let drop_parameters depth n argstk =
s.
@assumes [t] is an irreducible term, and not a constructor. [ind] is the inductive
of the constructor term [c]
- @raises Not_found if the inductive is not a primitive record, or if the
+ @raise Not_found if the inductive is not a primitive record, or if the
constructor is partially applied.
*)
let eta_expand_ind_stack env ind m s (f, s') =
diff --git a/kernel/cClosure.mli b/kernel/cClosure.mli
index 71453a04b..b9c71d72a 100644
--- a/kernel/cClosure.mli
+++ b/kernel/cClosure.mli
@@ -216,7 +216,7 @@ val whd_stack :
s.
@assumes [t] is a rigid term, and not a constructor. [ind] is the inductive
of the constructor term [c]
- @raises Not_found if the inductive is not a primitive record, or if the
+ @raise Not_found if the inductive is not a primitive record, or if the
constructor is partially applied.
*)
val eta_expand_ind_stack : env -> inductive -> fconstr -> stack ->
diff --git a/kernel/environ.mli b/kernel/environ.mli
index 51388b8f3..4e6ac1e72 100644
--- a/kernel/environ.mli
+++ b/kernel/environ.mli
@@ -201,7 +201,7 @@ val lookup_modtype : ModPath.t -> env -> module_type_body
(** {5 Universe constraints } *)
(** Add universe constraints to the environment.
- @raises UniverseInconsistency
+ @raise UniverseInconsistency .
*)
val add_constraints : Univ.Constraint.t -> env -> env
diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli
index 97745771e..d4fba63fb 100644
--- a/kernel/uGraph.mli
+++ b/kernel/uGraph.mli
@@ -43,7 +43,7 @@ val check_constraint : t -> univ_constraint -> bool
val check_constraints : Constraint.t -> t -> bool
(** Adds a universe to the graph, ensuring it is >= or > Set.
- @raises AlreadyDeclared if the level is already declared in the graph. *)
+ @raise AlreadyDeclared if the level is already declared in the graph. *)
exception AlreadyDeclared
diff --git a/lib/lib.mllib b/lib/lib.mllib
index b2260ba09..089185942 100644
--- a/lib/lib.mllib
+++ b/lib/lib.mllib
@@ -15,7 +15,6 @@ CWarnings
Rtree
System
Explore
-RTree
CProfile
Future
Spawn
diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml
index 5163ec7b3..007d139a3 100644
--- a/plugins/ssr/ssrcommon.ml
+++ b/plugins/ssr/ssrcommon.ml
@@ -460,7 +460,7 @@ let red_product_skip_id env sigma c = match EConstr.kind sigma c with
let ssrevaltac ist gtac = Tacinterp.tactic_of_value ist gtac
-(** Open term to lambda-term coercion {{{ ************************************)
+(** Open term to lambda-term coercion *)(* {{{ ************************************)
(* This operation takes a goal gl and an open term (sigma, t), and *)
(* returns a term t' where all the new evars in sigma are abstracted *)
@@ -1000,7 +1000,7 @@ let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl =
try applyn ~with_evars ~with_shelve:true ?beta n (EConstr.of_constr oc) gl
with e when CErrors.noncritical e -> raise dependent_apply_error
-(** Profiling {{{ *************************************************************)
+(** Profiling *)(* {{{ *************************************************************)
type profiler = {
profile : 'a 'b. ('a -> 'b) -> 'a -> 'b;
reset : unit -> unit;
@@ -1128,7 +1128,7 @@ let interp_clr sigma = function
(** Basic tacticals *)
-(** Multipliers {{{ ***********************************************************)
+(** Multipliers *)(* {{{ ***********************************************************)
(* tactical *)
@@ -1168,7 +1168,7 @@ let tclMULT = function
let old_cleartac clr = check_hyps_uniq [] clr; Proofview.V82.of_tactic (Tactics.clear (hyps_ids clr))
let cleartac clr = check_hyps_uniq [] clr; Tactics.clear (hyps_ids clr)
-(** }}} *)
+(* }}} *)
(** Generalize tactic *)
diff --git a/plugins/ssr/ssrparser.ml4 b/plugins/ssr/ssrparser.ml4
index 70f73c1fe..2bed8b624 100644
--- a/plugins/ssr/ssrparser.ml4
+++ b/plugins/ssr/ssrparser.ml4
@@ -952,7 +952,7 @@ let pr_ssrhint _ _ = pr_hint
ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY pr_ssrhint
| [ ] -> [ nohint ]
END
-(** The "in" pseudo-tactical {{{ **********************************************)
+(** The "in" pseudo-tactical *)(* {{{ **********************************************)
(* We can't make "in" into a general tactical because this would create a *)
(* crippling conflict with the ltac let .. in construct. Hence, we add *)
@@ -1438,7 +1438,7 @@ let tactic_expr = Pltac.tactic_expr
let old_tac = V82.tactic
-(** Name generation {{{ *******************************************************)
+(** Name generation *)(* {{{ *******************************************************)
(* Since Coq now does repeated internal checks of its external lexical *)
(* rules, we now need to carve ssreflect reserved identifiers out of *)
@@ -1490,7 +1490,7 @@ let _ = add_internal_name (is_tagged perm_tag)
(* We must not anonymize context names discharged by the "in" tactical. *)
-(** Tactical extensions. {{{ **************************************************)
+(** Tactical extensions. *)(* {{{ **************************************************)
(* The TACTIC EXTEND facility can't be used for defining new user *)
(* tacticals, because: *)
diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml
index 21ad6e6ce..cdf1c6f95 100644
--- a/plugins/ssr/ssrtacticals.ml
+++ b/plugins/ssr/ssrtacticals.ml
@@ -59,7 +59,7 @@ let tclSEQAT ist atac1 dir (ivar, ((_, atacs2), atac3)) =
| L2R, pad, tacs2 -> Tacticals.tclTHENSFIRSTn tac1 (Array.of_list (pad @ tacs2)) tac3
| R2L, pad, tacs2 -> Tacticals.tclTHENSLASTn tac1 tac3 (Array.of_list (tacs2 @ pad))
-(** The "in" pseudo-tactical {{{ **********************************************)
+(** The "in" pseudo-tactical *)(* {{{ **********************************************)
let hidden_goal_tag = "the_hidden_goal"
diff --git a/plugins/ssr/ssrvernac.ml4 b/plugins/ssr/ssrvernac.ml4
index f37452613..e3941c1c5 100644
--- a/plugins/ssr/ssrvernac.ml4
+++ b/plugins/ssr/ssrvernac.ml4
@@ -49,7 +49,7 @@ let frozen_lexer = CLexer.get_keyword_state () ;;
(* global syntactic changes and vernacular commands *)
-(** Alternative notations for "match" and anonymous arguments. {{{ ************)
+(** Alternative notations for "match" and anonymous arguments. *)(* {{{ ************)
(* Syntax: *)
(* if <term> is <pattern> then ... else ... *)
@@ -127,7 +127,7 @@ GEXTEND Gram
END
(* }}} *)
-(** Vernacular commands: Prenex Implicits and Search {{{ **********************)
+(** Vernacular commands: Prenex Implicits and Search *)(* {{{ **********************)
(* This should really be implemented as an extension to the implicit *)
(* arguments feature, but unfortuately that API is sealed. The current *)
@@ -461,7 +461,7 @@ END
(* }}} *)
-(** View hint database and View application. {{{ ******************************)
+(** View hint database and View application. *)(* {{{ ******************************)
(* There are three databases of lemmas used to mediate the application *)
(* of reflection lemmas: one for forward chaining, one for backward *)
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 1f1a63dac..62c35d6df 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -70,7 +70,7 @@ let _ =
Goptions.optwrite = debug }
let pp s = !pp_ref s
-(** Utils {{{ *****************************************************************)
+(** Utils *)(* {{{ *****************************************************************)
let env_size env = List.length (Environ.named_context env)
let safeDestApp c =
match kind c with App (f, a) -> f, a | _ -> c, [| |]
@@ -179,7 +179,7 @@ let nf_evar sigma c =
(* }}} *)
-(** Profiling {{{ *************************************************************)
+(** Profiling *)(* {{{ *************************************************************)
type profiler = {
profile : 'a 'b. ('a -> 'b) -> 'a -> 'b;
reset : unit -> unit;
diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli
index cd5676f28..07d0f9757 100644
--- a/plugins/ssrmatching/ssrmatching.mli
+++ b/plugins/ssrmatching/ssrmatching.mli
@@ -74,7 +74,7 @@ val interp_cpattern :
pattern
(** The set of occurrences to be matched. The boolean is set to true
- * to signal the complement of this set (i.e. {-1 3}) *)
+ * to signal the complement of this set (i.e. \{-1 3\}) *)
type occ = (bool * int list) option
(** [subst e p t i]. [i] is the number of binders
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 16d75685d..3b56513f5 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -278,7 +278,7 @@ val check_conv : ?pb:conv_pb -> ?ts:transparent_state -> env -> evar_map -> con
(** [infer_conv] Adds necessary universe constraints to the evar map.
pb defaults to CUMUL and ts to a full transparent state.
- @raises UniverseInconsistency iff catch_incon is set to false,
+ @raise UniverseInconsistency iff catch_incon is set to false,
otherwise returns false in that case.
*)
val infer_conv : ?catch_incon:bool -> ?pb:conv_pb -> ?ts:transparent_state ->
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index a59046a67..b012a7ecd 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -511,10 +511,10 @@ let coq_eqdec ~sum ~rev =
mkPattern (mkGAppRef sum args)
)
-(** { ?X2 = ?X3 :> ?X1 } + { ~ ?X2 = ?X3 :> ?X1 } *)
+(** [{ ?X2 = ?X3 :> ?X1 } + { ~ ?X2 = ?X3 :> ?X1 }] *)
let coq_eqdec_inf_pattern = coq_eqdec ~sum:coq_sumbool_ref ~rev:false
-(** { ~ ?X2 = ?X3 :> ?X1 } + { ?X2 = ?X3 :> ?X1 } *)
+(** [{ ~ ?X2 = ?X3 :> ?X1 } + { ?X2 = ?X3 :> ?X1 }] *)
let coq_eqdec_inf_rev_pattern = coq_eqdec ~sum:coq_sumbool_ref ~rev:true
(** %coq_or_ref (?X2 = ?X3 :> ?X1) (~ ?X2 = ?X3 :> ?X1) *)