aboutsummaryrefslogtreecommitdiffhomepage
diff options
context:
space:
mode:
-rw-r--r--ide/texmacspp.ml58
-rw-r--r--interp/constrexpr_ops.ml126
-rw-r--r--interp/constrexpr_ops.mli4
-rw-r--r--interp/constrextern.ml78
-rw-r--r--interp/constrintern.ml110
-rw-r--r--interp/implicit_quantifiers.ml22
-rw-r--r--interp/topconstr.ml94
-rw-r--r--intf/constrexpr.mli60
-rw-r--r--lib/loc.ml5
-rw-r--r--lib/loc.mli5
-rw-r--r--parsing/egramcoq.ml18
-rw-r--r--parsing/g_constr.ml4111
-rw-r--r--parsing/g_proofs.ml42
-rw-r--r--parsing/g_vernac.ml424
-rw-r--r--plugins/decl_mode/g_decl_mode.ml4387
-rw-r--r--plugins/funind/glob_term_to_relation.ml31
-rw-r--r--plugins/funind/indfun.ml105
-rw-r--r--plugins/funind/merge.ml4
-rw-r--r--plugins/ltac/g_ltac.ml48
-rw-r--r--plugins/ltac/g_tactic.ml412
-rw-r--r--plugins/ltac/pptactic.ml4
-rw-r--r--plugins/ltac/rewrite.ml12
-rw-r--r--plugins/ltac/tacintern.ml10
-rw-r--r--plugins/ltac/tacinterp.ml2
-rw-r--r--plugins/ssrmatching/ssrmatching.ml432
-rw-r--r--printing/ppconstr.ml117
-rw-r--r--printing/ppvernac.ml4
-rw-r--r--vernac/classes.ml8
-rw-r--r--vernac/command.ml19
-rw-r--r--vernac/metasyntax.ml4
-rw-r--r--vernac/record.ml7
31 files changed, 933 insertions, 550 deletions
diff --git a/ide/texmacspp.ml b/ide/texmacspp.ml
index 19be13be7..77dca0d06 100644
--- a/ide/texmacspp.ml
+++ b/ide/texmacspp.ml
@@ -219,7 +219,7 @@ let rec pp_bindlist bl =
(fun (loc, name) ->
xmlCst (string_of_name name) loc) loc_names) in
match e with
- | CHole _ -> names
+ | _, CHole _ -> names
| _ -> names @ [pp_expr e])
bl) in
match tlist with
@@ -231,7 +231,9 @@ and pp_local_binder lb = (* don't know what it is for now *)
match lb with
| CLocalDef ((loc, nam), ce, ty) ->
let attrs = ["name", string_of_name nam] in
- let value = match ty with Some t -> CCast (Loc.merge (constr_loc ce) (constr_loc t),ce, CastConv t) | None -> ce in
+ let value = match ty with
+ Some t -> Loc.tag ~loc:(Loc.merge (constr_loc ce) (constr_loc t)) @@ CCast (ce, CastConv t)
+ | None -> ce in
pp_expr ~attr:attrs value
| CLocalAssum (namll, _, ce) ->
let ppl =
@@ -388,42 +390,42 @@ and pp_local_binder_list lbl =
and pp_const_expr_list cel =
let l = List.map pp_expr cel in
Element ("recurse", (backstep_loc l), l)
-and pp_expr ?(attr=[]) e =
+and pp_expr ?(attr=[]) (loc, e) =
match e with
| CRef (r, _) ->
xmlCst ~attr
(Libnames.string_of_reference r) (Libnames.loc_of_reference r)
- | CProdN (loc, bl, e) ->
+ | CProdN (bl, e) ->
xmlApply loc
(xmlOperator "forall" loc :: [pp_bindlist bl] @ [pp_expr e])
- | CApp (loc, (_, hd), args) ->
+ | CApp ((_, hd), args) ->
xmlApply ~attr loc (pp_expr hd :: List.map (fun (e,_) -> pp_expr e) args)
- | CAppExpl (loc, (_, r, _), args) ->
+ | CAppExpl ((_, r, _), args) ->
xmlApply ~attr loc
(xmlCst (Libnames.string_of_reference r)
(Libnames.loc_of_reference r) :: List.map pp_expr args)
- | CNotation (loc, notation, ([],[],[])) ->
+ | CNotation (notation, ([],[],[])) ->
xmlOperator notation loc
- | CNotation (loc, notation, (args, cell, lbll)) ->
+ | CNotation (notation, (args, cell, lbll)) ->
let fmts = Notation.find_notation_extra_printing_rules notation in
let oper = xmlOperator notation loc ~pprules:fmts in
let cels = List.map pp_const_expr_list cell in
let lbls = List.map pp_local_binder_list lbll in
let args = List.map pp_expr args in
xmlApply loc (oper :: (List.sort compare_begin_att (args @ cels @ lbls)))
- | CSort(loc, s) ->
+ | CSort(s) ->
xmlOperator (string_of_glob_sort s) loc
- | CDelimiters (loc, scope, ce) ->
+ | CDelimiters (scope, ce) ->
xmlApply loc (xmlOperator "delimiter" ~attr:["name", scope] loc ::
[pp_expr ce])
- | CPrim (loc, tok) -> pp_token loc tok
- | CGeneralization (loc, kind, _, e) ->
+ | CPrim tok -> pp_token loc tok
+ | CGeneralization (kind, _, e) ->
let kind= match kind with
| Explicit -> "explicit"
| Implicit -> "implicit" in
xmlApply loc
(xmlOperator "generalization" ~attr:["kind", kind] loc :: [pp_expr e])
- | CCast (loc, e, tc) ->
+ | CCast (e, tc) ->
begin match tc with
| CastConv t | CastVM t |CastNative t ->
xmlApply loc
@@ -434,21 +436,21 @@ and pp_expr ?(attr=[]) e =
(xmlOperator ":" loc ~attr:["kind", "CastCoerce"] ::
[pp_expr e])
end
- | CEvar (loc, ek, cel) ->
+ | CEvar (ek, cel) ->
let ppcel = List.map (fun (id,e) -> xmlAssign id (pp_expr e)) cel in
xmlApply loc
(xmlOperator "evar" loc ~attr:["id", string_of_id ek] ::
ppcel)
- | CPatVar (loc, id) -> xmlPatvar (string_of_id id) loc
- | CHole (loc, _, _, _) -> xmlCst ~attr "_" loc
- | CIf (loc, test, (_, ret), th, el) ->
+ | CPatVar id -> xmlPatvar (string_of_id id) loc
+ | CHole (_, _, _) -> xmlCst ~attr "_" loc
+ | CIf (test, (_, ret), th, el) ->
let return = match ret with
| None -> []
| Some r -> [xmlReturn [pp_expr r]] in
xmlApply loc
(xmlOperator "if" loc ::
return @ [pp_expr th] @ [pp_expr el])
- | CLetTuple (loc, names, (_, ret), value, body) ->
+ | CLetTuple (names, (_, ret), value, body) ->
let return = match ret with
| None -> []
| Some r -> [xmlReturn [pp_expr r]] in
@@ -457,7 +459,7 @@ and pp_expr ?(attr=[]) e =
return @
(List.map (fun (loc, var) -> xmlCst (string_of_name var) loc) names) @
[pp_expr value; pp_expr body])
- | CCases (loc, sty, ret, cel, bel) ->
+ | CCases (sty, ret, cel, bel) ->
let return = match ret with
| None -> []
| Some r -> [xmlReturn [pp_expr r]] in
@@ -466,17 +468,21 @@ and pp_expr ?(attr=[]) e =
(return @
[Element ("scrutinees", [], List.map pp_case_expr cel)] @
[pp_branch_expr_list bel]))
- | CRecord (_, _) -> assert false
- | CLetIn (loc, (varloc, var), value, typ, body) ->
- let value = match typ with Some t -> CCast (Loc.merge (constr_loc value) (constr_loc t),value, CastConv t) | None -> value in
+ | CRecord _ -> assert false
+
+ | CLetIn ((varloc, var), value, typ, body) ->
+ let (loc, value) = match typ with
+ | Some t ->
+ Loc.tag ~loc:(Loc.merge (constr_loc value) (constr_loc t)) (CCast (value, CastConv t))
+ | None -> value in
xmlApply loc
(xmlOperator "let" loc ::
- [xmlCst (string_of_name var) varloc; pp_expr value; pp_expr body])
- | CLambdaN (loc, bl, e) ->
+ [xmlCst (string_of_name var) varloc; pp_expr (Loc.tag ~loc value); pp_expr body])
+ | CLambdaN (bl, e) ->
xmlApply loc
(xmlOperator "lambda" loc :: [pp_bindlist bl] @ [pp_expr e])
- | CCoFix (_, _, _) -> assert false
- | CFix (loc, lid, fel) ->
+ | CCoFix (_, _) -> assert false
+ | CFix (lid, fel) ->
xmlApply loc
(xmlOperator "fix" loc ::
List.flatten (List.map
diff --git a/interp/constrexpr_ops.ml b/interp/constrexpr_ops.ml
index 3ba5d985f..4f23dd2ab 100644
--- a/interp/constrexpr_ops.ml
+++ b/interp/constrexpr_ops.ml
@@ -97,79 +97,79 @@ let eq_universes u1 u2 =
| Some l, Some l' -> l = l'
| _, _ -> false
-let rec constr_expr_eq e1 e2 =
+let rec constr_expr_eq (_loc1, e1) (_loc2, e2) =
if e1 == e2 then true
else match e1, e2 with
| CRef (r1,u1), CRef (r2,u2) -> eq_reference r1 r2 && eq_universes u1 u2
- | CFix(_,id1,fl1), CFix(_,id2,fl2) ->
+ | CFix(id1,fl1), CFix(id2,fl2) ->
eq_located Id.equal id1 id2 &&
List.equal fix_expr_eq fl1 fl2
- | CCoFix(_,id1,fl1), CCoFix(_,id2,fl2) ->
+ | CCoFix(id1,fl1), CCoFix(id2,fl2) ->
eq_located Id.equal id1 id2 &&
List.equal cofix_expr_eq fl1 fl2
- | CProdN(_,bl1,a1), CProdN(_,bl2,a2) ->
+ | CProdN(bl1,a1), CProdN(bl2,a2) ->
List.equal binder_expr_eq bl1 bl2 &&
constr_expr_eq a1 a2
- | CLambdaN(_,bl1,a1), CLambdaN(_,bl2,a2) ->
+ | CLambdaN(bl1,a1), CLambdaN(bl2,a2) ->
List.equal binder_expr_eq bl1 bl2 &&
constr_expr_eq a1 a2
- | CLetIn(_,(_,na1),a1,t1,b1), CLetIn(_,(_,na2),a2,t2,b2) ->
+ | CLetIn((_,na1),a1,t1,b1), CLetIn((_,na2),a2,t2,b2) ->
Name.equal na1 na2 &&
constr_expr_eq a1 a2 &&
Option.equal constr_expr_eq t1 t2 &&
constr_expr_eq b1 b2
- | CAppExpl(_,(proj1,r1,_),al1), CAppExpl(_,(proj2,r2,_),al2) ->
+ | CAppExpl((proj1,r1,_),al1), CAppExpl((proj2,r2,_),al2) ->
Option.equal Int.equal proj1 proj2 &&
eq_reference r1 r2 &&
List.equal constr_expr_eq al1 al2
- | CApp(_,(proj1,e1),al1), CApp(_,(proj2,e2),al2) ->
+ | CApp((proj1,e1),al1), CApp((proj2,e2),al2) ->
Option.equal Int.equal proj1 proj2 &&
constr_expr_eq e1 e2 &&
List.equal args_eq al1 al2
- | CRecord (_, l1), CRecord (_, l2) ->
+ | CRecord l1, CRecord l2 ->
let field_eq (r1, e1) (r2, e2) =
eq_reference r1 r2 && constr_expr_eq e1 e2
in
List.equal field_eq l1 l2
- | CCases(_,_,r1,a1,brl1), CCases(_,_,r2,a2,brl2) ->
+ | CCases(_,r1,a1,brl1), CCases(_,r2,a2,brl2) ->
(** Don't care about the case_style *)
Option.equal constr_expr_eq r1 r2 &&
List.equal case_expr_eq a1 a2 &&
List.equal branch_expr_eq brl1 brl2
- | CLetTuple (_, n1, (m1, e1), t1, b1), CLetTuple (_, n2, (m2, e2), t2, b2) ->
+ | CLetTuple (n1, (m1, e1), t1, b1), CLetTuple (n2, (m2, e2), t2, b2) ->
List.equal (eq_located Name.equal) n1 n2 &&
Option.equal (eq_located Name.equal) m1 m2 &&
Option.equal constr_expr_eq e1 e2 &&
constr_expr_eq t1 t2 &&
constr_expr_eq b1 b2
- | CIf (_, e1, (n1, r1), t1, f1), CIf (_, e2, (n2, r2), t2, f2) ->
+ | CIf (e1, (n1, r1), t1, f1), CIf (e2, (n2, r2), t2, f2) ->
constr_expr_eq e1 e2 &&
Option.equal (eq_located Name.equal) n1 n2 &&
Option.equal constr_expr_eq r1 r2 &&
constr_expr_eq t1 t2 &&
constr_expr_eq f1 f2
| CHole _, CHole _ -> true
- | CPatVar(_,i1), CPatVar(_,i2) ->
+ | CPatVar i1, CPatVar i2 ->
Id.equal i1 i2
- | CEvar (_, id1, c1), CEvar (_, id2, c2) ->
+ | CEvar (id1, c1), CEvar (id2, c2) ->
Id.equal id1 id2 && List.equal instance_eq c1 c2
- | CSort(_,s1), CSort(_,s2) ->
+ | CSort s1, CSort s2 ->
Miscops.glob_sort_eq s1 s2
- | CCast(_,a1,(CastConv b1|CastVM b1)), CCast(_,a2,(CastConv b2|CastVM b2)) ->
+ | CCast(a1,(CastConv b1|CastVM b1)), CCast(a2,(CastConv b2|CastVM b2)) ->
constr_expr_eq a1 a2 &&
constr_expr_eq b1 b2
- | CCast(_,a1,CastCoerce), CCast(_,a2, CastCoerce) ->
+ | CCast(a1,CastCoerce), CCast(a2, CastCoerce) ->
constr_expr_eq a1 a2
- | CNotation(_, n1, s1), CNotation(_, n2, s2) ->
+ | CNotation(n1, s1), CNotation(n2, s2) ->
String.equal n1 n2 &&
constr_notation_substitution_eq s1 s2
- | CPrim(_,i1), CPrim(_,i2) ->
+ | CPrim i1, CPrim i2 ->
prim_token_eq i1 i2
- | CGeneralization (_, bk1, ak1, e1), CGeneralization (_, bk2, ak2, e2) ->
+ | CGeneralization (bk1, ak1, e1), CGeneralization (bk2, ak2, e2) ->
binding_kind_eq bk1 bk2 &&
Option.equal abstraction_kind_eq ak1 ak2 &&
constr_expr_eq e1 e2
- | CDelimiters(_,s1,e1), CDelimiters(_,s2,e2) ->
+ | CDelimiters(s1,e1), CDelimiters(s2,e2) ->
String.equal s1 s2 &&
constr_expr_eq e1 e2
| _ -> false
@@ -228,29 +228,7 @@ and constr_notation_substitution_eq (e1, el1, bl1) (e2, el2, bl2) =
and instance_eq (x1,c1) (x2,c2) =
Id.equal x1 x2 && constr_expr_eq c1 c2
-let constr_loc = function
- | CRef (Ident (loc,_),_) -> loc
- | CRef (Qualid (loc,_),_) -> loc
- | CFix (loc,_,_) -> loc
- | CCoFix (loc,_,_) -> loc
- | CProdN (loc,_,_) -> loc
- | CLambdaN (loc,_,_) -> loc
- | CLetIn (loc,_,_,_,_) -> loc
- | CAppExpl (loc,_,_) -> loc
- | CApp (loc,_,_) -> loc
- | CRecord (loc,_) -> loc
- | CCases (loc,_,_,_,_) -> loc
- | CLetTuple (loc,_,_,_,_) -> loc
- | CIf (loc,_,_,_,_) -> loc
- | CHole (loc,_,_,_) -> loc
- | CPatVar (loc,_) -> loc
- | CEvar (loc,_,_) -> loc
- | CSort (loc,_) -> loc
- | CCast (loc,_,_) -> loc
- | CNotation (loc,_,_) -> loc
- | CGeneralization (loc,_,_,_) -> loc
- | CPrim (loc,_) -> loc
- | CDelimiters (loc,_,_) -> loc
+let constr_loc (l,_) = l
let cases_pattern_expr_loc (l,_) = l
@@ -270,18 +248,18 @@ let local_binders_loc bll = match bll with
(** Pseudo-constructors *)
-let mkIdentC id = CRef (Ident (Loc.ghost, id),None)
-let mkRefC r = CRef (r,None)
-let mkCastC (a,k) = CCast (Loc.ghost,a,k)
-let mkLambdaC (idl,bk,a,b) = CLambdaN (Loc.ghost,[idl,bk,a],b)
-let mkLetInC (id,a,t,b) = CLetIn (Loc.ghost,id,a,t,b)
-let mkProdC (idl,bk,a,b) = CProdN (Loc.ghost,[idl,bk,a],b)
+let mkIdentC id = Loc.tag @@ CRef (Ident (Loc.ghost, id),None)
+let mkRefC r = Loc.tag @@ CRef (r,None)
+let mkCastC (a,k) = Loc.tag @@ CCast (a,k)
+let mkLambdaC (idl,bk,a,b) = Loc.tag @@ CLambdaN ([idl,bk,a],b)
+let mkLetInC (id,a,t,b) = Loc.tag @@ CLetIn (id,a,t,b)
+let mkProdC (idl,bk,a,b) = Loc.tag @@ CProdN ([idl,bk,a],b)
let mkAppC (f,l) =
let l = List.map (fun x -> (x,None)) l in
match f with
- | CApp (_,g,l') -> CApp (Loc.ghost, g, l' @ l)
- | _ -> CApp (Loc.ghost, (None, f), l)
+ | _loc, CApp (g,l') -> Loc.tag @@ CApp (g, l' @ l)
+ | _ -> Loc.tag @@ CApp ((None, f), l)
let add_name_in_env env n =
match snd n with
@@ -290,47 +268,47 @@ let add_name_in_env env n =
let (fresh_var, fresh_var_hook) = Hook.make ~default:(fun _ _ -> assert false) ()
-let expand_binders mkC loc bl c =
- let rec loop loc bl c =
+let expand_binders ~loc mkC bl c =
+ let rec loop ~loc bl c =
match bl with
| [] -> ([], c)
| b :: bl ->
match b with
| CLocalDef ((loc1,_) as n, oty, b) ->
- let env, c = loop (Loc.merge loc1 loc) bl c in
+ let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in
let env = add_name_in_env env n in
- (env, CLetIn (loc,n,oty,b,c))
+ (env, Loc.tag ~loc @@ CLetIn (n,oty,b,c))
| CLocalAssum ((loc1,_)::_ as nl, bk, t) ->
- let env, c = loop (Loc.merge loc1 loc) bl c in
+ let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in
let env = List.fold_left add_name_in_env env nl in
- (env, mkC loc (nl,bk,t) c)
+ (env, mkC ~loc (nl,bk,t) c)
| CLocalAssum ([],_,_) -> loop loc bl c
| CLocalPattern (loc1, p, ty) ->
- let env, c = loop (Loc.merge loc1 loc) bl c in
+ let env, c = loop ~loc:(Loc.merge loc1 loc) bl c in
let ni = Hook.get fresh_var env c in
let id = (loc1, Name ni) in
let ty = match ty with
| Some ty -> ty
- | None -> CHole (loc1, None, IntroAnonymous, None)
+ | None -> Loc.tag ~loc:loc1 @@ CHole (None, IntroAnonymous, None)
in
- let e = CRef (Libnames.Ident (loc1, ni), None) in
- let c =
+ let e = Loc.tag @@ CRef (Libnames.Ident (loc1, ni), None) in
+ let c = Loc.tag ~loc @@
CCases
- (loc, LetPatternStyle, None, [(e,None,None)],
- [(loc1, ([(loc1,[p])], c))])
+ (LetPatternStyle, None, [(e,None,None)],
+ [(Loc.tag ~loc:loc1 ([(loc1,[p])], c))])
in
- (ni :: env, mkC loc ([id],Default Explicit,ty) c)
+ (ni :: env, mkC ~loc ([id],Default Explicit,ty) c)
in
let (_, c) = loop loc bl c in
c
-let mkCProdN loc bll c =
- let mk loc b c = CProdN (loc,[b],c) in
- expand_binders mk loc bll c
+let mkCProdN ~loc bll c =
+ let mk ~loc b c = Loc.tag ~loc @@ CProdN ([b],c) in
+ expand_binders ~loc mk bll c
-let mkCLambdaN loc bll c =
- let mk loc b c = CLambdaN (loc,[b],c) in
- expand_binders mk loc bll c
+let mkCLambdaN ~loc bll c =
+ let mk ~loc b c = Loc.tag ~loc @@ CLambdaN ([b],c) in
+ expand_binders ~loc mk bll c
(* Deprecated *)
let abstract_constr_expr c bl = mkCLambdaN (local_binders_loc bl) bl c
@@ -343,14 +321,14 @@ let coerce_reference_to_id = function
(str "This expression should be a simple identifier.")
let coerce_to_id = function
- | CRef (Ident (loc,id),_) -> (loc,id)
+ | _loc, CRef (Ident (loc,id),_) -> (loc,id)
| a -> CErrors.user_err ~loc:(constr_loc a)
~hdr:"coerce_to_id"
(str "This expression should be a simple identifier.")
let coerce_to_name = function
- | CRef (Ident (loc,id),_) -> (loc,Name id)
- | CHole (loc,_,_,_) -> (loc,Anonymous)
+ | _loc, CRef (Ident (loc,id),_) -> (loc,Name id)
+ | loc, CHole (_,_,_) -> (loc,Anonymous)
| a -> CErrors.user_err
~loc:(constr_loc a) ~hdr:"coerce_to_name"
(str "This expression should be a name.")
diff --git a/interp/constrexpr_ops.mli b/interp/constrexpr_ops.mli
index f6d97e107..ae5ec2be5 100644
--- a/interp/constrexpr_ops.mli
+++ b/interp/constrexpr_ops.mli
@@ -49,10 +49,10 @@ val mkLambdaC : Name.t located list * binder_kind * constr_expr * constr_expr ->
val mkLetInC : Name.t located * constr_expr * constr_expr option * constr_expr -> constr_expr
val mkProdC : Name.t located list * binder_kind * constr_expr * constr_expr -> constr_expr
-val mkCLambdaN : Loc.t -> local_binder_expr list -> constr_expr -> constr_expr
+val mkCLambdaN : loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr
(** Same as [abstract_constr_expr], with location *)
-val mkCProdN : Loc.t -> local_binder_expr list -> constr_expr -> constr_expr
+val mkCProdN : loc:Loc.t -> local_binder_expr list -> constr_expr -> constr_expr
(** Same as [prod_constr_expr], with location *)
(** @deprecated variant of mkCLambdaN *)
diff --git a/interp/constrextern.ml b/interp/constrextern.ml
index 7a229856e..255de8500 100644
--- a/interp/constrextern.ml
+++ b/interp/constrextern.ml
@@ -144,7 +144,7 @@ module PrintingConstructor = Goptions.MakeRefTable(PrintingRecordConstructor)
let insert_delimiters e = function
| None -> e
- | Some sc -> CDelimiters (Loc.ghost,sc,e)
+ | Some sc -> Loc.tag @@ CDelimiters (sc,e)
let insert_pat_delimiters loc p = function
| None -> p
@@ -157,7 +157,7 @@ let insert_pat_alias loc p = function
(**********************************************************************)
(* conversion of references *)
-let extern_evar loc n l = CEvar (loc,n,l)
+let extern_evar loc n l = Loc.tag @@ CEvar (n,l)
(** We allow customization of the global_reference printer.
For instance, in the debugger the tables of global references
@@ -236,7 +236,7 @@ let expand_curly_brackets loc mknot ntn l =
(* side effect *)
mknot (loc,!ntn',l)
-let destPrim = function CPrim(_,t) -> Some t | _ -> None
+let destPrim = function _loc, CPrim t -> Some t | _ -> None
let destPatPrim = function _loc, CPatPrim t -> Some t | _ -> None
let make_notation_gen loc ntn mknot mkprim destprim l =
@@ -259,11 +259,11 @@ let make_notation_gen loc ntn mknot mkprim destprim l =
let make_notation loc ntn (terms,termlists,binders as subst) =
if not (List.is_empty termlists) || not (List.is_empty binders) then
- CNotation (loc,ntn,subst)
+ Loc.tag ~loc @@ CNotation (ntn,subst)
else
make_notation_gen loc ntn
- (fun (loc,ntn,l) -> CNotation (loc,ntn,(l,[],[])))
- (fun (loc,p) -> CPrim (loc,p))
+ (fun (loc,ntn,l) -> Loc.tag ~loc @@ CNotation (ntn,(l,[],[])))
+ (fun (loc,p) -> Loc.tag ~loc @@ CPrim p)
destPrim terms
let make_pat_notation loc ntn (terms,termlists as subst) args =
@@ -462,11 +462,11 @@ let is_projection nargs = function
else None
with Not_found -> None)
| _ -> None
-
+
let is_hole = function CHole _ | CEvar _ -> true | _ -> false
let is_significant_implicit a =
- not (is_hole a)
+ not (is_hole (snd a))
let is_needed_for_correct_partial_application tail imp =
List.is_empty tail && not (maximal_insertion_of imp)
@@ -512,16 +512,16 @@ 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
- CApp (loc,(ip,f),args1@args2)
+ Loc.tag ~loc @@ CApp ((ip,f),args1@args2)
| None ->
let args = exprec 1 (args,impl) in
- if List.is_empty args then f else CApp (loc, (None, f), args)
+ if List.is_empty args then f else Loc.tag ~loc @@ CApp ((None, f), args)
in
try expl ()
with Expl ->
- let f',us = match f with 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
- CAppExpl (loc, (ip, f', us), List.map Lazy.force args)
+ Loc.tag ~loc @@ CAppExpl ((ip, f', us), List.map Lazy.force args)
let is_start_implicit = function
| imp :: _ -> is_status_implicit imp && maximal_insertion_of imp
@@ -530,23 +530,23 @@ let is_start_implicit = function
let extern_global loc impl f us =
if not !Constrintern.parsing_explicit && is_start_implicit impl
then
- CAppExpl (loc, (None, f, us), [])
+ Loc.tag ~loc @@ CAppExpl ((None, f, us), [])
else
- CRef (f,us)
+ Loc.tag ~loc @@ CRef (f,us)
let extern_app loc inctx impl (cf,f) us args =
if List.is_empty args then
(* If coming from a notation "Notation a := @b" *)
- CAppExpl (loc, (None, f, us), [])
+ Loc.tag ~loc @@ 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
- CAppExpl (loc, (is_projection (List.length args) cf,f,us), args)
+ Loc.tag ~loc @@ CAppExpl ((is_projection (List.length args) cf,f,us), args)
else
- explicitize loc inctx impl (cf,CRef (f,us)) args
+ explicitize loc inctx impl (cf, Loc.tag ~loc @@ CRef (f,us)) args
let rec fill_arg_scopes args subscopes scopes = match args, subscopes with
| [], _ -> []
@@ -600,7 +600,7 @@ let extern_possible_prim_token scopes r =
let (sc,n) = uninterp_prim_token r in
match availability_of_prim_token n sc scopes with
| None -> None
- | Some key -> Some (insert_delimiters (CPrim (loc_of_glob_constr r,n)) key)
+ | Some key -> Some (insert_delimiters (Loc.tag ~loc:(loc_of_glob_constr r) @@ CPrim n) key)
with No_match ->
None
@@ -608,7 +608,7 @@ let extern_optimal_prim_token scopes r r' =
let c = extern_possible_prim_token scopes r in
let c' = if r==r' then None else extern_possible_prim_token scopes r' in
match c,c' with
- | Some n, (Some (CDelimiters _) | None) | _, Some n -> n
+ | Some n, (Some ((_, CDelimiters _)) | None) | _, Some n -> n
| _ -> raise No_match
(**********************************************************************)
@@ -647,16 +647,16 @@ let rec extern inctx scopes vars r =
extern_global loc (select_stronger_impargs (implicits_of_global ref))
(extern_reference loc vars ref) (extern_universes us)
- | GVar (loc,id) -> CRef (Ident (loc,id),None)
+ | GVar (loc,id) -> Loc.tag ~loc @@ CRef (Ident (loc,id),None)
- | GEvar (loc,n,[]) when !print_meta_as_hole -> CHole (loc, None, Misctypes.IntroAnonymous, None)
+ | GEvar (loc,n,[]) when !print_meta_as_hole -> Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None)
| GEvar (loc,n,l) ->
extern_evar loc n (List.map (on_snd (extern false scopes vars)) l)
- | GPatVar (loc,(b,n)) ->
- if !print_meta_as_hole then CHole (loc, None, Misctypes.IntroAnonymous, None) else
- if b then CPatVar (loc,n) else CEvar (loc,n,[])
+ | GPatVar (loc,(b,n)) -> Loc.tag ~loc @@
+ if !print_meta_as_hole then CHole (None, Misctypes.IntroAnonymous, None) else
+ if b then CPatVar n else CEvar (n,[])
| GApp (loc,f,args) ->
(match f with
@@ -701,7 +701,7 @@ 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
- CRecord (loc, List.rev (ip projs locals args []))
+ Loc.tag ~loc @@ CRecord (List.rev (ip projs locals args []))
with
| Not_found | No_match | Exit ->
let args = extern_args (extern true) vars args in
@@ -715,19 +715,19 @@ let rec extern inctx scopes vars r =
(List.map (fun c -> lazy (sub_extern true scopes vars c)) args))
| GLetIn (loc,na,b,t,c) ->
- CLetIn (loc,(loc,na),sub_extern false scopes vars b,
+ Loc.tag ~loc @@ 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) ->
let t = extern_typ scopes vars t in
let (idl,c) = factorize_prod scopes (add_vname vars na) na bk t c in
- CProdN (loc,[(Loc.ghost,na)::idl,Default bk,t],c)
+ Loc.tag ~loc @@ CProdN ([(Loc.ghost,na)::idl,Default bk,t],c)
| GLambda (loc,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
- CLambdaN (loc,[(Loc.ghost,na)::idl,Default bk,t],c)
+ Loc.tag ~loc @@ CLambdaN ([(Loc.ghost,na)::idl,Default bk,t],c)
| GCases (loc,sty,rtntypopt,tml,eqns) ->
let vars' =
@@ -757,17 +757,17 @@ let rec extern inctx scopes vars r =
tml
in
let eqns = List.map (extern_eqn inctx scopes vars) eqns in
- CCases (loc,sty,rtntypopt',tml,eqns)
+ Loc.tag ~loc @@ CCases (sty,rtntypopt',tml,eqns)
| GLetTuple (loc,nal,(na,typopt),tm,b) ->
- CLetTuple (loc,List.map (fun na -> (Loc.ghost,na)) nal,
+ Loc.tag ~loc @@ 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) ->
- CIf (loc,sub_extern false scopes vars c,
+ Loc.tag ~loc @@ 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)
@@ -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
- CFix (loc,(loc,idv.(n)),Array.to_list listdecl)
+ Loc.tag ~loc @@ CFix ((loc,idv.(n)),Array.to_list listdecl)
| GCoFix n ->
let listdecl =
Array.mapi (fun i fi ->
@@ -803,14 +803,14 @@ 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
- CCoFix (loc,(loc,idv.(n)),Array.to_list listdecl))
+ Loc.tag ~loc @@ CCoFix ((loc,idv.(n)),Array.to_list listdecl))
- | GSort (loc,s) -> CSort (loc,extern_glob_sort s)
+ | GSort (loc,s) -> Loc.tag ~loc @@ CSort (extern_glob_sort s)
- | GHole (loc,e,naming,_) -> CHole (loc, Some e, naming, None) (** TODO: extern tactics. *)
+ | GHole (loc,e,naming,_) -> Loc.tag ~loc @@ CHole (Some e, naming, None) (** TODO: extern tactics. *)
| GCast (loc,c, c') ->
- CCast (loc,sub_extern true scopes vars c,
+ Loc.tag ~loc @@ CCast (sub_extern true scopes vars c,
Miscops.map_cast_type (extern_typ scopes vars) c')
and extern_typ (_,scopes) =
@@ -821,7 +821,7 @@ and sub_extern inctx (_,scopes) = extern inctx (None,scopes)
and factorize_prod scopes vars na bk aty c =
let c = extern_typ scopes vars c in
match na, c with
- | Name id, CProdN (loc,[nal,Default bk',ty],c)
+ | Name id, (loc, CProdN ([nal,Default bk',ty],c))
when binding_kind_eq bk bk' && constr_expr_eq aty ty
&& not (occur_var_constr_expr id ty) (* avoid na in ty escapes scope *) ->
nal,c
@@ -831,7 +831,7 @@ and factorize_prod scopes vars na bk aty c =
and factorize_lambda inctx scopes vars na bk aty c =
let c = sub_extern inctx scopes vars c in
match c with
- | CLambdaN (loc,[nal,Default bk',ty],c)
+ | loc, CLambdaN ([nal,Default bk',ty],c)
when binding_kind_eq bk bk' && constr_expr_eq aty ty
&& not (occur_name na ty) (* avoid na in ty escapes scope *) ->
nal,c
@@ -940,7 +940,7 @@ and extern_notation (tmp_scope,scopes as allscopes) vars t = function
extern true (scopt,scl@scopes) vars c, None)
terms in
let a = CRef (Qualid (loc, shortest_qualid_of_syndef vars kn),None) in
- if List.is_empty l then a else CApp (loc,(None,a),l) in
+ Loc.tag ~loc @@ if List.is_empty l then a else CApp ((None, Loc.tag a),l) in
if List.is_empty args then e
else
let args = fill_arg_scopes args argsscopes scopes in
diff --git a/interp/constrintern.ml b/interp/constrintern.ml
index 6bf6772c6..4af89e1ef 100644
--- a/interp/constrintern.ml
+++ b/interp/constrintern.ml
@@ -217,7 +217,7 @@ let contract_notation ntn (l,ll,bll) =
let ntn' = ref ntn in
let rec contract_squash n = function
| [] -> []
- | CNotation (_,"{ _ }",([a],[],[])) :: l ->
+ | (_loc, CNotation ("{ _ }",([a],[],[]))) :: l ->
ntn' := expand_notation_string !ntn' n;
contract_squash n (a::l)
| a :: l ->
@@ -407,7 +407,7 @@ let intern_generalized_binder ?(global_level=false) intern_type lvar
let name =
let id =
match ty with
- | CApp (_, (_, CRef (Ident (loc,id),_)), _) -> id
+ | _, CApp ((_, (_, CRef (Ident (loc,id),_))), _) -> id
| _ -> default_non_dependent_ident
in Implicit_quantifiers.make_fresh ids' (Global.env ()) id
in Name name
@@ -448,7 +448,7 @@ let intern_local_pattern intern lvar env p =
List.fold_left
(fun env (loc, i) ->
let bk = Default Implicit in
- let ty = CHole (loc, None, Misctypes.IntroAnonymous, None) in
+ let ty = Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None) in
let n = Name i in
let env, _ = intern_assumption intern lvar env [(loc, n)] bk ty in
env)
@@ -479,7 +479,7 @@ let intern_local_binder_aux ?(global_level=false) intern lvar (env,bl) = functio
let tyc =
match ty with
| Some ty -> ty
- | None -> CHole(loc,None,Misctypes.IntroAnonymous,None)
+ | None -> Loc.tag ~loc @@ CHole(None,Misctypes.IntroAnonymous,None)
in
let env = intern_local_pattern intern lvar env p in
let il = List.map snd (free_vars_of_pat [] p) in
@@ -592,15 +592,15 @@ let rec subordinate_letins letins = function
let terms_of_binders bl =
let rec term_of_pat = function
- | PatVar (loc,Name id) -> CRef (Ident (loc,id), None)
+ | PatVar (loc,Name id) -> Loc.tag ~loc @@ CRef (Ident (loc,id), None)
| PatVar (loc,Anonymous) -> error "Cannot turn \"_\" into a term."
| PatCstr (loc,c,l,_) ->
let r = Qualid (loc,qualid_of_path (path_of_global (ConstructRef c))) in
- let hole = CHole (loc,None,Misctypes.IntroAnonymous,None) in
+ let hole = Loc.tag ~loc @@ CHole (None,Misctypes.IntroAnonymous,None) in
let params = List.make (Inductiveops.inductive_nparams (fst c)) hole in
- CAppExpl (loc,(None,r,None),params @ List.map term_of_pat l) in
+ Loc.tag ~loc @@ CAppExpl ((None,r,None),params @ List.map term_of_pat l) in
let rec extract_variables = function
- | GLocalAssum (loc,Name id,_,_)::l -> CRef (Ident (loc,id), None) :: extract_variables l
+ | GLocalAssum (loc,Name id,_,_)::l -> (Loc.tag ~loc @@ CRef (Ident (loc,id), None)) :: extract_variables l
| GLocalDef (loc,Name id,_,_,_)::l -> extract_variables l
| GLocalDef (loc,Anonymous,_,_,_)::l
| GLocalAssum (loc,Anonymous,_,_)::l -> error "Cannot turn \"_\" into a term."
@@ -754,7 +754,7 @@ let intern_var genv (ltacvars,ntnvars) namedctx loc id us =
try
let ty,expl_impls,impls,argsc = Id.Map.find id genv.impls in
let expl_impls = List.map
- (fun id -> CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in
+ (fun id -> Loc.tag ~loc @@ CRef (Ident (loc,id),None), Some (loc,ExplByName id)) expl_impls in
let tys = string_of_ty ty in
Dumpglob.dump_reference loc "<>" (Id.to_string id) tys;
gvar (loc,id) us, make_implicits_list impls, argsc, expl_impls
@@ -1515,15 +1515,15 @@ let extract_explicit_arg imps args =
(* Main loop *)
let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
- let rec intern env = function
- | CRef (ref,us) as x ->
+ let rec intern env = Loc.with_loc (fun ~loc -> function
+ | CRef (ref,us) ->
let (c,imp,subscopes,l),_ =
- intern_applied_reference intern env (Environ.named_context globalenv)
- lvar us [] ref
+ intern_applied_reference intern env (Environ.named_context globalenv)
+ lvar us [] ref
in
- apply_impargs c env imp subscopes l (constr_loc x)
+ apply_impargs c env imp subscopes l loc
- | CFix (loc, (locid,iddef), dl) ->
+ | CFix ((locid,iddef), dl) ->
let lf = List.map (fun ((_, id),_,_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
@@ -1564,7 +1564,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
Array.map (fun (_,bl,_,_) -> bl) idl,
Array.map (fun (_,_,ty,_) -> ty) idl,
Array.map (fun (_,_,_,bd) -> bd) idl)
- | CCoFix (loc, (locid,iddef), dl) ->
+ | CCoFix ((locid,iddef), dl) ->
let lf = List.map (fun ((_, id),_,_,_) -> id) dl in
let dl = Array.of_list dl in
let n =
@@ -1589,33 +1589,33 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
Array.map (fun (bl,_,_) -> bl) idl,
Array.map (fun (_,ty,_) -> ty) idl,
Array.map (fun (_,_,bd) -> bd) idl)
- | CProdN (loc,[],c2) ->
+ | CProdN ([],c2) ->
intern_type env c2
- | CProdN (loc,(nal,bk,ty)::bll,c2) ->
- iterate_prod loc env bk ty (CProdN (loc, bll, c2)) nal
- | CLambdaN (loc,[],c2) ->
+ | CProdN ((nal,bk,ty)::bll,c2) ->
+ iterate_prod loc env bk ty (Loc.tag ~loc @@ CProdN (bll, c2)) nal
+ | CLambdaN ([],c2) ->
intern env c2
- | CLambdaN (loc,(nal,bk,ty)::bll,c2) ->
- iterate_lam loc (reset_tmp_scope env) bk ty (CLambdaN (loc, bll, c2)) nal
- | CLetIn (loc,na,c1,t,c2) ->
+ | CLambdaN ((nal,bk,ty)::bll,c2) ->
+ iterate_lam loc (reset_tmp_scope env) bk ty (Loc.tag ~loc @@ CLambdaN (bll, c2)) nal
+ | 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,
intern (push_name_env ntnvars (impls_term_list inc1) env na) c2)
- | CNotation (loc,"- _",([CPrim (_,Numeral p)],[],[]))
+ | CNotation ("- _",([_, CPrim (Numeral p)],[],[]))
when Bigint.is_strictly_pos p ->
- intern env (CPrim (loc,Numeral (Bigint.neg p)))
- | CNotation (_,"( _ )",([a],[],[])) -> intern env a
- | CNotation (loc,ntn,args) ->
+ intern env (Loc.tag ~loc @@ CPrim (Numeral (Bigint.neg p)))
+ | CNotation ("( _ )",([a],[],[])) -> intern env a
+ | CNotation (ntn,args) ->
intern_notation intern env ntnvars loc ntn args
- | CGeneralization (loc,b,a,c) ->
+ | CGeneralization (b,a,c) ->
intern_generalization intern env ntnvars loc b a c
- | CPrim (loc, p) ->
+ | CPrim p ->
fst (Notation.interp_prim_token loc p (env.tmp_scope,env.scopes))
- | CDelimiters (loc, key, e) ->
+ | CDelimiters (key, e) ->
intern {env with tmp_scope = None;
scopes = find_delimiters_scope loc key :: env.scopes} e
- | CAppExpl (loc, (isproj,ref,us), args) ->
+ | CAppExpl ((isproj,ref,us), args) ->
let (f,_,args_scopes,_),args =
let args = List.map (fun a -> (a,None)) args in
intern_applied_reference intern env (Environ.named_context globalenv)
@@ -1624,42 +1624,42 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(* Rem: GApp(_,f,[]) stands for @f *)
GApp (loc, f, intern_args env args_scopes (List.map fst args))
- | CApp (loc, (isproj,f), args) ->
+ | CApp ((isproj,f), args) ->
let f,args = match f with
(* Compact notations like "t.(f args') args" *)
- | CApp (_,(Some _,f), args') when not (Option.has_some isproj) ->
+ | _loc, CApp ((Some _,f), args') when not (Option.has_some isproj) ->
f,args'@args
(* Don't compact "(f args') args" to resolve implicits separately *)
| _ -> f,args in
let (c,impargs,args_scopes,l),args =
match f with
- | CRef (ref,us) ->
+ | _loc, CRef (ref,us) ->
intern_applied_reference intern env
(Environ.named_context globalenv) lvar us args ref
- | CNotation (loc,ntn,([],[],[])) ->
+ | _loc, CNotation (ntn,([],[],[])) ->
let c = intern_notation intern env ntnvars loc ntn ([],[],[]) in
let x, impl, scopes, l = find_appl_head_data c in
(x,impl,scopes,l), args
| x -> (intern env f,[],[],[]), args in
- apply_impargs c env impargs args_scopes
+ apply_impargs c env impargs args_scopes
(merge_impargs l args) loc
- | CRecord (loc, fs) ->
+ | CRecord fs ->
let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in
let fields =
sort_fields ~complete:true loc fs
- (fun _idx -> CHole (loc, Some (Evar_kinds.QuestionMark st),
- Misctypes.IntroAnonymous, None))
+ (fun _idx -> Loc.tag ~loc @@ CHole (Some (Evar_kinds.QuestionMark st),
+ Misctypes.IntroAnonymous, None))
in
begin
match fields with
| None -> user_err ~loc ~hdr:"intern" (str"No constructor inference.")
| Some (n, constrname, args) ->
- let pars = List.make n (CHole (loc, None, Misctypes.IntroAnonymous, None)) in
- let app = CAppExpl (loc, (None, constrname,None), List.rev_append pars args) in
+ let pars = List.make n (Loc.tag ~loc @@ CHole (None, Misctypes.IntroAnonymous, None)) in
+ let app = Loc.tag ~loc @@ CAppExpl ((None, constrname,None), List.rev_append pars args) in
intern env app
end
- | CCases (loc, sty, rtnpo, tms, eqns) ->
+ | CCases (sty, rtnpo, tms, eqns) ->
let as_in_vars = List.fold_left (fun acc (_,na,inb) ->
Option.fold_left (fun acc tt -> Id.Set.union (ids_of_cases_indtype tt) acc)
(Option.fold_left (fun acc (_,y) -> name_fold Id.Set.add y acc) acc na)
@@ -1701,7 +1701,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
in
let eqns' = List.map (intern_eqn (List.length tms) env) eqns in
GCases (loc, sty, rtnpo, tms, List.flatten eqns')
- | CLetTuple (loc, nal, (na,po), b, c) ->
+ | CLetTuple (nal, (na,po), b, c) ->
let env' = reset_tmp_scope env in
(* "in" is None so no match to add *)
let ((b',(na',_)),_,_) = intern_case_item env' Id.Set.empty (b,na,None) in
@@ -1711,7 +1711,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
intern_type env'' u) po in
GLetTuple (loc, 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 (loc, c, (na,po), b1, b2) ->
+ | CIf (c, (na,po), b1, b2) ->
let env' = reset_tmp_scope env in
let ((c',(na',_)),_,_) = intern_case_item env' Id.Set.empty (c,na,None) in (* no "in" no match to ad too *)
let p' = Option.map (fun p ->
@@ -1719,7 +1719,7 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
(Loc.ghost,na') in
intern_type env'' p) po in
GIf (loc, c', (na', p'), intern env b1, intern env b2)
- | CHole (loc, k, naming, solve) ->
+ | CHole (k, naming, solve) ->
let k = match k with
| None ->
let st = Evar_kinds.Define (not (Program.get_proofs_transparency ())) in
@@ -1745,22 +1745,22 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
in
GHole (loc, k, naming, solve)
(* Parsing pattern variables *)
- | CPatVar (loc, n) when allow_patvar ->
+ | CPatVar n when allow_patvar ->
GPatVar (loc, (true,n))
- | CEvar (loc, n, []) when allow_patvar ->
+ | CEvar (n, []) when allow_patvar ->
GPatVar (loc, (false,n))
(* end *)
(* Parsing existential variables *)
- | CEvar (loc, n, l) ->
+ | CEvar (n, l) ->
GEvar (loc, n, List.map (on_snd (intern env)) l)
- | CPatVar (loc, _) ->
+ | CPatVar _ ->
raise (InternalizationError (loc,IllegalMetavariable))
(* end *)
- | CSort (loc, s) ->
+ | CSort s ->
GSort(loc,s)
- | CCast (loc, c1, c2) ->
+ | CCast (c1, c2) ->
GCast (loc,intern env c1, Miscops.map_cast_type (intern_type env) c2)
-
+ )
and intern_type env = intern (set_type_scope env)
and intern_local_binder env bind =
@@ -1887,17 +1887,17 @@ let internalize globalenv env allow_patvar (_, ntnvars as lvar) c =
intern_args env subscopes rargs
in aux 1 l subscopes eargs rargs
- and apply_impargs c env imp subscopes l loc =
+ and apply_impargs c env imp subscopes l loc =
let imp = select_impargs_size (List.length (List.filter (fun (_,x) -> x == None) l)) imp in
let l = intern_impargs c env imp subscopes l in
smart_gapp c loc l
and smart_gapp f loc = function
| [] -> f
- | l -> match f with
+ | 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)
-
+
and intern_args env subscopes = function
| [] -> []
| a::args ->
diff --git a/interp/implicit_quantifiers.ml b/interp/implicit_quantifiers.ml
index 7f11c0a3b..d2bebfb54 100644
--- a/interp/implicit_quantifiers.ml
+++ b/interp/implicit_quantifiers.ml
@@ -92,11 +92,11 @@ let free_vars_of_constr_expr c ?(bound=Id.Set.empty) l =
else ungeneralizable loc id
else l
in
- let rec aux bdvars l c = match c with
+ let rec aux bdvars l (loc, c) = match c with
| CRef (Ident (loc,id),_) -> found loc id bdvars l
- | CNotation (_, "{ _ : _ | _ }", (CRef (Ident (_, id),_) :: _, [], [])) when not (Id.Set.mem id bdvars) ->
- Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l c
- | c -> Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l c
+ | CNotation ("{ _ : _ | _ }", ((_, CRef (Ident (_, id),_)) :: _, [], [])) when not (Id.Set.mem id bdvars) ->
+ Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux (Id.Set.add id bdvars) l (loc, c)
+ | c -> Topconstr.fold_constr_expr_with_binders (fun a l -> Id.Set.add a l) aux bdvars l (loc, c)
in aux bound l c
let ids_of_names l =
@@ -251,18 +251,18 @@ let combine_params avoid fn applied needed =
let combine_params_freevar =
fun avoid (_, decl) ->
let id' = next_name_away_from (RelDecl.get_name decl) avoid in
- (CRef (Ident (Loc.ghost, id'),None), Id.Set.add id' avoid)
+ (Loc.tag @@ CRef (Ident (Loc.ghost, id'),None), Id.Set.add id' avoid)
-let destClassApp cl =
+let destClassApp (loc, cl) =
match cl with
- | CApp (loc, (None, CRef (ref, inst)), l) -> loc, ref, List.map fst l, inst
- | CAppExpl (loc, (None, ref, inst), l) -> loc, ref, l, inst
+ | CApp ((None, (_loc, CRef (ref, inst))), l) -> loc, ref, List.map fst l, inst
+ | CAppExpl ((None, ref, inst), l) -> loc, ref, l, inst
| CRef (ref, inst) -> loc_of_reference ref, ref, [], inst
| _ -> raise Not_found
-let destClassAppExpl cl =
+let destClassAppExpl (loc, cl) =
match cl with
- | CApp (loc, (None, CRef (ref, inst)), l) -> loc, ref, l, inst
+ | CApp ((None, (_loc, CRef (ref, inst))), l) -> loc, ref, l, inst
| CRef (ref, inst) -> loc_of_reference ref, ref, [], inst
| _ -> raise Not_found
@@ -295,7 +295,7 @@ let implicit_application env ?(allow_partial=true) f ty =
end;
let pars = List.rev (List.combine ci rd) in
let args, avoid = combine_params avoid f par pars in
- CAppExpl (loc, (None, id, inst), args), avoid
+ Loc.tag ~loc @@ CAppExpl ((None, id, inst), args), avoid
in c, avoid
let implicits_of_glob_constr ?(with_products=true) l =
diff --git a/interp/topconstr.ml b/interp/topconstr.ml
index 172caa2ca..c3e341d74 100644
--- a/interp/topconstr.ml
+++ b/interp/topconstr.ml
@@ -103,49 +103,50 @@ let rec fold_local_binders g f n acc b = function
| [] ->
f n acc b
-let fold_constr_expr_with_binders g f n acc = function
- | CAppExpl (loc,(_,_,_),l) -> List.fold_left (f n) acc l
- | CApp (loc,(_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
- | CProdN (_,l,b) | CLambdaN (_,l,b) -> fold_constr_expr_binders g f n acc b l
- | CLetIn (_,na,a,t,b) ->
+let fold_constr_expr_with_binders g f n acc = Loc.with_loc (fun ~loc -> function
+ | CAppExpl ((_,_,_),l) -> List.fold_left (f n) acc l
+ | CApp ((_,t),l) -> List.fold_left (f n) (f n acc t) (List.map fst l)
+ | CProdN (l,b) | CLambdaN (l,b) -> fold_constr_expr_binders g f n acc b l
+ | CLetIn (na,a,t,b) ->
f (name_fold g (snd na) n) (Option.fold_left (f n) (f n acc a) t) b
- | CCast (loc,a,(CastConv b|CastVM b|CastNative b)) -> f n (f n acc a) b
- | CCast (loc,a,CastCoerce) -> f n acc a
- | CNotation (_,_,(l,ll,bll)) ->
+ | CCast (a,(CastConv b|CastVM b|CastNative b)) -> f n (f n acc a) b
+ | CCast (a,CastCoerce) -> f n acc a
+ | CNotation (_,(l,ll,bll)) ->
(* The following is an approximation: we don't know exactly if
an ident is binding nor to which subterms bindings apply *)
let acc = List.fold_left (f n) acc (l@List.flatten ll) in
- List.fold_left (fun acc bl -> fold_local_binders g f n acc (CHole (Loc.ghost,None,IntroAnonymous,None)) bl) acc bll
- | CGeneralization (_,_,_,c) -> f n acc c
- | CDelimiters (loc,_,a) -> f n acc a
+ List.fold_left (fun acc bl -> fold_local_binders g f n acc (Loc.tag @@ CHole (None,IntroAnonymous,None)) bl) acc bll
+ | CGeneralization (_,_,c) -> f n acc c
+ | CDelimiters (_,a) -> f n acc a
| CHole _ | CEvar _ | CPatVar _ | CSort _ | CPrim _ | CRef _ ->
acc
- | CRecord (loc,l) -> List.fold_left (fun acc (id, c) -> f n acc c) acc l
- | CCases (loc,sty,rtnpo,al,bl) ->
+ | CRecord l -> List.fold_left (fun acc (id, c) -> f n acc c) acc l
+ | CCases (sty,rtnpo,al,bl) ->
let ids = ids_of_cases_tomatch al in
let acc = Option.fold_left (f (Id.Set.fold g ids n)) acc rtnpo in
let acc = List.fold_left (f n) acc (List.map (fun (fst,_,_) -> fst) al) in
List.fold_right (fun (loc,(patl,rhs)) acc ->
let ids = ids_of_pattern_list patl in
f (Id.Set.fold g ids n) acc rhs) bl acc
- | CLetTuple (loc,nal,(ona,po),b,c) ->
+ | CLetTuple (nal,(ona,po),b,c) ->
let n' = List.fold_right (Loc.down_located (name_fold g)) nal n in
f (Option.fold_right (Loc.down_located (name_fold g)) ona n') (f n acc b) c
- | CIf (_,c,(ona,po),b1,b2) ->
+ | CIf (c,(ona,po),b1,b2) ->
let acc = f n (f n (f n acc b1) b2) c in
Option.fold_left
(f (Option.fold_right (Loc.down_located (name_fold g)) ona n)) acc po
- | CFix (loc,_,l) ->
+ | CFix (_,l) ->
let n' = List.fold_right (fun ((_,id),_,_,_,_) -> g id) l n in
List.fold_right (fun (_,(_,o),lb,t,c) acc ->
fold_local_binders g f n'
(fold_local_binders g f n acc t lb) c lb) l acc
- | CCoFix (loc,_,_) ->
+ | CCoFix (_,_) ->
Feedback.msg_warning (strbrk "Capture check in multiple binders not done"); acc
+ )
let free_vars_of_constr_expr c =
let rec aux bdvars l = function
- | CRef (Ident (_,id),_) -> if Id.List.mem id bdvars then l else Id.Set.add id l
+ | _loc, CRef (Ident (_,id),_) -> if Id.List.mem id bdvars then l else Id.Set.add id l
| c -> fold_constr_expr_with_binders (fun a l -> a::l) aux bdvars l c
in aux [] Id.Set.empty c
@@ -209,60 +210,61 @@ let map_local_binders f g e bl =
let (e,rbl) = List.fold_left h (e,[]) bl in
(e, List.rev rbl)
-let map_constr_expr_with_binders g f e = function
- | CAppExpl (loc,r,l) -> CAppExpl (loc,r,List.map (f e) l)
- | CApp (loc,(p,a),l) ->
- CApp (loc,(p,f e a),List.map (fun (a,i) -> (f e a,i)) l)
- | CProdN (loc,bl,b) ->
- let (e,bl) = map_binders f g e bl in CProdN (loc,bl,f e b)
- | CLambdaN (loc,bl,b) ->
- let (e,bl) = map_binders f g e bl in CLambdaN (loc,bl,f e b)
- | CLetIn (loc,na,a,t,b) ->
- CLetIn (loc,na,f e a,Option.map (f e) t,f (name_fold g (snd na) e) b)
- | CCast (loc,a,c) -> CCast (loc,f e a, Miscops.map_cast_type (f e) c)
- | CNotation (loc,n,(l,ll,bll)) ->
+let map_constr_expr_with_binders g f e = Loc.map (function
+ | CAppExpl (r,l) -> CAppExpl (r,List.map (f e) l)
+ | CApp ((p,a),l) ->
+ CApp ((p,f e a),List.map (fun (a,i) -> (f e a,i)) l)
+ | CProdN (bl,b) ->
+ let (e,bl) = map_binders f g e bl in CProdN (bl,f e b)
+ | CLambdaN (bl,b) ->
+ let (e,bl) = map_binders f g e bl in CLambdaN (bl,f e b)
+ | CLetIn (na,a,t,b) ->
+ CLetIn (na,f e a,Option.map (f e) t,f (name_fold g (snd na) e) b)
+ | CCast (a,c) -> CCast (f e a, Miscops.map_cast_type (f e) c)
+ | CNotation (n,(l,ll,bll)) ->
(* This is an approximation because we don't know what binds what *)
- CNotation (loc,n,(List.map (f e) l,List.map (List.map (f e)) ll,
+ CNotation (n,(List.map (f e) l,List.map (List.map (f e)) ll,
List.map (fun bl -> snd (map_local_binders f g e bl)) bll))
- | CGeneralization (loc,b,a,c) -> CGeneralization (loc,b,a,f e c)
- | CDelimiters (loc,s,a) -> CDelimiters (loc,s,f e a)
+ | CGeneralization (b,a,c) -> CGeneralization (b,a,f e c)
+ | CDelimiters (s,a) -> CDelimiters (s,f e a)
| CHole _ | CEvar _ | CPatVar _ | CSort _
| CPrim _ | CRef _ as x -> x
- | CRecord (loc,l) -> CRecord (loc,List.map (fun (id, c) -> (id, f e c)) l)
- | CCases (loc,sty,rtnpo,a,bl) ->
+ | CRecord l -> CRecord (List.map (fun (id, c) -> (id, f e c)) l)
+ | CCases (sty,rtnpo,a,bl) ->
let bl = List.map (fun (loc,(patl,rhs)) ->
let ids = ids_of_pattern_list patl in
(loc,(patl,f (Id.Set.fold g ids e) rhs))) bl in
let ids = ids_of_cases_tomatch a in
let po = Option.map (f (Id.Set.fold g ids e)) rtnpo in
- CCases (loc, sty, po, List.map (fun (tm,x,y) -> f e tm,x,y) a,bl)
- | CLetTuple (loc,nal,(ona,po),b,c) ->
+ CCases (sty, po, List.map (fun (tm,x,y) -> f e tm,x,y) a,bl)
+ | CLetTuple (nal,(ona,po),b,c) ->
let e' = List.fold_right (Loc.down_located (name_fold g)) nal e in
let e'' = Option.fold_right (Loc.down_located (name_fold g)) ona e in
- CLetTuple (loc,nal,(ona,Option.map (f e'') po),f e b,f e' c)
- | CIf (loc,c,(ona,po),b1,b2) ->
+ CLetTuple (nal,(ona,Option.map (f e'') po),f e b,f e' c)
+ | CIf (c,(ona,po),b1,b2) ->
let e' = Option.fold_right (Loc.down_located (name_fold g)) ona e in
- CIf (loc,f e c,(ona,Option.map (f e') po),f e b1,f e b2)
- | CFix (loc,id,dl) ->
- CFix (loc,id,List.map (fun (id,n,bl,t,d) ->
+ CIf (f e c,(ona,Option.map (f e') po),f e b1,f e b2)
+ | CFix (id,dl) ->
+ CFix (id,List.map (fun (id,n,bl,t,d) ->
let (e',bl') = map_local_binders f g e bl in
let t' = f e' t in
(* Note: fix names should be inserted before the arguments... *)
let e'' = List.fold_left (fun e ((_,id),_,_,_,_) -> g id e) e' dl in
let d' = f e'' d in
(id,n,bl',t',d')) dl)
- | CCoFix (loc,id,dl) ->
- CCoFix (loc,id,List.map (fun (id,bl,t,d) ->
+ | CCoFix (id,dl) ->
+ CCoFix (id,List.map (fun (id,bl,t,d) ->
let (e',bl') = map_local_binders f g e bl in
let t' = f e' t in
let e'' = List.fold_left (fun e ((_,id),_,_,_) -> g id e) e' dl in
let d' = f e'' d in
(id,bl',t',d')) dl)
+ )
(* Used in constrintern *)
let rec replace_vars_constr_expr l = function
- | CRef (Ident (loc,id),us) as x ->
- (try CRef (Ident (loc,Id.Map.find id l),us) with Not_found -> x)
+ | loc, CRef (Ident (loc_id,id),us) as x ->
+ (try loc, CRef (Ident (loc_id,Id.Map.find id l),us) with Not_found -> x)
| c -> map_constr_expr_with_binders Id.Map.remove
replace_vars_constr_expr l c
diff --git a/intf/constrexpr.mli b/intf/constrexpr.mli
index c1de0ce24..c1ea71df4 100644
--- a/intf/constrexpr.mli
+++ b/intf/constrexpr.mli
@@ -68,38 +68,38 @@ and cases_pattern_notation_substitution =
cases_pattern_expr list * (** for constr subterms *)
cases_pattern_expr list list (** for recursive notations *)
-and constr_expr =
- | CRef of reference * instance_expr option
- | CFix of Loc.t * Id.t located * fix_expr list
- | CCoFix of Loc.t * Id.t located * cofix_expr list
- | CProdN of Loc.t * binder_expr list * constr_expr
- | CLambdaN of Loc.t * binder_expr list * constr_expr
- | CLetIn of Loc.t * Name.t located * constr_expr * constr_expr option * constr_expr
- | CAppExpl of Loc.t * (proj_flag * reference * instance_expr option) * constr_expr list
- | CApp of Loc.t * (proj_flag * constr_expr) *
- (constr_expr * explicitation located option) list
- | CRecord of Loc.t * (reference * constr_expr) list
+and constr_expr_r =
+ | CRef of reference * instance_expr option
+ | CFix of Id.t located * fix_expr list
+ | CCoFix of Id.t located * cofix_expr list
+ | CProdN of binder_expr list * constr_expr
+ | CLambdaN of binder_expr list * constr_expr
+ | CLetIn of Name.t located * constr_expr * constr_expr option * constr_expr
+ | CAppExpl of (proj_flag * reference * instance_expr option) * constr_expr list
+ | CApp of (proj_flag * constr_expr) *
+ (constr_expr * explicitation located option) list
+ | CRecord of (reference * constr_expr) list
(* representation of the "let" and "match" constructs *)
- | CCases of Loc.t (* position of the "match" keyword *)
- * case_style (* determines whether this value represents "let" or "match" construct *)
- * constr_expr option (* return-clause *)
- * case_expr list
- * branch_expr list (* branches *)
-
- | CLetTuple of Loc.t * Name.t located list * (Name.t located option * constr_expr option) *
- constr_expr * constr_expr
- | CIf of Loc.t * constr_expr * (Name.t located option * constr_expr option)
- * constr_expr * constr_expr
- | CHole of Loc.t * Evar_kinds.t option * intro_pattern_naming_expr * Genarg.raw_generic_argument option
- | CPatVar of Loc.t * patvar
- | CEvar of Loc.t * Glob_term.existential_name * (Id.t * constr_expr) list
- | CSort of Loc.t * glob_sort
- | CCast of Loc.t * constr_expr * constr_expr cast_type
- | CNotation of Loc.t * notation * constr_notation_substitution
- | CGeneralization of Loc.t * binding_kind * abstraction_kind option * constr_expr
- | CPrim of Loc.t * prim_token
- | CDelimiters of Loc.t * string * constr_expr
+ | CCases of case_style (* determines whether this value represents "let" or "match" construct *)
+ * constr_expr option (* return-clause *)
+ * case_expr list
+ * branch_expr list (* branches *)
+
+ | CLetTuple of Name.t located list * (Name.t located option * constr_expr option) *
+ constr_expr * constr_expr
+ | CIf of constr_expr * (Name.t located option * constr_expr option)
+ * constr_expr * constr_expr
+ | CHole of Evar_kinds.t option * intro_pattern_naming_expr * Genarg.raw_generic_argument option
+ | CPatVar of patvar
+ | CEvar of Glob_term.existential_name * (Id.t * constr_expr) list
+ | CSort of glob_sort
+ | CCast of constr_expr * constr_expr cast_type
+ | CNotation of notation * constr_notation_substitution
+ | CGeneralization of binding_kind * abstraction_kind option * constr_expr
+ | CPrim of prim_token
+ | CDelimiters of string * constr_expr
+and constr_expr = constr_expr_r located
and case_expr = constr_expr (* expression that is being matched *)
* Name.t located option (* as-clause *)
diff --git a/lib/loc.ml b/lib/loc.ml
index 39f2d7dfb..8ae8fe25d 100644
--- a/lib/loc.ml
+++ b/lib/loc.ml
@@ -63,6 +63,11 @@ type 'a located = t * 'a
let to_pair x = x
let tag ?loc x = Option.default ghost loc, x
+let with_loc f (loc, x) = f ~loc x
+
+let map f (l,x) = (l, f x)
+let map_with_loc f (loc, x) = (loc, f ~loc x)
+
let located_fold_left f x (_,a) = f x a
let located_iter2 f (_,a) (_,b) = f a b
let down_located f (_,a) = f a
diff --git a/lib/loc.mli b/lib/loc.mli
index fef1d8938..7fc8efaa8 100644
--- a/lib/loc.mli
+++ b/lib/loc.mli
@@ -61,6 +61,11 @@ type 'a located = t * 'a
val to_pair : 'a located -> t * 'a
val tag : ?loc:t -> 'a -> 'a located
+val with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b
+
+val map : ('a -> 'b) -> 'a located -> 'b located
+val map_with_loc : (loc:t -> 'a -> 'b) -> 'a located -> 'b located
+
val located_fold_left : ('a -> 'b -> 'a) -> 'a -> 'b located -> 'a
val down_located : ('a -> 'b) -> 'a located -> 'b
diff --git a/parsing/egramcoq.ml b/parsing/egramcoq.ml
index a973a539a..c8fe96f74 100644
--- a/parsing/egramcoq.ml
+++ b/parsing/egramcoq.ml
@@ -313,13 +313,13 @@ let interp_entry forpat e = match e with
| ETBinderList (true, _) -> assert false
| ETBinderList (false, tkl) -> TTAny (TTBinderListF tkl)
-let constr_expr_of_name (loc,na) = match na with
- | Anonymous -> CHole (loc,None,Misctypes.IntroAnonymous,None)
- | Name id -> CRef (Ident (loc,id), None)
+let constr_expr_of_name (loc,na) = Loc.tag ~loc @@ match na with
+ | Anonymous -> CHole (None,Misctypes.IntroAnonymous,None)
+ | Name id -> CRef (Ident (loc,id), None)
-let cases_pattern_expr_of_name (loc,na) = match na with
- | Anonymous -> Loc.tag ~loc @@ CPatAtom None
- | Name id -> Loc.tag ~loc @@ CPatAtom (Some (Ident (loc,id)))
+let cases_pattern_expr_of_name (loc,na) = Loc.tag ~loc @@ match na with
+ | Anonymous -> CPatAtom None
+ | Name id -> CPatAtom (Some (Ident (loc,id)))
type 'r env = {
constrs : 'r list;
@@ -342,12 +342,12 @@ match e with
| TTBinderListF _ -> { subst with binders = (List.flatten v, false) :: subst.binders }
| TTBigint ->
begin match forpat with
- | ForConstr -> push_constr subst (CPrim (Loc.ghost, Numeral v))
+ | ForConstr -> push_constr subst (Loc.tag @@ CPrim (Numeral v))
| ForPattern -> push_constr subst (Loc.tag @@ CPatPrim (Numeral v))
end
| TTReference ->
begin match forpat with
- | ForConstr -> push_constr subst (CRef (v, None))
+ | ForConstr -> push_constr subst (Loc.tag @@ CRef (v, None))
| ForPattern -> push_constr subst (Loc.tag @@ CPatAtom (Some v))
end
| TTConstrList _ -> { subst with constrlists = v :: subst.constrlists }
@@ -431,7 +431,7 @@ let rec pure_sublevels : type a b c. int option -> (a, b, c) rule -> int list =
let make_act : type r. r target -> _ -> r gen_eval = function
| ForConstr -> fun notation loc env ->
let env = (env.constrs, env.constrlists, List.map fst env.binders) in
- CNotation (loc, notation , env)
+ Loc.tag ~loc @@ CNotation (notation , env)
| ForPattern -> fun notation loc env ->
let invalid = List.exists (fun (_, b) -> not b) env.binders in
let () = if invalid then Topconstr.error_invalid_pattern_notation ~loc () in
diff --git a/parsing/g_constr.ml4 b/parsing/g_constr.ml4
index 6ca8134c0..9bf00b0b1 100644
--- a/parsing/g_constr.ml4
+++ b/parsing/g_constr.ml4
@@ -34,11 +34,11 @@ let mk_cast = function
(c,(_,None)) -> c
| (c,(_,Some ty)) ->
let loc = Loc.merge (constr_loc c) (constr_loc ty)
- in CCast(loc, c, CastConv ty)
+ in Loc.tag ~loc @@ CCast(c, CastConv ty)
let binder_of_name expl (loc,na) =
CLocalAssum ([loc, na], Default expl,
- CHole (loc, Some (Evar_kinds.BinderType na), IntroAnonymous, None))
+ Loc.tag ~loc @@ CHole (Some (Evar_kinds.BinderType na), IntroAnonymous, None))
let binders_of_names l =
List.map (binder_of_name Explicit) l
@@ -46,7 +46,7 @@ let binders_of_names l =
let mk_fixb (id,bl,ann,body,(loc,tyc)) =
let ty = match tyc with
Some ty -> ty
- | None -> CHole (loc, None, IntroAnonymous, None) in
+ | None -> Loc.tag @@ CHole (None, IntroAnonymous, None) in
(id,ann,bl,ty,body)
let mk_cofixb (id,bl,ann,body,(loc,tyc)) =
@@ -56,16 +56,16 @@ let mk_cofixb (id,bl,ann,body,(loc,tyc)) =
(Pp.str"Annotation forbidden in cofix expression.")) (fst ann) in
let ty = match tyc with
Some ty -> ty
- | None -> CHole (loc, None, IntroAnonymous, None) in
+ | None -> Loc.tag @@ CHole (None, IntroAnonymous, None) in
(id,bl,ty,body)
let mk_fix(loc,kw,id,dcls) =
if kw then
let fb = List.map mk_fixb dcls in
- CFix(loc,id,fb)
+ Loc.tag ~loc @@ CFix(id,fb)
else
let fb = List.map mk_cofixb dcls in
- CCoFix(loc,id,fb)
+ Loc.tag ~loc @@ CCoFix(id,fb)
let mk_single_fix (loc,kw,dcl) =
let (id,_,_,_,_) = dcl in mk_fix(loc,kw,id,[dcl])
@@ -159,62 +159,62 @@ GEXTEND Gram
;
constr:
[ [ c = operconstr LEVEL "8" -> c
- | "@"; f=global; i = instance -> CAppExpl(!@loc,(None,f,i),[]) ] ]
+ | "@"; f=global; i = instance -> Loc.tag ~loc:!@loc @@ CAppExpl((None,f,i),[]) ] ]
;
operconstr:
[ "200" RIGHTA
[ c = binder_constr -> c ]
| "100" RIGHTA
[ c1 = operconstr; "<:"; c2 = binder_constr ->
- CCast(!@loc,c1, CastVM c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastVM c2)
| c1 = operconstr; "<:"; c2 = SELF ->
- CCast(!@loc,c1, CastVM c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastVM c2)
| c1 = operconstr; "<<:"; c2 = binder_constr ->
- CCast(!@loc,c1, CastNative c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastNative c2)
| c1 = operconstr; "<<:"; c2 = SELF ->
- CCast(!@loc,c1, CastNative c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastNative c2)
| c1 = operconstr; ":";c2 = binder_constr ->
- CCast(!@loc,c1, CastConv c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastConv c2)
| c1 = operconstr; ":"; c2 = SELF ->
- CCast(!@loc,c1, CastConv c2)
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastConv c2)
| c1 = operconstr; ":>" ->
- CCast(!@loc,c1, CastCoerce) ]
+ Loc.tag ~loc:(!@loc) @@ CCast(c1, CastCoerce) ]
| "99" RIGHTA [ ]
| "90" RIGHTA [ ]
| "10" LEFTA
- [ f=operconstr; args=LIST1 appl_arg -> CApp(!@loc,(None,f),args)
- | "@"; f=global; i = instance; args=LIST0 NEXT -> CAppExpl(!@loc,(None,f,i),args)
+ [ f=operconstr; args=LIST1 appl_arg -> Loc.tag ~loc:(!@loc) @@ CApp((None,f),args)
+ | "@"; f=global; i = instance; args=LIST0 NEXT -> Loc.tag ~loc:(!@loc) @@ CAppExpl((None,f,i),args)
| "@"; (locid,id) = pattern_identref; args=LIST1 identref ->
- let args = List.map (fun x -> CRef (Ident x,None), None) args in
- CApp(!@loc,(None,CPatVar(locid,id)),args) ]
+ let args = List.map (fun x -> Loc.tag @@ CRef (Ident x,None), None) args in
+ Loc.tag ~loc:(!@loc) @@ CApp((None, Loc.tag ~loc:locid @@ CPatVar id),args) ]
| "9"
[ ".."; c = operconstr LEVEL "0"; ".." ->
- CAppExpl (!@loc,(None,Ident (!@loc,ldots_var),None),[c]) ]
+ Loc.tag ~loc:(!@loc) @@ CAppExpl ((None, Ident (!@loc,ldots_var),None),[c]) ]
| "8" [ ]
| "1" LEFTA
[ c=operconstr; ".("; f=global; args=LIST0 appl_arg; ")" ->
- CApp(!@loc,(Some (List.length args+1),CRef (f,None)),args@[c,None])
+ Loc.tag ~loc:(!@loc) @@ CApp((Some (List.length args+1), Loc.tag @@ CRef (f,None)),args@[c,None])
| c=operconstr; ".("; "@"; f=global;
args=LIST0 (operconstr LEVEL "9"); ")" ->
- CAppExpl(!@loc,(Some (List.length args+1),f,None),args@[c])
- | c=operconstr; "%"; key=IDENT -> CDelimiters (!@loc,key,c) ]
+ Loc.tag ~loc:(!@loc) @@ CAppExpl((Some (List.length args+1),f,None),args@[c])
+ | c=operconstr; "%"; key=IDENT -> Loc.tag ~loc:(!@loc) @@ CDelimiters (key,c) ]
| "0"
[ c=atomic_constr -> c
| c=match_constr -> c
| "("; c = operconstr LEVEL "200"; ")" ->
- (match c with
- CPrim (_,Numeral z) when Bigint.is_pos_or_zero z ->
- CNotation(!@loc,"( _ )",([c],[],[]))
+ (match snd c with
+ CPrim (Numeral z) when Bigint.is_pos_or_zero z ->
+ Loc.tag ~loc:(!@loc) @@ CNotation("( _ )",([c],[],[]))
| _ -> c)
| "{|"; c = record_declaration; "|}" -> c
| "`{"; c = operconstr LEVEL "200"; "}" ->
- CGeneralization (!@loc, Implicit, None, c)
+ Loc.tag ~loc:(!@loc) @@ CGeneralization (Implicit, None, c)
| "`("; c = operconstr LEVEL "200"; ")" ->
- CGeneralization (!@loc, Explicit, None, c)
+ Loc.tag ~loc:(!@loc) @@ CGeneralization (Explicit, None, c)
] ]
;
record_declaration:
- [ [ fs = record_fields -> CRecord (!@loc, fs) ] ]
+ [ [ fs = record_fields -> Loc.tag ~loc:(!@loc) @@ CRecord fs ] ]
;
record_fields:
@@ -236,36 +236,36 @@ GEXTEND Gram
| "let"; id=name; bl = binders; ty = type_cstr; ":=";
c1 = operconstr LEVEL "200"; "in"; c2 = operconstr LEVEL "200" ->
let ty,c1 = match ty, c1 with
- | (_,None), CCast(loc,c, CastConv t) -> (constr_loc t,Some t), c (* Tolerance, see G_vernac.def_body *)
+ | (_,None), (_, CCast(c, CastConv t)) -> (Loc.tag ~loc:(constr_loc t) @@ Some t), c (* Tolerance, see G_vernac.def_body *)
| _, _ -> ty, c1 in
- CLetIn(!@loc,id,mkCLambdaN (constr_loc c1) bl c1,
- Option.map (mkCProdN (fst ty) bl) (snd ty), c2)
+ Loc.tag ~loc:!@loc @@ CLetIn(id,mkCLambdaN (constr_loc c1) bl c1,
+ Option.map (mkCProdN ~loc:(fst ty) bl) (snd ty), c2)
| "let"; fx = single_fix; "in"; c = operconstr LEVEL "200" ->
let fixp = mk_single_fix fx in
- let (li,id) = match fixp with
- CFix(_,id,_) -> id
- | CCoFix(_,id,_) -> id
+ let (li,id) = match snd fixp with
+ CFix(id,_) -> id
+ | CCoFix(id,_) -> id
| _ -> assert false in
- CLetIn(!@loc,(li,Name id),fixp,None,c)
+ Loc.tag ~loc:!@loc @@ CLetIn((li,Name id),fixp,None,c)
| "let"; lb = ["("; l=LIST0 name SEP ","; ")" -> l | "()" -> []];
po = return_type;
":="; c1 = operconstr LEVEL "200"; "in";
c2 = operconstr LEVEL "200" ->
- CLetTuple (!@loc,lb,po,c1,c2)
+ Loc.tag ~loc:!@loc @@ CLetTuple (lb,po,c1,c2)
| "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200";
"in"; c2 = operconstr LEVEL "200" ->
- CCases (!@loc, LetPatternStyle, None, [c1, None, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
+ Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, None, [c1, None, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
| "let"; "'"; p=pattern; ":="; c1 = operconstr LEVEL "200";
rt = case_type; "in"; c2 = operconstr LEVEL "200" ->
- CCases (!@loc, LetPatternStyle, Some rt, [c1, aliasvar p, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
+ Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, None], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
| "let"; "'"; p=pattern; "in"; t = pattern LEVEL "200";
":="; c1 = operconstr LEVEL "200"; rt = case_type;
"in"; c2 = operconstr LEVEL "200" ->
- CCases (!@loc, LetPatternStyle, Some rt, [c1, aliasvar p, Some t], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
+ Loc.tag ~loc:!@loc @@ CCases (LetPatternStyle, Some rt, [c1, aliasvar p, Some t], [Loc.tag ~loc:!@loc ([(!@loc, [p])], c2)])
| "if"; c=operconstr LEVEL "200"; po = return_type;
"then"; b1=operconstr LEVEL "200";
"else"; b2=operconstr LEVEL "200" ->
- CIf (!@loc, c, po, b1, b2)
+ Loc.tag ~loc:(!@loc) @@ CIf (c, po, b1, b2)
| c=fix_constr -> c ] ]
;
appl_arg:
@@ -274,14 +274,14 @@ GEXTEND Gram
| c=operconstr LEVEL "9" -> (c,None) ] ]
;
atomic_constr:
- [ [ g=global; i=instance -> CRef (g,i)
- | s=sort -> CSort (!@loc,s)
- | n=INT -> CPrim (!@loc, Numeral (Bigint.of_string n))
- | s=string -> CPrim (!@loc, String s)
- | "_" -> CHole (!@loc, None, IntroAnonymous, None)
- | "?"; "["; id=ident; "]" -> CHole (!@loc, None, IntroIdentifier id, None)
- | "?"; "["; id=pattern_ident; "]" -> CHole (!@loc, None, IntroFresh id, None)
- | id=pattern_ident; inst = evar_instance -> CEvar(!@loc,id,inst) ] ]
+ [ [ g=global; i=instance -> Loc.tag ~loc:!@loc @@ CRef (g,i)
+ | s=sort -> Loc.tag ~loc:!@loc @@ CSort s
+ | n=INT -> Loc.tag ~loc:!@loc @@ CPrim (Numeral (Bigint.of_string n))
+ | s=string -> Loc.tag ~loc:!@loc @@ CPrim (String s)
+ | "_" -> Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None)
+ | "?"; "["; id=ident; "]" -> Loc.tag ~loc:!@loc @@ CHole (None, IntroIdentifier id, None)
+ | "?"; "["; id=pattern_ident; "]" -> Loc.tag ~loc:!@loc @@ CHole (None, IntroFresh id, None)
+ | id=pattern_ident; inst = evar_instance -> Loc.tag ~loc:!@loc @@ CEvar(id,inst) ] ]
;
inst:
[ [ id = ident; ":="; c = lconstr -> (id,c) ] ]
@@ -322,7 +322,7 @@ GEXTEND Gram
;
match_constr:
[ [ "match"; ci=LIST1 case_item SEP ","; ty=OPT case_type; "with";
- br=branches; "end" -> CCases(!@loc,RegularStyle,ty,ci,br) ] ]
+ br=branches; "end" -> Loc.tag ~loc:!@loc @@ CCases(RegularStyle,ty,ci,br) ] ]
;
case_item:
[ [ c=operconstr LEVEL "100";
@@ -410,7 +410,8 @@ GEXTEND Gram
| nal=LIST1 name; ":"; c=lconstr; "}" ->
(fun na -> CLocalAssum (na::nal,Default Implicit,c))
| nal=LIST1 name; "}" ->
- (fun na -> CLocalAssum (na::nal,Default Implicit,CHole (Loc.join_loc (fst na) !@loc, Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None)))
+ (fun na -> CLocalAssum (na::nal,Default Implicit,
+ Loc.tag ~loc:(Loc.join_loc (fst na) !@loc) @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None)))
| ":"; c=lconstr; "}" ->
(fun na -> CLocalAssum ([na],Default Implicit,c))
] ]
@@ -444,7 +445,7 @@ GEXTEND Gram
binders_of_names (id::idl) @ bl
| id1 = name; ".."; id2 = name ->
[CLocalAssum ([id1;(!@loc,Name ldots_var);id2],
- Default Explicit,CHole (!@loc, None, IntroAnonymous, None))]
+ Default Explicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))]
| bl = closed_binder; bl' = binders ->
bl@bl'
] ]
@@ -453,7 +454,7 @@ GEXTEND Gram
[ [ l = LIST0 binder -> List.flatten l ] ]
;
binder:
- [ [ id = name -> [CLocalAssum ([id],Default Explicit,CHole (!@loc, None, IntroAnonymous, None))]
+ [ [ id = name -> [CLocalAssum ([id],Default Explicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))]
| bl = closed_binder -> bl ] ]
;
closed_binder:
@@ -463,18 +464,18 @@ GEXTEND Gram
[CLocalAssum ([id],Default Explicit,c)]
| "("; id=name; ":="; c=lconstr; ")" ->
(match c with
- | CCast(_,c, CastConv t) -> [CLocalDef (id,c,Some t)]
+ | (_, CCast(c, CastConv t)) -> [CLocalDef (id,c,Some t)]
| _ -> [CLocalDef (id,c,None)])
| "("; id=name; ":"; t=lconstr; ":="; c=lconstr; ")" ->
[CLocalDef (id,c,Some t)]
| "{"; id=name; "}" ->
- [CLocalAssum ([id],Default Implicit,CHole (!@loc, None, IntroAnonymous, None))]
+ [CLocalAssum ([id],Default Implicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))]
| "{"; id=name; idl=LIST1 name; ":"; c=lconstr; "}" ->
[CLocalAssum (id::idl,Default Implicit,c)]
| "{"; id=name; ":"; c=lconstr; "}" ->
[CLocalAssum ([id],Default Implicit,c)]
| "{"; id=name; idl=LIST1 name; "}" ->
- List.map (fun id -> CLocalAssum ([id],Default Implicit,CHole (!@loc, None, IntroAnonymous, None))) (id::idl)
+ List.map (fun id -> CLocalAssum ([id],Default Implicit, Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, None))) (id::idl)
| "`("; tc = LIST1 typeclass_constraint SEP "," ; ")" ->
List.map (fun (n, b, t) -> CLocalAssum ([n], Generalized (Implicit, Explicit, b), t)) tc
| "`{"; tc = LIST1 typeclass_constraint SEP "," ; "}" ->
diff --git a/parsing/g_proofs.ml4 b/parsing/g_proofs.ml4
index 7ca2e4a4f..488995201 100644
--- a/parsing/g_proofs.ml4
+++ b/parsing/g_proofs.ml4
@@ -120,7 +120,7 @@ GEXTEND Gram
;
constr_body:
[ [ ":="; c = lconstr -> c
- | ":"; t = lconstr; ":="; c = lconstr -> CCast(!@loc,c,CastConv t) ] ]
+ | ":"; t = lconstr; ":="; c = lconstr -> Loc.tag ~loc:!@loc @@ CCast(c,CastConv t) ] ]
;
mode:
[ [ l = LIST1 [ "+" -> ModeInput
diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4
index a43f1127a..3d0322b10 100644
--- a/parsing/g_vernac.ml4
+++ b/parsing/g_vernac.ml4
@@ -229,19 +229,19 @@ GEXTEND Gram
if List.exists (function CLocalPattern _ -> true | _ -> false) bl
then
(* FIXME: "red" will be applied to types in bl and Cast with remain *)
- let c = mkCLambdaN (!@loc) bl c in
+ let c = mkCLambdaN ~loc:!@loc bl c in
DefineBody ([], red, c, None)
else
(match c with
- | CCast(_,c, CastConv t) -> DefineBody (bl, red, c, Some t)
+ | _, CCast(c, CastConv t) -> DefineBody (bl, red, c, Some t)
| _ -> DefineBody (bl, red, c, None))
| bl = binders; ":"; t = lconstr; ":="; red = reduce; c = lconstr ->
let ((bl, c), tyo) =
if List.exists (function CLocalPattern _ -> true | _ -> false) bl
then
(* FIXME: "red" will be applied to types in bl and Cast with remain *)
- let c = CCast (!@loc, c, CastConv t) in
- (([],mkCLambdaN (!@loc) bl c), None)
+ let c = Loc.tag ~loc:!@loc @@ CCast (c, CastConv t) in
+ (([],mkCLambdaN ~loc:!@loc bl c), None)
else ((bl, c), Some t)
in
DefineBody (bl, red, c, tyo)
@@ -305,7 +305,7 @@ GEXTEND Gram
;
type_cstr:
[ [ ":"; c=lconstr -> c
- | -> CHole (!@loc, None, Misctypes.IntroAnonymous, None) ] ]
+ | -> Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None) ] ]
;
(* Inductive schemes *)
scheme:
@@ -352,16 +352,16 @@ GEXTEND Gram
t = lconstr -> fun id -> (oc,AssumExpr (id,mkCProdN (!@loc) l t))
| l = binders; oc = of_type_with_opt_coercion;
t = lconstr; ":="; b = lconstr -> fun id ->
- (oc,DefExpr (id,mkCLambdaN (!@loc) l b,Some (mkCProdN (!@loc) l t)))
+ (oc,DefExpr (id,mkCLambdaN ~loc:!@loc l b,Some (mkCProdN (!@loc) l t)))
| l = binders; ":="; b = lconstr -> fun id ->
- match b with
- | CCast(_,b, (CastConv t|CastVM t|CastNative t)) ->
- (None,DefExpr(id,mkCLambdaN (!@loc) l b,Some (mkCProdN (!@loc) l t)))
+ match snd b with
+ | CCast(b', (CastConv t|CastVM t|CastNative t)) ->
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b',Some (mkCProdN (!@loc) l t)))
| _ ->
- (None,DefExpr(id,mkCLambdaN (!@loc) l b,None)) ] ]
+ (None,DefExpr(id,mkCLambdaN ~loc:!@loc l b,None)) ] ]
;
record_binder:
- [ [ id = name -> (None,AssumExpr(id,CHole (!@loc, None, Misctypes.IntroAnonymous, None)))
+ [ [ id = name -> (None,AssumExpr(id, Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))
| id = name; f = record_binder_body -> f id ] ]
;
assum_list:
@@ -380,7 +380,7 @@ GEXTEND Gram
t= [ coe = of_type_with_opt_coercion; c = lconstr ->
fun l id -> (not (Option.is_empty coe),(id,mkCProdN (!@loc) l c))
| ->
- fun l id -> (false,(id,mkCProdN (!@loc) l (CHole (!@loc, None, Misctypes.IntroAnonymous, None)))) ]
+ fun l id -> (false,(id,mkCProdN (!@loc) l (Loc.tag ~loc:!@loc @@ CHole (None, Misctypes.IntroAnonymous, None)))) ]
-> t l
]]
;
diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4
new file mode 100644
index 000000000..2415080f3
--- /dev/null
+++ b/plugins/decl_mode/g_decl_mode.ml4
@@ -0,0 +1,387 @@
+(************************************************************************)
+(* 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 7dc869131..4b942c989 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -1247,16 +1247,15 @@ let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_
in
List.rev !l
-let rec rebuild_return_type rt =
+let rec rebuild_return_type (loc, rt) =
match rt with
- | Constrexpr.CProdN(loc,n,t') ->
- Constrexpr.CProdN(loc,n,rebuild_return_type t')
- | Constrexpr.CLetIn(loc,na,v,t,t') ->
- Constrexpr.CLetIn(loc,na,v,t,rebuild_return_type t')
- | _ -> Constrexpr.CProdN(Loc.ghost,[[Loc.ghost,Anonymous],
- Constrexpr.Default Decl_kinds.Explicit,rt],
- Constrexpr.CSort(Loc.ghost,GType []))
-
+ | Constrexpr.CProdN(n,t') ->
+ Loc.tag ~loc @@ Constrexpr.CProdN(n,rebuild_return_type t')
+ | Constrexpr.CLetIn(na,v,t,t') ->
+ Loc.tag ~loc @@ Constrexpr.CLetIn(na,v,t,rebuild_return_type t')
+ | _ -> Loc.tag ~loc @@ Constrexpr.CProdN([[Loc.ghost,Anonymous],
+ Constrexpr.Default Decl_kinds.Explicit,Loc.tag ~loc rt],
+ Loc.tag @@ Constrexpr.CSort(GType []))
let do_build_inductive
evd (funconstants: Term.pconstant list) (funsargs: (Name.t * glob_constr * glob_constr option) list list)
@@ -1307,13 +1306,12 @@ let do_build_inductive
(fun (n,t,typ) acc ->
match typ with
| Some typ ->
- Constrexpr.CLetIn(Loc.ghost,(Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
+ Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ),
acc)
| None ->
- Constrexpr.CProdN
- (Loc.ghost,
- [[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
+ Loc.tag @@ Constrexpr.CProdN
+ ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
acc
)
)
@@ -1375,13 +1373,12 @@ let do_build_inductive
(fun (n,t,typ) acc ->
match typ with
| Some typ ->
- Constrexpr.CLetIn(Loc.ghost,(Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
+ Loc.tag @@ Constrexpr.CLetIn((Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ),
acc)
| None ->
- Constrexpr.CProdN
- (Loc.ghost,
- [[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
+ Loc.tag @@ Constrexpr.CProdN
+ ([[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
acc
)
)
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index daa5cbb3f..6ee755323 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -469,8 +469,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas
in
let unbounded_eq =
let f_app_args =
- Constrexpr.CAppExpl
- (Loc.ghost,
+ Loc.tag @@ Constrexpr.CAppExpl(
(None,(Ident (Loc.ghost,fname)),None) ,
(List.map
(function
@@ -481,7 +480,7 @@ let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas
)
)
in
- Constrexpr.CApp (Loc.ghost,(None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(qualid_of_string "Logic.eq")))),
+ Loc.tag @@ Constrexpr.CApp ((None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(qualid_of_string "Logic.eq")))),
[(f_app_args,None);(body,None)])
in
let eq = Constrexpr_ops.mkCProdN Loc.ghost args unbounded_eq in
@@ -589,15 +588,15 @@ let rec rebuild_bl (aux,assoc) bl typ =
| [], _ -> (List.rev aux,replace_vars_constr_expr assoc typ,assoc)
| (Constrexpr.CLocalAssum(nal,bk,_))::bl',typ ->
rebuild_nal (aux,assoc) bk bl' nal (List.length nal) typ
- | (Constrexpr.CLocalDef(na,_,_))::bl',Constrexpr.CLetIn(_,_,nat,ty,typ') ->
+ | (Constrexpr.CLocalDef(na,_,_))::bl',(_loc, Constrexpr.CLetIn(_,nat,ty,typ')) ->
rebuild_bl ((Constrexpr.CLocalDef(na,replace_vars_constr_expr assoc nat,Option.map (replace_vars_constr_expr assoc) ty (* ??? *))::aux),assoc)
bl' typ'
| _ -> assert false
and rebuild_nal (aux,assoc) bk bl' nal lnal typ =
- match nal,typ with
+ match nal, snd typ with
| [], _ -> rebuild_bl (aux,assoc) bl' typ
- | _,CProdN(_,[],typ) -> rebuild_nal (aux,assoc) bk bl' nal lnal typ
- | _,CProdN(_,(nal',bk',nal't)::rest,typ') ->
+ | _,CProdN([],typ) -> rebuild_nal (aux,assoc) bk bl' nal lnal typ
+ | _,CProdN((nal',bk',nal't)::rest,typ') ->
let lnal' = List.length nal' in
if lnal' >= lnal
then
@@ -607,15 +606,15 @@ let rec rebuild_bl (aux,assoc) bl typ =
rebuild_bl ((assum :: aux), nassoc) bl'
(if List.is_empty new_nal' && List.is_empty rest
then typ'
- else if List.is_empty new_nal'
- then CProdN(Loc.ghost,rest,typ')
- else CProdN(Loc.ghost,((new_nal',bk',nal't)::rest),typ'))
+ else Loc.tag @@ if List.is_empty new_nal'
+ then CProdN(rest,typ')
+ else CProdN(((new_nal',bk',nal't)::rest),typ'))
else
let captured_nal,non_captured_nal = List.chop lnal' nal in
let nassoc = make_assoc assoc nal' captured_nal in
let assum = CLocalAssum(captured_nal,bk,replace_vars_constr_expr assoc nal't) in
rebuild_nal ((assum :: aux), nassoc)
- bk bl' non_captured_nal (lnal - lnal') (CProdN(Loc.ghost,rest,typ'))
+ bk bl' non_captured_nal (lnal - lnal') (Loc.tag @@ CProdN(rest,typ'))
| _ -> assert false
let rebuild_bl (aux,assoc) bl typ = rebuild_bl (aux,assoc) bl typ
@@ -726,67 +725,65 @@ let do_generate_principle pconstants on_error register_built interactive_proof
in
()
-let rec add_args id new_args b =
- match b with
- | CRef (r,_) ->
- begin match r with
+let rec add_args id new_args = Loc.map (function
+ | CRef (r,_) as b ->
+ begin match r with
| Libnames.Ident(loc,fname) when Id.equal fname id ->
- CAppExpl(Loc.ghost,(None,r,None),new_args)
+ CAppExpl((None,r,None),new_args)
| _ -> b
end
| CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo")
- | CProdN(loc,nal,b1) ->
- CProdN(loc,
- List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ | CProdN(nal,b1) ->
+ CProdN(List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
add_args id new_args b1)
- | CLambdaN(loc,nal,b1) ->
- CLambdaN(loc,
- List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
+ | CLambdaN(nal,b1) ->
+ CLambdaN(List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
add_args id new_args b1)
- | CLetIn(loc,na,b1,t,b2) ->
- CLetIn(loc,na,add_args id new_args b1,Option.map (add_args id new_args) t,add_args id new_args b2)
- | CAppExpl(loc,(pf,r,us),exprl) ->
+ | CLetIn(na,b1,t,b2) ->
+ CLetIn(na,add_args id new_args b1,Option.map (add_args id new_args) t,add_args id new_args b2)
+ | CAppExpl((pf,r,us),exprl) ->
begin
match r with
| Libnames.Ident(loc,fname) when Id.equal fname id ->
- CAppExpl(loc,(pf,r,us),new_args@(List.map (add_args id new_args) exprl))
- | _ -> CAppExpl(loc,(pf,r,us),List.map (add_args id new_args) exprl)
+ CAppExpl((pf,r,us),new_args@(List.map (add_args id new_args) exprl))
+ | _ -> CAppExpl((pf,r,us),List.map (add_args id new_args) exprl)
end
- | CApp(loc,(pf,b),bl) ->
- CApp(loc,(pf,add_args id new_args b),
+ | CApp((pf,b),bl) ->
+ CApp((pf,add_args id new_args b),
List.map (fun (e,o) -> add_args id new_args e,o) bl)
- | CCases(loc,sty,b_option,cel,cal) ->
- CCases(loc,sty,Option.map (add_args id new_args) b_option,
+ | CCases(sty,b_option,cel,cal) ->
+ CCases(sty,Option.map (add_args id new_args) b_option,
List.map (fun (b,na,b_option) ->
add_args id new_args b,
na, b_option) cel,
- List.map (fun (loc,(cpl,e)) -> Loc.tag ~loc:loc @@ (cpl,add_args id new_args e)) cal
+ List.map (fun (loc,(cpl,e)) -> Loc.tag ~loc @@ (cpl,add_args id new_args e)) cal
)
- | CLetTuple(loc,nal,(na,b_option),b1,b2) ->
- CLetTuple(loc,nal,(na,Option.map (add_args id new_args) b_option),
+ | CLetTuple(nal,(na,b_option),b1,b2) ->
+ CLetTuple(nal,(na,Option.map (add_args id new_args) b_option),
add_args id new_args b1,
add_args id new_args b2
)
- | CIf(loc,b1,(na,b_option),b2,b3) ->
- CIf(loc,add_args id new_args b1,
+ | CIf(b1,(na,b_option),b2,b3) ->
+ CIf(add_args id new_args b1,
(na,Option.map (add_args id new_args) b_option),
add_args id new_args b2,
add_args id new_args b3
)
- | CHole _ -> b
- | CPatVar _ -> b
- | CEvar _ -> b
- | CSort _ -> b
- | CCast(loc,b1,b2) ->
- CCast(loc,add_args id new_args b1,
+ | CHole _
+ | CPatVar _
+ | CEvar _
+ | CPrim _
+ | CSort _ as b -> b
+ | CCast(b1,b2) ->
+ CCast(add_args id new_args b1,
Miscops.map_cast_type (add_args id new_args) b2)
- | CRecord (loc, pars) ->
- CRecord (loc, List.map (fun (e,o) -> e, add_args id new_args o) pars)
+ | CRecord pars ->
+ CRecord (List.map (fun (e,o) -> e, add_args id new_args o) pars)
| CNotation _ -> anomaly ~label:"add_args " (Pp.str "CNotation")
| CGeneralization _ -> anomaly ~label:"add_args " (Pp.str "CGeneralization")
- | CPrim _ -> b
| CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters")
+ )
exception Stop of Constrexpr.constr_expr
@@ -797,8 +794,8 @@ let rec chop_n_arrow n t =
if n <= 0
then t (* If we have already removed all the arrows then return the type *)
else (* If not we check the form of [t] *)
- match t with
- | Constrexpr.CProdN(_,nal_ta',t') -> (* If we have a forall, to result are possible :
+ match snd t with
+ | Constrexpr.CProdN(nal_ta',t') -> (* If we have a forall, to result are possible :
either we need to discard more than the number of arrows contained
in this product declaration then we just recall [chop_n_arrow] on
the remaining number of arrow to chop and [t'] we discard it and
@@ -816,8 +813,8 @@ let rec chop_n_arrow n t =
then
aux (n - nal_l) nal_ta'
else
- let new_t' =
- Constrexpr.CProdN(Loc.ghost,
+ let new_t' = Loc.tag @@
+ Constrexpr.CProdN(
((snd (List.chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
@@ -832,8 +829,8 @@ let rec chop_n_arrow n t =
let rec get_args b t : Constrexpr.local_binder_expr list *
Constrexpr.constr_expr * Constrexpr.constr_expr =
- match b with
- | Constrexpr.CLambdaN (loc, (nal_ta), b') ->
+ match snd b with
+ | Constrexpr.CLambdaN ((nal_ta), b') ->
begin
let n =
(List.fold_left (fun n (nal,_,_) ->
@@ -872,8 +869,8 @@ let make_graph (f_ref:global_reference) =
in
let (nal_tas,b,t) = get_args extern_body extern_type in
let expr_list =
- match b with
- | Constrexpr.CFix(loc,l_id,fixexprl) ->
+ match snd b with
+ | Constrexpr.CFix(l_id,fixexprl) ->
let l =
List.map
(fun (id,(n,recexp),bl,t,b) ->
@@ -885,7 +882,7 @@ let make_graph (f_ref:global_reference) =
| Constrexpr.CLocalDef (na,_,_)-> []
| Constrexpr.CLocalAssum (nal,_,_) ->
List.map
- (fun (loc,n) ->
+ (fun (loc,n) -> Loc.tag ~loc @@
CRef(Libnames.Ident(loc, Nameops.out_name n),None))
nal
| Constrexpr.CLocalPattern _ -> assert false
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index f1ca57585..29de56478 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -73,7 +73,7 @@ let isVarf f x =
in global environment. *)
let ident_global_exist id =
try
- let ans = CRef (Libnames.Ident (Loc.ghost,id), None) in
+ let ans = Loc.tag @@ CRef (Libnames.Ident (Loc.ghost,id), None) in
let _ = ignore (Constrintern.intern_constr (Global.env()) ans) in
true
with e when CErrors.noncritical e -> false
@@ -835,7 +835,7 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let c = RelDecl.get_type decl in
let typ = Constrextern.extern_constr false env Evd.empty c in
let newenv = Environ.push_rel (LocalAssum (nm,c)) env in
- CProdN (Loc.ghost, [[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv)
+ Loc.tag @@ CProdN ([[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv)
(concl,Global.env())
(shift.funresprms2 @ shift.funresprms1
@ shift.args2 @ shift.args1 @ shift.otherprms2 @ shift.otherprms1) in
diff --git a/plugins/ltac/g_ltac.ml4 b/plugins/ltac/g_ltac.ml4
index ca5d198c2..dffd90be3 100644
--- a/plugins/ltac/g_ltac.ml4
+++ b/plugins/ltac/g_ltac.ml4
@@ -187,7 +187,7 @@ GEXTEND Gram
(* Tactic arguments to the right of an application *)
tactic_arg_compat:
[ [ a = tactic_arg -> a
- | c = Constr.constr -> (match c with CRef (r,None) -> Reference r | c -> ConstrMayEval (ConstrTerm c))
+ | c = Constr.constr -> (match c with _loc, CRef (r,None) -> Reference r | c -> ConstrMayEval (ConstrTerm c))
(* Unambiguous entries: tolerated w/o "ltac:" modifier *)
| "()" -> TacGeneric (genarg_of_unit ()) ] ]
;
@@ -255,10 +255,10 @@ GEXTEND Gram
let t, ty =
match mpv with
| Term t -> (match t with
- | CCast (loc, t, (CastConv ty | CastVM ty | CastNative ty)) -> Term t, Some (Term ty)
+ | _loc, CCast (t, (CastConv ty | CastVM ty | CastNative ty)) -> Term t, Some (Term ty)
| _ -> mpv, None)
| _ -> mpv, None
- in Def (na, t, Option.default (Term (CHole (Loc.ghost, None, IntroAnonymous, None))) ty)
+ in Def (na, t, Option.default (Term (Loc.tag @@ CHole (None, IntroAnonymous, None))) ty)
] ]
;
match_context_rule:
@@ -353,7 +353,7 @@ GEXTEND Gram
operconstr: LEVEL "0"
[ [ IDENT "ltac"; ":"; "("; tac = Pltac.tactic_expr; ")" ->
let arg = Genarg.in_gen (Genarg.rawwit Tacarg.wit_tactic) tac in
- CHole (!@loc, None, IntroAnonymous, Some arg) ] ]
+ Loc.tag ~loc:!@loc @@ CHole (None, IntroAnonymous, Some arg) ] ]
;
END
diff --git a/plugins/ltac/g_tactic.ml4 b/plugins/ltac/g_tactic.ml4
index 4b3ca80af..1674bb4ca 100644
--- a/plugins/ltac/g_tactic.ml4
+++ b/plugins/ltac/g_tactic.ml4
@@ -130,14 +130,14 @@ let mk_fix_tac (loc,id,bl,ann,ty) =
(try List.index Names.Name.equal (snd x) ids
with Not_found -> error "No such fix variable.")
| _ -> error "Cannot guess decreasing argument of fix." in
- (id,n,CProdN(loc,bl,ty))
+ (id,n, Loc.tag ~loc @@ CProdN(bl,ty))
let mk_cofix_tac (loc,id,bl,ann,ty) =
let _ = Option.map (fun (aloc,_) ->
user_err ~loc:aloc
~hdr:"Constr:mk_cofix_tac"
(Pp.str"Annotation forbidden in cofix expression.")) ann in
- (id,CProdN(loc,bl,ty))
+ (id,Loc.tag ~loc @@ CProdN(bl,ty))
(* Functions overloaded by quotifier *)
let destruction_arg_of_constr (c,lbind as clbind) = match lbind with
@@ -154,12 +154,12 @@ let mkTacCase with_evar = function
(* Reinterpret numbers as a notation for terms *)
| [(clear,ElimOnAnonHyp n),(None,None),None],None ->
TacCase (with_evar,
- (clear,(CPrim (Loc.ghost, Numeral (Bigint.of_int n)),
+ (clear,(Loc.tag @@ CPrim (Numeral (Bigint.of_int n)),
NoBindings)))
(* Reinterpret ident as notations for variables in the context *)
(* because we don't know if they are quantified or not *)
| [(clear,ElimOnIdent id),(None,None),None],None ->
- TacCase (with_evar,(clear,(CRef (Ident id,None),NoBindings)))
+ TacCase (with_evar,(clear,(Loc.tag @@ CRef (Ident id,None),NoBindings)))
| ic ->
if List.exists (function ((_, ElimOnAnonHyp _),_,_) -> true | _ -> false) (fst ic)
then
@@ -169,7 +169,7 @@ let mkTacCase with_evar = function
let rec mkCLambdaN_simple_loc loc bll c =
match bll with
| ((loc1,_)::_ as idl,bk,t) :: bll ->
- CLambdaN (loc,[idl,bk,t],mkCLambdaN_simple_loc (Loc.merge loc1 loc) bll c)
+ Loc.tag ~loc @@ CLambdaN ([idl,bk,t],mkCLambdaN_simple_loc (Loc.merge loc1 loc) bll c)
| ([],_,_) :: bll -> mkCLambdaN_simple_loc loc bll c
| [] -> c
@@ -440,7 +440,7 @@ GEXTEND Gram
| -> true ]]
;
simple_binder:
- [ [ na=name -> ([na],Default Explicit,CHole (!@loc, Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))
+ [ [ na=name -> ([na],Default Explicit, Loc.tag ~loc:!@loc @@ CHole (Some (Evar_kinds.BinderType (snd na)), IntroAnonymous, None))
| "("; nal=LIST1 name; ":"; c=lconstr; ")" -> (nal,Default Explicit,c)
] ]
;
diff --git a/plugins/ltac/pptactic.ml b/plugins/ltac/pptactic.ml
index 39ae1f41d..ad76ef9c6 100644
--- a/plugins/ltac/pptactic.ml
+++ b/plugins/ltac/pptactic.ml
@@ -340,8 +340,8 @@ type 'a extra_genarg_printer =
let strip_prod_binders_expr n ty =
let rec strip_ty acc n ty =
- match ty with
- Constrexpr.CProdN(_,bll,a) ->
+ match snd ty with
+ Constrexpr.CProdN(bll,a) ->
let nb =
List.fold_left (fun i (nal,_,_) -> i + List.length nal) 0 bll in
let bll = List.map (fun (x, _, y) -> x, y) bll in
diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml
index b84be4600..1f75f88d4 100644
--- a/plugins/ltac/rewrite.ml
+++ b/plugins/ltac/rewrite.ml
@@ -1787,18 +1787,18 @@ let rec strategy_of_ast = function
(* By default the strategy for "rewrite_db" is top-down *)
-let mkappc s l = CAppExpl (Loc.ghost,(None,(Libnames.Ident (Loc.ghost,Id.of_string s)),None),l)
+let mkappc s l = Loc.tag @@ CAppExpl ((None,(Libnames.Ident (Loc.ghost,Id.of_string s)),None),l)
let declare_an_instance n s args =
(((Loc.ghost,Name n),None), Explicit,
- CAppExpl (Loc.ghost, (None, Qualid (Loc.ghost, qualid_of_string s),None),
+ Loc.tag @@ CAppExpl ((None, Qualid (Loc.ghost, qualid_of_string s),None),
args))
let declare_instance a aeq n s = declare_an_instance n s [a;aeq]
let anew_instance global binders instance fields =
new_instance (Flags.is_universe_polymorphism ())
- binders instance (Some (true, CRecord (Loc.ghost,fields)))
+ binders instance (Some (true, Loc.tag @@ CRecord (fields)))
~global ~generalize:false ~refine:false Hints.empty_hint_info
let declare_instance_refl global binders a aeq n lemma =
@@ -1859,7 +1859,7 @@ let declare_relation ?(binders=[]) a aeq n refl symm trans =
(Ident (Loc.ghost,Id.of_string "Equivalence_Symmetric"), lemma2);
(Ident (Loc.ghost,Id.of_string "Equivalence_Transitive"), lemma3)])
-let cHole = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None)
+let cHole = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None)
let proper_projection sigma r ty =
let rel_vect n m = Array.init m (fun i -> mkRel(n+m-i)) in
@@ -2013,13 +2013,13 @@ let add_morphism glob binders m s n =
let instance_id = add_suffix n "_Proper" in
let instance =
(((Loc.ghost,Name instance_id),None), Explicit,
- CAppExpl (Loc.ghost,
+ Loc.tag @@ CAppExpl (
(None, Qualid (Loc.ghost, Libnames.qualid_of_string "Coq.Classes.Morphisms.Proper"),None),
[cHole; s; m]))
in
let tac = Tacinterp.interp (make_tactic "add_morphism_tactic") in
ignore(new_instance ~global:glob poly binders instance
- (Some (true, CRecord (Loc.ghost,[])))
+ (Some (true, Loc.tag @@ CRecord []))
~generalize:false ~tac ~hook:(declare_projection n instance_id) Hints.empty_hint_info)
(** Bind to "rewrite" too *)
diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml
index 3f83f104e..75f890c96 100644
--- a/plugins/ltac/tacintern.ml
+++ b/plugins/ltac/tacintern.ml
@@ -110,13 +110,13 @@ 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 (CRef (r,None))
+ GVar (dloc,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 (CRef (r,None))
+ GVar (dloc,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),
- if strict then None else Some (CRef (r,None))
+ if strict then None else Some (Loc.tag @@ CRef (r,None))
let intern_move_location ist = function
| MoveAfter id -> MoveAfter (intern_hyp ist id)
@@ -273,7 +273,7 @@ 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 (CRef (Ident (dloc,id), None)) with
+ match intern_constr ist (Loc.tag @@ CRef (Ident (dloc,id), None)) with
| GVar (loc,id),_ -> clear,ElimOnIdent (loc,id)
| c -> clear,ElimOnConstr (c,NoBindings)
else
@@ -363,7 +363,7 @@ let intern_typed_pattern_or_ref_with_occurrences ist (l,p) =
Inr (bound_names,(c,None),dummy_pat) in
(l, match p with
| Inl r -> interp_ref r
- | Inr (CAppExpl(_,(None,r,None),[])) ->
+ | Inr (_, CAppExpl((None,r,None),[])) ->
(* We interpret similarly @ref and ref *)
interp_ref (AN r)
| Inr c ->
diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml
index 50f43931e..de6c44b2b 100644
--- a/plugins/ltac/tacinterp.ml
+++ b/plugins/ltac/tacinterp.ml
@@ -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 (CRef (Ident (loc,id),None))) in
+ let c = (GVar (loc,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/ssrmatching/ssrmatching.ml4 b/plugins/ssrmatching/ssrmatching.ml4
index f3555ebc4..1a5ef825d 100644
--- a/plugins/ssrmatching/ssrmatching.ml4
+++ b/plugins/ssrmatching/ssrmatching.ml4
@@ -149,15 +149,15 @@ let add_genarg tag pr =
(** Constructors for cast type *)
let dC t = CastConv t
(** Constructors for constr_expr *)
-let isCVar = function CRef (Ident _, _) -> true | _ -> false
-let destCVar = function CRef (Ident (_, id), _) -> id | _ ->
+let isCVar = function _loc, CRef (Ident _, _) -> true | _ -> false
+let destCVar = function _loc, CRef (Ident (_, id), _) -> id | _ ->
CErrors.anomaly (str"not a CRef")
-let mkCHole loc = CHole (loc, None, IntroAnonymous, None)
-let mkCLambda loc name ty t =
- CLambdaN (loc, [[loc, name], Default Explicit, ty], t)
-let mkCLetIn loc name bo t =
- CLetIn (loc, (loc, name), bo, None, t)
-let mkCCast loc t ty = CCast (loc,t, dC ty)
+let mkCHole loc = Loc.tag ~loc @@ CHole (None, IntroAnonymous, None)
+let mkCLambda loc name ty t = Loc.tag ~loc @@
+ CLambdaN ([[loc, name], Default Explicit, ty], t)
+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)
@@ -952,8 +952,8 @@ let glob_cpattern gs p =
| _, (_, None) as x -> x
| k, (v, Some t) as orig ->
if k = 'x' then glob_ssrterm gs ('(', (v, Some t)) else
- match t with
- | CNotation(_, "( _ in _ )", ([t1; t2], [], [])) ->
+ match snd (Loc.to_pair t) with
+ | CNotation("( _ in _ )", ([t1; t2], [], [])) ->
(try match glob t1, glob t2 with
| (r1, None), (r2, None) -> encode k "In" [r1;r2]
| (r1, Some _), (r2, Some _) when isCVar t1 ->
@@ -961,11 +961,11 @@ let glob_cpattern gs p =
| (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2]
| _ -> CErrors.anomaly (str"where are we?")
with _ when isCVar t1 -> encode k "In" [bind_in t1 t2])
- | CNotation(_, "( _ in _ in _ )", ([t1; t2; t3], [], [])) ->
+ | CNotation("( _ in _ in _ )", ([t1; t2; t3], [], [])) ->
check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3]
- | CNotation(_, "( _ as _ )", ([t1; t2], [], [])) ->
+ | CNotation("( _ as _ )", ([t1; t2], [], [])) ->
encode k "As" [fst (glob t1); fst (glob t2)]
- | CNotation(_, "( _ as _ in _ )", ([t1; t2; t3], [], [])) ->
+ | CNotation("( _ as _ in _ )", ([t1; t2; t3], [], [])) ->
check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3]
| _ -> glob_ssrterm gs orig
;;
@@ -1021,8 +1021,8 @@ type pattern = Evd.evar_map * (constr, constr) ssrpattern
let id_of_cpattern = function
- | _,(_,Some (CRef (Ident (_, x), _))) -> Some x
- | _,(_,Some (CAppExpl (_, (_, Ident (_, x), _), []))) -> Some x
+ | _,(_,Some (_loc, CRef (Ident (_, x), _))) -> Some x
+ | _,(_,Some (_loc, CAppExpl ((_, Ident (_, x), _), []))) -> Some x
| _,(GRef (_, VarRef x, _) ,None) -> Some x
| _ -> None
let id_of_Cterm t = match id_of_cpattern t with
@@ -1378,7 +1378,7 @@ let pf_fill_occ_term gl occ t =
let cpattern_of_id id = ' ', (GRef (dummy_loc, VarRef id, None), None)
let is_wildcard = function
- | _,(_,Some (CHole _)|GHole _,None) -> true
+ | _,(_,Some (_, CHole _)|GHole _,None) -> true
| _ -> false
(* "ssrpattern" *)
diff --git a/printing/ppconstr.ml b/printing/ppconstr.ml
index 966ba6734..a99c0951f 100644
--- a/printing/ppconstr.ml
+++ b/printing/ppconstr.ml
@@ -213,11 +213,11 @@ let tag_var = tag Tag.variable
str "(" ++ pr_id id ++ str ":=" ++ pr ltop a ++ str ")"
let pr_opt_type pr = function
- | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt ()
+ | _loc, CHole (_,Misctypes.IntroAnonymous,_) -> mt ()
| t -> cut () ++ str ":" ++ pr t
let pr_opt_type_spc pr = function
- | CHole (_,_,Misctypes.IntroAnonymous,_) -> mt ()
+ | _loc, CHole (_,Misctypes.IntroAnonymous,_) -> mt ()
| t -> str " :" ++ pr_sep_com (fun()->brk(1,2)) (pr ltop) t
let pr_lident (loc,id) =
@@ -352,7 +352,7 @@ let tag_var = tag Tag.variable
end
| Default b ->
match t with
- | CHole (_,_,Misctypes.IntroAnonymous,_) ->
+ | _loc, CHole (_,Misctypes.IntroAnonymous,_) ->
let s = prlist_with_sep spc pr_lname nal in
hov 1 (surround_implicit b s)
| _ ->
@@ -394,39 +394,40 @@ let tag_var = tag Tag.variable
(* | CLetIn (loc,na,b,c) as x ->
let bl,c = extract_prod_binders c in
if bl = [] then [], x else CLocalDef (na,b) :: bl, c*)
- | CProdN (loc,[],c) ->
+ | _loc, CProdN ([],c) ->
extract_prod_binders c
- | CProdN (loc,[[_,Name id],bk,t],
- CCases (_,LetPatternStyle,None, [CRef (Ident (_,id'),None),None,None],[(_,([_,[p]],b))]))
+ | loc, CProdN ([[_,Name id],bk,t],
+ (_loc', CCases (LetPatternStyle,None, [(_loc'', CRef (Ident (_,id'),None)),None,None],[(_,([_,[p]],b))])))
when Id.equal id id' && not (Id.Set.mem id (Topconstr.free_vars_of_constr_expr b)) ->
let bl,c = extract_prod_binders b in
- CLocalPattern (loc,p,None) :: bl, c
- | CProdN (loc,(nal,bk,t)::bl,c) ->
- let bl,c = extract_prod_binders (CProdN(loc,bl,c)) in
+ CLocalPattern (loc, p,None) :: bl, c
+ | loc, CProdN ((nal,bk,t)::bl,c) ->
+ let bl,c = extract_prod_binders (loc, CProdN(bl,c)) in
CLocalAssum (nal,bk,t) :: bl, c
| c -> [], c
- let rec extract_lam_binders = function
+ let rec extract_lam_binders (loc, ce) = match ce with
(* | CLetIn (loc,na,b,c) as x ->
let bl,c = extract_lam_binders c in
if bl = [] then [], x else CLocalDef (na,b) :: bl, c*)
- | CLambdaN (loc,[],c) ->
+ | CLambdaN ([],c) ->
extract_lam_binders c
- | CLambdaN (loc,[[_,Name id],bk,t],
- CCases (_,LetPatternStyle,None, [CRef (Ident (_,id'),None),None,None],[(_,([_,[p]],b))]))
+ | CLambdaN ([[_,Name id],bk,t],
+ (_loc, CCases (LetPatternStyle,None, [(_loc', CRef (Ident (_,id'),None)),None,None],[(_,([_,[p]],b))])))
when Id.equal id id' && not (Id.Set.mem id (Topconstr.free_vars_of_constr_expr b)) ->
let bl,c = extract_lam_binders b in
CLocalPattern (loc,p,None) :: bl, c
- | CLambdaN (loc,(nal,bk,t)::bl,c) ->
- let bl,c = extract_lam_binders (CLambdaN(loc,bl,c)) in
+ | CLambdaN ((nal,bk,t)::bl,c) ->
+ let bl,c = extract_lam_binders (loc, CLambdaN(bl,c)) in
CLocalAssum (nal,bk,t) :: bl, c
- | c -> [], c
+ | c -> [], Loc.tag ~loc c
- let split_lambda = function
- | CLambdaN (loc,[[na],bk,t],c) -> (na,t,c)
- | CLambdaN (loc,([na],bk,t)::bl,c) -> (na,t,CLambdaN(loc,bl,c))
- | CLambdaN (loc,(na::nal,bk,t)::bl,c) -> (na,t,CLambdaN(loc,(nal,bk,t)::bl,c))
+ let split_lambda = Loc.with_loc (fun ~loc -> function
+ | CLambdaN ([[na],bk,t],c) -> (na,t,c)
+ | CLambdaN (([na],bk,t)::bl,c) -> (na,t,Loc.tag ~loc @@ CLambdaN(bl,c))
+ | CLambdaN ((na::nal,bk,t)::bl,c) -> (na,t,Loc.tag ~loc @@ CLambdaN((nal,bk,t)::bl,c))
| _ -> anomaly (Pp.str "ill-formed fixpoint body")
+ )
let rename na na' t c =
match (na,na') with
@@ -435,12 +436,13 @@ let tag_var = tag Tag.variable
| (_,Name id), (_,Anonymous) -> (na,t,c)
| _ -> (na',t,c)
- let split_product na' = function
- | CProdN (loc,[[na],bk,t],c) -> rename na na' t c
- | CProdN (loc,([na],bk,t)::bl,c) -> rename na na' t (CProdN(loc,bl,c))
- | CProdN (loc,(na::nal,bk,t)::bl,c) ->
- rename na na' t (CProdN(loc,(nal,bk,t)::bl,c))
+ let split_product na' = Loc.with_loc (fun ~loc -> function
+ | CProdN ([[na],bk,t],c) -> rename na na' t c
+ | CProdN (([na],bk,t)::bl,c) -> rename na na' t (Loc.tag ~loc @@ CProdN(bl,c))
+ | CProdN ((na::nal,bk,t)::bl,c) ->
+ rename na na' t (Loc.tag ~loc @@ CProdN((nal,bk,t)::bl,c))
| _ -> anomaly (Pp.str "ill-formed fixpoint body")
+ )
let rec split_fix n typ def =
if Int.equal n 0 then ([],typ,def)
@@ -506,7 +508,7 @@ let tag_var = tag Tag.variable
let pr_case_type pr po =
match po with
- | None | Some (CHole (_,_,Misctypes.IntroAnonymous,_)) -> mt()
+ | None | Some (_, CHole (_,Misctypes.IntroAnonymous,_)) -> mt()
| Some p ->
spc() ++ hov 2 (keyword "return" ++ pr_sep_com spc (pr lsimpleconstr) p)
@@ -543,25 +545,25 @@ let tag_var = tag Tag.variable
let pr_fun_sep = spc () ++ str "=>"
let pr_dangling_with_for sep pr inherited a =
- match a with
- | (CFix (_,_,[_])|CCoFix(_,_,[_])) ->
+ match snd a with
+ | (CFix (_,[_])|CCoFix(_,[_])) ->
pr sep (latom,E) a
| _ ->
pr sep inherited a
let pr pr sep inherited a =
let return (cmds, prec) = (tag_constr_expr a cmds, prec) in
- let (strm, prec) = match a with
+ let (strm, prec) = match snd @@ Loc.to_pair a with
| CRef (r, us) ->
return (pr_cref r us, latom)
- | CFix (_,id,fix) ->
+ | CFix (id,fix) ->
return (
hov 0 (keyword "fix" ++ spc () ++
pr_recursive
(pr_fixdecl (pr mt) (pr_dangling_with_for mt pr)) (snd id) fix),
lfix
)
- | CCoFix (_,id,cofix) ->
+ | CCoFix (id,cofix) ->
return (
hov 0 (keyword "cofix" ++ spc () ++
pr_recursive
@@ -586,7 +588,8 @@ let tag_var = tag Tag.variable
pr_fun_sep ++ pr spc ltop a),
llambda
)
- | CLetIn (_,(_,Name x),(CFix(_,(_,x'),[_])|CCoFix(_,(_,x'),[_]) as fx), t, b)
+ | CLetIn ((_,Name x), ((_loc, CFix((_,x'),[_]))
+ | (_loc, CCoFix((_,x'),[_])) as fx), t, b)
when Id.equal x x' ->
return (
hv 0 (
@@ -596,7 +599,7 @@ let tag_var = tag Tag.variable
pr spc ltop b),
lletin
)
- | CLetIn (_,x,a,t,b) ->
+ | CLetIn (x,a,t,b) ->
return (
hv 0 (
hov 2 (keyword "let" ++ spc () ++ pr_lname x
@@ -606,7 +609,7 @@ let tag_var = tag Tag.variable
pr spc ltop b),
lletin
)
- | CAppExpl (_,(Some i,f,us),l) ->
+ | CAppExpl ((Some i,f,us),l) ->
let l1,l2 = List.chop i l in
let c,l1 = List.sep_last l1 in
let p = pr_proj (pr mt) pr_appexpl c (f,us) l1 in
@@ -614,16 +617,16 @@ let tag_var = tag Tag.variable
return (p ++ prlist (pr spc (lapp,L)) l2, lapp)
else
return (p, lproj)
- | CAppExpl (_,(None,Ident (_,var),us),[t])
- | CApp (_,(_,CRef(Ident(_,var),us)),[t,None])
+ | CAppExpl ((None,Ident (_,var),us),[t])
+ | CApp ((_,(_, CRef(Ident(_,var),us))),[t,None])
when Id.equal var Notation_ops.ldots_var ->
return (
hov 0 (str ".." ++ pr spc (latom,E) t ++ spc () ++ str ".."),
larg
)
- | CAppExpl (_,(None,f,us),l) ->
+ | CAppExpl ((None,f,us),l) ->
return (pr_appexpl (pr mt) (f,us) l, lapp)
- | CApp (_,(Some i,f),l) ->
+ | CApp ((Some i,f),l) ->
let l1,l2 = List.chop i l in
let c,l1 = List.sep_last l1 in
assert (Option.is_empty (snd c));
@@ -635,14 +638,14 @@ let tag_var = tag Tag.variable
)
else
return (p, lproj)
- | CApp (_,(None,a),l) ->
+ | CApp ((None,a),l) ->
return (pr_app (pr mt) a l, lapp)
- | CRecord (_,l) ->
+ | CRecord l ->
return (
hv 0 (str"{|" ++ pr_record_body_gen (pr spc) l ++ str" |}"),
latom
)
- | CCases (_,LetPatternStyle,rtntypopt,[c,as_clause,in_clause],[(_,([(loc,[p])],b))]) ->
+ | CCases (LetPatternStyle,rtntypopt,[c,as_clause,in_clause],[(_,([(loc,[p])],b))]) ->
return (
hv 0 (
keyword "let" ++ spc () ++ str"'" ++
@@ -653,7 +656,7 @@ let tag_var = tag Tag.variable
spc () ++ keyword "in" ++ pr spc ltop b)),
lletpattern
)
- | CCases(_,_,rtntypopt,c,eqns) ->
+ | CCases(_,rtntypopt,c,eqns) ->
return (
v 0
(hv 0 (keyword "match" ++ brk (1,2) ++
@@ -666,7 +669,7 @@ let tag_var = tag Tag.variable
++ keyword "end"),
latom
)
- | CLetTuple (_,nal,(na,po),c,b) ->
+ | CLetTuple (nal,(na,po),c,b) ->
return (
hv 0 (
hov 2 (keyword "let" ++ spc () ++
@@ -679,7 +682,7 @@ let tag_var = tag Tag.variable
pr spc ltop b),
lletin
)
- | CIf (_,c,(na,po),b1,b2) ->
+ | CIf (c,(na,po),b1,b2) ->
(* On force les parenthèses autour d'un "if" sous-terme (même si le
parsing est lui plus tolérant) *)
return (
@@ -693,19 +696,19 @@ let tag_var = tag Tag.variable
lif
)
- | CHole (_,_,Misctypes.IntroIdentifier id,_) ->
+ | CHole (_,Misctypes.IntroIdentifier id,_) ->
return (str "?[" ++ pr_id id ++ str "]", latom)
- | CHole (_,_,Misctypes.IntroFresh id,_) ->
+ | CHole (_,Misctypes.IntroFresh id,_) ->
return (str "?[?" ++ pr_id id ++ str "]", latom)
- | CHole (_,_,_,_) ->
+ | CHole (_,_,_) ->
return (str "_", latom)
- | CEvar (_,n,l) ->
+ | CEvar (n,l) ->
return (pr_evar (pr mt) n l, latom)
- | CPatVar (_,p) ->
+ | CPatVar p ->
return (str "@?" ++ pr_patvar p, latom)
- | CSort (_,s) ->
+ | CSort s ->
return (pr_glob_sort s, latom)
- | CCast (_,a,b) ->
+ | CCast (a,b) ->
return (
hv 0 (pr mt (lcast,L) a ++ spc () ++
match b with
@@ -715,15 +718,15 @@ let tag_var = tag Tag.variable
| CastCoerce -> str ":>"),
lcast
)
- | CNotation (_,"( _ )",([t],[],[])) ->
+ | CNotation ("( _ )",([t],[],[])) ->
return (pr (fun()->str"(") (max_int,L) t ++ str")", latom)
- | CNotation (_,s,env) ->
+ | CNotation (s,env) ->
pr_notation (pr mt) (pr_binders_gen (pr mt ltop)) s env
- | CGeneralization (_,bk,ak,c) ->
+ | CGeneralization (bk,ak,c) ->
return (pr_generalization bk ak (pr mt ltop c), latom)
- | CPrim (_,p) ->
+ | CPrim p ->
return (pr_prim_token p, prec_of_prim_token p)
- | CDelimiters (_,sc,a) ->
+ | CDelimiters (sc,a) ->
return (pr_delimiters sc (pr mt (ldelim,E) a), ldelim)
in
let loc = constr_loc a in
@@ -751,7 +754,7 @@ let tag_var = tag Tag.variable
let pr prec c = pr prec (transf (Global.env()) c)
let pr_simpleconstr = function
- | CAppExpl (_,(None,f,us),[]) -> str "@" ++ pr_cref f us
+ | _lock, CAppExpl ((None,f,us),[]) -> str "@" ++ pr_cref f us
| c -> pr lsimpleconstr c
let default_term_pr = {
diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml
index 2e064454c..27faa7c5c 100644
--- a/printing/ppvernac.ml
+++ b/printing/ppvernac.ml
@@ -254,7 +254,7 @@ open Decl_kinds
prlist_strict (pr_module_vardecls pr_c) l
let pr_type_option pr_c = function
- | CHole (loc, k, Misctypes.IntroAnonymous, _) -> mt()
+ | _loc, CHole (k, Misctypes.IntroAnonymous, _) -> mt()
| _ as c -> brk(0,2) ++ str" :" ++ pr_c c
let pr_decl_notation prc ((loc,ntn),c,scopt) =
@@ -885,7 +885,7 @@ open Decl_kinds
(match bk with Implicit -> str "! " | Explicit -> mt ()) ++
pr_constr cl ++ pr_hint_info pr_constr_pattern_expr info ++
(match props with
- | Some (true,CRecord (_,l)) -> spc () ++ str":=" ++ spc () ++ str"{" ++ pr_record_body l ++ str "}"
+ | Some (true, (_loc, CRecord l)) -> spc () ++ str":=" ++ spc () ++ str"{" ++ pr_record_body l ++ str "}"
| Some (true,_) -> assert false
| Some (false,p) -> spc () ++ str":=" ++ spc () ++ pr_constr p
| None -> mt()))
diff --git a/vernac/classes.ml b/vernac/classes.ml
index 833719965..ffe03bfb7 100644
--- a/vernac/classes.ml
+++ b/vernac/classes.ml
@@ -147,14 +147,14 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
(fun avoid (clname, _) ->
match clname with
| Some (cl, b) ->
- let t = CHole (Loc.ghost, None, Misctypes.IntroAnonymous, None) in
+ let t = Loc.tag @@ CHole (None, Misctypes.IntroAnonymous, None) in
t, avoid
| None -> failwith ("new instance: under-applied typeclass"))
cl
| Explicit -> cl, Id.Set.empty
in
let tclass =
- if generalize then CGeneralization (Loc.ghost, Implicit, Some AbsPi, tclass)
+ if generalize then Loc.tag @@ CGeneralization (Implicit, Some AbsPi, tclass)
else tclass
in
let k, u, cty, ctx', ctx, len, imps, subst =
@@ -217,7 +217,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
else (
let props =
match props with
- | Some (true, CRecord (loc, fs)) ->
+ | Some (true, (_loc, CRecord fs)) ->
if List.length fs > List.length k.cl_props then
mismatched_props env' (List.map snd fs) k.cl_props;
Some (Inl fs)
@@ -261,7 +261,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) p
k.cl_projs;
c :: props, rest'
with Not_found ->
- (CHole (Loc.ghost, None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None) :: props), rest
+ ((Loc.tag @@ CHole (None(* Some Evar_kinds.GoalEvar *), Misctypes.IntroAnonymous, None)) :: props), rest
else props, rest)
([], props) k.cl_props
in
diff --git a/vernac/command.ml b/vernac/command.ml
index 45ff57955..1f1464856 100644
--- a/vernac/command.ml
+++ b/vernac/command.ml
@@ -53,18 +53,19 @@ let rec under_binders env sigma f n c =
mkLetIn (x,b,t,under_binders (push_rel (LocalDef (x,b,t)) env) sigma f (n-1) c)
| _ -> assert false
-let rec complete_conclusion a cs = function
- | CProdN (loc,bl,c) -> CProdN (loc,bl,complete_conclusion a cs c)
- | CLetIn (loc,na,b,t,c) -> CLetIn (loc,na,b,t,complete_conclusion a cs c)
- | CHole (loc, k, _, _) ->
+let rec complete_conclusion a cs = Loc.map_with_loc (fun ~loc -> function
+ | CProdN (bl,c) -> CProdN (bl,complete_conclusion a cs c)
+ | CLetIn (na,b,t,c) -> CLetIn (na,b,t,complete_conclusion a cs c)
+ | CHole (k, _, _) ->
let (has_no_args,name,params) = a in
if not has_no_args then
- user_err ~loc
+ user_err ~loc
(strbrk"Cannot infer the non constant arguments of the conclusion of "
++ pr_id cs ++ str ".");
- let args = List.map (fun id -> CRef(Ident(loc,id),None)) params in
- CAppExpl (loc,(None,Ident(loc,name),None),List.rev args)
+ let args = List.map (fun id -> Loc.tag ~loc @@ CRef(Ident(loc,id),None)) params in
+ CAppExpl ((None,Ident(loc,name),None),List.rev args)
| c -> c
+ )
(* Commands of the interface *)
@@ -682,7 +683,7 @@ let extract_params indl =
let extract_inductive indl =
List.map (fun (((_,indname),pl),_,ar,lc) -> {
ind_name = indname; ind_univs = pl;
- ind_arity = Option.cata (fun x -> x) (CSort (Loc.ghost,GType [])) ar;
+ ind_arity = Option.cata (fun x -> x) (Loc.tag @@ CSort (GType [])) ar;
ind_lc = List.map (fun (_,((_,id),t)) -> (id,t)) lc
}) indl
@@ -1354,7 +1355,7 @@ let do_program_fixpoint local poly l =
| [(n, CMeasureRec (m, r))], [((((_,id),pl),_,bl,typ,def),ntn)] ->
build_wellfounded (id, pl, n, bl, typ, out_def def) poly
- (Option.default (CRef (lt_ref,None)) r) m ntn
+ (Option.default (Loc.tag @@ CRef (lt_ref,None)) r) m ntn
| _, _ when List.for_all (fun (n, ro) -> ro == CStructRec) g ->
let fixl,ntns = extract_fixpoint_components true l in
diff --git a/vernac/metasyntax.ml b/vernac/metasyntax.ml
index f805eeaa9..98b2c3729 100644
--- a/vernac/metasyntax.ml
+++ b/vernac/metasyntax.ml
@@ -1416,7 +1416,7 @@ let add_notation_extra_printing_rule df k v =
(* Infix notations *)
-let inject_var x = CRef (Ident (Loc.ghost, Id.of_string x),None)
+let inject_var x = Loc.tag @@ CRef (Ident (Loc.ghost, Id.of_string x),None)
let add_infix local ((loc,inf),modifiers) pr sc =
check_infix_modifiers modifiers;
@@ -1477,7 +1477,7 @@ let add_class_scope scope cl =
(* Check if abbreviation to a name and avoid early insertion of
maximal implicit arguments *)
let try_interp_name_alias = function
- | [], CRef (ref,_) -> intern_reference ref
+ | [], (_loc, CRef (ref,_)) -> intern_reference ref
| _ -> raise Not_found
let add_syntactic_definition ident (vars,c) local onlyparse =
diff --git a/vernac/record.ml b/vernac/record.ml
index 8b4b7606f..37ce231f9 100644
--- a/vernac/record.ml
+++ b/vernac/record.ml
@@ -93,7 +93,8 @@ let compute_constructor_level evars env l =
let binder_of_decl = function
| Vernacexpr.AssumExpr(n,t) -> (n,None,t)
- | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c | None -> CHole (fst n, None, Misctypes.IntroAnonymous, None))
+ | Vernacexpr.DefExpr(n,c,t) -> (n,Some c, match t with Some c -> c
+ | None -> Loc.tag ~loc:(fst n) @@ CHole (None, Misctypes.IntroAnonymous, None))
let binders_of_decls = List.map binder_of_decl
@@ -120,8 +121,8 @@ let typecheck_params_and_fields def id pl t ps nots fs =
| Some t ->
let env = push_rel_context newps env0 in
let poly =
- match t with
- | CSort (_, Misctypes.GType []) -> true | _ -> false in
+ match snd t with
+ | CSort (Misctypes.GType []) -> true | _ -> false in
let s = interp_type_evars env evars ~impls:empty_internalization_env t in
let sred = Reductionops.whd_all env !evars s in
let s = EConstr.Unsafe.to_constr s in