(***********************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* tactic) -> ('b -> tactic) (* Inutile ?! réécrit plus loin let tactic_com tac t x = tac (pf_interp_constr x t) x let tactic_com_sort tac t x = tac (pf_interp_type x t) x let tactic_com_list tac tl x = let translate = pf_interp_constr x in tac (List.map translate tl) x let tactic_bind_list tac tl x = let translate = pf_interp_constr x in tac (List.map (fun (b,c)->(b,translate c)) tl) x let tactic_com_bind_list tac (c,tl) x = let translate = pf_interp_constr x in tac (translate c,List.map (fun (b,c')->(b,translate c')) tl) x let tactic_com_bind_list_list tac args gl = let translate (c,tl) = (pf_interp_constr gl c, List.map (fun (b,c')->(b,pf_interp_constr gl c')) tl) in tac (List.map translate args) gl *) (********************************************************) (* Functions for hiding the implementation of a tactic. *) (********************************************************) (* hide_tactic s tac pr registers a tactic s under the name s *) let hide_tactic s tac = add_tactic s tac; (fun args -> vernac_tactic(s,args)) (* overwriting_register_tactic s tac pr registers a tactic s under the name s even if a tactic of the same name is already registered *) let overwrite_hidden_tactic s tac = overwriting_add_tactic s tac; (fun args -> vernac_tactic(s,args)) let tactic_com = fun tac t x -> tac (pf_interp_constr x t) x let tactic_opencom = fun tac t x -> tac (pf_interp_openconstr x t) x let tactic_com_sort = fun tac t x -> tac (pf_interp_type x t) x let tactic_com_list = fun tac tl x -> let translate = pf_interp_constr x in tac (List.map translate tl) x let tactic_bind_list = fun tac tl x -> let translate = pf_interp_constr x in tac (List.map (fun (b,c)->(b,translate c)) tl) x let tactic_com_bind_list = fun tac (c,tl) x -> let translate = pf_interp_constr x in tac (translate c,List.map (fun (b,c')->(b,translate c')) tl) x let tactic_com_bind_list_list = fun tac args gl -> let translate (c,tl) = (pf_interp_constr gl c, List.map (fun (b,c')->(b,pf_interp_constr gl c')) tl) in tac (List.map translate args) gl (* Some useful combinators for hiding tactic implementations *) type 'a hide_combinator = string -> ('a -> tactic) -> ('a -> tactic) let hide_atomic_tactic s tac = add_tactic s (function [] -> tac | _ -> assert false); vernac_tactic(s,[]) let overwrite_hidden_atomic_tactic s tac = overwriting_tactic s (function [] -> tac | _ -> assert false); vernac_tactic(s,[]) let hide_constr_comarg_tactic s tac = let tacfun = function | [Constr c] -> tac c | [Command com] -> tactic_com tac com | _ -> anomaly "hide_constr_comarg_tactic : neither CONSTR nor COMMAND" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[Constr c]), fun com -> vernac_tactic(s,[Command com])) let overwrite_hidden_constr_comarg_tactic s tac = let tacfun = function | [Constr c] -> tac c | [Command com] -> (fun gls -> tac (pf_interp_constr gls com) gls) | _ -> anomaly "overwrite_hidden_constr_comarg_tactic : neither CONSTR nor COMMAND" in overwriting_tactic s tacfun; (fun c -> vernac_tactic(s,[(Constr c)]), fun c -> vernac_tactic(s,[(Command c)])) let hide_constr_tactic s tac = let tacfun = function | [Constr c] -> tac c | [Command com] -> tactic_com tac com | _ -> anomaly "hide_constr_tactic : neither CONSTR nor COMMAND" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[(Constr c)])) let hide_openconstr_tactic s tac = let tacfun = function | [OpenConstr c] -> tac c | [Command com] -> tactic_opencom tac com | _ -> anomaly "hide_openconstr_tactic : neither OPENCONSTR nor COMMAND" in add_tactic s tacfun; (fun c -> vernac_tactic(s,[(OpenConstr c)])) let hide_numarg_tactic s tac = let tacfun = (function [Integer n] -> tac n | _ -> assert false) in add_tactic s tacfun; fun n -> vernac_tactic(s,[Integer n]) let hide_ident_tactic s tac = let tacfun = (function [Identifier id] -> tac id | _ -> assert false) in add_tactic s tacfun; fun id -> vernac_tactic(s,[Identifier id]) let hide_string_tactic s tac = let tacfun = (function [Quoted_string str] -> tac str | _ -> assert false) in add_tactic s tacfun; fun str -> vernac_tactic(s,[Quoted_string str]) let hide_identl_tactic s tac = let tacfun = (function [Clause idl] -> tac idl | _ -> assert false) in add_tactic s tacfun; fun idl -> vernac_tactic(s,[Clause idl]) let hide_constrl_tactic s tac = let tacfun = function | ((Command com)::_) as al -> tactic_com_list tac (List.map (function (Command com) -> com | _ -> assert false) al) | ((Constr com)::_) as al -> tac (List.map (function (Constr c) -> c | _ -> assert false) al) | _ -> anomaly "hide_constrl_tactic : neither CONSTR nor COMMAND" in add_tactic s tacfun; fun ids -> vernac_tactic(s,(List.map (fun id -> Constr id) ids)) let hide_bindl_tactic s tac = let tacfun = function | [Bindings al] -> tactic_bind_list tac al | [Cbindings al] -> tac al | _ -> anomaly "hide_bindl_tactic : neither BINDINGS nor CBINDINGS" in add_tactic s tacfun; fun bindl -> vernac_tactic(s,[Cbindings bindl]) let hide_cbindl_tactic s tac = let tacfun = function | [Command com; Bindings al] -> tactic_com_bind_list tac (com,al) | [Constr c; Cbindings al] -> tac (c,al) | _ -> anomaly "hide_cbindl_tactic : neither CONSTR nor COMMAND" in add_tactic s tacfun; fun (c,bindl) -> vernac_tactic(s,[Constr c; Cbindings bindl]) let hide_cbindll_tactic s tac = let rec getcombinds = function | ((Command com)::(Bindings al)::l) -> (com,al)::(getcombinds l) | [] -> [] | _ -> anomaly "hide_cbindll_tactic : not the expected form" in let rec getconstrbinds = function | ((Constr c)::(Cbindings al)::l) -> (c,al)::(getconstrbinds l) | [] -> [] | _ -> anomaly "hide_cbindll_tactic : not the expected form" in let rec putconstrbinds = function | (c,binds)::l -> (Constr c)::(Cbindings binds)::(putconstrbinds l) | [] -> [] in let tacfun = function | ((Command com)::_) as args -> tactic_com_bind_list_list tac (getcombinds args) | ((Constr com)::_) as args -> tac (getconstrbinds args) | _ -> anomaly "hide_cbindll_tactic : neither CONSTR nor COMMAND" in add_tactic s tacfun; fun l -> vernac_tactic(s,putconstrbinds l) (* Pretty-printers *) open Pp open Printer let pr_com sigma goal com = prterm (rename_bound_var (ids_of_named_context goal.evar_hyps) (Astterm.interp_constr sigma (Evarutil.evar_env goal) com)) let pr_one_binding sigma goal = function | (Dep id,com) -> [< pr_id id ; 'sTR":=" ; pr_com sigma goal com >] | (NoDep n,com) -> [< 'iNT n ; 'sTR":=" ; pr_com sigma goal com >] | (Com,com) -> [< pr_com sigma goal com >] let pr_bindings sigma goal lb = let prf = pr_one_binding sigma goal in match lb with | [] -> [< prlist_with_sep pr_spc prf lb >] | _ -> [<'sTR"with";'sPC;prlist_with_sep pr_spc prf lb >] let rec pr_list f = function | [] -> [<>] | a::l1 -> [< (f a) ; pr_list f l1>] let pr_gls gls = hOV 0 [< pr_decls (sig_sig gls) ; 'fNL ; pr_seq (sig_it gls) >] let pr_glls glls = hOV 0 [< pr_decls (sig_sig glls) ; 'fNL ; prlist_with_sep pr_fnl pr_seq (sig_it glls) >] let pr_tactic = Refiner.pr_tactic