aboutsummaryrefslogtreecommitdiffhomepage
path: root/pretyping/typeclasses.ml
blob: 4c944f4608e96f11c2b557eb89ef404cd61cfc8c (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
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
319
320
321
322
323
324
325
326
327
328
329
330
331
332
333
334
335
336
337
338
339
340
341
342
343
344
345
346
347
348
349
350
351
352
353
354
355
356
357
358
359
360
361
362
363
364
365
366
367
368
369
370
371
372
373
374
375
376
377
378
379
380
381
382
383
384
385
386
387
388
389
390
(************************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team     *)
(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
(*   \VV/  **************************************************************)
(*    //   *      This file is distributed under the terms of the       *)
(*         *       GNU Lesser General Public License Version 2.1        *)
(************************************************************************)

(*i $Id:$ i*)

(*i*)
open Names
open Libnames
open Decl_kinds
open Term
open Sign
open Evd
open Environ
open Nametab
open Mod_subst
open Util
open Typeclasses_errors
(*i*)

let mismatched_params env n m = mismatched_ctx_inst env Parameters n m
(* let mismatched_defs env n m = mismatched_ctx_inst env Definitions n m *)
let mismatched_props env n m = mismatched_ctx_inst env Properties n m

type rels = constr list

(* This module defines type-classes *)
type typeclass = {
  (* The class implementation *)
  cl_impl : global_reference; 

  (* Context in which the definitions are typed. Includes both typeclass parameters and superclasses. *)
  cl_context : ((global_reference * bool) option * named_declaration) list; 

  cl_params: int;

  (* Context of definitions and properties on defs, will not be shared *)
  cl_props : named_context;
  
  (* The method implementaions as projections. *)
  cl_projs : constant list;
}

type typeclasses = (global_reference, typeclass) Gmap.t

type instance = {
  is_class: typeclass;
  is_pri: int option;
  is_impl: constant; 
}

type instances = (global_reference, instance list) Gmap.t
    
let classes : typeclasses ref = ref Gmap.empty

let methods : (constant, global_reference) Gmap.t ref = ref Gmap.empty
  
let instances : instances ref = ref Gmap.empty
  
let freeze () = !classes, !methods, !instances

let unfreeze (cl,m,is) = 
  classes:=cl;
  methods:=m;
  instances:=is
    
let init () =
  classes:= Gmap.empty; 
  methods:= Gmap.empty;
  instances:= Gmap.empty
    
let _ = 
  Summary.declare_summary "classes_and_instances"
    { Summary.freeze_function = freeze;
      Summary.unfreeze_function = unfreeze;
      Summary.init_function = init;
      Summary.survive_module = false;
      Summary.survive_section = true }

let gmap_merge old ne =
  Gmap.fold (fun k v acc -> Gmap.add k v acc) ne old

let gmap_list_merge old upd ne =
  let ne = 
    Gmap.fold (fun k v acc -> 
      let oldv = try Gmap.find k old with Not_found -> [] in
      let v' = 
	List.fold_left (fun acc x -> 
	  if not (List.exists (fun y -> y.is_impl = x.is_impl) v) then
	    (x :: acc) else acc)
	v oldv
      in Gmap.add k v' acc)
      ne Gmap.empty
  in
    Gmap.fold (fun k v acc -> 
      let newv = try Gmap.find k ne with Not_found -> [] in
      let v' = 
	List.fold_left (fun acc x -> 
	  if not (List.exists (fun y -> y.is_impl = x.is_impl) acc) then x :: acc else acc)
	  newv v
      in Gmap.add k v' acc)
      old ne

let add_instance_hint_ref = ref (fun id pri -> assert false)
let register_add_instance_hint = (:=) add_instance_hint_ref
let add_instance_hint id = !add_instance_hint_ref id

let qualid_of_con c = 
  Qualid (dummy_loc, shortest_qualid_of_global Idset.empty (ConstRef c))

let cache (_, (cl, m, inst)) =
  classes := gmap_merge !classes cl;
  methods := gmap_merge !methods m;
  instances := gmap_list_merge !instances 
    (fun i -> add_instance_hint (qualid_of_con i.is_impl))
    inst

open Libobject

let subst (_,subst,(cl,m,inst)) = 
  let do_subst_con c = fst (Mod_subst.subst_con subst c)
  and do_subst c = Mod_subst.subst_mps subst c
  and do_subst_gr gr = fst (subst_global subst gr)
  in
  let do_subst_named ctx = 
    list_smartmap (fun (na, b, t) ->
      (na, Option.smartmap do_subst b, do_subst t))
      ctx
  in
  let do_subst_ctx ctx = 
    list_smartmap (fun (cl, (na, b, t)) ->
      (Option.smartmap (fun (gr,b) -> do_subst_gr gr, b) cl,
      (na, Option.smartmap do_subst b, do_subst t)))
      ctx
  in
  let do_subst_projs projs = list_smartmap do_subst_con projs in
  let subst_class k cl classes = 
    let k = do_subst_gr k in
    let cl' = { cl with cl_impl = k;
		cl_context = do_subst_ctx cl.cl_context;
		cl_props = do_subst_named cl.cl_props;
		cl_projs = do_subst_projs cl.cl_projs; }
    in 
    let cl' = if cl' = cl then cl else cl' in
      Gmap.add k cl' classes
  in
  let classes = Gmap.fold subst_class cl Gmap.empty in
  let subst_inst k insts instances =
    let k = do_subst_gr k in
    let cl = Gmap.find k classes in
    let insts' = 
      list_smartmap (fun is -> 
	let is' = 
	  { is_class = cl; 
	    is_pri = is.is_pri;
	    is_impl = do_subst_con is.is_impl }
	in if is' = is then is else is') insts
    in  Gmap.add k insts' instances
  in
  let instances = Gmap.fold subst_inst inst Gmap.empty in
    (classes, m, instances)

let discharge (_,(cl,m,inst)) =
  let discharge_named (cl, r) =
    Option.smartmap (fun (gr, b) -> Lib.discharge_global gr, b) cl, r
  in
  let subst_class cl = 
    { cl with cl_impl = Lib.discharge_global cl.cl_impl;
      cl_context = list_smartmap discharge_named cl.cl_context;
      cl_projs = list_smartmap Lib.discharge_con cl.cl_projs }
  in
  let classes = Gmap.map subst_class cl in
  let subst_inst insts =
    List.map (fun is ->
      { is_impl = Lib.discharge_con is.is_impl;
	is_pri = is.is_pri;
	is_class = Gmap.find (Lib.discharge_global is.is_class.cl_impl) classes; })
      insts
  in
  let instances = Gmap.map subst_inst inst in
    Some (classes, m, instances)
  
let (input,output) = 
  declare_object
    { (default_object "type classes state") with
      cache_function = cache;
      load_function = (fun _ -> cache);
      open_function = (fun _ -> cache);
      classify_function = (fun (_,x) -> Substitute x);
      discharge_function = discharge;
      subst_function = subst;
      export_function = (fun x -> Some x) }

let update () = 
  Lib.add_anonymous_leaf (input (!classes, !methods, !instances))

let add_class c = 
  classes := Gmap.add c.cl_impl c !classes;
  methods := List.fold_left (fun acc x -> Gmap.add x c.cl_impl acc) !methods c.cl_projs;
  update ()
    
let class_info c = 
  try Gmap.find c !classes
  with _ -> not_a_class (Global.env()) (constr_of_global c)

let instance_constructor cl = 
  match cl.cl_impl with
    | IndRef ind -> (fun args -> applistc (mkConstruct (ind, 1)) args), mkInd ind
    | ConstRef cst -> list_last, mkConst cst
    | _ -> assert false

let typeclasses () = Gmap.fold (fun _ l c -> l :: c) !classes []

let gmapl_add x y m =
  try
    let l = Gmap.find x m in
    Gmap.add x (y::Util.list_except y l) m
  with Not_found ->
    Gmap.add x [y] m

let instances_of c = 
  try Gmap.find c.cl_impl !instances with Not_found -> []

let add_instance i =
  let cl = Gmap.find i.is_class.cl_impl !classes in
    instances := gmapl_add cl.cl_impl { i with is_class = cl } !instances;
    add_instance_hint (qualid_of_con i.is_impl) i.is_pri;
    update ()

let instances r = 
  let cl = class_info r in instances_of cl

let solve_instanciation_problem = ref (fun _ _ _ _ -> assert false)
let solve_instanciations_problem = ref (fun _ _ _ _ -> assert false)

let resolve_typeclass env ev evi (evd, defined as acc) =
  try
    assert(evi.evar_body = Evar_empty);
    !solve_instanciation_problem env evd ev evi
  with Exit -> acc

let resolve_one_typeclass env types =
  try
    let evi = Evd.make_evar (Environ.named_context_val env) types in
    let ev = 1 in
    let evm = Evd.add Evd.empty ev evi in
    let evd = Evd.create_evar_defs evm in
    let evd', b = !solve_instanciation_problem env evd ev evi in
      if b then 
	let evm' = Evd.evars_of evd' in
	  match Evd.evar_body (Evd.find evm' ev) with
	      Evar_empty -> raise Not_found
	    | Evar_defined c -> c
      else raise Not_found
  with Exit -> raise Not_found

let resolve_one_typeclass_evd env evd types =
  try
    let ev = Evarutil.e_new_evar evd env types in
    let (ev,_) = destEvar ev in
    let evd', b =
	!solve_instanciation_problem env !evd ev (Evd.find (Evd.evars_of !evd) ev)
    in
      evd := evd';
      if b then 
	let evm' = Evd.evars_of evd' in
	  match Evd.evar_body (Evd.find evm' ev) with
	      Evar_empty -> raise Not_found
	    | Evar_defined c -> Evarutil.nf_evar evm' c
      else raise Not_found
  with Exit -> raise Not_found

let method_typeclass ref = 
  match ref with 
    | ConstRef c -> 
	class_info (Gmap.find c !methods)
    | _ -> raise Not_found
      
let is_class gr = 
  Gmap.fold (fun k v acc -> acc || v.cl_impl = gr) !classes false
  
let is_implicit_arg k = 
  match k with
      ImplicitArg (ref, (n, id)) -> true
    | InternalHole -> true
    | _ -> false

let destClassApp c = 
  match kind_of_term c with
    | App (ki, args) when isInd ki -> 
	Some (destInd ki, args)
    | _ when isInd c -> Some (destInd c, [||])
    | _ -> None

let is_independent evm ev = 
  Evd.fold (fun ev' evi indep -> indep && 
    (ev = ev' || 
	evi.evar_body <> Evar_empty ||
	not (Termops.occur_evar ev evi.evar_concl)))
    evm true
    

(*   try !solve_instanciations_problem env (Evarutil.nf_evar_defs evd) *)
(*   with _ -> *)
(*   let evm = Evd.evars_of evd in *)
(*   let tc_evars =  *)
(*     let f ev evi acc = *)
(*       let (l, k) = Evd.evar_source ev evd in *)
(* 	if not check || is_implicit_arg k then *)
(* 	  match destClassApp evi.evar_concl with *)
(* 	    | Some (i, args) when is_class i ->  *)
(* 		Evd.add acc ev evi *)
(* 	    | _ -> acc *)
(* 	else acc *)
(*     in Evd.fold f evm Evd.empty *)
(*   in *)
(*   let rec sat evars = *)
(*     let evm = Evd.evars_of evars in *)
(*     let (evars', progress) =  *)
(*       Evd.fold  *)
(* 	(fun ev evi acc ->  *)
(* 	  if (Evd.mem tc_evars ev || not (Evd.mem evm ev))  *)
(* 	    && evi.evar_body = Evar_empty then *)
(* 	    resolve_typeclass env ev evi acc  *)
(* 	  else acc) *)
(* 	evm (evars, false)  *)
(*     in *)
(*       if not progress then evars' *)
(*       else *)
(* 	sat (Evarutil.nf_evar_defs evars') *)
(*   in sat (Evarutil.nf_evar_defs evd) *)
  
let class_of_constr c = 
  let extract_cl c =
    try Some (class_info (global_of_constr c)) with _ -> None
  in
    match kind_of_term c with
	App (c, _) -> extract_cl c
      | _ -> extract_cl c

let dest_class_app c =
  let cl c = class_info (global_of_constr c) in
    match kind_of_term c with
	App (c, args) -> cl c, args
      | _ -> cl c, [||]

(* To embed a boolean for resolvability status.
   This is essentially a hack to mark which evars correspond to 
   goals and do not need to be resolved when we have nested [resolve_all_evars] 
   calls (e.g. when doing apply in an External hint in typeclass_instances).
   Would be solved by having real evars-as-goals. *)

let ((bool_in : bool -> Dyn.t),
    (bool_out : Dyn.t -> bool)) = Dyn.create "bool"
  
let is_resolvable evi =
  match evi.evar_extra with
      Some t -> if Dyn.tag t = "bool" then bool_out t else true
    | None -> true
	
let mark_unresolvable evi = 
  { evi with evar_extra = Some (bool_in false) }

let has_typeclasses evd =
  Evd.fold (fun ev evi has -> has || 
    (evi.evar_body = Evar_empty && class_of_constr evi.evar_concl <> None 
	&& is_resolvable evi))
    evd false

let resolve_typeclasses ?(onlyargs=false) ?(all=true) env sigma evd =
  if not (has_typeclasses sigma) then evd
  else
    !solve_instanciations_problem env (Evarutil.nf_evar_defs evd) onlyargs all

type substitution = (identifier * constr) list

let substitution_of_named_context isevars env id n subst l = 
  List.fold_right
    (fun (na, _, t) subst -> 
      let t' = replace_vars subst t in
      let b = Evarutil.e_new_evar isevars env ~src:(dummy_loc, ImplicitArg (VarRef id, (n, Some na))) t' in
	(na, b) :: subst)
    l subst

let nf_substitution sigma subst = 
  List.map (function (na, c) -> na, Reductionops.nf_evar sigma c) subst