aboutsummaryrefslogtreecommitdiffhomepage
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
parentbe83b52cf50ed4c596e40cfd52da03258a7a4a18 (diff)
[location] Switch glob_constr to Loc.located
-rw-r--r--dev/top_printers.ml2
-rw-r--r--interp/constrextern.ml165
-rw-r--r--interp/constrintern.ml147
-rw-r--r--interp/implicit_quantifiers.ml33
-rw-r--r--interp/notation.ml23
-rw-r--r--interp/notation_ops.ml289
-rw-r--r--intf/glob_term.mli40
-rw-r--r--lib/loc.ml1
-rw-r--r--lib/loc.mli1
-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
-rw-r--r--pretyping/cases.ml26
-rw-r--r--pretyping/detyping.ml208
-rw-r--r--pretyping/detyping.mli6
-rw-r--r--pretyping/glob_ops.ml268
-rw-r--r--pretyping/patternops.ml45
-rw-r--r--pretyping/pretyping.ml36
-rw-r--r--tactics/hipattern.ml20
-rw-r--r--vernac/command.ml14
38 files changed, 1082 insertions, 1470 deletions
diff --git a/dev/top_printers.ml b/dev/top_printers.ml
index 474cc85c1..52cf8cf97 100644
--- a/dev/top_printers.ml
+++ b/dev/top_printers.ml
@@ -26,7 +26,7 @@ open Clenv
let _ = Detyping.print_evar_arguments := true
let _ = Detyping.print_universes := true
let _ = set_bool_option_value ["Printing";"Matching"] false
-let _ = Detyping.set_detype_anonymous (fun _ _ -> raise Not_found)
+let _ = Detyping.set_detype_anonymous (fun ?loc _ -> raise Not_found)
(* std_ppcmds *)
let pp x = Pp.pp_with !Topfmt.std_ft x
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index d45f3a9f1..bbc98dd28 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -157,7 +157,7 @@ let insert_pat_alias loc p = function
(**********************************************************************)
(* conversion of references *)
-let extern_evar loc n l = Loc.tag @@ CEvar (n,l)
+let extern_evar n l = CEvar (n,l)
(** We allow customization of the global_reference printer.
For instance, in the debugger the tables of global references
@@ -475,7 +475,7 @@ exception Expl
(* Implicit args indexes are in ascending order *)
(* inctx is useful only if there is a last argument to be deduced from ctxt *)
-let explicitize loc inctx impl (cf,f) args =
+let explicitize inctx impl (cf,f) args =
let impl = if !Constrintern.parsing_explicit then [] else impl in
let n = List.length args in
let rec exprec q = function
@@ -512,41 +512,41 @@ let explicitize loc inctx impl (cf,f) args =
let args1 = exprec 1 (args1,impl1) in
let args2 = exprec (i+1) (args2,impl2) in
let ip = Some (List.length args1) in
- Loc.tag ~loc @@ CApp ((ip,f),args1@args2)
+ CApp ((ip,f),args1@args2)
| None ->
let args = exprec 1 (args,impl) in
- if List.is_empty args then f else Loc.tag ~loc @@ CApp ((None, f), args)
+ if List.is_empty args then snd f else CApp ((None, f), args)
in
try expl ()
with Expl ->
- let f',us = match f with _loc, CRef (f,us) -> f,us | _ -> assert false in
+ let f',us = match f with (_loc, CRef (f,us)) -> f,us | _ -> assert false in
let ip = if !print_projections then ip else None in
- Loc.tag ~loc @@ CAppExpl ((ip, f', us), List.map Lazy.force args)
+ CAppExpl ((ip, f', us), List.map Lazy.force args)
let is_start_implicit = function
| imp :: _ -> is_status_implicit imp && maximal_insertion_of imp
| [] -> false
-let extern_global loc impl f us =
+let extern_global impl f us =
if not !Constrintern.parsing_explicit && is_start_implicit impl
then
- Loc.tag ~loc @@ CAppExpl ((None, f, us), [])
+ CAppExpl ((None, f, us), [])
else
- Loc.tag ~loc @@ CRef (f,us)
+ CRef (f,us)
-let extern_app loc inctx impl (cf,f) us args =
+let extern_app inctx impl (cf,f) us args =
if List.is_empty args then
(* If coming from a notation "Notation a := @b" *)
- Loc.tag ~loc @@ CAppExpl ((None, f, us), [])
+ CAppExpl ((None, f, us), [])
else if not !Constrintern.parsing_explicit &&
((!Flags.raw_print ||
(!print_implicits && not !print_implicits_explicit_args)) &&
List.exists is_status_implicit impl)
then
let args = List.map Lazy.force args in
- Loc.tag ~loc @@ CAppExpl ((is_projection (List.length args) cf,f,us), args)
+ CAppExpl ((is_projection (List.length args) cf,f,us), args)
else
- explicitize loc inctx impl (cf, Loc.tag ~loc @@ CRef (f,us)) args
+ explicitize inctx impl (cf, Loc.tag @@ CRef (f,us)) args
let rec fill_arg_scopes args subscopes scopes = match args, subscopes with
| [], _ -> []
@@ -560,7 +560,7 @@ let extern_args extern env args =
List.map map args
let match_coercion_app = function
- | GApp (loc,GRef (_,r,_),args) -> Some (loc, r, 0, args)
+ | (loc, GApp ((_, GRef (r,_)),args)) -> Some (loc, r, 0, args)
| _ -> None
let rec remove_coercions inctx c =
@@ -582,13 +582,13 @@ let rec remove_coercions inctx c =
been confused with ordinary application or would have need
a surrounding context and the coercion to funclass would
have been made explicit to match *)
- if List.is_empty l then a' else GApp (loc,a',l)
+ if List.is_empty l then a' else Loc.tag ~loc @@ GApp (a',l)
| _ -> c
with Not_found -> c)
| _ -> c
let rec flatten_application = function
- | GApp (loc,GApp(_,a,l'),l) -> flatten_application (GApp (loc,a,l'@l))
+ | (loc, GApp ((_, GApp(a,l')),l)) -> flatten_application (Loc.tag ~loc @@ GApp (a,l'@l))
| a -> a
(**********************************************************************)
@@ -616,7 +616,7 @@ let extern_optimal_prim_token scopes r r' =
let extended_glob_local_binder_of_decl loc = function
| (p,bk,None,t) -> GLocalAssum (loc,p,bk,t)
- | (p,bk,Some x,GHole (_, _, Misctypes.IntroAnonymous, None)) -> GLocalDef (loc,p,bk,x,None)
+ | (p,bk,Some x,(_,GHole ( _, Misctypes.IntroAnonymous, None))) -> GLocalDef (loc,p,bk,x,None)
| (p,bk,Some x,t) -> GLocalDef (loc,p,bk,x,Some t)
(**********************************************************************)
@@ -642,25 +642,25 @@ let rec extern inctx scopes vars r =
let r'' = flatten_application r' in
if !Flags.raw_print || !print_no_symbol then raise No_match;
extern_notation scopes vars r'' (uninterp_notations r'')
- with No_match -> match r' with
- | GRef (loc,ref,us) ->
- extern_global loc (select_stronger_impargs (implicits_of_global ref))
+ with No_match -> Loc.map_with_loc (fun ~loc -> function
+ | GRef (ref,us) ->
+ extern_global (select_stronger_impargs (implicits_of_global ref))
(extern_reference loc vars ref) (extern_universes us)
- | GVar (loc,id) -> Loc.tag ~loc @@ CRef (Ident (loc,id),None)
+ | GVar id -> CRef (Ident (loc,id),None)
- | GEvar (loc,n,[]) when !print_meta_as_hole -> Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None)
+ | GEvar (n,[]) when !print_meta_as_hole -> CHole (None, Misctypes.IntroAnonymous, None)
- | GEvar (loc,n,l) ->
- extern_evar loc n (List.map (on_snd (extern false scopes vars)) l)
+ | GEvar (n,l) ->
+ extern_evar n (List.map (on_snd (extern false scopes vars)) l)
- | GPatVar (loc,(b,n)) -> Loc.tag ~loc @@
+ | GPatVar (b,n) ->
if !print_meta_as_hole then CHole (None, Misctypes.IntroAnonymous, None) else
if b then CPatVar n else CEvar (n,[])
- | GApp (loc,f,args) ->
+ | GApp (f,args) ->
(match f with
- | GRef (rloc,ref,us) ->
+ | (rloc, GRef (ref,us)) ->
let subscopes = find_arguments_scope ref in
let args = fill_arg_scopes args subscopes (snd scopes) in
begin
@@ -701,42 +701,42 @@ let rec extern inctx scopes vars r =
let head = extern true scopes vars arg in
ip q locs' tail ((extern_reference loc Id.Set.empty (ConstRef c), head) :: acc)
in
- Loc.tag ~loc @@ CRecord (List.rev (ip projs locals args []))
+ CRecord (List.rev (ip projs locals args []))
with
| Not_found | No_match | Exit ->
let args = extern_args (extern true) vars args in
- extern_app loc inctx
+ extern_app inctx
(select_stronger_impargs (implicits_of_global ref))
(Some ref,extern_reference rloc vars ref) (extern_universes us) args
end
-
+
| _ ->
- explicitize loc inctx [] (None,sub_extern false scopes vars f)
+ explicitize inctx [] (None,sub_extern false scopes vars f)
(List.map (fun c -> lazy (sub_extern true scopes vars c)) args))
- | GLetIn (loc,na,b,t,c) ->
- Loc.tag ~loc @@ CLetIn ((loc,na),sub_extern false scopes vars b,
+ | GLetIn (na,b,t,c) ->
+ CLetIn ((loc,na),sub_extern false scopes vars b,
Option.map (extern_typ scopes vars) t,
extern inctx scopes (add_vname vars na) c)
- | GProd (loc,na,bk,t,c) ->
+ | GProd (na,bk,t,c) ->
let t = extern_typ scopes vars t in
let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in
- Loc.tag ~loc @@ CProdN ([(Loc.ghost,na)::idl,Default bk,t],c)
+ CProdN ([(Loc.ghost,na)::idl,Default bk,t],c)
- | GLambda (loc,na,bk,t,c) ->
+ | GLambda (na,bk,t,c) ->
let t = extern_typ scopes vars t in
let (idl,c) = factorize_lambda inctx scopes (add_vname vars na) na bk t c in
- Loc.tag ~loc @@ CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c)
+ CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c)
- | GCases (loc,sty,rtntypopt,tml,eqns) ->
+ | GCases (sty,rtntypopt,tml,eqns) ->
let vars' =
List.fold_right (name_fold Id.Set.add)
(cases_predicate_names tml) vars in
let rtntypopt' = Option.map (extern_typ scopes vars') rtntypopt in
let tml = List.map (fun (tm,(na,x)) ->
let na' = match na,tm with
- | Anonymous, GVar (_, id) ->
+ | Anonymous, (_, GVar id) ->
begin match rtntypopt with
| None -> None
| Some ntn ->
@@ -745,7 +745,7 @@ let rec extern inctx scopes vars r =
else None
end
| Anonymous, _ -> None
- | Name id, GVar (_,id') when Id.equal id id' -> None
+ | Name id, (_, GVar id') when Id.equal id id' -> None
| Name _, _ -> Some (Loc.ghost,na) in
(sub_extern false scopes vars tm,
na',
@@ -757,22 +757,22 @@ let rec extern inctx scopes vars r =
tml
in
let eqns = List.map (extern_eqn inctx scopes vars) eqns in
- Loc.tag ~loc @@ CCases (sty,rtntypopt',tml,eqns)
+ CCases (sty,rtntypopt',tml,eqns)
- | GLetTuple (loc,nal,(na,typopt),tm,b) ->
- Loc.tag ~loc @@ CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal,
+ | GLetTuple (nal,(na,typopt),tm,b) ->
+ CLetTuple (List.map (fun na -> (Loc.ghost,na)) nal,
(Option.map (fun _ -> (Loc.ghost,na)) typopt,
Option.map (extern_typ scopes (add_vname vars na)) typopt),
sub_extern false scopes vars tm,
extern inctx scopes (List.fold_left add_vname vars nal) b)
- | GIf (loc,c,(na,typopt),b1,b2) ->
- Loc.tag ~loc @@ CIf (sub_extern false scopes vars c,
+ | GIf (c,(na,typopt),b1,b2) ->
+ CIf (sub_extern false scopes vars c,
(Option.map (fun _ -> (Loc.ghost,na)) typopt,
Option.map (extern_typ scopes (add_vname vars na)) typopt),
sub_extern inctx scopes vars b1, sub_extern inctx scopes vars b2)
- | GRec (loc,fk,idv,blv,tyv,bv) ->
+ | GRec (fk,idv,blv,tyv,bv) ->
let vars' = Array.fold_right Id.Set.add idv vars in
(match fk with
| GFix (nv,n) ->
@@ -792,7 +792,7 @@ let rec extern inctx scopes vars r =
((Loc.ghost, fi), (n, ro), bl, extern_typ scopes vars0 ty,
extern false scopes vars1 def)) idv
in
- Loc.tag ~loc @@ CFix ((loc,idv.(n)),Array.to_list listdecl)
+ CFix ((loc,idv.(n)),Array.to_list listdecl)
| GCoFix n ->
let listdecl =
Array.mapi (fun i fi ->
@@ -803,15 +803,16 @@ let rec extern inctx scopes vars r =
((Loc.ghost, fi),bl,extern_typ scopes vars0 tyv.(i),
sub_extern false scopes vars1 bv.(i))) idv
in
- Loc.tag ~loc @@ CCoFix ((loc,idv.(n)),Array.to_list listdecl))
+ CCoFix ((loc,idv.(n)),Array.to_list listdecl))
- | GSort (loc,s) -> Loc.tag ~loc @@ CSort (extern_glob_sort s)
+ | GSort s -> CSort (extern_glob_sort s)
- | GHole (loc,e,naming,_) -> Loc.tag ~loc @@ CHole (Some e, naming, None) (** TODO: extern tactics. *)
+ | GHole (e,naming,_) -> CHole (Some e, naming, None) (** TODO: extern tactics. *)
- | GCast (loc,c, c') ->
- Loc.tag ~loc @@ CCast (sub_extern true scopes vars c,
+ | GCast (c, c') ->
+ CCast (sub_extern true scopes vars c,
Miscops.map_cast_type (extern_typ scopes vars) c')
+ ) r'
and extern_typ (_,scopes) =
extern true (Notation.current_type_scope_name (),scopes)
@@ -867,7 +868,7 @@ and extern_local_binder scopes vars = function
let (assums,ids,l) = extern_local_binder scopes vars l in
(assums,ids, CLocalPattern(Loc.ghost,p,ty) :: l)
-and extern_eqn inctx scopes vars (loc,ids,pl,c) =
+and extern_eqn inctx scopes vars (loc,(ids,pl,c)) =
Loc.tag ~loc ([loc,List.map (extern_cases_pattern_in_scope scopes vars) pl],
extern inctx scopes vars c)
@@ -878,13 +879,13 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
try
if List.mem keyrule !print_non_active_notations then raise No_match;
(* Adjusts to the number of arguments expected by the notation *)
- let (t,args,argsscopes,argsimpls) = match t,n with
- | GApp (_,f,args), Some n
+ let (t,args,argsscopes,argsimpls) = match snd t,n with
+ | GApp (f,args), Some n
when List.length args >= n ->
let args1, args2 = List.chop n args in
let subscopes, impls =
- match f with
- | GRef (_,ref,us) ->
+ match snd f with
+ | GRef (ref,us) ->
let subscopes =
try List.skipn n (find_arguments_scope ref)
with Failure _ -> [] in
@@ -896,15 +897,15 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
subscopes,impls
| _ ->
[], [] in
- (if Int.equal n 0 then f else GApp (Loc.ghost,f,args1)),
+ (if Int.equal n 0 then f else Loc.tag @@ GApp (f,args1)),
args2, subscopes, impls
- | GApp (_,(GRef (_,ref,us) as f),args), None ->
+ | GApp ((_, GRef (ref,us) as f),args), None ->
let subscopes = find_arguments_scope ref in
let impls =
select_impargs_size
(List.length args) (implicits_of_global ref) in
f, args, subscopes, impls
- | GRef (_,ref,us), Some 0 -> GApp (Loc.ghost,t,[]), [], [], []
+ | GRef (ref,us), Some 0 -> Loc.tag @@ GApp (t,[]), [], [], []
| _, None -> t, [], [], []
| _ -> raise No_match in
(* Try matching ... *)
@@ -945,7 +946,7 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
else
let args = fill_arg_scopes args argsscopes scopes in
let args = extern_args (extern true) vars args in
- explicitize loc false argsimpls (None,e) args
+ Loc.tag ~loc @@ explicitize false argsimpls (None,e) args
with
No_match -> extern_notation allscopes vars t rules
@@ -965,8 +966,6 @@ let extern_glob_type vars c =
(******************************************************************)
(* Main translation function from constr -> constr_expr *)
-let loc = Loc.ghost (* for constr and pattern, locations are lost *)
-
let extern_constr_gen lax goal_concl_style scopt env sigma t =
(* "goal_concl_style" means do alpha-conversion using the "goal" convention *)
(* i.e.: avoid using the names of goal/section/rel variables and the short *)
@@ -1008,11 +1007,11 @@ let extern_closed_glob ?lax goal_concl_style env sigma t =
let any_any_branch =
(* | _ => _ *)
- (loc,[],[Loc.tag ~loc @@ PatVar Anonymous],GHole (loc,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None))
+ Loc.tag ([],[Loc.tag @@ PatVar Anonymous], Loc.tag @@ GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None))
-let rec glob_of_pat env sigma = function
- | PRef ref -> GRef (loc,ref,None)
- | PVar id -> GVar (loc,id)
+let rec glob_of_pat env sigma pat = Loc.tag @@ match pat with
+ | PRef ref -> GRef (ref,None)
+ | PVar id -> GVar id
| PEvar (evk,l) ->
let test decl = function PVar id' -> Id.equal (NamedDecl.get_id decl) id' | _ -> false in
let l = Evd.evar_instance_array test (Evd.find sigma evk) l in
@@ -1020,36 +1019,36 @@ let rec glob_of_pat env sigma = function
| None -> Id.of_string "__"
| Some id -> id
in
- GEvar (loc,id,List.map (on_snd (glob_of_pat env sigma)) l)
+ GEvar (id,List.map (on_snd (glob_of_pat env sigma)) l)
| PRel n ->
let id = try match lookup_name_of_rel n env with
| Name id -> id
| Anonymous ->
anomaly ~label:"glob_constr_of_pattern" (Pp.str "index to an anonymous variable")
with Not_found -> Id.of_string ("_UNBOUND_REL_"^(string_of_int n)) in
- GVar (loc,id)
- | PMeta None -> GHole (loc,Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None)
- | PMeta (Some n) -> GPatVar (loc,(false,n))
- | PProj (p,c) -> GApp (loc,GRef (loc, ConstRef (Projection.constant p),None),
+ GVar id
+ | PMeta None -> GHole (Evar_kinds.InternalHole, Misctypes.IntroAnonymous,None)
+ | PMeta (Some n) -> GPatVar (false,n)
+ | PProj (p,c) -> GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p),None),
[glob_of_pat env sigma c])
| PApp (f,args) ->
- GApp (loc,glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args)
+ GApp (glob_of_pat env sigma f,Array.map_to_list (glob_of_pat env sigma) args)
| PSoApp (n,args) ->
- GApp (loc,GPatVar (loc,(true,n)),
+ GApp (Loc.tag @@ GPatVar (true,n),
List.map (glob_of_pat env sigma) args)
| PProd (na,t,c) ->
- GProd (loc,na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c)
+ GProd (na,Explicit,glob_of_pat env sigma t,glob_of_pat (na::env) sigma c)
| PLetIn (na,b,t,c) ->
- GLetIn (loc,na,glob_of_pat env sigma b, Option.map (glob_of_pat env sigma) t,
+ GLetIn (na,glob_of_pat env sigma b, Option.map (glob_of_pat env sigma) t,
glob_of_pat (na::env) sigma c)
| PLambda (na,t,c) ->
- GLambda (loc,na,Explicit,glob_of_pat env sigma t, glob_of_pat (na::env) sigma c)
+ GLambda (na,Explicit,glob_of_pat env sigma t, glob_of_pat (na::env) sigma c)
| PIf (c,b1,b2) ->
- GIf (loc, glob_of_pat env sigma c, (Anonymous,None),
+ GIf (glob_of_pat env sigma c, (Anonymous,None),
glob_of_pat env sigma b1, glob_of_pat env sigma b2)
| PCase ({cip_style=LetStyle; cip_ind_tags=None},PMeta None,tm,[(0,n,b)]) ->
let nal,b = it_destRLambda_or_LetIn_names n (glob_of_pat env sigma b) in
- GLetTuple (loc,nal,(Anonymous,None),glob_of_pat env sigma tm,b)
+ GLetTuple (nal,(Anonymous,None),glob_of_pat env sigma tm,b)
| PCase (info,p,tm,bl) ->
let mat = match bl, info.cip_ind with
| [], _ -> []
@@ -1066,10 +1065,10 @@ let rec glob_of_pat env sigma = function
return_type_of_predicate ind nargs (glob_of_pat env sigma p)
| _ -> anomaly (Pp.str "PCase with non-trivial predicate but unknown inductive")
in
- GCases (loc,RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat)
- | PFix f -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *)
- | PCoFix c -> Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c))
- | PSort s -> GSort (loc,s)
+ GCases (RegularStyle,rtn,[glob_of_pat env sigma tm,indnames],mat)
+ | PFix f -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkFix f)) (** FIXME bad env *)
+ | PCoFix c -> Loc.obj @@ Detyping.detype_names false [] env (Global.env()) sigma (EConstr.of_constr (mkCoFix c))
+ | PSort s -> GSort s
let extern_constr_pattern env sigma pat =
extern true (None,[]) Id.Set.empty (glob_of_pat env sigma pat)
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index f814205dc..cc7203ac0 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -304,12 +304,12 @@ let reset_tmp_scope env = {env with tmp_scope = None}
let rec it_mkGProd loc2 env body =
match env with
- (loc1, (na, bk, t)) :: tl -> it_mkGProd loc2 tl (GProd (Loc.merge loc1 loc2, na, bk, t, body))
+ (loc1, (na, bk, t)) :: tl -> it_mkGProd loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GProd (na, bk, t, body))
| [] -> body
let rec it_mkGLambda loc2 env body =
match env with
- (loc1, (na, bk, t)) :: tl -> it_mkGLambda loc2 tl (GLambda (Loc.merge loc1 loc2, na, bk, t, body))
+ (loc1, (na, bk, t)) :: tl -> it_mkGLambda loc2 tl (Loc.tag ~loc:(Loc.merge loc1 loc2) @@ GLambda (na, bk, t, body))
| [] -> body
(**********************************************************************)
@@ -322,14 +322,14 @@ let build_impls = function
let impls_type_list ?(args = []) =
let rec aux acc = function
- |GProd (_,na,bk,_,c) -> aux ((build_impls bk na)::acc) c
+ |_, GProd (na,bk,_,c) -> aux ((build_impls bk na)::acc) c
|_ -> (Variable,[],List.append args (List.rev acc),[])
in aux []
let impls_term_list ?(args = []) =
let rec aux acc = function
- |GLambda (_,na,bk,_,c) -> aux ((build_impls bk na)::acc) c
- |GRec (_, fix_kind, nas, args, tys, bds) ->
+ |_, GLambda (na,bk,_,c) -> aux ((build_impls bk na)::acc) c
+ |_, GRec (fix_kind, nas, args, tys, bds) ->
let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in
let acc' = List.fold_left (fun a (na, bk, _, _) -> (build_impls bk na)::a) acc args.(nb) in
aux acc' bds.(nb)
@@ -346,12 +346,12 @@ let rec check_capture ty = function
()
let locate_if_hole loc na = function
- | GHole (_,_,naming,arg) ->
+ | _, GHole (_,naming,arg) ->
(try match na with
| Name id -> glob_constr_of_notation_constr loc
(Reserve.find_reserved_type id)
| Anonymous -> raise Not_found
- with Not_found -> GHole (loc, Evar_kinds.BinderType na, naming, arg))
+ with Not_found -> Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na, naming, arg))
| x -> x
let reset_hidden_inductive_implicit_test env =
@@ -397,7 +397,7 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
env fvs in
let bl = List.map
(fun (id, loc) ->
- (loc, (Name id, b, GHole (loc, Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None))))
+ (loc, (Name id, b, Loc.tag ~loc @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None))))
fvs
in
let na = match na with
@@ -458,7 +458,7 @@ let glob_local_binder_of_extended = function
| GLocalAssum (loc,na,bk,t) -> (na,bk,None,t)
| GLocalDef (loc,na,bk,c,Some t) -> (na,bk,Some c,t)
| GLocalDef (loc,na,bk,c,None) ->
- let t = GHole(loc,Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in
+ let t = Loc.tag ~loc @@ GHole(Evar_kinds.BinderType na,Misctypes.IntroAnonymous,None) in
(na,bk,Some c,t)
| GLocalPattern (loc,_,_,_,_) ->
Loc.raise ~loc (Stream.Error "pattern with quote not allowed here.")
@@ -517,10 +517,12 @@ let intern_generalization intern env lvar loc bk ak c =
in
if pi then
(fun (id, loc') acc ->
- GProd (Loc.merge loc' loc, Name id, bk, GHole (loc', Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc))
+ Loc.tag ~loc:(Loc.merge loc' loc) @@
+ GProd (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc))
else
(fun (id, loc') acc ->
- GLambda (Loc.merge loc' loc, Name id, bk, GHole (loc', Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc))
+ Loc.tag ~loc:(Loc.merge loc' loc) @@
+ GLambda (Name id, bk, Loc.tag ~loc:loc' @@ GHole (Evar_kinds.BinderType (Name id), Misctypes.IntroAnonymous, None), acc))
in
List.fold_right (fun (id, loc as lid) (env, acc) ->
let env' = push_name_env lvar (Variable,[],[],[]) env (loc, Name id) in
@@ -572,10 +574,10 @@ let make_letins =
(fun a c ->
match a with
| LPLetIn (loc,(na,b,t)) ->
- GLetIn(loc,na,b,t,c)
+ Loc.tag ~loc @@ GLetIn(na,b,t,c)
| LPCases (loc,(cp,il),id) ->
- let tt = (GVar(loc,id),(Name id,None)) in
- GCases(loc,Misctypes.LetPatternStyle,None,[tt],[(loc,il,[cp],c)]))
+ let tt = (Loc.tag ~loc @@ GVar id, (Name id,None)) in
+ Loc.tag ~loc @@ GCases(Misctypes.LetPatternStyle,None,[tt],[(loc,(il,[cp],c))]))
let rec subordinate_letins letins = function
(* binders come in reverse order; the non-let are returned in reverse order together *)
@@ -660,7 +662,7 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
let bindings = Id.Map.map mk_env terms in
Some (Genintern.generic_substitute_notation bindings arg)
in
- GHole (loc, knd, naming, arg)
+ Loc.tag ~loc @@ GHole (knd, naming, arg)
| NBinderList (x,y,iter,terminator) ->
(try
(* All elements of the list are in scopes (scopt,subscopes) *)
@@ -678,22 +680,22 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
let a,letins = snd (Option.get binderopt) in
let e = make_letins letins (aux subst' infos c') in
let (loc,(na,bk,t)) = a in
- GProd (loc,na,bk,t,e)
+ Loc.tag ~loc @@ GProd (na,bk,t,e)
| NLambda (Name id,NHole _,c') when option_mem_assoc id binderopt ->
let a,letins = snd (Option.get binderopt) in
let (loc,(na,bk,t)) = a in
- GLambda (loc,na,bk,t,make_letins letins (aux subst' infos c'))
+ Loc.tag ~loc @@ GLambda (na,bk,t,make_letins letins (aux subst' infos c'))
(* Two special cases to keep binder name synchronous with BinderType *)
| NProd (na,NHole(Evar_kinds.BinderType na',naming,arg),c')
when Name.equal na na' ->
let subinfos,na = traverse_binder subst avoid subinfos na in
- let ty = GHole (loc,Evar_kinds.BinderType na,naming,arg) in
- GProd (loc,na,Explicit,ty,aux subst' subinfos c')
+ let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
+ Loc.tag ~loc @@ GProd (na,Explicit,ty,aux subst' subinfos c')
| NLambda (na,NHole(Evar_kinds.BinderType na',naming,arg),c')
when Name.equal na na' ->
let subinfos,na = traverse_binder subst avoid subinfos na in
- let ty = GHole (loc,Evar_kinds.BinderType na,naming,arg) in
- GLambda (loc,na,Explicit,ty,aux subst' subinfos c')
+ let ty = Loc.tag ~loc @@ GHole (Evar_kinds.BinderType na,naming,arg) in
+ Loc.tag ~loc @@ GLambda (na,Explicit,ty,aux subst' subinfos c')
| t ->
glob_constr_of_notation_constr_with_binders loc
(traverse_binder subst avoid) (aux subst') subinfos t
@@ -705,11 +707,12 @@ let instantiate_notation_constr loc intern ntnvars subst infos c =
intern {env with tmp_scope = scopt;
scopes = subscopes @ env.scopes} a
with Not_found ->
+ Loc.tag ~loc (
try
- GVar (loc, Id.Map.find id renaming)
+ GVar (Id.Map.find id renaming)
with Not_found ->
(* Happens for local notation joint with inductive/fixpoint defs *)
- GVar (loc,id)
+ GVar id)
in aux (terms,None,None) infos c
let split_by_type ids =
@@ -744,7 +747,7 @@ let string_of_ty = function
| Variable -> "var"
let gvar (loc, id) us = match us with
-| None -> GVar (loc, id)
+| None -> Loc.tag ~loc @@ GVar id
| Some _ ->
user_err ~loc (str "Variable " ++ pr_id id ++
str " cannot have a universe instance")
@@ -786,25 +789,25 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
let impls = implicits_of_global ref in
let scopes = find_arguments_scope ref in
Dumpglob.dump_reference loc "<>" (string_of_qualid (Decls.variable_secpath id)) "var";
- GRef (loc, ref, us), impls, scopes, []
+ Loc.tag ~loc @@ GRef (ref, us), impls, scopes, []
with e when CErrors.noncritical e ->
(* [id] a goal variable *)
gvar (loc,id) us, [], [], []
let find_appl_head_data c =
- match c with
- | GRef (loc,ref,_) as x ->
+ match Loc.obj c with
+ | GRef (ref,_) ->
let impls = implicits_of_global ref in
let scopes = find_arguments_scope ref in
- x, impls, scopes, []
- | GApp (_,GRef (_,ref,_),l) as x
+ c, impls, scopes, []
+ | GApp ((_, GRef (ref,_)),l)
when l != [] && Flags.version_strictly_greater Flags.V8_2 ->
let n = List.length l in
let impls = implicits_of_global ref in
let scopes = find_arguments_scope ref in
- x, List.map (drop_first_implicits n) impls,
+ c, List.map (drop_first_implicits n) impls,
List.skipn_at_least n scopes,[]
- | x -> x,[],[],[]
+ | _ -> c,[],[],[]
let error_not_enough_arguments loc =
user_err ~loc (str "Abbreviation is not applied enough.")
@@ -836,7 +839,7 @@ let intern_reference ref =
(* Is it a global reference or a syntactic definition? *)
let intern_qualid loc qid intern env lvar us args =
match intern_extended_global_of_qualid (loc,qid) with
- | TrueGlobal ref -> GRef (loc, ref, us), true, args
+ | TrueGlobal ref -> (Loc.tag ~loc @@ GRef (ref, us)), true, args
| SynDef sp ->
let (ids,c) = Syntax_def.search_syntactic_definition sp in
let nids = List.length ids in
@@ -850,9 +853,9 @@ let intern_qualid loc qid intern env lvar us args =
let c = instantiate_notation_constr loc intern lvar subst infos c in
let c = match us, c with
| None, _ -> c
- | Some _, GRef (loc, ref, None) -> GRef (loc, ref, us)
- | Some _, GApp (loc, GRef (loc', ref, None), arg) ->
- GApp (loc, GRef (loc', ref, us), arg)
+ | Some _, (loc, GRef (ref, None)) -> Loc.tag ~loc @@ GRef (ref, us)
+ | Some _, (loc, GApp ((loc', GRef (ref, None)), arg)) ->
+ Loc.tag ~loc @@ GApp (Loc.tag ~loc:loc' @@ GRef (ref, us), arg)
| Some _, _ ->
user_err ~loc (str "Notation " ++ pr_qualid qid ++
str " cannot have a universe instance, its expanded head
@@ -863,7 +866,7 @@ let intern_qualid loc qid intern env lvar us args =
(* Rule out section vars since these should have been found by intern_var *)
let intern_non_secvar_qualid loc qid intern env lvar us args =
match intern_qualid loc qid intern env lvar us args with
- | GRef (_, VarRef _, _),_,_ -> raise Not_found
+ | (_, GRef (VarRef _, _)),_,_ -> raise Not_found
| r -> r
let intern_applied_reference intern env namedctx (_, ntnvars as lvar) us args = function
@@ -1470,8 +1473,8 @@ let get_implicit_name n imps =
Some (Impargs.name_of_implicit (List.nth imps (n-1)))
let set_hole_implicit i b = function
- | GRef (loc,r,_) | GApp (_,GRef (loc,r,_),_) -> (loc,Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None)
- | GVar (loc,id) -> (loc,Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None)
+ | (loc, GRef (r,_)) | (_, GApp ((loc, (GRef (r,_))),_)) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (r,i,b),Misctypes.IntroAnonymous,None)
+ | (loc, GVar id) -> Loc.tag ~loc (Evar_kinds.ImplicitArg (VarRef id,i,b),Misctypes.IntroAnonymous,None)
| _ -> anomaly (Pp.str "Only refs have implicits")
let exists_implicit_name id =
@@ -1558,7 +1561,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
push_name_env ntnvars (impls_type_list ~args:fix_args tyi)
en (Loc.ghost, Name name)) 0 env' lf in
(a,b,c,intern {env'' with tmp_scope = None} bd)) dl idl_temp in
- GRec (loc,GFix
+ Loc.tag ~loc @@
+ GRec (GFix
(Array.map (fun (ro,_,_,_) -> ro) idl,n),
Array.of_list lf,
Array.map (fun (_,bl,_,_) -> bl) idl,
@@ -1584,7 +1588,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
push_name_env ntnvars (impls_type_list ~args:cofix_args tyi)
en (Loc.ghost, Name name)) 0 env' lf in
(b,c,intern {env'' with tmp_scope = None} bd)) dl idl_tmp in
- GRec (loc,GCoFix n,
+ Loc.tag ~loc @@
+ GRec (GCoFix n,
Array.of_list lf,
Array.map (fun (bl,_,_) -> bl) idl,
Array.map (fun (_,ty,_) -> ty) idl,
@@ -1600,7 +1605,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
| CLetIn (na,c1,t,c2) ->
let inc1 = intern (reset_tmp_scope env) c1 in
let int = Option.map (intern_type env) t in
- GLetIn (loc, snd na, inc1, int,
+ Loc.tag ~loc @@
+ GLetIn (snd na, inc1, int,
intern (push_name_env ntnvars (impls_term_list inc1) env na) c2)
| CNotation ("- _",([_, CPrim (Numeral p)],[],[]))
when Bigint.is_strictly_pos p ->
@@ -1622,7 +1628,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
lvar us args ref
in
(* Rem: GApp(_,f,[]) stands for @f *)
- GApp (loc, f, intern_args env args_scopes (List.map fst args))
+ Loc.tag ~loc @@
+ GApp (f, intern_args env args_scopes (List.map fst args))
| CApp ((isproj,f), args) ->
let f,args = match f with
@@ -1687,20 +1694,21 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(* Build a return predicate by expansion of the patterns of the "in" clause *)
let thevars, thepats = List.split l in
let sub_rtn = (* Some (GSort (Loc.ghost,GType None)) *) None in
- let sub_tms = List.map (fun id -> GVar (Loc.ghost,id),(Name id,None)) thevars (* "match v1,..,vn" *) in
- let main_sub_eqn =
- (Loc.ghost,[],thepats, (* "|p1,..,pn" *)
+ let sub_tms = List.map (fun id -> (Loc.tag @@ GVar id),(Name id,None)) thevars (* "match v1,..,vn" *) in
+ let main_sub_eqn = Loc.tag @@
+ ([],thepats, (* "|p1,..,pn" *)
Option.cata (intern_type env')
- (GHole(Loc.ghost,Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None))
+ (Loc.tag ~loc @@ GHole(Evar_kinds.CasesType false,Misctypes.IntroAnonymous,None))
rtnpo) (* "=> P" if there were a return predicate P, and "=> _" otherwise *) in
let catch_all_sub_eqn =
if List.for_all (irrefutable globalenv) thepats then [] else
- [Loc.ghost,[],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *)
- GHole(Loc.ghost,Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None)] (* "=> _" *) in
- Some (GCases(Loc.ghost,Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn))
+ [Loc.tag @@ ([],List.make (List.length thepats) (Loc.tag @@ PatVar Anonymous), (* "|_,..,_" *)
+ Loc.tag @@ GHole(Evar_kinds.ImpossibleCase,Misctypes.IntroAnonymous,None))] (* "=> _" *) in
+ Some (Loc.tag @@ GCases(Term.RegularStyle,sub_rtn,sub_tms,main_sub_eqn::catch_all_sub_eqn))
in
let eqns' = List.map (intern_eqn (List.length tms) env) eqns in
- GCases (loc, sty, rtnpo, tms, List.flatten eqns')
+ Loc.tag ~loc @@
+ GCases (sty, rtnpo, tms, List.flatten eqns')
| CLetTuple (nal, (na,po), b, c) ->
let env' = reset_tmp_scope env in
(* "in" is None so no match to add *)
@@ -1709,7 +1717,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env')
(Loc.ghost,na') in
intern_type env'' u) po in
- GLetTuple (loc, List.map snd nal, (na', p'), b',
+ Loc.tag ~loc @@
+ GLetTuple (List.map snd nal, (na', p'), b',
intern (List.fold_left (push_name_env ntnvars (Variable,[],[],[])) (reset_hidden_inductive_implicit_test env) nal) c)
| CIf (c, (na,po), b1, b2) ->
let env' = reset_tmp_scope env in
@@ -1718,7 +1727,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
let env'' = push_name_env ntnvars (Variable,[],[],[]) (reset_hidden_inductive_implicit_test env)
(Loc.ghost,na') in
intern_type env'' p) po in
- GIf (loc, c', (na', p'), intern env b1, intern env b2)
+ Loc.tag ~loc @@
+ GIf (c', (na', p'), intern env b1, intern env b2)
| CHole (k, naming, solve) ->
let k = match k with
| None ->
@@ -1743,23 +1753,29 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
let (_, glb) = Genintern.generic_intern ist gen in
Some glb
in
- GHole (loc, k, naming, solve)
+ Loc.tag ~loc @@
+ GHole (k, naming, solve)
(* Parsing pattern variables *)
| CPatVar n when allow_patvar ->
- GPatVar (loc, (true,n))
+ Loc.tag ~loc @@
+ GPatVar (true,n)
| CEvar (n, []) when allow_patvar ->
- GPatVar (loc, (false,n))
+ Loc.tag ~loc @@
+ GPatVar (false,n)
(* end *)
(* Parsing existential variables *)
| CEvar (n, l) ->
- GEvar (loc, n, List.map (on_snd (intern env)) l)
+ Loc.tag ~loc @@
+ GEvar (n, List.map (on_snd (intern env)) l)
| CPatVar _ ->
raise (InternalizationError (loc,IllegalMetavariable))
(* end *)
| CSort s ->
- GSort(loc,s)
+ Loc.tag ~loc @@
+ GSort s
| CCast (c1, c2) ->
- GCast (loc,intern env c1, Miscops.map_cast_type (intern_type env) c2)
+ Loc.tag ~loc @@
+ GCast (intern env c1, Miscops.map_cast_type (intern_type env) c2)
)
and intern_type env = intern (set_type_scope env)
@@ -1790,15 +1806,15 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
List.map (fun (asubst,pl) ->
let rhs = replace_vars_constr_expr asubst rhs in
let rhs' = intern {env with ids = env_ids} rhs in
- (loc,eqn_ids,pl,rhs')) pll
+ (loc,(eqn_ids,pl,rhs'))) pll
and intern_case_item env forbidden_names_for_gen (tm,na,t) =
(* the "match" part *)
let tm' = intern env tm in
(* the "as" part *)
let extra_id,na = match tm', na with
- | GVar (loc,id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id)
- | GRef (loc, VarRef id, _), None -> Some id,(loc,Name id)
+ | (loc , GVar id), None when not (Id.Map.mem id (snd lvar)) -> Some id,(loc,Name id)
+ | (loc, GRef (VarRef id, _)), None -> Some id,(loc,Name id)
| _, None -> None,(Loc.ghost,Anonymous)
| _, Some (loc,na) -> None,(loc,na) in
(* the "in" part *)
@@ -1870,8 +1886,9 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(* with implicit arguments if maximal insertion is set *)
[]
else
- GHole (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c) ::
- aux (n+1) impl' subscopes' eargs rargs
+ (Loc.map (fun (a,b,c) -> GHole(a,b,c))
+ (set_hole_implicit (n,get_implicit_name n l) (force_inference_of imp) c)
+ ) :: aux (n+1) impl' subscopes' eargs rargs
end
| (imp::impl', a::rargs') ->
intern enva a :: aux (n+1) impl' subscopes' eargs rargs'
@@ -1895,8 +1912,8 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
and smart_gapp f loc = function
| [] -> f
| l -> match f with
- | GApp (loc', g, args) -> GApp (Loc.merge loc' loc, g, args@l)
- | _ -> GApp (Loc.merge (loc_of_glob_constr f) loc, f, l)
+ | (loc', GApp (g, args)) -> Loc.tag ~loc:(Loc.merge loc' loc) @@ GApp (g, args@l)
+ | _ -> Loc.tag ~loc:(Loc.merge (loc_of_glob_constr f) loc) @@ GApp (f, l)
and intern_args env subscopes = function
| [] -> []
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index d2bebfb54..51152bb24 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -125,37 +125,38 @@ let add_name_to_ids set na =
| Name id -> Id.Set.add id set
let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.empty) =
- let rec vars bound vs = function
- | GVar (loc,id) ->
+ let rec vars bound vs (loc, t) = match t with
+ | GVar id ->
if is_freevar bound (Global.env ()) id then
if Id.List.mem_assoc id vs then vs
else (id, loc) :: vs
else vs
- | GApp (loc,f,args) -> List.fold_left (vars bound) vs (f::args)
- | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) ->
+
+ | GApp (f,args) -> List.fold_left (vars bound) vs (f::args)
+ | GLambda (na,_,ty,c) | GProd (na,_,ty,c) ->
let vs' = vars bound vs ty in
let bound' = add_name_to_ids bound na in
vars bound' vs' c
- | GLetIn (loc,na,b,ty,c) ->
+ | GLetIn (na,b,ty,c) ->
let vs' = vars bound vs b in
let vs'' = Option.fold_left (vars bound) vs' ty in
let bound' = add_name_to_ids bound na in
vars bound' vs'' c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ | GCases (sty,rtntypopt,tml,pl) ->
let vs1 = vars_option bound vs rtntypopt in
let vs2 = List.fold_left (fun vs (tm,_) -> vars bound vs tm) vs1 tml in
List.fold_left (vars_pattern bound) vs2 pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (nal,rtntyp,b,c) ->
let vs1 = vars_return_type bound vs rtntyp in
let vs2 = vars bound vs1 b in
let bound' = List.fold_left add_name_to_ids bound nal in
vars bound' vs2 c
- | GIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
let vs1 = vars_return_type bound vs rtntyp in
let vs2 = vars bound vs1 c in
let vs3 = vars bound vs2 b1 in
vars bound vs3 b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (fk,idl,bl,tyl,bv) ->
let bound' = Array.fold_right Id.Set.add idl bound in
let vars_fix i vs fid =
let vs1,bound1 =
@@ -173,11 +174,11 @@ let generalizable_vars_of_glob_constr ?(bound=Id.Set.empty) ?(allowed=Id.Set.emp
vars bound1 vs2 bv.(i)
in
Array.fold_left_i vars_fix vs idl
- | GCast (loc,c,k) -> let v = vars bound vs c in
+ | GCast (c,k) -> let v = vars bound vs c in
(match k with CastConv t | CastVM t -> vars bound v t | _ -> v)
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs
- and vars_pattern bound vs (loc,idl,p,c) =
+ and vars_pattern bound vs (loc,(idl,p,c)) =
let bound' = List.fold_right Id.Set.add idl bound in
vars bound' vs c
@@ -309,12 +310,12 @@ let implicits_of_glob_constr ?(with_products=true) l =
(ExplByPos (i, name), (true, true, true)) :: l
| _ -> l
in
- let rec aux i c =
+ let rec aux i (loc, c) =
let abs na bk b =
add_impl i na bk (aux (succ i) b)
in
match c with
- | GProd (loc, na, bk, t, b) ->
+ | GProd (na, bk, t, b) ->
if with_products then abs na bk b
else
let () = match bk with
@@ -323,9 +324,9 @@ let implicits_of_glob_constr ?(with_products=true) l =
pr_name na ++ strbrk " and following binders")
| _ -> ()
in []
- | GLambda (loc, na, bk, t, b) -> abs na bk b
- | GLetIn (loc, na, b, t, c) -> aux i c
- | GRec (_, fix_kind, nas, args, tys, bds) ->
+ | GLambda (na, bk, t, b) -> abs na bk b
+ | GLetIn (na, b, t, c) -> aux i b
+ | GRec (fix_kind, nas, args, tys, bds) ->
let nb = match fix_kind with |GFix (_, n) -> n | GCoFix n -> n in
List.fold_left_i (fun i l (na,bk,_,_) -> add_impl i na bk l) i (aux (List.length args.(nb) + i) bds.(nb)) args.(nb)
| _ -> []
diff --git a/interp/notation.ml b/interp/notation.ml
index aef089299..3bcec3001 100644
--- a/interp/notation.ml
+++ b/interp/notation.ml
@@ -264,12 +264,12 @@ let notations_key_table = ref (KeyMap.empty : notation_rule list KeyMap.t)
let prim_token_key_table = ref KeyMap.empty
let glob_prim_constr_key = function
- | GApp (_,GRef (_,ref,_),_) | GRef (_,ref,_) -> RefKey (canonical_gr ref)
+ | _, GApp ((_, GRef (ref,_)),_) | _, GRef (ref,_) -> RefKey (canonical_gr ref)
| _ -> Oth
let glob_constr_keys = function
- | GApp (_,GRef (_,ref,_),_) -> [RefKey (canonical_gr ref); Oth]
- | GRef (_,ref,_) -> [RefKey (canonical_gr ref)]
+ | _, GApp ((_, GRef (ref,_)),_) -> [RefKey (canonical_gr ref); Oth]
+ | _, GRef (ref,_) -> [RefKey (canonical_gr ref)]
| _ -> [Oth]
let cases_pattern_key = function
@@ -471,13 +471,14 @@ let interp_prim_token =
(** [rcp_of_glob] : from [glob_constr] to [raw_cases_pattern_expr] *)
-let rec rcp_of_glob looked_for = function
- | GVar (loc,id) -> Loc.tag ~loc @@ RCPatAtom (Some id)
- | GHole (loc,_,_,_) -> Loc.tag ~loc @@ RCPatAtom (None)
- | GRef (loc,g,_) -> looked_for g; Loc.tag ~loc @@ RCPatCstr (g,[],[])
- | GApp (loc,GRef (_,g,_),l) ->
- looked_for g; Loc.tag ~loc @@ RCPatCstr (g, List.map (rcp_of_glob looked_for) l,[])
+let rec rcp_of_glob looked_for gt = Loc.map (function
+ | GVar id -> RCPatAtom (Some id)
+ | GHole (_,_,_) -> RCPatAtom None
+ | GRef (g,_) -> looked_for g; RCPatCstr (g,[],[])
+ | GApp ((_, GRef (g,_)),l) ->
+ looked_for g; RCPatCstr (g, List.map (rcp_of_glob looked_for) l,[])
| _ -> raise Not_found
+ ) gt
let interp_prim_token_cases_pattern_expr loc looked_for p =
interp_prim_token_gen (rcp_of_glob looked_for) loc p
@@ -521,8 +522,8 @@ let uninterp_prim_token_ind_pattern ind args =
if not b then raise Notation_ops.No_match;
let args' = List.map
(fun x -> snd (glob_constr_of_closed_cases_pattern x)) args in
- let ref = GRef (Loc.ghost,ref,None) in
- match numpr (GApp (Loc.ghost,ref,args')) with
+ let ref = Loc.tag @@ GRef (ref,None) in
+ match numpr (Loc.tag @@ GApp (ref,args')) with
| None -> raise Notation_ops.No_match
| Some n -> (sc,n)
with Not_found -> raise Notation_ops.No_match
diff --git a/interp/notation_ops.ml b/interp/notation_ops.ml
index a25fd81f3..32c900504 100644
--- a/interp/notation_ops.ml
+++ b/interp/notation_ops.ml
@@ -24,19 +24,19 @@ open Notation_term
let on_true_do b f c = if b then (f c; b) else b
-let compare_glob_constr f add t1 t2 = match t1,t2 with
- | GRef (_,r1,_), GRef (_,r2,_) -> eq_gr r1 r2
- | GVar (_,v1), GVar (_,v2) -> on_true_do (Id.equal v1 v2) add (Name v1)
- | GApp (_,f1,l1), GApp (_,f2,l2) -> f f1 f2 && List.for_all2eq f l1 l2
- | GLambda (_,na1,bk1,ty1,c1), GLambda (_,na2,bk2,ty2,c2)
+let compare_glob_constr f add (_l1, t1) (_l2, t2) = match t1,t2 with
+ | GRef (r1,_), GRef (r2,_) -> eq_gr r1 r2
+ | GVar v1, GVar v2 -> on_true_do (Id.equal v1 v2) add (Name v1)
+ | GApp (f1,l1), GApp (f2,l2) -> f f1 f2 && List.for_all2eq f l1 l2
+ | GLambda (na1,bk1,ty1,c1), GLambda (na2,bk2,ty2,c2)
when Name.equal na1 na2 && Constrexpr_ops.binding_kind_eq bk1 bk2 ->
on_true_do (f ty1 ty2 && f c1 c2) add na1
- | GProd (_,na1,bk1,ty1,c1), GProd (_,na2,bk2,ty2,c2)
+ | GProd (na1,bk1,ty1,c1), GProd (na2,bk2,ty2,c2)
when Name.equal na1 na2 && Constrexpr_ops.binding_kind_eq bk1 bk2 ->
on_true_do (f ty1 ty2 && f c1 c2) add na1
| GHole _, GHole _ -> true
- | GSort (_,s1), GSort (_,s2) -> Miscops.glob_sort_eq s1 s2
- | GLetIn (_,na1,b1,t1,c1), GLetIn (_,na2,b2,t2,c2) when Name.equal na1 na2 ->
+ | GSort s1, GSort s2 -> Miscops.glob_sort_eq s1 s2
+ | GLetIn (na1,b1,t1,c1), GLetIn (na2,b2,t2,c2) when Name.equal na1 na2 ->
on_true_do (f b1 b2 && f c1 c2) add na1
| (GCases _ | GRec _
| GPatVar _ | GEvar _ | GLetTuple _ | GIf _ | GCast _),_
@@ -129,49 +129,51 @@ let rec cases_pattern_fold_map loc g e = Loc.with_unloc (function
let subst_binder_type_vars l = function
| Evar_kinds.BinderType (Name id) ->
let id =
- try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id
+ try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id
with Not_found -> id in
Evar_kinds.BinderType (Name id)
| e -> e
-let rec subst_glob_vars l = function
- | GVar (_,id) as r -> (try Id.List.assoc id l with Not_found -> r)
- | GProd (loc,Name id,bk,t,c) ->
+let rec subst_glob_vars l gc = Loc.map (function
+ | GVar id as r -> (try snd @@ Id.List.assoc id l with Not_found -> r)
+ | GProd (Name id,bk,t,c) ->
let id =
- try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id
+ try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id
with Not_found -> id in
- GProd (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
- | GLambda (loc,Name id,bk,t,c) ->
+ GProd (Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
+ | GLambda (Name id,bk,t,c) ->
let id =
- try match Id.List.assoc id l with GVar(_,id') -> id' | _ -> id
+ try match snd @@ Id.List.assoc id l with GVar id' -> id' | _ -> id
with Not_found -> id in
- GLambda (loc,Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
- | GHole (loc,x,naming,arg) -> GHole (loc,subst_binder_type_vars l x,naming,arg)
- | r -> map_glob_constr (subst_glob_vars l) r (* assume: id is not binding *)
+ GLambda (Name id,bk,subst_glob_vars l t,subst_glob_vars l c)
+ | GHole (x,naming,arg) -> GHole (subst_binder_type_vars l x,naming,arg)
+ | _ -> snd @@ map_glob_constr (subst_glob_vars l) gc (* assume: id is not binding *)
+ ) gc
let ldots_var = Id.of_string ".."
-let glob_constr_of_notation_constr_with_binders loc g f e = function
- | NVar id -> GVar (loc,id)
- | NApp (a,args) -> GApp (loc,f e a, List.map (f e) args)
+let glob_constr_of_notation_constr_with_binders loc g f e nc =
+ let lt x = Loc.tag ~loc x in lt @@ match nc with
+ | NVar id -> GVar id
+ | NApp (a,args) -> GApp (f e a, List.map (f e) args)
| NList (x,y,iter,tail,swap) ->
let t = f e tail in let it = f e iter in
- let innerl = (ldots_var,t)::(if swap then [] else [x,GVar(loc,y)]) in
- let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in
- let outerl = (ldots_var,inner)::(if swap then [x,GVar(loc,y)] else []) in
- subst_glob_vars outerl it
+ let innerl = (ldots_var,t)::(if swap then [] else [x, lt @@ GVar y]) in
+ let inner = lt @@ GApp (lt @@ GVar (ldots_var),[subst_glob_vars innerl it]) in
+ let outerl = (ldots_var,inner)::(if swap then [x, lt @@ GVar y] else []) in
+ Loc.obj @@ subst_glob_vars outerl it
| NBinderList (x,y,iter,tail) ->
let t = f e tail in let it = f e iter in
- let innerl = [(ldots_var,t);(x,GVar(loc,y))] in
- let inner = GApp (loc,GVar (loc,ldots_var),[subst_glob_vars innerl it]) in
+ let innerl = [(ldots_var,t);(x, lt @@ GVar y)] in
+ let inner = lt @@ GApp (lt @@ GVar ldots_var,[subst_glob_vars innerl it]) in
let outerl = [(ldots_var,inner)] in
- subst_glob_vars outerl it
+ Loc.obj @@ subst_glob_vars outerl it
| NLambda (na,ty,c) ->
- let e',na = g e na in GLambda (loc,na,Explicit,f e ty,f e' c)
+ let e',na = g e na in GLambda (na,Explicit,f e ty,f e' c)
| NProd (na,ty,c) ->
- let e',na = g e na in GProd (loc,na,Explicit,f e ty,f e' c)
+ let e',na = g e na in GProd (na,Explicit,f e ty,f e' c)
| NLetIn (na,b,t,c) ->
- let e',na = g e na in GLetIn (loc,na,f e b,Option.map (f e) t,f e' c)
+ let e',na = g e na in GLetIn (na,f e b,Option.map (f e) t,f e' c)
| NCases (sty,rtntypopt,tml,eqnl) ->
let e',tml' = List.fold_right (fun (tm,(na,t)) (e',tml') ->
let e',t' = match t with
@@ -186,25 +188,25 @@ let glob_constr_of_notation_constr_with_binders loc g f e = function
let eqnl' = List.map (fun (patl,rhs) ->
let ((idl,e),patl) =
List.fold_map (cases_pattern_fold_map loc fold) ([],e) patl in
- (loc,idl,patl,f e rhs)) eqnl in
- GCases (loc,sty,Option.map (f e') rtntypopt,tml',eqnl')
+ lt (idl,patl,f e rhs)) eqnl in
+ GCases (sty,Option.map (f e') rtntypopt,tml',eqnl')
| NLetTuple (nal,(na,po),b,c) ->
let e',nal = List.fold_map g e nal in
let e'',na = g e na in
- GLetTuple (loc,nal,(na,Option.map (f e'') po),f e b,f e' c)
+ GLetTuple (nal,(na,Option.map (f e'') po),f e b,f e' c)
| NIf (c,(na,po),b1,b2) ->
let e',na = g e na in
- GIf (loc,f e c,(na,Option.map (f e') po),f e b1,f e b2)
+ GIf (f e c,(na,Option.map (f e') po),f e b1,f e b2)
| NRec (fk,idl,dll,tl,bl) ->
let e,dll = Array.fold_map (List.fold_map (fun e (na,oc,b) ->
let e,na = g e na in
(e,(na,Explicit,Option.map (f e) oc,f e b)))) e dll in
let e',idl = Array.fold_map (to_id g) e idl in
- GRec (loc,fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl)
- | NCast (c,k) -> GCast (loc,f e c,Miscops.map_cast_type (f e) k)
- | NSort x -> GSort (loc,x)
- | NHole (x, naming, arg) -> GHole (loc, x, naming, arg)
- | NRef x -> GRef (loc,x,None)
+ GRec (fk,idl,dll,Array.map (f e) tl,Array.map (f e') bl)
+ | NCast (c,k) -> GCast (f e c,Miscops.map_cast_type (f e) k)
+ | NSort x -> GSort x
+ | NHole (x, naming, arg) -> GHole (x, naming, arg)
+ | NRef x -> GRef (x,None)
let glob_constr_of_notation_constr loc x =
let rec aux () x =
@@ -220,13 +222,13 @@ let add_name r = function Anonymous -> () | Name id -> add_id r id
let split_at_recursive_part c =
let sub = ref None in
let rec aux = function
- | GApp (loc0,GVar(loc,v),c::l) when Id.equal v ldots_var ->
+ | loc0, GApp ((loc,GVar v),c::l) when Id.equal v ldots_var ->
begin match !sub with
| None ->
let () = sub := Some c in
begin match l with
- | [] -> GVar (loc, ldots_var)
- | _ :: _ -> GApp (loc0, GVar (loc, ldots_var), l)
+ | [] -> Loc.tag ~loc @@ GVar ldots_var
+ | _ :: _ -> Loc.tag ~loc:loc0 @@ GApp (Loc.tag ~loc @@ GVar ldots_var, l)
end
| Some _ ->
(* Not narrowed enough to find only one recursive part *)
@@ -237,13 +239,13 @@ let split_at_recursive_part c =
match !sub with
| None -> (* No recursive pattern found *) raise Not_found
| Some c ->
- match outer_iterator with
- | GVar (_,v) when Id.equal v ldots_var -> (* Not enough context *) raise Not_found
+ match Loc.obj outer_iterator with
+ | GVar v when Id.equal v ldots_var -> (* Not enough context *) raise Not_found
| _ -> outer_iterator, c
let subtract_loc loc1 loc2 = Loc.make_loc (fst (Loc.unloc loc1),fst (Loc.unloc loc2)-1)
-let check_is_hole id = function GHole _ -> () | t ->
+let check_is_hole id = function _, GHole _ -> () | t ->
user_err ~loc:(loc_of_glob_constr t)
(strbrk "In recursive notation with binders, " ++ pr_id id ++
strbrk " is expected to come without type.")
@@ -257,19 +259,19 @@ type recursive_pattern_kind =
let compare_recursive_parts found f f' (iterator,subc) =
let diff = ref None in
let terminator = ref None in
- let rec aux c1 c2 = match c1,c2 with
- | GVar(_,v), term when Id.equal v ldots_var ->
+ let rec aux (l1, c1) (l2, c2) = match c1, c2 with
+ | GVar v, term when Id.equal v ldots_var ->
(* We found the pattern *)
assert (match !terminator with None -> true | Some _ -> false);
- terminator := Some term;
+ terminator := Some (l2, term);
true
- | GApp (_,GVar(_,v),l1), GApp (_,term,l2) when Id.equal v ldots_var ->
+ | GApp ((_, GVar v),l1), GApp (term, l2) when Id.equal v ldots_var ->
(* We found the pattern, but there are extra arguments *)
(* (this allows e.g. alternative (recursive) notation of application) *)
assert (match !terminator with None -> true | Some _ -> false);
terminator := Some term;
List.for_all2eq aux l1 l2
- | GVar (_,x), GVar (_,y) when not (Id.equal x y) ->
+ | GVar x, GVar y when not (Id.equal x y) ->
(* We found the position where it differs *)
let lassoc = match !terminator with None -> false | Some _ -> true in
let x,y = if lassoc then y,x else x,y in
@@ -279,8 +281,8 @@ let compare_recursive_parts found f f' (iterator,subc) =
true
| Some _ -> false
end
- | GLambda (_,Name x,_,t_x,c), GLambda (_,Name y,_,t_y,term)
- | GProd (_,Name x,_,t_x,c), GProd (_,Name y,_,t_y,term) ->
+ | GLambda (Name x,_,t_x,c), GLambda (Name y,_,t_y,term)
+ | GProd (Name x,_,t_x,c), GProd (Name y,_,t_y,term) ->
(* We found a binding position where it differs *)
begin match !diff with
| None ->
@@ -289,7 +291,7 @@ let compare_recursive_parts found f f' (iterator,subc) =
| Some _ -> false
end
| _ ->
- compare_glob_constr aux (add_name found) c1 c2 in
+ compare_glob_constr aux (add_name found) (l1, c1) (l2, c2) in
if aux iterator subc then
match !diff with
| None ->
@@ -312,13 +314,13 @@ let compare_recursive_parts found f f' (iterator,subc) =
(pi1 !found, (x,y) :: pi2 !found, pi3 !found),x,y,lassoc in
let iterator =
f' (if lassoc then iterator
- else subst_glob_vars [x,GVar(Loc.ghost,y)] iterator) in
+ else subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in
(* found have been collected by compare_constr *)
found := newfound;
NList (x,y,iterator,f (Option.get !terminator),lassoc)
| Some (x,y,RecursiveBinders (t_x,t_y)) ->
let newfound = (pi1 !found, pi2 !found, (x,y) :: pi3 !found) in
- let iterator = f' (subst_glob_vars [x,GVar(Loc.ghost,y)] iterator) in
+ let iterator = f' (subst_glob_vars [x, Loc.tag @@ GVar y] iterator) in
(* found have been collected by compare_constr *)
found := newfound;
check_is_hole x t_x;
@@ -336,22 +338,22 @@ let notation_constr_and_vars_of_glob_constr a =
try compare_recursive_parts found aux aux' (split_at_recursive_part c)
with Not_found ->
found := keepfound;
- match c with
- | GApp (_,GVar (loc,f),[c]) when Id.equal f ldots_var ->
+ match snd c with
+ | GApp ((loc, GVar f),[c]) when Id.equal f ldots_var ->
(* Fall on the second part of the recursive pattern w/o having
found the first part *)
user_err ~loc
(str "Cannot find where the recursive pattern starts.")
- | c ->
+ | _c ->
aux' c
- and aux' = function
- | GVar (_,id) -> add_id found id; NVar id
- | GApp (_,g,args) -> NApp (aux g, List.map aux args)
- | GLambda (_,na,bk,ty,c) -> add_name found na; NLambda (na,aux ty,aux c)
- | GProd (_,na,bk,ty,c) -> add_name found na; NProd (na,aux ty,aux c)
- | GLetIn (_,na,b,t,c) -> add_name found na; NLetIn (na,aux b,Option.map aux t,aux c)
- | GCases (_,sty,rtntypopt,tml,eqnl) ->
- let f (_,idl,pat,rhs) = List.iter (add_id found) idl; (pat,aux rhs) in
+ and aux' x = Loc.with_unloc (function
+ | GVar id -> add_id found id; NVar id
+ | GApp (g,args) -> NApp (aux g, List.map aux args)
+ | GLambda (na,bk,ty,c) -> add_name found na; NLambda (na,aux ty,aux c)
+ | GProd (na,bk,ty,c) -> add_name found na; NProd (na,aux ty,aux c)
+ | GLetIn (na,b,t,c) -> add_name found na; NLetIn (na,aux b,Option.map aux t, aux c)
+ | GCases (sty,rtntypopt,tml,eqnl) ->
+ let f (_,(idl,pat,rhs)) = List.iter (add_id found) idl; (pat,aux rhs) in
NCases (sty,Option.map aux rtntypopt,
List.map (fun (tm,(na,x)) ->
add_name found na;
@@ -359,29 +361,29 @@ let notation_constr_and_vars_of_glob_constr a =
(fun (_,(_,nl)) -> List.iter (add_name found) nl) x;
(aux tm,(na,Option.map (fun (_,(ind,nal)) -> (ind,nal)) x))) tml,
List.map f eqnl)
- | GLetTuple (loc,nal,(na,po),b,c) ->
+ | GLetTuple (nal,(na,po),b,c) ->
add_name found na;
List.iter (add_name found) nal;
NLetTuple (nal,(na,Option.map aux po),aux b,aux c)
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
add_name found na;
NIf (aux c,(na,Option.map aux po),aux b1,aux b2)
- | GRec (_,fk,idl,dll,tl,bl) ->
+ | GRec (fk,idl,dll,tl,bl) ->
Array.iter (add_id found) idl;
let dll = Array.map (List.map (fun (na,bk,oc,b) ->
if bk != Explicit then
error "Binders marked as implicit not allowed in notations.";
add_name found na; (na,Option.map aux oc,aux b))) dll in
NRec (fk,idl,dll,Array.map aux tl,Array.map aux bl)
- | GCast (_,c,k) -> NCast (aux c,Miscops.map_cast_type aux k)
- | GSort (_,s) -> NSort s
- | GHole (_,w,naming,arg) ->
+ | GCast (c,k) -> NCast (aux c,Miscops.map_cast_type aux k)
+ | GSort s -> NSort s
+ | GHole (w,naming,arg) ->
if arg != None then has_ltac := true;
NHole (w, naming, arg)
- | GRef (_,r,_) -> NRef r
+ | GRef (r,_) -> NRef r
| GEvar _ | GPatVar _ ->
error "Existential variables not allowed in notations."
-
+ ) x
in
let t = aux a in
(* Side effect *)
@@ -590,8 +592,8 @@ let abstract_return_type_context pi mklam tml rtno =
let abstract_return_type_context_glob_constr =
abstract_return_type_context (fun (_,(_,nal)) -> nal)
- (fun na c ->
- GLambda(Loc.ghost,na,Explicit,GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c))
+ (fun na c -> Loc.tag @@
+ GLambda(na,Explicit,Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None),c))
let abstract_return_type_context_notation_constr =
abstract_return_type_context snd
@@ -663,18 +665,19 @@ let add_binding_env alp (terms,onlybinders,termlists,binderlists) var v =
let add_bindinglist_env (terms,onlybinders,termlists,binderlists) x bl =
(terms,onlybinders,termlists,(x,bl)::binderlists)
-let rec pat_binder_of_term = function
- | GVar (loc, id) -> Loc.tag ~loc @@ PatVar (Name id)
- | GApp (loc, GRef (_,ConstructRef cstr,_), l) ->
+let rec pat_binder_of_term t = Loc.map (function
+ | GVar id -> PatVar (Name id)
+ | GApp ((_, GRef (ConstructRef cstr,_)), l) ->
let nparams = Inductiveops.inductive_nparams (fst cstr) in
let _,l = List.chop nparams l in
- Loc.tag ~loc @@ PatCstr (cstr, List.map pat_binder_of_term l, Anonymous)
+ PatCstr (cstr, List.map pat_binder_of_term l, Anonymous)
| _ -> raise No_match
+ ) t
let bind_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var v =
try
let v' = Id.List.assoc var terms in
- match v, v' with
+ match Loc.obj v, Loc.obj v' with
| GHole _, _ -> sigma
| _, GHole _ ->
let sigma = Id.List.remove_assoc var terms,onlybinders,termlists,binderlists in
@@ -688,7 +691,7 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var
try
let vl' = Id.List.assoc var termlists in
let unify_term v v' =
- match v, v' with
+ match Loc.obj v, Loc.obj v' with
| GHole _, _ -> v'
| _, GHole _ -> v
| _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v' else raise No_match in
@@ -704,8 +707,8 @@ let bind_termlist_env alp (terms,onlybinders,termlists,binderlists as sigma) var
let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sigma) var id =
try
- match Id.List.assoc var terms with
- | GVar (_,id') ->
+ match Loc.obj @@ Id.List.assoc var terms with
+ | GVar id' ->
(if not (Id.equal id id') then (fst alp,(id,id')::snd alp) else alp),
sigma
| _ -> anomaly (str "A term which can be a binder has to be a variable")
@@ -713,7 +716,7 @@ let bind_term_as_binding_env alp (terms,onlybinders,termlists,binderlists as sig
(* The matching against a term allowing to find the instance has not been found yet *)
(* If it will be a different name, we shall unfortunately fail *)
(* TODO: look at the consequences for alp *)
- alp, add_env alp sigma var (GVar (Loc.ghost,id))
+ alp, add_env alp sigma var (Loc.tag @@ GVar id)
let bind_binding_as_term_env alp (terms,onlybinders,termlists,binderlists as sigma) var id =
try
@@ -782,7 +785,7 @@ let bind_bindinglist_env alp (terms,onlybinders,termlists,binderlists as sigma)
let unify_pat alp p p' =
try fold_cases_pattern_eq unify_name alp p p' with Failure _ -> raise No_match in
let unify_term alp v v' =
- match v, v' with
+ match Loc.obj v, Loc.obj v' with
| GHole _, _ -> v'
| _, GHole _ -> v
| _, _ -> if glob_constr_eq (alpha_rename (snd alp) v) v' then v else raise No_match in
@@ -831,7 +834,7 @@ let bind_bindinglist_as_term_env alp (terms,onlybinders,termlists,binderlists) v
else raise No_match in
let unify_term_binder c b' =
match c, b' with
- | GVar (loc, id), GLocalAssum (_, na', bk', t') ->
+ | (_, GVar id), GLocalAssum (loc, na', bk', t') ->
GLocalAssum (loc, unify_id id na', bk', t')
| c, GLocalPattern (loc, (p',ids), id, bk', t') ->
let p = pat_binder_of_term c in
@@ -892,21 +895,22 @@ let rec match_cases_pattern_binders metas acc (_, pat1) (_, pat2) =
let glue_letin_with_decls = true
-let rec match_iterated_binders islambda decls = function
- | GLambda (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b)]))
+let rec match_iterated_binders islambda decls bi = Loc.with_loc (fun ~loc -> function
+ | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))])))
when islambda && Id.equal p e ->
- match_iterated_binders islambda (GLocalPattern (loc,(cp,ids),p,bk,t)::decls) b
- | GLambda (loc,na,bk,t,b) when islambda ->
- match_iterated_binders islambda (GLocalAssum (loc,na,bk,t)::decls) b
- | GProd (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b)]))
+ match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b
+ | GLambda (na,bk,t,b) when islambda ->
+ match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b
+ | GProd (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b))])))
when not islambda && Id.equal p e ->
- match_iterated_binders islambda (GLocalPattern (loc,(cp,ids),p,bk,t)::decls) b
- | GProd (loc,(Name _ as na),bk,t,b) when not islambda ->
- match_iterated_binders islambda (GLocalAssum (loc,na,bk,t)::decls) b
- | GLetIn (loc,na,c,t,b) when glue_letin_with_decls ->
+ match_iterated_binders islambda (GLocalPattern(loc,(cp,ids),p,bk,t)::decls) b
+ | GProd ((Name _ as na),bk,t,b) when not islambda ->
+ match_iterated_binders islambda (GLocalAssum(loc,na,bk,t)::decls) b
+ | GLetIn (na,c,t,b) when glue_letin_with_decls ->
match_iterated_binders islambda
(GLocalDef (loc,na,Explicit (*?*), c,t)::decls) b
- | b -> (decls,b)
+ | b -> (decls, Loc.tag ~loc b)
+ ) bi
let remove_sigma x (terms,onlybinders,termlists,binderlists) =
(Id.List.remove_assoc x terms,onlybinders,termlists,binderlists)
@@ -967,91 +971,92 @@ let does_not_come_from_already_eta_expanded_var =
(* The following test is then an approximation of what can be done *)
(* optimally (whether other looping situations can occur remains to be *)
(* checked). *)
- function GVar _ -> false | _ -> true
+ function _loc, GVar _ -> false | _ -> true
let rec match_ inner u alp metas sigma a1 a2 =
- match (a1,a2) with
+ let loc, a1_val = Loc.to_pair a1 in
+ match a1_val, a2 with
(* Matching notation variable *)
- | r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 r1
- | GVar (_,id1), NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1
- | r1, NVar id2 when is_bindinglist_meta id2 metas -> bind_term_env alp sigma id2 r1
+ | r1, NVar id2 when is_term_meta id2 metas -> bind_term_env alp sigma id2 a1
+ | GVar id1, NVar id2 when is_onlybinding_meta id2 metas -> bind_binding_as_term_env alp sigma id2 id1
+ | r1, NVar id2 when is_bindinglist_meta id2 metas -> bind_term_env alp sigma id2 a1
(* Matching recursive notations for terms *)
| r1, NList (x,y,iter,termin,lassoc) ->
- match_termlist (match_hd u alp) alp metas sigma r1 x y iter termin lassoc
+ match_termlist (match_hd u alp) alp metas sigma a1 x y iter termin lassoc
(* "λ p, let 'cp = p in t" -> "λ 'cp, t" *)
- | GLambda (loc,Name p,bk,t1,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])),
+ | GLambda (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))),
NBinderList (x,_,NLambda (Name _id2,_,b2),termin) when Id.equal p e ->
let (decls,b) = match_iterated_binders true [GLocalPattern(loc,(cp,ids),p,bk,t1)] b1 in
let alp,sigma = bind_bindinglist_env alp sigma x decls in
match_in u alp metas sigma b termin
(* Matching recursive notations for binders: ad hoc cases supporting let-in *)
- | GLambda (loc,na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)->
+ | GLambda (na1,bk,t1,b1), NBinderList (x,_,NLambda (Name _id2,_,b2),termin)->
let (decls,b) = match_iterated_binders true [GLocalAssum (loc,na1,bk,t1)] b1 in
(* TODO: address the possibility that termin is a Lambda itself *)
let alp,sigma = bind_bindinglist_env alp sigma x decls in
match_in u alp metas sigma b termin
(* "∀ p, let 'cp = p in t" -> "∀ 'cp, t" *)
- | GProd (loc,Name p,bk,t1,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])),
+ | GProd (Name p,bk,t1,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))),
NBinderList (x,_,NProd (Name _id2,_,b2),(NVar v as termin)) when Id.equal p e ->
let (decls,b) = match_iterated_binders true [GLocalPattern (loc,(cp,ids),p,bk,t1)] b1 in
let alp,sigma = bind_bindinglist_env alp sigma x decls in
match_in u alp metas sigma b termin
- | GProd (loc,na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin)
+ | GProd (na1,bk,t1,b1), NBinderList (x,_,NProd (Name _id2,_,b2),termin)
when na1 != Anonymous ->
let (decls,b) = match_iterated_binders false [GLocalAssum (loc,na1,bk,t1)] b1 in
(* TODO: address the possibility that termin is a Prod itself *)
let alp,sigma = bind_bindinglist_env alp sigma x decls in
match_in u alp metas sigma b termin
(* Matching recursive notations for binders: general case *)
- | r, NBinderList (x,y,iter,termin) ->
- match_binderlist_with_app (match_hd u) alp metas sigma r x y iter termin
+ | _r, NBinderList (x,y,iter,termin) ->
+ match_binderlist_with_app (match_hd u) alp metas sigma a1 x y iter termin
(* Matching individual binders as part of a recursive pattern *)
- | GLambda (loc,Name p,bk,t,GCases (_,LetPatternStyle,None,[(GVar(_,e),_)],[(_,ids,[cp],b1)])),
+ | GLambda (Name p,bk,t,(_, GCases (LetPatternStyle,None,[((_, GVar e),_)],[(_,(ids,[cp],b1))]))),
NLambda (Name id,_,b2)
when is_bindinglist_meta id metas ->
let alp,sigma = bind_bindinglist_env alp sigma id [GLocalPattern (loc,(cp,ids),p,bk,t)] in
match_in u alp metas sigma b1 b2
- | GLambda (loc,na,bk,t,b1), NLambda (Name id,_,b2)
+ | GLambda (na,bk,t,b1), NLambda (Name id,_,b2)
when is_bindinglist_meta id metas ->
let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in
match_in u alp metas sigma b1 b2
- | GProd (loc,na,bk,t,b1), NProd (Name id,_,b2)
+ | GProd (na,bk,t,b1), NProd (Name id,_,b2)
when is_bindinglist_meta id metas && na != Anonymous ->
let alp,sigma = bind_bindinglist_env alp sigma id [GLocalAssum (loc,na,bk,t)] in
match_in u alp metas sigma b1 b2
(* Matching compositionally *)
- | GVar (_,id1), NVar id2 when alpha_var id1 id2 (fst alp) -> sigma
- | GRef (_,r1,_), NRef r2 when (eq_gr r1 r2) -> sigma
- | GApp (loc,f1,l1), NApp (f2,l2) ->
+ | GVar id1, NVar id2 when alpha_var id1 id2 (fst alp) -> sigma
+ | GRef (r1,_), NRef r2 when (eq_gr r1 r2) -> sigma
+ | GApp (f1,l1), NApp (f2,l2) ->
let n1 = List.length l1 and n2 = List.length l2 in
let f1,l1,f2,l2 =
if n1 < n2 then
let l21,l22 = List.chop (n2-n1) l2 in f1,l1, NApp (f2,l21), l22
else if n1 > n2 then
- let l11,l12 = List.chop (n1-n2) l1 in GApp (loc,f1,l11),l12, f2,l2
+ let l11,l12 = List.chop (n1-n2) l1 in Loc.tag ~loc @@ GApp (f1,l11),l12, f2,l2
else f1,l1, f2, l2 in
let may_use_eta = does_not_come_from_already_eta_expanded_var f1 in
List.fold_left2 (match_ may_use_eta u alp metas)
(match_in u alp metas sigma f1 f2) l1 l2
- | GLambda (_,na1,_,t1,b1), NLambda (na2,t2,b2) ->
+ | GLambda (na1,_,t1,b1), NLambda (na2,t2,b2) ->
match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2
- | GProd (_,na1,_,t1,b1), NProd (na2,t2,b2) ->
+ | GProd (na1,_,t1,b1), NProd (na2,t2,b2) ->
match_binders u alp metas na1 na2 (match_in u alp metas sigma t1 t2) b1 b2
- | GLetIn (_,na1,b1,_,c1), NLetIn (na2,b2,None,c2)
- | GLetIn (_,na1,b1,None,c1), NLetIn (na2,b2,_,c2) ->
+ | GLetIn (na1,b1,_,c1), NLetIn (na2,b2,None,c2)
+ | GLetIn (na1,b1,None,c1), NLetIn (na2,b2,_,c2) ->
match_binders u alp metas na1 na2 (match_in u alp metas sigma b1 b2) c1 c2
- | GLetIn (_,na1,b1,Some t1,c1), NLetIn (na2,b2,Some t2,c2) ->
+ | GLetIn (na1,b1,Some t1,c1), NLetIn (na2,b2,Some t2,c2) ->
match_binders u alp metas na1 na2
(match_in u alp metas (match_in u alp metas sigma b1 b2) t1 t2) c1 c2
- | GCases (_,sty1,rtno1,tml1,eqnl1), NCases (sty2,rtno2,tml2,eqnl2)
+ | GCases (sty1,rtno1,tml1,eqnl1), NCases (sty2,rtno2,tml2,eqnl2)
when sty1 == sty2
&& Int.equal (List.length tml1) (List.length tml2)
&& Int.equal (List.length eqnl1) (List.length eqnl2) ->
@@ -1065,17 +1070,17 @@ let rec match_ inner u alp metas sigma a1 a2 =
(fun s (tm1,_) (tm2,_) ->
match_in u alp metas s tm1 tm2) sigma tml1 tml2 in
List.fold_left2 (match_equations u alp metas) sigma eqnl1 eqnl2
- | GLetTuple (_,nal1,(na1,to1),b1,c1), NLetTuple (nal2,(na2,to2),b2,c2)
+ | GLetTuple (nal1,(na1,to1),b1,c1), NLetTuple (nal2,(na2,to2),b2,c2)
when Int.equal (List.length nal1) (List.length nal2) ->
let sigma = match_opt (match_binders u alp metas na1 na2) sigma to1 to2 in
let sigma = match_in u alp metas sigma b1 b2 in
let (alp,sigma) =
List.fold_left2 (match_names metas) (alp,sigma) nal1 nal2 in
match_in u alp metas sigma c1 c2
- | GIf (_,a1,(na1,to1),b1,c1), NIf (a2,(na2,to2),b2,c2) ->
+ | GIf (a1,(na1,to1),b1,c1), NIf (a2,(na2,to2),b2,c2) ->
let sigma = match_opt (match_binders u alp metas na1 na2) sigma to1 to2 in
List.fold_left2 (match_in u alp metas) sigma [a1;b1;c1] [a2;b2;c2]
- | GRec (_,fk1,idl1,dll1,tl1,bl1), NRec (fk2,idl2,dll2,tl2,bl2)
+ | GRec (fk1,idl1,dll1,tl1,bl1), NRec (fk2,idl2,dll2,tl2,bl2)
when match_fix_kind fk1 fk2 && Int.equal (Array.length idl1) (Array.length idl2) &&
Array.for_all2 (fun l1 l2 -> Int.equal (List.length l1) (List.length l2)) dll1 dll2
->
@@ -1089,13 +1094,13 @@ let rec match_ inner u alp metas sigma a1 a2 =
let alp,sigma = Array.fold_right2 (fun id1 id2 alsig ->
match_names metas alsig (Name id1) (Name id2)) idl1 idl2 (alp,sigma) in
Array.fold_left2 (match_in u alp metas) sigma bl1 bl2
- | GCast(_,c1,CastConv t1), NCast (c2,CastConv t2)
- | GCast(_,c1,CastVM t1), NCast (c2,CastVM t2) ->
+ | GCast(c1,CastConv t1), NCast (c2,CastConv t2)
+ | GCast(c1,CastVM t1), NCast (c2,CastVM t2) ->
match_in u alp metas (match_in u alp metas sigma c1 c2) t1 t2
- | GCast(_,c1, CastCoerce), NCast(c2, CastCoerce) ->
+ | GCast(c1, CastCoerce), NCast(c2, CastCoerce) ->
match_in u alp metas sigma c1 c2
- | GSort (_,GType _), NSort (GType _) when not u -> sigma
- | GSort (_,s1), NSort s2 when Miscops.glob_sort_eq s1 s2 -> sigma
+ | GSort (GType _), NSort (GType _) when not u -> sigma
+ | GSort s1, NSort s2 when Miscops.glob_sort_eq s1 s2 -> sigma
| GPatVar _, NHole _ -> (*Don't hide Metas, they bind in ltac*) raise No_match
| a, NHole _ -> sigma
@@ -1105,11 +1110,11 @@ let rec match_ inner u alp metas sigma a1 a2 =
otherwise how to ensure it corresponds to a well-typed eta-expansion;
we make an exception for types which are metavariables: this is useful e.g.
to print "{x:_ & P x}" knowing that notation "{x & P x}" is not defined. *)
- | b1, NLambda (Name id as na,(NHole _ | NVar _ as t2),b2) when inner ->
+ | _b1, NLambda (Name id as na,(NHole _ | NVar _ as t2),b2) when inner ->
let avoid =
- free_glob_vars b1 @ (* as in Namegen: *) glob_visible_short_qualid b1 in
+ free_glob_vars a1 @ (* as in Namegen: *) glob_visible_short_qualid a1 in
let id' = Namegen.next_ident_away id avoid in
- let t1 = GHole(Loc.ghost,Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in
+ let t1 = Loc.tag @@ GHole(Evar_kinds.BinderType (Name id'),Misctypes.IntroAnonymous,None) in
let sigma = match t2 with
| NHole _ -> sigma
| NVar id2 -> bind_term_env alp sigma id2 t1
@@ -1119,7 +1124,7 @@ let rec match_ inner u alp metas sigma a1 a2 =
bind_bindinglist_env alp sigma id [GLocalAssum (Loc.ghost,Name id',Explicit,t1)]
else
match_names metas (alp,sigma) (Name id') na in
- match_in u alp metas sigma (mkGApp Loc.ghost b1 (GVar (Loc.ghost,id'))) b2
+ match_in u alp metas sigma (mkGApp Loc.ghost a1 (Loc.tag @@ GVar id')) b2
| (GRec _ | GEvar _), _
| _,_ -> raise No_match
@@ -1132,7 +1137,7 @@ and match_binders u alp metas na1 na2 sigma b1 b2 =
let (alp,sigma) = match_names metas (alp,sigma) na1 na2 in
match_in u alp metas sigma b1 b2
-and match_equations u alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) =
+and match_equations u alp metas sigma (_,(_,patl1,rhs1)) (patl2,rhs2) =
(* patl1 and patl2 have the same length because they respectively
correspond to some tml1 and tml2 that have the same length *)
let (alp,sigma) =
@@ -1140,9 +1145,9 @@ and match_equations u alp metas sigma (_,_,patl1,rhs1) (patl2,rhs2) =
(alp,sigma) patl1 patl2 in
match_in u alp metas sigma rhs1 rhs2
-let term_of_binder = function
- | Name id -> GVar (Loc.ghost,id)
- | Anonymous -> GHole (Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)
+let term_of_binder bi = Loc.tag @@ match bi with
+ | Name id -> GVar id
+ | Anonymous -> GHole (Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None)
type glob_decl2 =
(name, cases_pattern) Util.union * Decl_kinds.binding_kind *
@@ -1157,7 +1162,7 @@ let match_notation_constr u c (metas,pat) =
with Not_found ->
(* Happens for binders bound to Anonymous *)
(* Find a better way to propagate Anonymous... *)
- GVar (Loc.ghost,x) in
+ Loc.tag @@GVar x in
List.fold_right (fun (x,(scl,typ)) (terms',termlists',binders') ->
match typ with
| NtnTypeConstr ->
diff --git a/intf/glob_term.mli b/intf/glob_term.mli
index 7a43c44f9..a8311e093 100644
--- a/intf/glob_term.mli
+++ b/intf/glob_term.mli
@@ -31,29 +31,29 @@ type cases_pattern_r =
and cases_pattern = cases_pattern_r Loc.located
(** Representation of an internalized (or in other words globalized) term. *)
-type glob_constr =
- | GRef of (Loc.t * global_reference * glob_level list option)
+type glob_constr_r =
+ | GRef of global_reference * glob_level list option
(** An identifier that represents a reference to an object defined
either in the (global) environment or in the (local) context. *)
- | GVar of (Loc.t * Id.t)
+ | GVar of Id.t
(** An identifier that cannot be regarded as "GRef".
Bound variables are typically represented this way. *)
- | GEvar of Loc.t * existential_name * (Id.t * glob_constr) list
- | GPatVar of Loc.t * (bool * patvar) (** Used for patterns only *)
- | GApp of Loc.t * glob_constr * glob_constr list
- | GLambda of Loc.t * Name.t * binding_kind * glob_constr * glob_constr
- | GProd of Loc.t * Name.t * binding_kind * glob_constr * glob_constr
- | GLetIn of Loc.t * Name.t * glob_constr * glob_constr option * glob_constr
- | GCases of Loc.t * case_style * glob_constr option * tomatch_tuples * cases_clauses
- (** [GCases(l,style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in [MatchStyle]) *)
- | GLetTuple of Loc.t * Name.t list * (Name.t * glob_constr option) *
- glob_constr * glob_constr
- | GIf of Loc.t * glob_constr * (Name.t * glob_constr option) * glob_constr * glob_constr
- | GRec of Loc.t * fix_kind * Id.t array * glob_decl list array *
- glob_constr array * glob_constr array
- | GSort of Loc.t * glob_sort
- | GHole of (Loc.t * Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option)
- | GCast of Loc.t * glob_constr * glob_constr cast_type
+ | GEvar of existential_name * (Id.t * glob_constr) list
+ | GPatVar of bool * patvar (** Used for patterns only *)
+ | GApp of glob_constr * glob_constr list
+ | GLambda of Name.t * binding_kind * glob_constr * glob_constr
+ | GProd of Name.t * binding_kind * glob_constr * glob_constr
+ | GLetIn of Name.t * glob_constr * glob_constr option * glob_constr
+ | GCases of case_style * glob_constr option * tomatch_tuples * cases_clauses
+ (** [GCases(style,r,tur,cc)] = "match 'tur' return 'r' with 'cc'" (in [MatchStyle]) *)
+ | GLetTuple of Name.t list * (Name.t * glob_constr option) * glob_constr * glob_constr
+ | GIf of glob_constr * (Name.t * glob_constr option) * glob_constr * glob_constr
+ | GRec of fix_kind * Id.t array * glob_decl list array *
+ glob_constr array * glob_constr array
+ | GSort of glob_sort
+ | GHole of Evar_kinds.t * intro_pattern_naming_expr * Genarg.glob_generic_argument option
+ | GCast of glob_constr * glob_constr cast_type
+and glob_constr = glob_constr_r Loc.located
and glob_decl = Name.t * binding_kind * glob_constr option * glob_constr
@@ -74,7 +74,7 @@ and tomatch_tuple = (glob_constr * predicate_pattern)
and tomatch_tuples = tomatch_tuple list
-and cases_clause = (Loc.t * Id.t list * cases_pattern list * glob_constr)
+and cases_clause = (Id.t list * cases_pattern list * glob_constr) Loc.located
(** [(p,il,cl,t)] = "|'cl' => 't'". Precondition: the free variables
of [t] are members of [il]. *)
and cases_clauses = cases_clause list
diff --git a/lib/loc.ml b/lib/loc.ml
index 8d7432ff4..2a785fac4 100644
--- a/lib/loc.ml
+++ b/lib/loc.ml
@@ -62,6 +62,7 @@ type 'a located = t * 'a
let to_pair x = x
let tag ?loc x = Option.default ghost loc, x
+let obj (_,x) = x
let with_loc f (loc, x) = f ~loc x
let with_unloc f (_,x) = f x
diff --git a/lib/loc.mli b/lib/loc.mli
index 3f484bc4c..10f63a8dd 100644
--- a/lib/loc.mli
+++ b/lib/loc.mli
@@ -60,6 +60,7 @@ type 'a located = t * 'a
val to_pair : 'a located -> t * 'a
val tag : ?loc:t -> 'a -> 'a located
+val obj : 'a located -> 'a
val with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b
val with_unloc : ('a -> 'b) -> 'a located -> 'b
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)
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 531485935..347c49f44 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -348,8 +348,8 @@ let find_tomatch_tycon evdref env loc = function
empty_tycon,None
let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
- let loc = Some (loc_of_glob_constr tomatch) in
- let tycon,realnames = find_tomatch_tycon evdref env loc indopt in
+ let loc = loc_of_glob_constr tomatch in
+ let tycon,realnames = find_tomatch_tycon evdref env (Some loc) indopt in
let j = typing_fun tycon env evdref tomatch in
let evd, j = Coercion.inh_coerce_to_base (loc_of_glob_constr tomatch) env !evdref j in
evdref := evd;
@@ -357,7 +357,7 @@ let coerce_row typing_fun evdref env pats (tomatch,(_,indopt)) =
let t =
try try_find_ind env !evdref typ realnames
with Not_found ->
- unify_tomatch_with_patterns evdref env loc typ pats realnames in
+ unify_tomatch_with_patterns evdref env (Some loc) typ pats realnames in
(j.uj_val,t)
let coerce_to_indtype typing_fun evdref env matx tomatchl =
@@ -1535,7 +1535,7 @@ substituer après par les initiaux *)
* and linearizing the _ patterns.
* Syntactic correctness has already been done in astterm *)
let matx_of_eqns env eqns =
- let build_eqn (loc,ids,lpat,rhs) =
+ let build_eqn (loc,(ids,lpat,rhs)) =
let initial_lpat,initial_rhs = lpat,rhs in
let initial_rhs = rhs in
let rhs =
@@ -2059,8 +2059,8 @@ let mk_JMeq evdref typ x typ' y =
let mk_JMeq_refl evdref typ x =
papp evdref coq_JMeq_refl [| typ; x |]
-let hole =
- GHole (Loc.ghost, Evar_kinds.QuestionMark (Evar_kinds.Define false),
+let hole = Loc.tag @@
+ GHole (Evar_kinds.QuestionMark (Evar_kinds.Define false),
Misctypes.IntroAnonymous, None)
let constr_of_pat env evdref arsign pat avoid =
@@ -2160,13 +2160,13 @@ let vars_of_ctx sigma ctx =
match decl with
| LocalDef (na,t',t) when is_topvar sigma t' ->
prev,
- (GApp (Loc.ghost,
- (GRef (Loc.ghost, delayed_force coq_eq_refl_ref, None)),
- [hole; GVar (Loc.ghost, prev)])) :: vars
+ (Loc.tag @@ GApp (
+ (Loc.tag @@ GRef (delayed_force coq_eq_refl_ref, None)),
+ [hole; Loc.tag @@ GVar prev])) :: vars
| _ ->
match RelDecl.get_name decl with
Anonymous -> invalid_arg "vars_of_ctx"
- | Name n -> n, GVar (Loc.ghost, n) :: vars)
+ | Name n -> n, (Loc.tag @@ GVar n) :: vars)
ctx (Id.of_string "vars_of_ctx_error", [])
in List.rev y
@@ -2289,13 +2289,13 @@ let constrs_of_pats typing_fun env evdref eqns tomatchs sign neqs arity =
let branch_name = Id.of_string ("program_branch_" ^ (string_of_int !i)) in
let branch_decl = LocalDef (Name branch_name, lift !i bbody, lift !i btype) in
let branch =
- let bref = GVar (Loc.ghost, branch_name) in
+ let bref = Loc.tag @@ GVar branch_name in
match vars_of_ctx !evdref rhs_rels with
[] -> bref
- | l -> GApp (Loc.ghost, bref, l)
+ | l -> Loc.tag @@ GApp (bref, l)
in
let branch = match ineqs with
- Some _ -> GApp (Loc.ghost, branch, [ hole ])
+ Some _ -> Loc.tag @@ GApp (branch, [ hole ])
| None -> branch
in
incr i;
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index f3018ac64..05d6a1ad4 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -331,20 +331,20 @@ let is_nondep_branch sigma c l =
let extract_nondep_branches test c b l =
let rec strip l r =
- match r,l with
- | r, [] -> r
- | GLambda (_,_,_,_,t), false::l -> strip l t
- | GLetIn (_,_,_,_,t), true::l -> strip l t
+ match snd r,l with
+ | r', [] -> r
+ | GLambda (_,_,_,t), false::l -> strip l t
+ | GLetIn (_,_,_,t), true::l -> strip l t
(* FIXME: do we need adjustment? *)
| _,_ -> assert false in
if test c l then Some (strip l b) else None
let it_destRLambda_or_LetIn_names l c =
let rec aux l nal c =
- match c, l with
+ match snd c, l with
| _, [] -> (List.rev nal,c)
- | GLambda (_,na,_,_,c), false::l -> aux l (na::nal) c
- | GLetIn (_,na,_,_,c), true::l -> aux l (na::nal) c
+ | GLambda (na,_,_,c), false::l -> aux l (na::nal) c
+ | GLetIn (na,_,_,c), true::l -> aux l (na::nal) c
| _, true::l -> (* let-expansion *) aux l (Anonymous :: nal) c
| _, false::l ->
(* eta-expansion *)
@@ -355,11 +355,11 @@ let it_destRLambda_or_LetIn_names l c =
x
in
let x = next (free_glob_vars c) in
- let a = GVar (dl,x) in
+ let a = Loc.tag @@ GVar x in
aux l (Name x :: nal)
(match c with
- | GApp (loc,p,l) -> GApp (loc,p,l@[a])
- | _ -> (GApp (dl,c,[a])))
+ | loc, GApp (p,l) -> (loc, GApp (p,l@[a]))
+ | _ -> Loc.tag @@ GApp (c,[a]))
in aux l [] c
let detype_case computable detype detype_eqns testdep avoid data p c bl =
@@ -375,12 +375,12 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
| None -> Anonymous, None, None
| Some p ->
let nl,typ = it_destRLambda_or_LetIn_names k p in
- let n,typ = match typ with
- | GLambda (_,x,_,t,c) -> x, c
+ let n,typ = match snd typ with
+ | GLambda (x,_,t,c) -> x, c
| _ -> Anonymous, typ in
let aliastyp =
if List.for_all (Name.equal Anonymous) nl then None
- else Some (dl,(indsp,nl)) in
+ else Some (Loc.tag (indsp,nl)) in
n, aliastyp, Some typ
in
let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in
@@ -397,25 +397,25 @@ let detype_case computable detype detype_eqns testdep avoid data p c bl =
else
st
with Not_found -> st
- in
+ in Loc.tag @@
match tag, aliastyp with
| LetStyle, None ->
let bl' = Array.map detype bl in
let (nal,d) = it_destRLambda_or_LetIn_names constagsl.(0) bl'.(0) in
- GLetTuple (dl,nal,(alias,pred),tomatch,d)
+ GLetTuple (nal,(alias,pred),tomatch,d)
| IfStyle, None ->
let bl' = Array.map detype bl in
let nondepbrs =
Array.map3 (extract_nondep_branches testdep) bl bl' constagsl in
if Array.for_all ((!=) None) nondepbrs then
- GIf (dl,tomatch,(alias,pred),
+ GIf (tomatch,(alias,pred),
Option.get nondepbrs.(0),Option.get nondepbrs.(1))
else
let eqnl = detype_eqns constructs constagsl bl in
- GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl)
| _ ->
let eqnl = detype_eqns constructs constagsl bl in
- GCases (dl,tag,pred,[tomatch,(alias,aliastyp)],eqnl)
+ GCases (tag,pred,[tomatch,(alias,aliastyp)],eqnl)
let detype_sort sigma = function
| Prop Null -> GProp
@@ -423,7 +423,7 @@ let detype_sort sigma = function
| Type u ->
GType
(if !print_universes
- then [dl, Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)]
+ then [Loc.tag @@ Pp.string_of_ppcmds (Univ.Universe.pr_with (Termops.pr_evd_level sigma) u)]
else [])
type binder_kind = BProd | BLambda | BLetIn
@@ -431,36 +431,36 @@ type binder_kind = BProd | BLambda | BLetIn
(**********************************************************************)
(* Main detyping function *)
-let detype_anonymous = ref (fun loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable"))
+let detype_anonymous = ref (fun ?loc n -> anomaly ~label:"detype" (Pp.str "index to an anonymous variable"))
let set_detype_anonymous f = detype_anonymous := f
let detype_level sigma l =
- GType (Some (dl, Pp.string_of_ppcmds (Termops.pr_evd_level sigma l)))
+ GType (Some (Loc.tag @@ Pp.string_of_ppcmds (Termops.pr_evd_level sigma l)))
let detype_instance sigma l =
let l = EInstance.kind sigma l in
if Univ.Instance.is_empty l then None
else Some (List.map (detype_level sigma) (Array.to_list (Univ.Instance.to_array l)))
-let rec detype flags avoid env sigma t =
+let rec detype flags avoid env sigma t = Loc.tag @@
match EConstr.kind sigma (collapse_appl sigma t) with
| Rel n ->
(try match lookup_name_of_rel n (fst env) with
- | Name id -> GVar (dl, id)
- | Anonymous -> !detype_anonymous dl n
+ | Name id -> GVar id
+ | Anonymous -> snd @@ !detype_anonymous n
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in GVar (dl, Id.of_string s))
+ in GVar (Id.of_string s))
| Meta n ->
(* Meta in constr are not user-parsable and are mapped to Evar *)
(* using numbers to be unparsable *)
- GEvar (dl, Id.of_string ("M" ^ string_of_int n), [])
+ GEvar (Id.of_string ("M" ^ string_of_int n), [])
| Var id ->
- (try let _ = Global.lookup_named id in GRef (dl, VarRef id, None)
- with Not_found -> GVar (dl, id))
- | Sort s -> GSort (dl,detype_sort sigma (ESorts.kind sigma s))
+ (try let _ = Global.lookup_named id in GRef (VarRef id, None)
+ with Not_found -> GVar id)
+ | Sort s -> GSort (detype_sort sigma (ESorts.kind sigma s))
| Cast (c1,REVERTcast,c2) when not !Flags.raw_print ->
- detype flags avoid env sigma c1
+ snd (detype flags avoid env sigma c1)
| Cast (c1,k,c2) ->
let d1 = detype flags avoid env sigma c1 in
let d2 = detype flags avoid env sigma c2 in
@@ -469,34 +469,34 @@ let rec detype flags avoid env sigma t =
| NATIVEcast -> CastNative d2
| _ -> CastConv d2
in
- GCast(dl,d1,cast)
- | Prod (na,ty,c) -> detype_binder flags BProd avoid env sigma na None ty c
- | Lambda (na,ty,c) -> detype_binder flags BLambda avoid env sigma na None ty c
- | LetIn (na,b,ty,c) -> detype_binder flags BLetIn avoid env sigma na (Some b) ty c
+ GCast(d1,cast)
+ | Prod (na,ty,c) -> snd @@ detype_binder flags BProd avoid env sigma na None ty c
+ | Lambda (na,ty,c) -> snd @@ detype_binder flags BLambda avoid env sigma na None ty c
+ | LetIn (na,b,ty,c) -> snd @@ detype_binder flags BLetIn avoid env sigma na (Some b) ty c
| App (f,args) ->
let mkapp f' args' =
- match f' with
- | GApp (dl',f',args'') ->
- GApp (dl,f',args''@args')
- | _ -> GApp (dl,f',args')
+ match snd f' with
+ | GApp (f',args'') ->
+ GApp (f',args''@args')
+ | _ -> GApp (f',args')
in
mkapp (detype flags avoid env sigma f)
(Array.map_to_list (detype flags avoid env sigma) args)
- | Const (sp,u) -> GRef (dl, ConstRef sp, detype_instance sigma u)
+ | Const (sp,u) -> GRef (ConstRef sp, detype_instance sigma u)
| Proj (p,c) ->
let noparams () =
let pb = Environ.lookup_projection p (snd env) in
let pars = pb.Declarations.proj_npars in
- let hole = GHole(Loc.ghost,Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in
+ let hole = Loc.tag @@ GHole(Evar_kinds.InternalHole,Misctypes.IntroAnonymous,None) in
let args = List.make pars hole in
- GApp (dl, GRef (dl, ConstRef (Projection.constant p), None),
+ GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None),
(args @ [detype flags avoid env sigma c]))
in
if fst flags || !Flags.in_debugger || !Flags.in_toplevel then
try noparams ()
with _ ->
(* lax mode, used by debug printers only *)
- GApp (dl, GRef (dl, ConstRef (Projection.constant p), None),
+ GApp (Loc.tag @@ GRef (ConstRef (Projection.constant p), None),
[detype flags avoid env sigma c])
else
if print_primproj_compatibility () && Projection.unfolded p then
@@ -514,12 +514,12 @@ let rec detype flags avoid env sigma t =
substl (c :: List.rev args) body'
with Retyping.RetypeError _ | Not_found ->
anomaly (str"Cannot detype an unfolded primitive projection.")
- in detype flags avoid env sigma c'
+ in snd @@ detype flags avoid env sigma c'
else
if print_primproj_params () then
try
let c = Retyping.expand_projection (snd env) sigma p c [] in
- detype flags avoid env sigma c
+ snd @@ detype flags avoid env sigma c
with Retyping.RetypeError _ -> noparams ()
else noparams ()
@@ -546,14 +546,15 @@ let rec detype flags avoid env sigma t =
Id.of_string ("X" ^ string_of_int (Evar.repr evk)),
(Array.map_to_list (fun c -> (Id.of_string "__",c)) cl)
in
- GEvar (dl,id,
+ GEvar (id,
List.map (on_snd (detype flags avoid env sigma)) l)
| Ind (ind_sp,u) ->
- GRef (dl, IndRef ind_sp, detype_instance sigma u)
+ GRef (IndRef ind_sp, detype_instance sigma u)
| Construct (cstr_sp,u) ->
- GRef (dl, ConstructRef cstr_sp, detype_instance sigma u)
+ GRef (ConstructRef cstr_sp, detype_instance sigma u)
| Case (ci,p,c,bl) ->
let comp = computable sigma p (List.length (ci.ci_pp_info.ind_tags)) in
+ snd @@
detype_case comp (detype flags avoid env sigma)
(detype_eqns flags avoid env sigma ci comp)
(is_nondep_branch sigma) avoid
@@ -574,7 +575,7 @@ and detype_fix flags avoid env sigma (vn,_ as nvn) (names,tys,bodies) =
let v = Array.map3
(fun c t i -> share_names flags (i+1) [] def_avoid def_env sigma c (lift n t))
bodies tys vn in
- GRec(dl,GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
+ GRec(GFix (Array.map (fun i -> Some i, GStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
@@ -590,7 +591,7 @@ and detype_cofix flags avoid env sigma n (names,tys,bodies) =
let v = Array.map2
(fun c t -> share_names flags 0 [] def_avoid def_env sigma c (lift ntys t))
bodies tys in
- GRec(dl,GCoFix n,Array.of_list (List.rev lfi),
+ GRec(GCoFix n,Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
@@ -635,7 +636,7 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
try
if !Flags.raw_print || not (reverse_matching ()) then raise Exit;
let mat = build_tree Anonymous (snd flags) (avoid,env) sigma ci bl in
- List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype flags avoid env sigma c))
+ List.map (fun (pat,((avoid,env),c)) -> Loc.tag ([],[pat],detype flags avoid env sigma c))
mat
with e when CErrors.noncritical e ->
Array.to_list
@@ -644,7 +645,7 @@ and detype_eqns flags avoid env sigma ci computable constructs consnargsl bl =
and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs branch =
let make_pat x avoid env b body ty ids =
if force_wildcard () && noccurn sigma 1 b then
- Loc.tag @@ PatVar (Anonymous),avoid,(add_name Anonymous body ty env),ids
+ Loc.tag @@ PatVar Anonymous,avoid,(add_name Anonymous body ty env),ids
else
let flag = if isgoal then RenamingForGoal else RenamingForCasesPattern (fst env,b) in
let na,avoid' = compute_displayed_name_in sigma flag avoid x b in
@@ -652,9 +653,9 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
in
let rec buildrec ids patlist avoid env l b =
match EConstr.kind sigma b, l with
- | _, [] ->
- (dl, Id.Set.elements ids,
- [Loc.tag ~loc:dl @@ PatCstr(constr, List.rev patlist,Anonymous)],
+ | _, [] -> Loc.tag @@
+ (Id.Set.elements ids,
+ [Loc.tag @@ PatCstr(constr, List.rev patlist,Anonymous)],
detype flags avoid env sigma b)
| Lambda (x,t,b), false::l ->
let pat,new_avoid,new_env,new_ids = make_pat x avoid env b None t ids in
@@ -668,7 +669,7 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
buildrec ids patlist avoid env l c
| _, true::l ->
- let pat = Loc.tag ~loc:dl @@ PatVar Anonymous in
+ let pat = Loc.tag @@ PatVar Anonymous in
buildrec ids (pat::patlist) avoid env l b
| _, false::l ->
@@ -683,21 +684,21 @@ and detype_eqn (lax,isgoal as flags) avoid env sigma constr construct_nargs bran
in
buildrec Id.Set.empty [] avoid env construct_nargs branch
-and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c =
+and detype_binder (lax,isgoal as flags) bk avoid env sigma na body ty c = Loc.tag @@
let flag = if isgoal then RenamingForGoal else RenamingElsewhereFor (fst env,c) in
let na',avoid' = match bk with
| BLetIn -> compute_displayed_let_name_in sigma flag avoid na c
| _ -> compute_displayed_name_in sigma flag avoid na c in
let r = detype flags avoid' (add_name na' body ty env) sigma c in
match bk with
- | BProd -> GProd (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r)
- | BLambda -> GLambda (dl, na',Explicit,detype (lax,false) avoid env sigma ty, r)
+ | BProd -> GProd (na',Explicit,detype (lax,false) avoid env sigma ty, r)
+ | BLambda -> GLambda (na',Explicit,detype (lax,false) avoid env sigma ty, r)
| BLetIn ->
let c = detype (lax,false) avoid env sigma (Option.get body) in
(* Heuristic: we display the type if in Prop *)
let s = try Retyping.get_sort_family_of (snd env) sigma ty with _ when !Flags.in_debugger || !Flags.in_toplevel -> InType (* Can fail because of sigma missing in debugger *) in
let t = if s != InProp then None else Some (detype (lax,false) avoid env sigma ty) in
- GLetIn (dl, na', c, t, r)
+ GLetIn (na', c, t, r)
let detype_rel_context ?(lax=false) where avoid env sigma sign =
let where = Option.map (fun c -> EConstr.it_mkLambda_or_LetIn c sign) where in
@@ -741,11 +742,11 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
| Name id -> Name (convert_id cl id)
| Anonymous -> Anonymous
in
- let rec detype_closed_glob cl = function
- | GVar (loc,id) ->
+ let rec detype_closed_glob cl cg = Loc.map (function
+ | GVar id ->
(* if [id] is bound to a name. *)
begin try
- GVar(loc,Id.Map.find id cl.idents)
+ GVar(Id.Map.find id cl.idents)
(* if [id] is bound to a typed term *)
with Not_found -> try
(* assumes [detype] does not raise [Not_found] exceptions *)
@@ -755,38 +756,39 @@ let detype_closed_glob ?lax isgoal avoid env sigma t =
[Printer.pr_constr_under_binders_env] does. *)
let assums = List.map (fun id -> LocalAssum (Name id,(* dummy *) mkProp)) b in
let env = push_rel_context assums env in
- detype ?lax isgoal avoid env sigma c
+ snd @@ detype ?lax isgoal avoid env sigma c
(* if [id] is bound to a [closed_glob_constr]. *)
with Not_found -> try
let {closure;term} = Id.Map.find id cl.untyped in
- detype_closed_glob closure term
+ snd @@ detype_closed_glob closure term
(* Otherwise [id] stands for itself *)
with Not_found ->
- GVar(loc,id)
+ GVar id
end
- | GLambda (loc,id,k,t,c) ->
+ | GLambda (id,k,t,c) ->
let id = convert_name cl id in
- GLambda(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c)
- | GProd (loc,id,k,t,c) ->
+ GLambda(id,k,detype_closed_glob cl t, detype_closed_glob cl c)
+ | GProd (id,k,t,c) ->
let id = convert_name cl id in
- GProd(loc,id,k,detype_closed_glob cl t, detype_closed_glob cl c)
- | GLetIn (loc,id,b,t,e) ->
+ GProd(id,k,detype_closed_glob cl t, detype_closed_glob cl c)
+ | GLetIn (id,b,t,e) ->
let id = convert_name cl id in
- GLetIn(loc,id,detype_closed_glob cl b, Option.map (detype_closed_glob cl) t, detype_closed_glob cl e)
- | GLetTuple (loc,ids,(n,r),b,e) ->
+ GLetIn(id,detype_closed_glob cl b, Option.map (detype_closed_glob cl) t, detype_closed_glob cl e)
+ | GLetTuple (ids,(n,r),b,e) ->
let ids = List.map (convert_name cl) ids in
let n = convert_name cl n in
- GLetTuple (loc,ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e)
- | GCases (loc,sty,po,tml,eqns) ->
+ GLetTuple (ids,(n,r),detype_closed_glob cl b, detype_closed_glob cl e)
+ | GCases (sty,po,tml,eqns) ->
let (tml,eqns) =
Glob_ops.map_pattern_binders (fun na -> convert_name cl na) tml eqns
in
let (tml,eqns) =
Glob_ops.map_pattern (fun c -> detype_closed_glob cl c) tml eqns
in
- GCases(loc,sty,po,tml,eqns)
+ GCases(sty,po,tml,eqns)
| c ->
- Glob_ops.map_glob_constr (detype_closed_glob cl) c
+ snd @@ Glob_ops.map_glob_constr (detype_closed_glob cl) cg
+ ) cg
in
detype_closed_glob t.closure t.term
@@ -804,41 +806,41 @@ let rec subst_cases_pattern subst (loc, pat) = Loc.tag ~loc @@
let (f_subst_genarg, subst_genarg_hook) = Hook.make ()
-let rec subst_glob_constr subst raw =
+let rec subst_glob_constr subst (loc, raw) = Loc.tag ~loc @@
match raw with
- | GRef (loc,ref,u) ->
+ | GRef (ref,u) ->
let ref',t = subst_global subst ref in
if ref' == ref then raw else
- detype false [] (Global.env()) Evd.empty (EConstr.of_constr t)
+ snd @@ detype false [] (Global.env()) Evd.empty (EConstr.of_constr t)
| GVar _ -> raw
| GEvar _ -> raw
| GPatVar _ -> raw
- | GApp (loc,r,rl) ->
+ | GApp (r,rl) ->
let r' = subst_glob_constr subst r
and rl' = List.smartmap (subst_glob_constr subst) rl in
if r' == r && rl' == rl then raw else
- GApp(loc,r',rl')
+ GApp(r',rl')
- | GLambda (loc,n,bk,r1,r2) ->
+ | GLambda (n,bk,r1,r2) ->
let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- GLambda (loc,n,bk,r1',r2')
+ GLambda (n,bk,r1',r2')
- | GProd (loc,n,bk,r1,r2) ->
+ | GProd (n,bk,r1,r2) ->
let r1' = subst_glob_constr subst r1 and r2' = subst_glob_constr subst r2 in
if r1' == r1 && r2' == r2 then raw else
- GProd (loc,n,bk,r1',r2')
+ GProd (n,bk,r1',r2')
- | GLetIn (loc,n,r1,t,r2) ->
+ | GLetIn (n,r1,t,r2) ->
let r1' = subst_glob_constr subst r1 in
- let t' = Option.smartmap (subst_glob_constr subst) t in
let r2' = subst_glob_constr subst r2 in
+ let t' = Option.smartmap (subst_glob_constr subst) t in
if r1' == r1 && t == t' && r2' == r2 then raw else
- GLetIn (loc,n,r1',t',r2')
+ GLetIn (n,r1',t',r2')
- | GCases (loc,sty,rtno,rl,branches) ->
+ | GCases (sty,rtno,rl,branches) ->
let rtno' = Option.smartmap (subst_glob_constr subst) rtno
and rl' = List.smartmap (fun (a,x as y) ->
let a' = subst_glob_constr subst a in
@@ -849,33 +851,33 @@ let rec subst_glob_constr subst raw =
if sp == sp' then t else (loc,((sp',i),y))) topt in
if a == a' && topt == topt' then y else (a',(n,topt'))) rl
and branches' = List.smartmap
- (fun (loc,idl,cpl,r as branch) ->
+ (fun (loc,(idl,cpl,r) as branch) ->
let cpl' =
List.smartmap (subst_cases_pattern subst) cpl
and r' = subst_glob_constr subst r in
if cpl' == cpl && r' == r then branch else
- (loc,idl,cpl',r'))
+ (loc,(idl,cpl',r')))
branches
in
if rtno' == rtno && rl' == rl && branches' == branches then raw else
- GCases (loc,sty,rtno',rl',branches')
+ GCases (sty,rtno',rl',branches')
- | GLetTuple (loc,nal,(na,po),b,c) ->
+ | GLetTuple (nal,(na,po),b,c) ->
let po' = Option.smartmap (subst_glob_constr subst) po
and b' = subst_glob_constr subst b
and c' = subst_glob_constr subst c in
if po' == po && b' == b && c' == c then raw else
- GLetTuple (loc,nal,(na,po'),b',c')
+ GLetTuple (nal,(na,po'),b',c')
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
let po' = Option.smartmap (subst_glob_constr subst) po
and b1' = subst_glob_constr subst b1
and b2' = subst_glob_constr subst b2
and c' = subst_glob_constr subst c in
if c' == c && po' == po && b1' == b1 && b2' == b2 then raw else
- GIf (loc,c',(na,po'),b1',b2')
+ GIf (c',(na,po'),b1',b2')
- | GRec (loc,fix,ida,bl,ra1,ra2) ->
+ | GRec (fix,ida,bl,ra1,ra2) ->
let ra1' = Array.smartmap (subst_glob_constr subst) ra1
and ra2' = Array.smartmap (subst_glob_constr subst) ra2 in
let bl' = Array.smartmap
@@ -885,11 +887,11 @@ let rec subst_glob_constr subst raw =
if ty'==ty && obd'==obd then dcl else (na,k,obd',ty')))
bl in
if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
- GRec (loc,fix,ida,bl',ra1',ra2')
+ GRec (fix,ida,bl',ra1',ra2')
| GSort _ -> raw
- | GHole (loc, knd, naming, solve) ->
+ | GHole (knd, naming, solve) ->
let nknd = match knd with
| Evar_kinds.ImplicitArg (ref, i, b) ->
let nref, _ = subst_global subst ref in
@@ -898,12 +900,12 @@ let rec subst_glob_constr subst raw =
in
let nsolve = Option.smartmap (Hook.get f_subst_genarg subst) solve in
if nsolve == solve && nknd == knd then raw
- else GHole (loc, nknd, naming, nsolve)
+ else GHole (nknd, naming, nsolve)
- | GCast (loc,r1,k) ->
+ | GCast (r1,k) ->
let r1' = subst_glob_constr subst r1 in
let k' = Miscops.smartmap_cast_type (subst_glob_constr subst) k in
- if r1' == r1 && k' == k then raw else GCast (loc,r1',k')
+ if r1' == r1 && k' == k then raw else GCast (r1',k')
(* Utilities to transform kernel cases to simple pattern-matching problem *)
@@ -914,7 +916,7 @@ let simple_cases_matrix_of_branches ind brs =
let p = Loc.tag @@ PatCstr ((ind,i+1),List.map mkPatVar nal,Anonymous) in
let map name = try Some (Nameops.out_name name) with Failure _ -> None in
let ids = List.map_filter map nal in
- (Loc.ghost,ids,[p],c))
+ Loc.tag @@ (ids,[p],c))
brs
let return_type_of_predicate ind nrealargs_tags pred =
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index 4c6f9129f..84da3652f 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -38,7 +38,7 @@ val detype : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> constr -> glob
val detype_case :
bool -> (constr -> glob_constr) ->
(constructor array -> bool list array -> constr array ->
- (Loc.t * Id.t list * cases_pattern list * glob_constr) list) ->
+ (Id.t list * cases_pattern list * glob_constr) Loc.located list) ->
(constr -> bool list -> bool) ->
Id.t list -> inductive * case_style * bool list array * bool list ->
constr option -> constr -> constr array -> glob_constr
@@ -54,7 +54,9 @@ val detype_closed_glob : ?lax:bool -> bool -> Id.t list -> env -> evar_map -> cl
val lookup_name_as_displayed : env -> evar_map -> constr -> Id.t -> int option
val lookup_index_as_renamed : env -> evar_map -> constr -> int -> int option
-val set_detype_anonymous : (Loc.t -> int -> glob_constr) -> unit
+(* XXX: This is a hack and should go away *)
+val set_detype_anonymous : (?loc:Loc.t -> int -> glob_constr) -> unit
+
val force_wildcard : unit -> bool
val synthetize_type : unit -> bool
diff --git a/pretyping/glob_ops.ml b/pretyping/glob_ops.ml
index 4cccaaf8f..25ece5b8e 100644
--- a/pretyping/glob_ops.ml
+++ b/pretyping/glob_ops.ml
@@ -22,10 +22,10 @@ let cases_predicate_names tml =
| (tm,(na,None)) -> [na]
| (tm,(na,Some (_,(_,nal)))) -> na::nal) tml)
-let mkGApp loc p t =
- match p with
- | GApp (loc,f,l) -> GApp (loc,f,l@[t])
- | _ -> GApp (loc,p,[t])
+let mkGApp loc p t = Loc.tag ~loc @@
+ match snd p with
+ | GApp (f,l) -> GApp (f,l@[t])
+ | _ -> GApp (p,[t])
let map_glob_decl_left_to_right f (na,k,obd,ty) =
let comp1 = Option.map f obd in
@@ -59,46 +59,46 @@ let cast_type_eq eq t1 t2 = match t1, t2 with
| CastNative t1, CastNative t2 -> eq t1 t2
| _ -> false
-let rec glob_constr_eq c1 c2 = match c1, c2 with
-| GRef (_, gr1, _), GRef (_, gr2, _) -> eq_gr gr1 gr2
-| GVar (_, id1), GVar (_, id2) -> Id.equal id1 id2
-| GEvar (_, id1, arg1), GEvar (_, id2, arg2) ->
+let rec glob_constr_eq (_loc1, c1) (_loc2, c2) = match c1, c2 with
+| GRef (gr1, _), GRef (gr2, _) -> eq_gr gr1 gr2
+| GVar id1, GVar id2 -> Id.equal id1 id2
+| GEvar (id1, arg1), GEvar (id2, arg2) ->
Id.equal id1 id2 &&
List.equal instance_eq arg1 arg2
-| GPatVar (_, (b1, pat1)), GPatVar (_, (b2, pat2)) ->
+| GPatVar (b1, pat1), GPatVar (b2, pat2) ->
(b1 : bool) == b2 && Id.equal pat1 pat2
-| GApp (_, f1, arg1), GApp (_, f2, arg2) ->
+| GApp (f1, arg1), GApp (f2, arg2) ->
glob_constr_eq f1 f2 && List.equal glob_constr_eq arg1 arg2
-| GLambda (_, na1, bk1, t1, c1), GLambda (_, na2, bk2, t2, c2) ->
+| GLambda (na1, bk1, t1, c1), GLambda (na2, bk2, t2, c2) ->
Name.equal na1 na2 && binding_kind_eq bk1 bk2 &&
glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GProd (_, na1, bk1, t1, c1), GProd (_, na2, bk2, t2, c2) ->
+| GProd (na1, bk1, t1, c1), GProd (na2, bk2, t2, c2) ->
Name.equal na1 na2 && binding_kind_eq bk1 bk2 &&
glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GLetIn (_, na1, b1, t1, c1), GLetIn (_, na2, b2, t2, c2) ->
+| GLetIn (na1, b1, t1, c1), GLetIn (na2, b2, t2, c2) ->
Name.equal na1 na2 && glob_constr_eq b1 b2 && Option.equal glob_constr_eq t1 t2 && glob_constr_eq c1 c2
-| GCases (_, st1, c1, tp1, cl1), GCases (_, st2, c2, tp2, cl2) ->
+| GCases (st1, c1, tp1, cl1), GCases (st2, c2, tp2, cl2) ->
case_style_eq st1 st2 && Option.equal glob_constr_eq c1 c2 &&
List.equal tomatch_tuple_eq tp1 tp2 &&
List.equal cases_clause_eq cl1 cl2
-| GLetTuple (_, na1, (n1, p1), c1, t1), GLetTuple (_, na2, (n2, p2), c2, t2) ->
+| GLetTuple (na1, (n1, p1), c1, t1), GLetTuple (na2, (n2, p2), c2, t2) ->
List.equal Name.equal na1 na2 && Name.equal n1 n2 &&
Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 &&
glob_constr_eq t1 t2
-| GIf (_, m1, (pat1, p1), c1, t1), GIf (_, m2, (pat2, p2), c2, t2) ->
+| GIf (m1, (pat1, p1), c1, t1), GIf (m2, (pat2, p2), c2, t2) ->
glob_constr_eq m1 m2 && Name.equal pat1 pat2 &&
Option.equal glob_constr_eq p1 p2 && glob_constr_eq c1 c2 &&
glob_constr_eq t1 t2
-| GRec (_, kn1, id1, decl1, c1, t1), GRec (_, kn2, id2, decl2, c2, t2) ->
+| GRec (kn1, id1, decl1, c1, t1), GRec (kn2, id2, decl2, c2, t2) ->
fix_kind_eq kn1 kn2 && Array.equal Id.equal id1 id2 &&
Array.equal (fun l1 l2 -> List.equal glob_decl_eq l1 l2) decl1 decl2 &&
Array.equal glob_constr_eq c1 c2 &&
Array.equal glob_constr_eq t1 t2
-| GSort (_, s1), GSort (_, s2) -> Miscops.glob_sort_eq s1 s2
-| GHole (_, kn1, nam1, gn1), GHole (_, kn2, nam2, gn2) ->
+| GSort s1, GSort s2 -> Miscops.glob_sort_eq s1 s2
+| GHole (kn1, nam1, gn1), GHole (kn2, nam2, gn2) ->
Option.equal (==) gn1 gn2 (** Only thing sensible *) &&
Miscops.intro_pattern_naming_eq nam1 nam2
-| GCast (_, c1, t1), GCast (_, c2, t2) ->
+| GCast (c1, t1), GCast (c2, t2) ->
glob_constr_eq c1 c2 && cast_type_eq glob_constr_eq t1 t2
| _ -> false
@@ -109,7 +109,7 @@ and tomatch_tuple_eq (c1, p1) (c2, p2) =
let eq_pred (n1, o1) (n2, o2) = Name.equal n1 n2 && Option.equal eqp o1 o2 in
glob_constr_eq c1 c2 && eq_pred p1 p2
-and cases_clause_eq (_, id1, p1, c1) (_, id2, p2, c2) =
+and cases_clause_eq (_, (id1, p1, c1)) (_, (id2, p2, c2)) =
List.equal Id.equal id1 id2 && List.equal cases_pattern_eq p1 p2 &&
glob_constr_eq c1 c2
@@ -137,80 +137,82 @@ and fix_recursion_order_eq o1 o2 = match o1, o2 with
and instance_eq (x1,c1) (x2,c2) =
Id.equal x1 x2 && glob_constr_eq c1 c2
-let map_glob_constr_left_to_right f = function
- | GApp (loc,g,args) ->
+let map_glob_constr_left_to_right f = Loc.map (function
+ | GApp (g,args) ->
let comp1 = f g in
let comp2 = Util.List.map_left f args in
- GApp (loc,comp1,comp2)
- | GLambda (loc,na,bk,ty,c) ->
+ GApp (comp1,comp2)
+ | GLambda (na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- GLambda (loc,na,bk,comp1,comp2)
- | GProd (loc,na,bk,ty,c) ->
+ GLambda (na,bk,comp1,comp2)
+ | GProd (na,bk,ty,c) ->
let comp1 = f ty in
let comp2 = f c in
- GProd (loc,na,bk,comp1,comp2)
- | GLetIn (loc,na,b,t,c) ->
+ GProd (na,bk,comp1,comp2)
+ | GLetIn (na,b,t,c) ->
let comp1 = f b in
let compt = Option.map f t in
let comp2 = f c in
- GLetIn (loc,na,comp1,compt,comp2)
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ GLetIn (na,comp1,compt,comp2)
+ | GCases (sty,rtntypopt,tml,pl) ->
let comp1 = Option.map f rtntypopt in
let comp2 = Util.List.map_left (fun (tm,x) -> (f tm,x)) tml in
- let comp3 = Util.List.map_left (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl in
- GCases (loc,sty,comp1,comp2,comp3)
- | GLetTuple (loc,nal,(na,po),b,c) ->
+ let comp3 = Util.List.map_left (fun (loc,(idl,p,c)) -> (loc,(idl,p,f c))) pl in
+ GCases (sty,comp1,comp2,comp3)
+ | GLetTuple (nal,(na,po),b,c) ->
let comp1 = Option.map f po in
let comp2 = f b in
let comp3 = f c in
- GLetTuple (loc,nal,(na,comp1),comp2,comp3)
- | GIf (loc,c,(na,po),b1,b2) ->
+ GLetTuple (nal,(na,comp1),comp2,comp3)
+ | GIf (c,(na,po),b1,b2) ->
let comp1 = Option.map f po in
let comp2 = f b1 in
let comp3 = f b2 in
- GIf (loc,f c,(na,comp1),comp2,comp3)
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ GIf (f c,(na,comp1),comp2,comp3)
+ | GRec (fk,idl,bl,tyl,bv) ->
let comp1 = Array.map (Util.List.map_left (map_glob_decl_left_to_right f)) bl in
let comp2 = Array.map f tyl in
let comp3 = Array.map f bv in
- GRec (loc,fk,idl,comp1,comp2,comp3)
- | GCast (loc,c,k) ->
+ GRec (fk,idl,comp1,comp2,comp3)
+ | GCast (c,k) ->
let comp1 = f c in
let comp2 = Miscops.map_cast_type f k in
- GCast (loc,comp1,comp2)
+ GCast (comp1,comp2)
| (GVar _ | GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) as x -> x
+ )
let map_glob_constr = map_glob_constr_left_to_right
let fold_return_type f acc (na,tyopt) = Option.fold_left f acc tyopt
-let fold_glob_constr f acc = function
+let fold_glob_constr f acc = Loc.with_unloc (function
| GVar _ -> acc
- | GApp (_,c,args) -> List.fold_left f (f acc c) args
- | GLambda (_,_,_,b,c) | GProd (_,_,_,b,c) ->
+ | GApp (c,args) -> List.fold_left f (f acc c) args
+ | GLambda (_,_,b,c) | GProd (_,_,b,c) ->
f (f acc b) c
- | GLetIn (_,_,b,t,c) ->
+ | GLetIn (_,b,t,c) ->
f (Option.fold_left f (f acc b) t) c
- | GCases (_,_,rtntypopt,tml,pl) ->
- let fold_pattern acc (_,idl,p,c) = f acc c in
+ | GCases (_,rtntypopt,tml,pl) ->
+ let fold_pattern acc (_,(idl,p,c)) = f acc c in
List.fold_left fold_pattern
(List.fold_left f (Option.fold_left f acc rtntypopt) (List.map fst tml))
pl
- | GLetTuple (_,_,rtntyp,b,c) ->
+ | GLetTuple (_,rtntyp,b,c) ->
f (f (fold_return_type f acc rtntyp) b) c
- | GIf (_,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
f (f (f (fold_return_type f acc rtntyp) c) b1) b2
- | GRec (_,_,_,bl,tyl,bv) ->
+ | GRec (_,_,bl,tyl,bv) ->
let acc = Array.fold_left
(List.fold_left (fun acc (na,k,bbd,bty) ->
f (Option.fold_left f acc bbd) bty)) acc bl in
Array.fold_left f (Array.fold_left f acc tyl) bv
- | GCast (_,c,k) ->
+ | GCast (c,k) ->
let acc = match k with
| CastConv t | CastVM t | CastNative t -> f acc t | CastCoerce -> acc in
f acc c
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> acc
+ )
let iter_glob_constr f = fold_glob_constr (fun () -> f) ()
@@ -219,25 +221,25 @@ let same_id na id = match na with
| Name id' -> Id.equal id id'
let occur_glob_constr id =
- let rec occur = function
- | GVar (loc,id') -> Id.equal id id'
- | GApp (loc,f,args) -> (occur f) || (List.exists occur args)
- | GLambda (loc,na,bk,ty,c) ->
+ let rec occur gt = Loc.with_unloc (function
+ | GVar (id') -> Id.equal id id'
+ | GApp (f,args) -> (occur f) || (List.exists occur args)
+ | GLambda (na,bk,ty,c) ->
(occur ty) || (not (same_id na id) && (occur c))
- | GProd (loc,na,bk,ty,c) ->
+ | GProd (na,bk,ty,c) ->
(occur ty) || (not (same_id na id) && (occur c))
- | GLetIn (loc,na,b,t,c) ->
+ | GLetIn (na,b,t,c) ->
(Option.fold_left (fun b t -> occur t || b) (occur b) t) || (not (same_id na id) && (occur c))
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ | GCases (sty,rtntypopt,tml,pl) ->
(occur_option rtntypopt)
|| (List.exists (fun (tm,_) -> occur tm) tml)
|| (List.exists occur_pattern pl)
- | GLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (nal,rtntyp,b,c) ->
occur_return_type rtntyp id
|| (occur b) || (not (List.mem_f Name.equal (Name id) nal) && (occur c))
- | GIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
occur_return_type rtntyp id || (occur c) || (occur b1) || (occur b2)
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (fk,idl,bl,tyl,bv) ->
not (Array.for_all4 (fun fid bl ty bd ->
let rec occur_fix = function
[] -> not (occur ty) && (Id.equal fid id || not(occur bd))
@@ -249,11 +251,11 @@ let occur_glob_constr id =
(match na with Name id' -> Id.equal id id' | _ -> not (occur_fix bl)) in
occur_fix bl)
idl bl tyl bv)
- | GCast (loc,c,k) -> (occur c) || (match k with CastConv t
+ | GCast (c,k) -> (occur c) || (match k with CastConv t
| CastVM t | CastNative t -> occur t | CastCoerce -> false)
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> false
-
- and occur_pattern (loc,idl,p,c) = not (Id.List.mem id idl) && (occur c)
+ ) gt
+ and occur_pattern (loc,(idl,p,c)) = not (Id.List.mem id idl) && (occur c)
and occur_option = function None -> false | Some p -> occur p
@@ -268,33 +270,33 @@ let add_name_to_ids set na =
| Name id -> Id.Set.add id set
let free_glob_vars =
- let rec vars bounded vs = function
- | GVar (loc,id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs
- | GApp (loc,f,args) -> List.fold_left (vars bounded) vs (f::args)
- | GLambda (loc,na,_,ty,c) | GProd (loc,na,_,ty,c) ->
+ let rec vars bounded vs = Loc.with_unloc @@ (function
+ | GVar (id') -> if Id.Set.mem id' bounded then vs else Id.Set.add id' vs
+ | GApp (f,args) -> List.fold_left (vars bounded) vs (f::args)
+ | GLambda (na,_,ty,c) | GProd (na,_,ty,c) ->
let vs' = vars bounded vs ty in
let bounded' = add_name_to_ids bounded na in
vars bounded' vs' c
- | GLetIn (loc,na,b,ty,c) ->
+ | GLetIn (na,b,ty,c) ->
let vs' = vars bounded vs b in
let vs'' = Option.fold_left (vars bounded) vs' ty in
let bounded' = add_name_to_ids bounded na in
vars bounded' vs'' c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ | GCases (sty,rtntypopt,tml,pl) ->
let vs1 = vars_option bounded vs rtntypopt in
let vs2 = List.fold_left (fun vs (tm,_) -> vars bounded vs tm) vs1 tml in
List.fold_left (vars_pattern bounded) vs2 pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (nal,rtntyp,b,c) ->
let vs1 = vars_return_type bounded vs rtntyp in
let vs2 = vars bounded vs1 b in
let bounded' = List.fold_left add_name_to_ids bounded nal in
vars bounded' vs2 c
- | GIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
let vs1 = vars_return_type bounded vs rtntyp in
let vs2 = vars bounded vs1 c in
let vs3 = vars bounded vs2 b1 in
vars bounded vs3 b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (fk,idl,bl,tyl,bv) ->
let bounded' = Array.fold_right Id.Set.add idl bounded in
let vars_fix i vs fid =
let vs1,bounded1 =
@@ -312,11 +314,12 @@ let free_glob_vars =
vars bounded1 vs2 bv.(i)
in
Array.fold_left_i vars_fix vs idl
- | GCast (loc,c,k) -> let v = vars bounded vs c in
+ | GCast (c,k) -> let v = vars bounded vs c in
(match k with CastConv t | CastVM t | CastNative t -> vars bounded v t | _ -> v)
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> vs
+ )
- and vars_pattern bounded vs (loc,idl,p,c) =
+ and vars_pattern bounded vs (loc,(idl,p,c)) =
let bounded' = List.fold_right Id.Set.add idl bounded in
vars bounded' vs c
@@ -332,7 +335,7 @@ let free_glob_vars =
let glob_visible_short_qualid c =
let rec aux acc = function
- | GRef (_,c,_) ->
+ | _, GRef (c,_) ->
let qualid = Nametab.shortest_qualid_of_global Id.Set.empty c in
let dir,id = Libnames.repr_qualid qualid in
if DirPath.is_empty dir then id :: acc else acc
@@ -351,26 +354,26 @@ let add_and_check_ident id set =
Id.Set.add id set
let bound_glob_vars =
- let rec vars bound = function
- | GLambda (_,na,_,_,_) | GProd (_,na,_,_,_) | GLetIn (_,na,_,_,_) as c ->
+ let rec vars bound = Loc.with_loc (fun ~loc -> function
+ | GLambda (na,_,_,_) | GProd (na,_,_,_) | GLetIn (na,_,_,_) as c ->
let bound = name_fold add_and_check_ident na bound in
- fold_glob_constr vars bound c
- | GCases (loc,sty,rtntypopt,tml,pl) ->
+ fold_glob_constr vars bound (loc, c)
+ | GCases (sty,rtntypopt,tml,pl) ->
let bound = vars_option bound rtntypopt in
let bound =
List.fold_left (fun bound (tm,_) -> vars bound tm) bound tml in
List.fold_left vars_pattern bound pl
- | GLetTuple (loc,nal,rtntyp,b,c) ->
+ | GLetTuple (nal,rtntyp,b,c) ->
let bound = vars_return_type bound rtntyp in
let bound = vars bound b in
let bound = List.fold_right (name_fold add_and_check_ident) nal bound in
vars bound c
- | GIf (loc,c,rtntyp,b1,b2) ->
+ | GIf (c,rtntyp,b1,b2) ->
let bound = vars_return_type bound rtntyp in
let bound = vars bound c in
let bound = vars bound b1 in
vars bound b2
- | GRec (loc,fk,idl,bl,tyl,bv) ->
+ | GRec (fk,idl,bl,tyl,bv) ->
let bound = Array.fold_right Id.Set.add idl bound in
let vars_fix i bound fid =
let bound =
@@ -388,9 +391,10 @@ let bound_glob_vars =
in
Array.fold_left_i vars_fix bound idl
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _ | GVar _) -> bound
- | GApp _ | GCast _ as c -> fold_glob_constr vars bound c
+ | GApp _ | GCast _ as c -> fold_glob_constr vars bound (loc, c)
+ )
- and vars_pattern bound (loc,idl,p,c) =
+ and vars_pattern bound (loc,(idl,p,c)) =
let bound = List.fold_right add_and_check_ident idl bound in
vars bound c
@@ -435,14 +439,14 @@ let rec map_case_pattern_binders f = Loc.map (function
else PatCstr(c,rps,rna)
)
-let map_cases_branch_binders f ((loc,il,cll,rhs) as x) : cases_clause =
+let map_cases_branch_binders f ((loc,(il,cll,rhs)) as x) : cases_clause =
(* spiwack: not sure if I must do something with the list of idents.
It is intended to be a superset of the free variable of the
right-hand side, if I understand correctly. But I'm not sure when
or how they are used. *)
let r = List.smartmap (fun cl -> map_case_pattern_binders f cl) cll in
if r == cll then x
- else loc,il,r,rhs
+ else loc,(il,r,rhs)
let map_pattern_binders f tomatch branches =
CList.smartmap (fun tm -> map_tomatch_binders f tm) tomatch,
@@ -452,29 +456,14 @@ let map_pattern_binders f tomatch branches =
let map_tomatch f (c,pp) : tomatch_tuple = f c , pp
-let map_cases_branch f (loc,il,cll,rhs) : cases_clause =
- loc , il , cll , f rhs
+let map_cases_branch f (loc,(il,cll,rhs)) : cases_clause =
+ loc , (il , cll , f rhs)
let map_pattern f tomatch branches =
List.map (fun tm -> map_tomatch f tm) tomatch,
List.map (fun br -> map_cases_branch f br) branches
-let loc_of_glob_constr = function
- | GRef (loc,_,_) -> loc
- | GVar (loc,_) -> loc
- | GEvar (loc,_,_) -> loc
- | GPatVar (loc,_) -> loc
- | GApp (loc,_,_) -> loc
- | GLambda (loc,_,_,_,_) -> loc
- | GProd (loc,_,_,_,_) -> loc
- | GLetIn (loc,_,_,_,_) -> loc
- | GCases (loc,_,_,_,_) -> loc
- | GLetTuple (loc,_,_,_,_) -> loc
- | GIf (loc,_,_,_,_) -> loc
- | GRec (loc,_,_,_,_,_) -> loc
- | GSort (loc,_) -> loc
- | GHole (loc,_,_,_) -> loc
- | GCast (loc,_,_) -> loc
+let loc_of_glob_constr (loc, _) = loc
(**********************************************************************)
(* Alpha-renaming *)
@@ -506,73 +495,74 @@ let rename_var l id =
if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming
else id
-let rec rename_glob_vars l = function
- | GVar (loc,id) as r ->
+let rec rename_glob_vars l = Loc.map_with_loc (fun ~loc -> function
+ | GVar id as r ->
let id' = rename_var l id in
- if id == id' then r else GVar (loc,id')
- | GRef (_,VarRef id,_) as r ->
+ if id == id' then r else GVar id'
+ | GRef (VarRef id,_) as r ->
if List.exists (fun (_,id') -> Id.equal id id') l then raise UnsoundRenaming
else r
- | GProd (loc,na,bk,t,c) ->
+ | GProd (na,bk,t,c) ->
let na',l' = update_subst na l in
- GProd (loc,na,bk,rename_glob_vars l t,rename_glob_vars l' c)
- | GLambda (loc,na,bk,t,c) ->
+ GProd (na,bk,rename_glob_vars l t,rename_glob_vars l' c)
+ | GLambda (na,bk,t,c) ->
let na',l' = update_subst na l in
- GLambda (loc,na',bk,rename_glob_vars l t,rename_glob_vars l' c)
- | GLetIn (loc,na,b,t,c) ->
+ GLambda (na',bk,rename_glob_vars l t,rename_glob_vars l' c)
+ | GLetIn (na,b,t,c) ->
let na',l' = update_subst na l in
- GLetIn (loc,na',rename_glob_vars l b,Option.map (rename_glob_vars l) t,rename_glob_vars l' c)
+ GLetIn (na',rename_glob_vars l b,Option.map (rename_glob_vars l) t,rename_glob_vars l' c)
(* Lazy strategy: we fail if a collision with renaming occurs, rather than renaming further *)
- | GCases (loc,ci,po,tomatchl,cls) ->
+ | GCases (ci,po,tomatchl,cls) ->
let test_pred_pat (na,ino) =
test_na l na; Option.iter (fun (_,(_,nal)) -> List.iter (test_na l) nal) ino in
let test_clause idl = List.iter (test_id l) idl in
let po = Option.map (rename_glob_vars l) po in
let tomatchl = Util.List.map_left (fun (tm,x) -> test_pred_pat x; (rename_glob_vars l tm,x)) tomatchl in
- let cls = Util.List.map_left (fun (loc,idl,p,c) -> test_clause idl; (loc,idl,p,rename_glob_vars l c)) cls in
- GCases (loc,ci,po,tomatchl,cls)
- | GLetTuple (loc,nal,(na,po),c,b) ->
+ let cls = Util.List.map_left (fun (loc,(idl,p,c)) -> test_clause idl; (loc,(idl,p,rename_glob_vars l c))) cls in
+ GCases (ci,po,tomatchl,cls)
+ | GLetTuple (nal,(na,po),c,b) ->
List.iter (test_na l) (na::nal);
- GLetTuple (loc,nal,(na,Option.map (rename_glob_vars l) po),
+ GLetTuple (nal,(na,Option.map (rename_glob_vars l) po),
rename_glob_vars l c,rename_glob_vars l b)
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
test_na l na;
- GIf (loc,rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po),
+ GIf (rename_glob_vars l c,(na,Option.map (rename_glob_vars l) po),
rename_glob_vars l b1,rename_glob_vars l b2)
- | GRec (loc,k,idl,decls,bs,ts) ->
+ | GRec (k,idl,decls,bs,ts) ->
Array.iter (test_id l) idl;
- GRec (loc,k,idl,
+ GRec (k,idl,
Array.map (List.map (fun (na,k,bbd,bty) ->
test_na l na; (na,k,Option.map (rename_glob_vars l) bbd,rename_glob_vars l bty))) decls,
Array.map (rename_glob_vars l) bs,
Array.map (rename_glob_vars l) ts)
- | r -> map_glob_constr (rename_glob_vars l) r
+ (* XXX: This located use case should be improved. *)
+ | r -> snd @@ map_glob_constr (rename_glob_vars l) (loc, r)
+ )
(**********************************************************************)
(* Conversion from glob_constr to cases pattern, if possible *)
-let rec cases_pattern_of_glob_constr na = function
- | GVar (loc,id) ->
+let rec cases_pattern_of_glob_constr na = Loc.map (function
+ | GVar id ->
begin match na with
| Name _ ->
(* Unable to manage the presence of both an alias and a variable *)
raise Not_found
- | Anonymous -> Loc.tag ~loc @@ PatVar (Name id)
+ | Anonymous -> PatVar (Name id)
end
- | GHole (loc,_,_,_) -> Loc.tag ~loc @@ PatVar na
- | GRef (loc,ConstructRef cstr,_) ->
- Loc.tag ~loc @@ PatCstr (cstr,[],na)
- | GApp (loc,GRef (_,ConstructRef cstr,_),l) ->
- Loc.tag ~loc @@ PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
+ | GHole (_,_,_) -> PatVar na
+ | GRef (ConstructRef cstr,_) -> PatCstr (cstr,[],na)
+ | GApp ((_loc, GRef (ConstructRef cstr,_)),l) ->
+ PatCstr (cstr,List.map (cases_pattern_of_glob_constr Anonymous) l,na)
| _ -> raise Not_found
+ )
(* Turn a closed cases pattern into a glob_constr *)
-let rec glob_constr_of_closed_cases_pattern_aux x = Loc.with_loc (fun ~loc -> function
- | PatCstr (cstr,[],Anonymous) ->
- GRef (loc,ConstructRef cstr,None)
- | PatCstr (cstr,l,Anonymous) ->
- let ref = GRef (loc,ConstructRef cstr,None) in
- GApp (loc,ref, List.map glob_constr_of_closed_cases_pattern_aux l)
+let rec glob_constr_of_closed_cases_pattern_aux x = Loc.map_with_loc (fun ~loc -> function
+ | PatCstr (cstr,[],Anonymous) -> GRef (ConstructRef cstr,None)
+ | PatCstr (cstr,l,Anonymous) ->
+ let ref = Loc.tag ~loc @@ GRef (ConstructRef cstr,None) in
+ GApp (ref, List.map glob_constr_of_closed_cases_pattern_aux l)
| _ -> raise Not_found
) x
diff --git a/pretyping/patternops.ml b/pretyping/patternops.ml
index 48ae93f3e..6696e174b 100644
--- a/pretyping/patternops.ml
+++ b/pretyping/patternops.ml
@@ -324,46 +324,46 @@ let warn_cast_in_pattern =
CWarnings.create ~name:"cast-in-pattern" ~category:"automation"
(fun () -> Pp.strbrk "Casts are ignored in patterns")
-let rec pat_of_raw metas vars = function
- | GVar (_,id) ->
+let rec pat_of_raw metas vars = Loc.with_loc (fun ~loc -> function
+ | GVar id ->
(try PRel (List.index Name.equal (Name id) vars)
with Not_found -> PVar id)
- | GPatVar (_,(false,n)) ->
+ | GPatVar (false,n) ->
metas := n::!metas; PMeta (Some n)
- | GRef (_,gr,_) ->
+ | GRef (gr,_) ->
PRef (canonical_gr gr)
(* Hack to avoid rewriting a complete interpretation of patterns *)
- | GApp (_, GPatVar (_,(true,n)), cl) ->
+ | GApp ((_, GPatVar (true,n)), cl) ->
metas := n::!metas; PSoApp (n, List.map (pat_of_raw metas vars) cl)
- | GApp (_,c,cl) ->
+ | GApp (c,cl) ->
PApp (pat_of_raw metas vars c,
Array.of_list (List.map (pat_of_raw metas vars) cl))
- | GLambda (_,na,bk,c1,c2) ->
+ | GLambda (na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLambda (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | GProd (_,na,bk,c1,c2) ->
+ | GProd (na,bk,c1,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PProd (na, pat_of_raw metas vars c1,
pat_of_raw metas (na::vars) c2)
- | GLetIn (_,na,c1,t,c2) ->
+ | GLetIn (na,c1,t,c2) ->
name_iter (fun n -> metas := n::!metas) na;
PLetIn (na, pat_of_raw metas vars c1,
Option.map (pat_of_raw metas vars) t,
pat_of_raw metas (na::vars) c2)
- | GSort (_,s) ->
+ | GSort s ->
PSort s
| GHole _ ->
PMeta None
- | GCast (_,c,_) ->
+ | GCast (c,_) ->
warn_cast_in_pattern ();
pat_of_raw metas vars c
- | GIf (_,c,(_,None),b1,b2) ->
+ | GIf (c,(_,None),b1,b2) ->
PIf (pat_of_raw metas vars c,
pat_of_raw metas vars b1,pat_of_raw metas vars b2)
- | GLetTuple (loc,nal,(_,None),b,c) ->
- let mkGLambda c na =
- GLambda (loc,na,Explicit,GHole (loc,Evar_kinds.InternalHole, IntroAnonymous, None),c) in
+ | GLetTuple (nal,(_,None),b,c) ->
+ let mkGLambda c na = Loc.tag ~loc @@
+ GLambda (na,Explicit, Loc.tag @@ GHole (Evar_kinds.InternalHole, IntroAnonymous, None),c) in
let c = List.fold_left mkGLambda c nal in
let cip =
{ cip_style = LetStyle;
@@ -374,9 +374,9 @@ let rec pat_of_raw metas vars = function
let tags = List.map (fun _ -> false) nal (* Approximation which can be without let-ins... *) in
PCase (cip, PMeta None, pat_of_raw metas vars b,
[0,tags,pat_of_raw metas vars c])
- | GCases (loc,sty,p,[c,(na,indnames)],brs) ->
+ | GCases (sty,p,[c,(na,indnames)],brs) ->
let get_ind = function
- | (_,_,[_, PatCstr((ind,_),_,_)],_)::_ -> Some ind
+ | (_,(_,[_, PatCstr((ind,_),_,_)],_))::_ -> Some ind
| _ -> None
in
let ind_tags,ind = match indnames with
@@ -389,7 +389,7 @@ let rec pat_of_raw metas vars = function
| Some p, Some (_,(_,nal)) ->
let nvars = na :: List.rev nal @ vars in
rev_it_mkPLambda nal (mkPLambda na (pat_of_raw metas nvars p))
- | (None | Some (GHole _)), _ -> PMeta None
+ | (None | Some (_, GHole _)), _ -> PMeta None
| Some p, None ->
user_err ~loc (strbrk "Clause \"in\" expected in patterns over \"match\" expressions with an explicit \"return\" clause.")
in
@@ -404,7 +404,8 @@ let rec pat_of_raw metas vars = function
one non-trivial branch. These facts are used in [Constrextern]. *)
PCase (info, pred, pat_of_raw metas vars c, brs)
- | r -> err ~loc:(loc_of_glob_constr r) (Pp.str "Non supported pattern.")
+ | r -> err ~loc (Pp.str "Non supported pattern.")
+ )
and pats_of_glob_branches loc metas vars ind brs =
let get_arg = function
@@ -415,8 +416,8 @@ and pats_of_glob_branches loc metas vars ind brs =
in
let rec get_pat indexes = function
| [] -> false, []
- | [(_,_,[_, PatVar(Anonymous)],GHole _)] -> true, [] (* ends with _ => _ *)
- | (_,_,[_, PatCstr((indsp,j),lv,_)],br) :: brs ->
+ | [(_,(_,[_, PatVar(Anonymous)],(_,GHole _)))] -> true, [] (* ends with _ => _ *)
+ | (_,(_,[_, PatCstr((indsp,j),lv,_)],br)) :: brs ->
let () = match ind with
| Some sp when eq_ind sp indsp -> ()
| _ ->
@@ -431,7 +432,7 @@ and pats_of_glob_branches loc metas vars ind brs =
let ext,pats = get_pat (Int.Set.add (j-1) indexes) brs in
let tags = List.map (fun _ -> false) lv (* approximation, w/o let-in *) in
ext, ((j-1, tags, pat) :: pats)
- | (loc,_,_,_) :: _ -> err ~loc (Pp.str "Non supported pattern.")
+ | (loc,(_,_,_)) :: _ -> err ~loc (Pp.str "Non supported pattern.")
in
get_pat Int.Set.empty brs
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index ae87cd8c0..5f9f4bb08 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -567,23 +567,23 @@ let (f_genarg_interp, genarg_interp_hook) = Hook.make ()
(* in environment [env], with existential variables [evdref] and *)
(* the type constraint tycon *)
-let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) t =
+let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdref (lvar : ltac_var_map) (loc, t) =
let inh_conv_coerce_to_tycon = inh_conv_coerce_to_tycon resolve_tc in
let pretype_type = pretype_type k0 resolve_tc in
let pretype = pretype k0 resolve_tc in
let open Context.Rel.Declaration in
match t with
- | GRef (loc,ref,u) ->
+ | GRef (ref,u) ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_ref loc evdref env ref u)
tycon
- | GVar (loc, id) ->
+ | GVar id ->
inh_conv_coerce_to_tycon loc env evdref
(pretype_id (fun e r l t -> pretype tycon e r l t) k0 loc env evdref lvar id)
tycon
- | GEvar (loc, id, inst) ->
+ | GEvar (id, inst) ->
(* Ne faudrait-il pas s'assurer que hyps est bien un
sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
let evk =
@@ -596,7 +596,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let j = (Retyping.get_judgment_of env.ExtraEnv.env !evdref c) in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | GPatVar (loc,(someta,n)) ->
+ | GPatVar (someta,n) ->
let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
@@ -605,7 +605,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let k = Evar_kinds.MatchingVar (someta,n) in
{ uj_val = e_new_evar env evdref ~src:(loc,k) ty; uj_type = ty }
- | GHole (loc, k, naming, None) ->
+ | GHole (k, naming, None) ->
let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
@@ -614,7 +614,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
new_type_evar env evdref loc in
{ uj_val = e_new_evar env evdref ~src:(loc,k) ~naming ty; uj_type = ty }
- | GHole (loc, k, _naming, Some arg) ->
+ | GHole (k, _naming, Some arg) ->
let env = ltac_interp_name_env k0 lvar env !evdref in
let ty =
match tycon with
@@ -626,7 +626,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let () = evdref := sigma in
{ uj_val = c; uj_type = ty }
- | GRec (loc,fixkind,names,bl,lar,vdef) ->
+ | GRec (fixkind,names,bl,lar,vdef) ->
let rec type_bl env ctxt = function
[] -> ctxt
| (na,bk,None,ty)::bl ->
@@ -711,11 +711,11 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
in
inh_conv_coerce_to_tycon loc env evdref fixj tycon
- | GSort (loc,s) ->
+ | GSort s ->
let j = pretype_sort loc evdref s in
inh_conv_coerce_to_tycon loc env evdref j tycon
- | GApp (loc,f,args) ->
+ | GApp (f,args) ->
let fj = pretype empty_tycon env evdref lvar f in
let floc = loc_of_glob_constr f in
let length = List.length args in
@@ -794,7 +794,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | GLambda(loc,name,bk,c1,c2) ->
+ | GLambda(name,bk,c1,c2) ->
let tycon' = evd_comb1
(fun evd tycon ->
match tycon with
@@ -816,7 +816,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let resj = judge_of_abstraction env.ExtraEnv.env (orelse_name name name') j j' in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | GProd(loc,name,bk,c1,c2) ->
+ | GProd(name,bk,c1,c2) ->
let j = pretype_type empty_valcon env evdref lvar c1 in
(* The name specified by ltac is used also to create bindings. So
the substitution must also be applied on variables before they are
@@ -840,7 +840,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
iraise (e, info) in
inh_conv_coerce_to_tycon loc env evdref resj tycon
- | GLetIn(loc,name,c1,t,c2) ->
+ | GLetIn(name,c1,t,c2) ->
let tycon1 =
match t with
| Some t ->
@@ -861,7 +861,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
{ uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
uj_type = subst1 j.uj_val j'.uj_type }
- | GLetTuple (loc,nal,(na,po),c,d) ->
+ | GLetTuple (nal,(na,po),c,d) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env.ExtraEnv.env !evdref cj.uj_type
@@ -954,7 +954,7 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
obj ind p cj.uj_val fj.uj_val
in { uj_val = v; uj_type = ccl })
- | GIf (loc,c,(na,po),b1,b2) ->
+ | GIf (c,(na,po),b1,b2) ->
let cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
try find_rectype env.ExtraEnv.env !evdref cj.uj_type
@@ -1022,12 +1022,12 @@ let rec pretype k0 resolve_tc (tycon : type_constraint) (env : ExtraEnv.t) evdre
let cj = { uj_val = v; uj_type = p } in
inh_conv_coerce_to_tycon loc env evdref cj tycon
- | GCases (loc,sty,po,tml,eqns) ->
+ | GCases (sty,po,tml,eqns) ->
Cases.compile_cases loc sty
((fun vtyc env evdref -> pretype vtyc (make_env env !evdref) evdref lvar),evdref)
tycon env.ExtraEnv.env (* loc *) (po,tml,eqns)
- | GCast (loc,c,k) ->
+ | GCast (c,k) ->
let cj =
match k with
| CastCoerce ->
@@ -1097,7 +1097,7 @@ and pretype_instance k0 resolve_tc env evdref lvar loc hyps evk update =
(* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
and pretype_type k0 resolve_tc valcon (env : ExtraEnv.t) evdref lvar = function
- | GHole (loc, knd, naming, None) ->
+ | loc, GHole (knd, naming, None) ->
let rec is_Type c = match EConstr.kind !evdref c with
| Sort s ->
begin match ESorts.kind !evdref s with
diff --git a/tactics/hipattern.ml b/tactics/hipattern.ml
index 851554b83..2c331ba56 100644
--- a/tactics/hipattern.ml
+++ b/tactics/hipattern.ml
@@ -253,16 +253,16 @@ open Decl_kinds
open Evar_kinds
let mkPattern c = snd (Patternops.pattern_of_glob_constr c)
-let mkGApp f args = GApp (Loc.ghost, f, args)
-let mkGHole =
- GHole (Loc.ghost, QuestionMark (Define false), Misctypes.IntroAnonymous, None)
-let mkGProd id c1 c2 =
- GProd (Loc.ghost, Name (Id.of_string id), Explicit, c1, c2)
-let mkGArrow c1 c2 =
- GProd (Loc.ghost, Anonymous, Explicit, c1, c2)
-let mkGVar id = GVar (Loc.ghost, Id.of_string id)
-let mkGPatVar id = GPatVar(Loc.ghost, (false, Id.of_string id))
-let mkGRef r = GRef (Loc.ghost, Lazy.force r, None)
+let mkGApp f args = Loc.tag @@ GApp (f, args)
+let mkGHole = Loc.tag @@
+ GHole (QuestionMark (Define false), Misctypes.IntroAnonymous, None)
+let mkGProd id c1 c2 = Loc.tag @@
+ GProd (Name (Id.of_string id), Explicit, c1, c2)
+let mkGArrow c1 c2 = Loc.tag @@
+ GProd (Anonymous, Explicit, c1, c2)
+let mkGVar id = Loc.tag @@ GVar (Id.of_string id)
+let mkGPatVar id = Loc.tag @@ GPatVar((false, Id.of_string id))
+let mkGRef r = Loc.tag @@ GRef (Lazy.force r, None)
let mkGAppRef r args = mkGApp (mkGRef r) args
(** forall x : _, _ x x *)
diff --git a/vernac/command.ml b/vernac/command.ml
index 1f1464856..446afb578 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -422,13 +422,13 @@ let prepare_param = function
let rec check_anonymous_type ind =
let open Glob_term in
- match ind with
- | GSort (_, GType []) -> true
- | GProd (_, _, _, _, e)
- | GLetIn (_, _, _, _, e)
- | GLambda (_, _, _, _, e)
- | GApp (_, e, _)
- | GCast (_, e, _) -> check_anonymous_type e
+ match snd ind with
+ | GSort (GType []) -> true
+ | GProd ( _, _, _, e)
+ | GLetIn (_, _, _, e)
+ | GLambda (_, _, _, e)
+ | GApp (e, _)
+ | GCast (e, _) -> check_anonymous_type e
| _ -> false
let make_conclusion_flexible evdref ty poly =