From 91ff75cf42ebc883e2cfcdc4928154315984beb8 Mon Sep 17 00:00:00 2001 From: Pierre-Marie Pédrot Date: Tue, 11 Apr 2017 14:51:56 +0200 Subject: Removing tactic compatibility layer in Micromega. --- plugins/micromega/coq_micromega.ml | 57 ++++++++++++++++++-------------------- 1 file changed, 27 insertions(+), 30 deletions(-) (limited to 'plugins/micromega') diff --git a/plugins/micromega/coq_micromega.ml b/plugins/micromega/coq_micromega.ml index 4b87e6e2e..eb26c5431 100644 --- a/plugins/micromega/coq_micromega.ml +++ b/plugins/micromega/coq_micromega.ml @@ -901,16 +901,13 @@ struct coq_Qeq, Mc.OpEq ] - let has_typ gl t1 typ = - let ty = Retyping.get_type_of (Tacmach.pf_env gl) (Tacmach.project gl) t1 in - EConstr.eq_constr (Tacmach.project gl) ty typ - + type gl = { env : Environ.env; sigma : Evd.evar_map } let is_convertible gl t1 t2 = - Reductionops.is_conv (Tacmach.pf_env gl) (Tacmach.project gl) t1 t2 + Reductionops.is_conv gl.env gl.sigma t1 t2 let parse_zop gl (op,args) = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in match EConstr.kind sigma op with | Const (x,_) -> (assoc_const sigma op zop_table, args.(0) , args.(1)) | Ind((n,0),_) -> @@ -920,7 +917,7 @@ struct | _ -> failwith "parse_zop" let parse_rop gl (op,args) = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in match EConstr.kind sigma op with | Const (x,_) -> (assoc_const sigma op rop_table, args.(0) , args.(1)) | Ind((n,0),_) -> @@ -930,7 +927,7 @@ struct | _ -> failwith "parse_zop" let parse_qop gl (op,args) = - (assoc_const (Tacmach.project gl) op qop_table, args.(0) , args.(1)) + (assoc_const gl.sigma op qop_table, args.(0) , args.(1)) let is_constant sigma t = (* This is an approx *) match EConstr.kind sigma t with @@ -1154,7 +1151,7 @@ struct rop_spec let parse_arith parse_op parse_expr env cstr gl = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in if debug then Feedback.msg_debug (Pp.str "parse_arith: " ++ Printer.pr_leconstr cstr ++ fnl ()); match EConstr.kind sigma cstr with @@ -1199,7 +1196,7 @@ struct *) let parse_formula gl parse_atom env tg term = - let sigma = Tacmach.project gl in + let sigma = gl.sigma in let parse_atom env tg t = try @@ -1208,7 +1205,7 @@ struct with e when CErrors.noncritical e -> (X(t),env,tg) in let is_prop term = - let sort = Retyping.get_sort_of (Tacmach.pf_env gl) (Tacmach.project gl) term in + let sort = Retyping.get_sort_of gl.env gl.sigma term in Sorts.is_prop sort in let rec xparse_formula env tg term = @@ -1720,7 +1717,6 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* let vm = dump_varmap (spec.typ) (vm_of_list env) in (* todo : directly generate the proof term - or generalize before conversion? *) Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in Tacticals.New.tclTHENLIST [ Tactics.change_concl @@ -1730,7 +1726,7 @@ let micromega_order_change spec cert cert_typ env ff (*: unit Proofview.tactic* ("__varmap", vm, Term.mkApp(Lazy.force coq_VarMap, [|spec.typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl)) + (Tacmach.New.pf_concl gl)) ] end } @@ -1967,11 +1963,13 @@ let micromega_tauto negate normalise unsat deduce spec prover env polys1 polys2 Some (ids,ff',res') - (** * Parse the proof environment, and call micromega_tauto *) +let fresh_id avoid id gl = + Tactics.fresh_id_in_env avoid id (Proofview.Goal.env gl) + let micromega_gen parse_arith (negate:'cst atom -> 'cst mc_cnf) @@ -1979,17 +1977,17 @@ let micromega_gen unsat deduce spec dumpexpr prover tac = Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in - let sigma = Tacmach.project gl in - let concl = Tacmach.pf_concl gl in - let hyps = Tacmach.pf_hyps_types gl in + let sigma = Tacmach.New.project gl in + let concl = Tacmach.New.pf_concl gl in + let hyps = Tacmach.New.pf_hyps_types gl in try - let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in + let gl0 = { env = Tacmach.New.pf_env gl; sigma } in + let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let dumpexpr = Lazy.force dumpexpr in - match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl with + match micromega_tauto negate normalise unsat deduce spec prover env hyps concl gl0 with | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Some (ids,ff',res') -> let (arith_goal,props,vars,ff_arith) = make_goal_of_formula sigma dumpexpr ff' in @@ -1998,7 +1996,7 @@ let micromega_gen let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in - let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in + let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; @@ -2057,7 +2055,6 @@ let micromega_order_changer cert env ff = let ff = dump_formula formula_typ (dump_cstr coeff dump_coeff) ff in let vm = dump_varmap (typ) (vm_of_list env) in Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in Tacticals.New.tclTHENLIST [ (Tactics.change_concl @@ -2069,7 +2066,7 @@ let micromega_order_changer cert env ff = [["Coq" ; "micromega" ; "VarMap"] ; ["VarMap"]] "t"), [|typ|])); ("__wit", cert, cert_typ) ] - (Tacmach.pf_concl gl))); + (Tacmach.New.pf_concl gl))); (* Tacticals.New.tclTHENLIST (List.map (fun id -> (Tactics.introduction id)) ids)*) ] end } @@ -2088,20 +2085,20 @@ let micromega_genr prover tac = dump_proof = dump_psatz coq_Q dump_q } in Proofview.Goal.nf_enter { enter = begin fun gl -> - let gl = Tacmach.New.of_old (fun x -> x) gl in - let sigma = Tacmach.project gl in - let concl = Tacmach.pf_concl gl in - let hyps = Tacmach.pf_hyps_types gl in + let sigma = Tacmach.New.project gl in + let concl = Tacmach.New.pf_concl gl in + let hyps = Tacmach.New.pf_hyps_types gl in try - let (hyps,concl,env) = parse_goal gl parse_arith Env.empty hyps concl in + let gl0 = { env = Tacmach.New.pf_env gl; sigma } in + let (hyps,concl,env) = parse_goal gl0 parse_arith Env.empty hyps concl in let env = Env.elements env in let spec = Lazy.force spec in let hyps' = List.map (fun (n,f) -> (n, map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) f)) hyps in let concl' = map_atoms (Micromega.map_Formula Micromega.q_of_Rcst) concl in - match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl with + match micromega_tauto negate normalise unsat deduce spec prover env hyps' concl' gl0 with | None -> Tacticals.New.tclFAIL 0 (Pp.str " Cannot find witness") | Some (ids,ff',res') -> let (ff,ids) = formula_hyps_concl @@ -2114,7 +2111,7 @@ let micromega_genr prover tac = let intro_vars = Tacticals.New.tclTHENLIST (List.map intro vars) in let intro_props = Tacticals.New.tclTHENLIST (List.map intro props) in let ipat_of_name id = Some (Loc.ghost, Misctypes.IntroNaming (Misctypes.IntroIdentifier id)) in - let goal_name = Tactics.fresh_id [] (Names.Id.of_string "__arith") gl in + let goal_name = fresh_id [] (Names.Id.of_string "__arith") gl in let env' = List.map (fun (id,i) -> Term.mkVar id,i) vars in let tac_arith = Tacticals.New.tclTHENLIST [ intro_props ; intro_vars ; -- cgit v1.2.3 From a5f33b5fda592c2dfaed0c16d3773b35e7acab28 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:52:33 +0200 Subject: Disambiguate Polynomial.Hyp and Mfourier.Hyp -> Assum --- plugins/micromega/mfourier.ml | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) (limited to 'plugins/micromega') diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml index f4f9b3c2f..377994415 100644 --- a/plugins/micromega/mfourier.ml +++ b/plugins/micromega/mfourier.ml @@ -99,7 +99,7 @@ module PSet = ISet module System = Hashtbl.Make(Vect) type proof = -| Hyp of int +| Assum of int | Elim of var * proof * proof | And of proof * proof @@ -134,7 +134,7 @@ exception SystemContradiction of proof let hyps prf = let rec hyps prf acc = match prf with - | Hyp i -> ISet.add i acc + | Assum i -> ISet.add i acc | Elim(_,prf1,prf2) | And(prf1,prf2) -> hyps prf1 (hyps prf2 acc) in hyps prf ISet.empty @@ -143,7 +143,7 @@ let hyps prf = (** Pretty printing *) let rec pp_proof o prf = match prf with - | Hyp i -> Printf.fprintf o "H%i" i + | Assum i -> Printf.fprintf o "H%i" i | Elim(v, prf1,prf2) -> Printf.fprintf o "E(%i,%a,%a)" v pp_proof prf1 pp_proof prf2 | And(prf1,prf2) -> Printf.fprintf o "A(%a,%a)" pp_proof prf1 pp_proof prf2 @@ -270,7 +270,7 @@ let norm_cstr {coeffs = v ; op = o ; cst = c} idx = (match o with | Eq -> Some c , Some c | Ge -> Some c , None) ; - prf = Hyp idx } + prf = Assum idx } (** [load_system l] takes a list of constraints of type [cstr_compat] @@ -285,7 +285,7 @@ let load_system l = let vars = List.fold_left (fun vrs (cstr,i) -> match norm_cstr cstr i with - | Contradiction -> raise (SystemContradiction (Hyp i)) + | Contradiction -> raise (SystemContradiction (Assum i)) | Redundant -> vrs | Cstr(vect,info) -> xadd_cstr vect info sys ; @@ -867,7 +867,7 @@ let mk_proof hyps prf = let rec mk_proof prf = match prf with - | Hyp i -> [ ([i, Int 1] , List.nth hyps i) ] + | Assum i -> [ ([i, Int 1] , List.nth hyps i) ] | Elim(v,prf1,prf2) -> let prfsl = mk_proof prf1 -- cgit v1.2.3 From 95239fa33c5da60080833dabc3ced246680dfdf9 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 19:56:05 +0200 Subject: Rename Sos_lib.(||) -> parser_or to avoid (deprecated) Pervasives.or --- plugins/micromega/sos.ml | 19 +++++++++++-------- plugins/micromega/sos_lib.ml | 2 +- 2 files changed, 12 insertions(+), 9 deletions(-) (limited to 'plugins/micromega') diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml index cc89e2b9d..7e716258c 100644 --- a/plugins/micromega/sos.ml +++ b/plugins/micromega/sos.ml @@ -282,14 +282,14 @@ let poly_variables (p:poly) = (* Order monomials for human presentation. *) (* ------------------------------------------------------------------------- *) -let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 or x1 = x2 && k1 > k2;; +let humanorder_varpow (x1,k1) (x2,k2) = x1 < x2 || x1 = x2 && k1 > k2;; let humanorder_monomial = let rec ord l1 l2 = match (l1,l2) with _,[] -> true | [],_ -> false - | h1::t1,h2::t2 -> humanorder_varpow h1 h2 or h1 = h2 && ord t1 t2 in - fun m1 m2 -> m1 = m2 or + | h1::t1,h2::t2 -> humanorder_varpow h1 h2 || h1 = h2 && ord t1 t2 in + fun m1 m2 -> m1 = m2 || ord (sort humanorder_varpow (graph m1)) (sort humanorder_varpow (graph m2));; @@ -466,6 +466,7 @@ let token s = >> (fun ((_,t),_) -> t);; let decimal = + let (||) = parser_or in let numeral = some isnum in let decimalint = atleast 1 numeral >> ((o) Num.num_of_string implode) in let decimalfrac = atleast 1 numeral @@ -492,6 +493,7 @@ let parse_decimal = mkparser decimal;; (* ------------------------------------------------------------------------- *) let parse_sdpaoutput,parse_csdpoutput = + let (||) = parser_or in let vector = token "{" ++ listof decimal (token ",") "decimal" ++ token "}" >> (fun ((_,v),_) -> vector_of_list v) in @@ -512,6 +514,7 @@ let parse_sdpaoutput,parse_csdpoutput = (* ------------------------------------------------------------------------- *) let sdpa_run_succeeded = + let (||) = parser_or in let rec skipupto dscr prs inp = (dscr ++ prs >> snd || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in @@ -602,7 +605,7 @@ let run_csdp dbg obj mats = let csdp obj mats = let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (* Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline() *) @@ -653,7 +656,7 @@ let linear_program_basic a = let mats = List.map (fun j -> diagonal (column j a)) (1--n) and obj = vector_const (Int 1) m in let rv,res = run_csdp false obj mats in - if rv = 1 or rv = 2 then false + if rv = 1 || rv = 2 then false else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; @@ -667,7 +670,7 @@ let linear_program a b = let mats = diagonal b :: List.map (fun j -> diagonal (column j a)) (1--n) and obj = vector_const (Int 1) m in let rv,res = run_csdp false obj mats in - if rv = 1 or rv = 2 then false + if rv = 1 || rv = 2 then false else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; @@ -968,7 +971,7 @@ let run_csdp dbg nblocks blocksizes obj mats = let csdp nblocks blocksizes obj mats = let rv,res = run_csdp (!debugging) nblocks blocksizes obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (*Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline() *) @@ -1439,7 +1442,7 @@ let run_csdp dbg obj mats = let csdp obj mats = let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 or rv = 2 then failwith "csdp: Problem is infeasible" + (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" else if rv = 3 then () (* (Format.print_string "csdp warning: Reduced accuracy"; Format.print_newline()) *) diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml index f54914f25..9ee3db6e0 100644 --- a/plugins/micromega/sos_lib.ml +++ b/plugins/micromega/sos_lib.ml @@ -525,7 +525,7 @@ let isspace,issep,isbra,issymb,isalpha,isnum,isalnum = and isalnum c = Array.get ctable (charcode c) >= 16 in isspace,issep,isbra,issymb,isalpha,isnum,isalnum;; -let (||) parser1 parser2 input = +let parser_or parser1 parser2 input = try parser1 input with Noparse -> parser2 input;; -- cgit v1.2.3 From e1d2a898feacbe4bd363818f259bce5fd74c2ee7 Mon Sep 17 00:00:00 2001 From: Gaetan Gilbert Date: Fri, 21 Apr 2017 20:09:23 +0200 Subject: Micromega: do not use Filename.temp_dir_path, remove unused values --- plugins/micromega/sos.ml | 257 ++----------------------------------------- plugins/micromega/sos_lib.ml | 2 +- 2 files changed, 11 insertions(+), 248 deletions(-) (limited to 'plugins/micromega') diff --git a/plugins/micromega/sos.ml b/plugins/micromega/sos.ml index 7e716258c..e1ceabe9e 100644 --- a/plugins/micromega/sos.ml +++ b/plugins/micromega/sos.ml @@ -21,8 +21,6 @@ let debugging = ref false;; exception Sanity;; -exception Unsolvable;; - (* ------------------------------------------------------------------------- *) (* Turn a rational into a decimal string with d sig digits. *) (* ------------------------------------------------------------------------- *) @@ -99,28 +97,11 @@ let vector_const c n = if c =/ Int 0 then vector_0 n else (n,itlist (fun k -> k |-> c) (1--n) undefined :vector);; -let vector_1 = vector_const (Int 1);; - let vector_cmul c (v:vector) = let n = dim v in if c =/ Int 0 then vector_0 n else n,mapf (fun x -> c */ x) (snd v) -let vector_neg (v:vector) = (fst v,mapf minus_num (snd v) :vector);; - -let vector_add (v1:vector) (v2:vector) = - let m = dim v1 and n = dim v2 in - if m <> n then failwith "vector_add: incompatible dimensions" else - (n,combine (+/) (fun x -> x =/ Int 0) (snd v1) (snd v2) :vector);; - -let vector_sub v1 v2 = vector_add v1 (vector_neg v2);; - -let vector_dot (v1:vector) (v2:vector) = - let m = dim v1 and n = dim v2 in - if m <> n then failwith "vector_add: incompatible dimensions" else - foldl (fun a i x -> x +/ a) (Int 0) - (combine ( */ ) (fun x -> x =/ Int 0) (snd v1) (snd v2));; - let vector_of_list l = let n = List.length l in (n,itlist2 (|->) (1--n) l undefined :vector);; @@ -133,13 +114,6 @@ let matrix_0 (m,n) = ((m,n),undefined:matrix);; let dimensions (m:matrix) = fst m;; -let matrix_const c (m,n as mn) = - if m <> n then failwith "matrix_const: needs to be square" - else if c =/ Int 0 then matrix_0 mn - else (mn,itlist (fun k -> (k,k) |-> c) (1--n) undefined :matrix);; - -let matrix_1 = matrix_const (Int 1);; - let matrix_cmul c (m:matrix) = let (i,j) = dimensions m in if c =/ Int 0 then matrix_0 (i,j) @@ -152,8 +126,6 @@ let matrix_add (m1:matrix) (m2:matrix) = if d1 <> d2 then failwith "matrix_add: incompatible dimensions" else (d1,combine (+/) (fun x -> x =/ Int 0) (snd m1) (snd m2) :matrix);; -let matrix_sub m1 m2 = matrix_add m1 (matrix_neg m2);; - let row k (m:matrix) = let i,j = dimensions m in (j, @@ -166,20 +138,10 @@ let column k (m:matrix) = foldl (fun a (i,j) c -> if j = k then (i |-> c) a else a) undefined (snd m) : vector);; -let transp (m:matrix) = - let i,j = dimensions m in - ((j,i),foldl (fun a (i,j) c -> ((j,i) |-> c) a) undefined (snd m) :matrix);; - let diagonal (v:vector) = let n = dim v in ((n,n),foldl (fun a i c -> ((i,i) |-> c) a) undefined (snd v) : matrix);; -let matrix_of_list l = - let m = List.length l in - if m = 0 then matrix_0 (0,0) else - let n = List.length (List.hd l) in - (m,n),itern 1 l (fun v i -> itern 1 v (fun c j -> (i,j) |-> c)) undefined;; - (* ------------------------------------------------------------------------- *) (* Monomials. *) (* ------------------------------------------------------------------------- *) @@ -195,24 +157,8 @@ let monomial_var x = (x |=> 1 :monomial);; let (monomial_mul:monomial->monomial->monomial) = combine (+) (fun x -> false);; -let monomial_pow (m:monomial) k = - if k = 0 then monomial_1 - else mapf (fun x -> k * x) m;; - -let monomial_divides (m1:monomial) (m2:monomial) = - foldl (fun a x k -> tryapplyd m2 x 0 >= k && a) true m1;; - -let monomial_div (m1:monomial) (m2:monomial) = - let m = combine (+) (fun x -> x = 0) m1 (mapf (fun x -> -x) m2) in - if foldl (fun a x k -> k >= 0 && a) true m then m - else failwith "monomial_div: non-divisible";; - let monomial_degree x (m:monomial) = tryapplyd m x 0;; -let monomial_lcm (m1:monomial) (m2:monomial) = - (itlist (fun x -> x |-> max (monomial_degree x m1) (monomial_degree x m2)) - (union (dom m1) (dom m2)) undefined :monomial);; - let monomial_multidegree (m:monomial) = foldl (fun a x k -> k + a) 0 m;; let monomial_variables m = dom m;; @@ -252,12 +198,6 @@ let poly_cmmul (c,m) (p:poly) = let poly_mul (p1:poly) (p2:poly) = foldl (fun a m c -> poly_add (poly_cmmul (c,m) p2) a) poly_0 p1;; -let poly_div (p1:poly) (p2:poly) = - if not(poly_isconst p2) then failwith "poly_div: non-constant" else - let c = eval undefined p2 in - if c =/ Int 0 then failwith "poly_div: division by zero" - else poly_cmul (Int 1 // c) p1;; - let poly_square p = poly_mul p p;; let rec poly_pow p k = @@ -266,10 +206,6 @@ let rec poly_pow p k = else let q = poly_square(poly_pow p (k / 2)) in if k mod 2 = 1 then poly_mul p q else q;; -let poly_exp p1 p2 = - if not(poly_isconst p2) then failwith "poly_exp: not a constant" else - poly_pow p1 (Num.int_of_num (eval undefined p2));; - let degree x (p:poly) = foldl (fun a m c -> max (monomial_degree x m) a) 0 p;; let multidegree (p:poly) = @@ -297,42 +233,8 @@ let humanorder_monomial = (* Conversions to strings. *) (* ------------------------------------------------------------------------- *) -let string_of_vector min_size max_size (v:vector) = - let n_raw = dim v in - if n_raw = 0 then "[]" else - let n = max min_size (min n_raw max_size) in - let xs = List.map ((o) string_of_num (element v)) (1--n) in - "[" ^ end_itlist (fun s t -> s ^ ", " ^ t) xs ^ - (if n_raw > max_size then ", ...]" else "]");; - -let string_of_matrix max_size (m:matrix) = - let i_raw,j_raw = dimensions m in - let i = min max_size i_raw and j = min max_size j_raw in - let rstr = List.map (fun k -> string_of_vector j j (row k m)) (1--i) in - "["^end_itlist(fun s t -> s^";\n "^t) rstr ^ - (if j > max_size then "\n ...]" else "]");; - let string_of_vname (v:vname): string = (v: string);; -let rec string_of_term t = - match t with - Opp t1 -> "(- " ^ string_of_term t1 ^ ")" -| Add (t1, t2) -> - "(" ^ (string_of_term t1) ^ " + " ^ (string_of_term t2) ^ ")" -| Sub (t1, t2) -> - "(" ^ (string_of_term t1) ^ " - " ^ (string_of_term t2) ^ ")" -| Mul (t1, t2) -> - "(" ^ (string_of_term t1) ^ " * " ^ (string_of_term t2) ^ ")" -| Inv t1 -> "(/ " ^ string_of_term t1 ^ ")" -| Div (t1, t2) -> - "(" ^ (string_of_term t1) ^ " / " ^ (string_of_term t2) ^ ")" -| Pow (t1, n1) -> - "(" ^ (string_of_term t1) ^ " ^ " ^ (string_of_int n1) ^ ")" -| Zero -> "0" -| Var v -> "x" ^ (string_of_vname v) -| Const x -> string_of_num x;; - - let string_of_varpow x k = if k = 1 then string_of_vname x else string_of_vname x^"^"^string_of_int k;; @@ -363,6 +265,7 @@ let string_of_poly (p:poly) = (* Printers. *) (* ------------------------------------------------------------------------- *) +(* let print_vector v = Format.print_string(string_of_vector 0 20 v);; let print_matrix m = Format.print_string(string_of_matrix 20 m);; @@ -371,7 +274,6 @@ let print_monomial m = Format.print_string(string_of_monomial m);; let print_poly m = Format.print_string(string_of_poly m);; -(* #install_printer print_vector;; #install_printer print_matrix;; #install_printer print_monomial;; @@ -410,19 +312,6 @@ let sdpa_of_vector (v:vector) = let strs = List.map (o (decimalize 20) (element v)) (1--n) in end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";; -(* ------------------------------------------------------------------------- *) -(* String for block diagonal matrix numbered k. *) -(* ------------------------------------------------------------------------- *) - -let sdpa_of_blockdiagonal k m = - let pfx = string_of_int k ^" " in - let ents = - foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in - let entss = sort (increasing fst) ents in - itlist (fun ((b,i,j),c) a -> - pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ - " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; - (* ------------------------------------------------------------------------- *) (* String for a matrix numbered k, in SDPA sparse format. *) (* ------------------------------------------------------------------------- *) @@ -486,13 +375,11 @@ let mkparser p s = let x,rst = p(explode s) in if rst = [] then x else failwith "mkparser: unparsed input";; -let parse_decimal = mkparser decimal;; - (* ------------------------------------------------------------------------- *) (* Parse back a vector. *) (* ------------------------------------------------------------------------- *) -let parse_sdpaoutput,parse_csdpoutput = +let _parse_sdpaoutput, parse_csdpoutput = let (||) = parser_or in let vector = token "{" ++ listof decimal (token ",") "decimal" ++ token "}" @@ -509,25 +396,11 @@ let parse_sdpaoutput,parse_csdpoutput = (a " " ++ a "\n" ++ ignore) >> ((o) vector_of_list fst) in mkparser sdpaoutput,mkparser csdpoutput;; -(* ------------------------------------------------------------------------- *) -(* Also parse the SDPA output to test success (CSDP yields a return code). *) -(* ------------------------------------------------------------------------- *) - -let sdpa_run_succeeded = - let (||) = parser_or in - let rec skipupto dscr prs inp = - (dscr ++ prs >> snd - || some (fun c -> true) ++ skipupto dscr prs >> snd) inp in - let prs = skipupto (word "phase.value" ++ token "=") - (possibly (a "p") ++ possibly (a "d") ++ - (word "OPT" || word "FEAS")) in - fun s -> try ignore (prs (explode s)); true with Noparse -> false;; - (* ------------------------------------------------------------------------- *) (* The default parameters. Unfortunately this goes to a fixed file. *) (* ------------------------------------------------------------------------- *) -let sdpa_default_parameters = +let _sdpa_default_parameters = "100 unsigned int maxIteration;\ \n1.0E-7 double 0.0 < epsilonStar;\ \n1.0E2 double 0.0 < lambdaStar;\ @@ -558,7 +431,7 @@ let sdpa_alt_parameters = \n1.0E-7 double 0.0 < epsilonDash;\ \n";; -let sdpa_params = sdpa_alt_parameters;; +let _sdpa_params = sdpa_alt_parameters;; (* ------------------------------------------------------------------------- *) (* CSDP parameters; so far I'm sticking with the defaults. *) @@ -591,10 +464,10 @@ let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in @@ -603,16 +476,6 @@ let run_csdp dbg obj mats = else (Sys.remove input_file; Sys.remove output_file)); rv,res);; -let csdp obj mats = - let rv,res = run_csdp (!debugging) obj mats in - (if rv = 1 || rv = 2 then failwith "csdp: Problem is infeasible" - else if rv = 3 then () - (* Format.print_string "csdp warning: Reduced accuracy"; - Format.print_newline() *) - else if rv <> 0 then failwith("csdp: error "^string_of_int rv) - else ()); - res;; - (* ------------------------------------------------------------------------- *) (* Try some apparently sensible scaling first. Note that this is purely to *) (* get a cleaner translation to floating-point, and doesn't affect any of *) @@ -660,20 +523,6 @@ let linear_program_basic a = else if rv = 0 then true else failwith "linear_program: An error occurred in the SDP solver";; -(* ------------------------------------------------------------------------- *) -(* Alternative interface testing A x >= b for matrix A, vector b. *) -(* ------------------------------------------------------------------------- *) - -let linear_program a b = - let m,n = dimensions a in - if dim b <> m then failwith "linear_program: incompatible dimensions" else - let mats = diagonal b :: List.map (fun j -> diagonal (column j a)) (1--n) - and obj = vector_const (Int 1) m in - let rv,res = run_csdp false obj mats in - if rv = 1 || rv = 2 then false - else if rv = 0 then true - else failwith "linear_program: An error occurred in the SDP solver";; - (* ------------------------------------------------------------------------- *) (* Test whether a point is in the convex hull of others. Rather than use *) (* computational geometry, express as linear inequalities and call CSDP. *) @@ -718,40 +567,6 @@ let equation_eval assig eq = let value v = apply assig v in foldl (fun a v c -> a +/ value(v) */ c) (Int 0) eq;; -(* ------------------------------------------------------------------------- *) -(* Eliminate among linear equations: return unconstrained variables and *) -(* assignments for the others in terms of them. We give one pseudo-variable *) -(* "one" that's used for a constant term. *) -(* ------------------------------------------------------------------------- *) - -let failstore = ref [];; - -let eliminate_equations = - let rec extract_first p l = - match l with - [] -> failwith "extract_first" - | h::t -> if p(h) then h,t else - let k,s = extract_first p t in - k,h::s in - let rec eliminate vars dun eqs = - match vars with - [] -> if forall is_undefined eqs then dun - else (failstore := [vars,dun,eqs]; raise Unsolvable) - | v::vs -> - try let eq,oeqs = extract_first (fun e -> defined e v) eqs in - let a = apply eq v in - let eq' = equation_cmul (Int(-1) // a) (undefine v eq) in - let elim e = - let b = tryapplyd e v (Int 0) in - if b =/ Int 0 then e else - equation_add e (equation_cmul (minus_num b // a) eq) in - eliminate vs ((v |-> eq') (mapf elim dun)) (List.map elim oeqs) - with Failure _ -> eliminate vs dun eqs in - fun one vars eqs -> - let assig = eliminate vars undefined eqs in - let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in - setify vs,assig;; - (* ------------------------------------------------------------------------- *) (* Eliminate all variables, in an essentially arbitrary order. *) (* ------------------------------------------------------------------------- *) @@ -782,18 +597,6 @@ let eliminate_all_equations one = let vs = foldl (fun a x f -> subtract (dom f) [one] @ a) [] assig in setify vs,assig;; -(* ------------------------------------------------------------------------- *) -(* Solve equations by assigning arbitrary numbers. *) -(* ------------------------------------------------------------------------- *) - -let solve_equations one eqs = - let vars,assigs = eliminate_all_equations one eqs in - let vfn = itlist (fun v -> (v |-> Int 0)) vars (one |=> Int(-1)) in - let ass = - combine (+/) (fun c -> false) (mapf (equation_eval vfn) assigs) vfn in - if forall (fun e -> equation_eval ass e =/ Int 0) eqs - then undefine one ass else raise Sanity;; - (* ------------------------------------------------------------------------- *) (* Hence produce the "relevant" monomials: those whose squares lie in the *) (* Newton polytope of the monomials in the input. (This is enough according *) @@ -900,19 +703,6 @@ let epoly_pmul p q acc = (m |-> equation_add (equation_cmul c e) es) b) a q) acc p;; -(* ------------------------------------------------------------------------- *) -(* Usual operations on equation-parametrized poly. *) -(* ------------------------------------------------------------------------- *) - -let epoly_cmul c l = - if c =/ Int 0 then undefined else mapf (equation_cmul c) l;; - -let epoly_neg = epoly_cmul (Int(-1));; - -let epoly_add = combine equation_add is_undefined;; - -let epoly_sub p q = epoly_add p (epoly_neg q);; - (* ------------------------------------------------------------------------- *) (* Convert regular polynomial. Note that we treat (0,0,0) as -1. *) (* ------------------------------------------------------------------------- *) @@ -956,11 +746,11 @@ let run_csdp dbg nblocks blocksizes obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_blockproblem "" nblocks blocksizes obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in @@ -991,8 +781,6 @@ let bmatrix_cmul c bm = let bmatrix_neg = bmatrix_cmul (Int(-1));; -let bmatrix_sub m1 m2 = bmatrix_add m1 (bmatrix_neg m2);; - (* ------------------------------------------------------------------------- *) (* Smash a block matrix into components. *) (* ------------------------------------------------------------------------- *) @@ -1104,15 +892,6 @@ let real_positivnullstellensatz_general linf d eqs leqs pol = if not(is_undefined sanity) then raise Sanity else cfs,List.map (fun (a,b) -> snd a,b) msq;; -(* ------------------------------------------------------------------------- *) -(* Iterative deepening. *) -(* ------------------------------------------------------------------------- *) - -let rec deepen f n = - try print_string "Searching with depth limit "; - print_int n; print_newline(); f n - with Failure _ -> deepen f (n + 1);; - (* ------------------------------------------------------------------------- *) (* The ordering so we can create canonical HOL polynomials. *) (* ------------------------------------------------------------------------- *) @@ -1139,10 +918,6 @@ let monomial_order = if deg1 < deg2 then false else if deg1 > deg2 then true else lexorder mon1 mon2;; -let dest_poly p = - List.map (fun (m,c) -> c,dest_monomial m) - (sort (fun (m1,_) (m2,_) -> monomial_order m1 m2) (graph p));; - (* ------------------------------------------------------------------------- *) (* Map back polynomials and their composites to HOL. *) (* ------------------------------------------------------------------------- *) @@ -1376,9 +1151,6 @@ let rec allpermutations l = itlist (fun h acc -> List.map (fun t -> h::t) (allpermutations (subtract l [h])) @ acc) l [];; -let allvarorders l = - List.map (fun vlis x -> index x vlis) (allpermutations l);; - let changevariables_monomial zoln (m:monomial) = foldl (fun a x k -> (List.assoc x zoln |-> k) a) monomial_1 m;; @@ -1395,15 +1167,6 @@ let sdpa_of_vector (v:vector) = let strs = List.map (o (decimalize 20) (element v)) (1--n) in end_itlist (fun x y -> x ^ " " ^ y) strs ^ "\n";; -let sdpa_of_blockdiagonal k m = - let pfx = string_of_int k ^" " in - let ents = - foldl (fun a (b,i,j) c -> if i > j then a else ((b,i,j),c)::a) [] m in - let entss = sort (increasing fst) ents in - itlist (fun ((b,i,j),c) a -> - pfx ^ string_of_int b ^ " " ^ string_of_int i ^ " " ^ string_of_int j ^ - " " ^ decimalize 20 c ^ "\n" ^ a) entss "";; - let sdpa_of_matrix k (m:matrix) = let pfx = string_of_int k ^ " 1 " in let ms = foldr (fun (i,j) c a -> if i > j then a else ((i,j),c)::a) @@ -1428,10 +1191,10 @@ let run_csdp dbg obj mats = let input_file = Filename.temp_file "sos" ".dat-s" in let output_file = String.sub input_file 0 (String.length input_file - 6) ^ ".out" - and params_file = Filename.concat (!temp_path) "param.csdp" in + and params_file = Filename.concat temp_path "param.csdp" in file_of_string input_file (sdpa_of_problem "" obj mats); file_of_string params_file csdp_params; - let rv = Sys.command("cd "^(!temp_path)^"; csdp "^input_file ^ + let rv = Sys.command("cd "^temp_path^"; csdp "^input_file ^ " " ^ output_file ^ (if dbg then "" else "> /dev/null")) in let op = string_of_file output_file in diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml index 9ee3db6e0..6b8b820ac 100644 --- a/plugins/micromega/sos_lib.ml +++ b/plugins/micromega/sos_lib.ml @@ -571,7 +571,7 @@ let finished input = (* ------------------------------------------------------------------------- *) -let temp_path = ref Filename.temp_dir_name;; +let temp_path = Filename.get_temp_dir_name ();; (* ------------------------------------------------------------------------- *) (* Convenient conversion between files and (lists of) strings. *) -- 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/micromega') 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