From 2826683746569b9d78aa01e319315ab554e1619b Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:36:45 +0200 Subject: Fix omitted labels in function calls --- plugins/ssrmatching/ssrmatching.ml4 | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) (limited to 'plugins/ssrmatching') diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f3555ebc4..fc7963b2d 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -1261,7 +1261,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in let rp = mk_upat_for env0 sigma0 (ise, rp) all_ok in let find_T, end_T = mk_tpattern_matcher ?raise_NoMatch sigma0 occ rp in - let concl = find_T env0 concl0 1 do_subst in + let concl = find_T env0 concl0 1 ~k:do_subst in let _ = end_T () in concl | Some (sigma, (X_In_T (hole, p) | In_X_In_T (hole, p))) -> @@ -1273,11 +1273,11 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = (* we start from sigma, so hole is considered a rigid head *) let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in - let concl = find_T env0 concl0 1 (fun env c _ h -> + let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in fs p_sigma (find_X env (fs sigma p) h - (fun env _ -> do_subst env e_body))) in + ~k:(fun env _ -> do_subst env e_body))) in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_In_X_In_T (e, hole, p)) -> @@ -1289,11 +1289,11 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let find_X, end_X = mk_tpattern_matcher sigma noindex holep in let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in - let concl = find_T env0 concl0 1 (fun env c _ h -> + let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in - fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> - find_E env e_body h do_subst))) in + fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> + find_E env e_body h ~k:do_subst))) in let _ = end_E () in let _ = end_X () in let _ = end_T () in concl | Some (sigma, E_As_X_In_T (e, hole, p)) -> @@ -1306,10 +1306,10 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in let find_X, end_X = mk_tpattern_matcher sigma occ holep in - let concl = find_TE env0 concl0 1 (fun env c _ h -> + let concl = find_TE env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in - fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> + fs p_sigma (find_X env (fs sigma p) h ~k:(fun env c _ h -> let e_sigma = unify_HO env sigma (EConstr.of_constr e_body) (EConstr.of_constr e) in let e_body = fs e_sigma e in do_subst env e_body e_body h))) in @@ -1352,7 +1352,7 @@ let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = let ise, u = mk_tpattern env sigma0 (ise,EConstr.Unsafe.to_constr t) ok L2R p in let find_U, end_U = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in - let concl = find_U env concl h (fun _ _ _ -> mkRel) in + let concl = find_U env concl h ~k:(fun _ _ _ -> mkRel) in let rdx, _, (sigma, uc, p) = end_U () in sigma, uc, EConstr.of_constr p, EConstr.of_constr concl, EConstr.of_constr rdx -- cgit v1.2.3 From 02d2f34e5c84f0169e884c07054a6fbfef9f365c Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:04:58 +0200 Subject: Remove some unused values and types --- checker/closure.ml | 16 -------- checker/declarations.ml | 11 ----- checker/inductive.ml | 7 ---- checker/typeops.ml | 3 -- checker/univ.ml | 7 ---- engine/termops.ml | 24 ----------- ide/coqOps.ml | 3 -- ide/ideutils.ml | 11 ----- interp/notation_ops.ml | 4 -- kernel/names.ml | 2 - kernel/uGraph.mli | 3 ++ kernel/univ.ml | 4 -- lib/feedback.ml | 2 - lib/stateid.ml | 1 - parsing/cLexer.ml4 | 8 ---- parsing/egramcoq.ml | 4 -- parsing/g_vernac.ml4 | 5 --- parsing/pcoq.ml | 7 ---- plugins/funind/functional_principles_proofs.ml | 6 --- plugins/ltac/extratactics.ml4 | 2 - plugins/ltac/profile_ltac.ml | 1 - plugins/ltac/tacentries.ml | 3 -- plugins/ltac/tacintern.ml | 6 --- plugins/ltac/tacinterp.ml | 6 --- plugins/nsatz/ideal.ml | 57 -------------------------- plugins/nsatz/nsatz.ml | 28 ------------- plugins/setoid_ring/newring.ml | 2 - plugins/ssrmatching/ssrmatching.ml4 | 21 ---------- plugins/ssrmatching/ssrmatching.mli | 1 - pretyping/unification.ml | 3 -- printing/ppconstr.ml | 4 -- printing/ppvernac.ml | 4 -- printing/printer.ml | 5 --- proofs/clenv.ml | 5 --- proofs/logic.ml | 5 --- proofs/tacmach.ml | 3 -- stm/stm.ml | 7 ---- tactics/hipattern.ml | 1 - tactics/tactics.ml | 2 - tools/coqdoc/output.mli | 1 - toplevel/vernac.ml | 1 - vernac/auto_ind_decl.ml | 2 - vernac/metasyntax.ml | 2 +- vernac/record.ml | 2 +- vernac/topfmt.ml | 12 ------ 45 files changed, 5 insertions(+), 309 deletions(-) (limited to 'plugins/ssrmatching') diff --git a/checker/closure.ml b/checker/closure.ml index cef1d31a6..b8294e795 100644 --- a/checker/closure.ml +++ b/checker/closure.ml @@ -651,22 +651,6 @@ let drop_parameters depth n argstk = (** Projections and eta expansion *) -let rec get_parameters depth n argstk = - match argstk with - Zapp args::s -> - let q = Array.length args in - if n > q then Array.append args (get_parameters depth (n-q) s) - else if Int.equal n q then [||] - else Array.sub args 0 n - | Zshift(k)::s -> - get_parameters (depth-k) n s - | [] -> (* we know that n < stack_args_size(argstk) (if well-typed term) *) - if Int.equal n 0 then [||] - else raise Not_found (* Trying to eta-expand a partial application..., should do - eta expansion first? *) - | _ -> assert false - (* strip_update_shift_app only produces Zapp and Zshift items *) - let eta_expand_ind_stack env ind m s (f, s') = let mib = lookup_mind (fst ind) env in match mib.mind_record with diff --git a/checker/declarations.ml b/checker/declarations.ml index 56d437c16..ad93146d5 100644 --- a/checker/declarations.ml +++ b/checker/declarations.ml @@ -463,13 +463,6 @@ let is_opaque cb = match cb.const_body with let opaque_univ_context cb = force_lazy_constr_univs cb.const_body -let subst_rel_declaration sub (id,copt,t as x) = - let copt' = Option.smartmap (subst_mps sub) copt in - let t' = subst_mps sub t in - if copt == copt' && t == t' then x else (id,copt',t') - -let subst_rel_context sub = List.smartmap (subst_rel_declaration sub) - let subst_recarg sub r = match r with | Norec -> r | (Mrec(kn,i)|Imbr (kn,i)) -> let kn' = subst_ind sub kn in @@ -517,10 +510,6 @@ let subst_decl_arity f g sub ar = if x' == x then ar else TemplateArity x' -let map_decl_arity f g = function - | RegularArity a -> RegularArity (f a) - | TemplateArity a -> TemplateArity (g a) - let subst_rel_declaration sub = Term.map_rel_decl (subst_mps sub) diff --git a/checker/inductive.ml b/checker/inductive.ml index ae2f7de8a..8f23a38af 100644 --- a/checker/inductive.ml +++ b/checker/inductive.ml @@ -436,13 +436,6 @@ let eq_recarg r1 r2 = match r1, r2 with let eq_wf_paths = Rtree.equal eq_recarg -let pp_recarg = function - | Norec -> Pp.str "Norec" - | Mrec i -> Pp.str ("Mrec "^MutInd.to_string (fst i)) - | Imbr i -> Pp.str ("Imbr "^MutInd.to_string (fst i)) - -let pp_wf_paths = Rtree.pp_tree pp_recarg - let inter_recarg r1 r2 = match r1, r2 with | Norec, Norec -> Some r1 | Mrec i1, Mrec i2 diff --git a/checker/typeops.ml b/checker/typeops.ml index 173e19ce1..02d801741 100644 --- a/checker/typeops.ml +++ b/checker/typeops.ml @@ -85,9 +85,6 @@ let type_of_constant_knowing_parameters env cst paramtyps = let type_of_constant_type env t = type_of_constant_type_knowing_parameters env t [||] -let type_of_constant env cst = - type_of_constant_knowing_parameters env cst [||] - let judge_of_constant_knowing_parameters env (kn,u as cst) paramstyp = let _cb = try lookup_constant kn env diff --git a/checker/univ.ml b/checker/univ.ml index 668f3a058..fb1a0faa7 100644 --- a/checker/univ.ml +++ b/checker/univ.ml @@ -87,7 +87,6 @@ module HList = struct val exists : (elt -> bool) -> t -> bool val for_all : (elt -> bool) -> t -> bool val for_all2 : (elt -> elt -> bool) -> t -> t -> bool - val remove : elt -> t -> t val to_list : t -> elt list end @@ -128,12 +127,6 @@ module HList = struct | Nil -> [] | Cons (x, _, l) -> x :: to_list l - let rec remove x = function - | Nil -> nil - | Cons (y, _, l) -> - if H.eq x y then l - else cons y (remove x l) - end end diff --git a/engine/termops.ml b/engine/termops.ml index 64f4c6dc5..29dcddbb0 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -612,30 +612,6 @@ let adjust_app_array_size f1 l1 f2 l2 = let extras,restl1 = Array.chop (len1-len2) l1 in (mkApp (f1,extras), restl1, f2, l2) -(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate - subterms of [c]; it carries an extra data [l] (typically a name - list) which is processed by [g na] (which typically cons [na] to - [l]) at each binder traversal (with name [na]); it is not recursive - and the order with which subterms are processed is not specified *) - -let map_constr_with_named_binders g f l c = match kind_of_term c with - | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ - | Construct _) -> c - | Cast (c,k,t) -> mkCast (f l c, k, f l t) - | Prod (na,t,c) -> mkProd (na, f l t, f (g na l) c) - | Lambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c) - | LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c) - | App (c,al) -> mkApp (f l c, Array.map (f l) al) - | Proj (p,c) -> mkProj (p, f l c) - | Evar (e,al) -> mkEvar (e, Array.map (f l) al) - | Case (ci,p,c,bl) -> mkCase (ci, f l p, f l c, Array.map (f l) bl) - | Fix (ln,(lna,tl,bl)) -> - let l' = Array.fold_left (fun l na -> g na l) l lna in - mkFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl)) - | CoFix(ln,(lna,tl,bl)) -> - let l' = Array.fold_left (fun l na -> g na l) l lna in - mkCoFix (ln,(lna,Array.map (f l) tl,Array.map (f l') bl)) - (* [map_constr_with_binders_left_to_right g f n c] maps [f n] on the immediate subterms of [c]; it carries an extra data [n] (typically a lift index) which is processed by [g] (which typically add 1 to diff --git a/ide/coqOps.ml b/ide/coqOps.ml index 259557f40..b180aa556 100644 --- a/ide/coqOps.ml +++ b/ide/coqOps.ml @@ -128,9 +128,6 @@ end = struct end open SentenceId -let log_pp msg : unit task = - Coq.lift (fun () -> Minilib.log_pp msg) - let log msg : unit task = Coq.lift (fun () -> Minilib.log msg) diff --git a/ide/ideutils.ml b/ide/ideutils.ml index d7d38a5ec..a08ab07b5 100644 --- a/ide/ideutils.ml +++ b/ide/ideutils.ml @@ -35,17 +35,6 @@ let flash_info = let flash_context = status#new_context ~name:"Flash" in (fun ?(delay=5000) s -> flash_context#flash ~delay s) -let xml_to_string xml = - let open Xml_datatype in - let buf = Buffer.create 1024 in - let rec iter = function - | PCData s -> Buffer.add_string buf s - | Element (_, _, children) -> - List.iter iter children - in - let () = iter xml in - Buffer.contents buf - let insert_with_tags (buf : #GText.buffer_skel) mark rmark tags text = (** FIXME: LablGTK2 does not export the C insert_with_tags function, so that it has to reimplement its own helper function. Unluckily, it relies on diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml index 8b4fadb5a..d08fb107b 100644 --- a/interp/notation_ops.ml +++ b/interp/notation_ops.ml @@ -1141,10 +1141,6 @@ let term_of_binder = function | Name id -> GVar (Loc.ghost,id) | Anonymous -> GHole (Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) -type glob_decl2 = - (name, cases_pattern) Util.union * Decl_kinds.binding_kind * - glob_constr option * glob_constr - let match_notation_constr u c (metas,pat) = let terms,binders,termlists,binderlists = match_ false u ([],[]) metas ([],[],[],[]) c pat in diff --git a/kernel/names.ml b/kernel/names.ml index 5c10badbe..811b4a62a 100644 --- a/kernel/names.ml +++ b/kernel/names.ml @@ -542,7 +542,6 @@ module KerPair = struct end module SyntacticOrd = struct - type t = kernel_pair let compare x y = match x, y with | Same knx, Same kny -> KerName.compare knx kny | Dual (knux,kncx), Dual (knuy,kncy) -> @@ -865,7 +864,6 @@ struct let hash (c, b) = (if b then 0 else 1) + Constant.hash c module SyntacticOrd = struct - type t = constant * bool let compare (c, b) (c', b') = if b = b' then Constant.SyntacticOrd.compare c c' else -1 let equal (c, b as x) (c', b' as x') = diff --git a/kernel/uGraph.mli b/kernel/uGraph.mli index e95cf4d1c..c8ac7df5c 100644 --- a/kernel/uGraph.mli +++ b/kernel/uGraph.mli @@ -61,3 +61,6 @@ val pr_universes : (Level.t -> Pp.std_ppcmds) -> universes -> Pp.std_ppcmds val dump_universes : (constraint_type -> string -> string -> unit) -> universes -> unit + +(** {6 Debugging} *) +val check_universes_invariants : universes -> unit diff --git a/kernel/univ.ml b/kernel/univ.ml index 09f884ecd..afe9cbe8d 100644 --- a/kernel/univ.ml +++ b/kernel/univ.ml @@ -440,10 +440,6 @@ struct let set = make Level.set let type1 = hcons (Level.set, 1) - let is_prop = function - | (l,0) -> Level.is_prop l - | _ -> false - let is_small = function | (l,0) -> Level.is_small l | _ -> false diff --git a/lib/feedback.ml b/lib/feedback.ml index df6fe3a62..0846e419b 100644 --- a/lib/feedback.ml +++ b/lib/feedback.ml @@ -40,8 +40,6 @@ type feedback = { contents : feedback_content; } -let default_route = 0 - (** Feeders *) let feeders : (int, feedback -> unit) Hashtbl.t = Hashtbl.create 7 diff --git a/lib/stateid.ml b/lib/stateid.ml index ae25735c5..c153f0e80 100644 --- a/lib/stateid.ml +++ b/lib/stateid.ml @@ -32,7 +32,6 @@ let compare = Int.compare module Self = struct type t = int let compare = compare - let equal = equal end module Set = Set.Make(Self) diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4 index f25b394f0..18942ac29 100644 --- a/parsing/cLexer.ml4 +++ b/parsing/cLexer.ml4 @@ -114,9 +114,6 @@ module Error = struct | UnsupportedUnicode x -> Printf.sprintf "Unsupported Unicode character (0x%x)" x) - (* Require to fix the Camlp4 signature *) - let print ppf x = Pp.pp_with ppf (Pp.str (to_string x)) - end open Error @@ -153,10 +150,6 @@ let bump_loc_line_last loc bol_pos = in Ploc.encl loc loc' -let set_loc_file loc fname = - Ploc.make_loc fname (Ploc.line_nb loc) (Ploc.bol_pos loc) - (Ploc.first_pos loc, Ploc.last_pos loc) (Ploc.comment loc) - (* For some reason, the [Ploc.after] function of Camlp5 does not update line numbers, so we define our own function that does it. *) let after loc = @@ -434,7 +427,6 @@ let push_char c = real_push_char c let push_string s = Buffer.add_string current_comment s -let push_bytes s = Buffer.add_bytes current_comment s let null_comment s = let rec null i = diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 496b20002..65e99d1b9 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -80,10 +80,6 @@ let create_pos = function | None -> Extend.First | Some lev -> Extend.After (constr_level lev) -type gram_level = - gram_position option * gram_assoc option * string option * - (** for reinitialization: *) gram_reinit option - let find_position_gen current ensure assoc lev = match lev with | None -> diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index 011565d86..3438635cf 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -756,11 +756,6 @@ GEXTEND Gram implicit_status = MaximallyImplicit}) items ] ]; - name_or_bang: [ - [ b = OPT "!"; id = name -> - not (Option.is_empty b), id - ] - ]; (* Same as [argument_spec_block], but with only implicit status and names *) more_implicits_block: [ [ name = name -> [(snd name, Vernacexpr.NotImplicit)] diff --git a/parsing/pcoq.ml b/parsing/pcoq.ml index dad98e2e9..3f014c4c8 100644 --- a/parsing/pcoq.ml +++ b/parsing/pcoq.ml @@ -171,7 +171,6 @@ module Symbols : sig val sopt : G.symbol -> G.symbol val snterml : G.internal_entry * string -> G.symbol val snterm : G.internal_entry -> G.symbol - val snterml_level : G.symbol -> string end = struct let stoken tok = @@ -199,10 +198,6 @@ end = struct let slist1 x = Gramext.Slist1 x let sopt x = Gramext.Sopt x - let snterml_level = function - | Gramext.Snterml (_, l) -> l - | _ -> failwith "snterml_level" - end let camlp4_verbosity silent f x = @@ -211,8 +206,6 @@ let camlp4_verbosity silent f x = f x; warning_verbose := a -let camlp4_verbose f x = camlp4_verbosity (not !Flags.quiet) f x - (** Grammar extensions *) (** NB: [extend_statment = diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml index 8dae17d69..df81bc78c 100644 --- a/plugins/funind/functional_principles_proofs.ml +++ b/plugins/funind/functional_principles_proofs.ml @@ -19,12 +19,6 @@ open Context.Rel.Declaration module RelDecl = Context.Rel.Declaration -let local_assum (na, t) = - RelDecl.LocalAssum (na, EConstr.Unsafe.to_constr t) - -let local_def (na, b, t) = - RelDecl.LocalDef (na, EConstr.Unsafe.to_constr b, EConstr.Unsafe.to_constr t) - (* let msgnl = Pp.msgnl *) (* diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 38fdfb759..35cfe8b54 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -52,8 +52,6 @@ let replace_in_clause_maybe_by ist c1 c2 cl tac = let replace_term ist dir_opt c cl = with_delayed_uconstr ist c (fun c -> replace_term dir_opt c cl) -let clause = Pltac.clause_dft_concl - TACTIC EXTEND replace ["replace" uconstr(c1) "with" constr(c2) clause(cl) by_arg_tac(tac) ] -> [ replace_in_clause_maybe_by ist c1 c2 cl tac ] diff --git a/plugins/ltac/profile_ltac.ml b/plugins/ltac/profile_ltac.ml index bcb28f77c..a853576f2 100644 --- a/plugins/ltac/profile_ltac.ml +++ b/plugins/ltac/profile_ltac.ml @@ -136,7 +136,6 @@ let feedback_results results = let format_sec x = (Printf.sprintf "%.3fs" x) let format_ratio x = (Printf.sprintf "%.1f%%" (100. *. x)) let padl n s = ws (max 0 (n - utf8_length s)) ++ str s -let padr n s = str s ++ ws (max 0 (n - utf8_length s)) let padr_with c n s = let ulength = utf8_length s in str (utf8_sub s 0 n) ++ str (String.make (max 0 (n - ulength)) c) diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index 8cda73b4b..ef1d69d35 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -88,9 +88,6 @@ let rec parse_user_entry s sep = else Uentry s -let arg_list = function Rawwit t -> Rawwit (ListArg t) -let arg_opt = function Rawwit t -> Rawwit (OptArg t) - let interp_entry_name interp symb = let rec eval = function | Ulist1 e -> Ulist1 (eval e) diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 3f83f104e..75227def0 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -118,12 +118,6 @@ let intern_constr_reference strict ist = function GRef (loc,locate_global_with_alias lqid,None), if strict then None else Some (CRef (r,None)) -let intern_move_location ist = function - | MoveAfter id -> MoveAfter (intern_hyp ist id) - | MoveBefore id -> MoveBefore (intern_hyp ist id) - | MoveFirst -> MoveFirst - | MoveLast -> MoveLast - (* Internalize an isolated reference in position of tactic *) let intern_isolated_global_tactic_reference r = diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index 50f43931e..fcdf7bb2c 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -436,12 +436,6 @@ let interp_hyp_list_as_list ist env sigma (loc,id as x) = let interp_hyp_list ist env sigma l = List.flatten (List.map (interp_hyp_list_as_list ist env sigma) l) -let interp_move_location ist env sigma = function - | MoveAfter id -> MoveAfter (interp_hyp ist env sigma id) - | MoveBefore id -> MoveBefore (interp_hyp ist env sigma id) - | MoveFirst -> MoveFirst - | MoveLast -> MoveLast - let interp_reference ist env sigma = function | ArgArg (_,r) -> r | ArgVar (loc, id) -> diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml index b1ff59e78..41f2edfcf 100644 --- a/plugins/nsatz/ideal.ml +++ b/plugins/nsatz/ideal.ml @@ -153,7 +153,6 @@ module Make (P:Polynom.S) = struct type coef = P.t let coef0 = P.of_num (Num.Int 0) let coef1 = P.of_num (Num.Int 1) - let coefm1 = P.of_num (Num.Int (-1)) let string_of_coef c = "["^(P.to_string c)^"]" (*********************************************************************** @@ -252,59 +251,6 @@ let string_of_pol zeroP hdP tlP coefterm monterm string_of_coef in (stringP p true) - - -let print_pol zeroP hdP tlP coefterm monterm string_of_coef - dimmon string_of_exp lvar p = - - let rec print_mon m coefone = - let s=ref [] in - for i=1 to (dimmon m) do - (match (string_of_exp m i) with - "0" -> () - | "1" -> s:= (!s) @ [(getvar lvar (i-1))] - | e -> s:= (!s) @ [((getvar lvar (i-1)) ^ "^" ^ e)]); - done; - (match !s with - [] -> if coefone - then print_string "1" - else () - | l -> if coefone - then print_string (String.concat "*" l) - else (print_string "*"; - print_string (String.concat "*" l))) - and print_term t start = let a = coefterm t and m = monterm t in - match (string_of_coef a) with - "0" -> () - | "1" ->(match start with - true -> print_mon m true - |false -> (print_string "+ "; - print_mon m true)) - | "-1" ->(print_string "-";print_space();print_mon m true) - | c -> if (String.get c 0)='-' - then (print_string "- "; - print_string (String.sub c 1 - ((String.length c)-1)); - print_mon m false) - else (match start with - true -> (print_string c;print_mon m false) - |false -> (print_string "+ "; - print_string c;print_mon m false)) - and printP p start = - if (zeroP p) - then (if start - then print_string("0") - else ()) - else (print_term (hdP p) start; - if start then open_hovbox 0; - print_space(); - print_cut(); - printP (tlP p) false) - in open_hovbox 3; - printP p true; - print_flush() - - let stringP metadata (p : poly) = string_of_pol (fun p -> match p with [] -> true | _ -> false) @@ -595,9 +541,6 @@ let addS x l = l @ [x] (* oblige de mettre en queue sinon le certificat decon critical pairs/s-polynomials *) -let ordcpair ((i1,j1),m1) ((i2,j2),m2) = - compare_mon m1 m2 - module CPair = struct type t = (int * int) * Monomial.t diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index a5b016611..632b9dac1 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -22,7 +22,6 @@ open Utile let num_0 = Int 0 and num_1 = Int 1 and num_2 = Int 2 -and num_10 = Int 10 let numdom r = let r' = Ratio.normalize_ratio (ratio_of_num r) in @@ -35,7 +34,6 @@ module BigInt = struct type t = big_int let of_int = big_int_of_int let coef0 = of_int 0 - let coef1 = of_int 1 let of_num = Num.big_int_of_num let to_num = Num.num_of_big_int let equal = eq_big_int @@ -49,7 +47,6 @@ module BigInt = struct let div = div_big_int let modulo = mod_big_int let to_string = string_of_big_int - let to_int x = int_of_big_int x let hash x = try (int_of_big_int x) with Failure _ -> 1 @@ -61,15 +58,6 @@ module BigInt = struct then a else if lt a b then pgcd b a else pgcd b (modulo a b) - - (* signe du pgcd = signe(a)*signe(b) si non nuls. *) - let pgcd2 a b = - if equal a coef0 then b - else if equal b coef0 then a - else let c = pgcd (abs a) (abs b) in - if ((lt coef0 a)&&(lt b coef0)) - ||((lt coef0 b)&&(lt a coef0)) - then opp c else c end (* @@ -146,8 +134,6 @@ let mul = function | (Const n,q) when eq_num n num_1 -> q | (p,q) -> Mul(p,q) -let unconstr = mkRel 1 - let tpexpr = lazy (gen_constant "CC" ["setoid_ring";"Ring_polynom"] "PExpr") let ttconst = lazy (gen_constant "CC" ["setoid_ring";"Ring_polynom"] "PEc") @@ -271,20 +257,6 @@ let set_nvars_term nvars t = | Pow (t1,n) -> aux t1 nvars in aux t nvars -let string_of_term p = - let rec aux p = - match p with - | Zero -> "0" - | Const r -> string_of_num r - | Var v -> "x"^v - | Opp t1 -> "(-"^(aux t1)^")" - | Add (t1,t2) -> "("^(aux t1)^"+"^(aux t2)^")" - | Sub (t1,t2) -> "("^(aux t1)^"-"^(aux t2)^")" - | Mul (t1,t2) -> "("^(aux t1)^"*"^(aux t2)^")" - | Pow (t1,n) -> (aux t1)^"^"^(string_of_int n) - in aux p - - (*********************************************************************** Coefficients: recursive polynomials *) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index dd68eac24..6e072e831 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -279,8 +279,6 @@ let my_constant c = let my_reference c = lazy (Coqlib.gen_reference_in_modules "Ring" plugin_modules c) -let new_ring_path = - DirPath.make (List.map Id.of_string ["Ring_tac";plugin_dir;"Coq"]) let znew_ring_path = DirPath.make (List.map Id.of_string ["InitialRing";plugin_dir;"Coq"]) let zltac s = diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index fc7963b2d..60c199bf5 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -90,8 +90,6 @@ let pp s = !pp_ref s let env_size env = List.length (Environ.named_context env) let safeDestApp c = match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] -let get_index = function ArgArg i -> i | _ -> - CErrors.anomaly (str"Uninterpreted index") (* Toplevel constr must be globalized twice ! *) let glob_constr ist genv = function | _, Some ce -> @@ -304,8 +302,6 @@ let unif_EQ_args env sigma pa a = let unif_HO env ise p c = Evarconv.the_conv_x env p c ise -let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise - let unif_HO_args env ise0 pa i ca = let n = Array.length pa in let rec loop ise j = @@ -371,11 +367,6 @@ let unif_end env sigma0 ise0 pt ok = let s, uc', t = nf_open_term sigma0 ise2 t in s, Evd.union_evar_universe_context uc uc', t -let pf_unif_HO gl sigma pt p c = - let env = pf_env gl in - let ise = unif_HO env (create_evar_defs sigma) p c in - unif_end env (project gl) ise pt (fun _ -> true) - let unify_HO env sigma0 t1 t2 = let sigma = unif_HO env sigma0 t1 t2 in let sigma, uc, _ = unif_end env sigma0 sigma t2 (fun _ -> true) in @@ -440,10 +431,6 @@ let all_ok _ _ = true let proj_nparams c = try 1 + Recordops.find_projection_nparams (ConstRef c) with _ -> 0 -let isFixed c = match kind_of_term c with - | Var _ | Ind _ | Construct _ | Const _ | Proj _ -> true - | _ -> false - let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false @@ -917,13 +904,6 @@ let pp_pattern (sigma, p) = let pr_cpattern = pr_term let pr_rpattern _ _ _ = pr_pattern -let pr_option f = function None -> mt() | Some x -> f x -let pr_ssrpattern _ _ _ = pr_option pr_pattern -let pr_pattern_squarep = pr_option (fun r -> str "[" ++ pr_pattern r ++ str "]") -let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep -let pr_pattern_roundp = pr_option (fun r -> str "(" ++ pr_pattern r ++ str ")") -let pr_ssrpattern_roundp _ _ _ = pr_pattern_roundp - let wit_rpatternty = add_genarg "rpatternty" pr_pattern let glob_ssrterm gs = function @@ -1045,7 +1025,6 @@ let interp_wit wit ist gl x = let arg = interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in sigma, Value.cast (topwit wit) arg -let interp_constr = interp_wit wit_constr let interp_open_constr ist gl gc = interp_wit wit_open_constr ist gl gc let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 894cdb943..1f984b160 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -226,7 +226,6 @@ val loc_of_cpattern : cpattern -> Loc.t val id_of_pattern : pattern -> Names.variable option val is_wildcard : cpattern -> bool val cpattern_of_id : Names.variable -> cpattern -val cpattern_of_id : Names.variable -> cpattern val pr_constr_pat : constr -> Pp.std_ppcmds val pf_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma val pf_unsafe_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 532cc8baa..4a3836d86 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -1535,9 +1535,6 @@ let indirectly_dependent sigma c d decls = way to see that the second hypothesis depends indirectly over 2 *) List.exists (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) decls -let indirect_dependency sigma d decls = - decls |> List.filter (fun d' -> dependent_in_decl sigma (EConstr.mkVar (NamedDecl.get_id d')) d) |> List.hd |> NamedDecl.get_id - let finish_evar_resolution ?(flags=Pretyping.all_and_fail_flags) env current_sigma (pending,c) = let current_sigma = Sigma.to_evar_map current_sigma in let sigma = Pretyping.solve_remaining_evars flags env current_sigma pending in diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml index 38eeda9b9..3041ac259 100644 --- a/printing/ppconstr.ml +++ b/printing/ppconstr.ml @@ -212,10 +212,6 @@ let tag_var = tag Tag.variable | Some (_,ExplByName id) -> str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")" - let pr_opt_type pr = function - | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt () - | t -> cut () ++ str ":" ++ pr t - let pr_opt_type_spc pr = function | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt () | t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index e4a87739b..58475630e 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -395,10 +395,6 @@ open Decl_kinds pr_opt (pr_guard_annot pr_lconstr_expr bl) guard ++ str":" ++ pr_spc_lconstr c) - let pr_priority = function - | None -> mt () - | Some i -> spc () ++ str "|" ++ spc () ++ int i - (**************************************) (* Pretty printer for vernac commands *) (**************************************) diff --git a/printing/printer.ml b/printing/printer.ml index 35ddf2e8c..93d221f03 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -370,11 +370,6 @@ let pr_context_limit_compact ?n env sigma = env ~init:(mt ()) in (sign_env ++ db_env) -(* compact printing an env (variables and de Bruijn). Separator: three - spaces between simple hyps, and newline otherwise *) -let pr_context_unlimited_compact env sigma = - pr_context_limit_compact env sigma - (* The number of printed hypothesis in a goal *) (* If [None], no limit *) let print_hyps_limit = ref (None : int option) diff --git a/proofs/clenv.ml b/proofs/clenv.ml index f9ebc4233..605914a01 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -27,11 +27,6 @@ open Unification open Misctypes open Sigma.Notations -(* Abbreviations *) - -let pf_env = Refiner.pf_env -let pf_type_of gls c = Typing.unsafe_type_of (pf_env gls) gls.sigma c - (******************************************************************) (* Clausal environments *) diff --git a/proofs/logic.ml b/proofs/logic.ml index e6024785d..9ab8c34d8 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -97,11 +97,6 @@ let check_typability env sigma c = (instead of iterating on the list of identifier to be removed, which forces the user to give them in order). *) -let clear_hyps env sigma ids sign cl = - let evdref = ref (Evd.clear_metas sigma) in - let (hyps,cl) = Evarutil.clear_hyps_in_evi env evdref sign cl ids in - (hyps, cl, !evdref) - let clear_hyps2 env sigma ids sign t cl = let evdref = ref (Evd.clear_metas sigma) in let (hyps,t,cl) = Evarutil.clear_hyps2_in_evi env evdref sign t cl ids in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index b55f8ef11..97c5cda77 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -179,9 +179,6 @@ module New = struct let pf_unsafe_type_of gl t = pf_apply unsafe_type_of gl t - let pf_get_type_of gl t = - pf_apply (Retyping.get_type_of ~lax:true) gl t - let pf_type_of gl t = pf_apply type_of gl t diff --git a/stm/stm.ml b/stm/stm.ml index 3738aa94a..f773d60c5 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1041,13 +1041,6 @@ end = struct (* {{{ *) | `Stop x -> x | `Cont acc -> next acc - let back_safe () = - let id = - fold_until (fun n (id,_,_,_,_) -> - if n >= 0 && State.is_cached_and_valid id then `Stop id else `Cont (succ n)) - 0 (VCS.get_branch_pos (VCS.current_branch ())) in - backto id - let undo_vernac_classifier v = try match v with diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml index 851554b83..15b40b42d 100644 --- a/tactics/hipattern.ml +++ b/tactics/hipattern.ml @@ -40,7 +40,6 @@ let mkmeta n = Nameops.make_ident "X" (Some n) let meta1 = mkmeta 1 let meta2 = mkmeta 2 let meta3 = mkmeta 3 -let meta4 = mkmeta 4 let op2bool = function Some _ -> true | None -> false diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 9c2a1d850..53f8e4d5f 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5039,8 +5039,6 @@ end module New = struct open Proofview.Notations - let exact_proof c = exact_proof c - open Genredexpr open Locus diff --git a/tools/coqdoc/output.mli b/tools/coqdoc/output.mli index 853bc29aa..235f2588c 100644 --- a/tools/coqdoc/output.mli +++ b/tools/coqdoc/output.mli @@ -64,7 +64,6 @@ val keyword : string -> loc -> unit val ident : string -> loc option -> unit val sublexer : char -> loc -> unit val sublexer_in_doc : char -> unit -val initialize : unit -> unit val proofbox : unit -> unit diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index 48cd70f69..9f6f77ef1 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -41,7 +41,6 @@ let vernac_echo loc in_chan = let open Loc in (* vernac parses the given stream, executes interpfun on the syntax tree it * parses, and is verbose on "primitives" commands if verbosely is true *) -let chan_beautify = ref stdout let beautify_suffix = ".beautified" let set_formatter_translator ch = diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index b9c4c6cc5..015552d68 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -28,8 +28,6 @@ open Proofview.Notations module RelDecl = Context.Rel.Declaration -let out_punivs = Univ.out_punivs - (**********************************************************************) (* Generic synthesis of boolean equality *) diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml index f86a0ef92..bb5be4cb0 100644 --- a/vernac/metasyntax.ml +++ b/vernac/metasyntax.ml @@ -527,7 +527,7 @@ let warn_skip_spaces_curly = (fun () ->strbrk "Skipping spaces inside curly brackets") let rec drop_spacing = function - | UnpCut _ as u :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt + | UnpCut _ :: fmt -> warn_skip_spaces_curly (); drop_spacing fmt | UnpTerminal s' :: fmt when String.equal s' (String.make (String.length s') ' ') -> warn_skip_spaces_curly (); drop_spacing fmt | fmt -> fmt diff --git a/vernac/record.ml b/vernac/record.ml index 8b4b7606f..53722b8f6 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -216,7 +216,7 @@ let warning_or_error coe indsp err = (pr_id fi ++ strbrk " cannot be defined because it is not typable.") in if coe then user_err ~hdr:"structure" st; - Flags.if_verbose Feedback.msg_info (hov 0 st) + warn_cannot_define_projection (hov 0 st) type field_status = | NoProjection of Name.t diff --git a/vernac/topfmt.ml b/vernac/topfmt.ml index a1835959c..e44b585d7 100644 --- a/vernac/topfmt.ml +++ b/vernac/topfmt.ml @@ -106,8 +106,6 @@ module Tag = struct end -type logger = ?loc:Loc.t -> level -> std_ppcmds -> unit - let msgnl_with ?pre_hdr fmt strm = pp_with fmt (strm ++ fnl ()); Format.pp_print_flush fmt () @@ -283,16 +281,6 @@ let print_err_exn ?extra any = let msg = CErrors.iprint (e, info) ++ fnl () in std_logger ~pre_hdr Feedback.Error msg -(* Output to file, used only in extraction so a candidate for removal *) -let ft_logger old_logger ft ?loc level mesg = - let id x = x in - match level with - | Debug -> msgnl_with ft (make_body id dbg_hdr mesg) - | Info -> msgnl_with ft (make_body id info_hdr mesg) - | Notice -> msgnl_with ft mesg - | Warning -> old_logger ?loc level mesg - | Error -> old_logger ?loc level mesg - let with_output_to_file fname func input = (* XXX FIXME: redirect std_ft *) (* let old_logger = !logger in *) -- cgit v1.2.3 From 9be835c4f16b3bc08ff9540a6854ced2d8185cb2 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:09:05 +0200 Subject: Remove unused constructors --- parsing/cLexer.ml4 | 5 +---- plugins/ssrmatching/ssrmatching.ml4 | 2 -- stm/stm.ml | 1 - 3 files changed, 1 insertion(+), 7 deletions(-) (limited to 'plugins/ssrmatching') diff --git a/parsing/cLexer.ml4 b/parsing/cLexer.ml4 index 18942ac29..462905808 100644 --- a/parsing/cLexer.ml4 +++ b/parsing/cLexer.ml4 @@ -99,7 +99,6 @@ module Error = struct | Unterminated_string | Undefined_token | Bad_token of string - | UnsupportedUnicode of int exception E of t @@ -110,9 +109,7 @@ module Error = struct | Unterminated_comment -> "Unterminated comment" | Unterminated_string -> "Unterminated string" | Undefined_token -> "Undefined token" - | Bad_token tok -> Format.sprintf "Bad token %S" tok - | UnsupportedUnicode x -> - Printf.sprintf "Unsupported Unicode character (0x%x)" x) + | Bad_token tok -> Format.sprintf "Bad token %S" tok) end open Error diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index 60c199bf5..f146fbb11 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -435,8 +435,6 @@ let isRigid c = match kind_of_term c with | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true | _ -> false -exception UndefPat - let hole_var = mkVar (id_of_string "_") let pr_constr_pat c0 = let rec wipe_evar c = diff --git a/stm/stm.ml b/stm/stm.ml index f773d60c5..8c1185b6d 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1274,7 +1274,6 @@ end = struct (* {{{ *) | RespBuiltProof of Proof_global.closed_proof_output * float | RespError of error | RespStates of (Stateid.t * State.partial_state) list - | RespDone let name = ref "proofworker" let extra_env () = !async_proofs_workers_extra_env -- cgit v1.2.3 From 4e84e83911c1cf7613a35b921b1e68e097f84b5a Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:11:47 +0200 Subject: Remove unused [open] statements --- dev/top_printers.ml | 1 - engine/eConstr.ml | 1 - engine/eConstr.mli | 1 - engine/evarutil.ml | 1 - engine/evd.ml | 4 ---- engine/proofview.mli | 1 - engine/termops.ml | 1 - engine/universes.ml | 1 - interp/constrintern.ml | 2 -- interp/constrintern.mli | 1 - interp/stdarg.ml | 2 -- interp/stdarg.mli | 1 - interp/syntax_def.ml | 2 -- intf/tactypes.mli | 1 - library/goptions.ml | 1 - library/libobject.ml | 1 - parsing/egramcoq.ml | 1 - parsing/egramcoq.mli | 8 -------- parsing/g_prim.ml4 | 1 - parsing/g_proofs.ml4 | 1 - plugins/firstorder/instances.ml | 1 - plugins/funind/functional_principles_proofs.mli | 1 - plugins/funind/g_indfun.ml4 | 1 - plugins/funind/indfun_common.ml | 1 - plugins/funind/invfun.ml | 1 - plugins/funind/merge.ml | 1 - plugins/ltac/evar_tactics.ml | 1 - plugins/ltac/extratactics.ml4 | 1 - plugins/ltac/g_auto.ml4 | 1 - plugins/ltac/g_class.ml4 | 3 --- plugins/ltac/pltac.ml | 1 - plugins/ltac/rewrite.mli | 1 - plugins/ltac/taccoerce.mli | 1 - plugins/ltac/tacentries.ml | 1 - plugins/ltac/tacenv.mli | 1 - plugins/ltac/tacinterp.ml | 1 - plugins/ltac/tacinterp.mli | 1 - plugins/ltac/tactic_debug.mli | 1 - plugins/micromega/coq_micromega.ml | 1 - plugins/nsatz/ideal.ml | 2 -- plugins/setoid_ring/newring.ml | 1 - plugins/setoid_ring/newring.mli | 3 --- plugins/ssrmatching/ssrmatching.ml4 | 14 -------------- plugins/ssrmatching/ssrmatching.mli | 1 - pretyping/cbv.mli | 1 - pretyping/classops.ml | 1 - pretyping/classops.mli | 1 - pretyping/coercion.ml | 1 - pretyping/coercion.mli | 1 - pretyping/detyping.ml | 1 - pretyping/evarconv.ml | 1 - pretyping/evarconv.mli | 1 - pretyping/evardefine.ml | 1 - pretyping/find_subterm.mli | 1 - pretyping/inductiveops.ml | 1 - pretyping/patternops.ml | 1 - pretyping/patternops.mli | 1 - pretyping/pretype_errors.ml | 1 - pretyping/pretyping.ml | 1 - pretyping/program.ml | 1 - pretyping/tacred.mli | 1 - pretyping/typeclasses_errors.ml | 1 - pretyping/typeclasses_errors.mli | 1 - printing/prettyp.mli | 1 - proofs/clenvtac.mli | 1 - proofs/goal.ml | 1 - proofs/proof_type.mli | 3 --- proofs/refiner.ml | 1 - proofs/tacmach.mli | 1 - stm/stm.mli | 3 --- tactics/auto.ml | 1 - tactics/auto.mli | 1 - tactics/btermdn.mli | 1 - tactics/class_tactics.ml | 3 --- tactics/class_tactics.mli | 1 - tactics/contradiction.mli | 1 - tactics/eauto.mli | 2 -- tactics/elim.ml | 1 - tactics/elim.mli | 1 - tactics/eqdecide.ml | 1 - tactics/equality.ml | 1 - tactics/equality.mli | 1 - tactics/hints.mli | 1 - tactics/hipattern.mli | 1 - tactics/inv.ml | 1 - tactics/inv.mli | 1 - tactics/leminv.mli | 1 - tactics/tactics.ml | 2 -- toplevel/coqloop.mli | 2 -- vernac/auto_ind_decl.ml | 2 -- vernac/classes.ml | 2 -- vernac/obligations.mli | 1 - vernac/search.ml | 2 -- vernac/search.mli | 1 - 94 files changed, 137 deletions(-) (limited to 'plugins/ssrmatching') diff --git a/dev/top_printers.ml b/dev/top_printers.ml index 918e98f77..7caaf2d9d 100644 --- a/dev/top_printers.ml +++ b/dev/top_printers.ml @@ -503,7 +503,6 @@ VERNAC COMMAND EXTEND PrintConstr END *) -open Pcoq open Genarg open Stdarg open Egramml diff --git a/engine/eConstr.ml b/engine/eConstr.ml index bb9075e74..54d3ce6cf 100644 --- a/engine/eConstr.ml +++ b/engine/eConstr.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open CSig open CErrors open Util open Names diff --git a/engine/eConstr.mli b/engine/eConstr.mli index 3a9469e55..693b592fd 100644 --- a/engine/eConstr.mli +++ b/engine/eConstr.mli @@ -9,7 +9,6 @@ open CSig open Names open Constr -open Context open Environ type t diff --git a/engine/evarutil.ml b/engine/evarutil.ml index 1624dc93e..fba466107 100644 --- a/engine/evarutil.ml +++ b/engine/evarutil.ml @@ -10,7 +10,6 @@ open CErrors open Util open Names open Term -open Vars open Termops open Namegen open Pre_env diff --git a/engine/evd.ml b/engine/evd.ml index 5419a10a8..6b1e1a855 100644 --- a/engine/evd.ml +++ b/engine/evd.ml @@ -14,8 +14,6 @@ open Nameops open Term open Vars open Environ -open Globnames -open Context.Named.Declaration module RelDecl = Context.Rel.Declaration module NamedDecl = Context.Named.Declaration @@ -360,8 +358,6 @@ module EvMap = Evar.Map module EvNames : sig -open Misctypes - type t val empty : t diff --git a/engine/proofview.mli b/engine/proofview.mli index a3b0304b1..da8a8fecd 100644 --- a/engine/proofview.mli +++ b/engine/proofview.mli @@ -13,7 +13,6 @@ state and returning a value of type ['a]. *) open Util -open Term open EConstr (** Main state of tactics *) diff --git a/engine/termops.ml b/engine/termops.ml index 29dcddbb0..19e62f8e6 100644 --- a/engine/termops.ml +++ b/engine/termops.ml @@ -1427,7 +1427,6 @@ let dependency_closure env sigma sign hyps = List.rev lh let global_app_of_constr sigma c = - let open Univ in let open Globnames in match EConstr.kind sigma c with | Const (c, u) -> (ConstRef c, u), None diff --git a/engine/universes.ml b/engine/universes.ml index ab561784c..1900112dd 100644 --- a/engine/universes.ml +++ b/engine/universes.ml @@ -13,7 +13,6 @@ open Term open Environ open Univ open Globnames -open Decl_kinds let pr_with_global_universes l = try Nameops.pr_id (LMap.find l (snd (Global.global_universe_names ()))) diff --git a/interp/constrintern.ml b/interp/constrintern.ml index d75487ecf..389880c5c 100644 --- a/interp/constrintern.ml +++ b/interp/constrintern.ml @@ -2045,8 +2045,6 @@ let interp_binder_evars env evdref na t = let t' = locate_if_hole (loc_of_glob_constr t) na t in understand_tcc_evars env evdref ~expected_type:IsType t' -open Environ - let my_intern_constr env lvar acc c = internalize env acc false lvar c diff --git a/interp/constrintern.mli b/interp/constrintern.mli index 758d4e650..fdd50c6a1 100644 --- a/interp/constrintern.mli +++ b/interp/constrintern.mli @@ -18,7 +18,6 @@ open Constrexpr open Notation_term open Pretyping open Misctypes -open Decl_kinds (** Translation from front abstract syntax of term to untyped terms (glob_constr) *) diff --git a/interp/stdarg.ml b/interp/stdarg.ml index 341ff5662..5920b0d50 100644 --- a/interp/stdarg.ml +++ b/interp/stdarg.ml @@ -6,9 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Loc open Misctypes -open Tactypes open Genarg open Geninterp diff --git a/interp/stdarg.mli b/interp/stdarg.mli index 113fe40ba..ac40a2328 100644 --- a/interp/stdarg.mli +++ b/interp/stdarg.mli @@ -10,7 +10,6 @@ open Loc open Names -open Term open EConstr open Libnames open Globnames diff --git a/interp/syntax_def.ml b/interp/syntax_def.ml index c3f4c4f30..ed7b0b70d 100644 --- a/interp/syntax_def.ml +++ b/interp/syntax_def.ml @@ -106,5 +106,3 @@ let search_syntactic_definition kn = let def = out_pat pat in verbose_compat kn def v; def - -open Goptions diff --git a/intf/tactypes.mli b/intf/tactypes.mli index 02cfc44e2..ef90b911c 100644 --- a/intf/tactypes.mli +++ b/intf/tactypes.mli @@ -13,7 +13,6 @@ open Loc open Names open Constrexpr -open Glob_term open Pattern open Misctypes diff --git a/library/goptions.ml b/library/goptions.ml index 1c08b9539..c111113ca 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -235,7 +235,6 @@ with Not_found -> then error "Sorry, this option name is already used." open Libobject -open Lib let warn_deprecated_option = CWarnings.create ~name:"deprecated-option" ~category:"deprecated" diff --git a/library/libobject.ml b/library/libobject.ml index 8757ca08c..897591762 100644 --- a/library/libobject.ml +++ b/library/libobject.ml @@ -8,7 +8,6 @@ open Libnames open Pp -open Util module Dyn = Dyn.Make(struct end) diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml index 65e99d1b9..86c66ec5f 100644 --- a/parsing/egramcoq.ml +++ b/parsing/egramcoq.ml @@ -10,7 +10,6 @@ open CErrors open Util open Pcoq open Constrexpr -open Notation open Notation_term open Extend open Libnames diff --git a/parsing/egramcoq.mli b/parsing/egramcoq.mli index 6dda3817a..0a0430ba6 100644 --- a/parsing/egramcoq.mli +++ b/parsing/egramcoq.mli @@ -6,14 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Names -open Constrexpr -open Notation_term -open Pcoq -open Extend -open Genarg -open Egramml - (** Mapping of grammar productions to camlp4 actions *) (** This is the part specific to Coq-level Notation and Tactic Notation. diff --git a/parsing/g_prim.ml4 b/parsing/g_prim.ml4 index 2af4ed533..abb463f82 100644 --- a/parsing/g_prim.ml4 +++ b/parsing/g_prim.ml4 @@ -8,7 +8,6 @@ open Names open Libnames -open Tok (* necessary for camlp4 *) open Pcoq open Pcoq.Prim diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4 index 7ca2e4a4f..68b8be6b8 100644 --- a/parsing/g_proofs.ml4 +++ b/parsing/g_proofs.ml4 @@ -9,7 +9,6 @@ open Constrexpr open Vernacexpr open Misctypes -open Tok open Pcoq open Pcoq.Prim diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index 62f811546..2b624b983 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -17,7 +17,6 @@ open Tacmach.New open Tactics open Tacticals.New open Proofview.Notations -open Termops open Reductionops open Formula open Sequent diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli index 7ddc84d01..61752aa33 100644 --- a/plugins/funind/functional_principles_proofs.mli +++ b/plugins/funind/functional_principles_proofs.mli @@ -1,5 +1,4 @@ open Names -open Term val prove_princ_for_struct : Evd.evar_map ref -> diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 index 0dccd25d7..b5eacee81 100644 --- a/plugins/funind/g_indfun.ml4 +++ b/plugins/funind/g_indfun.ml4 @@ -8,7 +8,6 @@ (*i camlp4deps: "grammar/grammar.cma" i*) open Ltac_plugin open Util -open Term open Pp open Constrexpr open Indfun_common diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml index 7b0d7d27d..c6f5c2745 100644 --- a/plugins/funind/indfun_common.ml +++ b/plugins/funind/indfun_common.ml @@ -508,7 +508,6 @@ let list_rewrite (rev:bool) (eqs: (EConstr.constr*bool) list) = (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));; let decompose_lam_n sigma n = - let open EConstr in if n < 0 then CErrors.error "decompose_lam_n: integer parameter must be positive"; let rec lamdec_rec l n c = if Int.equal n 0 then l,c diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index 94ec0a898..29472cdef 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -7,7 +7,6 @@ (************************************************************************) open Ltac_plugin -open Tacexpr open Declarations open CErrors open Util diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml index f1ca57585..0af0898a0 100644 --- a/plugins/funind/merge.ml +++ b/plugins/funind/merge.ml @@ -19,7 +19,6 @@ open Pp open Names open Term open Vars -open Termops open Declarations open Glob_term open Glob_termops diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5d3f6df03..bc9c300e2 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -9,7 +9,6 @@ open Util open Names open Term -open EConstr open CErrors open Evar_refiner open Tacmach diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4 index 35cfe8b54..21419d1f9 100644 --- a/plugins/ltac/extratactics.ml4 +++ b/plugins/ltac/extratactics.ml4 @@ -21,7 +21,6 @@ open Tacexpr open Glob_ops open CErrors open Util -open Evd open Termops open Equality open Misctypes diff --git a/plugins/ltac/g_auto.ml4 b/plugins/ltac/g_auto.ml4 index dfa8331ff..50e8255a6 100644 --- a/plugins/ltac/g_auto.ml4 +++ b/plugins/ltac/g_auto.ml4 @@ -16,7 +16,6 @@ open Pcoq.Constr open Pltac open Hints open Tacexpr -open Proofview.Notations open Names DECLARE PLUGIN "g_auto" diff --git a/plugins/ltac/g_class.ml4 b/plugins/ltac/g_class.ml4 index ff5e7d5ff..23ce368ee 100644 --- a/plugins/ltac/g_class.ml4 +++ b/plugins/ltac/g_class.ml4 @@ -8,9 +8,7 @@ (*i camlp4deps: "grammar/grammar.cma" i*) -open Misctypes open Class_tactics -open Pltac open Stdarg open Tacarg open Names @@ -95,7 +93,6 @@ END (** TODO: DEPRECATE *) (* A progress test that allows to see if the evars have changed *) open Term -open Proofview.Goal open Proofview.Notations let rec eq_constr_mod_evars sigma x y = diff --git a/plugins/ltac/pltac.ml b/plugins/ltac/pltac.ml index 1d21118ae..7e979d269 100644 --- a/plugins/ltac/pltac.ml +++ b/plugins/ltac/pltac.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Names open Pcoq (* Main entry for extensions *) diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli index 7a20838a2..6683d753b 100644 --- a/plugins/ltac/rewrite.mli +++ b/plugins/ltac/rewrite.mli @@ -14,7 +14,6 @@ open Constrexpr open Tacexpr open Misctypes open Evd -open Proof_type open Tacinterp (** TODO: document and clean me! *) diff --git a/plugins/ltac/taccoerce.mli b/plugins/ltac/taccoerce.mli index 9c4ac5265..4a44f86d9 100644 --- a/plugins/ltac/taccoerce.mli +++ b/plugins/ltac/taccoerce.mli @@ -8,7 +8,6 @@ open Util open Names -open Term open EConstr open Misctypes open Pattern diff --git a/plugins/ltac/tacentries.ml b/plugins/ltac/tacentries.ml index ef1d69d35..32750383b 100644 --- a/plugins/ltac/tacentries.ml +++ b/plugins/ltac/tacentries.ml @@ -15,7 +15,6 @@ open Genarg open Extend open Pcoq open Egramml -open Egramcoq open Vernacexpr open Libnames open Nameops diff --git a/plugins/ltac/tacenv.mli b/plugins/ltac/tacenv.mli index 94e14223a..d1e2a7bbe 100644 --- a/plugins/ltac/tacenv.mli +++ b/plugins/ltac/tacenv.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Genarg open Names open Tacexpr open Geninterp diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index fcdf7bb2c..b8c021f18 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -25,7 +25,6 @@ open Refiner open Tacmach.New open Tactic_debug open Constrexpr -open Term open Termops open Tacexpr open Genarg diff --git a/plugins/ltac/tacinterp.mli b/plugins/ltac/tacinterp.mli index 1e5f6bd42..494f36a95 100644 --- a/plugins/ltac/tacinterp.mli +++ b/plugins/ltac/tacinterp.mli @@ -8,7 +8,6 @@ open Names open Tactic_debug -open Term open EConstr open Tacexpr open Genarg diff --git a/plugins/ltac/tactic_debug.mli b/plugins/ltac/tactic_debug.mli index 7745d9b7b..0b4d35a22 100644 --- a/plugins/ltac/tactic_debug.mli +++ b/plugins/ltac/tactic_debug.mli @@ -10,7 +10,6 @@ open Environ open Pattern open Names open Tacexpr -open Term open EConstr open Evd diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index eb26c5431..a36607ec3 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -331,7 +331,6 @@ module M = struct open Coqlib - open Term open Constr open EConstr diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml index 41f2edfcf..a120d4efb 100644 --- a/plugins/nsatz/ideal.ml +++ b/plugins/nsatz/ideal.ml @@ -196,8 +196,6 @@ module Hashpol = Hashtbl.Make( (* A pretty printer for polynomials, with Maple-like syntax. *) -open Format - let getvar lv i = try (List.nth lv i) with Failure _ -> (List.fold_left (fun r x -> r^" "^x) "lv= " lv) diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index 6e072e831..d59ffee54 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -19,7 +19,6 @@ open Environ open Libnames open Globnames open Glob_term -open Tacticals open Tacexpr open Coqlib open Mod_subst diff --git a/plugins/setoid_ring/newring.mli b/plugins/setoid_ring/newring.mli index 4367d021c..d9d32c681 100644 --- a/plugins/setoid_ring/newring.mli +++ b/plugins/setoid_ring/newring.mli @@ -7,13 +7,10 @@ (************************************************************************) open Names -open Constr open EConstr open Libnames open Globnames open Constrexpr -open Tacexpr -open Proof_type open Newring_ast val protect_tac_in : string -> Id.t -> unit Proofview.tactic diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4 index f146fbb11..72c70750c 100644 --- a/plugins/ssrmatching/ssrmatching.ml4 +++ b/plugins/ssrmatching/ssrmatching.ml4 @@ -21,30 +21,21 @@ open Pp open Pcoq open Genarg open Stdarg -open Tacarg open Term open Vars -open Topconstr open Libnames open Tactics open Tacticals open Termops -open Namegen open Recordops open Tacmach -open Coqlib open Glob_term open Util open Evd -open Extend -open Goptions open Tacexpr -open Proofview.Notations open Tacinterp open Pretyping open Constr -open Pltac -open Extraargs open Ppconstr open Printer @@ -54,14 +45,9 @@ open Decl_kinds open Evar_kinds open Constrexpr open Constrexpr_ops -open Notation_term -open Notation_ops -open Locus -open Locusops DECLARE PLUGIN "ssrmatching_plugin" -type loc = Loc.t let dummy_loc = Loc.ghost let errorstrm = CErrors.user_err ~hdr:"ssrmatching" let loc_error loc msg = CErrors.user_err ~loc ~hdr:msg (str msg) diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index 1f984b160..638b4e254 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -4,7 +4,6 @@ open Genarg open Tacexpr open Environ -open Tacmach open Evd open Proof_type open Term diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli index b014af2c7..eb25994be 100644 --- a/pretyping/cbv.mli +++ b/pretyping/cbv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Environ open CClosure diff --git a/pretyping/classops.ml b/pretyping/classops.ml index e9b3d197b..32da81f96 100644 --- a/pretyping/classops.ml +++ b/pretyping/classops.ml @@ -17,7 +17,6 @@ open Nametab open Environ open Libobject open Term -open Termops open Mod_subst (* usage qque peu general: utilise aussi dans record *) diff --git a/pretyping/classops.mli b/pretyping/classops.mli index 0d741a5a5..c4238e8b0 100644 --- a/pretyping/classops.mli +++ b/pretyping/classops.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Environ open EConstr open Evd diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 542db7fdf..c26e7458e 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -22,7 +22,6 @@ open Environ open EConstr open Vars open Reductionops -open Typeops open Pretype_errors open Classops open Evarutil diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli index bc63d092d..ea3d3f0fa 100644 --- a/pretyping/coercion.mli +++ b/pretyping/coercion.mli @@ -8,7 +8,6 @@ open Evd open Names -open Term open Environ open EConstr open Glob_term diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 8ba408679..483e2b432 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -13,7 +13,6 @@ open CErrors open Util open Names open Term -open Environ open EConstr open Vars open Inductiveops diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index 4bb66b8e9..305eae15a 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -21,7 +21,6 @@ open Recordops open Evarutil open Evardefine open Evarsolve -open Globnames open Evd open Pretype_errors open Sigma.Notations diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli index fc07f0fbe..7cee1e8a7 100644 --- a/pretyping/evarconv.mli +++ b/pretyping/evarconv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Environ open Reductionops diff --git a/pretyping/evardefine.ml b/pretyping/evardefine.ml index c5ae684e3..5fd104c78 100644 --- a/pretyping/evardefine.ml +++ b/pretyping/evardefine.ml @@ -11,7 +11,6 @@ open Pp open Names open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/pretyping/find_subterm.mli b/pretyping/find_subterm.mli index e3d3b74f1..d22f94e4e 100644 --- a/pretyping/find_subterm.mli +++ b/pretyping/find_subterm.mli @@ -7,7 +7,6 @@ (************************************************************************) open Locus -open Term open Evd open Pretype_errors open Environ diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml index 5b42add28..429e5005e 100644 --- a/pretyping/inductiveops.ml +++ b/pretyping/inductiveops.ml @@ -459,7 +459,6 @@ let extract_mrectype sigma t = | _ -> raise Not_found let find_mrectype_vect env sigma c = - let open EConstr in let (t, l) = Termops.decompose_app_vect sigma (whd_all env sigma c) in match EConstr.kind sigma t with | Ind ind -> (ind, l) diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml index b16d04495..33a68589c 100644 --- a/pretyping/patternops.ml +++ b/pretyping/patternops.ml @@ -20,7 +20,6 @@ open Mod_subst open Misctypes open Decl_kinds open Pattern -open Evd open Environ let case_info_pattern_eq i1 i2 = diff --git a/pretyping/patternops.mli b/pretyping/patternops.mli index 5694d345c..791fd74ed 100644 --- a/pretyping/patternops.mli +++ b/pretyping/patternops.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr open Globnames open Glob_term diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 24f6d1689..f9cf6b83b 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Util open Names open Term open Environ diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 497bde340..68ef97659 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -33,7 +33,6 @@ open EConstr open Vars open Reductionops open Type_errors -open Typeops open Typing open Globnames open Nameops diff --git a/pretyping/program.ml b/pretyping/program.ml index caa5a5c8a..42acc5705 100644 --- a/pretyping/program.ml +++ b/pretyping/program.ml @@ -10,7 +10,6 @@ open Pp open CErrors open Util open Names -open Term let make_dir l = DirPath.make (List.rev_map Id.of_string l) diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli index 76d0bc241..c31212e26 100644 --- a/pretyping/tacred.mli +++ b/pretyping/tacred.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Environ open Evd open EConstr diff --git a/pretyping/typeclasses_errors.ml b/pretyping/typeclasses_errors.ml index 2db0e9e88..754dacd19 100644 --- a/pretyping/typeclasses_errors.ml +++ b/pretyping/typeclasses_errors.ml @@ -8,7 +8,6 @@ (*i*) open Names -open Term open EConstr open Environ open Constrexpr diff --git a/pretyping/typeclasses_errors.mli b/pretyping/typeclasses_errors.mli index 9bd430e4d..558575ccc 100644 --- a/pretyping/typeclasses_errors.mli +++ b/pretyping/typeclasses_errors.mli @@ -8,7 +8,6 @@ open Loc open Names -open Term open EConstr open Environ open Constrexpr diff --git a/printing/prettyp.mli b/printing/prettyp.mli index 38e111034..6841781cc 100644 --- a/printing/prettyp.mli +++ b/printing/prettyp.mli @@ -8,7 +8,6 @@ open Pp open Names -open Term open Environ open Reductionops open Libnames diff --git a/proofs/clenvtac.mli b/proofs/clenvtac.mli index 5b7164705..26069207e 100644 --- a/proofs/clenvtac.mli +++ b/proofs/clenvtac.mli @@ -8,7 +8,6 @@ (** Legacy components of the previous proof engine. *) -open Term open Clenv open EConstr open Unification diff --git a/proofs/goal.ml b/proofs/goal.ml index 9046f4534..fc8e635a0 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -8,7 +8,6 @@ open Util open Pp -open Term open Sigma.Notations module NamedDecl = Context.Named.Declaration diff --git a/proofs/proof_type.mli b/proofs/proof_type.mli index 03bc5e471..e59db9e42 100644 --- a/proofs/proof_type.mli +++ b/proofs/proof_type.mli @@ -11,9 +11,6 @@ open Evd open Names open Term -open Glob_term -open Nametab -open Misctypes (** This module defines the structure of proof tree and the tactic type. So, it is used by [Proof_tree] and [Refiner] *) diff --git a/proofs/refiner.ml b/proofs/refiner.ml index bc12b3ba7..259e96a27 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -10,7 +10,6 @@ open Pp open CErrors open Util open Evd -open Environ open Proof_type open Logic diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 627a8e0e7..e6e60e27f 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -15,7 +15,6 @@ open Proof_type open Redexpr open Pattern open Locus -open Misctypes (** Operations for handling terms under a local typing context. *) diff --git a/stm/stm.mli b/stm/stm.mli index 30e9629c6..d2bee4496 100644 --- a/stm/stm.mli +++ b/stm/stm.mli @@ -6,10 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Vernacexpr open Names -open Feedback -open Loc (** state-transaction-machine interface *) diff --git a/tactics/auto.ml b/tactics/auto.ml index 42230dff1..324c551d0 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -12,7 +12,6 @@ open Pp open Util open CErrors open Names -open Vars open Termops open EConstr open Environ diff --git a/tactics/auto.mli b/tactics/auto.mli index 32710e347..9ed9f0ae2 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -9,7 +9,6 @@ (** This files implements auto and related automation tactics *) open Names -open Term open EConstr open Clenv open Pattern diff --git a/tactics/btermdn.mli b/tactics/btermdn.mli index 2a5e7c345..27f624f71 100644 --- a/tactics/btermdn.mli +++ b/tactics/btermdn.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open Pattern open Names diff --git a/tactics/class_tactics.ml b/tactics/class_tactics.ml index c430e776a..6c724a1eb 100644 --- a/tactics/class_tactics.ml +++ b/tactics/class_tactics.ml @@ -18,7 +18,6 @@ open Names open Term open Termops open EConstr -open Reduction open Proof_type open Tacticals open Tacmach @@ -1219,7 +1218,6 @@ module Search = struct let intro_tac info kont gl = let open Proofview in - let open Proofview.Notations in let env = Goal.env gl in let sigma = Goal.sigma gl in let s = Sigma.to_evar_map sigma in @@ -1257,7 +1255,6 @@ module Search = struct let search_tac_gl ?st only_classes dep hints depth i sigma gls gl : unit Proofview.tactic = let open Proofview in - let open Proofview.Notations in if false (* In 8.6, still allow non-class goals only_classes && not (is_class_type sigma (Goal.concl gl)) *) then Tacticals.New.tclZEROMSG (str"Not a subgoal for a class") else diff --git a/tactics/class_tactics.mli b/tactics/class_tactics.mli index 738cc0feb..4ab29f260 100644 --- a/tactics/class_tactics.mli +++ b/tactics/class_tactics.mli @@ -9,7 +9,6 @@ (** This files implements typeclasses eauto *) open Names -open Constr open EConstr open Tacmach diff --git a/tactics/contradiction.mli b/tactics/contradiction.mli index 510b135b0..2cf5a6829 100644 --- a/tactics/contradiction.mli +++ b/tactics/contradiction.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr open Misctypes diff --git a/tactics/eauto.mli b/tactics/eauto.mli index e2006ac1e..c952f4e72 100644 --- a/tactics/eauto.mli +++ b/tactics/eauto.mli @@ -6,9 +6,7 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Term open EConstr -open Proof_type open Hints open Tactypes diff --git a/tactics/elim.ml b/tactics/elim.ml index e37ec6bce..855cb206f 100644 --- a/tactics/elim.ml +++ b/tactics/elim.ml @@ -8,7 +8,6 @@ open Util open Names -open Term open Termops open EConstr open Inductiveops diff --git a/tactics/elim.mli b/tactics/elim.mli index dc1af79ba..fb7cc7b83 100644 --- a/tactics/elim.mli +++ b/tactics/elim.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Tacticals open Misctypes diff --git a/tactics/eqdecide.ml b/tactics/eqdecide.ml index bac3980d2..5012b0ef7 100644 --- a/tactics/eqdecide.ml +++ b/tactics/eqdecide.ml @@ -25,7 +25,6 @@ open Constr_matching open Misctypes open Tactypes open Hipattern -open Pretyping open Proofview.Notations open Tacmach.New open Coqlib diff --git a/tactics/equality.ml b/tactics/equality.ml index 25c28cf4a..cc7701ad5 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -14,7 +14,6 @@ open Names open Nameops open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/tactics/equality.mli b/tactics/equality.mli index 5467b4af2..d979c580a 100644 --- a/tactics/equality.mli +++ b/tactics/equality.mli @@ -8,7 +8,6 @@ (*i*) open Names -open Term open Evd open EConstr open Environ diff --git a/tactics/hints.mli b/tactics/hints.mli index 467fd46d5..3a0339ff5 100644 --- a/tactics/hints.mli +++ b/tactics/hints.mli @@ -9,7 +9,6 @@ open Pp open Util open Names -open Term open EConstr open Environ open Globnames diff --git a/tactics/hipattern.mli b/tactics/hipattern.mli index dd09c3a4d..82a3d47b5 100644 --- a/tactics/hipattern.mli +++ b/tactics/hipattern.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open Evd open EConstr open Coqlib diff --git a/tactics/inv.ml b/tactics/inv.ml index 904a17417..266cac5c7 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -13,7 +13,6 @@ open Names open Nameops open Term open Termops -open Environ open EConstr open Vars open Namegen diff --git a/tactics/inv.mli b/tactics/inv.mli index 446a62f6d..5835e763d 100644 --- a/tactics/inv.mli +++ b/tactics/inv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Misctypes open Tactypes diff --git a/tactics/leminv.mli b/tactics/leminv.mli index 26d4ac994..a343fc81a 100644 --- a/tactics/leminv.mli +++ b/tactics/leminv.mli @@ -7,7 +7,6 @@ (************************************************************************) open Names -open Term open EConstr open Constrexpr open Misctypes diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 53f8e4d5f..f3f7d16b7 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -5037,8 +5037,6 @@ end (** Tacticals defined directly in term of Proofview *) module New = struct - open Proofview.Notations - open Genredexpr open Locus diff --git a/toplevel/coqloop.mli b/toplevel/coqloop.mli index 66bbf43f6..13e860a88 100644 --- a/toplevel/coqloop.mli +++ b/toplevel/coqloop.mli @@ -6,8 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp - (** The Coq toplevel loop. *) (** A buffer for the character read from a channel. We store the command diff --git a/vernac/auto_ind_decl.ml b/vernac/auto_ind_decl.ml index 015552d68..25091f783 100644 --- a/vernac/auto_ind_decl.ml +++ b/vernac/auto_ind_decl.ml @@ -9,7 +9,6 @@ (* This file is about the automatic generation of schemes about decidable equality, created by Vincent Siles, Oct 2007 *) -open Tacmach open CErrors open Util open Pp @@ -716,7 +715,6 @@ let compute_lb_goal ind lnamesparrec nparrec = ))), eff let compute_lb_tact mode lb_scheme_key ind lnamesparrec nparrec = - let open EConstr in let list_id = list_id lnamesparrec in let avoid = ref [] in let first_intros = diff --git a/vernac/classes.ml b/vernac/classes.ml index 833719965..d515b0c9b 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -66,8 +66,6 @@ let _ = Hook.set Typeclasses.classes_transparent_state_hook (fun () -> Hints.Hint_db.transparent_state (Hints.searchtable_map typeclasses_db)) -open Vernacexpr - (** TODO: add subinstances *) let existing_instance glob g info = let c = global g in diff --git a/vernac/obligations.mli b/vernac/obligations.mli index 11366fe91..a276f9f9a 100644 --- a/vernac/obligations.mli +++ b/vernac/obligations.mli @@ -12,7 +12,6 @@ open Evd open Names open Pp open Globnames -open Vernacexpr open Decl_kinds (** Forward declaration. *) diff --git a/vernac/search.ml b/vernac/search.ml index 6279b17ae..5b6e9a9c3 100644 --- a/vernac/search.ml +++ b/vernac/search.ml @@ -14,11 +14,9 @@ open Declarations open Libobject open Environ open Pattern -open Printer open Libnames open Globnames open Nametab -open Goptions module NamedDecl = Context.Named.Declaration diff --git a/vernac/search.mli b/vernac/search.mli index 82b79f75d..e34522d8a 100644 --- a/vernac/search.mli +++ b/vernac/search.mli @@ -6,7 +6,6 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) -open Pp open Names open Term open Environ -- cgit v1.2.3