aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins
diff options
context:
space:
mode:
Diffstat (limited to 'plugins')
-rw-r--r--plugins/cc/ccalgo.ml16
-rw-r--r--plugins/cc/ccproof.ml4
-rw-r--r--plugins/extraction/modutil.ml4
-rw-r--r--plugins/extraction/table.ml2
-rw-r--r--plugins/firstorder/instances.ml2
-rw-r--r--plugins/firstorder/rules.ml2
-rw-r--r--plugins/funind/functional_principles_proofs.ml10
-rw-r--r--plugins/funind/functional_principles_types.ml4
-rw-r--r--plugins/funind/glob_term_to_relation.ml2
-rw-r--r--plugins/funind/glob_termops.ml4
-rw-r--r--plugins/funind/indfun.ml12
-rw-r--r--plugins/funind/indfun_common.ml4
-rw-r--r--plugins/funind/invfun.ml10
-rw-r--r--plugins/funind/recdef.ml18
-rw-r--r--plugins/ltac/pptactic.ml2
-rw-r--r--plugins/ltac/profile_ltac.ml4
-rw-r--r--plugins/ltac/rewrite.ml4
-rw-r--r--plugins/ltac/tacenv.ml4
-rw-r--r--plugins/ltac/tacinterp.ml2
-rw-r--r--plugins/omega/coq_omega.ml4
-rw-r--r--plugins/quote/quote.ml2
-rw-r--r--plugins/rtauto/proof_search.ml10
-rw-r--r--plugins/setoid_ring/newring.ml2
-rw-r--r--plugins/ssrmatching/ssrmatching.ml426
24 files changed, 77 insertions, 77 deletions
diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml
index 5dea4631c..ba398c385 100644
--- a/plugins/cc/ccalgo.ml
+++ b/plugins/cc/ccalgo.ml
@@ -61,7 +61,7 @@ module ST=struct
let enter t sign st=
if IntPairTable.mem st.toterm sign then
- anomaly ~label:"enter" (Pp.str "signature already entered")
+ anomaly ~label:"enter" (Pp.str "signature already entered.")
else
IntPairTable.replace st.toterm sign t;
IntTable.replace st.tosign t sign
@@ -321,7 +321,7 @@ let find uf i= find_aux uf [] i
let get_representative uf i=
match uf.map.(i).clas with
Rep r -> r
- | _ -> anomaly ~label:"get_representative" (Pp.str "not a representative")
+ | _ -> anomaly ~label:"get_representative" (Pp.str "not a representative.")
let get_constructors uf i= uf.map.(i).constructors
@@ -339,7 +339,7 @@ let rec find_oldest_pac uf i pac=
let get_constructor_info uf i=
match uf.map.(i).term with
Constructor cinfo->cinfo
- | _ -> anomaly ~label:"get_constructor" (Pp.str "not a constructor")
+ | _ -> anomaly ~label:"get_constructor" (Pp.str "not a constructor.")
let size uf i=
(get_representative uf i).weight
@@ -384,7 +384,7 @@ let term uf i=uf.map.(i).term
let subterms uf i=
match uf.map.(i).vertex with
Node(j,k) -> (j,k)
- | _ -> anomaly ~label:"subterms" (Pp.str "not a node")
+ | _ -> anomaly ~label:"subterms" (Pp.str "not a node.")
let signature uf i=
let j,k=subterms uf i in (find uf j,find uf k)
@@ -485,7 +485,7 @@ let build_subst uf subst =
(fun i ->
try term uf i
with e when CErrors.noncritical e ->
- anomaly (Pp.str "incomplete matching"))
+ anomaly (Pp.str "incomplete matching."))
subst
let rec inst_pattern subst = function
@@ -750,7 +750,7 @@ let process_constructor_mark t i rep pac state =
state.combine;
f (n-1) q1 q2
| _-> anomaly ~label:"add_pacs"
- (Pp.str "weird error in injection subterms merge")
+ (Pp.str "weird error in injection subterms merge.")
in f cinfo.ci_nhyps opac.args pac.args
| Partial_applied | Partial _ ->
(* add_pac state.uf.map.(i) pac t; *)
@@ -841,7 +841,7 @@ let complete_one_class state i=
let ct = app (term state.uf i) typ pac.arity in
state.uf.epsilons <- pac :: state.uf.epsilons;
ignore (add_term state ct)
- | _ -> anomaly (Pp.str "wrong incomplete class")
+ | _ -> anomaly (Pp.str "wrong incomplete class.")
let complete state =
Int.Set.iter (complete_one_class state) state.pa_classes
@@ -981,7 +981,7 @@ let find_instances state =
Control.check_for_interrupt ();
do_match state res pb_stack
done;
- anomaly (Pp.str "get out of here !")
+ anomaly (Pp.str "get out of here!")
with Stack.Empty -> () in
!res
diff --git a/plugins/cc/ccproof.ml b/plugins/cc/ccproof.ml
index f58847caf..642ceba3d 100644
--- a/plugins/cc/ccproof.ml
+++ b/plugins/cc/ccproof.ml
@@ -47,7 +47,7 @@ let rec ptrans p1 p3=
{p_lhs=p1.p_lhs;
p_rhs=p3.p_rhs;
p_rule=Trans (p1,p3)}
- else anomaly (Pp.str "invalid cc transitivity")
+ else anomaly (Pp.str "invalid cc transitivity.")
let rec psym p =
match p.p_rule with
@@ -85,7 +85,7 @@ let rec nth_arg t n=
if n>0 then
nth_arg t1 (n-1)
else t2
- | _ -> anomaly ~label:"nth_arg" (Pp.str "not enough args")
+ | _ -> anomaly ~label:"nth_arg" (Pp.str "not enough args.")
let pinject p c n a =
{p_lhs=nth_arg p.p_lhs (n-a);
diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml
index 60fe8e762..b67b9931e 100644
--- a/plugins/extraction/modutil.ml
+++ b/plugins/extraction/modutil.ml
@@ -19,7 +19,7 @@ open Mlutil
let rec msid_of_mt = function
| MTident mp -> mp
| MTwith(mt,_)-> msid_of_mt mt
- | _ -> anomaly ~label:"extraction" (Pp.str "the With operator isn't applied to a name")
+ | _ -> anomaly ~label:"extraction" (Pp.str "the With operator isn't applied to a name.")
(*s Apply some functions upon all [ml_decl] and [ml_spec] found in a
[ml_structure]. *)
@@ -231,7 +231,7 @@ let get_decl_in_structure r struc =
| _ -> error_not_visible r
in go ll sel
with Not_found ->
- anomaly (Pp.str "reference not found in extracted structure")
+ anomaly (Pp.str "reference not found in extracted structure.")
(*s Optimization of a [ml_structure]. *)
diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml
index a369cbdf3..29dd8ff4f 100644
--- a/plugins/extraction/table.ml
+++ b/plugins/extraction/table.ml
@@ -261,7 +261,7 @@ let safe_basename_of_global r =
let last_chance r =
try Nametab.basename_of_global r
with Not_found ->
- anomaly (Pp.str "Inductive object unknown to extraction and not globally visible")
+ anomaly (Pp.str "Inductive object unknown to extraction and not globally visible.")
in
match r with
| ConstRef kn -> Label.to_id (con_label kn)
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index 4c6355f61..04bca584f 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -77,7 +77,7 @@ let match_one_quantified_hyp sigma setref seq lf=
Left(Lforall(i,dom,triv))|Right(Rexists(i,dom,triv))->
if do_sequent sigma setref triv lf.id seq i dom lf.atoms then
setref:=IS.add ((Phantom dom),lf.id) !setref
- | _ -> anomaly (Pp.str "can't happen")
+ | _ -> anomaly (Pp.str "can't happen.")
let give_instances sigma lf seq=
let setref=ref IS.empty in
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index 8c6b5b91d..f745dbeb4 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -40,7 +40,7 @@ let wrap n b continue seq =
let rec aux i nc ctx=
if i<=0 then seq else
match nc with
- []->anomaly (Pp.str "Not the expected number of hyps")
+ []->anomaly (Pp.str "Not the expected number of hyps.")
| nd::q->
let id = NamedDecl.get_id nd in
if occur_var env sigma id (pf_concl gls) ||
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index 0041797de..c2f705898 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -397,7 +397,7 @@ let rewrite_until_var arg_num eq_ids : tactic =
then tclIDTAC g
else
match eq_ids with
- | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property");
+ | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property.");
| eq_id::eq_ids ->
tclTHEN
(tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar eq_id))))
@@ -605,7 +605,7 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
observe (str "cannot compute new term value : " ++ pr_gls g' ++ fnl () ++ str "last hyp is" ++
pr_leconstr_env (pf_env g') (project g') new_term_value_eq
);
- anomaly (Pp.str "cannot compute new term value")
+ anomaly (Pp.str "cannot compute new term value.")
in
let fun_body =
mkLambda(Anonymous,
@@ -838,7 +838,7 @@ let build_proof
h_reduce_with_zeta Locusops.onConcl;
build_proof do_finalize new_infos
] g
- | Rel _ -> anomaly (Pp.str "Free var in goal conclusion !")
+ | Rel _ -> anomaly (Pp.str "Free var in goal conclusion!")
and build_proof do_finalize dyn_infos g =
(* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *)
observe_tac_stream (str "build_proof with " ++ Printer.pr_leconstr dyn_infos.info ) (build_proof_aux do_finalize dyn_infos) g
@@ -1032,7 +1032,7 @@ let do_replace (evd:Evd.evar_map ref) params rec_arg_num rev_args_id f fun_num a
{finfos with
equation_lemma = Some (match Nametab.locate (qualid_of_ident equation_lemma_id) with
ConstRef c -> c
- | _ -> CErrors.anomaly (Pp.str "Not a constant")
+ | _ -> CErrors.anomaly (Pp.str "Not a constant.")
)
}
| _ -> ()
@@ -1255,7 +1255,7 @@ let prove_princ_for_struct (evd:Evd.evar_map ref) interactive_proof fun_num fnam
try
let pte =
try destVar (project gl) pte
- with DestKO -> anomaly (Pp.str "Property is not a variable")
+ with DestKO -> anomaly (Pp.str "Property is not a variable.")
in
let fix_info = Id.Map.find pte ptes_to_fix in
let nb_args = fix_info.nb_realargs in
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 942527167..a0eb9e2b2 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -44,7 +44,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
let id = Namegen.next_ident_away x avoid in
Hashtbl.add tbl id x;
RelDecl.set_name (Name id) decl :: change_predicates_names (id::avoid) predicates
- | Anonymous -> anomaly (Pp.str "Anonymous property binder "))
+ | Anonymous -> anomaly (Pp.str "Anonymous property binder."))
in
let avoid = (Termops.ids_of_context env_with_params ) in
let princ_type_info =
@@ -399,7 +399,7 @@ let get_funs_constant mp dp =
let const = make_con mp dp (Label.of_id id) in
const,i
| Anonymous ->
- anomaly (Pp.str "Anonymous fix")
+ anomaly (Pp.str "Anonymous fix.")
)
na
| _ -> [|const,0|]
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 68e097fe9..aa8f918ae 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1115,7 +1115,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude)
else
CAst.make @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
- | _ -> anomaly (Pp.str "Should not have an anonymous function here")
+ | _ -> anomaly (Pp.str "Should not have an anonymous function here.")
(* We have renamed all the anonymous functions during alpha_renaming phase *)
end
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 0361e8cb1..b99f5a923 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -532,7 +532,7 @@ let rec are_unifiable_aux = function
else
let eqs' =
try (List.combine cpl1 cpl2) @ eqs
- with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux")
+ with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux.")
in
are_unifiable_aux eqs'
@@ -555,7 +555,7 @@ let rec eq_cases_pattern_aux = function
else
let eqs' =
try (List.combine cpl1 cpl2) @ eqs
- with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux")
+ with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux.")
in
eq_cases_pattern_aux eqs'
| _ -> raise NotUnifiable
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 4946285e1..143ce8288 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -734,7 +734,7 @@ let rec add_args id new_args = CAst.map (function
CAppExpl((None,r,None),new_args)
| _ -> b
end
- | CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo")
+ | CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo.")
| CProdN(nal,b1) ->
CProdN(List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
add_args id new_args b1)
@@ -782,9 +782,9 @@ let rec add_args id new_args = CAst.map (function
Miscops.map_cast_type (add_args id new_args) b2)
| CRecord pars ->
CRecord (List.map (fun (e,o) -> e, add_args id new_args o) pars)
- | CNotation _ -> anomaly ~label:"add_args " (Pp.str "CNotation")
- | CGeneralization _ -> anomaly ~label:"add_args " (Pp.str "CGeneralization")
- | CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters")
+ | CNotation _ -> anomaly ~label:"add_args " (Pp.str "CNotation.")
+ | CGeneralization _ -> anomaly ~label:"add_args " (Pp.str "CGeneralization.")
+ | CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters.")
)
exception Stop of Constrexpr.constr_expr
@@ -826,7 +826,7 @@ let rec chop_n_arrow n t =
chop_n_arrow new_n t'
with Stop t -> t
end
- | _ -> anomaly (Pp.str "Not enough products")
+ | _ -> anomaly (Pp.str "Not enough products.")
let rec get_args b t : Constrexpr.local_binder_expr list *
@@ -856,7 +856,7 @@ let make_graph (f_ref:global_reference) =
| _ -> raise (UserError (None, str "Not a function reference") )
in
(match Global.body_of_constant_body c_body with
- | None -> error "Cannot build a graph over an axiom !"
+ | None -> error "Cannot build a graph over an axiom!"
| Some body ->
let env = Global.env () in
let sigma = Evd.from_env env in
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 2476478ab..a73425543 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -369,7 +369,7 @@ let in_Function : function_info -> Libobject.obj =
let find_or_none id =
try Some
- (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> CErrors.anomaly (Pp.str "Not a constant")
+ (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> CErrors.anomaly (Pp.str "Not a constant.")
)
with Not_found -> None
@@ -397,7 +397,7 @@ let add_Function is_general f =
and prop_lemma = find_or_none (Nameops.add_suffix f_id "_ind")
and graph_ind =
match Nametab.locate (qualid_of_ident (mk_rel_id f_id))
- with | IndRef ind -> ind | _ -> CErrors.anomaly (Pp.str "Not an inductive")
+ with | IndRef ind -> ind | _ -> CErrors.anomaly (Pp.str "Not an inductive.")
in
let finfos =
{ function_constant = f;
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 12232dd83..f373e486c 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -140,7 +140,7 @@ let generate_type evd g_to_f f graph i =
let ctxt,_ = decompose_prod_assum !evd graph_arity in
let fun_ctxt,res_type =
match ctxt with
- | [] | [_] -> anomaly (Pp.str "Not a valid context")
+ | [] | [_] -> anomaly (Pp.str "Not a valid context.")
| decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl
in
let rec args_from_decl i accu = function
@@ -292,7 +292,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(fun (_,pat) acc ->
match pat with
| IntroNaming (IntroIdentifier id) -> id::acc
- | _ -> anomaly (Pp.str "Not an identifier")
+ | _ -> anomaly (Pp.str "Not an identifier.")
)
(List.nth intro_pats (pred i))
[]
@@ -401,7 +401,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
Array.map
(fun ((_,(ctxt,concl))) ->
match ctxt with
- | [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
+ | [] | [_] | [_;_] -> anomaly (Pp.str "bad context.")
| hres::res::decl::ctxt ->
let res = EConstr.it_mkLambda_or_LetIn
(EConstr.it_mkProd_or_LetIn concl [hres;res])
@@ -708,7 +708,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
then
let eq_lemma =
try Option.get (infos).equation_lemma
- with Option.IsNone -> anomaly (Pp.str "Cannot find equation lemma")
+ with Option.IsNone -> anomaly (Pp.str "Cannot find equation lemma.")
in
tclTHENSEQ[
tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids;
@@ -938,7 +938,7 @@ let revert_graph kn post_tac hid g =
let info =
try find_Function_of_graph ind'
with Not_found -> (* The graphs are mutually recursive but we cannot find one of them !*)
- anomaly (Pp.str "Cannot retrieve infos about a mutual block")
+ anomaly (Pp.str "Cannot retrieve infos about a mutual block.")
in
(* if we can find a completeness lemma for this function
then we can come back to the functional form. If not, we do nothing
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index 62eba9513..e6f199dba 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -76,7 +76,7 @@ let def_of_const t =
| _ -> raise Not_found)
with Not_found ->
anomaly (str "Cannot find definition of constant " ++
- (Id.print (Label.to_id (con_label (fst sp)))))
+ (Id.print (Label.to_id (con_label (fst sp)))) ++ str ".")
)
|_ -> assert false
@@ -95,7 +95,7 @@ let constant sl s = constr_of_global (find_reference sl s)
let const_of_ref = function
ConstRef kn -> kn
- | _ -> anomaly (Pp.str "ConstRef expected")
+ | _ -> anomaly (Pp.str "ConstRef expected.")
let nf_zeta env =
@@ -442,7 +442,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
travel jinfo new_continuation_tac
{expr_info with info = b; is_final=false} g
end
- | Rel _ -> anomaly (Pp.str "Free var in goal conclusion !")
+ | Rel _ -> anomaly (Pp.str "Free var in goal conclusion!")
| Prod _ ->
begin
try
@@ -486,7 +486,7 @@ let rec travel_aux jinfo continuation_tac (expr_info:constr infos) g =
travel_args jinfo
expr_info.is_main_branch new_continuation_tac new_infos g
| Case _ -> user_err ~hdr:"Recdef.travel" (str "the term " ++ Printer.pr_leconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
- | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr expr_info.info)
+ | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_leconstr expr_info.info ++ Pp.str ".")
end
| Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t} g
| Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ ->
@@ -1165,7 +1165,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
let f_id =
match f_name with
| Name f_id -> next_ident_away_in_goal f_id ids
- | Anonymous -> anomaly (Pp.str "Anonymous function")
+ | Anonymous -> anomaly (Pp.str "Anonymous function.")
in
let n_names_types,_ = decompose_lam_n sigma nb_args body1 in
let n_ids,ids =
@@ -1175,7 +1175,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
| Name id ->
let n_id = next_ident_away_in_goal id ids in
n_id::n_ids,n_id::ids
- | _ -> anomaly (Pp.str "anonymous argument")
+ | _ -> anomaly (Pp.str "anonymous argument.")
)
([],(f_id::ids))
n_names_types
@@ -1302,7 +1302,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
| None ->
try add_suffix current_proof_name "_subproof"
with e when CErrors.noncritical e ->
- anomaly (Pp.str "open_new_goal with an unamed theorem")
+ anomaly (Pp.str "open_new_goal with an unamed theorem.")
in
let na = next_global_ident_away name [] in
if Termops.occur_existential sigma gls_type then
@@ -1313,7 +1313,7 @@ let open_new_goal build_proof sigma using_lemmas ref_ goal_name (gls_type,decomp
let na_global = Smartlocate.global_with_alias na_ref in
match na_global with
ConstRef c -> is_opaque_constant c
- | _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant")
+ | _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant.")
in
let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in
ref_ := Value (EConstr.Unsafe.to_constr lemma);
@@ -1464,7 +1464,7 @@ let (com_eqn : int -> Id.t ->
let opacity =
match terminate_ref with
| ConstRef c -> is_opaque_constant c
- | _ -> anomaly ~label:"terminate_lemma" (Pp.str "not a constant")
+ | _ -> anomaly ~label:"terminate_lemma" (Pp.str "not a constant.")
in
let (evmap, env) = Lemmas.get_current_context() in
let evmap = Evd.from_ctx (Evd.evar_universe_context evmap) in
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 580c21d40..d66298344 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -582,7 +582,7 @@ type 'a extra_genarg_printer =
hv 0
(pr_let_clause (if recflag then "let rec" else "let") pr hd ++
prlist (fun t -> spc () ++ pr_let_clause "with" pr t) tl)
- | [] -> anomaly (Pp.str "LetIn must declare at least one binding")
+ | [] -> anomaly (Pp.str "LetIn must declare at least one binding.")
let pr_seq_body pr tl =
hv 0 (str "[ " ++
diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml
index 3ff7b53c7..b237e917d 100644
--- a/plugins/ltac/profile_ltac.ml
+++ b/plugins/ltac/profile_ltac.ml
@@ -113,7 +113,7 @@ let rec to_ltacprof_tactic m xml =
children = List.fold_left to_ltacprof_tactic M.empty xs;
} in
M.add name node m
- | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof_tactic XML")
+ | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof_tactic XML.")
let to_ltacprof_results xml =
let open Xml_datatype in
@@ -125,7 +125,7 @@ let to_ltacprof_results xml =
max_total = 0.0;
local = 0.0;
children = List.fold_left to_ltacprof_tactic M.empty xs }
- | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof XML")
+ | _ -> CErrors.anomaly Pp.(str "Malformed ltacprof XML.")
let feedback_results results =
Feedback.(feedback
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index dadcfb9f2..f028abde9 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -236,7 +236,7 @@ end) = struct
let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in
if Option.is_empty obj then evars, mkProd(na, ty, b), arg', (ty, None) :: cstrs
else user_err Pp.(str "build_signature: no constraint can apply on a dependent argument")
- | _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products")
+ | _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products.")
| _, [] ->
(match finalcstr with
| None | Some (_, None) ->
@@ -1424,7 +1424,7 @@ module Strategies =
let unfolded =
try Tacred.try_red_product env sigma c
with e when CErrors.noncritical e ->
- user_err Pp.(str "fold: the term is not unfoldable !")
+ user_err Pp.(str "fold: the term is not unfoldable!")
in
try
let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in
diff --git a/plugins/ltac/tacenv.ml b/plugins/ltac/tacenv.ml
index 2d2dcd0c0..efb7e780d 100644
--- a/plugins/ltac/tacenv.ml
+++ b/plugins/ltac/tacenv.ml
@@ -24,7 +24,7 @@ let register_alias key tac =
let interp_alias key =
try KNmap.find key !alias_map
- with Not_found -> CErrors.anomaly (str "Unknown tactic alias: " ++ KerName.print key)
+ with Not_found -> CErrors.anomaly (str "Unknown tactic alias: " ++ KerName.print key ++ str ".")
let check_alias key = KNmap.mem key !alias_map
@@ -55,7 +55,7 @@ let register_ml_tactic ?(overwrite = false) s (t : ml_tactic array) =
if overwrite then
tac_tab := MLTacMap.remove s !tac_tab
else
- CErrors.anomaly (str "Cannot redeclare tactic " ++ pr_tacname s)
+ CErrors.anomaly (str "Cannot redeclare tactic " ++ pr_tacname s ++ str ".")
in
tac_tab := MLTacMap.add s t !tac_tab
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 594c4fa15..2c09a8b69 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -379,7 +379,7 @@ let try_interp_ltac_var coerce ist env (loc,id) =
let interp_ltac_var coerce ist env locid =
try try_interp_ltac_var coerce ist env locid
- with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time")
+ with Not_found -> anomaly (str "Detected '" ++ Id.print (snd locid) ++ str "' as ltac var at interning time.")
let interp_ident ist env sigma id =
try try_interp_ltac_var (coerce_var_to_ident false env sigma) ist (Some (env,sigma)) (Loc.tag id)
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index d7408e88e..0cd18ae50 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -362,7 +362,7 @@ let coq_True = lazy (init_constant "True")
let evaluable_ref_of_constr s c = match EConstr.kind Evd.empty (Lazy.force c) with
| Const (kn,u) when Tacred.is_evaluable (Global.env()) (EvalConstRef kn) ->
EvalConstRef kn
- | _ -> anomaly ~label:"Coq_omega" (Pp.str (s^" is not an evaluable constant"))
+ | _ -> anomaly ~label:"Coq_omega" (Pp.str (s^" is not an evaluable constant."))
let sp_Zsucc = lazy (evaluable_ref_of_constr "Z.succ" coq_Zsucc)
let sp_Zpred = lazy (evaluable_ref_of_constr "Z.pred" coq_Zpred)
@@ -630,7 +630,7 @@ let compile name kind =
let id = new_id () in
tag_hypothesis name id;
{kind = kind; body = List.rev accu; constant = n; id = id}
- | _ -> anomaly (Pp.str "compile_equation")
+ | _ -> anomaly (Pp.str "compile_equation.")
in
loop []
diff --git a/plugins/quote/quote.ml b/plugins/quote/quote.ml
index ba8356b52..59ed8439b 100644
--- a/plugins/quote/quote.ml
+++ b/plugins/quote/quote.ml
@@ -423,7 +423,7 @@ let quote_terms env sigma ivs lc =
| None ->
begin match ivs.constant_lhs with
| Some c_lhs -> subst_meta [1, c] c_lhs
- | None -> anomaly (Pp.str "invalid inversion scheme for quote")
+ | None -> anomaly (Pp.str "invalid inversion scheme for quote.")
end
| Some var_lhs ->
begin match ivs.constant_lhs with
diff --git a/plugins/rtauto/proof_search.ml b/plugins/rtauto/proof_search.ml
index 4eef1b0a7..153a6a49a 100644
--- a/plugins/rtauto/proof_search.ml
+++ b/plugins/rtauto/proof_search.ml
@@ -145,7 +145,7 @@ let add_step s sub =
| SI_Or_r,[p] -> I_Or_r p
| SE_Or i,[p1;p2] -> E_Or(i,p1,p2)
| SD_Or i,[p] -> D_Or(i,p)
- | _,_ -> anomaly ~label:"add_step" (Pp.str "wrong arity")
+ | _,_ -> anomaly ~label:"add_step" (Pp.str "wrong arity.")
type 'a with_deps =
{dep_it:'a;
@@ -167,7 +167,7 @@ type state =
let project = function
Complete prf -> prf
- | Incomplete (_,_) -> anomaly (Pp.str "not a successful state")
+ | Incomplete (_,_) -> anomaly (Pp.str "not a successful state.")
let pop n prf =
let nprf=
@@ -361,7 +361,7 @@ let search_norev seq=
(Arrow(f2,f3)))
f1;
add_hyp (embed nseq) f3]):: !goals
- | _ -> anomaly ~label:"search_no_rev" (Pp.str "can't happen") in
+ | _ -> anomaly ~label:"search_no_rev" (Pp.str "can't happen.") in
Int.Map.iter add_one seq.norev_hyps;
List.rev !goals
@@ -386,7 +386,7 @@ let search_in_rev_hyps seq=
| Arrow (Disjunct (f1,f2),f0) ->
[make_step (SD_Or(i)),
[add_hyp (add_hyp (embed nseq) (Arrow(f1,f0))) (Arrow (f2,f0))]]
- | _ -> anomaly ~label:"search_in_rev_hyps" (Pp.str "can't happen")
+ | _ -> anomaly ~label:"search_in_rev_hyps" (Pp.str "can't happen.")
with
Not_found -> search_norev seq
@@ -464,7 +464,7 @@ let branching = function
| _::next ->
s_info.nd_branching<-s_info.nd_branching+List.length next in
List.map (append stack) successors
- | Complete prf -> anomaly (Pp.str "already succeeded")
+ | Complete prf -> anomaly (Pp.str "already succeeded.")
open Pp
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 38f05978d..cca5cde15 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -47,7 +47,7 @@ let tag_arg tag_rec map subs i c =
let global_head_of_constr sigma c =
let f, args = decompose_app sigma c in
try fst (Termops.global_of_constr sigma f)
- with Not_found -> CErrors.anomaly (str "global_head_of_constr")
+ with Not_found -> CErrors.anomaly (str "global_head_of_constr.")
let global_of_constr_nofail c =
try global_of_constr c
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 6b752fb4b..67e6c7e93 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -133,7 +133,7 @@ let dC t = CastConv t
(** Constructors for constr_expr *)
let isCVar = function { CAst.v = CRef (Ident _, _) } -> true | _ -> false
let destCVar = function { CAst.v = CRef (Ident (_, id), _) } -> id | _ ->
- CErrors.anomaly (str"not a CRef")
+ CErrors.anomaly (str"not a CRef.")
let mkCHole ~loc = CAst.make ?loc @@ CHole (None, IntroAnonymous, None)
let mkCLambda ?loc name ty t = CAst.make ?loc @@
CLambdaN ([[Loc.tag ?loc name], Default Explicit, ty], t)
@@ -150,8 +150,8 @@ let mkRLambda n s t = CAst.make @@ GLambda (n, Explicit, s, t)
let combineCG t1 t2 f g = match t1, t2 with
| (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None)
| (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2))
- | _, (_, (_, None)) -> CErrors.anomaly (str"have: mixed C-G constr")
- | _ -> CErrors.anomaly (str"have: mixed G-C constr")
+ | _, (_, (_, None)) -> CErrors.anomaly (str"have: mixed C-G constr.")
+ | _ -> CErrors.anomaly (str"have: mixed G-C constr.")
let loc_ofCG = function
| (_, (s, None)) -> Glob_ops.loc_of_glob_constr s
| (_, (_, Some s)) -> Constrexpr_ops.constr_loc s
@@ -620,12 +620,12 @@ let match_upats_FO upats env sigma0 ise orig_c =
let pt' = pi1 pt', pi2 pt', EConstr.Unsafe.to_constr (pi3 pt') in
raise (FoundUnif (ungen_upat lhs pt' u))
with FoundUnif (s,_,_) as sig_u when dont_impact_evars s -> raise sig_u
- | Not_found -> CErrors.anomaly (str"incomplete ise in match_upats_FO")
+ | Not_found -> CErrors.anomaly (str"incomplete ise in match_upats_FO.")
| e when CErrors.noncritical e -> () in
List.iter one_match fpats
done;
iter_constr_LR loop f; Array.iter loop a in
- try loop orig_c with Invalid_argument _ -> CErrors.anomaly (str"IN FO")
+ try loop orig_c with Invalid_argument _ -> CErrors.anomaly (str"IN FO.")
let prof_FO = mk_profiler "match_upats_FO";;
let match_upats_FO upats env sigma0 ise c =
@@ -696,11 +696,11 @@ let fixed_upat = function
let do_once r f = match !r with Some _ -> () | None -> r := Some (f ())
let assert_done r =
- match !r with Some x -> x | None -> CErrors.anomaly (str"do_once never called")
+ match !r with Some x -> x | None -> CErrors.anomaly (str"do_once never called.")
let assert_done_multires r =
match !r with
- | None -> CErrors.anomaly (str"do_once never called")
+ | None -> CErrors.anomaly (str"do_once never called.")
| Some (n, xs) ->
r := Some (n+1,xs);
try List.nth xs n with Failure _ -> raise NoMatch
@@ -757,7 +757,7 @@ let source () = match upats_origin, upats with
| Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++
pr_constr_pat rule ++ spc()
| _, [] | None, _::_::_ ->
- CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin") in
+ CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in
let on_instance, instances =
let instances = ref [] in
(fun x ->
@@ -795,7 +795,7 @@ let rec uniquize = function
errorstrm (source () ++ str "does not match any subterm of the goal")
| NoProgress when (not raise_NoMatch) ->
let dir = match upats_origin with Some (d,_) -> d | _ ->
- CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin") in
+ CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in
errorstrm (str"all matches of "++source()++
str"are equal to the " ++ pr_dir_side (inv_dir dir))
| NoProgress -> raise NoMatch);
@@ -833,7 +833,7 @@ let rec uniquize = function
let sigma, uc, ({up_f = pf; up_a = pa} as u) =
match !upat_that_matched with
| Some (_,x) -> List.hd x | None when raise_NoMatch -> raise NoMatch
- | None -> CErrors.anomaly (str"companion function never called") in
+ | None -> CErrors.anomaly (str"companion function never called.") in
let p' = mkApp (pf, pa) in
if max_occ <= !nocc then p', u.up_dir, (sigma, uc, u.up_t)
else errorstrm (str"Only " ++ int !nocc ++ str" < " ++ int max_occ ++
@@ -920,7 +920,7 @@ let glob_cpattern gs p =
| (r1, Some _), (r2, Some _) when isCVar t1 ->
encode k "In" [r1; r2; bind_in t1 t2]
| (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2]
- | _ -> CErrors.anomaly (str"where are we?")
+ | _ -> CErrors.anomaly (str"where are we?.")
with _ when isCVar t1 -> encode k "In" [bind_in t1 t2])
| CNotation("( _ in _ in _ )", ([t1; t2; t3], [], [])) ->
check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3]
@@ -1094,7 +1094,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
(Value.cast (topwit (Option.get wit_ssrpatternarg)) v)
| it -> g t with e when CErrors.noncritical e -> g t in
let decodeG t f g = decode ist (mkG t) f g in
- let bad_enc id _ = CErrors.anomaly (str"bad encoding for pattern "++str id) in
+ let bad_enc id _ = CErrors.anomaly (str"bad encoding for pattern "++str id++str".") in
let cleanup_XinE h x rp sigma =
let h_k = match kind_of_term h with Evar (k,_) -> k | _ -> assert false in
let to_clean, update = (* handle rename if x is already used *)
@@ -1280,7 +1280,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst =
let redex_of_pattern ?(resolve_typeclasses=false) env (sigma, p) =
let e = match p with
- | In_T _ | In_X_In_T _ -> CErrors.anomaly (str"pattern without redex")
+ | In_T _ | In_X_In_T _ -> CErrors.anomaly (str"pattern without redex.")
| T e | X_In_T (e, _) | E_As_X_In_T (e, _, _) | E_In_X_In_T (e, _, _) -> e in
let sigma =
if not resolve_typeclasses then sigma