aboutsummaryrefslogtreecommitdiffhomepage
path: root/interp/constrextern.ml
blob: 8cbddd87c168a40b1d22690e4797d2f22538ff08 (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
391
392
393
394
395
396
397
398
399
400
401
402
403
404
405
406
407
408
409
410
411
412
413
414
415
416
417
418
419
420
421
422
423
424
425
426
427
428
429
430
431
432
433
434
435
436
437
438
439
440
441
442
443
444
445
446
447
448
449
450
451
452
453
454
455
456
457
458
459
460
461
462
463
464
465
466
467
468
469
470
(***********************************************************************)
(*  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$ *)

(*i*)
open Pp
open Util
open Univ
open Names
open Nameops
open Term
open Termops
open Inductive
open Sign
open Environ
open Libnames
open Declare
open Impargs
open Topconstr
open Rawterm
open Pattern
open Nametab
open Symbols
(*i*)

(* Translation from rawconstr to front constr *)

(**********************************************************************)
(* Parametrization                                                    *)

(* This governs printing of local context of references *)
let print_arguments = ref false

(* If true, prints local context of evars, whatever print_arguments *)
let print_evar_arguments = ref false

(* This governs printing of implicit arguments.  When
   [print_implicits] is on then [print_implicits_explicit_args] tells
   how implicit args are printed. If on, implicit args are printed
   prefixed by "!" otherwise the function and not the arguments is
   prefixed by "!" *)
let print_implicits = ref false
let print_implicits_explicit_args = ref false

(* This forces printing of coercions *)
let print_coercions = ref false

(* This forces printing universe names of Type{.} *)
let print_universes = ref false

(* This suppresses printing of numeral and symbols *)
let print_no_symbol = ref false

let print_meta_as_hole = ref false

let with_option o f x =
  let old = !o in o:=true;
  try let r = f x in o := old; r
  with e -> o := old; raise e

let with_arguments f = with_option print_arguments f
let with_implicits f = with_option print_implicits f
let with_coercions f = with_option print_coercions f
let with_universes f = with_option print_universes f
let without_symbols f = with_option print_no_symbol f
let with_meta_as_hole f = with_option print_meta_as_hole f

(**********************************************************************)
(* Various externalisation functions *)

let insert_delimiters e = function
  | None -> e
  | Some sc -> CDelimiters (dummy_loc,sc,e)

let insert_pat_delimiters e = function
  | None -> e
  | Some sc -> CPatDelimiters (dummy_loc,sc,e)

let loc = dummy_loc (* shorter... *)

(**********************************************************************)
(* conversion of references                                           *)

let ids_of_ctxt ctxt =
  Array.to_list
    (Array.map
       (function c -> match kind_of_term c with
	  | Var id -> id
	  | _ ->
       error
       "Termast: arbitrary substitution of references not yet implemented")
     ctxt)

let idopt_of_name = function 
  | Name id -> Some id
  | Anonymous -> None

let extern_evar loc n =
  warning
    "Existential variable turned into meta-variable during externalization";
  CMeta (loc,n)

let raw_string_of_ref = function
  | ConstRef kn -> 
      "CONST("^(string_of_kn kn)^")"
  | IndRef (kn,i) ->
      "IND("^(string_of_kn kn)^","^(string_of_int i)^")"
  | ConstructRef ((kn,i),j) -> 
      "CONSTRUCT("^
      (string_of_kn kn)^","^(string_of_int i)^","^(string_of_int j)^")"
  | VarRef id -> 
      "SECVAR("^(string_of_id id)^")"

let extern_reference loc vars r =
  try Qualid (loc,shortest_qualid_of_global vars r)
  with Not_found ->
    (* happens in debugger *)
    Ident (loc,id_of_string (raw_string_of_ref r))

(**********************************************************************)
(* conversion of terms and cases patterns                             *)

let rec extern_cases_pattern_in_scope scopes vars pat =
  try 
    if !print_no_symbol then raise No_match;
    let (sc,n) = Symbols.uninterp_cases_numeral pat in
    match Symbols.availability_of_numeral sc scopes with
    | None -> raise No_match
    | Some key -> insert_pat_delimiters (CPatNumeral (loc,n)) key
  with No_match ->
  match pat with
  | PatVar (_,Name id) -> CPatAtom (loc,Some (Ident (loc,id)))
  | PatVar (_,Anonymous) -> CPatAtom (loc, None) 
  | PatCstr(_,cstrsp,args,na) ->
      let args = List.map (extern_cases_pattern_in_scope scopes vars) args in
      let p = CPatCstr
        (loc,extern_reference loc vars (ConstructRef cstrsp),args) in
      (match na with
	| Name id -> CPatAlias (loc,p,id)
	| Anonymous -> p)
	
let occur_name na aty =
  match na with
    | Name id -> occur_var_constr_expr id aty
    | Anonymous -> false

(* Implicit args indexes are in ascending order *)
let explicitize impl f args =
  let n = List.length args in
  let rec exprec q = function
    | a::args, imp::impl when is_status_implicit imp ->
        let tail = exprec (q+1) (args,impl) in
        let visible =
          (!print_implicits & !print_implicits_explicit_args)
          or not (is_inferable_implicit n imp) in
        if visible then (a,Some q) :: tail else tail
    | a::args, _::impl -> (a,None) :: exprec (q+1) (args,impl)
    | args, [] -> List.map (fun a -> (a,None)) args (*In case of polymorphism*)
    | [], _ -> [] in
  let args = exprec 1 (args,impl) in
  if args = [] then f else CApp (dummy_loc, f, args)

let rec skip_coercion dest_ref (f,args as app) =
  if !print_coercions then app
  else
    try
      match dest_ref f with
	| Some r ->
	    (match Classops.hide_coercion r with
	       | Some n ->
		   if n >= List.length args then app
		   else (* We skip a coercion *)
		     let _,fargs = list_chop n args in
	       	     skip_coercion dest_ref (List.hd fargs,List.tl fargs)
	       | None -> app)
	| None -> app
    with Not_found -> app

let extern_app impl f args =
  if (!print_implicits & not !print_implicits_explicit_args) then 
    CAppExpl (dummy_loc, f, args)
  else
    explicitize impl (CRef f) args

let rec extern_args extern scopes env args subscopes =
  match args with
    | [] -> []
    | a::args ->
	let argscopes, subscopes = match subscopes with
	  | [] -> scopes, []
	  | scopt::subscopes -> option_cons scopt scopes, subscopes in
	extern argscopes env a :: extern_args extern scopes env args subscopes

let rec extern scopes vars r =
  try 
    if !print_no_symbol then raise No_match;
    extern_numeral scopes r (Symbols.uninterp_numeral r)
  with No_match ->

  try 
    if !print_no_symbol then raise No_match;
    extern_symbol scopes vars r (Symbols.uninterp_notations r)
  with No_match -> match r with

  | RRef (_,ref) -> CRef (extern_reference loc vars ref)

  | RVar (_,id) -> CRef (Ident (loc,id))

  | REvar (_,n) -> extern_evar loc n

  | RMeta (_,n) -> if !print_meta_as_hole then CHole loc else CMeta (loc,n)

  | RApp (_,f,args) ->
      let (f,args) =
	skip_coercion (function RRef(_,r) -> Some r | _ -> None) (f,args) in
      (match f with
	 | REvar (loc,ev) -> extern_evar loc ev (* we drop args *)
	 | RRef (loc,ref) ->
	     let subscopes = Symbols.find_arguments_scope ref in
	     let args = extern_args extern scopes vars args subscopes in
	     extern_app (implicits_of_global ref)
               (extern_reference loc vars ref)
	       args
	 | _       -> 
	     explicitize [] (extern scopes vars f)
               (List.map (extern scopes vars) args))

  | RProd (_,Anonymous,t,c) ->
      (* Anonymous product are never factorized *)
      CArrow (loc,extern scopes vars t, extern scopes vars c)

  | RLetIn (_,na,t,c) ->
      CLetIn (loc,(loc,na),extern scopes vars t,
              extern scopes (add_vname vars na) c)

  | RProd (_,na,t,c) ->
      let t = extern scopes vars t in
      let (idl,c) = factorize_prod scopes (add_vname vars na) t c in
      CProdN (loc,[(loc,na)::idl,t],c)

  | RLambda (_,na,t,c) ->
      let t = extern scopes vars t in
      let (idl,c) = factorize_lambda scopes (add_vname vars na) t c in
      CLambdaN (loc,[(loc,na)::idl,t],c)

  | RCases (_,typopt,tml,eqns) ->
      let pred = option_app (extern scopes vars) typopt in
      let tml = List.map (extern scopes vars) tml in
      let eqns = List.map (extern_eqn scopes vars) eqns in 
      CCases (loc,pred,tml,eqns)
	
  | ROrderedCase (_,cs,typopt,tm,bv) ->
      let bv = Array.to_list (Array.map (extern scopes vars) bv) in
      COrderedCase
	(loc,cs,option_app (extern scopes vars) typopt,
         extern scopes vars tm,bv)

  | RRec (_,fk,idv,tyv,bv) ->
      let vars' = Array.fold_right Idset.add idv vars in
      (match fk with
	 | RFix (nv,n) ->
	     let listdecl = 
	       Array.mapi (fun i fi ->
		 (fi,nv.(i),extern scopes vars tyv.(i),
                  extern scopes vars' bv.(i))) idv
	     in 
	     CFix (loc,(loc,idv.(n)),Array.to_list listdecl)
	 | RCoFix n -> 
	     let listdecl =
               Array.mapi (fun i fi ->
		 (fi,extern scopes vars tyv.(i),
                  extern scopes vars' bv.(i))) idv
	     in
	     CCoFix (loc,(loc,idv.(n)),Array.to_list listdecl))

  | RSort (_,s) ->
      let s = match s with
	 | RProp _ -> s
	 | RType (Some _) when !print_universes -> s
	 | RType _ -> RType None in
      CSort (loc,s)

  | RHole (_,e) -> CHole loc

  | RCast (_,c,t) -> CCast (loc,extern scopes vars c,extern scopes vars t)

  | RDynamic (_,d) -> CDynamic (loc,d)
	
and factorize_prod scopes vars aty = function
  | RProd (_,Name id,ty,c)
      when aty = extern scopes vars ty
	& not (occur_var_constr_expr id aty) (* avoid na in ty escapes scope *)
	-> let (nal,c) = factorize_prod scopes (Idset.add id vars) aty c in
           ((loc,Name id)::nal,c)
  | c -> ([],extern scopes vars c)

and factorize_lambda scopes vars aty = function
  | RLambda (_,na,ty,c)
      when aty = extern scopes vars ty
	& not (occur_name na aty) (* To avoid na in ty' escapes scope *)
	-> let (nal,c) = factorize_lambda scopes (add_vname vars na) aty c in
           ((loc,na)::nal,c)
  | c -> ([],extern scopes vars c)

and extern_eqn scopes vars (loc,ids,pl,c) =
  (loc,List.map (extern_cases_pattern_in_scope scopes vars) pl,
   extern scopes vars c)

and extern_numeral scopes t (sc,n) =
  match Symbols.availability_of_numeral sc scopes with
    | None -> raise No_match
    | Some key -> insert_delimiters (CNumeral (loc,n)) key

and extern_symbol scopes vars t = function
  | [] -> raise No_match
  | (keyrule,pat,n as rule)::rules ->
      try
	(* Adjusts to the number of arguments expected by the notation *)
	let (t,args) = match t,n with
	  | RApp (_,f,args), Some n when List.length args > n ->
	      let args1, args2 = list_chop n args in
	      RApp (loc,f,args1), args2
	  | _ -> t,[] in
	(* Try matching ... *)
	let subst = match_aconstr t pat in
	(* Try availability of interpretation ... *)
        let e =
          match keyrule with
          | NotationRule (sc,ntn) ->
	      (match Symbols.availability_of_notation (sc,ntn) scopes with
                  (* Uninterpretation is not allowed in current context *)
              | None -> raise No_match
                  (* Uninterpretation is allowed in current context *)
	      | Some (scopt,key) ->
	          let scopes = option_cons scopt scopes in
	          let l =
		    List.map (fun (c,scl) -> extern (scl@scopes) vars c)
                      subst in
	          insert_delimiters (CNotation (loc,ntn,l)) key)
          | SynDefRule kn ->
              CRef (Qualid (loc, shortest_qualid_of_syndef kn)) in
 	if args = [] then e 
	else explicitize [] e (List.map (extern scopes vars) args)
      with
	  No_match -> extern_symbol scopes vars t rules

let extern_rawconstr vars c = 
  extern (Symbols.current_scopes()) vars c

let extern_cases_pattern vars p = 
  extern_cases_pattern_in_scope (Symbols.current_scopes()) vars p

(******************************************************************)
(* Main translation function from constr -> constr_expr *)
       
let extern_constr at_top env t =
  let vars = vars_of_env env in
  let avoid = if at_top then ids_of_context env else [] in
  extern (Symbols.current_scopes()) vars
    (Detyping.detype env avoid (names_of_rel_context env) t)

(******************************************************************)
(* Main translation function from pattern -> constr_expr *)

let rec extern_pattern tenv vars env = function
  | PRef ref -> CRef (extern_reference loc vars ref)

  | PVar id -> CRef (Ident (loc,id))

  | PEvar n -> extern_evar loc n

  | PRel n ->
      CRef (Ident (loc,
        try match lookup_name_of_rel n env with
	 | Name id   -> id
	 | Anonymous ->
	     anomaly "ast_of_pattern: index to an anonymous variable"
       with Not_found ->
	 id_of_string ("[REL "^(string_of_int n)^"]")))

  | PMeta None -> CHole loc

  | PMeta (Some n) -> CMeta (loc,n)

  | PApp (f,args) ->
      let (f,args) = 
	skip_coercion (function PRef r -> Some r | _ -> None)
	  (f,Array.to_list args) in
      let args = List.map (extern_pattern tenv vars env) args in
      (match f with
	 | PRef ref ->
	     extern_app (implicits_of_global ref)
               (extern_reference loc vars ref)
	       args
	 | _       -> explicitize [] (extern_pattern tenv vars env f) args)

  | PSoApp (n,args) ->
      let args = List.map (extern_pattern tenv vars env) args in
      (* [-n] is the trick to embed a so patten into a regular application *)
      (* see constrintern.ml and g_constr.ml4 *)
      explicitize [] (CMeta (loc,-n)) args

  | PProd (Anonymous,t,c) ->
      (* Anonymous product are never factorized *)
      CArrow (loc,extern_pattern tenv vars env t,
              extern_pattern tenv vars env c)

  | PLetIn (na,t,c) ->
      CLetIn (loc,(loc,na),extern_pattern tenv vars env t,
              extern_pattern tenv (add_vname vars na) (na::env) c)

  | PProd (na,t,c) ->
      let t = extern_pattern tenv vars env t in
      let (idl,c) =
        factorize_prod_pattern tenv (add_vname vars na) (na::env) t c in
      CProdN (loc,[(loc,na)::idl,t],c)

  | PLambda (na,t,c) ->
      let t = extern_pattern tenv vars env t in
      let (idl,c) =
        factorize_lambda_pattern tenv (add_vname vars na) (na::env) t c in
      CLambdaN (loc,[(loc,na)::idl,t],c)

  | PCase (cs,typopt,tm,bv) ->
      let bv = Array.to_list (Array.map (extern_pattern tenv vars env) bv) in
      COrderedCase
	(loc,cs,option_app (extern_pattern tenv vars env) typopt,
	 extern_pattern tenv vars env tm,bv)

  | PFix f ->
      extern (Symbols.current_scopes()) vars
        (Detyping.detype tenv [] env (mkFix f))

  | PCoFix c ->
      extern (Symbols.current_scopes()) vars
        (Detyping.detype tenv [] env (mkCoFix c))

  | PSort s ->
      let s = match s with
	 | RProp _ -> s
	 | RType (Some _) when !print_universes -> s
	 | RType _ -> RType None in
      CSort (loc,s)

and factorize_prod_pattern tenv vars env aty = function
  | PProd (Name id as na,ty,c)
      when aty = extern_pattern tenv vars env ty
      & not (occur_var_constr_expr id aty) (*To avoid na in ty escapes scope*)
	-> let (nal,c) =
          factorize_prod_pattern tenv (add_vname vars na) (na::env) aty c in
	   ((loc,Name id)::nal,c)
  | c -> ([],extern_pattern tenv vars env c)

and factorize_lambda_pattern tenv vars env aty = function
  | PLambda (na,ty,c)
      when aty = extern_pattern tenv vars env ty
	& not (occur_name na aty) (* To avoid na in ty' escapes scope *)
	-> let (nal,c) =
          factorize_lambda_pattern tenv (add_vname vars na) (na::env) aty c
        in ((loc,na)::nal,c)
  | c -> ([],extern_pattern tenv vars env c)

let extern_pattern tenv env pat =
  extern_pattern tenv (vars_of_env tenv) env pat