aboutsummaryrefslogtreecommitdiffhomepage
path: root/interp/constrintern.ml
blob: 74740abcef2973c0ea61cc08ad1798f120fc75f5 (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
471
472
473
474
475
476
477
478
479
480
481
482
483
484
485
486
487
488
489
490
491
492
493
494
495
496
497
498
499
500
501
502
503
504
505
506
507
508
509
510
511
512
513
514
515
516
517
518
519
520
521
522
523
524
525
526
527
528
529
530
531
532
533
534
535
536
537
538
539
540
541
542
543
544
545
546
547
548
549
550
551
552
553
554
555
556
557
558
559
560
561
562
563
564
565
566
567
568
569
570
571
572
573
574
575
576
577
578
579
580
581
582
583
584
585
586
587
588
589
590
591
592
593
594
595
596
597
598
599
600
601
602
603
604
605
606
607
608
609
610
611
612
613
614
615
616
617
618
619
620
621
622
623
624
625
626
627
628
629
630
631
632
633
634
635
636
637
638
639
640
641
642
643
644
645
646
647
648
649
650
651
652
653
654
655
656
657
658
659
660
661
662
663
664
665
666
667
668
669
670
671
672
673
674
675
676
677
678
679
680
681
682
683
684
685
686
687
688
689
690
691
692
693
694
695
696
(***********************************************************************)
(*  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 Options
open Names
open Nameops
open Libnames
open Impargs
open Rawterm
open Pattern
open Pretyping
open Topconstr
open Nametab
open Symbols

type implicits_env = (identifier * Impargs.implicits_list) list

let interning_grammar = ref false

let for_grammar f x =
  interning_grammar := true;
  let a = f x in
  interning_grammar := false;
  a

(**********************************************************************)
(* Internalisation errors                                             *)

type internalisation_error =
  | VariableCapture of identifier
  | WrongExplicitImplicit
  | NegativeMetavariable
  | NotAConstructor of reference
  | UnboundFixName of bool * identifier
  | NonLinearPattern of identifier
  | BadPatternsNumber of int * int
  | BadExplicitationNumber of int * int option

exception InternalisationError of loc * internalisation_error

let explain_variable_capture id =
  str "The variable " ++ pr_id id ++ str " occurs in its type"

let explain_wrong_explicit_implicit =
  str "Found an explicitely given implicit argument but was expecting" ++
  fnl () ++ str "a regular one"

let explain_negative_metavariable =
  str "Metavariable numbers must be positive"

let explain_not_a_constructor ref =
  str "Unknown constructor: " ++ pr_reference ref

let explain_unbound_fix_name is_cofix id =
  str "The name" ++ spc () ++ pr_id id ++ 
  spc () ++ str "is not bound in the corresponding" ++ spc () ++
  str (if is_cofix then "co" else "") ++ str "fixpoint definition"

let explain_non_linear_pattern id =
  str "The variable " ++ pr_id id ++ str " is bound several times in pattern"

let explain_bad_patterns_number n1 n2 =
  let s = if n1 > 1 then "s" else "" in
  str "Expecting " ++ int n1 ++ str " pattern" ++ str s ++ str " but found "
    ++ int n2

let explain_bad_explicitation_number n po =
  let s = match po with
    | None -> "a regular argument"
    | Some p -> string_of_int p in
  str "Bad explicitation number: found " ++ int n ++ 
  str" but was expecting " ++ str s

let explain_internalisation_error = function
  | VariableCapture id -> explain_variable_capture id
  | WrongExplicitImplicit -> explain_wrong_explicit_implicit
  | NegativeMetavariable -> explain_negative_metavariable
  | NotAConstructor ref -> explain_not_a_constructor ref
  | UnboundFixName (iscofix,id) -> explain_unbound_fix_name iscofix id
  | NonLinearPattern id -> explain_non_linear_pattern id
  | BadPatternsNumber (n1,n2) -> explain_bad_patterns_number n1 n2
  | BadExplicitationNumber (n,po) -> explain_bad_explicitation_number n po

(**********************************************************************)
(* Dump of globalization (to be used by coqdoc)                       *)

let add_glob loc ref = 
(*i
  let sp = Nametab.sp_of_global (Global.env ()) ref in
  let dir,_ = repr_path sp in
  let rec find_module d = 
    try 
      let qid = let dir,id = split_dirpath d in make_qualid dir id in
      let _ = Nametab.locate_loaded_library qid in d
    with Not_found -> find_module (dirpath_prefix d) 
  in
  let s = string_of_dirpath (find_module dir) in
  i*)
  let sp = Nametab.sp_of_global None ref in
  let id = let _,id = repr_path sp in string_of_id id in
  let dp = string_of_dirpath (Declare.library_part ref) in
  dump_string (Printf.sprintf "R%d %s.%s\n" (fst loc) dp id)

(**********************************************************************)
(* Remembering the parsing scope of variables in notations            *)

let set_var_scope loc id (_,_,scopes) (_,_,varscopes) =
  try 
    let idscopes = List.assoc id varscopes in
    if !idscopes <> None & !idscopes <> Some scopes then
      user_err_loc (loc,"set_var_scope",
        pr_id id ++ str " already occurs in a different scope")
    else
      idscopes := Some scopes
  with Not_found ->
    if_verbose warning ("Could not globalize " ^ (string_of_id id))

(**********************************************************************)
(* Discriminating between bound variables and global references       *)

(* [vars1] is a set of name to avoid (used for the tactic language);
   [vars2] is the set of global variables, env is the set of variables
   abstracted until this point *)

let intern_var (env,impls,scopes) ((vars1,unbndltacvars),vars2,_) loc id =
  (* Is [id] bound in current env or ltac vars bound to constr *)
  if Idset.mem id env or List.mem id vars1
  then
    RVar (loc,id), (try List.assoc id impls with Not_found -> []), []
  else
  (* Is [id] bound to a free name in ltac (this is an ltac error message) *)
  try
    match List.assoc id unbndltacvars with
      | None -> user_err_loc (loc,"intern_var",
	  pr_id id ++ str " ist not bound to a term")
      | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0
  with Not_found ->
  (* Is [id] a section variable *)
  let _ = Sign.lookup_named id vars2 in
  (* Car Fixpoint met les fns définies temporairement comme vars de sect *)
  let imps, args_scopes =
    try
      let ref = VarRef id in
      implicits_of_global ref, find_arguments_scope ref
    with _ -> [], [] in
  RRef (loc, VarRef id), imps, args_scopes

(* Is it a global reference or a syntactic definition? *)
let intern_qualid loc qid =
  try match Nametab.extended_locate qid with
  | TrueGlobal ref ->
    if !dump then add_glob loc ref;
    RRef (loc, ref), implicits_of_global ref, find_arguments_scope ref
  | SyntacticDef sp ->
      (Syntax_def.search_syntactic_definition loc sp),[],[]
  with Not_found ->
    error_global_not_found_loc loc qid

let intern_reference env lvar = function
  | Qualid (loc, qid) ->
      intern_qualid loc qid
  | Ident (loc, id) ->
      (* For old ast syntax compatibility *)
      if (string_of_id id).[0] = '$' then RVar (loc,id),[],[] else
      (* End old ast syntax compatibility *)
      try intern_var env lvar loc id
      with Not_found -> 
      try intern_qualid loc (make_short_qualid id)
      with e ->
	(* Extra allowance for grammars *)
	if !interning_grammar then begin
	  set_var_scope loc id env lvar;
	  RVar (loc,id), [], []
	end
	else raise e

let interp_reference vars r =
  let (r,_,_) = intern_reference (Idset.empty,[],[]) (vars,[],[]) r in r

let apply_scope_env (ids,impls,scopes as env) = function
  | [] -> env, []
  | (Some sc)::scl -> (ids,impls,sc::scopes), scl
  | None::scl -> env, scl

(**********************************************************************)
(* Cases                                                              *)

(* Check linearity of pattern-matching *)
let rec has_duplicate = function 
  | [] -> None
  | x::l -> if List.mem x l then (Some x) else has_duplicate l

let loc_of_lhs lhs = 
 join_loc (cases_pattern_loc (List.hd lhs)) (cases_pattern_loc (list_last lhs))

let check_linearity lhs ids =
  match has_duplicate ids with
    | Some id ->
	raise (InternalisationError (loc_of_lhs lhs,NonLinearPattern id))
    | None ->
	()

(* Warns if some pattern variable starts with uppercase *)
let check_uppercase loc ids =
(* A quoi ça sert ? Pour l'extraction vers ML ? Maintenant elle est externe
  let is_uppercase_var v =
    match (string_of_id v).[0] with 'A'..'Z' -> true | _  -> false
  in
  let warning_uppercase loc uplid =
    let vars = h 0 (prlist_with_sep pr_coma pr_id uplid) in
    let (s1,s2) = if List.length uplid = 1 then (" ","s ") else ("s "," ") in
    warn (str ("the variable"^s1) ++ vars ++
	  str (" start"^s2^"with an upper case letter in pattern")) in
  let uplid = List.filter is_uppercase_var ids in
  if uplid <> [] then warning_uppercase loc uplid
*)
  ()

(* Match the number of pattern against the number of matched args *)
let check_number_of_pattern loc n l =
  let p = List.length l in
  if n<>p then raise (InternalisationError (loc,BadPatternsNumber (n,p)))

(* Manage multiple aliases *)

  (* [merge_aliases] returns the sets of all aliases encountered at this
     point and a substitution mapping extra aliases to the first one *)
let merge_aliases (ids,subst as aliases) id =
  ids@[id], if ids=[] then subst else (id, List.hd ids)::subst

let alias_of = function
  | ([],_) -> Anonymous
  | (id::_,_) -> Name id

let message_redundant_alias (id1,id2) =
  if_verbose warning 
   ("Alias variable "^(string_of_id id1)^" is merged with "^(string_of_id id2))

(* Differentiating between constructors and matching variables *)
type pattern_qualid_kind =
  | ConstrPat of constructor
  | VarPat of identifier

let find_constructor ref =
  let (loc,qid) = qualid_of_reference ref in
  try match extended_locate qid with
    | SyntacticDef sp ->
	(match Syntax_def.search_syntactic_definition loc sp with
	  | RRef (_,(ConstructRef c as x)) ->
	      if !dump then add_glob loc x; 
 	      c
	  | _ ->
	      raise (InternalisationError (loc,NotAConstructor ref)))
    | TrueGlobal r ->
        let rec unf = function
          | ConstRef cst ->
              (try
		let v = Environ.constant_value (Global.env()) cst in
                unf (reference_of_constr v)
              with
                  Environ.NotEvaluableConst _ | Not_found ->
		    raise (InternalisationError (loc,NotAConstructor ref)))
          | ConstructRef c -> 
	      if !dump then add_glob loc r; 
	      c
          | _ -> raise (InternalisationError (loc,NotAConstructor ref))
        in unf r
  with Not_found ->
    raise (InternalisationError (loc,NotAConstructor ref))

let find_pattern_variable = function
  | Ident (loc,id) -> id
  | Qualid (loc,_) as x -> raise (InternalisationError(loc,NotAConstructor x))

let maybe_constructor ref =
  try ConstrPat (find_constructor ref)
  with InternalisationError _ -> VarPat (find_pattern_variable ref)

let rec intern_cases_pattern scopes aliases = function
  | CPatAlias (loc, p, id) ->
      let aliases' = merge_aliases aliases id in
      intern_cases_pattern scopes aliases' p
  | CPatCstr (loc, head, pl) ->
      let c = find_constructor head in
      let (idsl,pl') =
	List.split (List.map (intern_cases_pattern scopes ([],[])) pl)
      in
      (aliases::(List.flatten idsl), PatCstr (loc,c,pl',alias_of aliases))
  | CPatNumeral (loc, n) ->
      ([aliases],
      Symbols.interp_numeral_as_pattern loc n (alias_of aliases) scopes)
  | CPatDelimiters (loc, key, e) ->
      intern_cases_pattern (find_delimiters_scope loc key::scopes) aliases e
  | CPatAtom (loc, Some head) ->
      (match maybe_constructor head with
	 | ConstrPat c ->
	     ([aliases], PatCstr (loc,c,[],alias_of aliases))
	 | VarPat id ->
	     let aliases = merge_aliases aliases id in
	     ([aliases], PatVar (loc,alias_of aliases)))
  | CPatAtom (loc, None) ->
      ([aliases], PatVar (loc,alias_of aliases))

(**********************************************************************)
(* Fix and CoFix                                                      *)

let rec intern_fix = function
  | [] -> ([],[],[],[])
  | (fi, n, c, t)::rest->
      let (lf,ln,lc,lt) = intern_fix rest in
      (fi::lf, n::ln, c::lc, t::lt)

let rec intern_cofix = function
  | [] -> ([],[],[])
  | (fi, c, t)::rest -> 
      let (lf,lc,lt) = intern_cofix rest in
      (fi::lf, c::lc, t::lt)

(**********************************************************************)
(* Utilities for binders                                              *)

let check_capture loc ty = function
  | Name id when occur_var_constr_expr id ty ->
      raise (InternalisationError (loc,VariableCapture id))
  | _ ->
      ()

let locate_if_isevar loc na = function
  | RHole _ -> RHole (loc, AbstractionType na)
  | x -> x

(**********************************************************************)
(* Utilities for application                                          *)

let set_hole_implicit i = function
  | RRef (loc,r) -> (loc,ImplicitArg (r,i))
  | RVar (loc,id) -> (loc,ImplicitArg (VarRef id,i))
  | _ -> anomaly "Only refs have implicits"

(**********************************************************************)
(* Syntax extensions                                                  *)

let coerce_to_id = function
  | CRef (Ident (_,id)) -> id 
  | c ->
      user_err_loc (constr_loc c, "subst_rawconstr",
        str"This expression should be a simple identifier")

let traverse_binder subst id (ids,impls,scopes as env) =
  let id = try coerce_to_id (fst (List.assoc id subst)) with Not_found -> id in
  id,(Idset.add id ids,impls,scopes)

let rec subst_rawconstr loc interp subst (ids,impls,scopes as env) = function
  | AVar id ->
      let (a,subscopes) =
	try List.assoc id subst
	with Not_found -> (CRef (Ident (loc,id)),[]) in
      interp (ids,impls,subscopes@scopes) a
  | t ->
      map_aconstr_with_binders_loc loc (traverse_binder subst)
      (subst_rawconstr loc interp subst) env t

(**********************************************************************)
(* Main loop                                                          *)

let internalise isarity sigma env allow_soapp lvar c =
  let rec intern isarity (ids,impls,scopes as env) = function
    | CRef ref as x ->
	let (c,imp,subscopes) = intern_reference env lvar ref in
	(match intern_impargs c env imp subscopes [] with
          | [] -> c
          | l -> RApp (constr_loc x, c, l))
    | CFix (loc, (locid,iddef), ldecl) ->
	let (lf,ln,lc,lt) = intern_fix ldecl in
	let n =
	  try
	    (list_index iddef lf) -1
          with Not_found ->
	    raise (InternalisationError (locid,UnboundFixName (false,iddef)))
	in
	let ids' = List.fold_right Idset.add lf ids in
	let defl =
	  Array.of_list (List.map (intern false (ids',impls,scopes)) lt) in
	let arityl = Array.of_list (List.map (intern true env) lc) in
	RRec (loc,RFix (Array.of_list ln,n), Array.of_list lf, arityl, defl)
    | CCoFix (loc, (locid,iddef), ldecl) ->
	let (lf,lc,lt) = intern_cofix ldecl in
	let n =
          try 
	    (list_index iddef lf) -1
          with Not_found ->
	    raise (InternalisationError (locid,UnboundFixName (true,iddef)))
	in
	let ids' = List.fold_right Idset.add lf ids in
	let defl =
	  Array.of_list (List.map (intern false (ids',impls,scopes)) lt) in
	let arityl = Array.of_list (List.map (intern true env) lc) in
	RRec (loc,RCoFix n, Array.of_list lf, arityl, defl)
    | CArrow (loc,c1,c2) ->
        RProd (loc, Anonymous, intern true env c1, intern true env c2)
    | CProdN (loc,[],c2) ->
        intern true env c2
    | CProdN (loc,(nal,ty)::bll,c2) ->
        iterate_prod loc env ty (CProdN (loc, bll, c2)) nal
    | CLambdaN (loc,[],c2) ->
        intern isarity env c2
    | CLambdaN (loc,(nal,ty)::bll,c2) ->
	iterate_lam loc isarity env ty (CLambdaN (loc, bll, c2)) nal
    | CLetIn (loc,(_,na),c1,c2) ->
	RLetIn (loc, na, intern false env c1,
          intern false (name_fold Idset.add na ids,impls,scopes) c2)
    | CNotation (loc,ntn,args) ->
	let scopes = if isarity then Symbols.type_scope::scopes else scopes in
	let (ids,c) = Symbols.interp_notation ntn scopes in
	let subst = List.map2 (fun (id,scl) a -> (id,(a,scl))) ids args in
	subst_rawconstr loc (intern false) subst env c
    | CNumeral (loc, n) ->
	Symbols.interp_numeral loc n scopes
    | CDelimiters (loc, key, e) ->
	intern isarity (ids,impls,find_delimiters_scope loc key::scopes) e
    | CAppExpl (loc, ref, args) ->
        let (f,_,args_scopes) = intern_reference env lvar ref in
	RApp (loc, f, intern_args env args_scopes args)
    | CApp (loc, f, args) ->
	let (c, impargs, args_scopes) =
	  match f with
	    | CRef ref -> intern_reference env lvar ref
	    | _ -> (intern false env f, [], [])
        in
	  RApp (loc, c, intern_impargs c env impargs args_scopes args)
    | CCases (loc, po, tms, eqns) ->
	RCases (loc, option_app (intern true env) po,
		List.map (intern false env) tms,
		List.map (intern_eqn (List.length tms) env) eqns)
    | COrderedCase (loc, tag, po, c, cl) ->
	ROrderedCase (loc, tag, option_app (intern true env) po,
	          intern false env c,
		  Array.of_list (List.map (intern false env) cl))
    | CHole loc -> 
	RHole (loc, QuestionMark)
    | CMeta (loc, n) when n >=0 or allow_soapp ->
	RMeta (loc, n)
    | CMeta (loc, _) ->
	raise (InternalisationError (loc,NegativeMetavariable))
    | CSort (loc, s) ->
	RSort(loc,s)
    | CCast (loc, c1, c2) ->
	RCast (loc,intern false env c1,intern true env c2)

    | CDynamic (loc,d) -> RDynamic (loc,d)

  and intern_eqn n (ids,impls,scopes as env) (loc,lhs,rhs) =
	let (idsl_substl_list,pl) =
	  List.split (List.map (intern_cases_pattern scopes ([],[])) lhs) in
	let idsl, substl = List.split (List.flatten idsl_substl_list) in
	let eqn_ids = List.flatten idsl in
	let subst = List.flatten substl in 
	(* Linearity implies the order in ids is irrelevant *)
	check_linearity lhs eqn_ids;
	check_uppercase loc eqn_ids;
	check_number_of_pattern loc n pl;
	let rhs = replace_vars_constr_expr subst rhs in
	List.iter message_redundant_alias subst;
	let env_ids = List.fold_right Idset.add eqn_ids ids in
	(loc, eqn_ids,pl,intern false (env_ids,impls,scopes) rhs)

  and iterate_prod loc2 (ids,impls,scopes as env) ty body = function
    | (loc1,na)::nal ->
	if nal <> [] then check_capture loc1 ty na;
	let ids' = name_fold Idset.add na ids in
	let body = iterate_prod loc2 (ids',impls,scopes) ty body nal in
	let ty = locate_if_isevar loc1 na (intern true env ty) in
	RProd (join_loc loc1 loc2, na, ty, body)
    | [] -> intern true env body

  and iterate_lam loc2 isarity (ids,impls,scopes as env) ty body = function
    | (loc1,na)::nal ->
	if nal <> [] then check_capture loc1 ty na;
	let ids' = name_fold Idset.add na ids in
	let body = iterate_lam loc2 isarity (ids',impls,scopes) ty body nal in
	let ty = locate_if_isevar loc1 na (intern true env ty) in
	RLambda (join_loc loc1 loc2, na, ty, body)
    | [] -> intern isarity env body

  and intern_impargs c env l subscopes args =
    let rec aux n l subscopes args = 
      let (enva,subscopes') = apply_scope_env env subscopes in
      match (l,args) with 
      | (imp::l', (a,Some j)::args') ->
	  if is_status_implicit imp & j>=n then
	    if j=n then 
	      (intern false enva a)::(aux (n+1) l' subscopes' args')
	    else 
	      (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes' args)
	  else 
	    let e = if is_status_implicit imp then Some n else None in
	    raise
	      (InternalisationError(constr_loc a,BadExplicitationNumber (j,e)))
      | (imp::l',(a,None)::args') -> 
	  if is_status_implicit imp then 
	    (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes' args)
	  else 
	    (intern false enva a)::(aux (n+1) l' subscopes' args')
      | ([],args) -> intern_tailargs env subscopes args
      | (_::l',[]) ->
          if List.for_all is_status_implicit l then
            (RHole (set_hole_implicit n c))::(aux (n+1) l' subscopes args)
          else []
    in 
    aux 1 l subscopes args

  and intern_tailargs env subscopes = function
    | (a,Some _)::args' ->
	raise (InternalisationError (constr_loc a, WrongExplicitImplicit))
    | (a,None)::args ->
      let (enva,subscopes) = apply_scope_env env subscopes in
      (intern false enva a) :: (intern_tailargs env subscopes args)
    | [] -> []

  and intern_args env subscopes = function
    | [] -> []
    | a::args ->
        let (enva,subscopes) = apply_scope_env env subscopes in
        (intern false enva a) :: (intern_args env subscopes args)

  in 
  try 
    intern isarity env c
  with
      InternalisationError (loc,e) ->
	user_err_loc (loc,"internalize",explain_internalisation_error e)

(**************************************************************************)
(* Functions to translate constr_expr into rawconstr                       *)
(**************************************************************************)

let extract_ids env =
  List.fold_right Idset.add 
    (Termops.ids_of_rel_context (Environ.rel_context env))
    Idset.empty

let interp_rawconstr_gen isarity sigma env impls allow_soapp ltacvar c =
  internalise isarity sigma (extract_ids env, impls, [])
    allow_soapp (ltacvar,Environ.named_context env, []) c

let interp_rawconstr sigma env c =
  interp_rawconstr_gen false sigma env [] false ([],[]) c

let interp_rawtype sigma env c =
  interp_rawconstr_gen true sigma env [] false ([],[]) c

let interp_rawtype_with_implicits sigma env impls c =
  interp_rawconstr_gen true sigma env impls false ([],[]) c

(*
(* The same as interp_rawconstr but with a list of variables which must not be
   globalized *)

let interp_rawconstr_wo_glob sigma env lvar c =
  interp_rawconstr_gen sigma env [] false lvar c
*)

(*********************************************************************)
(* Functions to parse and interpret constructions *)

let interp_constr sigma env c =
  understand sigma env (interp_rawconstr sigma env c)

let interp_openconstr sigma env c =
  understand_gen_tcc sigma env [] [] None (interp_rawconstr sigma env c)

let interp_casted_openconstr sigma env c typ =
  understand_gen_tcc sigma env [] [] (Some typ) (interp_rawconstr sigma env c)

let interp_type sigma env c =
  understand_type sigma env (interp_rawtype sigma env c)

let interp_type_with_implicits sigma env impls c =
  understand_type sigma env (interp_rawtype_with_implicits sigma env impls c)

let judgment_of_rawconstr sigma env c =
  understand_judgment sigma env (interp_rawconstr sigma env c)

let type_judgment_of_rawconstr sigma env c =
  understand_type_judgment sigma env (interp_rawconstr sigma env c)

(* To retype a list of key*constr with undefined key *)
let retype_list sigma env lst =
  List.fold_right (fun (x,csr) a ->
    try (x,Retyping.get_judgment_of env sigma csr)::a with
    | Anomaly _ -> a) lst []

(*  List.map (fun (x,csr) -> (x,Retyping.get_judgment_of env sigma csr)) lst*)

type ltac_sign = 
  identifier list * (identifier * identifier option) list

type ltac_env = 
  (identifier * Term.constr) list * (identifier * identifier option) list

(* Interprets a constr according to two lists *)
(* of instantiations (variables and metas)    *)
(* Note: typ is retyped *)
let interp_constr_gen sigma env (ltacvar,unbndltacvar) lmeta c exptyp =
  let c = interp_rawconstr_gen false sigma env [] false 
    (List.map fst ltacvar,unbndltacvar) c
  and rtype lst = retype_list sigma env lst in
  understand_gen sigma env (rtype ltacvar) (rtype lmeta) exptyp c;;

(*Interprets a casted constr according to two lists of instantiations
  (variables and metas)*)
let interp_openconstr_gen sigma env (ltacvar,unbndltacvar) lmeta c exptyp =
  let c = interp_rawconstr_gen false sigma env [] false
    (List.map fst ltacvar,unbndltacvar) c
  and rtype lst = retype_list sigma env lst in
  understand_gen_tcc sigma env (rtype ltacvar) (rtype lmeta) exptyp c;;

let interp_casted_constr sigma env c typ = 
  understand_gen sigma env [] [] (Some typ) (interp_rawconstr sigma env c)

(* To process patterns, we need a translation without typing at all. *)

let rec pat_of_raw metas vars lvar = function
  | RVar (_,id) ->
      (try PRel (list_index (Name id) vars)
       with Not_found ->
       try List.assoc id lvar
       with Not_found -> PVar id)
  | RMeta (_,n) ->
      metas := n::!metas; PMeta (Some n)
  | RRef (_,r) ->
      PRef r
  (* Hack pour ne pas réécrire une interprétation complète des patterns*)
  | RApp (_, RMeta (_,n), cl) when n<0 ->
      PSoApp (- n, List.map (pat_of_raw metas vars lvar) cl)
  | RApp (_,c,cl) -> 
      PApp (pat_of_raw metas vars lvar c,
	    Array.of_list (List.map (pat_of_raw metas vars lvar) cl))
  | RLambda (_,na,c1,c2) ->
      PLambda (na, pat_of_raw metas vars lvar c1,
	       pat_of_raw metas (na::vars) lvar c2)
  | RProd (_,na,c1,c2) ->
      PProd (na, pat_of_raw metas vars lvar c1,
	       pat_of_raw metas (na::vars) lvar c2)
  | RLetIn (_,na,c1,c2) ->
      PLetIn (na, pat_of_raw metas vars lvar c1,
	       pat_of_raw metas (na::vars) lvar c2)
  | RSort (_,s) ->
      PSort s
  | RHole _ ->
      PMeta None
  | RCast (_,c,t) ->
      if_verbose warning "Cast not taken into account in constr pattern";
      pat_of_raw metas vars lvar c
  | ROrderedCase (_,st,po,c,br) ->
      PCase (st,option_app (pat_of_raw metas vars lvar) po,
             pat_of_raw metas vars lvar c,
             Array.map (pat_of_raw metas vars lvar) br)
  | r ->
      let loc = loc_of_rawconstr r in
      user_err_loc (loc,"pattern_of_rawconstr", str "Not supported pattern")

let pattern_of_rawconstr lvar c =
  let metas = ref [] in
  let p = pat_of_raw metas [] lvar c in
  (!metas,p)

let interp_constrpattern_gen sigma env (ltacvar,unbndltacvar) c =
  let c = interp_rawconstr_gen false sigma env [] true
    (List.map fst ltacvar,unbndltacvar) c in
  let nlvar = List.map (fun (id,c) -> (id,pattern_of_constr c)) ltacvar in
  pattern_of_rawconstr nlvar c

let interp_constrpattern sigma env c =
  interp_constrpattern_gen sigma env ([],[]) c

let interp_aconstr vars a =
  let env = Global.env () in
  (* [vl] is intended to remember the scope of the free variables of [a] *)
  let vl = List.map (fun id -> (id,ref None)) vars in
  let c = for_grammar (internalise false Evd.empty (extract_ids env, [], [])
    false (([],[]),Environ.named_context env,vl)) a in
  (* Translate and check that [c] has all its free variables bound in [vars] *)
  let a = aconstr_of_rawconstr vars c in
  (* Returns [a] and the ordered list of variables with their scopes *)
  (* Variables occuring in binders have no relevant scope since bound *)
  List.map (fun (id,r) -> (id,match !r with None -> [] | Some l -> l)) vl, a