(***********************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) (* head_const d | LetIn (_,_,_,d) -> head_const d | App (f,_) -> head_const f | Cast (d,_) -> head_const d | _ -> c let crible (fn : global_reference -> env -> constr -> unit) ref = let env = Global.env () in let imported = Library.opened_modules() in let const = constr_of_reference ref in let crible_rec sp lobj = match object_tag lobj with | "VARIABLE" -> (try let ((idc,_,typ),_) = get_variable (basename sp) in if (head_const typ) = const then fn (VarRef idc) env typ with Not_found -> (* we are in a section *) ()) | "CONSTANT" | "PARAMETER" -> let {const_type=typ} = Global.lookup_constant sp in if (head_const typ) = const then fn (ConstRef sp) env typ | "INDUCTIVE" -> let mib = Global.lookup_mind sp in let arities = array_map_to_list (fun mip -> (Name mip.mind_typename, None, mip.mind_nf_arity)) mib.mind_packets in (match kind_of_term const with | Ind ((sp',tyi) as indsp) -> if sp=sp' then print_constructors indsp fn env mib.mind_packets.(tyi) | _ -> ()) | _ -> () in try Library.iter_all_segments false crible_rec with Not_found -> errorlabstrm "search" [< pr_global ref; 'sPC; 'sTR "not declared" >] (* Fine Search. By Yves Bertot. *) exception No_section_path let rec head c = let c = strip_outer_cast c in match kind_of_term c with | Prod (_,_,c) -> head c | _ -> c let constr_to_section_path c = match kind_of_term c with | Const sp -> sp | _ -> raise No_section_path let xor a b = (a or b) & (not (a & b)) let plain_display ref a c = let pc = prterm_env a c in let pr = pr_global ref in mSG [< hOV 2 [< pr; 'sTR":"; 'sPC; pc >]; 'fNL>] let filter_by_module (module_list:dir_path list) (accept:bool) (ref:global_reference) (env:env) _ = try let sp = sp_of_global env ref in let sl = dirpath sp in let rec filter_aux = function | m :: tl -> (not (is_dirpath_prefix_of m sl)) && (filter_aux tl) | [] -> true in xor accept (filter_aux module_list) with No_section_path -> false let gref_eq = IndRef (make_path Coqlib.logic_module (id_of_string "eq"), 0) let gref_eqT = IndRef (make_path Coqlib.logic_type_module (id_of_string "eqT"), 0) let mk_rewrite_pattern1 eq pattern = PApp (PRef eq, [| PMeta None; pattern; PMeta None |]) let mk_rewrite_pattern2 eq pattern = PApp (PRef eq, [| PMeta None; PMeta None; pattern |]) let pattern_filter pat _ a c = try try Pattern.is_matching pat (head c) with _ -> Pattern.is_matching pat (head (Typing.type_of (Global.env()) Evd.empty c)) with UserError _ -> false let filtered_search filter_function display_function ref = crible (fun s a c -> if filter_function s a c then display_function s a c) ref let rec id_from_pattern = function | PRef gr -> gr (* should be appear as a PRef (VarRef sp) !! | PVar id -> Nametab.locate (make_qualid [] (string_of_id id)) *) | PApp (p,_) -> id_from_pattern p | _ -> error "the pattern is not simple enough" let raw_pattern_search extra_filter display_function pat = let name = id_from_pattern pat in filtered_search (fun s a c -> (pattern_filter pat s a c) & extra_filter s a c) display_function name let raw_search_rewrite extra_filter display_function pattern = filtered_search (fun s a c -> ((pattern_filter (mk_rewrite_pattern1 gref_eq pattern) s a c) || (pattern_filter (mk_rewrite_pattern2 gref_eq pattern) s a c)) && extra_filter s a c) display_function gref_eq; filtered_search (fun s a c -> ((pattern_filter (mk_rewrite_pattern1 gref_eqT pattern) s a c) || (pattern_filter (mk_rewrite_pattern2 gref_eqT pattern) s a c)) && extra_filter s a c) display_function gref_eqT let text_pattern_search extra_filter = raw_pattern_search extra_filter plain_display let text_search_rewrite extra_filter = raw_search_rewrite extra_filter plain_display let filter_by_module_from_list = function | [], _ -> (fun _ _ _ -> true) | l, outside -> filter_by_module l (not outside) let search_by_head ref inout = filtered_search (filter_by_module_from_list inout) plain_display ref let search_rewrite pat inout = text_search_rewrite (filter_by_module_from_list inout) pat let search_pattern pat inout = text_pattern_search (filter_by_module_from_list inout) pat