aboutsummaryrefslogtreecommitdiffhomepage
path: root/parsing/search.ml
blob: fc5792fa0be1012b02dd6db6862ee79600396e95 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
(***********************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team    *)
(* <O___,, *        INRIA-Rocquencourt  &  LRI-CNRS-Orsay              *)
(*   \VV/  *************************************************************)
(*    //   *      This file is distributed under the terms of the      *)
(*         *       GNU Lesser General Public License Version 2.1       *)
(***********************************************************************)

(* $Id$ *)

open Pp
open Util
open Names
open Nameops
open Term
open Rawterm
open Declarations
open Libobject
open Declare
open Coqast
open Environ
open Pattern
open Printer
open Libnames
open Nametab

(* The functions print_constructors and crible implement the behavior needed
   for the Coq searching commands.
   These functions take as first argument the procedure
   that will be called to treat each entry.  This procedure receives the name
   of the object, the assumptions that will make it possible to print its type,
   and the constr term that represent its type. *)

let print_constructors indsp fn env mip =
  let lc = mip.mind_user_lc in
  for i = 1 to Array.length lc do
    fn (ConstructRef (indsp,i)) env
      (Retyping.get_type_of env Evd.empty
	 (Pretyping.understand Evd.empty env
	    (RRef(dummy_loc, ConstructRef(indsp,i)))))
  done

let rec head_const c = match kind_of_term c with
  | Prod (_,_,d) -> 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_libraries() 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" ->
	  let kn = locate_constant (qualid_of_sp sp) in
	  let {const_type=typ} = Global.lookup_constant kn in
	  if (head_const typ) = const then fn (ConstRef kn) env typ
      | "INDUCTIVE" -> 
	  let kn = locate_mind (qualid_of_sp sp) in
          let mib = Global.lookup_mind kn 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 ((kn',tyi) as ind) -> 
		 if kn=kn' then
		   print_constructors ind fn env mib.mind_packets.(tyi)
	     | _ -> ())
      | _ -> ()
  in 
    try 
      Declaremods.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) _ _ =
  try
    let sp = sp_of_global 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 (Libnames.encode_kn Coqlib.logic_module (id_of_string "eq"), 0)
let gref_eqT =
  IndRef (Libnames.encode_kn 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


(* General search, not restricted to head constant *)

let gen_crible (fn : global_reference -> env -> constr -> unit) =
  let env = Global.env () in
  let imported = Library.opened_libraries() in
  let crible_rec (sp,_) lobj = match object_tag lobj with
    | "VARIABLE" ->
	(try 
	   let (idc,_,typ) = get_variable (basename sp) in 
           fn (VarRef idc) env typ
	 with Not_found -> (* we are in a section *) ())
    | "CONSTANT" ->
	let kn = locate_constant (qualid_of_sp sp) in
	let {const_type=typ} = Global.lookup_constant kn in
	fn (ConstRef kn) env typ
    | "INDUCTIVE" -> 
	let kn = locate_mind (qualid_of_sp sp) in
        let mib = Global.lookup_mind kn in 
        Array.iteri 
	  (fun i -> print_constructors (kn,i) fn env) mib.mind_packets
    | _ -> ()
  in 
  try 
    Declaremods.iter_all_segments false crible_rec
  with Not_found -> 
    ()

let gen_filtered_search filter_function display_function =
  gen_crible 
    (fun s a c -> if filter_function s a c then display_function s a c) 

let raw_search_about filter_modules display_function ref =
  let c = constr_of_reference ref in
  let filter ref' env typ =
    filter_modules ref' env typ &&
    Termops.occur_term c typ
  in
  gen_filtered_search filter display_function

let search_about ref inout = 
  raw_search_about (filter_by_module_from_list inout) plain_display ref