aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins
diff options
context:
space:
mode:
authorGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-01-17 14:23:53 +0100
committerGravatar Emilio Jesus Gallego Arias <e+git@x80.org>2017-04-24 23:58:23 +0200
commit158f40db9482ead89befbf9bc9ad45ff8a60b75f (patch)
tree92587db07ddf50e2db16b270966115fa3d66d64a /plugins
parentbe83b52cf50ed4c596e40cfd52da03258a7a4a18 (diff)
[location] Switch glob_constr to Loc.located
Diffstat (limited to 'plugins')
-rw-r--r--plugins/decl_mode/g_decl_mode.ml4387
-rw-r--r--plugins/funind/glob_term_to_relation.ml102
-rw-r--r--plugins/funind/glob_termops.ml328
-rw-r--r--plugins/funind/glob_termops.mli7
-rw-r--r--plugins/funind/indfun.ml20
-rw-r--r--plugins/funind/indfun_common.ml10
-rw-r--r--plugins/funind/merge.ml52
-rw-r--r--plugins/funind/recdef.ml18
-rw-r--r--plugins/ltac/extratactics.ml412
-rw-r--r--plugins/ltac/g_rewrite.ml42
-rw-r--r--plugins/ltac/pptactic.ml4
-rw-r--r--plugins/ltac/tacintern.ml22
-rw-r--r--plugins/ltac/tacinterp.ml4
-rw-r--r--plugins/setoid_ring/newring.ml2
-rw-r--r--plugins/ssrmatching/ssrmatching.ml439
-rw-r--r--plugins/syntax/ascii_syntax.ml14
-rw-r--r--plugins/syntax/nat_syntax.ml17
-rw-r--r--plugins/syntax/numbers_syntax.ml84
-rw-r--r--plugins/syntax/r_syntax.ml32
-rw-r--r--plugins/syntax/string_syntax.ml16
-rw-r--r--plugins/syntax/z_syntax.ml56
21 files changed, 410 insertions, 818 deletions
diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4
deleted file mode 100644
index 2415080f3..000000000
--- a/plugins/decl_mode/g_decl_mode.ml4
+++ /dev/null
@@ -1,387 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i camlp4deps: "grammar/grammar.cma" i*)
-
-DECLARE PLUGIN "decl_mode_plugin"
-
-open Ltac_plugin
-open Compat
-open Pp
-open Decl_expr
-open Names
-open Pcoq
-open Vernacexpr
-open Tok (* necessary for camlp4 *)
-
-open Pcoq.Constr
-open Pltac
-open Ppdecl_proof
-
-let pr_goal gs =
- let (g,sigma) = Goal.V82.nf_evar (Tacmach.project gs) (Evd.sig_it gs) in
- let env = Goal.V82.env sigma g in
- let concl = Goal.V82.concl sigma g in
- let goal =
- Printer.pr_context_of env sigma ++ cut () ++
- str "============================" ++ cut () ++
- str "thesis :=" ++ cut () ++
- Printer.pr_goal_concl_style_env env sigma concl in
- str " *** Declarative Mode ***" ++ fnl () ++ fnl () ++
- str " " ++ v 0 goal
-
-let pr_subgoals ?(pr_first=true) _ sigma _ _ _ gll =
- match gll with
- | [goal] when pr_first ->
- pr_goal { Evd.it = goal ; sigma = sigma }
- | _ ->
- (* spiwack: it's not very nice to have to call proof global
- here, a more robust solution would be to add a hook for
- [Printer.pr_open_subgoals] in proof modes, in order to
- compute the end command. Yet a more robust solution would be
- to have focuses give explanations of their unfocusing
- behaviour. *)
- let p = Proof_global.give_me_the_proof () in
- let close_cmd = Decl_mode.get_end_command p in
- str "Subproof completed, now type " ++ str close_cmd ++ str "."
-
-let interp_proof_instr _ { Evd.it = gl ; sigma = sigma }=
- Decl_interp.interp_proof_instr
- (Decl_mode.get_info sigma gl)
- (Goal.V82.env sigma gl)
- (sigma)
-
-let vernac_decl_proof () =
- let pf = Proof_global.give_me_the_proof () in
- if Proof.is_done pf then
- CErrors.error "Nothing left to prove here."
- else
- begin
- Decl_proof_instr.go_to_proof_mode () ;
- Proof_global.set_proof_mode "Declarative"
- end
-
-(* spiwack: some bureaucracy is not performed here *)
-let vernac_return () =
- begin
- Decl_proof_instr.return_from_tactic_mode () ;
- Proof_global.set_proof_mode "Declarative"
- end
-
-let vernac_proof_instr instr =
- Decl_proof_instr.proof_instr instr
-
-(* Before we can write an new toplevel command (see below)
- which takes a [proof_instr] as argument, we need to declare
- how to parse it, print it, globalise it and interprete it.
- Normally we could do that easily through ARGUMENT EXTEND,
- but as the parsing is fairly complicated we will do it manually to
- indirect through the [proof_instr] grammar entry. *)
-(* spiwack: proposal: doing that directly from argextend.ml4, maybe ? *)
-
-(* Only declared at raw level, because only used in vernac commands. *)
-let wit_proof_instr : (raw_proof_instr, glob_proof_instr, proof_instr) Genarg.genarg_type =
- Genarg.make0 "proof_instr"
-
-(* We create a new parser entry [proof_mode]. The Declarative proof mode
- will replace the normal parser entry for tactics with this one. *)
-let proof_mode : vernac_expr Gram.entry =
- Gram.entry_create "vernac:proof_command"
-(* Auxiliary grammar entry. *)
-let proof_instr : raw_proof_instr Gram.entry =
- Pcoq.create_generic_entry Pcoq.utactic "proof_instr" (Genarg.rawwit wit_proof_instr)
-
-let _ = Pptactic.declare_extra_genarg_pprule wit_proof_instr
- pr_raw_proof_instr pr_glob_proof_instr pr_proof_instr
-
-let classify_proof_instr = function
- | { instr = Pescape |Pend B_proof } -> VtProofMode "Classic", VtNow
- | _ -> Vernac_classifier.classify_as_proofstep
-
-(* We use the VERNAC EXTEND facility with a custom non-terminal
- to populate [proof_mode] with a new toplevel interpreter.
- The "-" indicates that the rule does not start with a distinguished
- string. *)
-VERNAC proof_mode EXTEND ProofInstr
- [ - proof_instr(instr) ] => [classify_proof_instr instr] -> [ vernac_proof_instr instr ]
-END
-
-(* It is useful to use GEXTEND directly to call grammar entries that have been
- defined previously VERNAC EXTEND. In this case we allow, in proof mode,
- the use of commands like Check or Print. VERNAC EXTEND does quite a bit of
- bureaucracy for us, but it is not needed in this sort of case, and it would require
- to have an ARGUMENT EXTEND version of the "proof_mode" grammar entry. *)
-GEXTEND Gram
- GLOBAL: proof_mode ;
-
- proof_mode: LAST
- [ [ c=G_vernac.subgoal_command -> c (Some (Vernacexpr.SelectNth 1)) ] ]
- ;
-END
-
-(* We register a new proof mode here *)
-
-let _ =
- Proof_global.register_proof_mode { Proof_global.
- name = "Declarative" ; (* name for identifying and printing *)
- (* function [set] goes from No Proof Mode to
- Declarative Proof Mode performing side effects *)
- set = begin fun () ->
- (* We set the command non terminal to
- [proof_mode] (which we just defined). *)
- Pcoq.set_command_entry proof_mode ;
- (* We substitute the goal printer, by the one we built
- for the proof mode. *)
- Printer.set_printer_pr { Printer.default_printer_pr with
- Printer.pr_goal = pr_goal;
- pr_subgoals = pr_subgoals; }
- end ;
- (* function [reset] goes back to No Proof Mode from
- Declarative Proof Mode *)
- reset = begin fun () ->
- (* We restore the command non terminal to
- [noedit_mode]. *)
- Pcoq.set_command_entry Pcoq.Vernac_.noedit_mode ;
- (* We restore the goal printer to default *)
- Printer.set_printer_pr Printer.default_printer_pr
- end
- }
-
-VERNAC COMMAND EXTEND DeclProof
-[ "proof" ] => [ VtProofMode "Declarative", VtNow ] -> [ vernac_decl_proof () ]
-END
-VERNAC COMMAND EXTEND DeclReturn
-[ "return" ] => [ VtProofMode "Declarative", VtNow ] -> [ vernac_return () ]
-END
-
-let none_is_empty = function
- None -> []
- | Some l -> l
-
-GEXTEND Gram
-GLOBAL: proof_instr;
- thesis :
- [[ "thesis" -> Plain
- | "thesis"; "for"; i=ident -> (For i)
- ]];
- statement :
- [[ i=ident ; ":" ; c=constr -> {st_label=Name i;st_it=c}
- | i=ident -> {st_label=Anonymous;
- st_it= Loc.tag ~loc:!@loc @@ Constrexpr.CRef (Libnames.Ident (!@loc, i), None)}
- | c=constr -> {st_label=Anonymous;st_it=c}
- ]];
- constr_or_thesis :
- [[ t=thesis -> Thesis t ] |
- [ c=constr -> This c
- ]];
- statement_or_thesis :
- [
- [ t=thesis -> {st_label=Anonymous;st_it=Thesis t} ]
- |
- [ i=ident ; ":" ; cot=constr_or_thesis -> {st_label=Name i;st_it=cot}
- | i=ident -> {st_label=Anonymous;
- st_it=This (Loc.tag ~loc:!@loc @@ Constrexpr.CRef (Libnames.Ident (!@loc, i), None))}
- | c=constr -> {st_label=Anonymous;st_it=This c}
- ]
- ];
- justification_items :
- [[ -> Some []
- | "by"; l=LIST1 constr SEP "," -> Some l
- | "by"; "*" -> None ]]
- ;
- justification_method :
- [[ -> None
- | "using"; tac = tactic -> Some tac ]]
- ;
- simple_cut_or_thesis :
- [[ ls = statement_or_thesis;
- j = justification_items;
- taco = justification_method
- -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- simple_cut :
- [[ ls = statement;
- j = justification_items;
- taco = justification_method
- -> {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- elim_type:
- [[ IDENT "induction" -> ET_Induction
- | IDENT "cases" -> ET_Case_analysis ]]
- ;
- block_type :
- [[ IDENT "claim" -> B_claim
- | IDENT "focus" -> B_focus
- | IDENT "proof" -> B_proof
- | et=elim_type -> B_elim et ]]
- ;
- elim_obj:
- [[ IDENT "on"; c=constr -> Real c
- | IDENT "of"; c=simple_cut -> Virtual c ]]
- ;
- elim_step:
- [[ IDENT "consider" ;
- h=consider_vars ; IDENT "from" ; c=constr -> Pconsider (c,h)
- | IDENT "per"; et=elim_type; obj=elim_obj -> Pper (et,obj)
- | IDENT "suffices"; ls=suff_clause;
- j = justification_items;
- taco = justification_method
- -> Psuffices {cut_stat=ls;cut_by=j;cut_using=taco} ]]
- ;
- rew_step :
- [[ "~=" ; c=simple_cut -> (Rhs,c)
- | "=~" ; c=simple_cut -> (Lhs,c)]]
- ;
- cut_step:
- [[ "then"; tt=elim_step -> Pthen tt
- | "then"; c=simple_cut_or_thesis -> Pthen (Pcut c)
- | IDENT "thus"; tt=rew_step -> Pthus (let s,c=tt in Prew (s,c))
- | IDENT "thus"; c=simple_cut_or_thesis -> Pthus (Pcut c)
- | IDENT "hence"; c=simple_cut_or_thesis -> Phence (Pcut c)
- | tt=elim_step -> tt
- | tt=rew_step -> let s,c=tt in Prew (s,c);
- | IDENT "have"; c=simple_cut_or_thesis -> Pcut c;
- | IDENT "claim"; c=statement -> Pclaim c;
- | IDENT "focus"; IDENT "on"; c=statement -> Pfocus c;
- | "end"; bt = block_type -> Pend bt;
- | IDENT "escape" -> Pescape ]]
- ;
- (* examiner s'il est possible de faire R _ et _ R pour R une relation qcq*)
- loc_id:
- [[ id=ident -> fun x -> (!@loc,(id,x)) ]];
- hyp:
- [[ id=loc_id -> id None ;
- | id=loc_id ; ":" ; c=constr -> id (Some c)]]
- ;
- consider_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=consider_vars -> (Hvar name) :: v
- | name=hyp;
- IDENT "such"; IDENT "that"; h=consider_hyps -> (Hvar name)::h
- ]]
- ;
- consider_hyps:
- [[ st=statement; IDENT "and"; h=consider_hyps -> Hprop st::h
- | st=statement; IDENT "and";
- IDENT "consider" ; v=consider_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]]
- ;
- assume_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=assume_vars -> (Hvar name) :: v
- | name=hyp;
- IDENT "such"; IDENT "that"; h=assume_hyps -> (Hvar name)::h
- ]]
- ;
- assume_hyps:
- [[ st=statement; IDENT "and"; h=assume_hyps -> Hprop st::h
- | st=statement; IDENT "and";
- IDENT "we"; IDENT "have" ; v=assume_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]]
- ;
- assume_clause:
- [[ IDENT "we" ; IDENT "have" ; v=assume_vars -> v
- | h=assume_hyps -> h ]]
- ;
- suff_vars:
- [[ name=hyp; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
- [Hvar name],c
- | name=hyp; ","; v=suff_vars ->
- let (q,c) = v in ((Hvar name) :: q),c
- | name=hyp;
- IDENT "such"; IDENT "that"; h=suff_hyps ->
- let (q,c) = h in ((Hvar name) :: q),c
- ]];
- suff_hyps:
- [[ st=statement; IDENT "and"; h=suff_hyps ->
- let (q,c) = h in (Hprop st::q),c
- | st=statement; IDENT "and";
- IDENT "to" ; IDENT "have" ; v=suff_vars ->
- let (q,c) = v in (Hprop st::q),c
- | st=statement; IDENT "to"; IDENT "show" ; c = constr_or_thesis ->
- [Hprop st],c
- ]]
- ;
- suff_clause:
- [[ IDENT "to" ; IDENT "have" ; v=suff_vars -> v
- | h=suff_hyps -> h ]]
- ;
- let_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=let_vars -> (Hvar name) :: v
- | name=hyp; IDENT "be";
- IDENT "such"; IDENT "that"; h=let_hyps -> (Hvar name)::h
- ]]
- ;
- let_hyps:
- [[ st=statement; IDENT "and"; h=let_hyps -> Hprop st::h
- | st=statement; IDENT "and"; "let"; v=let_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]];
- given_vars:
- [[ name=hyp -> [Hvar name]
- | name=hyp; ","; v=given_vars -> (Hvar name) :: v
- | name=hyp; IDENT "such"; IDENT "that"; h=given_hyps -> (Hvar name)::h
- ]]
- ;
- given_hyps:
- [[ st=statement; IDENT "and"; h=given_hyps -> Hprop st::h
- | st=statement; IDENT "and"; IDENT "given"; v=given_vars -> Hprop st::v
- | st=statement -> [Hprop st]
- ]];
- suppose_vars:
- [[name=hyp -> [Hvar name]
- |name=hyp; ","; v=suppose_vars -> (Hvar name) :: v
- |name=hyp; OPT[IDENT "be"];
- IDENT "such"; IDENT "that"; h=suppose_hyps -> (Hvar name)::h
- ]]
- ;
- suppose_hyps:
- [[ st=statement_or_thesis; IDENT "and"; h=suppose_hyps -> Hprop st::h
- | st=statement_or_thesis; IDENT "and"; IDENT "we"; IDENT "have";
- v=suppose_vars -> Hprop st::v
- | st=statement_or_thesis -> [Hprop st]
- ]]
- ;
- suppose_clause:
- [[ IDENT "we"; IDENT "have"; v=suppose_vars -> v;
- | h=suppose_hyps -> h ]]
- ;
- intro_step:
- [[ IDENT "suppose" ; h=assume_clause -> Psuppose h
- | IDENT "suppose" ; IDENT "it"; IDENT "is" ; c=pattern LEVEL "0" ;
- po=OPT[ "with"; p=LIST1 hyp SEP ","-> p ] ;
- ho=OPT[ IDENT "and" ; h=suppose_clause -> h ] ->
- Pcase (none_is_empty po,c,none_is_empty ho)
- | "let" ; v=let_vars -> Plet v
- | IDENT "take"; witnesses = LIST1 constr SEP "," -> Ptake witnesses
- | IDENT "assume"; h=assume_clause -> Passume h
- | IDENT "given"; h=given_vars -> Pgiven h
- | IDENT "define"; id=ident; args=LIST0 hyp;
- "as"; body=constr -> Pdefine(id,args,body)
- | IDENT "reconsider"; id=ident; "as" ; typ=constr -> Pcast (This id,typ)
- | IDENT "reconsider"; t=thesis; "as" ; typ=constr -> Pcast (Thesis t ,typ)
- ]]
- ;
- emphasis :
- [[ -> 0
- | "*" -> 1
- | "**" -> 2
- | "***" -> 3
- ]]
- ;
- bare_proof_instr:
- [[ c = cut_step -> c ;
- | i = intro_step -> i ]]
- ;
- proof_instr :
- [[ e=emphasis;i=bare_proof_instr;"." -> {emph=e;instr=i}]]
- ;
-END;;
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index 85d465a4b..07a0d5a50 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -274,10 +274,10 @@ let make_discr_match_el =
*)
let make_discr_match_brl i =
List.map_i
- (fun j (_,idl,patl,_) ->
+ (fun j (_,(idl,patl,_)) -> Loc.tag @@
if Int.equal j i
- then (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_True_ref))
- else (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_False_ref))
+ then (idl,patl, mkGRef (Lazy.force coq_True_ref))
+ else (idl,patl, mkGRef (Lazy.force coq_False_ref))
)
0
(*
@@ -464,13 +464,13 @@ let rec pattern_to_term_and_type env typ = Loc.with_unloc (function
*)
-let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
+let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
observe (str " Entering : " ++ Printer.pr_glob_constr rt);
match rt with
- | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ ->
+ | _, GRef _ | _, GVar _ | _, GEvar _ | _, GPatVar _ | _, GSort _ | _, GHole _ ->
(* do nothing (except changing type of course) *)
mk_result [] rt avoid
- | GApp(_,_,_) ->
+ | _, GApp(_,_) ->
let f,args = glob_decompose_app rt in
let args_res : (glob_constr list) build_entry_return =
List.fold_right (* create the arguments lists of constructors and combine them *)
@@ -482,20 +482,20 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
(mk_result [] [] avoid)
in
begin
- match f with
+ match Loc.obj f with
| GLambda _ ->
let rec aux t l =
match l with
| [] -> t
- | u::l ->
+ | u::l -> Loc.tag @@
match t with
- | GLambda(loc,na,_,nat,b) ->
- GLetIn(Loc.ghost,na,u,None,aux b l)
+ | loc, GLambda(na,_,nat,b) ->
+ GLetIn(na,u,None,aux b l)
| _ ->
- GApp(Loc.ghost,t,l)
+ GApp(t,l)
in
build_entry_lc env funnames avoid (aux f args)
- | GVar(_,id) when Id.Set.mem id funnames ->
+ | GVar id when Id.Set.mem id funnames ->
(* if we have [f t1 ... tn] with [f]$\in$[fnames]
then we create a fresh variable [res],
add [res] and its "value" (i.e. [res v1 ... vn]) to each
@@ -536,7 +536,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
args_res.result
}
| GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *)
- | GLetIn(_,n,v,t,b) ->
+ | GLetIn(n,v,t,b) ->
(* if we have [(let x := v in b) t1 ... tn] ,
we discard our work and compute the list of constructor for
[let x = v in (b t1 ... tn)] up to alpha conversion
@@ -550,7 +550,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let new_b =
replace_var_by_term
id
- (GVar(Loc.ghost,id))
+ (Loc.tag @@ GVar id)
b
in
(Name new_id,new_b,new_avoid)
@@ -568,7 +568,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
*)
let f_res = build_entry_lc env funnames args_res.to_avoid f in
combine_results combine_app f_res args_res
- | GCast(_,b,_) ->
+ | GCast(b,_) ->
(* for an applied cast we just trash the cast part
and restart the work.
@@ -579,7 +579,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
| GProd _ -> error "Cannot apply a type"
end (* end of the application treatement *)
- | GLambda(_,n,_,t,b) ->
+ | _, GLambda(n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -594,7 +594,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let new_env = raw_push_named (new_n,None,t) env in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_lam new_n) t_res b_res
- | GProd(_,n,_,t,b) ->
+ | _, GProd(n,_,t,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the type
@@ -604,13 +604,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let new_env = raw_push_named (n,None,t) env in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_prod n) t_res b_res
- | GLetIn(loc,n,v,typ,b) ->
+ | loc, GLetIn(n,v,typ,b) ->
(* we first compute the list of constructor
corresponding to the body of the function,
then the one corresponding to the value [t]
and combine the two result
*)
- let v = match typ with None -> v | Some t -> GCast (loc,v,CastConv t) in
+ let v = match typ with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in
let v_res = build_entry_lc env funnames avoid v in
let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in
let v_type = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr v_as_constr) in
@@ -622,13 +622,13 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
in
let b_res = build_entry_lc new_env funnames avoid b in
combine_results (combine_letin n) v_res b_res
- | GCases(_,_,_,el,brl) ->
+ | _, GCases(_,_,el,brl) ->
(* we create the discrimination function
and treat the case itself
*)
let make_discr = make_discr_match brl in
build_entry_lc_from_case env funnames make_discr el brl avoid
- | GIf(_,b,(na,e_option),lhs,rhs) ->
+ | _, GIf(b,(na,e_option),lhs,rhs) ->
let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in
let b_typ = Typing.unsafe_type_of env (Evd.from_env env) (EConstr.of_constr b_as_constr) in
let (ind,_) =
@@ -642,7 +642,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
assert (Int.equal (Array.length case_pats) 2);
let brl =
List.map_i
- (fun i x -> (Loc.ghost,[],[case_pats.(i)],x))
+ (fun i x -> Loc.tag ([],[case_pats.(i)],x))
0
[lhs;rhs]
in
@@ -651,7 +651,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
in
(* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *)
build_entry_lc env funnames avoid match_expr
- | GLetTuple(_,nal,_,b,e) ->
+ | _, GLetTuple(nal,_,b,e) ->
begin
let nal_as_glob_constr =
List.map
@@ -673,14 +673,14 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in
assert (Int.equal (Array.length case_pats) 1);
let br =
- (Loc.ghost,[],[case_pats.(0)],e)
+ (Loc.ghost,([],[case_pats.(0)],e))
in
let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in
build_entry_lc env funnames avoid match_expr
end
- | GRec _ -> error "Not handled GRec"
- | GCast(_,b,_) ->
+ | _, GRec _ -> error "Not handled GRec"
+ | _, GCast(b,_) ->
build_entry_lc env funnames avoid b
and build_entry_lc_from_case env funname make_discr
(el:tomatch_tuples)
@@ -740,7 +740,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
| [] -> (* computed_branches *) {result = [];to_avoid = avoid}
| br::brl' ->
(* alpha conversion to prevent name clashes *)
- let _,idl,patl,return = alpha_br avoid br in
+ let _,(idl,patl,return) = alpha_br avoid br in
let new_avoid = idl@avoid in (* for now we can no more use idl as an identifier *)
(* building a list of precondition stating that we are not in this branch
(will be used in the following recursive calls)
@@ -862,9 +862,9 @@ let is_res id =
-let same_raw_term rt1 rt2 =
+let same_raw_term (_,rt1) (_,rt2) =
match rt1,rt2 with
- | GRef(_,r1,_), GRef (_,r2,_) -> Globnames.eq_gr r1 r2
+ | GRef(r1,_), GRef (r2,_) -> Globnames.eq_gr r1 r2
| GHole _, GHole _ -> true
| _ -> false
let decompose_raw_eq lhs rhs =
@@ -897,15 +897,15 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
observe (str "rebuilding : " ++ pr_glob_constr rt);
let open Context.Rel.Declaration in
match rt with
- | GProd(_,n,k,t,b) ->
+ | _, GProd(n,k,t,b) ->
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t::crossed_types in
begin
match t with
- | GApp(_,(GVar(_,res_id) as res_rt),args') when is_res res_id ->
+ | _, GApp(((_, GVar(res_id)) as res_rt),args') when is_res res_id ->
begin
match args' with
- | (GVar(_,this_relname))::args' ->
+ | (_, (GVar this_relname))::args' ->
(*i The next call to mk_rel_id is
valid since we are constructing the graph
Ensures by: obvious
@@ -927,7 +927,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| _ -> (* the first args is the name of the function! *)
assert false
end
- | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;GVar(loc3,id);rt])
+ | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;(loc3, GVar id);rt])
when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous
->
begin
@@ -964,9 +964,8 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let params,arg' =
((Util.List.chop nparam args'))
in
- let rt_typ =
- GApp(Loc.ghost,
- GRef (Loc.ghost,Globnames.IndRef (fst ind),None),
+ let rt_typ = Loc.tag @@
+ GApp(Loc.tag @@ GRef (Globnames.IndRef (fst ind),None),
(List.map
(fun p -> Detyping.detype false []
env (Evd.from_env env)
@@ -976,7 +975,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(mkGHole ()))))
in
let eq' =
- GApp(loc1,GRef(loc2,jmeq,None),[ty;GVar(loc3,id);rt_typ;rt])
+ Loc.tag ~loc:loc1 @@ GApp(Loc.tag ~loc:loc2 @@GRef(jmeq,None),[ty;Loc.tag ~loc:loc3 @@ GVar id;rt_typ;rt])
in
observe (str "computing new type for jmeq : " ++ pr_glob_constr eq');
let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in
@@ -1045,7 +1044,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
mkGProd(n,t,new_b),id_to_exclude
else new_b, Id.Set.add id id_to_exclude
*)
- | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;rt1;rt2])
+ | loc1, GApp((loc2, GRef(eq_as_ref,_)),[ty;rt1;rt2])
when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous
->
begin
@@ -1096,7 +1095,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(Id.Set.filter not_free_in_t id_to_exclude)
| _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
end
- | GLambda(_,n,k,t,b) ->
+ | _, GLambda(n,k,t,b) ->
begin
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t :: crossed_types in
@@ -1115,14 +1114,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
then
new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude)
else
- GProd(Loc.ghost,n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
+ Loc.tag @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
| _ -> anomaly (Pp.str "Should not have an anonymous function here")
(* We have renamed all the anonymous functions during alpha_renaming phase *)
end
- | GLetIn(loc,n,v,t,b) ->
+ | loc, GLetIn(n,v,t,b) ->
begin
- let t = match t with None -> v | Some t -> GCast (loc,v,CastConv t) in
+ let t = match t with None -> v | Some t -> Loc.tag ~loc @@ GCast (v,CastConv t) in
let not_free_in_t id = not (is_free_in id t) in
let evd = (Evd.from_env env) in
let t',ctx = Pretyping.understand env evd t in
@@ -1138,10 +1137,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
match n with
| Name id when Id.Set.mem id id_to_exclude && depth >= nb_args ->
new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude)
- | _ -> GLetIn(Loc.ghost,n,t,None,new_b), (* HOPING IT WOULD WORK *)
+ | _ -> Loc.tag @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *)
Id.Set.filter not_free_in_t id_to_exclude
end
- | GLetTuple(_,nal,(na,rto),t,b) ->
+ | _, GLetTuple(nal,(na,rto),t,b) ->
assert (Option.is_empty rto);
begin
let not_free_in_t id = not (is_free_in id t) in
@@ -1164,7 +1163,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(* | Name id when Id.Set.mem id id_to_exclude -> *)
(* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *)
(* | _ -> *)
- GLetTuple(Loc.ghost,nal,(na,None),t,new_b),
+ Loc.tag @@ GLetTuple(nal,(na,None),t,new_b),
Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude')
end
@@ -1190,16 +1189,16 @@ let rebuild_cons env nb_args relname args crossed_types rt =
TODO: Find a valid way to deal with implicit arguments here!
*)
-let rec compute_cst_params relnames params = function
+let rec compute_cst_params relnames params gt = Loc.with_unloc (function
| GRef _ | GVar _ | GEvar _ | GPatVar _ -> params
- | GApp(_,GVar(_,relname'),rtl) when Id.Set.mem relname' relnames ->
+ | GApp((_, GVar relname'),rtl) when Id.Set.mem relname' relnames ->
compute_cst_params_from_app [] (params,rtl)
- | GApp(_,f,args) ->
+ | GApp(f,args) ->
List.fold_left (compute_cst_params relnames) params (f::args)
- | GLambda(_,_,_,t,b) | GProd(_,_,_,t,b) | GLetTuple(_,_,_,t,b) ->
+ | GLambda(_,_,t,b) | GProd(_,_,t,b) | GLetTuple(_,_,t,b) ->
let t_params = compute_cst_params relnames params t in
compute_cst_params relnames t_params b
- | GLetIn(_,_,v,t,b) ->
+ | GLetIn(_,v,t,b) ->
let v_params = compute_cst_params relnames params v in
let t_params = Option.fold_left (compute_cst_params relnames) v_params t in
compute_cst_params relnames t_params b
@@ -1210,10 +1209,11 @@ let rec compute_cst_params relnames params = function
| GHole _ -> params
| GIf _ | GRec _ | GCast _ ->
raise (UserError(Some "compute_cst_params", str "Not handled case"))
+ ) gt
and compute_cst_params_from_app acc (params,rtl) =
match params,rtl with
| _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *)
- | ((Name id,_,None) as param)::params',(GVar(_,id'))::rtl'
+ | ((Name id,_,None) as param)::params',(_, GVar id')::rtl'
when Id.compare id id' == 0 ->
compute_cst_params_from_app (param::acc) (params',rtl')
| _ -> List.rev acc
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 51ca8c471..01e607412 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -10,16 +10,16 @@ open Misctypes
Some basic functions to rebuild glob_constr
In each of them the location is Loc.ghost
*)
-let mkGRef ref = GRef(Loc.ghost,ref,None)
-let mkGVar id = GVar(Loc.ghost,id)
-let mkGApp(rt,rtl) = GApp(Loc.ghost,rt,rtl)
-let mkGLambda(n,t,b) = GLambda(Loc.ghost,n,Explicit,t,b)
-let mkGProd(n,t,b) = GProd(Loc.ghost,n,Explicit,t,b)
-let mkGLetIn(n,b,t,c) = GLetIn(Loc.ghost,n,b,t,c)
-let mkGCases(rto,l,brl) = GCases(Loc.ghost,Term.RegularStyle,rto,l,brl)
-let mkGSort s = GSort(Loc.ghost,s)
-let mkGHole () = GHole(Loc.ghost,Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None)
-let mkGCast(b,t) = GCast(Loc.ghost,b,CastConv t)
+let mkGRef ref = Loc.tag @@ GRef(ref,None)
+let mkGVar id = Loc.tag @@ GVar(id)
+let mkGApp(rt,rtl) = Loc.tag @@ GApp(rt,rtl)
+let mkGLambda(n,t,b) = Loc.tag @@ GLambda(n,Explicit,t,b)
+let mkGProd(n,t,b) = Loc.tag @@ GProd(n,Explicit,t,b)
+let mkGLetIn(n,b,t,c) = Loc.tag @@ GLetIn(n,b,t,c)
+let mkGCases(rto,l,brl) = Loc.tag @@ GCases(Term.RegularStyle,rto,l,brl)
+let mkGSort s = Loc.tag @@ GSort(s)
+let mkGHole () = Loc.tag @@ GHole(Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None)
+let mkGCast(b,t) = Loc.tag @@ GCast(b,CastConv t)
(*
Some basic functions to decompose glob_constrs
@@ -27,7 +27,7 @@ let mkGCast(b,t) = GCast(Loc.ghost,b,CastConv t)
*)
let glob_decompose_prod =
let rec glob_decompose_prod args = function
- | GProd(_,n,k,t,b) ->
+ | _, GProd(n,k,t,b) ->
glob_decompose_prod ((n,t)::args) b
| rt -> args,rt
in
@@ -35,9 +35,9 @@ let glob_decompose_prod =
let glob_decompose_prod_or_letin =
let rec glob_decompose_prod args = function
- | GProd(_,n,k,t,b) ->
+ | _, GProd(n,k,t,b) ->
glob_decompose_prod ((n,None,Some t)::args) b
- | GLetIn(_,n,b,t,c) ->
+ | _, GLetIn(n,b,t,c) ->
glob_decompose_prod ((n,Some b,t)::args) c
| rt -> args,rt
in
@@ -59,7 +59,7 @@ let glob_decompose_prod_n n =
if i<=0 then args,c
else
match c with
- | GProd(_,n,_,t,b) ->
+ | _, GProd(n,_,t,b) ->
glob_decompose_prod (i-1) ((n,t)::args) b
| rt -> args,rt
in
@@ -71,9 +71,9 @@ let glob_decompose_prod_or_letin_n n =
if i<=0 then args,c
else
match c with
- | GProd(_,n,_,t,b) ->
+ | _, GProd(n,_,t,b) ->
glob_decompose_prod (i-1) ((n,None,Some t)::args) b
- | GLetIn(_,n,b,t,c) ->
+ | _, GLetIn(n,b,t,c) ->
glob_decompose_prod (i-1) ((n,Some b,t)::args) c
| rt -> args,rt
in
@@ -84,7 +84,7 @@ let glob_decompose_app =
let rec decompose_rapp acc rt =
(* msgnl (str "glob_decompose_app on : "++ Printer.pr_glob_constr rt); *)
match rt with
- | GApp(_,rt,rtl) ->
+ | _, GApp(rt,rtl) ->
decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt
| rt -> rt,List.rev acc
in
@@ -120,75 +120,70 @@ let remove_name_from_mapping mapping na =
let change_vars =
let rec change_vars mapping rt =
- match rt with
- | GRef _ -> rt
- | GVar(loc,id) ->
+ Loc.map (function
+ | GRef _ as x -> x
+ | GVar id ->
let new_id =
try
Id.Map.find id mapping
with Not_found -> id
in
- GVar(loc,new_id)
- | GEvar _ -> rt
- | GPatVar _ -> rt
- | GApp(loc,rt',rtl) ->
- GApp(loc,
- change_vars mapping rt',
+ GVar(new_id)
+ | GEvar _ as x -> x
+ | GPatVar _ as x -> x
+ | GApp(rt',rtl) ->
+ GApp(change_vars mapping rt',
List.map (change_vars mapping) rtl
)
- | GLambda(loc,name,k,t,b) ->
- GLambda(loc,
- name,
+ | GLambda(name,k,t,b) ->
+ GLambda(name,
k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
- | GProd(loc,name,k,t,b) ->
- GProd(loc,
- name,
+ | GProd(name,k,t,b) ->
+ GProd( name,
k,
change_vars mapping t,
change_vars (remove_name_from_mapping mapping name) b
)
- | GLetIn(loc,name,def,typ,b) ->
- GLetIn(loc,
- name,
+ | GLetIn(name,def,typ,b) ->
+ GLetIn(name,
change_vars mapping def,
Option.map (change_vars mapping) typ,
change_vars (remove_name_from_mapping mapping name) b
)
- | GLetTuple(loc,nal,(na,rto),b,e) ->
+ | GLetTuple(nal,(na,rto),b,e) ->
let new_mapping = List.fold_left remove_name_from_mapping mapping nal in
- GLetTuple(loc,
- nal,
+ GLetTuple(nal,
(na, Option.map (change_vars mapping) rto),
change_vars mapping b,
change_vars new_mapping e
)
- | GCases(loc,sty,infos,el,brl) ->
- GCases(loc,sty,
+ | GCases(sty,infos,el,brl) ->
+ GCases(sty,
infos,
List.map (fun (e,x) -> (change_vars mapping e,x)) el,
List.map (change_vars_br mapping) brl
)
- | GIf(loc,b,(na,e_option),lhs,rhs) ->
- GIf(loc,
- change_vars mapping b,
+ | GIf(b,(na,e_option),lhs,rhs) ->
+ GIf(change_vars mapping b,
(na,Option.map (change_vars mapping) e_option),
change_vars mapping lhs,
change_vars mapping rhs
)
| GRec _ -> error "Local (co)fixes are not supported"
- | GSort _ -> rt
- | GHole _ -> rt
- | GCast(loc,b,c) ->
- GCast(loc,change_vars mapping b,
+ | GSort _ as x -> x
+ | GHole _ as x -> x
+ | GCast(b,c) ->
+ GCast(change_vars mapping b,
Miscops.map_cast_type (change_vars mapping) c)
- and change_vars_br mapping ((loc,idl,patl,res) as br) =
+ ) rt
+ and change_vars_br mapping ((loc,(idl,patl,res)) as br) =
let new_mapping = List.fold_right Id.Map.remove idl mapping in
if Id.Map.is_empty new_mapping
then br
- else (loc,idl,patl,change_vars new_mapping res)
+ else (loc,(idl,patl,change_vars new_mapping res))
in
change_vars
@@ -259,30 +254,30 @@ let raw_get_pattern_id pat acc =
let get_pattern_id pat = raw_get_pattern_id pat []
-let rec alpha_rt excluded rt =
- let new_rt =
+let rec alpha_rt excluded (loc, rt) =
+ let new_rt = Loc.tag ~loc @@
match rt with
| GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt
- | GLambda(loc,Anonymous,k,t,b) ->
+ | GLambda(Anonymous,k,t,b) ->
let new_id = Namegen.next_ident_away (Id.of_string "_x") excluded in
let new_excluded = new_id :: excluded in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- GLambda(loc,Name new_id,k,new_t,new_b)
- | GProd(loc,Anonymous,k,t,b) ->
+ GLambda(Name new_id,k,new_t,new_b)
+ | GProd(Anonymous,k,t,b) ->
let new_t = alpha_rt excluded t in
let new_b = alpha_rt excluded b in
- GProd(loc,Anonymous,k,new_t,new_b)
- | GLetIn(loc,Anonymous,b,t,c) ->
+ GProd(Anonymous,k,new_t,new_b)
+ | GLetIn(Anonymous,b,t,c) ->
let new_b = alpha_rt excluded b in
let new_t = Option.map (alpha_rt excluded) t in
let new_c = alpha_rt excluded c in
- GLetIn(loc,Anonymous,new_b,new_t,new_c)
- | GLambda(loc,Name id,k,t,b) ->
+ GLetIn(Anonymous,new_b,new_t,new_c)
+ | GLambda(Name id,k,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let t,b =
if Id.equal new_id id
- then t,b
+ then t, b
else
let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in
(t,replace b)
@@ -290,8 +285,8 @@ let rec alpha_rt excluded rt =
let new_excluded = new_id::excluded in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- GLambda(loc,Name new_id,k,new_t,new_b)
- | GProd(loc,Name id,k,t,b) ->
+ GLambda(Name new_id,k,new_t,new_b)
+ | GProd(Name id,k,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let new_excluded = new_id::excluded in
let t,b =
@@ -303,8 +298,8 @@ let rec alpha_rt excluded rt =
in
let new_t = alpha_rt new_excluded t in
let new_b = alpha_rt new_excluded b in
- GProd(loc,Name new_id,k,new_t,new_b)
- | GLetIn(loc,Name id,b,t,c) ->
+ GProd(Name new_id,k,new_t,new_b)
+ | GLetIn(Name id,b,t,c) ->
let new_id = Namegen.next_ident_away id excluded in
let c =
if Id.equal new_id id then c
@@ -314,10 +309,9 @@ let rec alpha_rt excluded rt =
let new_b = alpha_rt new_excluded b in
let new_t = Option.map (alpha_rt new_excluded) t in
let new_c = alpha_rt new_excluded c in
- GLetIn(loc,Name new_id,new_b,new_t,new_c)
-
+ GLetIn(Name new_id,new_b,new_t,new_c)
- | GLetTuple(loc,nal,(na,rto),t,b) ->
+ | GLetTuple(nal,(na,rto),t,b) ->
let rev_new_nal,new_excluded,mapping =
List.fold_left
(fun (nal,excluded,mapping) na ->
@@ -344,14 +338,14 @@ let rec alpha_rt excluded rt =
let new_t = alpha_rt new_excluded new_t in
let new_b = alpha_rt new_excluded new_b in
let new_rto = Option.map (alpha_rt new_excluded) new_rto in
- GLetTuple(loc,new_nal,(na,new_rto),new_t,new_b)
- | GCases(loc,sty,infos,el,brl) ->
+ GLetTuple(new_nal,(na,new_rto),new_t,new_b)
+ | GCases(sty,infos,el,brl) ->
let new_el =
List.map (function (rt,i) -> alpha_rt excluded rt, i) el
in
- GCases(loc,sty,infos,new_el,List.map (alpha_br excluded) brl)
- | GIf(loc,b,(na,e_o),lhs,rhs) ->
- GIf(loc,alpha_rt excluded b,
+ GCases(sty,infos,new_el,List.map (alpha_br excluded) brl)
+ | GIf(b,(na,e_o),lhs,rhs) ->
+ GIf(alpha_rt excluded b,
(na,Option.map (alpha_rt excluded) e_o),
alpha_rt excluded lhs,
alpha_rt excluded rhs
@@ -359,66 +353,65 @@ let rec alpha_rt excluded rt =
| GRec _ -> error "Not handled GRec"
| GSort _ -> rt
| GHole _ -> rt
- | GCast (loc,b,c) ->
- GCast(loc,alpha_rt excluded b,
+ | GCast (b,c) ->
+ GCast(alpha_rt excluded b,
Miscops.map_cast_type (alpha_rt excluded) c)
- | GApp(loc,f,args) ->
- GApp(loc,
- alpha_rt excluded f,
+ | GApp(f,args) ->
+ GApp(alpha_rt excluded f,
List.map (alpha_rt excluded) args
)
in
new_rt
-and alpha_br excluded (loc,ids,patl,res) =
+and alpha_br excluded (loc,(ids,patl,res)) =
let new_patl,new_excluded,mapping = alpha_patl excluded patl in
let new_ids = List.fold_right raw_get_pattern_id new_patl [] in
let new_excluded = new_ids@excluded in
let renamed_res = change_vars mapping res in
let new_res = alpha_rt new_excluded renamed_res in
- (loc,new_ids,new_patl,new_res)
+ (loc,(new_ids,new_patl,new_res))
(*
[is_free_in id rt] checks if [id] is a free variable in [rt]
*)
let is_free_in id =
- let rec is_free_in = function
+ let rec is_free_in (loc, gt) = match gt with
| GRef _ -> false
- | GVar(_,id') -> Id.compare id' id == 0
+ | GVar id' -> Id.compare id' id == 0
| GEvar _ -> false
| GPatVar _ -> false
- | GApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl)
- | GLambda(_,n,_,t,b) | GProd(_,n,_,t,b) ->
+ | GApp(rt,rtl) -> List.exists is_free_in (rt::rtl)
+ | GLambda(n,_,t,b) | GProd(n,_,t,b) ->
let check_in_b =
match n with
| Name id' -> not (Id.equal id' id)
| _ -> true
in
is_free_in t || (check_in_b && is_free_in b)
- | GLetIn(_,n,b,t,c) ->
+ | GLetIn(n,b,t,c) ->
let check_in_c =
match n with
| Name id' -> not (Id.equal id' id)
| _ -> true
in
is_free_in b || Option.cata is_free_in true t || (check_in_c && is_free_in c)
- | GCases(_,_,_,el,brl) ->
+ | GCases(_,_,el,brl) ->
(List.exists (fun (e,_) -> is_free_in e) el) ||
List.exists is_free_in_br brl
- | GLetTuple(_,nal,_,b,t) ->
+ | GLetTuple(nal,_,b,t) ->
let check_in_nal =
not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal)
in
is_free_in t || (check_in_nal && is_free_in b)
- | GIf(_,cond,_,br1,br2) ->
+ | GIf(cond,_,br1,br2) ->
is_free_in cond || is_free_in br1 || is_free_in br2
| GRec _ -> raise (UserError(None,str "Not handled GRec"))
| GSort _ -> false
| GHole _ -> false
- | GCast (_,b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t
- | GCast (_,b,CastCoerce) -> is_free_in b
- and is_free_in_br (_,ids,_,rt) =
+ | GCast (b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t
+ | GCast (b,CastCoerce) -> is_free_in b
+ and is_free_in_br (_,(ids,_,rt)) =
(not (Id.List.mem id ids)) && is_free_in rt
in
is_free_in
@@ -452,60 +445,55 @@ let rec pattern_to_term pt = Loc.with_unloc (function
let replace_var_by_term x_id term =
- let rec replace_var_by_pattern rt =
+ let rec replace_var_by_pattern (loc, rt) = Loc.tag ~loc @@
match rt with
| GRef _ -> rt
- | GVar(_,id) when Id.compare id x_id == 0 -> term
+ | GVar id when Id.compare id x_id == 0 -> Loc.obj term
| GVar _ -> rt
| GEvar _ -> rt
| GPatVar _ -> rt
- | GApp(loc,rt',rtl) ->
- GApp(loc,
- replace_var_by_pattern rt',
+ | GApp(rt',rtl) ->
+ GApp(replace_var_by_pattern rt',
List.map replace_var_by_pattern rtl
)
- | GLambda(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt
- | GLambda(loc,name,k,t,b) ->
- GLambda(loc,
- name,
+ | GLambda(Name id,_,_,_) when Id.compare id x_id == 0 -> rt
+ | GLambda(name,k,t,b) ->
+ GLambda(name,
k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | GProd(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt
- | GProd(loc,name,k,t,b) ->
- GProd(loc,
- name,
+ | GProd(Name id,_,_,_) when Id.compare id x_id == 0 -> rt
+ | GProd(name,k,t,b) ->
+ GProd( name,
k,
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | GLetIn(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt
- | GLetIn(loc,name,def,typ,b) ->
- GLetIn(loc,
- name,
+ | GLetIn(Name id,_,_,_) when Id.compare id x_id == 0 -> rt
+ | GLetIn(name,def,typ,b) ->
+ GLetIn(name,
replace_var_by_pattern def,
Option.map (replace_var_by_pattern) typ,
replace_var_by_pattern b
)
- | GLetTuple(_,nal,_,_,_)
+ | GLetTuple(nal,_,_,_)
when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal ->
rt
- | GLetTuple(loc,nal,(na,rto),def,b) ->
- GLetTuple(loc,
- nal,
+ | GLetTuple(nal,(na,rto),def,b) ->
+ GLetTuple(nal,
(na,Option.map replace_var_by_pattern rto),
replace_var_by_pattern def,
replace_var_by_pattern b
)
- | GCases(loc,sty,infos,el,brl) ->
- GCases(loc,sty,
+ | GCases(sty,infos,el,brl) ->
+ GCases(sty,
infos,
List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el,
List.map replace_var_by_pattern_br brl
)
- | GIf(loc,b,(na,e_option),lhs,rhs) ->
- GIf(loc, replace_var_by_pattern b,
+ | GIf(b,(na,e_option),lhs,rhs) ->
+ GIf(replace_var_by_pattern b,
(na,Option.map replace_var_by_pattern e_option),
replace_var_by_pattern lhs,
replace_var_by_pattern rhs
@@ -513,13 +501,13 @@ let replace_var_by_term x_id term =
| GRec _ -> raise (UserError(None,str "Not handled GRec"))
| GSort _ -> rt
| GHole _ -> rt
- | GCast(loc,b,c) ->
- GCast(loc,replace_var_by_pattern b,
+ | GCast(b,c) ->
+ GCast(replace_var_by_pattern b,
Miscops.map_cast_type replace_var_by_pattern c)
- and replace_var_by_pattern_br ((loc,idl,patl,res) as br) =
+ and replace_var_by_pattern_br ((loc,(idl,patl,res)) as br) =
if List.exists (fun id -> Id.compare id x_id == 0) idl
then br
- else (loc,idl,patl,replace_var_by_pattern res)
+ else (loc,(idl,patl,replace_var_by_pattern res))
in
replace_var_by_pattern
@@ -590,22 +578,22 @@ let id_of_name = function
(* TODO: finish Rec caes *)
let ids_of_glob_constr c =
- let rec ids_of_glob_constr acc c =
+ let rec ids_of_glob_constr acc (loc, c) =
let idof = id_of_name in
match c with
- | GVar (_,id) -> id::acc
- | GApp (loc,g,args) ->
+ | GVar id -> id::acc
+ | GApp (g,args) ->
ids_of_glob_constr [] g @ List.flatten (List.map (ids_of_glob_constr []) args) @ acc
- | GLambda (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
- | GProd (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
- | GLetIn (loc,na,b,t,c) -> idof na :: ids_of_glob_constr [] b @ Option.cata (ids_of_glob_constr []) [] t @ ids_of_glob_constr [] c @ acc
- | GCast (loc,c,(CastConv t|CastVM t|CastNative t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc
- | GCast (loc,c,CastCoerce) -> ids_of_glob_constr [] c @ acc
- | GIf (loc,c,(na,po),b1,b2) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] b1 @ ids_of_glob_constr [] b2 @ acc
- | GLetTuple (_,nal,(na,po),b,c) ->
+ | GLambda (na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
+ | GProd (na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
+ | GLetIn (na,b,t,c) -> idof na :: ids_of_glob_constr [] b @ Option.cata (ids_of_glob_constr []) [] t @ ids_of_glob_constr [] c @ acc
+ | GCast (c,(CastConv t|CastVM t|CastNative t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc
+ | GCast (c,CastCoerce) -> ids_of_glob_constr [] c @ acc
+ | GIf (c,(na,po),b1,b2) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] b1 @ ids_of_glob_constr [] b2 @ acc
+ | GLetTuple (nal,(na,po),b,c) ->
List.map idof nal @ ids_of_glob_constr [] b @ ids_of_glob_constr [] c @ acc
- | GCases (loc,sty,rtntypopt,tml,brchl) ->
- List.flatten (List.map (fun (_,idl,patl,c) -> idl @ ids_of_glob_constr [] c) brchl)
+ | GCases (sty,rtntypopt,tml,brchl) ->
+ List.flatten (List.map (fun (_,(idl,patl,c)) -> idl @ ids_of_glob_constr [] c) brchl)
| GRec _ -> failwith "Fix inside a constructor branch"
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> []
in
@@ -617,49 +605,46 @@ let ids_of_glob_constr c =
let zeta_normalize =
- let rec zeta_normalize_term rt =
+ let rec zeta_normalize_term (loc, rt) = Loc.tag ~loc @@
match rt with
| GRef _ -> rt
| GVar _ -> rt
| GEvar _ -> rt
| GPatVar _ -> rt
- | GApp(loc,rt',rtl) ->
- GApp(loc,
- zeta_normalize_term rt',
+ | GApp(rt',rtl) ->
+ GApp(zeta_normalize_term rt',
List.map zeta_normalize_term rtl
)
- | GLambda(loc,name,k,t,b) ->
- GLambda(loc,
- name,
+ | GLambda(name,k,t,b) ->
+ GLambda(name,
k,
zeta_normalize_term t,
zeta_normalize_term b
)
- | GProd(loc,name,k,t,b) ->
- GProd(loc,
- name,
+ | GProd(name,k,t,b) ->
+ GProd(name,
k,
zeta_normalize_term t,
zeta_normalize_term b
)
- | GLetIn(_,Name id,def,typ,b) ->
- zeta_normalize_term (replace_var_by_term id def b)
- | GLetIn(loc,Anonymous,def,typ,b) -> zeta_normalize_term b
- | GLetTuple(loc,nal,(na,rto),def,b) ->
- GLetTuple(loc,
- nal,
+ | GLetIn(Name id,def,typ,b) ->
+ Loc.obj @@ zeta_normalize_term (replace_var_by_term id def b)
+ | GLetIn(Anonymous,def,typ,b) ->
+ Loc.obj @@ zeta_normalize_term b
+ | GLetTuple(nal,(na,rto),def,b) ->
+ GLetTuple(nal,
(na,Option.map zeta_normalize_term rto),
zeta_normalize_term def,
zeta_normalize_term b
)
- | GCases(loc,sty,infos,el,brl) ->
- GCases(loc,sty,
+ | GCases(sty,infos,el,brl) ->
+ GCases(sty,
infos,
List.map (fun (e,x) -> (zeta_normalize_term e,x)) el,
List.map zeta_normalize_br brl
)
- | GIf(loc,b,(na,e_option),lhs,rhs) ->
- GIf(loc, zeta_normalize_term b,
+ | GIf(b,(na,e_option),lhs,rhs) ->
+ GIf(zeta_normalize_term b,
(na,Option.map zeta_normalize_term e_option),
zeta_normalize_term lhs,
zeta_normalize_term rhs
@@ -667,11 +652,11 @@ let zeta_normalize =
| GRec _ -> raise (UserError(None,str "Not handled GRec"))
| GSort _ -> rt
| GHole _ -> rt
- | GCast(loc,b,c) ->
- GCast(loc,zeta_normalize_term b,
+ | GCast(b,c) ->
+ GCast(zeta_normalize_term b,
Miscops.map_cast_type zeta_normalize_term c)
- and zeta_normalize_br (loc,idl,patl,res) =
- (loc,idl,patl,zeta_normalize_term res)
+ and zeta_normalize_br (loc,(idl,patl,res)) =
+ (loc,(idl,patl,zeta_normalize_term res))
in
zeta_normalize_term
@@ -687,33 +672,34 @@ let expand_as =
Id.Map.add id (pattern_to_term (loc, pat)) (List.fold_left add_as map patl)
| PatCstr(_,patl,_) -> List.fold_left add_as map patl
in
- let rec expand_as map rt =
+ let rec expand_as map (loc, rt) =
+ Loc.tag ~loc @@
match rt with
| GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> rt
- | GVar(_,id) ->
+ | GVar id ->
begin
try
- Id.Map.find id map
+ Loc.obj @@ Id.Map.find id map
with Not_found -> rt
end
- | GApp(loc,f,args) -> GApp(loc,expand_as map f,List.map (expand_as map) args)
- | GLambda(loc,na,k,t,b) -> GLambda(loc,na,k,expand_as map t, expand_as map b)
- | GProd(loc,na,k,t,b) -> GProd(loc,na,k,expand_as map t, expand_as map b)
- | GLetIn(loc,na,v,typ,b) -> GLetIn(loc,na, expand_as map v,Option.map (expand_as map) typ,expand_as map b)
- | GLetTuple(loc,nal,(na,po),v,b) ->
- GLetTuple(loc,nal,(na,Option.map (expand_as map) po),
+ | GApp(f,args) -> GApp(expand_as map f,List.map (expand_as map) args)
+ | GLambda(na,k,t,b) -> GLambda(na,k,expand_as map t, expand_as map b)
+ | GProd(na,k,t,b) -> GProd(na,k,expand_as map t, expand_as map b)
+ | GLetIn(na,v,typ,b) -> GLetIn(na, expand_as map v,Option.map (expand_as map) typ,expand_as map b)
+ | GLetTuple(nal,(na,po),v,b) ->
+ GLetTuple(nal,(na,Option.map (expand_as map) po),
expand_as map v, expand_as map b)
- | GIf(loc,e,(na,po),br1,br2) ->
- GIf(loc,expand_as map e,(na,Option.map (expand_as map) po),
+ | GIf(e,(na,po),br1,br2) ->
+ GIf(expand_as map e,(na,Option.map (expand_as map) po),
expand_as map br1, expand_as map br2)
| GRec _ -> error "Not handled GRec"
- | GCast(loc,b,c) ->
- GCast(loc,expand_as map b,
+ | GCast(b,c) ->
+ GCast(expand_as map b,
Miscops.map_cast_type (expand_as map) c)
- | GCases(loc,sty,po,el,brl) ->
- GCases(loc, sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el,
+ | GCases(sty,po,el,brl) ->
+ GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el,
List.map (expand_as_br map) brl)
- and expand_as_br map (loc,idl,cpl,rt) =
- (loc,idl,cpl, expand_as (List.fold_left add_as map cpl) rt)
+ and expand_as_br map (loc,(idl,cpl,rt)) =
+ (loc,(idl,cpl, expand_as (List.fold_left add_as map cpl) rt))
in
expand_as Id.Map.empty
diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli
index 84359a36b..25d79582f 100644
--- a/plugins/funind/glob_termops.mli
+++ b/plugins/funind/glob_termops.mli
@@ -82,11 +82,8 @@ val alpha_rt : Id.t list -> glob_constr -> glob_constr
(* same as alpha_rt but for case branches *)
val alpha_br : Id.t list ->
- Loc.t * Id.t list * Glob_term.cases_pattern list *
- Glob_term.glob_constr ->
- Loc.t * Id.t list * Glob_term.cases_pattern list *
- Glob_term.glob_constr
-
+ Glob_term.cases_clause ->
+ Glob_term.cases_clause
(* Reduction function *)
val replace_var_by_term :
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index 6ee755323..cad96946c 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -190,18 +190,18 @@ let build_newrecursive l =
let is_rec names =
let names = List.fold_right Id.Set.add names Id.Set.empty in
let check_id id names = Id.Set.mem id names in
- let rec lookup names = function
- | GVar(_,id) -> check_id id names
+ let rec lookup names (loc, gt) = match gt with
+ | GVar(id) -> check_id id names
| GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> false
- | GCast(_,b,_) -> lookup names b
+ | GCast(b,_) -> lookup names b
| GRec _ -> error "GRec not handled"
- | GIf(_,b,_,lhs,rhs) ->
+ | GIf(b,_,lhs,rhs) ->
(lookup names b) || (lookup names lhs) || (lookup names rhs)
- | GProd(_,na,_,t,b) | GLambda(_,na,_,t,b) ->
+ | GProd(na,_,t,b) | GLambda(na,_,t,b) ->
lookup names t || lookup (Nameops.name_fold Id.Set.remove na names) b
- | GLetIn(_,na,b,t,c) ->
+ | GLetIn(na,b,t,c) ->
lookup names b || Option.cata (lookup names) true t || lookup (Nameops.name_fold Id.Set.remove na names) c
- | GLetTuple(_,nal,_,t,b) -> lookup names t ||
+ | GLetTuple(nal,_,t,b) -> lookup names t ||
lookup
(List.fold_left
(fun acc na -> Nameops.name_fold Id.Set.remove na acc)
@@ -209,11 +209,11 @@ let is_rec names =
nal
)
b
- | GApp(_,f,args) -> List.exists (lookup names) (f::args)
- | GCases(_,_,_,el,brl) ->
+ | GApp(f,args) -> List.exists (lookup names) (f::args)
+ | GCases(_,_,el,brl) ->
List.exists (fun (e,_) -> lookup names e) el ||
List.exists (lookup_br names) brl
- and lookup_br names (_,idl,_,rt) =
+ and lookup_br names (_,(idl,_,rt)) =
let new_names = List.fold_right Id.Set.remove idl names in
lookup new_names rt
in
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 7b0d7d27d..8f0c98624 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -69,9 +69,9 @@ let chop_rlambda_n =
if n == 0
then List.rev acc,rt
else
- match rt with
- | Glob_term.GLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b
- | Glob_term.GLetIn(_,name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b
+ match Loc.obj rt with
+ | Glob_term.GLambda(name,k,t,b) -> chop_lambda_n ((name,t,None)::acc) (n-1) b
+ | Glob_term.GLetIn(name,v,t,b) -> chop_lambda_n ((name,v,t)::acc) (n-1) b
| _ ->
raise (CErrors.UserError(Some "chop_rlambda_n",
str "chop_rlambda_n: Not enough Lambdas"))
@@ -83,8 +83,8 @@ let chop_rprod_n =
if n == 0
then List.rev acc,rt
else
- match rt with
- | Glob_term.GProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
+ match Loc.obj rt with
+ | Glob_term.GProd(name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b
| _ -> raise (CErrors.UserError(Some "chop_rprod_n",str "chop_rprod_n: Not enough products"))
in
chop_prod_n []
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index 29de56478..9dc1d48df 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -66,7 +66,7 @@ let string_of_name = id_of_name %> Id.to_string
(** [isVarf f x] returns [true] if term [x] is of the form [(Var f)]. *)
let isVarf f x =
match x with
- | GVar (_,x) -> Id.equal x f
+ | _, GVar x -> Id.equal x f
| _ -> false
(** [ident_global_exist id] returns true if identifier [id] is linked
@@ -504,40 +504,40 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array
exception NoMerge
-let rec merge_app c1 c2 id1 id2 shift filter_shift_stable =
+let rec merge_app (loc1, c1) (loc2, c2) id1 id2 shift filter_shift_stable =
let lnk = Array.append shift.lnk1 shift.lnk2 in
- match c1 , c2 with
- | GApp(_,f1, arr1), GApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 ->
+ match c1, c2 with
+ | GApp(f1, arr1), GApp(f2,arr2) when isVarf id1 f1 && isVarf id2 f2 ->
let _ = prstr "\nICI1!\n" in
let args = filter_shift_stable lnk (arr1 @ arr2) in
- GApp (Loc.ghost,GVar (Loc.ghost,shift.ident) , args)
- | GApp(_,f1, arr1), GApp(_,f2,arr2) -> raise NoMerge
- | GLetIn(_,nme,bdy,typ,trm) , _ ->
+ Loc.tag @@ GApp ((Loc.tag @@ GVar shift.ident) , args)
+ | GApp(f1, arr1), GApp(f2,arr2) -> raise NoMerge
+ | GLetIn(nme,bdy,typ,trm) , _ ->
let _ = prstr "\nICI2!\n" in
- let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in
- GLetIn(Loc.ghost,nme,bdy,typ,newtrm)
- | _, GLetIn(_,nme,bdy,typ,trm) ->
+ let newtrm = merge_app trm (loc2, c2) id1 id2 shift filter_shift_stable in
+ Loc.tag @@ GLetIn(nme,bdy,typ,newtrm)
+ | _, GLetIn(nme,bdy,typ,trm) ->
let _ = prstr "\nICI3!\n" in
- let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in
- GLetIn(Loc.ghost,nme,bdy,typ,newtrm)
+ let newtrm = merge_app (loc1, c1) trm id1 id2 shift filter_shift_stable in
+ Loc.tag @@ GLetIn(nme,bdy,typ,newtrm)
| _ -> let _ = prstr "\nICI4!\n" in
raise NoMerge
-let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
+let rec merge_app_unsafe (l1, c1) (l2, c2) shift filter_shift_stable =
let lnk = Array.append shift.lnk1 shift.lnk2 in
match c1 , c2 with
- | GApp(_,f1, arr1), GApp(_,f2,arr2) ->
+ | GApp(f1, arr1), GApp(f2,arr2) ->
let args = filter_shift_stable lnk (arr1 @ arr2) in
- GApp (Loc.ghost,GVar(Loc.ghost,shift.ident) , args)
+ Loc.tag @@ GApp (Loc.tag @@ GVar shift.ident, args)
(* FIXME: what if the function appears in the body of the let? *)
- | GLetIn(_,nme,bdy,typ,trm) , _ ->
+ | GLetIn(nme,bdy,typ,trm) , _ ->
let _ = prstr "\nICI2 '!\n" in
- let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in
- GLetIn(Loc.ghost,nme,bdy,typ,newtrm)
- | _, GLetIn(_,nme,bdy,typ,trm) ->
+ let newtrm = merge_app_unsafe trm (l2, c2) shift filter_shift_stable in
+ Loc.tag @@ GLetIn(nme,bdy,typ,newtrm)
+ | _, GLetIn(nme,bdy,typ,trm) ->
let _ = prstr "\nICI3 '!\n" in
- let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in
- GLetIn(Loc.ghost,nme,bdy,typ,newtrm)
+ let newtrm = merge_app_unsafe (l1, c1) trm shift filter_shift_stable in
+ Loc.tag @@ GLetIn(nme,bdy,typ,newtrm)
| _ -> let _ = prstr "\nICI4 '!\n" in raise NoMerge
@@ -550,14 +550,14 @@ let rec merge_rec_hyps shift accrec
filter_shift_stable : (Name.t * glob_constr option * glob_constr option) list =
let mergeonehyp t reldecl =
match reldecl with
- | (nme,x,Some (GApp(_,i,args) as ind))
+ | (nme,x,Some (_, GApp(i,args) as ind))
-> nme,x, Some (merge_app_unsafe ind t shift filter_shift_stable)
| (nme,Some _,None) -> error "letins with recursive calls not treated yet"
| (nme,None,Some _) -> assert false
| (nme,None,None) | (nme,Some _,Some _) -> assert false in
match ltyp with
| [] -> []
- | (nme,None,Some (GApp(_,f, largs) as t)) :: lt when isVarf ind2name f ->
+ | (nme,None,Some (_, GApp(f, largs) as t)) :: lt when isVarf ind2name f ->
let rechyps = List.map (mergeonehyp t) accrec in
rechyps @ merge_rec_hyps shift accrec lt filter_shift_stable
| e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable
@@ -573,7 +573,7 @@ let find_app (nme:Id.t) ltyp =
(List.map
(fun x ->
match x with
- | _,None,Some (GApp(_,f,_)) when isVarf nme f -> raise (Found 0)
+ | _,None,Some (_, GApp(f,_)) when isVarf nme f -> raise (Found 0)
| _ -> ())
ltyp);
false
@@ -633,7 +633,7 @@ let rec merge_types shift accrec1
rechyps , concl
| (nme,None, Some t1)as e ::lt1 ->
(match t1 with
- | GApp(_,f,carr) when isVarf ind1name f ->
+ | _, GApp(f,carr) when isVarf ind1name f ->
merge_types shift (e::accrec1) lt1 concl1 ltyp2 concl2
| _ ->
let recres, recconcl2 =
@@ -864,7 +864,7 @@ let mkProd_reldecl (rdecl:Context.Rel.Declaration.t) (t2:glob_constr) =
| LocalAssum (nme,t) ->
let t = EConstr.of_constr t in
let traw = Detyping.detype false [] (Global.env()) Evd.empty t in
- GProd (Loc.ghost,nme,Explicit,traw,t2)
+ Loc.tag @@ GProd (nme,Explicit,traw,t2)
| LocalDef _ -> assert false
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index ee7b33227..c796fe7a2 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -172,7 +172,6 @@ let simpl_iter clause =
let (value_f:Constr.constr list -> global_reference -> Constr.constr) =
let open Term in
fun al fterm ->
- let d0 = Loc.ghost in
let rev_x_id_l =
(
List.fold_left
@@ -189,16 +188,15 @@ let (value_f:Constr.constr list -> global_reference -> Constr.constr) =
in
let env = Environ.push_rel_context context (Global.env ()) in
let glob_body =
- GCases
- (d0,RegularStyle,None,
- [GApp(d0, GRef(d0,fterm,None), List.rev_map (fun x_id -> GVar(d0, x_id)) rev_x_id_l),
+ Loc.tag @@
+ GCases
+ (RegularStyle,None,
+ [Loc.tag @@ GApp(Loc.tag @@ GRef(fterm,None), List.rev_map (fun x_id -> Loc.tag @@ GVar x_id) rev_x_id_l),
(Anonymous,None)],
- [d0, [v_id], [d0,PatCstr((destIndRef
- (delayed_force coq_sig_ref),1),
- [d0, PatVar(Name v_id);
- d0, PatVar(Anonymous)],
- Anonymous)],
- GVar(d0,v_id)])
+ [Loc.tag ([v_id], [Loc.tag @@ PatCstr ((destIndRef (delayed_force coq_sig_ref),1),
+ [Loc.tag @@ PatVar(Name v_id); Loc.tag @@ PatVar Anonymous],
+ Anonymous)],
+ Loc.tag @@ GVar v_id)])
in
let body = fst (understand env (Evd.from_env env) glob_body)(*FIXME*) in
it_mkLambda_or_LetIn body context
diff --git a/plugins/ltac/extratactics.ml4 b/plugins/ltac/extratactics.ml4
index 38fdfb759..232bd851f 100644
--- a/plugins/ltac/extratactics.ml4
+++ b/plugins/ltac/extratactics.ml4
@@ -631,15 +631,15 @@ let subst_var_with_hole occ tid t =
let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in
let locref = ref 0 in
let rec substrec = function
- | GVar (_,id) as x ->
+ | (_, GVar id) as x ->
if Id.equal id tid
then
(decr occref;
if Int.equal !occref 0 then x
else
(incr locref;
- GHole (Loc.make_loc (!locref,0),
- Evar_kinds.QuestionMark(Evar_kinds.Define true),
+ Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@
+ GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),
Misctypes.IntroAnonymous, None)))
else x
| c -> map_glob_constr_left_to_right substrec c in
@@ -651,13 +651,13 @@ let subst_hole_with_term occ tc t =
let locref = ref 0 in
let occref = ref occ in
let rec substrec = function
- | GHole (_,Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) ->
+ | _, GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s) ->
decr occref;
if Int.equal !occref 0 then tc
else
(incr locref;
- GHole (Loc.make_loc (!locref,0),
- Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s))
+ Loc.tag ~loc:(Loc.make_loc (!locref,0)) @@
+ GHole (Evar_kinds.QuestionMark(Evar_kinds.Define true),Misctypes.IntroAnonymous,s))
| c -> map_glob_constr_left_to_right substrec c
in
substrec t
diff --git a/plugins/ltac/g_rewrite.ml4 b/plugins/ltac/g_rewrite.ml4
index c50100bf5..1f40c67b5 100644
--- a/plugins/ltac/g_rewrite.ml4
+++ b/plugins/ltac/g_rewrite.ml4
@@ -123,7 +123,7 @@ TACTIC EXTEND rewrite_strat
END
let clsubstitute o c =
- let is_tac id = match fst (fst (snd c)) with GVar (_, id') when Id.equal id' id -> true | _ -> false in
+ let is_tac id = match fst (fst (snd c)) with (_, GVar id') when Id.equal id' id -> true | _ -> false in
Tacticals.onAllHypsAndConcl
(fun cl ->
match cl with
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index ad76ef9c6..aec2e37fd 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -1085,8 +1085,8 @@ type 'a extra_genarg_printer =
let strip_prod_binders_glob_constr n (ty,_) =
let rec strip_ty acc n ty =
if Int.equal n 0 then (List.rev acc, (ty,None)) else
- match ty with
- Glob_term.GProd(loc,na,Explicit,a,b) ->
+ match Loc.obj ty with
+ Glob_term.GProd(na,Explicit,a,b) ->
strip_ty (([Loc.ghost,na],(a,None))::acc) (n-1) b
| _ -> error "Cannot translate fix tactic: not enough products" in
strip_ty [] n ty
diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml
index 75f890c96..e7d4c1be9 100644
--- a/plugins/ltac/tacintern.ml
+++ b/plugins/ltac/tacintern.ml
@@ -31,8 +31,6 @@ open Locus
(** Globalization of tactic expressions :
Conversion from [raw_tactic_expr] to [glob_tactic_expr] *)
-let dloc = Loc.ghost
-
let error_tactic_expected ?loc =
user_err ?loc (str "Tactic expected.")
@@ -74,14 +72,14 @@ let intern_name l ist = function
let strict_check = ref false
-let adjust_loc loc = if !strict_check then dloc else loc
+let adjust_loc loc = if !strict_check then Loc.ghost else loc
(* Globalize a name which must be bound -- actually just check it is bound *)
let intern_hyp ist (loc,id as locid) =
if not !strict_check then
locid
else if find_ident id ist then
- (dloc,id)
+ Loc.tag id
else
Pretype_errors.error_var_not_found ~loc id
@@ -110,12 +108,12 @@ let intern_ltac_variable ist = function
let intern_constr_reference strict ist = function
| Ident (_,id) as r when not strict && find_hyp id ist ->
- GVar (dloc,id), Some (Loc.tag @@ CRef (r,None))
+ (Loc.tag @@ GVar id), Some (Loc.tag @@ CRef (r,None))
| Ident (_,id) as r when find_var id ist ->
- GVar (dloc,id), if strict then None else Some (Loc.tag @@ CRef (r,None))
+ (Loc.tag @@ GVar id), if strict then None else Some (Loc.tag @@ CRef (r,None))
| r ->
let loc,_ as lqid = qualid_of_reference r in
- GRef (loc,locate_global_with_alias lqid,None),
+ Loc.tag @@ GRef (locate_global_with_alias lqid,None),
if strict then None else Some (Loc.tag @@ CRef (r,None))
let intern_move_location ist = function
@@ -273,8 +271,8 @@ let intern_destruction_arg ist = function
| clear,ElimOnIdent (loc,id) ->
if !strict_check then
(* If in a defined tactic, no intros-until *)
- match intern_constr ist (Loc.tag @@ CRef (Ident (dloc,id), None)) with
- | GVar (loc,id),_ -> clear,ElimOnIdent (loc,id)
+ match intern_constr ist (Loc.tag @@ CRef (Ident (Loc.tag id), None)) with
+ | (loc, GVar id), _ -> clear,ElimOnIdent (loc,id)
| c -> clear,ElimOnConstr (c,NoBindings)
else
clear,ElimOnIdent (loc,id)
@@ -352,10 +350,10 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) =
| _ -> Qualid (loc,qualid_of_path (path_of_global (smart_global r))) in
let sign = { Constrintern.ltac_vars = ist.ltacvars; Constrintern.ltac_bound = Id.Set.empty } in
let c = Constrintern.interp_reference sign r in
- match c with
- | GRef (_,r,None) ->
+ match Loc.obj c with
+ | GRef (r,None) ->
Inl (ArgArg (evaluable_of_global_reference ist.genv r,None))
- | GVar (_,id) ->
+ | GVar id ->
let r = evaluable_of_global_reference ist.genv (VarRef id) in
Inl (ArgArg (r,None))
| _ ->
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index de6c44b2b..a8d8eda1d 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -699,7 +699,7 @@ let interp_typed_pattern ist env sigma (_,c,_) =
let interp_constr_in_compound_list inj_fun dest_fun interp_fun ist env sigma l =
let try_expand_ltac_var sigma x =
try match dest_fun x with
- | GVar (_,id), _ ->
+ | (_, GVar id), _ ->
let v = Id.Map.find id ist.lfun in
sigma, List.map inj_fun (coerce_to_constr_list env v)
| _ ->
@@ -1074,7 +1074,7 @@ let interp_destruction_arg ist gl arg =
if Tactics.is_quantified_hypothesis id gl then
keep,ElimOnIdent (loc,id)
else
- let c = (GVar (loc,id),Some (Loc.tag @@ CRef (Ident (loc,id),None))) in
+ let c = (Loc.tag ~loc @@ GVar id,Some (Loc.tag @@ CRef (Ident (loc,id),None))) in
let f = { delayed = fun env sigma ->
let sigma = Sigma.to_evar_map sigma in
let (sigma,c) = interp_open_constr ist env sigma c in
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index dd68eac24..bef8139be 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -131,7 +131,7 @@ let closed_term_ast l =
let l = List.map (fun gr -> ArgArg(Loc.ghost,gr)) l in
TacFun([Name(Id.of_string"t")],
TacML(Loc.ghost,tacname,
- [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (GVar(Loc.ghost,Id.of_string"t"),None));
+ [TacGeneric (Genarg.in_gen (Genarg.glbwit Stdarg.wit_constr) (Loc.tag @@ GVar(Id.of_string"t"),None));
TacGeneric (Genarg.in_gen (Genarg.glbwit (Genarg.wit_list Stdarg.wit_ref)) l)]))
(*
let _ = add_tacdef false ((Loc.ghost,Id.of_string"ring_closed_term"
diff --git a/plugins/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index 1a5ef825d..f8cccf714 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -62,7 +62,6 @@ open Locusops
DECLARE PLUGIN "ssrmatching_plugin"
type loc = Loc.t
-let dummy_loc = Loc.ghost
let errorstrm = CErrors.user_err ~hdr:"ssrmatching"
let loc_error loc msg = CErrors.user_err ~loc ~hdr:msg (str msg)
let ppnl = Feedback.msg_info
@@ -159,10 +158,10 @@ let mkCLetIn loc name bo t = Loc.tag ~loc @@
CLetIn ((loc, name), bo, None, t)
let mkCCast loc t ty = Loc.tag ~loc @@ CCast (t, dC ty)
(** Constructors for rawconstr *)
-let mkRHole = GHole (dummy_loc, InternalHole, IntroAnonymous, None)
-let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args)
-let mkRCast rc rt = GCast (dummy_loc, rc, dC rt)
-let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t)
+let mkRHole = Loc.tag @@ GHole (InternalHole, IntroAnonymous, None)
+let mkRApp f args = if args = [] then f else Loc.tag @@ GApp (f, args)
+let mkRCast rc rt = Loc.tag @@ GCast (rc, dC rt)
+let mkRLambda n s t = Loc.tag @@ GLambda (n, Explicit, s, t)
(* ssrterm conbinators *)
let combineCG t1 t2 f g = match t1, t2 with
@@ -944,7 +943,7 @@ let glob_cpattern gs p =
let name = Name (id_of_string ("_ssrpat_" ^ s)) in
k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in
let bind_in t1 t2 =
- let d = dummy_loc in let n = Name (destCVar t1) in
+ let d = Loc.ghost in let n = Name (destCVar t1) in
fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in
let check_var t2 = if not (isCVar t2) then
loc_error (constr_loc t2) "Only identifiers are allowed here" in
@@ -1023,7 +1022,7 @@ type pattern = Evd.evar_map * (constr, constr) ssrpattern
let id_of_cpattern = function
| _,(_,Some (_loc, CRef (Ident (_, x), _))) -> Some x
| _,(_,Some (_loc, CAppExpl ((_, Ident (_, x), _), []))) -> Some x
- | _,(GRef (_, VarRef x, _) ,None) -> Some x
+ | _,((_, GRef (VarRef x, _)) ,None) -> Some x
| _ -> None
let id_of_Cterm t = match id_of_cpattern t with
| Some x -> x
@@ -1123,8 +1122,8 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
let mkG ?(k=' ') x = k,(x,None) in
let decode ist t ?reccall f g =
try match (pf_intern_term ist gl t) with
- | GCast(_,GHole _,CastConv(GLambda(_,Name x,_,_,c))) -> f x (' ',(c,None))
- | GVar(_,id)
+ | _, GCast((_, GHole _),CastConv((_, GLambda(Name x,_,_,c)))) -> f x (' ',(c,None))
+ | _, GVar id
when Id.Map.mem id ist.lfun &&
not(Option.is_empty reccall) &&
not(Option.is_empty wit_ssrpatternarg) ->
@@ -1166,17 +1165,17 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
sigma new_evars in
sigma in
let red = let rec decode_red (ist,red) = match red with
- | T(k,(GCast (_,GHole _,(CastConv(GLambda (_,Name id,_,_,t)))),None))
+ | T(k,((_, GCast ((_, GHole _),CastConv((_, GLambda (Name id,_,_,t))))),None))
when let id = string_of_id id in let len = String.length id in
(len > 8 && String.sub id 0 8 = "_ssrpat_") ->
let id = string_of_id id in let len = String.length id in
(match String.sub id 8 (len - 8), t with
- | "In", GApp(_, _, [t]) -> decodeG t xInT (fun x -> T x)
- | "In", GApp(_, _, [e; t]) -> decodeG t (eInXInT (mkG e)) (bad_enc id)
- | "In", GApp(_, _, [e; t; e_in_t]) ->
+ | "In", (_, GApp( _, [t])) -> decodeG t xInT (fun x -> T x)
+ | "In", (_, GApp( _, [e; t])) -> decodeG t (eInXInT (mkG e)) (bad_enc id)
+ | "In", (_, GApp( _, [e; t; e_in_t])) ->
decodeG t (eInXInT (mkG e))
(fun _ -> decodeG e_in_t xInT (fun _ -> assert false))
- | "As", GApp(_, _, [e; t]) -> decodeG t (eAsXInT (mkG e)) (bad_enc id)
+ | "As", (_, GApp(_, [e; t])) -> decodeG t (eAsXInT (mkG e)) (bad_enc id)
| _ -> bad_enc id ())
| T t -> decode ist ~reccall:decode_red t xInT (fun x -> T x)
| In_T t -> decode ist t inXInT inT
@@ -1202,13 +1201,13 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
pp(lazy(str"typed as: " ++ pr_pattern_w_ids red));
let mkXLetIn loc x (a,(g,c)) = match c with
| Some b -> a,(g,Some (mkCLetIn loc x (mkCHole loc) b))
- | None -> a,(GLetIn (loc,x,(GHole (loc, BinderType x, IntroAnonymous, None)), None, g), None) in
+ | None -> a,(Loc.tag ~loc @@ GLetIn (x, Loc.tag ~loc @@ GHole (BinderType x, IntroAnonymous, None), None, g), None) in
match red with
| T t -> let sigma, t = interp_term ist gl t in sigma, T t
| In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t
| X_In_T (x, rp) | In_X_In_T (x, rp) ->
let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in
- let rp = mkXLetIn dummy_loc (Name x) rp in
+ let rp = mkXLetIn Loc.ghost (Name x) rp in
let sigma, rp = interp_term ist gl rp in
let _, h, _, rp = destLetIn rp in
let sigma = cleanup_XinE h x rp sigma in
@@ -1217,7 +1216,7 @@ let interp_pattern ?wit_ssrpatternarg ist gl red redty =
| E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) ->
let mk e x p =
match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in
- let rp = mkXLetIn dummy_loc (Name x) rp in
+ let rp = mkXLetIn Loc.ghost (Name x) rp in
let sigma, rp = interp_term ist gl rp in
let _, h, _, rp = destLetIn rp in
let sigma = cleanup_XinE h x rp sigma in
@@ -1375,10 +1374,10 @@ let pf_fill_occ_term gl occ t =
let cl,(_,t) = fill_occ_term env concl occ sigma0 t in
cl, t
-let cpattern_of_id id = ' ', (GRef (dummy_loc, VarRef id, None), None)
+let cpattern_of_id id = ' ', (Loc.tag @@ GRef (VarRef id, None), None)
-let is_wildcard = function
- | _,(_,Some (_, CHole _)|GHole _,None) -> true
+let is_wildcard : cpattern -> bool = function
+ | _,(_,Some (_, CHole _)| (_, GHole _),None) -> true
| _ -> false
(* "ssrpattern" *)
diff --git a/plugins/syntax/ascii_syntax.ml b/plugins/syntax/ascii_syntax.ml
index ed8cc6ab0..dc0b87793 100644
--- a/plugins/syntax/ascii_syntax.ml
+++ b/plugins/syntax/ascii_syntax.ml
@@ -37,13 +37,13 @@ let glob_Ascii = lazy (make_reference "Ascii")
open Lazy
-let interp_ascii dloc p =
+let interp_ascii loc p =
let rec aux n p =
if Int.equal n 0 then [] else
let mp = p mod 2 in
- GRef (dloc,(if Int.equal mp 0 then glob_false else glob_true),None)
+ (Loc.tag ~loc @@ GRef ((if Int.equal mp 0 then glob_false else glob_true),None))
:: (aux (n-1) (p/2)) in
- GApp (dloc,GRef(dloc,force glob_Ascii,None), aux 8 p)
+ Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef(force glob_Ascii,None), aux 8 p)
let interp_ascii_string dloc s =
let p =
@@ -59,12 +59,12 @@ let interp_ascii_string dloc s =
let uninterp_ascii r =
let rec uninterp_bool_list n = function
| [] when Int.equal n 0 -> 0
- | GRef (_,k,_)::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l)
- | GRef (_,k,_)::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l)
+ | (_, GRef (k,_))::l when Globnames.eq_gr k glob_true -> 1+2*(uninterp_bool_list (n-1) l)
+ | (_, GRef (k,_))::l when Globnames.eq_gr k glob_false -> 2*(uninterp_bool_list (n-1) l)
| _ -> raise Non_closed_ascii in
try
let aux = function
- | GApp (_,GRef (_,k,_),l) when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l
+ | _, GApp ((_, GRef (k,_)),l) when Globnames.eq_gr k (force glob_Ascii) -> uninterp_bool_list 8 l
| _ -> raise Non_closed_ascii in
Some (aux r)
with
@@ -80,4 +80,4 @@ let _ =
Notation.declare_string_interpreter "char_scope"
(ascii_path,ascii_module)
interp_ascii_string
- ([GRef (Loc.ghost,static_glob_Ascii,None)], uninterp_ascii_string, true)
+ ([Loc.tag @@ GRef (static_glob_Ascii,None)], uninterp_ascii_string, true)
diff --git a/plugins/syntax/nat_syntax.ml b/plugins/syntax/nat_syntax.ml
index ab262fea7..90d643b7f 100644
--- a/plugins/syntax/nat_syntax.ml
+++ b/plugins/syntax/nat_syntax.ml
@@ -33,14 +33,14 @@ let warn_large_nat =
strbrk "may vary from 5000 to 70000 depending on your system " ++
strbrk "limits and on the command executed).")
-let nat_of_int dloc n =
+let nat_of_int loc n =
if is_pos_or_zero n then begin
if less_than threshold n then warn_large_nat ();
- let ref_O = GRef (dloc, glob_O, None) in
- let ref_S = GRef (dloc, glob_S, None) in
+ let ref_O = Loc.tag ~loc @@ GRef (glob_O, None) in
+ let ref_S = Loc.tag ~loc @@ GRef (glob_S, None) in
let rec mk_nat acc n =
if n <> zero then
- mk_nat (GApp (dloc,ref_S, [acc])) (sub_1 n)
+ mk_nat (Loc.tag ~loc @@ GApp (ref_S, [acc])) (sub_1 n)
else
acc
in
@@ -55,10 +55,11 @@ let nat_of_int dloc n =
exception Non_closed_number
-let rec int_of_nat = function
- | GApp (_,GRef (_,s,_),[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a)
- | GRef (_,z,_) when Globnames.eq_gr z glob_O -> zero
+let rec int_of_nat x = Loc.with_unloc (function
+ | GApp ((_, GRef (s,_)),[a]) when Globnames.eq_gr s glob_S -> add_1 (int_of_nat a)
+ | GRef (z,_) when Globnames.eq_gr z glob_O -> zero
| _ -> raise Non_closed_number
+ ) x
let uninterp_nat p =
try
@@ -73,4 +74,4 @@ let _ =
Notation.declare_numeral_interpreter "nat_scope"
(nat_path,datatypes_module_name)
nat_of_int
- ([GRef (Loc.ghost,glob_S,None); GRef (Loc.ghost,glob_O,None)], uninterp_nat, true)
+ ([Loc.tag @@ GRef (glob_S,None); Loc.tag @@ GRef (glob_O,None)], uninterp_nat, true)
diff --git a/plugins/syntax/numbers_syntax.ml b/plugins/syntax/numbers_syntax.ml
index a25ddb062..8876d464a 100644
--- a/plugins/syntax/numbers_syntax.ml
+++ b/plugins/syntax/numbers_syntax.ml
@@ -86,10 +86,10 @@ exception Non_closed
(* parses a *non-negative* integer (from bigint.ml) into an int31
wraps modulo 2^31 *)
-let int31_of_pos_bigint dloc n =
- let ref_construct = GRef (dloc, int31_construct, None) in
- let ref_0 = GRef (dloc, int31_0, None) in
- let ref_1 = GRef (dloc, int31_1, None) in
+let int31_of_pos_bigint loc n =
+ let ref_construct = Loc.tag ~loc @@ GRef (int31_construct, None) in
+ let ref_0 = Loc.tag ~loc @@ GRef (int31_0, None) in
+ let ref_1 = Loc.tag ~loc @@ GRef (int31_1, None) in
let rec args counter n =
if counter <= 0 then
[]
@@ -97,7 +97,7 @@ let int31_of_pos_bigint dloc n =
let (q,r) = div2_with_rest n in
(if r then ref_1 else ref_0)::(args (counter-1) q)
in
- GApp (dloc, ref_construct, List.rev (args 31 n))
+ Loc.tag ~loc @@ GApp (ref_construct, List.rev (args 31 n))
let error_negative dloc =
CErrors.user_err ~loc:dloc ~hdr:"interp_int31" (Pp.str "int31 are only non-negative numbers.")
@@ -114,12 +114,12 @@ let bigint_of_int31 =
let rec args_parsing args cur =
match args with
| [] -> cur
- | (GRef (_,b,_))::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur)
- | (GRef (_,b,_))::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur))
+ | (_, GRef (b,_))::l when eq_gr b int31_0 -> args_parsing l (mult_2 cur)
+ | (_, GRef (b,_))::l when eq_gr b int31_1 -> args_parsing l (add_1 (mult_2 cur))
| _ -> raise Non_closed
in
function
- | GApp (_, GRef (_, c, _), args) when eq_gr c int31_construct -> args_parsing args zero
+ | _, GApp ((_, GRef (c, _)), args) when eq_gr c int31_construct -> args_parsing args zero
| _ -> raise Non_closed
let uninterp_int31 i =
@@ -132,7 +132,7 @@ let uninterp_int31 i =
let _ = Notation.declare_numeral_interpreter int31_scope
(int31_path, int31_module)
interp_int31
- ([GRef (Loc.ghost, int31_construct, None)],
+ ([Loc.tag @@ GRef (int31_construct, None)],
uninterp_int31,
true)
@@ -162,34 +162,34 @@ let height bi =
in hght 0 base
(* n must be a non-negative integer (from bigint.ml) *)
-let word_of_pos_bigint dloc hght n =
- let ref_W0 = GRef (dloc, zn2z_W0, None) in
- let ref_WW = GRef (dloc, zn2z_WW, None) in
+let word_of_pos_bigint loc hght n =
+ let ref_W0 = Loc.tag ~loc @@ GRef (zn2z_W0, None) in
+ let ref_WW = Loc.tag ~loc @@ GRef (zn2z_WW, None) in
let rec decomp hgt n =
if hgt <= 0 then
- int31_of_pos_bigint dloc n
+ int31_of_pos_bigint loc n
else if equal n zero then
- GApp (dloc, ref_W0, [GHole (dloc, Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)])
+ Loc.tag ~loc @@ GApp (ref_W0, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None)])
else
let (h,l) = split_at hgt n in
- GApp (dloc, ref_WW, [GHole (dloc, Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None);
+ Loc.tag ~loc @@ GApp (ref_WW, [Loc.tag ~loc @@ GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous, None);
decomp (hgt-1) h;
decomp (hgt-1) l])
in
decomp hght n
-let bigN_of_pos_bigint dloc n =
+let bigN_of_pos_bigint loc n =
let h = height n in
- let ref_constructor = GRef (dloc, bigN_constructor h, None) in
- let word = word_of_pos_bigint dloc h n in
+ let ref_constructor = Loc.tag ~loc @@ GRef (bigN_constructor h, None) in
+ let word = word_of_pos_bigint loc h n in
let args =
if h < n_inlined then [word]
- else [Nat_syntax_plugin.Nat_syntax.nat_of_int dloc (of_int (h-n_inlined));word]
+ else [Nat_syntax_plugin.Nat_syntax.nat_of_int loc (of_int (h-n_inlined));word]
in
- GApp (dloc, ref_constructor, args)
+ Loc.tag ~loc @@ GApp (ref_constructor, args)
-let bigN_error_negative dloc =
- CErrors.user_err ~loc:dloc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.")
+let bigN_error_negative loc =
+ CErrors.user_err ~loc ~hdr:"interp_bigN" (Pp.str "bigN are only non-negative numbers.")
let interp_bigN dloc n =
if is_pos_or_zero n then
@@ -203,14 +203,14 @@ let interp_bigN dloc n =
let bigint_of_word =
let rec get_height rc =
match rc with
- | GApp (_,GRef(_,c,_), [_;lft;rght]) when eq_gr c zn2z_WW ->
+ | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW ->
1+max (get_height lft) (get_height rght)
| _ -> 0
in
let rec transform hght rc =
match rc with
- | GApp (_,GRef(_,c,_),_) when eq_gr c zn2z_W0-> zero
- | GApp (_,GRef(_,c,_), [_;lft;rght]) when eq_gr c zn2z_WW->
+ | _, GApp ((_, GRef(c,_)),_) when eq_gr c zn2z_W0-> zero
+ | _, GApp ((_, GRef(c,_)), [_;lft;rght]) when eq_gr c zn2z_WW->
let new_hght = hght-1 in
add (mult (rank new_hght)
(transform new_hght lft))
@@ -223,8 +223,8 @@ let bigint_of_word =
let bigint_of_bigN rc =
match rc with
- | GApp (_,_,[one_arg]) -> bigint_of_word one_arg
- | GApp (_,_,[_;second_arg]) -> bigint_of_word second_arg
+ | _, GApp (_,[one_arg]) -> bigint_of_word one_arg
+ | _, GApp (_,[_;second_arg]) -> bigint_of_word second_arg
| _ -> raise Non_closed
let uninterp_bigN rc =
@@ -240,7 +240,7 @@ let uninterp_bigN rc =
let bigN_list_of_constructors =
let rec build i =
if i < n_inlined+1 then
- GRef (Loc.ghost, bigN_constructor i,None)::(build (i+1))
+ (Loc.tag @@ GRef (bigN_constructor i,None))::(build (i+1))
else
[]
in
@@ -256,18 +256,18 @@ let _ = Notation.declare_numeral_interpreter bigN_scope
(*** Parsing for bigZ in digital notation ***)
-let interp_bigZ dloc n =
- let ref_pos = GRef (dloc, bigZ_pos, None) in
- let ref_neg = GRef (dloc, bigZ_neg, None) in
+let interp_bigZ loc n =
+ let ref_pos = Loc.tag ~loc @@ GRef (bigZ_pos, None) in
+ let ref_neg = Loc.tag ~loc @@ GRef (bigZ_neg, None) in
if is_pos_or_zero n then
- GApp (dloc, ref_pos, [bigN_of_pos_bigint dloc n])
+ Loc.tag ~loc @@ GApp (ref_pos, [bigN_of_pos_bigint loc n])
else
- GApp (dloc, ref_neg, [bigN_of_pos_bigint dloc (neg n)])
+ Loc.tag ~loc @@ GApp (ref_neg, [bigN_of_pos_bigint loc (neg n)])
(* pretty printing functions for bigZ *)
let bigint_of_bigZ = function
- | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigZ_pos -> bigint_of_bigN one_arg
- | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigZ_neg ->
+ | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_pos -> bigint_of_bigN one_arg
+ | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigZ_neg ->
let opp_val = bigint_of_bigN one_arg in
if equal opp_val zero then
raise Non_closed
@@ -286,19 +286,19 @@ let uninterp_bigZ rc =
let _ = Notation.declare_numeral_interpreter bigZ_scope
(bigZ_path, bigZ_module)
interp_bigZ
- ([GRef (Loc.ghost, bigZ_pos, None);
- GRef (Loc.ghost, bigZ_neg, None)],
+ ([Loc.tag @@ GRef (bigZ_pos, None);
+ Loc.tag @@ GRef (bigZ_neg, None)],
uninterp_bigZ,
true)
(*** Parsing for bigQ in digital notation ***)
-let interp_bigQ dloc n =
- let ref_z = GRef (dloc, bigQ_z, None) in
- GApp (dloc, ref_z, [interp_bigZ dloc n])
+let interp_bigQ loc n =
+ let ref_z = Loc.tag ~loc @@ GRef (bigQ_z, None) in
+ Loc.tag ~loc @@ GApp (ref_z, [interp_bigZ loc n])
let uninterp_bigQ rc =
try match rc with
- | GApp (_, GRef(_,c,_), [one_arg]) when eq_gr c bigQ_z ->
+ | _, GApp ((_, GRef(c,_)), [one_arg]) when eq_gr c bigQ_z ->
Some (bigint_of_bigZ one_arg)
| _ -> None (* we don't pretty-print yet fractions *)
with Non_closed -> None
@@ -307,5 +307,5 @@ let uninterp_bigQ rc =
let _ = Notation.declare_numeral_interpreter bigQ_scope
(bigQ_path, bigQ_module)
interp_bigQ
- ([GRef (Loc.ghost, bigQ_z, None)], uninterp_bigQ,
+ ([Loc.tag @@ GRef (bigQ_z, None)], uninterp_bigQ,
true)
diff --git a/plugins/syntax/r_syntax.ml b/plugins/syntax/r_syntax.ml
index 8f065f528..1af3f6c5b 100644
--- a/plugins/syntax/r_syntax.ml
+++ b/plugins/syntax/r_syntax.ml
@@ -42,13 +42,13 @@ let glob_xO = ConstructRef path_of_xO
let glob_xH = ConstructRef path_of_xH
let pos_of_bignat dloc x =
- let ref_xI = GRef (dloc, glob_xI, None) in
- let ref_xH = GRef (dloc, glob_xH, None) in
- let ref_xO = GRef (dloc, glob_xO, None) in
+ let ref_xI = Loc.tag @@ GRef (glob_xI, None) in
+ let ref_xH = Loc.tag @@ GRef (glob_xH, None) in
+ let ref_xO = Loc.tag @@ GRef (glob_xO, None) in
let rec pos_of x =
match div2_with_rest x with
- | (q,false) -> GApp (dloc, ref_xO,[pos_of q])
- | (q,true) when not (Bigint.equal q zero) -> GApp (dloc,ref_xI,[pos_of q])
+ | (q,false) -> Loc.tag @@ GApp (ref_xO,[pos_of q])
+ | (q,true) when not (Bigint.equal q zero) -> Loc.tag @@ GApp (ref_xI,[pos_of q])
| (q,true) -> ref_xH
in
pos_of x
@@ -58,9 +58,9 @@ let pos_of_bignat dloc x =
(**********************************************************************)
let rec bignat_of_pos = function
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a)
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a))
- | GRef (_, a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a)
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a))
+ | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
| _ -> raise Non_closed_number
(**********************************************************************)
@@ -81,18 +81,18 @@ let z_of_int dloc n =
if not (Bigint.equal n zero) then
let sgn, n =
if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in
- GApp(dloc, GRef (dloc,sgn,None), [pos_of_bignat dloc n])
+ Loc.tag @@ GApp(Loc.tag @@ GRef (sgn,None), [pos_of_bignat dloc n])
else
- GRef (dloc, glob_ZERO, None)
+ Loc.tag @@ GRef (glob_ZERO, None)
(**********************************************************************)
(* Printing Z via scopes *)
(**********************************************************************)
let bigint_of_z = function
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a)
- | GRef (_, a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a)
+ | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
| _ -> raise Non_closed_number
(**********************************************************************)
@@ -108,14 +108,14 @@ let make_path dir id = Globnames.encode_con dir (Id.of_string id)
let glob_IZR = ConstRef (make_path (make_dir rdefinitions) "IZR")
let r_of_int dloc z =
- GApp (dloc, GRef(dloc,glob_IZR,None), [z_of_int dloc z])
+ Loc.tag @@ GApp (Loc.tag @@ GRef(glob_IZR,None), [z_of_int dloc z])
(**********************************************************************)
(* Printing R via scopes *)
(**********************************************************************)
let bigint_of_r = function
- | GApp (_,GRef (_,o,_), [a]) when Globnames.eq_gr o glob_IZR ->
+ | _, GApp ((_, GRef (o,_)), [a]) when Globnames.eq_gr o glob_IZR ->
bigint_of_z a
| _ -> raise Non_closed_number
@@ -128,6 +128,6 @@ let uninterp_r p =
let _ = Notation.declare_numeral_interpreter "R_scope"
(r_path,["Coq";"Reals";"Rdefinitions"])
r_of_int
- ([GRef (Loc.ghost,glob_IZR,None)],
+ ([Loc.tag @@ GRef (glob_IZR,None)],
uninterp_r,
false)
diff --git a/plugins/syntax/string_syntax.ml b/plugins/syntax/string_syntax.ml
index de0fa77ef..539670722 100644
--- a/plugins/syntax/string_syntax.ml
+++ b/plugins/syntax/string_syntax.ml
@@ -33,23 +33,23 @@ let glob_EmptyString = lazy (make_reference "EmptyString")
open Lazy
-let interp_string dloc s =
+let interp_string loc s =
let le = String.length s in
let rec aux n =
- if n = le then GRef (dloc, force glob_EmptyString, None) else
- GApp (dloc,GRef (dloc, force glob_String, None),
- [interp_ascii dloc (int_of_char s.[n]); aux (n+1)])
+ if n = le then Loc.tag ~loc @@ GRef (force glob_EmptyString, None) else
+ Loc.tag ~loc @@ GApp (Loc.tag ~loc @@ GRef (force glob_String, None),
+ [interp_ascii loc (int_of_char s.[n]); aux (n+1)])
in aux 0
let uninterp_string r =
try
let b = Buffer.create 16 in
let rec aux = function
- | GApp (_,GRef (_,k,_),[a;s]) when eq_gr k (force glob_String) ->
+ | _, GApp ((_, GRef (k,_)),[a;s]) when eq_gr k (force glob_String) ->
(match uninterp_ascii a with
| Some c -> Buffer.add_char b (Char.chr c); aux s
| _ -> raise Non_closed_string)
- | GRef (_,z,_) when eq_gr z (force glob_EmptyString) ->
+ | _, GRef (z,_) when eq_gr z (force glob_EmptyString) ->
Some (Buffer.contents b)
| _ ->
raise Non_closed_string
@@ -61,6 +61,6 @@ let _ =
Notation.declare_string_interpreter "string_scope"
(string_path,["Coq";"Strings";"String"])
interp_string
- ([GRef (Loc.ghost,static_glob_String,None);
- GRef (Loc.ghost,static_glob_EmptyString,None)],
+ ([Loc.tag @@ GRef (static_glob_String,None);
+ Loc.tag @@ GRef (static_glob_EmptyString,None)],
uninterp_string, true)
diff --git a/plugins/syntax/z_syntax.ml b/plugins/syntax/z_syntax.ml
index b7b5fb8a5..a00525f91 100644
--- a/plugins/syntax/z_syntax.ml
+++ b/plugins/syntax/z_syntax.ml
@@ -44,14 +44,14 @@ let glob_xI = ConstructRef path_of_xI
let glob_xO = ConstructRef path_of_xO
let glob_xH = ConstructRef path_of_xH
-let pos_of_bignat dloc x =
- let ref_xI = GRef (dloc, glob_xI, None) in
- let ref_xH = GRef (dloc, glob_xH, None) in
- let ref_xO = GRef (dloc, glob_xO, None) in
+let pos_of_bignat loc x =
+ let ref_xI = Loc.tag ~loc @@ GRef (glob_xI, None) in
+ let ref_xH = Loc.tag ~loc @@ GRef (glob_xH, None) in
+ let ref_xO = Loc.tag ~loc @@ GRef (glob_xO, None) in
let rec pos_of x =
match div2_with_rest x with
- | (q,false) -> GApp (dloc, ref_xO,[pos_of q])
- | (q,true) when not (Bigint.equal q zero) -> GApp (dloc,ref_xI,[pos_of q])
+ | (q,false) -> Loc.tag ~loc @@ GApp (ref_xO,[pos_of q])
+ | (q,true) when not (Bigint.equal q zero) -> Loc.tag ~loc @@ GApp (ref_xI,[pos_of q])
| (q,true) -> ref_xH
in
pos_of x
@@ -69,9 +69,9 @@ let interp_positive dloc n =
(**********************************************************************)
let rec bignat_of_pos = function
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a)
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a))
- | GRef (_, a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xO -> mult_2(bignat_of_pos a)
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_xI -> add_1(mult_2(bignat_of_pos a))
+ | _, GRef (a, _) when Globnames.eq_gr a glob_xH -> Bigint.one
| _ -> raise Non_closed_number
let uninterp_positive p =
@@ -87,9 +87,9 @@ let uninterp_positive p =
let _ = Notation.declare_numeral_interpreter "positive_scope"
(positive_path,binnums)
interp_positive
- ([GRef (Loc.ghost, glob_xI, None);
- GRef (Loc.ghost, glob_xO, None);
- GRef (Loc.ghost, glob_xH, None)],
+ ([Loc.tag @@ GRef (glob_xI, None);
+ Loc.tag @@ GRef (glob_xO, None);
+ Loc.tag @@ GRef (glob_xH, None)],
uninterp_positive,
true)
@@ -106,11 +106,11 @@ let glob_Npos = ConstructRef path_of_Npos
let n_path = make_path binnums "N"
-let n_of_binnat dloc pos_or_neg n =
+let n_of_binnat loc pos_or_neg n = Loc.tag ~loc @@
if not (Bigint.equal n zero) then
- GApp(dloc, GRef (dloc,glob_Npos,None), [pos_of_bignat dloc n])
+ GApp(Loc.tag @@ GRef (glob_Npos,None), [pos_of_bignat loc n])
else
- GRef (dloc, glob_N0, None)
+ GRef(glob_N0, None)
let error_negative dloc =
user_err ~loc:dloc ~hdr:"interp_N" (str "No negative numbers in type \"N\".")
@@ -124,8 +124,8 @@ let n_of_int dloc n =
(**********************************************************************)
let bignat_of_n = function
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a
- | GRef (_, a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_Npos -> bignat_of_pos a
+ | _, GRef (a,_) when Globnames.eq_gr a glob_N0 -> Bigint.zero
| _ -> raise Non_closed_number
let uninterp_n p =
@@ -138,8 +138,8 @@ let uninterp_n p =
let _ = Notation.declare_numeral_interpreter "N_scope"
(n_path,binnums)
n_of_int
- ([GRef (Loc.ghost, glob_N0, None);
- GRef (Loc.ghost, glob_Npos, None)],
+ ([Loc.tag @@ GRef (glob_N0, None);
+ Loc.tag @@ GRef (glob_Npos, None)],
uninterp_n,
true)
@@ -157,22 +157,22 @@ let glob_ZERO = ConstructRef path_of_ZERO
let glob_POS = ConstructRef path_of_POS
let glob_NEG = ConstructRef path_of_NEG
-let z_of_int dloc n =
+let z_of_int loc n =
if not (Bigint.equal n zero) then
let sgn, n =
if is_pos_or_zero n then glob_POS, n else glob_NEG, Bigint.neg n in
- GApp(dloc, GRef (dloc,sgn,None), [pos_of_bignat dloc n])
+ Loc.tag ~loc @@ GApp(Loc.tag ~loc @@ GRef(sgn,None), [pos_of_bignat loc n])
else
- GRef (dloc, glob_ZERO, None)
+ Loc.tag ~loc @@ GRef(glob_ZERO, None)
(**********************************************************************)
(* Printing Z via scopes *)
(**********************************************************************)
let bigint_of_z = function
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a
- | GApp (_, GRef (_,b,_),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a)
- | GRef (_, a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_POS -> bignat_of_pos a
+ | _, GApp ((_, GRef (b,_)),[a]) when Globnames.eq_gr b glob_NEG -> Bigint.neg (bignat_of_pos a)
+ | _, GRef (a, _) when Globnames.eq_gr a glob_ZERO -> Bigint.zero
| _ -> raise Non_closed_number
let uninterp_z p =
@@ -186,8 +186,8 @@ let uninterp_z p =
let _ = Notation.declare_numeral_interpreter "Z_scope"
(z_path,binnums)
z_of_int
- ([GRef (Loc.ghost, glob_ZERO, None);
- GRef (Loc.ghost, glob_POS, None);
- GRef (Loc.ghost, glob_NEG, None)],
+ ([Loc.tag @@ GRef (glob_ZERO, None);
+ Loc.tag @@ GRef (glob_POS, None);
+ Loc.tag @@ GRef (glob_NEG, None)],
uninterp_z,
true)